Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
authorhoelzl
Mon, 21 Jun 2010 19:33:51 +0200
changeset 37489 44e42d392c6e
parent 37482 6849464ab10e
child 37490 9de1add14bac
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
src/HOL/IsaMakefile
src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy
src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Derivative.thy
src/HOL/Multivariate_Analysis/Determinants.thy
src/HOL/Multivariate_Analysis/Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Fashoda.thy
src/HOL/Multivariate_Analysis/Integration.certs
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy
src/HOL/Multivariate_Analysis/Operator_Norm.thy
src/HOL/Multivariate_Analysis/Path_Connected.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Vec1.thy
--- a/src/HOL/IsaMakefile	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/IsaMakefile	Mon Jun 21 19:33:51 2010 +0200
@@ -1108,8 +1108,8 @@
   Multivariate_Analysis/Real_Integration.thy				\
   Multivariate_Analysis/Topology_Euclidean_Space.thy			\
   Multivariate_Analysis/document/root.tex				\
-  Multivariate_Analysis/normarith.ML Multivariate_Analysis/Vec1.thy	\
-  Library/Glbs.thy Library/Inner_Product.thy Library/Numeral_Type.thy	\
+  Multivariate_Analysis/normarith.ML Library/Glbs.thy	\
+  Library/Inner_Product.thy Library/Numeral_Type.thy	\
   Library/Convex.thy Library/FrechetDeriv.thy				\
   Library/Product_Vector.thy Library/Product_plus.thy
 	@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -23,7 +23,7 @@
 begin
 
 lemma brouwer_compactness_lemma:
-  assumes "compact s" "continuous_on s f" "\<not> (\<exists>x\<in>s. (f x = (0::real^'n)))"
+  assumes "compact s" "continuous_on s f" "\<not> (\<exists>x\<in>s. (f x = (0::_::euclidean_space)))"
   obtains d where "0 < d" "\<forall>x\<in>s. d \<le> norm(f x)" proof(cases "s={}") case False
   have "continuous_on s (norm \<circ> f)" by(rule continuous_on_intros continuous_on_norm assms(2))+
   then obtain x where x:"x\<in>s" "\<forall>y\<in>s. (norm \<circ> f) x \<le> (norm \<circ> f) y"
@@ -31,21 +31,21 @@
   have "(norm \<circ> f) x > 0" using assms(3) and x(1) by auto
   thus ?thesis apply(rule that) using x(2) unfolding o_def by auto qed(rule that[of 1], auto)
 
-lemma kuhn_labelling_lemma:
-  assumes "(\<forall>x::real^_. P x \<longrightarrow> P (f x))"  "\<forall>x. P x \<longrightarrow> (\<forall>i. Q i \<longrightarrow> 0 \<le> x$i \<and> x$i \<le> 1)"
-  shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>
-             (\<forall>x i. P x \<and> Q i \<and> (x$i = 0) \<longrightarrow> (l x i = 0)) \<and>
-             (\<forall>x i. P x \<and> Q i \<and> (x$i = 1) \<longrightarrow> (l x i = 1)) \<and>
-             (\<forall>x i. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x$i \<le> f(x)$i) \<and>
-             (\<forall>x i. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)$i \<le> x$i)" proof-
+lemma kuhn_labelling_lemma: fixes type::"'a::euclidean_space"
+  assumes "(\<forall>x::'a. P x \<longrightarrow> P (f x))"  "\<forall>x. P x \<longrightarrow> (\<forall>i<DIM('a). Q i \<longrightarrow> 0 \<le> x$$i \<and> x$$i \<le> 1)"
+  shows "\<exists>l. (\<forall>x.\<forall> i<DIM('a). l x i \<le> (1::nat)) \<and>
+             (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (x$$i = 0) \<longrightarrow> (l x i = 0)) \<and>
+             (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (x$$i = 1) \<longrightarrow> (l x i = 1)) \<and>
+             (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x$$i \<le> f(x)$$i) \<and>
+             (\<forall>x.\<forall> i<DIM('a). P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)$$i \<le> x$$i)" proof-
   have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" by auto
   have *:"\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> (x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x)" by auto
   show ?thesis unfolding and_forall_thm apply(subst choice_iff[THEN sym])+ proof(rule,rule) case goal1
-    let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x $ xa = 0 \<longrightarrow> y = (0::nat)) \<and>
-        (P x \<and> Q xa \<and> x $ xa = 1 \<longrightarrow> y = 1) \<and> (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x $ xa \<le> f x $ xa) \<and> (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x $ xa \<le> x $ xa)"
-    { assume "P x" "Q xa" hence "0 \<le> f x $ xa \<and> f x $ xa \<le> 1" using assms(2)[rule_format,of "f x" xa]
+    let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x $$ xa = 0 \<longrightarrow> y = (0::nat)) \<and>
+        (P x \<and> Q xa \<and> x $$ xa = 1 \<longrightarrow> y = 1) \<and> (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x $$ xa \<le> f x $$ xa) \<and> (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x $$ xa \<le> x $$ xa)"
+    { assume "P x" "Q xa" "xa<DIM('a)" hence "0 \<le> f x $$ xa \<and> f x $$ xa \<le> 1" using assms(2)[rule_format,of "f x" xa]
         apply(drule_tac assms(1)[rule_format]) by auto }
-    hence "?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed 
+    hence "xa<DIM('a) \<Longrightarrow> ?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed 
  
 subsection {* The key "counting" observation, somewhat abstracted. *}
 
@@ -1159,49 +1159,50 @@
         apply(drule_tac assms(1)[rule_format]) by auto }
     hence "?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed 
 
-lemma brouwer_cube: fixes f::"real^'n \<Rightarrow> real^'n"
-  assumes "continuous_on {0..1} f" "f ` {0..1} \<subseteq> {0..1}"
-  shows "\<exists>x\<in>{0..1}. f x = x" apply(rule ccontr) proof-
-  def n \<equiv> "CARD('n)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by auto
-  assume "\<not> (\<exists>x\<in>{0..1}. f x = x)" hence *:"\<not> (\<exists>x\<in>{0..1}. f x - x = 0)" by auto
+lemma brouwer_cube: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
+  assumes "continuous_on {0..\<chi>\<chi> i. 1} f" "f ` {0..\<chi>\<chi> i. 1} \<subseteq> {0..\<chi>\<chi> i. 1}"
+  shows "\<exists>x\<in>{0..\<chi>\<chi> i. 1}. f x = x" apply(rule ccontr) proof-
+  def n \<equiv> "DIM('a)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by(auto simp add:Suc_le_eq)
+  assume "\<not> (\<exists>x\<in>{0..\<chi>\<chi> i. 1}. f x = x)" hence *:"\<not> (\<exists>x\<in>{0..\<chi>\<chi> i. 1}. f x - x = 0)" by auto
   guess d apply(rule brouwer_compactness_lemma[OF compact_interval _ *]) 
     apply(rule continuous_on_intros assms)+ . note d=this[rule_format]
-  have *:"\<forall>x. x \<in> {0..1} \<longrightarrow> f x \<in> {0..1}"  "\<forall>x. x \<in> {0..1::real^'n} \<longrightarrow> (\<forall>i. True \<longrightarrow> 0 \<le> x $ i \<and> x $ i \<le> 1)"
+  have *:"\<forall>x. x \<in> {0..\<chi>\<chi> i. 1} \<longrightarrow> f x \<in> {0..\<chi>\<chi> i. 1}"  "\<forall>x. x \<in> {0..(\<chi>\<chi> i. 1)::'a} \<longrightarrow>
+    (\<forall>i<DIM('a). True \<longrightarrow> 0 \<le> x $$ i \<and> x $$ i \<le> 1)"
     using assms(2)[unfolded image_subset_iff Ball_def] unfolding mem_interval by auto
   guess label using kuhn_labelling_lemma[OF *] apply-apply(erule exE,(erule conjE)+) . note label = this[rule_format]
-  have lem1:"\<forall>x\<in>{0..1}.\<forall>y\<in>{0..1}.\<forall>i. label x i \<noteq> label y i
-            \<longrightarrow> abs(f x $ i - x $ i) \<le> norm(f y - f x) + norm(y - x)" proof(rule,rule,rule,rule)
-    fix x y assume xy:"x\<in>{0..1::real^'n}" "y\<in>{0..1::real^'n}" fix i::'n assume i:"label x i \<noteq> label y i"
+  have lem1:"\<forall>x\<in>{0..\<chi>\<chi> i. 1}.\<forall>y\<in>{0..\<chi>\<chi> i. 1}.\<forall>i<DIM('a). label x i \<noteq> label y i
+            \<longrightarrow> abs(f x $$ i - x $$ i) \<le> norm(f y - f x) + norm(y - x)" proof safe
+    fix x y::'a assume xy:"x\<in>{0..\<chi>\<chi> i. 1}" "y\<in>{0..\<chi>\<chi> i. 1}" fix i assume i:"label x i \<noteq> label y i" "i<DIM('a)"
     have *:"\<And>x y fx fy::real. (x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy)
              \<Longrightarrow> abs(fx - x) \<le> abs(fy - fx) + abs(y - x)" by auto
-    have "\<bar>(f x - x) $ i\<bar> \<le> abs((f y - f x)$i) + abs((y - x)$i)" unfolding vector_minus_component
+    have "\<bar>(f x - x) $$ i\<bar> \<le> abs((f y - f x)$$i) + abs((y - x)$$i)" unfolding euclidean_simps
       apply(rule *) apply(cases "label x i = 0") apply(rule disjI1,rule) prefer 3 proof(rule disjI2,rule)
-      assume lx:"label x i = 0" hence ly:"label y i = 1" using i label(1)[of y i] by auto
-      show "x $ i \<le> f x $ i" apply(rule label(4)[rule_format]) using xy lx by auto
-      show "f y $ i \<le> y $ i" apply(rule label(5)[rule_format]) using xy ly by auto next
+      assume lx:"label x i = 0" hence ly:"label y i = 1" using i label(1)[of i y] by auto
+      show "x $$ i \<le> f x $$ i" apply(rule label(4)[rule_format]) using xy lx i(2) by auto
+      show "f y $$ i \<le> y $$ i" apply(rule label(5)[rule_format]) using xy ly i(2) by auto next
       assume "label x i \<noteq> 0" hence l:"label x i = 1" "label y i = 0"
-        using i label(1)[of x i] label(1)[of y i] by auto
-      show "f x $ i \<le> x $ i" apply(rule label(5)[rule_format]) using xy l  by auto
-      show "y $ i \<le> f y $ i" apply(rule label(4)[rule_format]) using xy l  by auto qed 
+        using i label(1)[of i x] label(1)[of i y] by auto
+      show "f x $$ i \<le> x $$ i" apply(rule label(5)[rule_format]) using xy l i(2) by auto
+      show "y $$ i \<le> f y $$ i" apply(rule label(4)[rule_format]) using xy l i(2) by auto qed 
     also have "\<dots> \<le> norm (f y - f x) + norm (y - x)" apply(rule add_mono) by(rule component_le_norm)+
-    finally show "\<bar>f x $ i - x $ i\<bar> \<le> norm (f y - f x) + norm (y - x)" unfolding vector_minus_component . qed
-  have "\<exists>e>0. \<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. \<forall>z\<in>{0..1}. \<forall>i.
-    norm(x - z) < e \<and> norm(y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> abs((f(z) - z)$i) < d / (real n)" proof-
+    finally show "\<bar>f x $$ i - x $$ i\<bar> \<le> norm (f y - f x) + norm (y - x)" unfolding euclidean_simps . qed
+  have "\<exists>e>0. \<forall>x\<in>{0..\<chi>\<chi> i. 1}. \<forall>y\<in>{0..\<chi>\<chi> i. 1}. \<forall>z\<in>{0..\<chi>\<chi> i. 1}. \<forall>i<DIM('a).
+    norm(x - z) < e \<and> norm(y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow> abs((f(z) - z)$$i) < d / (real n)" proof-
     have d':"d / real n / 8 > 0" apply(rule divide_pos_pos)+ using d(1) unfolding n_def by auto
-    have *:"uniformly_continuous_on {0..1} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval])
+    have *:"uniformly_continuous_on {0..\<chi>\<chi> i. 1} f" by(rule compact_uniformly_continuous[OF assms(1) compact_interval])
     guess e using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] apply-apply(erule exE,(erule conjE)+) .
     note e=this[rule_format,unfolded dist_norm]
-    show ?thesis apply(rule_tac x="min (e/2) (d/real n/8)" in exI) apply(rule) defer
-      apply(rule,rule,rule,rule,rule) apply(erule conjE)+ proof-
-      show "0 < min (e / 2) (d / real n / 8)" using d' e by auto
-      fix x y z i assume as:"x \<in> {0..1}" "y \<in> {0..1}" "z \<in> {0..1}" "norm (x - z) < min (e / 2) (d / real n / 8)"
-        "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \<noteq> label y i"
+    show ?thesis apply(rule_tac x="min (e/2) (d/real n/8)" in exI)
+    proof safe show "0 < min (e / 2) (d / real n / 8)" using d' e by auto
+      fix x y z i assume as:"x \<in> {0..\<chi>\<chi> i. 1}" "y \<in> {0..\<chi>\<chi> i. 1}" "z \<in> {0..\<chi>\<chi> i. 1}"
+        "norm (x - z) < min (e / 2) (d / real n / 8)"
+        "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \<noteq> label y i" and i:"i<DIM('a)"
       have *:"\<And>z fz x fx n1 n2 n3 n4 d4 d::real. abs(fx - x) \<le> n1 + n2 \<Longrightarrow> abs(fx - fz) \<le> n3 \<Longrightarrow> abs(x - z) \<le> n4 \<Longrightarrow>
         n1 < d4 \<Longrightarrow> n2 < 2 * d4 \<Longrightarrow> n3 < d4 \<Longrightarrow> n4 < d4 \<Longrightarrow> (8 * d4 = d) \<Longrightarrow> abs(fz - z) < d" by auto
-      show "\<bar>(f z - z) $ i\<bar> < d / real n" unfolding vector_minus_component proof(rule *)
-        show "\<bar>f x $ i - x $ i\<bar> \<le> norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as by auto
-        show "\<bar>f x $ i - f z $ i\<bar> \<le> norm (f x - f z)" "\<bar>x $ i - z $ i\<bar> \<le> norm (x - z)"
-          unfolding vector_minus_component[THEN sym] by(rule component_le_norm)+
+      show "\<bar>(f z - z) $$ i\<bar> < d / real n" unfolding euclidean_simps proof(rule *)
+        show "\<bar>f x $$ i - x $$ i\<bar> \<le> norm (f y -f x) + norm (y - x)" apply(rule lem1[rule_format]) using as i  by auto
+        show "\<bar>f x $$ i - f z $$ i\<bar> \<le> norm (f x - f z)" "\<bar>x $$ i - z $$ i\<bar> \<le> norm (x - z)"
+          unfolding euclidean_component.diff[THEN sym] by(rule component_le_norm)+
         have tria:"norm (y - x) \<le> norm (y - z) + norm (x - z)" using dist_triangle[of y x z,unfolded dist_norm]
           unfolding norm_minus_commute by auto
         also have "\<dots> < e / 2 + e / 2" apply(rule add_strict_mono) using as(4,5) by auto
@@ -1213,83 +1214,99 @@
   guess p using real_arch_simple[of "1 + real n / e"] .. note p=this
   have "1 + real n / e > 0" apply(rule add_pos_pos) defer apply(rule divide_pos_pos) using e(1) n by auto
   hence "p > 0" using p by auto
-  guess b using ex_bij_betw_nat_finite_1[OF finite_UNIV[where 'a='n]] .. note b=this
+  def b \<equiv> "\<lambda>i. i - 1::nat" have b:"bij_betw b {1..n} {..<DIM('a)}"
+    unfolding bij_betw_def inj_on_def b_def n_def apply rule defer
+    apply safe defer unfolding image_iff apply(rule_tac x="Suc x" in bexI) by auto
   def b' \<equiv> "inv_into {1..n} b"
-  have b':"bij_betw b' UNIV {1..n}" using bij_betw_inv_into[OF b] unfolding b'_def n_def by auto
-  have bb'[simp]:"\<And>i. b (b' i) = i" unfolding b'_def apply(rule f_inv_into_f) unfolding n_def using b  
+  have b':"bij_betw b' {..<DIM('a)} {1..n}" using bij_betw_inv_into[OF b] unfolding b'_def n_def by auto
+  have bb'[simp]:"\<And>i. i<DIM('a) \<Longrightarrow> b (b' i) = i" unfolding b'_def apply(rule f_inv_into_f) using b  
     unfolding bij_betw_def by auto
   have b'b[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> b' (b i) = i" unfolding b'_def apply(rule inv_into_f_eq)
     using b unfolding n_def bij_betw_def by auto
   have *:"\<And>x::nat. x=0 \<or> x=1 \<longleftrightarrow> x\<le>1" by auto
+  have b'':"\<And>j. j\<in>{1..n} \<Longrightarrow> b j <DIM('a)" using b unfolding bij_betw_def by auto
   have q1:"0 < p" "0 < n"  "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow>
-    (\<forall>i\<in>{1..n}. (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0 \<or> (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
-    unfolding * using `p>0` `n>0` using label(1) by auto
-  have q2:"\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = 0 \<longrightarrow> (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0)"
-    "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow> (label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
+    (\<forall>i\<in>{1..n}. (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0 \<or> (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
+    unfolding * using `p>0` `n>0` using label(1)[OF b'']  by auto
+  have q2:"\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = 0 \<longrightarrow> (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0)"
+    "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. x i = p \<longrightarrow> (label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1)"
     apply(rule,rule,rule,rule) defer proof(rule,rule,rule,rule) fix x i 
     assume as:"\<forall>i\<in>{1..n}. x i \<le> p" "i \<in> {1..n}"
     { assume "x i = p \<or> x i = 0" 
-      have "(\<chi> i. real (x (b' i)) / real p) \<in> {0..1}" unfolding mem_interval Cart_lambda_beta proof(rule,rule)
-        fix j::'n have j:"b' j \<in> {1..n}" using b' unfolding n_def bij_betw_def by auto
-        show "0 $ j \<le> real (x (b' j)) / real p" unfolding zero_index
+      have "(\<chi>\<chi> i. real (x (b' i)) / real p) \<in> {0::'a..\<chi>\<chi> i. 1}" unfolding mem_interval 
+        apply safe unfolding euclidean_lambda_beta euclidean_component.zero
+      proof (simp_all only: if_P) fix j assume j':"j<DIM('a)"
+        hence j:"b' j \<in> {1..n}" using b' unfolding n_def bij_betw_def by auto
+        show "0 \<le> real (x (b' j)) / real p"
           apply(rule divide_nonneg_pos) using `p>0` using as(1)[rule_format,OF j] by auto
-        show "real (x (b' j)) / real p \<le> 1 $ j" unfolding one_index divide_le_eq_1
+        show "real (x (b' j)) / real p \<le> 1" unfolding divide_le_eq_1
           using as(1)[rule_format,OF j] by auto qed } note cube=this
-    { assume "x i = p" thus "(label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1" unfolding o_def
-        apply-apply(rule label(3)) using cube using as `p>0` by auto }
-    { assume "x i = 0" thus "(label (\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0" unfolding o_def
-        apply-apply(rule label(2)) using cube using as `p>0` by auto } qed
+    { assume "x i = p" thus "(label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 1" unfolding o_def
+        apply- apply(rule label(3)) apply(rule b'') using cube using as `p>0`
+      proof safe assume i:"i\<in>{1..n}"
+        show "((\<chi>\<chi> ia. real (x (b' ia)) / real (x i))::'a) $$ b i = 1"
+          unfolding euclidean_lambda_beta apply(subst if_P) apply(rule b''[OF i]) unfolding b'b[OF i] 
+          unfolding  `x i = p` using q1(1) by auto
+      qed auto }
+    { assume "x i = 0" thus "(label (\<chi>\<chi> i. real (x (b' i)) / real p) \<circ> b) i = 0" unfolding o_def
+        apply-apply(rule label(2)[OF b'']) using cube using as `p>0`
+      proof safe assume i:"i\<in>{1..n}"
+        show "((\<chi>\<chi> ia. real (x (b' ia)) / real p)::'a) $$ b i = 0"
+          unfolding euclidean_lambda_beta apply (subst if_P) apply(rule b''[OF i]) unfolding b'b[OF i] 
+          unfolding `x i = 0` using q1(1) by auto 
+      qed auto }
+  qed
   guess q apply(rule kuhn_lemma[OF q1 q2]) . note q=this
-  def z \<equiv> "\<chi> i. real (q (b' i)) / real p"
-  have "\<exists>i. d / real n \<le> abs((f z - z)$i)" proof(rule ccontr)
-    have "\<forall>i. q (b' i) \<in> {0..<p}" using q(1) b'[unfolded bij_betw_def] by auto 
-    hence "\<forall>i. q (b' i) \<in> {0..p}" apply-apply(rule,erule_tac x=i in allE) by auto
-    hence "z\<in>{0..1}" unfolding z_def mem_interval unfolding one_index zero_index Cart_lambda_beta
-      apply-apply(rule,rule) apply(rule divide_nonneg_pos) using `p>0` unfolding divide_le_eq_1 by auto
+  def z \<equiv> "(\<chi>\<chi> i. real (q (b' i)) / real p)::'a"
+  have "\<exists>i<DIM('a). d / real n \<le> abs((f z - z)$$i)" proof(rule ccontr)
+    have "\<forall>i<DIM('a). q (b' i) \<in> {0..<p}" using q(1) b'[unfolded bij_betw_def] by auto 
+    hence "\<forall>i<DIM('a). q (b' i) \<in> {0..p}" apply-apply(rule,erule_tac x=i in allE) by auto
+    hence "z\<in>{0..\<chi>\<chi> i.1}" unfolding z_def mem_interval apply safe unfolding euclidean_lambda_beta
+      unfolding euclidean_component.zero apply (simp_all only: if_P)
+      apply(rule divide_nonneg_pos) using `p>0` unfolding divide_le_eq_1 by auto
     hence d_fz_z:"d \<le> norm (f z - z)" apply(drule_tac d) .
-    case goal1 hence as:"\<forall>i. \<bar>f z $ i - z $ i\<bar> < d / real n" using `n>0` by(auto simp add:not_le)
-    have "norm (f z - z) \<le> (\<Sum>i\<in>UNIV. \<bar>f z $ i - z $ i\<bar>)" unfolding vector_minus_component[THEN sym] by(rule norm_le_l1)
-    also have "\<dots> < (\<Sum>(i::'n)\<in>UNIV. d / real n)" apply(rule setsum_strict_mono) using as by auto
-    also have "\<dots> = d" unfolding real_eq_of_nat n_def using n by auto
+    case goal1 hence as:"\<forall>i<DIM('a). \<bar>f z $$ i - z $$ i\<bar> < d / real n" using `n>0` by(auto simp add:not_le)
+    have "norm (f z - z) \<le> (\<Sum>i<DIM('a). \<bar>f z $$ i - z $$ i\<bar>)" unfolding euclidean_component.diff[THEN sym] by(rule norm_le_l1)
+    also have "\<dots> < (\<Sum>i<DIM('a). d / real n)" apply(rule setsum_strict_mono) using as by auto
+    also have "\<dots> = d" unfolding real_eq_of_nat n_def using n using DIM_positive[where 'a='a] by auto
     finally show False using d_fz_z by auto qed then guess i .. note i=this
-  have *:"b' i \<in> {1..n}" using b'[unfolded bij_betw_def] by auto
+  have *:"b' i \<in> {1..n}" using i using b'[unfolded bij_betw_def] by auto
   guess r using q(2)[rule_format,OF *] .. then guess s apply-apply(erule exE,(erule conjE)+) . note rs=this[rule_format]
-  have b'_im:"\<And>i. b' i \<in> {1..n}" using b' unfolding bij_betw_def by auto
-  def r' \<equiv> "\<chi> i. real (r (b' i)) / real p"
-  have "\<And>i. r (b' i) \<le> p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2])
+  have b'_im:"\<And>i. i<DIM('a) \<Longrightarrow>  b' i \<in> {1..n}" using b' unfolding bij_betw_def by auto
+  def r' \<equiv> "(\<chi>\<chi> i. real (r (b' i)) / real p)::'a"
+  have "\<And>i. i<DIM('a) \<Longrightarrow> r (b' i) \<le> p" apply(rule order_trans) apply(rule rs(1)[OF b'_im,THEN conjunct2])
     using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq)
-  hence "r' \<in> {0..1::real^'n}" unfolding r'_def mem_interval Cart_lambda_beta one_index zero_index
-    apply-apply(rule,rule,rule divide_nonneg_pos)
-    using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto
-  def s' \<equiv> "\<chi> i. real (s (b' i)) / real p"
-  have "\<And>i. s (b' i) \<le> p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2])
+  hence "r' \<in> {0..\<chi>\<chi> i. 1}"  unfolding r'_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component.zero
+    apply (simp only: if_P)
+    apply(rule divide_nonneg_pos) using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto
+  def s' \<equiv> "(\<chi>\<chi> i. real (s (b' i)) / real p)::'a"
+  have "\<And>i. i<DIM('a) \<Longrightarrow> s (b' i) \<le> p" apply(rule order_trans) apply(rule rs(2)[OF b'_im,THEN conjunct2])
     using q(1)[rule_format,OF b'_im] by(auto simp add: Suc_le_eq)
-  hence "s' \<in> {0..1::real^'n}" unfolding s'_def mem_interval Cart_lambda_beta one_index zero_index
-    apply-apply(rule,rule,rule divide_nonneg_pos)
-    using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto
-  have "z\<in>{0..1}" unfolding z_def mem_interval Cart_lambda_beta one_index zero_index 
-    apply(rule,rule,rule divide_nonneg_pos) using q(1)[rule_format,OF b'_im] `p>0` by(auto intro:less_imp_le)
+  hence "s' \<in> {0..\<chi>\<chi> i.1}" unfolding s'_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component.zero
+    apply (simp_all only: if_P) apply(rule divide_nonneg_pos) using rs(1)[OF b'_im] q(1)[rule_format,OF b'_im] `p>0` by auto
+  have "z\<in>{0..\<chi>\<chi> i.1}" unfolding z_def mem_interval apply safe unfolding euclidean_lambda_beta euclidean_component.zero
+    apply (simp_all only: if_P) apply(rule divide_nonneg_pos) using q(1)[rule_format,OF b'_im] `p>0` by(auto intro:less_imp_le)
   have *:"\<And>x. 1 + real x = real (Suc x)" by auto
-  { have "(\<Sum>i\<in>UNIV. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'n)\<in>UNIV. 1)" 
+  { have "(\<Sum>i<DIM('a). \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>i<DIM('a). 1)" 
       apply(rule setsum_mono) using rs(1)[OF b'_im] by(auto simp add:* field_simps)
     also have "\<dots> < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def
       by(auto simp add:field_simps)
-    finally have "(\<Sum>i\<in>UNIV. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . } moreover
-  { have "(\<Sum>i\<in>UNIV. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'n)\<in>UNIV. 1)" 
+    finally have "(\<Sum>i<DIM('a). \<bar>real (r (b' i)) - real (q (b' i))\<bar>) < e * real p" . } moreover
+  { have "(\<Sum>i<DIM('a). \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>i<DIM('a). 1)" 
       apply(rule setsum_mono) using rs(2)[OF b'_im] by(auto simp add:* field_simps)
     also have "\<dots> < e * real p" using p `e>0` `p>0` unfolding n_def real_of_nat_def
       by(auto simp add:field_simps)
-    finally have "(\<Sum>i\<in>UNIV. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . } ultimately
+    finally have "(\<Sum>i<DIM('a). \<bar>real (s (b' i)) - real (q (b' i))\<bar>) < e * real p" . } ultimately
   have "norm (r' - z) < e" "norm (s' - z) < e" unfolding r'_def s'_def z_def apply-
     apply(rule_tac[!] le_less_trans[OF norm_le_l1]) using `p>0`
     by(auto simp add:field_simps setsum_divide_distrib[THEN sym])
-  hence "\<bar>(f z - z) $ i\<bar> < d / real n" apply-apply(rule e(2)[OF `r'\<in>{0..1}` `s'\<in>{0..1}` `z\<in>{0..1}`])
-    using rs(3) unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' by auto
+  hence "\<bar>(f z - z) $$ i\<bar> < d / real n" apply-apply(rule e(2)[OF `r'\<in>{0..\<chi>\<chi> i.1}` `s'\<in>{0..\<chi>\<chi> i.1}` `z\<in>{0..\<chi>\<chi> i.1}`])
+    using rs(3) unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' using i by auto
   thus False using i by auto qed
 
 subsection {* Retractions. *}
 
-definition "retraction s t (r::real^'n\<Rightarrow>real^'n) \<longleftrightarrow>
+definition "retraction s t r \<longleftrightarrow>
   t \<subseteq> s \<and> continuous_on s r \<and> (r ` s \<subseteq> t) \<and> (\<forall>x\<in>t. r x = x)"
 
 definition retract_of (infixl "retract'_of" 12) where
@@ -1300,7 +1317,7 @@
 
 subsection {*preservation of fixpoints under (more general notion of) retraction. *}
 
-lemma invertible_fixpoint_property: fixes s::"(real^'n) set" and t::"(real^'m) set" 
+lemma invertible_fixpoint_property: fixes s::"('a::euclidean_space) set" and t::"('b::euclidean_space) set" 
   assumes "continuous_on t i" "i ` t \<subseteq> s" "continuous_on s r" "r ` s \<subseteq> t" "\<forall>y\<in>t. r (i y) = y"
   "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)" "continuous_on t g" "g ` t \<subseteq> t"
   obtains y where "y\<in>t" "g y = y" proof-
@@ -1313,7 +1330,7 @@
       unfolding assms(5)[rule_format,OF *] using assms(4) by auto qed
 
 lemma homeomorphic_fixpoint_property:
-  fixes s::"(real^'n) set" and t::"(real^'m) set" assumes "s homeomorphic t"
+  fixes s::"('a::euclidean_space) set" and t::"('b::euclidean_space) set" assumes "s homeomorphic t"
   shows "(\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)) \<longleftrightarrow>
          (\<forall>g. continuous_on t g \<and> g ` t \<subseteq> t \<longrightarrow> (\<exists>y\<in>t. g y = y))" proof-
   guess r using assms[unfolded homeomorphic_def homeomorphism_def] .. then guess i ..
@@ -1321,7 +1338,7 @@
     apply(rule_tac g=g in invertible_fixpoint_property[of t i s r]) prefer 10
     apply(rule_tac g=f in invertible_fixpoint_property[of s r t i]) by auto qed
 
-lemma retract_fixpoint_property:
+lemma retract_fixpoint_property: fixes f::"'a::euclidean_space => 'b::euclidean_space" and s::"'a set"
   assumes "t retract_of s"  "\<forall>f. continuous_on s f \<and> f ` s \<subseteq> s \<longrightarrow> (\<exists>x\<in>s. f x = x)"  "continuous_on t g" "g ` t \<subseteq> t"
   obtains y where "y \<in> t" "g y = y" proof- guess h using assms(1) unfolding retract_of_def .. 
   thus ?thesis unfolding retraction_def apply-
@@ -1330,18 +1347,19 @@
 
 subsection {*So the Brouwer theorem for any set with nonempty interior. *}
 
-lemma brouwer_weak: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma brouwer_weak: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
   assumes "compact s" "convex s" "interior s \<noteq> {}" "continuous_on s f" "f ` s \<subseteq> s"
   obtains x where "x \<in> s" "f x = x" proof-
-  have *:"interior {0..1::real^'n} \<noteq> {}" unfolding interior_closed_interval interval_eq_empty by auto
-  have *:"{0..1::real^'n} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] .
-  have "\<forall>f. continuous_on {0..1} f \<and> f ` {0..1} \<subseteq> {0..1} \<longrightarrow> (\<exists>x\<in>{0..1::real^'n}. f x = x)" using brouwer_cube by auto
+  have *:"interior {0::'a..\<chi>\<chi> i.1} \<noteq> {}" unfolding interior_closed_interval interval_eq_empty by auto
+  have *:"{0::'a..\<chi>\<chi> i.1} homeomorphic s" using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] .
+  have "\<forall>f. continuous_on {0::'a..\<chi>\<chi> i.1} f \<and> f ` {0::'a..\<chi>\<chi> i.1} \<subseteq> {0::'a..\<chi>\<chi> i.1} \<longrightarrow> (\<exists>x\<in>{0::'a..\<chi>\<chi> i.1}. f x = x)"
+    using brouwer_cube by auto
   thus ?thesis unfolding homeomorphic_fixpoint_property[OF *] apply(erule_tac x=f in allE)
     apply(erule impE) defer apply(erule bexE) apply(rule_tac x=y in that) using assms by auto qed
 
 subsection {* And in particular for a closed ball. *}
 
-lemma brouwer_ball: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma brouwer_ball: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a"
   assumes "0 < e" "continuous_on (cball a e) f" "f ` (cball a e) \<subseteq> (cball a e)"
   obtains x where "x \<in> cball a e" "f x = x"
   using brouwer_weak[OF compact_cball convex_cball,of a e f] unfolding interior_cball ball_eq_empty
@@ -1351,7 +1369,7 @@
   rather involved @{text "HOMEOMORPHIC_CONVEX_COMPACT"} theorem, just using
   a scaling and translation to put the set inside the unit cube. *}
 
-lemma brouwer: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma brouwer: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a"
   assumes "compact s" "convex s" "s \<noteq> {}" "continuous_on s f" "f ` s \<subseteq> s"
   obtains x where "x \<in> s" "f x = x" proof-
   have "\<exists>e>0. s \<subseteq> cball 0 e" using compact_imp_bounded[OF assms(1)] unfolding bounded_pos
@@ -1371,10 +1389,10 @@
   show thesis apply(rule_tac x="closest_point s x" in that) unfolding x(2)[unfolded o_def]
     apply(rule closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)]) using * by auto qed
 
-text {*So we get the no-retraction theorem. *}                                      
+text {*So we get the no-retraction theorem. *}
 
-lemma no_retraction_cball: assumes "0 < e" 
-  shows "\<not> (frontier(cball a e) retract_of (cball a e))" proof case goal1
+lemma no_retraction_cball: assumes "0 < e" fixes type::"'a::ordered_euclidean_space"
+  shows "\<not> (frontier(cball a e) retract_of (cball (a::'a) e))" proof case goal1
   have *:"\<And>xa. a - (2 *\<^sub>R a - xa) = -(a - xa)" using scaleR_left_distrib[of 1 1 a] by auto
   guess x apply(rule retract_fixpoint_property[OF goal1, of "\<lambda>x. scaleR 2 a - x"])
     apply(rule,rule,erule conjE) apply(rule brouwer_ball[OF assms]) apply assumption+
@@ -1387,20 +1405,20 @@
 
 subsection {*Bijections between intervals. *}
 
-definition "interval_bij = (\<lambda> (a,b) (u,v) (x::real^'n).
-    (\<chi> i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)"
+definition "interval_bij = (\<lambda> (a::'a,b::'a) (u::'a,v::'a) (x::'a::ordered_euclidean_space).
+    (\<chi>\<chi> i. u$$i + (x$$i - a$$i) / (b$$i - a$$i) * (v$$i - u$$i))::'a)"
 
 lemma interval_bij_affine:
- "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +
-            (\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i))"
-  apply rule unfolding Cart_eq interval_bij_def vector_component_simps
-  by(auto simp add: field_simps add_divide_distrib[THEN sym]) 
+ "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi>\<chi> i. (v$$i - u$$i) / (b$$i - a$$i) * x$$i) +
+            (\<chi>\<chi> i. u$$i - (v$$i - u$$i) / (b$$i - a$$i) * a$$i))"
+  apply rule apply(subst euclidean_eq,safe) unfolding euclidean_simps interval_bij_def euclidean_lambda_beta
+  by(auto simp add: field_simps add_divide_distrib[THEN sym])
 
 lemma continuous_interval_bij:
-  "continuous (at x) (interval_bij (a,b::real^'n) (u,v))" 
+  "continuous (at x) (interval_bij (a,b::'a::ordered_euclidean_space) (u,v::'a))" 
   unfolding interval_bij_affine apply(rule continuous_intros)
     apply(rule linear_continuous_at) unfolding linear_conv_bounded_linear[THEN sym]
-    unfolding linear_def unfolding Cart_eq unfolding Cart_lambda_beta defer
+    unfolding linear_def euclidean_eq[where 'a='a] apply safe unfolding euclidean_lambda_beta prefer 3
     apply(rule continuous_intros) by(auto simp add:field_simps add_divide_distrib[THEN sym])
 
 lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a,b) (u,v))"
@@ -1411,23 +1429,25 @@
   apply(cases "b=0") defer apply(rule divide_nonneg_pos) using assms by auto
 
 lemma in_interval_interval_bij: assumes "x \<in> {a..b}" "{u..v} \<noteq> {}"
-  shows "interval_bij (a,b) (u,v) x \<in> {u..v::real^'n}" 
-  unfolding interval_bij_def split_conv mem_interval Cart_lambda_beta proof(rule,rule) 
-  fix i::'n have "{a..b} \<noteq> {}" using assms by auto
-  hence *:"a$i \<le> b$i" "u$i \<le> v$i" using assms(2) unfolding interval_eq_empty not_ex apply-
+  shows "interval_bij (a,b) (u,v) x \<in> {u..v::'a::ordered_euclidean_space}" 
+  unfolding interval_bij_def split_conv mem_interval apply safe unfolding euclidean_lambda_beta
+proof (simp_all only: if_P)
+  fix i assume i:"i<DIM('a)" have "{a..b} \<noteq> {}" using assms by auto
+  hence *:"a$$i \<le> b$$i" "u$$i \<le> v$$i" using assms(2) unfolding interval_eq_empty not_ex apply-
     apply(erule_tac[!] x=i in allE)+ by auto
-  have x:"a$i\<le>x$i" "x$i\<le>b$i" using assms(1)[unfolded mem_interval] by auto
-  have "0 \<le> (x $ i - a $ i) / (b $ i - a $ i) * (v $ i - u $ i)"
+  have x:"a$$i\<le>x$$i" "x$$i\<le>b$$i" using assms(1)[unfolded mem_interval] using i by auto
+  have "0 \<le> (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i)"
     apply(rule mult_nonneg_nonneg) apply(rule divide_nonneg_nonneg)
     using * x by(auto simp add:field_simps)
-  thus "u $ i \<le> u $ i + (x $ i - a $ i) / (b $ i - a $ i) * (v $ i - u $ i)" using * by auto
-  have "((x $ i - a $ i) / (b $ i - a $ i)) * (v $ i - u $ i) \<le> 1 * (v $ i - u $ i)"
+  thus "u $$ i \<le> u $$ i + (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i)" using * by auto
+  have "((x $$ i - a $$ i) / (b $$ i - a $$ i)) * (v $$ i - u $$ i) \<le> 1 * (v $$ i - u $$ i)"
     apply(rule mult_right_mono) unfolding divide_le_eq_1 using * x by auto
-  thus "u $ i + (x $ i - a $ i) / (b $ i - a $ i) * (v $ i - u $ i) \<le> v $ i" using * by auto qed
+  thus "u $$ i + (x $$ i - a $$ i) / (b $$ i - a $$ i) * (v $$ i - u $$ i) \<le> v $$ i" using * by auto
+qed
 
-lemma interval_bij_bij: assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" 
+lemma interval_bij_bij: fixes x::"'a::ordered_euclidean_space" assumes "\<forall>i. a$$i < b$$i \<and> u$$i < v$$i" 
   shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
-  unfolding interval_bij_def split_conv Cart_eq Cart_lambda_beta
+  unfolding interval_bij_def split_conv euclidean_eq[where 'a='a]
   apply(rule,insert assms,erule_tac x=i in allE) by auto
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -0,0 +1,2334 @@
+
+header {*Instanciates the finite cartesian product of euclidean spaces as a euclidean space.*}
+
+theory Cartesian_Euclidean_Space
+imports Finite_Cartesian_Product Integration
+begin
+
+(* TODO:  real_vector^'n is instance of real_vector *)
+
+(* Some strange lemmas, are they really needed? *)
+
+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 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 setsum_mult_product:
+  "setsum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
+  unfolding sumr_group[of h B A, unfolded atLeast0LessThan, symmetric]
+proof (rule setsum_cong, simp, rule setsum_reindex_cong)
+  fix i show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
+  show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}"
+  proof safe
+    fix j assume "j \<in> {i * B..<i * B + B}"
+    thus "j \<in> (\<lambda>j. j + i * B) ` {..<B}"
+      by (auto intro!: image_eqI[of _ _ "j - i * B"])
+  qed simp
+qed simp
+
+subsection{* Basic componentwise operations on vectors. *}
+
+instantiation cart :: (times,finite) times
+begin
+  definition vector_mult_def : "op * \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))"
+  instance ..
+end
+
+instantiation cart :: (one,finite) one
+begin
+  definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"
+  instance ..
+end
+
+instantiation cart :: (ord,finite) ord
+begin
+  definition vector_le_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
+
+text{* The ordering on one-dimensional vectors is linear. *}
+
+class cart_one = assumes UNIV_one: "card (UNIV \<Colon> 'a set) = Suc 0"
+begin
+  subclass finite
+  proof from UNIV_one show "finite (UNIV :: 'a set)"
+      by (auto intro!: card_ge_0_finite) qed
+end
+
+instantiation cart :: (linorder,cart_one) linorder begin
+instance proof
+  guess a B using UNIV_one[where 'a='b] unfolding card_Suc_eq apply- by(erule exE)+
+  hence *:"UNIV = {a}" by auto
+  have "\<And>P. (\<forall>i\<in>UNIV. P i) \<longleftrightarrow> P a" unfolding * by auto hence all:"\<And>P. (\<forall>i. P i) \<longleftrightarrow> P a" by auto
+  fix x y z::"'a^'b::cart_one" note * = vector_le_def vector_less_def all Cart_eq
+  show "x\<le>x" "(x < y) = (x \<le> y \<and> \<not> y \<le> x)" "x\<le>y \<or> y\<le>x" unfolding * by(auto simp only:field_simps)
+  { assume "x\<le>y" "y\<le>z" thus "x\<le>z" unfolding * by(auto simp only:field_simps) }
+  { assume "x\<le>y" "y\<le>x" thus "x=y" unfolding * by(auto simp only:field_simps) }
+qed end
+
+text{* Constant Vectors *} 
+
+definition "vec x = (\<chi> i. x)"
+
+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))"
+
+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 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)
+
+lemma vec_inj[simp]: "vec x = vec y \<longleftrightarrow> x = y" by vector
+
+lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
+
+lemma vec_add: "vec(x + y) = vec x + vec y"  by (vector vec_def)
+lemma vec_sub: "vec(x - y) = vec x - vec y" by (vector vec_def)
+lemma vec_cmul: "vec(c * x) = c *s vec x " by (vector vec_def)
+lemma vec_neg: "vec(- x) = - vec x " by (vector vec_def)
+
+lemma vec_setsum: assumes fS: "finite S"
+  shows "vec(setsum f S) = setsum (vec o f) S"
+  apply (induct rule: finite_induct[OF fS])
+  apply (simp)
+  apply (auto simp add: vec_add)
+  done
+
+text{* Obvious "component-pushing". *}
+
+lemma vec_component [simp]: "vec x $ i = x"
+  by (vector vec_def)
+
+lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i"
+  by vector
+
+lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$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 cart :: (semigroup_mult,finite) semigroup_mult
+  apply (intro_classes) by (vector mult_assoc)
+
+instance cart :: (monoid_mult,finite) monoid_mult
+  apply (intro_classes) by vector+
+
+instance cart :: (ab_semigroup_mult,finite) ab_semigroup_mult
+  apply (intro_classes) by (vector mult_commute)
+
+instance cart :: (ab_semigroup_idem_mult,finite) ab_semigroup_idem_mult
+  apply (intro_classes) by (vector mult_idem)
+
+instance cart :: (comm_monoid_mult,finite) comm_monoid_mult
+  apply (intro_classes) by vector
+
+instance cart :: (semiring,finite) semiring
+  apply (intro_classes) by (vector field_simps)+
+
+instance cart :: (semiring_0,finite) semiring_0
+  apply (intro_classes) by (vector field_simps)+
+instance cart :: (semiring_1,finite) semiring_1
+  apply (intro_classes) by vector
+instance cart :: (comm_semiring,finite) comm_semiring
+  apply (intro_classes) by (vector field_simps)+
+
+instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes)
+instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
+instance cart :: (semiring_0_cancel,finite) semiring_0_cancel by (intro_classes)
+instance cart :: (comm_semiring_0_cancel,finite) comm_semiring_0_cancel by (intro_classes)
+instance cart :: (ring,finite) ring by (intro_classes)
+instance cart :: (semiring_1_cancel,finite) semiring_1_cancel by (intro_classes)
+instance cart :: (comm_semiring_1,finite) comm_semiring_1 by (intro_classes)
+
+instance cart :: (ring_1,finite) ring_1 ..
+
+instance cart :: (real_algebra,finite) real_algebra
+  apply intro_classes
+  apply (simp_all add: vector_scaleR_def field_simps)
+  apply vector
+  apply vector
+  done
+
+instance cart :: (real_algebra_1,finite) 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 one_index[simp]:
+  "(1 :: 'a::one ^'n)$i = 1" by vector
+
+instance cart :: (semiring_char_0,finite) 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 cart :: (comm_ring_1,finite) comm_ring_1 by intro_classes
+instance cart :: (ring_char_0,finite) ring_char_0 by intro_classes
+
+instance cart :: (real_vector,finite) real_vector ..
+
+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 field_simps)
+lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
+  by (vector field_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 field_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 field_simps)
+
+lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
+  by (simp add: Cart_eq)
+
+lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" 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 component_le_norm_cart: "\<bar>x$i\<bar> <= norm x"
+  apply (simp add: norm_vector_def)
+  apply (rule member_le_setL2, simp_all)
+  done
+
+lemma norm_bound_component_le_cart: "norm x <= e ==> \<bar>x$i\<bar> <= e"
+  by (metis component_le_norm_cart order_trans)
+
+lemma norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e"
+  by (metis component_le_norm_cart basic_trans_rules(21))
+
+lemma norm_le_l1_cart: "norm x <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
+  by (simp add: norm_vector_def setL2_le_setsum)
+
+lemma scalar_mult_eq_scaleR: "c *s x = c *\<^sub>R x"
+  unfolding vector_scaleR_def vector_scalar_mult_def by simp
+
+lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"
+  unfolding dist_norm scalar_mult_eq_scaleR
+  unfolding scaleR_right_diff_distrib[symmetric] by simp
+
+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_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)
+
+(* TODO: use setsum_norm_allsubsets_bound *)
+lemma setsum_norm_allsubsets_bound_cart:
+  fixes f:: "'a \<Rightarrow> real ^'n"
+  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_cart)
+  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_cart[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_cart[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
+
+subsection {* A bijection between 'n::finite and {..<CARD('n)} *}
+
+definition cart_bij_nat :: "nat \<Rightarrow> ('n::finite)" where
+  "cart_bij_nat = (SOME p. bij_betw p {..<CARD('n)} (UNIV::'n set) )"
+
+abbreviation "\<pi> \<equiv> cart_bij_nat"
+definition "\<pi>' = inv_into {..<CARD('n)} (\<pi>::nat \<Rightarrow> ('n::finite))"
+
+lemma bij_betw_pi:
+  "bij_betw \<pi> {..<CARD('n::finite)} (UNIV::('n::finite) set)"
+  using ex_bij_betw_nat_finite[of "UNIV::'n set"]
+  by (auto simp: cart_bij_nat_def atLeast0LessThan
+    intro!: someI_ex[of "\<lambda>x. bij_betw x {..<CARD('n)} (UNIV::'n set)"])
+
+lemma bij_betw_pi'[intro]: "bij_betw \<pi>' (UNIV::'n set) {..<CARD('n::finite)}"
+  using bij_betw_inv_into[OF bij_betw_pi] unfolding \<pi>'_def by auto
+
+lemma pi'_inj[intro]: "inj \<pi>'"
+  using bij_betw_pi' unfolding bij_betw_def by auto
+
+lemma pi'_range[intro]: "\<And>i::'n. \<pi>' i < CARD('n::finite)"
+  using bij_betw_pi' unfolding bij_betw_def by auto
+
+lemma \<pi>\<pi>'[simp]: "\<And>i::'n::finite. \<pi> (\<pi>' i) = i"
+  using bij_betw_pi by (auto intro!: f_inv_into_f simp: \<pi>'_def bij_betw_def)
+
+lemma \<pi>'\<pi>[simp]: "\<And>i. i\<in>{..<CARD('n::finite)} \<Longrightarrow> \<pi>' (\<pi> i::'n) = i"
+  using bij_betw_pi by (auto intro!: inv_into_f_eq simp: \<pi>'_def bij_betw_def)
+
+lemma \<pi>\<pi>'_alt[simp]: "\<And>i. i<CARD('n::finite) \<Longrightarrow> \<pi>' (\<pi> i::'n) = i"
+  by auto
+
+lemma \<pi>_inj_on: "inj_on (\<pi>::nat\<Rightarrow>'n::finite) {..<CARD('n)}"
+  using bij_betw_pi[where 'n='n] by (simp add: bij_betw_def)
+
+instantiation cart :: (real_basis,finite) real_basis
+begin
+
+definition "(basis i::'a^'b) =
+  (if i < (CARD('b) * DIM('a))
+  then (\<chi> j::'b. if j = \<pi>(i div DIM('a)) then basis (i mod DIM('a)) else 0)
+  else 0)"
+
+lemma basis_eq:
+  assumes "i < CARD('b)" and "j < DIM('a)"
+  shows "basis (j + i * DIM('a)) = (\<chi> k. if k = \<pi> i then basis j else 0)"
+proof -
+  have "j + i * DIM('a) <  DIM('a) * (i + 1)" using assms by (auto simp: field_simps)
+  also have "\<dots> \<le> DIM('a) * CARD('b)" using assms unfolding mult_le_cancel1 by auto
+  finally show ?thesis
+    unfolding basis_cart_def using assms by (auto simp: Cart_eq not_less field_simps)
+qed
+
+lemma basis_eq_pi':
+  assumes "j < DIM('a)"
+  shows "basis (j + \<pi>' i * DIM('a)) $ k = (if k = i then basis j else 0)"
+  apply (subst basis_eq)
+  using pi'_range assms by simp_all
+
+lemma split_times_into_modulo[consumes 1]:
+  fixes k :: nat
+  assumes "k < A * B"
+  obtains i j where "i < A" and "j < B" and "k = j + i * B"
+proof
+  have "A * B \<noteq> 0"
+  proof assume "A * B = 0" with assms show False by simp qed
+  hence "0 < B" by auto
+  thus "k mod B < B" using `0 < B` by auto
+next
+  have "k div B * B \<le> k div B * B + k mod B" by (rule le_add1)
+  also have "... < A * B" using assms by simp
+  finally show "k div B < A" by auto
+qed simp
+
+lemma split_CARD_DIM[consumes 1]:
+  fixes k :: nat
+  assumes k: "k < CARD('b) * DIM('a)"
+  obtains i and j::'b where "i < DIM('a)" "k = i + \<pi>' j * DIM('a)"
+proof -
+  from split_times_into_modulo[OF k] guess i j . note ij = this
+  show thesis
+  proof
+    show "j < DIM('a)" using ij by simp
+    show "k = j + \<pi>' (\<pi> i :: 'b) * DIM('a)"
+      using ij by simp
+  qed
+qed
+
+lemma linear_less_than_times:
+  fixes i j A B :: nat assumes "i < B" "j < A"
+  shows "j + i * A < B * A"
+proof -
+  have "i * A + j < (Suc i)*A" using `j < A` by simp
+  also have "\<dots> \<le> B * A" using `i < B` unfolding mult_le_cancel2 by simp
+  finally show ?thesis by simp
+qed
+
+instance
+proof
+  let ?b = "basis :: nat \<Rightarrow> 'a^'b"
+  let ?b' = "basis :: nat \<Rightarrow> 'a"
+
+  { fix D :: nat and f :: "nat \<Rightarrow> 'c::real_vector"
+    assume "inj_on f {..<D}"
+    hence eq: "\<And>i. i < D \<Longrightarrow> f ` {..<D} - {f i} = f`({..<D} - {i})"
+      and inj: "\<And>i. inj_on f ({..<D} - {i})"
+      by (auto simp: inj_on_def)
+    have *: "\<And>i. finite (f ` {..<D} - {i})" by simp
+    have "independent (f ` {..<D}) \<longleftrightarrow> (\<forall>a u. a < D \<longrightarrow> (\<Sum>i\<in>{..<D}-{a}. u (f i) *\<^sub>R f i) \<noteq> f a)"
+      unfolding dependent_def span_finite[OF *]
+      by (auto simp: eq setsum_reindex[OF inj]) }
+  note independentI = this
+
+  have setsum_basis:
+    "\<And>f. (\<Sum>x\<in>range basis. f (x::'a)) = f 0 + (\<Sum>i<DIM('a). f (basis i))"
+    unfolding range_basis apply (subst setsum.insert)
+    by (auto simp: basis_eq_0_iff setsum.insert setsum_reindex[OF basis_inj])
+
+  have inj: "inj_on ?b {..<CARD('b)*DIM('a)}"
+    by (auto intro!: inj_onI elim!: split_CARD_DIM split: split_if_asm
+             simp add: Cart_eq basis_eq_pi' all_conj_distrib basis_neq_0
+                       inj_on_iff[OF basis_inj])
+  moreover
+  hence indep: "independent (?b ` {..<CARD('b) * DIM('a)})"
+  proof (rule independentI[THEN iffD2], safe elim!: split_CARD_DIM del: notI)
+    fix j and i :: 'b and u :: "'a^'b \<Rightarrow> real" assume "j < DIM('a)"
+    let ?x = "j + \<pi>' i * DIM('a)"
+    show "(\<Sum>k\<in>{..<CARD('b) * DIM('a)} - {?x}. u(?b k) *\<^sub>R ?b k) \<noteq> ?b ?x"
+      unfolding Cart_eq not_all
+    proof
+      have "(\<lambda>j. j + \<pi>' i*DIM('a))`({..<DIM('a)}-{j}) =
+        {\<pi>' i*DIM('a)..<Suc (\<pi>' i) * DIM('a)} - {?x}"(is "?S = ?I - _")
+      proof safe
+        fix y assume "y \<in> ?I"
+        moreover def k \<equiv> "y - \<pi>' i*DIM('a)"
+        ultimately have "k < DIM('a)" and "y = k + \<pi>' i * DIM('a)" by auto
+        moreover assume "y \<notin> ?S"
+        ultimately show "y = j + \<pi>' i * DIM('a)" by auto
+      qed auto
+
+      have "(\<Sum>k\<in>{..<CARD('b) * DIM('a)} - {?x}. u(?b k) *\<^sub>R ?b k) $ i =
+          (\<Sum>k\<in>{..<CARD('b) * DIM('a)} - {?x}. u(?b k) *\<^sub>R ?b k $ i)" by simp
+      also have "\<dots> = (\<Sum>k\<in>?S. u(?b k) *\<^sub>R ?b k $ i)"
+        unfolding `?S = ?I - {?x}`
+      proof (safe intro!: setsum_mono_zero_cong_right)
+        fix y assume "y \<in> {\<pi>' i*DIM('a)..<Suc (\<pi>' i) * DIM('a)}"
+        moreover have "Suc (\<pi>' i) * DIM('a) \<le> CARD('b) * DIM('a)"
+          unfolding mult_le_cancel2 using pi'_range[of i] by simp
+        ultimately show "y < CARD('b) * DIM('a)" by simp
+      next
+        fix y assume "y < CARD('b) * DIM('a)"
+        with split_CARD_DIM guess l k . note y = this
+        moreover assume "u (?b y) *\<^sub>R ?b y $ i \<noteq> 0"
+        ultimately show "y \<in> {\<pi>' i*DIM('a)..<Suc (\<pi>' i) * DIM('a)}"
+          by (auto simp: basis_eq_pi' split: split_if_asm)
+      qed simp
+      also have "\<dots> = (\<Sum>k\<in>{..<DIM('a)} - {j}. u (?b (k + \<pi>' i*DIM('a))) *\<^sub>R (?b' k))"
+        by (subst setsum_reindex) (auto simp: basis_eq_pi' intro!: inj_onI)
+      also have "\<dots> \<noteq> ?b ?x $ i"
+      proof -
+        note independentI[THEN iffD1, OF basis_inj independent_basis, rule_format]
+        note this[of j "\<lambda>v. u (\<chi> ka::'b. if ka = i then v else (0\<Colon>'a))"]
+        thus ?thesis by (simp add: `j < DIM('a)` basis_eq pi'_range)
+      qed
+      finally show "(\<Sum>k\<in>{..<CARD('b) * DIM('a)} - {?x}. u(?b k) *\<^sub>R ?b k) $ i \<noteq> ?b ?x $ i" .
+    qed
+  qed
+  ultimately
+  show "\<exists>d>0. ?b ` {d..} = {0} \<and> independent (?b ` {..<d}) \<and> inj_on ?b {..<d}"
+    by (auto intro!: exI[of _ "CARD('b) * DIM('a)"] simp: basis_cart_def)
+
+  from indep have exclude_0: "0 \<notin> ?b ` {..<CARD('b) * DIM('a)}"
+    using dependent_0[of "?b ` {..<CARD('b) * DIM('a)}"] by blast
+
+  show "span (range ?b) = UNIV"
+  proof -
+    { fix x :: "'a^'b"
+      let "?if i y" = "(\<chi> k::'b. if k = i then ?b' y else (0\<Colon>'a))"
+      have The_if: "\<And>i j. j < DIM('a) \<Longrightarrow> (THE k. (?if i j) $ k \<noteq> 0) = i"
+        by (rule the_equality) (simp_all split: split_if_asm add: basis_neq_0)
+      { fix x :: 'a
+        have "x \<in> span (range basis)"
+          using span_basis by (auto simp: basis_range)
+        hence "\<exists>u. (\<Sum>x<DIM('a). u (?b' x) *\<^sub>R ?b' x) = x"
+          by (subst (asm) span_finite) (auto simp: setsum_basis) }
+      hence "\<forall>i. \<exists>u. (\<Sum>x<DIM('a). u (?b' x) *\<^sub>R ?b' x) = i" by auto
+      then obtain u where u: "\<forall>i. (\<Sum>x<DIM('a). u i (?b' x) *\<^sub>R ?b' x) = i"
+        by (auto dest: choice)
+      have "\<exists>u. \<forall>i. (\<Sum>j<DIM('a). u (?if i j) *\<^sub>R ?b' j) = x $ i"
+        apply (rule exI[of _ "\<lambda>v. let i = (THE i. v$i \<noteq> 0) in u (x$i) (v$i)"])
+        using The_if u by simp }
+    moreover
+    have "\<And>i::'b. {..<CARD('b)} \<inter> {x. i = \<pi> x} = {\<pi>' i}"
+      using pi'_range[where 'n='b] by auto
+    moreover
+    have "range ?b = {0} \<union> ?b ` {..<CARD('b) * DIM('a)}"
+      by (auto simp: image_def basis_cart_def)
+    ultimately
+    show ?thesis
+      by (auto simp add: Cart_eq setsum_reindex[OF inj] basis_range
+          setsum_mult_product basis_eq if_distrib setsum_cases span_finite
+          setsum_reindex[OF basis_inj])
+  qed
+qed
+
+end
+
+lemma DIM_cart[simp]: "DIM('a^'b) = CARD('b) * DIM('a::real_basis)"
+proof (safe intro!: dimension_eq elim!: split_times_into_modulo del: notI)
+  fix i j assume *: "i < CARD('b)" "j < DIM('a)"
+  hence A: "(i * DIM('a) + j) div DIM('a) = i"
+    by (subst div_add1_eq) simp
+  from * have B: "(i * DIM('a) + j) mod DIM('a) = j"
+    unfolding mod_mult_self3 by simp
+  show "basis (j + i * DIM('a)) \<noteq> (0::'a^'b)" unfolding basis_cart_def
+    using * basis_finite[where 'a='a]
+      linear_less_than_times[of i "CARD('b)" j "DIM('a)"]
+    by (auto simp: A B field_simps Cart_eq basis_eq_0_iff)
+qed (auto simp: basis_cart_def)
+
+lemma if_distr: "(if P then f else g) $ i = (if P then f $ i else g $ i)" by simp
+
+lemma split_dimensions'[consumes 1]:
+  assumes "k < DIM('a::real_basis^'b)"
+  obtains i j where "i < CARD('b::finite)" and "j < DIM('a::real_basis)" and "k = j + i * DIM('a::real_basis)"
+using split_times_into_modulo[OF assms[simplified]] .
+
+lemma cart_euclidean_bound[intro]:
+  assumes j:"j < DIM('a::{real_basis})"
+  shows "j + \<pi>' (i::'b::finite) * DIM('a) < CARD('b) * DIM('a::real_basis)"
+  using linear_less_than_times[OF pi'_range j, of i] .
+
+instance cart :: (real_basis_with_inner,finite) real_basis_with_inner ..
+
+lemma (in real_basis) forall_CARD_DIM:
+  "(\<forall>i<CARD('b) * DIM('a). P i) \<longleftrightarrow> (\<forall>(i::'b::finite) j. j<DIM('a) \<longrightarrow> P (j + \<pi>' i * DIM('a)))"
+   (is "?l \<longleftrightarrow> ?r")
+proof (safe elim!: split_times_into_modulo)
+  fix i :: 'b and j assume "j < DIM('a)"
+  note linear_less_than_times[OF pi'_range[of i] this]
+  moreover assume "?l"
+  ultimately show "P (j + \<pi>' i * DIM('a))" by auto
+next
+  fix i j assume "i < CARD('b)" "j < DIM('a)" and "?r"
+  from `?r`[rule_format, OF `j < DIM('a)`, of "\<pi> i"] `i < CARD('b)`
+  show "P (j + i * DIM('a))" by simp
+qed
+
+lemma (in real_basis) exists_CARD_DIM:
+  "(\<exists>i<CARD('b) * DIM('a). P i) \<longleftrightarrow> (\<exists>i::'b::finite. \<exists>j<DIM('a). P (j + \<pi>' i * DIM('a)))"
+  using forall_CARD_DIM[where 'b='b, of "\<lambda>x. \<not> P x"] by blast
+
+lemma forall_CARD:
+  "(\<forall>i<CARD('b). P i) \<longleftrightarrow> (\<forall>i::'b::finite. P (\<pi>' i))"
+  using forall_CARD_DIM[where 'a=real, of P] by simp
+
+lemma exists_CARD:
+  "(\<exists>i<CARD('b). P i) \<longleftrightarrow> (\<exists>i::'b::finite. P (\<pi>' i))"
+  using exists_CARD_DIM[where 'a=real, of P] by simp
+
+lemmas cart_simps = forall_CARD_DIM exists_CARD_DIM forall_CARD exists_CARD
+
+lemma cart_euclidean_nth[simp]:
+  fixes x :: "('a::real_basis_with_inner, 'b::finite) cart"
+  assumes j:"j < DIM('a)"
+  shows "x $$ (j + \<pi>' i * DIM('a)) = x $ i $$ j"
+  unfolding euclidean_component_def inner_vector_def basis_eq_pi'[OF j] if_distrib cond_application_beta
+  by (simp add: setsum_cases)
+
+lemma real_euclidean_nth:
+  fixes x :: "real^'n"
+  shows "x $$ \<pi>' i = (x $ i :: real)"
+  using cart_euclidean_nth[where 'a=real, of 0 x i] by simp
+
+lemmas nth_conv_component = real_euclidean_nth[symmetric]
+
+lemma mult_split_eq:
+  fixes A :: nat assumes "x < A" "y < A"
+  shows "x + i * A = y + j * A \<longleftrightarrow> x = y \<and> i = j"
+proof
+  assume *: "x + i * A = y + j * A"
+  { fix x y i j assume "i < j" "x < A" and *: "x + i * A = y + j * A"
+    hence "x + i * A < Suc i * A" using `x < A` by simp
+    also have "\<dots> \<le> j * A" using `i < j` unfolding mult_le_cancel2 by simp
+    also have "\<dots> \<le> y + j * A" by simp
+    finally have "i = j" using * by simp }
+  note eq = this
+
+  have "i = j"
+  proof (cases rule: linorder_cases)
+    assume "i < j" from eq[OF this `x < A` *] show "i = j" by simp
+  next
+    assume "j < i" from eq[OF this `y < A` *[symmetric]] show "i = j" by simp
+  qed simp
+  thus "x = y \<and> i = j" using * by simp
+qed simp
+
+instance cart :: (euclidean_space,finite) euclidean_space
+proof (default, safe elim!: split_dimensions')
+  let ?b = "basis :: nat \<Rightarrow> 'a^'b"
+  have if_distrib_op: "\<And>f P Q a b c d.
+    f (if P then a else b) (if Q then c else d) =
+      (if P then if Q then f a c else f a d else if Q then f b c else f b d)"
+    by simp
+
+  fix i j k l
+  assume "i < CARD('b)" "k < CARD('b)" "j < DIM('a)" "l < DIM('a)"
+  thus "?b (j + i * DIM('a)) \<bullet> ?b (l + k * DIM('a)) =
+    (if j + i * DIM('a) = l + k * DIM('a) then 1 else 0)"
+    using inj_on_iff[OF \<pi>_inj_on[where 'n='b], of k i]
+    by (auto simp add: basis_eq inner_vector_def if_distrib_op[of inner] setsum_cases basis_orthonormal mult_split_eq)
+qed
+
+instance cart :: (ordered_euclidean_space,finite) ordered_euclidean_space
+proof
+  fix x y::"'a^'b"
+  show "(x \<le> y) = (\<forall>i<DIM(('a, 'b) cart). x $$ i \<le> y $$ i)"
+    unfolding vector_le_def apply(subst eucl_le) by (simp add: cart_simps)
+  show"(x < y) = (\<forall>i<DIM(('a, 'b) cart). x $$ i < y $$ i)"
+    unfolding vector_less_def apply(subst eucl_less) by (simp add: cart_simps)
+qed
+
+subsection{* Basis vectors in coordinate directions. *}
+
+definition "cart_basis k = (\<chi> i. if i = k then 1 else 0)"
+
+lemma basis_component [simp]: "cart_basis k $ i = (if k=i then 1 else 0)"
+  unfolding cart_basis_def by simp
+
+lemma norm_basis[simp]:
+  shows "norm (cart_basis k :: real ^'n) = 1"
+  apply (simp add: cart_basis_def norm_eq_sqrt_inner) unfolding inner_vector_def
+  apply (vector delta_mult_idempotent)
+  using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"] by auto
+
+lemma norm_basis_1: "norm(cart_basis 1 :: real ^'n::{finite,one}) = 1"
+  by (rule norm_basis)
+
+lemma vector_choose_size: "0 <= c ==> \<exists>(x::real^'n). norm x = c"
+  by (rule exI[where x="c *\<^sub>R cart_basis arbitrary"]) simp
+
+lemma vector_choose_dist: assumes e: "0 <= e"
+  shows "\<exists>(y::real^'n). 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[intro]: "inj (cart_basis :: 'n \<Rightarrow> real ^'n)"
+  by (simp add: inj_on_def Cart_eq)
+
+lemma basis_expansion:
+  "setsum (\<lambda>i. (x$i) *s cart_basis i) UNIV = (x::('a::ring_1) ^'n)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
+  by (auto simp add: Cart_eq if_distrib setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
+
+lemma smult_conv_scaleR: "c *s x = scaleR c x"
+  unfolding vector_scalar_mult_def vector_scaleR_def by simp
+
+lemma basis_expansion':
+  "setsum (\<lambda>i. (x$i) *\<^sub>R cart_basis i) UNIV = x"
+  by (rule basis_expansion [where 'a=real, unfolded smult_conv_scaleR])
+
+lemma basis_expansion_unique:
+  "setsum (\<lambda>i. f i *s cart_basis i) UNIV = (x::('a::comm_ring_1) ^'n) \<longleftrightarrow> (\<forall>i. f i = x$i)"
+  by (simp add: Cart_eq setsum_delta if_distrib cong del: if_weak_cong)
+
+lemma dot_basis:
+  shows "cart_basis i \<bullet> x = x$i" "x \<bullet> (cart_basis i) = (x$i)"
+  by (auto simp add: inner_vector_def cart_basis_def cond_application_beta if_distrib setsum_delta
+           cong del: if_weak_cong)
+
+lemma inner_basis:
+  fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n"
+  shows "inner (cart_basis i) x = inner 1 (x $ i)"
+    and "inner x (cart_basis i) = inner (x $ i) 1"
+  unfolding inner_vector_def cart_basis_def
+  by (auto simp add: cond_application_beta if_distrib setsum_delta cong del: if_weak_cong)
+
+lemma basis_eq_0: "cart_basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> False"
+  by (auto simp add: Cart_eq)
+
+lemma basis_nonzero:
+  shows "cart_basis k \<noteq> (0:: 'a::semiring_1 ^'n)"
+  by (simp add: basis_eq_0)
+
+text {* some lemmas to map between Eucl and Cart *}
+lemma basis_real_n[simp]:"(basis (\<pi>' i)::real^'a) = cart_basis i"
+  unfolding basis_cart_def using pi'_range[where 'n='a]
+  by (auto simp: Cart_eq Cart_lambda_beta)
+
+subsection {* Orthogonality on cartesian products *}
+
+lemma orthogonal_basis:
+  shows "orthogonal (cart_basis i) x \<longleftrightarrow> x$i = (0::real)"
+  by (auto simp add: orthogonal_def inner_vector_def cart_basis_def if_distrib
+                     cond_application_beta setsum_delta cong del: if_weak_cong)
+
+lemma orthogonal_basis_basis:
+  shows "orthogonal (cart_basis i :: real^'n) (cart_basis j) \<longleftrightarrow> i \<noteq> j"
+  unfolding orthogonal_basis[of i] basis_component[of j] by simp
+
+subsection {* Linearity on cartesian products *}
+
+lemma linear_vmul_component:
+  assumes lf: "linear f"
+  shows "linear (\<lambda>x. f x $ k *\<^sub>R v)"
+  using lf
+  by (auto simp add: linear_def algebra_simps)
+
+
+subsection{* Adjoints on cartesian products *}
+
+text {* TODO: The following lemmas about adjoints should hold for any
+Hilbert space (i.e. complete inner product space).
+(see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint})
+*}
+
+lemma adjoint_works_lemma:
+  fixes f:: "real ^'n \<Rightarrow> real ^'m"
+  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:: "real ^ 'm"
+    let ?w = "(\<chi> i. (f (cart_basis i) \<bullet> y)) :: real ^ 'n"
+    {fix x
+      have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *\<^sub>R cart_basis i) ?N) \<bullet> y"
+        by (simp only: basis_expansion')
+      also have "\<dots> = (setsum (\<lambda>i. (x$i) *\<^sub>R f (cart_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: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_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:: "real ^'n \<Rightarrow> real ^'m"
+  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:: "real ^'n \<Rightarrow> real ^'m"
+  assumes lf: "linear f"
+  shows "linear (adjoint f)"
+  unfolding linear_def vector_eq_ldot[where 'a="real^'n", symmetric] apply safe
+  unfolding inner_simps smult_conv_scaleR adjoint_works[OF lf] by auto
+
+lemma adjoint_clauses:
+  fixes f:: "real ^'n \<Rightarrow> real ^'m"
+  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] inner_commute)
+
+lemma adjoint_adjoint:
+  fixes f:: "real ^'n \<Rightarrow> real ^'m"
+  assumes lf: "linear f"
+  shows "adjoint (adjoint f) = f"
+  by (rule adjoint_unique, simp add: adjoint_clauses [OF lf])
+
+
+subsection {* Matrix operations *}
+
+text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
+
+definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
+  where "m ** m' == (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
+
+definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
+  where "m *v x \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
+
+definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
+  where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
+
+definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
+definition transpose where 
+  "(transpose::'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 ** B) + (A ** C)"
+  by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps)
+
+lemma matrix_mul_lid:
+  fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
+  shows "mat 1 ** A = A"
+  apply (simp add: matrix_matrix_mult_def mat_def)
+  apply vector
+  by (auto simp only: if_distrib 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 ^ 'n"
+  shows "A ** mat 1 = A"
+  apply (simp add: matrix_matrix_mult_def mat_def)
+  apply vector
+  by (auto simp only: if_distrib 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)"
+  apply (vector matrix_vector_mult_def mat_def)
+  by (simp add: if_distrib cond_application_beta
+    setsum_delta' cong del: if_weak_cong)
+
+lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"
+  by (simp add: matrix_matrix_mult_def transpose_def Cart_eq mult_commute)
+
+lemma matrix_eq:
+  fixes A B :: "'a::semiring_1 ^ 'n ^ '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 cart_basis_def if_distrib cond_application_beta Cart_eq cong del: if_weak_cong)
+  apply (erule_tac x="cart_basis ia" in allE)
+  apply (erule_tac x="i" in allE)
+  by (auto simp add: cart_basis_def if_distrib cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong)
+
+lemma matrix_vector_mul_component:
+  shows "((A::real^_^_) *v x)$k = (A$k) \<bullet> x"
+  by (simp add: matrix_vector_mult_def inner_vector_def)
+
+lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
+  apply (simp add: inner_vector_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)
+  apply (subst setsum_commute)
+  by simp
+
+lemma transpose_mat: "transpose (mat n) = mat n"
+  by (vector transpose_def mat_def)
+
+lemma transpose_transpose: "transpose(transpose A) = A"
+  by (vector transpose_def)
+
+lemma row_transpose:
+  fixes A:: "'a::semiring_1^_^_"
+  shows "row i (transpose A) = column i A"
+  by (simp add: row_def column_def transpose_def Cart_eq)
+
+lemma column_transpose:
+  fixes A:: "'a::semiring_1^_^_"
+  shows "column i (transpose A) = row i A"
+  by (simp add: row_def column_def transpose_def Cart_eq)
+
+lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A"
+by (auto simp add: rows_def columns_def row_transpose intro: set_ext)
+
+lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" by (metis transpose_transpose rows_transpose)
+
+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 inner_vector_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) = (\<chi> j. setsum (\<lambda>i. (x$i) * (cart_basis i :: 'a^'n)$j) (UNIV :: 'n set))"
+  apply (subst basis_expansion[symmetric])
+  by (vector Cart_eq setsum_component)
+
+lemma linear_componentwise:
+  fixes f:: "real ^'m \<Rightarrow> real ^ _"
+  assumes lf: "linear f"
+  shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (cart_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) *\<^sub>R f (cart_basis i) ) ?M)$j"
+    unfolding vector_smult_component[symmetric] smult_conv_scaleR
+    unfolding setsum_component[of "(\<lambda>i.(x$i) *\<^sub>R f (cart_basis i :: real^'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(cart_basis j))$i)"
+
+lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))"
+  by (simp add: linear_def matrix_vector_mult_def Cart_eq field_simps setsum_right_distrib setsum_addf)
+
+lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::real ^ 'n)"
+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::real ^ 'n))" by (simp add: ext matrix_works)
+
+lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: real ^ 'n)) = A"
+  by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
+
+lemma matrix_compose:
+  assumes lf: "linear (f::real^'n \<Rightarrow> real^'m)"
+  and lg: "linear (g::real^'m \<Rightarrow> real^_)"
+  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^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
+  by (simp add: matrix_vector_mult_def transpose_def Cart_eq mult_commute)
+
+lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
+  apply (rule adjoint_unique)
+  apply (simp add: transpose_def inner_vector_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 :: real^'n \<Rightarrow> real ^'m)"
+  shows "matrix(adjoint f) = transpose(matrix f)"
+  apply (subst matrix_vector_mul[OF lf])
+  unfolding adjoint_matrix matrix_of_matrix_vector_mul ..
+
+section {* lambda_skolem on cartesian products *}
+
+(* FIXME: rename do choice_cart *)
+
+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
+
+subsection {* Standard bases are a spanning set, and obviously finite. *}
+
+lemma span_stdbasis:"span {cart_basis i :: real^'n | 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 finite_stdbasis: "finite {cart_basis i ::real^'n |i. i\<in> (UNIV:: 'n set)}" (is "finite ?S")
+proof-
+  have eq: "?S = cart_basis ` UNIV" by blast
+  show ?thesis unfolding eq by auto
+qed
+
+lemma card_stdbasis: "card {cart_basis i ::real^'n |i. i\<in> (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _")
+proof-
+  have eq: "?S = cart_basis ` UNIV" by blast
+  show ?thesis unfolding eq using card_image[OF basis_inj] by simp
+qed
+
+
+lemma independent_stdbasis_lemma:
+  assumes x: "(x::real ^ 'n) \<in> span (cart_basis ` S)"
+  and iS: "i \<notin> S"
+  shows "(x$i) = 0"
+proof-
+  let ?U = "UNIV :: 'n set"
+  let ?B = "cart_basis ` S"
+  let ?P = "\<lambda>(x::real^_). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0"
+ {fix x::"real^_" 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 {cart_basis i ::real^'n |i. i\<in> (UNIV :: 'n set)}"
+proof-
+  let ?I = "UNIV :: 'n set"
+  let ?b = "cart_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
+
+lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
+  unfolding inner_simps smult_conv_scaleR by auto
+
+lemma linear_eq_stdbasis_cart:
+  assumes lf: "linear (f::real^'m \<Rightarrow> _)" and lg: "linear g"
+  and fg: "\<forall>i. f (cart_basis i) = g(cart_basis i)"
+  shows "f = g"
+proof-
+  let ?U = "UNIV :: 'm set"
+  let ?I = "{cart_basis i:: real^'m|i. i \<in> ?U}"
+  {fix x assume x: "x \<in> (UNIV :: (real^'m) set)"
+    from equalityD2[OF span_stdbasis]
+    have IU: " (UNIV :: (real^'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
+
+lemma bilinear_eq_stdbasis_cart:
+  assumes bf: "bilinear (f:: real^'m \<Rightarrow> real^'n \<Rightarrow> _)"
+  and bg: "bilinear g"
+  and fg: "\<forall>i j. f (cart_basis i) (cart_basis j) = g (cart_basis i) (cart_basis j)"
+  shows "f = g"
+proof-
+  from fg have th: "\<forall>x \<in> {cart_basis i| i. i\<in> (UNIV :: 'm set)}. \<forall>y\<in>  {cart_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
+
+lemma left_invertible_transpose:
+  "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
+  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
+
+lemma right_invertible_transpose:
+  "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
+  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
+
+lemma matrix_left_invertible_injective:
+"(\<exists>B. (B::real^'m^'n) ** (A::real^'n^'m) = 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^'n) ** (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^'m) ** (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^'m"
+  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^'m"
+  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_transpose[symmetric]
+    matrix_left_invertible_independent_columns
+  by (simp add: column_transpose)
+
+lemma matrix_right_invertible_span_columns:
+  "(\<exists>(B::real ^'n^'m). (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
+          unfolding smult_conv_scaleR
+          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", folded smult_conv_scaleR])
+        show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
+          by (rule exI[where x=0], simp)
+      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] if_distrib 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: field_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^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
+  unfolding right_invertible_transpose[symmetric]
+  unfolding columns_transpose[symmetric]
+  unfolding matrix_right_invertible_span_columns
+ ..
+
+text {* The same result in terms of square matrices. *}
+
+lemma matrix_left_right_inverse:
+  fixes A A' :: "real ^'n^'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
+
+text {* 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 transpose_columnvector:
+ "transpose(columnvector v) = rowvector v"
+  by (simp add: transpose_def rowvector_def columnvector_def Cart_eq)
+
+lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"
+  by (simp add: transpose_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::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1"
+  by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vector_def)
+
+lemma dot_matrix_vector_mul:
+  fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
+  shows "(A *v x) \<bullet> (B *v y) =
+      (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
+unfolding dot_matrix_product transpose_columnvector[symmetric]
+  dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
+
+
+lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
+  unfolding infnorm_def apply(rule arg_cong[where f=Sup]) apply safe
+  apply(rule_tac x="\<pi> i" in exI) defer
+  apply(rule_tac x="\<pi>' i" in exI) unfolding nth_conv_component using pi'_range by auto
+
+lemma infnorm_set_image_cart:
+  "{abs(x$i) |i. i\<in> (UNIV :: _ set)} =
+  (\<lambda>i. abs(x$i)) ` (UNIV)" by blast
+
+lemma infnorm_set_lemma_cart:
+  shows "finite {abs((x::'a::abs ^'n)$i) |i. i\<in> (UNIV :: 'n set)}"
+  and "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
+  unfolding  infnorm_set_image_cart
+  by (auto intro: finite_imageI)
+
+lemma component_le_infnorm_cart:
+  shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
+  unfolding nth_conv_component
+  using component_le_infnorm[of x] .
+
+lemma dist_nth_le_cart: "dist (x $ i) (y $ i) \<le> dist x y"
+  unfolding dist_vector_def
+  by (rule member_le_setL2) simp_all
+
+instance cart :: (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_positive_orthant: "closed {x::real^'n. \<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_cart[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 Lim_component_cart:
+  fixes f :: "'a \<Rightarrow> 'b::metric_space ^ 'n"
+  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_cart])
+  done
+
+lemma bounded_component_cart: "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_cart], simp)
+done
+
+lemma compact_lemma_cart:
+  fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n"
+  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_cart)
+    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 cart :: (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_cart [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 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_def by (intro ballI tendsto_intros)
+
+lemma interval_cart: fixes a :: "'a::ord^'n" 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_le_def)
+
+lemma mem_interval_cart: fixes a :: "'a::ord^'n" 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_cart[of a b] by(auto simp add: expand_set_eq vector_less_def vector_le_def)
+
+lemma interval_eq_empty_cart: fixes a :: "real^'n" 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_cart 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  }
+    hence "{a <..< b} \<noteq> {}" using mem_interval_cart(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_cart 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  }
+    hence "{a .. b} \<noteq> {}" using mem_interval_cart(2)[of "?x" a b] by auto  }
+  ultimately show ?th2 by blast
+qed
+
+lemma interval_ne_empty_cart: fixes a :: "real^'n" 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_cart[of a b] by (auto simp add: not_less not_le)
+    (* BH: Why doesn't just "auto" work here? *)
+
+lemma subset_interval_imp_cart: fixes a :: "real^'n" 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_cart
+  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" shows
+ "{a .. a} = {a} \<and> {a<..<a} = {}"
+apply(auto simp add: expand_set_eq vector_less_def vector_le_def Cart_eq)
+apply (simp add: order_eq_iff)
+apply (auto simp add: not_less less_imp_le)
+done
+
+lemma interval_open_subset_closed_cart:  fixes a :: "'a::preorder^'n" 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_le_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_le_def Cart_eq)
+  }
+  ultimately
+  show "a \<le> x \<and> x \<le> b"
+    by(simp add: expand_set_eq vector_less_def vector_le_def Cart_eq)
+qed
+
+lemma subset_interval_cart: fixes a :: "real^'n" 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)
+  using subset_interval[of c d a b] by (simp_all add: cart_simps real_euclidean_nth)
+
+lemma disjoint_interval_cart: fixes a::"real^'n" 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)
+  using disjoint_interval[of a b c d] by (simp_all add: cart_simps real_euclidean_nth)
+
+lemma inter_interval_cart: fixes a :: "'a::linorder^'n" 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_cart
+  by auto
+
+lemma closed_interval_left_cart: fixes b::"real^'n"
+  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_cart[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_cart: fixes a::"real^'n"
+  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_cart[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
+
+lemma is_interval_cart:"is_interval (s::(real^'n) set) \<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)"
+  unfolding is_interval_def Ball_def by (simp add: cart_simps real_euclidean_nth)
+
+lemma closed_halfspace_component_le_cart:
+  shows "closed {x::real^'n. x$i \<le> a}"
+  using closed_halfspace_le[of "(cart_basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+
+lemma closed_halfspace_component_ge_cart:
+  shows "closed {x::real^'n. x$i \<ge> a}"
+  using closed_halfspace_ge[of a "(cart_basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+
+lemma open_halfspace_component_lt_cart:
+  shows "open {x::real^'n. x$i < a}"
+  using open_halfspace_lt[of "(cart_basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+
+lemma open_halfspace_component_gt_cart:
+  shows "open {x::real^'n. x$i  > a}"
+  using open_halfspace_gt[of a "(cart_basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+
+lemma Lim_component_le_cart: fixes f :: "'a \<Rightarrow> real^'n"
+  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 (cart_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 (cart_basis i) x \<le> b}" f net l] unfolding *
+    using closed_halfspace_le[of "(cart_basis i)::real^'n" b] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_ge_cart: fixes f :: "'a \<Rightarrow> real^'n"
+  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 (cart_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 (cart_basis i) x \<ge> b}" f net l] unfolding *
+    using closed_halfspace_ge[of b "(cart_basis i)::real^'n"] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_eq_cart: fixes f :: "'a \<Rightarrow> real^'n"
+  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_cart[OF net, of b i] and
+    Lim_component_le_cart[OF net, of i b] by auto
+
+lemma connected_ivt_component_cart: fixes x::"real^'n" shows
+ "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 "(cart_basis k)::real^'n" a] by (auto simp add: inner_basis)
+
+lemma subspace_substandard_cart:
+ "subspace {x::real^_. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
+  unfolding subspace_def by auto
+
+lemma closed_substandard_cart:
+ "closed {x::real^'n. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
+proof-
+  let ?D = "{i. P i}"
+  let ?Bs = "{{x::real^'n. inner (cart_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 (cart_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" .. }
+  hence "?A = \<Inter> ?Bs" by auto
+  thus ?thesis by(auto simp add: closed_Inter closed_hyperplane)
+qed
+
+lemma dim_substandard_cart:
+  shows "dim {x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" (is "dim ?A = _")
+proof- have *:"{x. \<forall>i<DIM((real, 'n) cart). i \<notin> \<pi>' ` d \<longrightarrow> x $$ i = 0} = 
+    {x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0}"apply safe
+    apply(erule_tac x="\<pi>' i" in allE) defer
+    apply(erule_tac x="\<pi> i" in allE)
+    unfolding image_iff real_euclidean_nth[symmetric] by (auto simp: pi'_inj[THEN inj_eq])
+  have " \<pi>' ` d \<subseteq> {..<DIM((real, 'n) cart)}" using pi'_range[where 'n='n] by auto
+  thus ?thesis using dim_substandard[of "\<pi>' ` d", where 'a="real^'n"] 
+    unfolding * using card_image[of "\<pi>'" d] using pi'_inj unfolding inj_on_def by auto
+qed
+
+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)
+by (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
+
+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: field_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 const_vector_cart:"((\<chi> i. d)::real^'n) = (\<chi>\<chi> i. d)"
+  apply(subst euclidean_eq)
+proof safe case goal1
+  hence *:"(basis i::real^'n) = cart_basis (\<pi> i)"
+    unfolding basis_real_n[THEN sym] by auto
+  have "((\<chi> i. d)::real^'n) $$ i = d" unfolding euclidean_component_def *
+    unfolding dot_basis by auto
+  thus ?case using goal1 by auto
+qed
+
+section "Convex Euclidean Space"
+
+lemma Cart_1:"(1::real^'n) = (\<chi>\<chi> i. 1)"
+  apply(subst euclidean_eq)
+proof safe case goal1 thus ?case using nth_conv_component[THEN sym,where i1="\<pi> i" and x1="1::real^'n"] by auto
+qed
+
+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]
+
+lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_le_def Cart_lambda_beta basis_component vector_uminus_component
+
+lemma convex_box_cart:
+  assumes "\<And>i. convex {x. P i x}"
+  shows "convex {x. \<forall>i. P i (x$i)}"
+  using assms unfolding convex_def by auto
+
+lemma convex_positive_orthant_cart: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
+  by (rule convex_box_cart) (simp add: atLeast_def[symmetric] convex_real_interval)
+
+lemma unit_interval_convex_hull_cart:
+  "{0::real^'n .. 1} = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}" (is "?int = convex hull ?points")
+  unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"]
+  apply(rule arg_cong[where f="\<lambda>x. convex hull x"]) apply(rule set_ext) unfolding mem_Collect_eq
+  apply safe apply(erule_tac x="\<pi>' i" in allE) unfolding nth_conv_component defer
+  apply(erule_tac x="\<pi> i" in allE) by auto
+
+lemma cube_convex_hull_cart:
+  assumes "0 < d" obtains s::"(real^'n) set" where "finite s" "{x - (\<chi> i. d) .. x + (\<chi> i. d)} = convex hull s" 
+proof- from cube_convex_hull[OF assms, where 'a="real^'n" and x=x] guess s . note s=this
+  show thesis apply(rule that[OF s(1)]) unfolding s(2)[THEN sym] const_vector_cart ..
+qed
+
+lemma std_simplex_cart:
+  "(insert (0::real^'n) { cart_basis i | i. i\<in>UNIV}) =
+  (insert 0 { basis i | i. i<DIM((real,'n) cart)})"
+  apply(rule arg_cong[where f="\<lambda>s. (insert 0 s)"])
+  unfolding basis_real_n[THEN sym] apply safe
+  apply(rule_tac x="\<pi>' i" in exI) defer
+  apply(rule_tac x="\<pi> i" in exI) using pi'_range[where 'n='n] by auto
+
+subsection "Brouwer Fixpoint"
+
+lemma kuhn_labelling_lemma_cart:
+  assumes "(\<forall>x::real^_. P x \<longrightarrow> P (f x))"  "\<forall>x. P x \<longrightarrow> (\<forall>i. Q i \<longrightarrow> 0 \<le> x$i \<and> x$i \<le> 1)"
+  shows "\<exists>l. (\<forall>x i. l x i \<le> (1::nat)) \<and>
+             (\<forall>x i. P x \<and> Q i \<and> (x$i = 0) \<longrightarrow> (l x i = 0)) \<and>
+             (\<forall>x i. P x \<and> Q i \<and> (x$i = 1) \<longrightarrow> (l x i = 1)) \<and>
+             (\<forall>x i. P x \<and> Q i \<and> (l x i = 0) \<longrightarrow> x$i \<le> f(x)$i) \<and>
+             (\<forall>x i. P x \<and> Q i \<and> (l x i = 1) \<longrightarrow> f(x)$i \<le> x$i)" proof-
+  have and_forall_thm:"\<And>P Q. (\<forall>x. P x) \<and> (\<forall>x. Q x) \<longleftrightarrow> (\<forall>x. P x \<and> Q x)" by auto
+  have *:"\<forall>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1 \<longrightarrow> (x \<noteq> 1 \<and> x \<le> y \<or> x \<noteq> 0 \<and> y \<le> x)" by auto
+  show ?thesis unfolding and_forall_thm apply(subst choice_iff[THEN sym])+ proof(rule,rule) case goal1
+    let ?R = "\<lambda>y. (P x \<and> Q xa \<and> x $ xa = 0 \<longrightarrow> y = (0::nat)) \<and>
+        (P x \<and> Q xa \<and> x $ xa = 1 \<longrightarrow> y = 1) \<and> (P x \<and> Q xa \<and> y = 0 \<longrightarrow> x $ xa \<le> f x $ xa) \<and> (P x \<and> Q xa \<and> y = 1 \<longrightarrow> f x $ xa \<le> x $ xa)"
+    { assume "P x" "Q xa" hence "0 \<le> f x $ xa \<and> f x $ xa \<le> 1" using assms(2)[rule_format,of "f x" xa]
+        apply(drule_tac assms(1)[rule_format]) by auto }
+    hence "?R 0 \<or> ?R 1" by auto thus ?case by auto qed qed 
+
+lemma interval_bij_cart:"interval_bij = (\<lambda> (a,b) (u,v) (x::real^'n).
+    (\<chi> i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)"
+  unfolding interval_bij_def apply(rule ext)+ apply safe
+  unfolding Cart_eq Cart_lambda_beta unfolding nth_conv_component
+  apply rule apply(subst euclidean_lambda_beta) using pi'_range by auto
+
+lemma interval_bij_affine_cart:
+ "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +
+            (\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i)::real^'n)"
+  apply rule unfolding Cart_eq interval_bij_cart vector_component_simps
+  by(auto simp add: field_simps add_divide_distrib[THEN sym]) 
+
+subsection "Derivative"
+
+lemma has_derivative_vmul_component_cart: fixes c::"real^'a \<Rightarrow> real^'b" and v::"real^'c"
+  assumes "(c has_derivative c') net"
+  shows "((\<lambda>x. c(x)$k *\<^sub>R v) has_derivative (\<lambda>x. (c' x)$k *\<^sub>R v)) net" 
+  using has_derivative_vmul_component[OF assms] 
+  unfolding nth_conv_component .
+
+lemma differentiable_at_imp_differentiable_on: "(\<forall>x\<in>(s::(real^'n) set). f differentiable at x) \<Longrightarrow> f differentiable_on s"
+  unfolding differentiable_on_def by(auto intro!: differentiable_at_withinI)
+
+definition "jacobian f net = matrix(frechet_derivative f net)"
+
+lemma jacobian_works: "(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow> (f has_derivative (\<lambda>h. (jacobian f net) *v h)) net"
+  apply rule unfolding jacobian_def apply(simp only: matrix_works[OF linear_frechet_derivative]) defer
+  apply(rule differentiableI) apply assumption unfolding frechet_derivative_works by assumption
+
+subsection {* Component of the differential must be zero if it exists at a local        *)
+(* maximum or minimum for that corresponding component. *}
+
+lemma differential_zero_maxmin_component: fixes f::"real^'a \<Rightarrow> real^'b"
+  assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))"
+  "f differentiable (at x)" shows "jacobian f (at x) $ k = 0"
+(* FIXME: reuse proof of generic differential_zero_maxmin_component*)
+
+proof(rule ccontr)
+  def D \<equiv> "jacobian f (at x)" assume "jacobian f (at x) $ k \<noteq> 0"
+  then obtain j where j:"D$k$j \<noteq> 0" unfolding Cart_eq D_def by auto
+  hence *:"abs (jacobian f (at x) $ k $ j) / 2 > 0" unfolding D_def by auto
+  note as = assms(3)[unfolded jacobian_works has_derivative_at_alt]
+  guess e' using as[THEN conjunct2,rule_format,OF *] .. note e' = this
+  guess d using real_lbound_gt_zero[OF assms(1) e'[THEN conjunct1]] .. note d = this
+  { fix c assume "abs c \<le> d" 
+    hence *:"norm (x + c *\<^sub>R cart_basis j - x) < e'" using norm_basis[of j] d by auto
+    have "\<bar>(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\<bar> \<le> norm (f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j))" 
+      by(rule component_le_norm_cart)
+    also have "\<dots> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>" using e'[THEN conjunct2,rule_format,OF *] and norm_basis[of j] unfolding D_def[symmetric] by auto
+    finally have "\<bar>(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\<bar> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>" by simp
+    hence "\<bar>f (x + c *\<^sub>R cart_basis j) $ k - f x $ k - c * D $ k $ j\<bar> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>"
+      unfolding vector_component_simps matrix_vector_mul_component unfolding smult_conv_scaleR[symmetric] 
+      unfolding inner_simps dot_basis smult_conv_scaleR by simp  } note * = this
+  have "x + d *\<^sub>R cart_basis j \<in> ball x e" "x - d *\<^sub>R cart_basis j \<in> ball x e"
+    unfolding mem_ball dist_norm using norm_basis[of j] d by auto
+  hence **:"((f (x - d *\<^sub>R cart_basis j))$k \<le> (f x)$k \<and> (f (x + d *\<^sub>R cart_basis j))$k \<le> (f x)$k) \<or>
+         ((f (x - d *\<^sub>R cart_basis j))$k \<ge> (f x)$k \<and> (f (x + d *\<^sub>R cart_basis j))$k \<ge> (f x)$k)" using assms(2) by auto
+  have ***:"\<And>y y1 y2 d dx::real. (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
+  show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\<bar>D $ k $ j\<bar> / 2 * \<bar>d\<bar>"]) 
+    using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left
+    unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding algebra_simps by (auto intro: mult_pos_pos)
+qed
+
+subsection {* Lemmas for working on real^1 *}
+
+lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
+  by (metis num1_eq_iff)
+
+lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
+  by auto (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 [simp]: "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)
+
+instantiation num1 :: cart_one begin
+instance proof
+  show "CARD(1) = Suc 0" by auto
+qed end
+
+(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *)
+
+abbreviation vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x \<equiv> vec x"
+
+abbreviation dest_vec1:: "'a ^1 \<Rightarrow> 'a"
+  where "dest_vec1 x \<equiv> (x$1)"
+
+lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
+  by (simp_all add:  Cart_eq)
+
+lemma vec1_component[simp]: "(vec1 x)$1 = x"
+  by (simp_all add:  Cart_eq)
+
+declare vec1_dest_vec1(1) [simp]
+
+lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))"
+  by (metis vec1_dest_vec1(1))
+
+lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))"
+  by (metis vec1_dest_vec1(1))
+
+lemma vec1_eq[simp]:  "vec1 x = vec1 y \<longleftrightarrow> x = y"
+  by (metis vec1_dest_vec1(2))
+
+lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y"
+  by (metis vec1_dest_vec1(1))
+
+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)
+
+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)
+
+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{* 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
+  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
+
+lemma range_vec1[simp]:"range vec1 = UNIV" apply(rule set_ext,rule) unfolding image_iff defer
+  apply(rule_tac x="dest_vec1 x" in bexI) by auto
+
+lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1"
+  by (simp)
+
+lemma dest_vec1_vec: "dest_vec1(vec x) = x"
+  by (simp)
+
+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
+  apply auto
+  done
+
+lemma norm_vec1 [simp]: "norm(vec1 x) = abs(x)"
+  by (simp add: vec_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(1) norm_vec1)
+
+lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component
+   vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def
+
+lemma bounded_linear_vec1:"bounded_linear (vec1::real\<Rightarrow>real^1)"
+  unfolding bounded_linear_def additive_def bounded_linear_axioms_def 
+  unfolding smult_conv_scaleR[THEN sym] unfolding vec1_dest_vec1_simps
+  apply(rule conjI) defer apply(rule conjI) defer apply(rule_tac x=1 in exI) by auto
+
+lemma linear_vmul_dest_vec1:
+  fixes f:: "real^_ \<Rightarrow> real^1"
+  shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
+  unfolding smult_conv_scaleR
+  by (rule linear_vmul_component)
+
+lemma linear_from_scalars:
+  assumes lf: "linear (f::real^1 \<Rightarrow> real^_)"
+  shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"
+  unfolding smult_conv_scaleR
+  apply (rule ext)
+  apply (subst matrix_works[OF lf, symmetric])
+  apply (auto simp add: Cart_eq matrix_vector_mult_def column_def mult_commute)
+  done
+
+lemma linear_to_scalars: assumes lf: "linear (f::real ^'n \<Rightarrow> real^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 row_def inner_vector_def mult_commute)
+  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 vec_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{* Lifting and dropping *}
+
+lemma continuous_on_o_dest_vec1: fixes f::"real \<Rightarrow> 'a::real_normed_vector"
+  assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)"
+  using assms unfolding continuous_on_iff apply safe
+  apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe
+  apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real 
+  apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:vector_le_def)
+
+lemma continuous_on_o_vec1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
+  assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)"
+  using assms unfolding continuous_on_iff apply safe
+  apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe
+  apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real 
+  apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:vector_le_def)
+
+lemma continuous_on_vec1:"continuous_on A (vec1::real\<Rightarrow>real^1)"
+  by(rule linear_continuous_on[OF bounded_linear_vec1])
+
+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_le_def)
+
+lemma vec1_interval:fixes a::"real" shows
+  "vec1 ` {a .. b} = {vec1 a .. vec1 b}"
+  "vec1 ` {a<..<b} = {vec1 a<..<vec1 b}"
+  apply(rule_tac[!] set_ext) unfolding image_iff vector_less_def unfolding mem_interval_cart
+  unfolding forall_1 unfolding vec1_dest_vec1_simps
+  apply rule defer apply(rule_tac x="dest_vec1 x" in bexI) prefer 3 apply rule defer
+  apply(rule_tac x="dest_vec1 x" in bexI) by auto
+
+(* Some special cases for intervals in R^1.                                  *)
+
+lemma interval_cases_1: fixes x :: "real^1" shows
+ "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
+  unfolding Cart_eq vector_less_def vector_le_def mem_interval_cart by(auto simp del:dest_vec1_eq)
+
+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)"
+  unfolding Cart_eq vector_less_def vector_le_def mem_interval_cart by(auto simp del:dest_vec1_eq)
+
+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_cart and ex_1 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_cart[of a b c d] unfolding forall_1 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"
+unfolding set_eq_subset[of "{a .. b}" "{c .. d}"]
+unfolding subset_interval_1(1)[of a b c d]
+unfolding subset_interval_1(1)[of c d a b]
+by auto
+
+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_cart and 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_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq)
+
+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_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq)
+
+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_cart[of f l net 1 b] 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_cart[of f l net b 1] by auto
+
+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  by auto
+qed
+
+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_le_def Cart_eq)
+
+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_le_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 open_dest_vec1_vimage: "open S \<Longrightarrow> open (dest_vec1 -` S)"
+unfolding open_vector_def forall_1 by auto
+
+lemma tendsto_dest_vec1 [tendsto_intros]:
+  "(f ---> l) net \<Longrightarrow> ((\<lambda>x. dest_vec1 (f x)) ---> dest_vec1 l) net"
+by(rule tendsto_Cart_nth)
+
+lemma continuous_dest_vec1: "continuous net f \<Longrightarrow> continuous net (\<lambda>x. dest_vec1 (f x))"
+  unfolding continuous_def by (rule tendsto_dest_vec1)
+
+lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))" 
+  apply safe defer apply(erule_tac x="vec1 x" in allE) 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 dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding dist_norm by auto
+
+lemma bounded_linear_vec1_dest_vec1: fixes f::"real \<Rightarrow> real"
+  shows "linear (vec1 \<circ> f \<circ> dest_vec1) = bounded_linear f" (is "?l = ?r") proof-
+  { assume ?l guess K using linear_bounded[OF `?l`] ..
+    hence "\<exists>K. \<forall>x. \<bar>f x\<bar> \<le> \<bar>x\<bar> * K" apply(rule_tac x=K in exI)
+      unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) }
+  thus ?thesis unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def
+    unfolding vec1_dest_vec1_simps by auto qed
+
+lemma vec1_le[simp]:fixes a::real shows "vec1 a \<le> vec1 b \<longleftrightarrow> a \<le> b"
+  unfolding vector_le_def by auto
+lemma vec1_less[simp]:fixes a::real shows "vec1 a < vec1 b \<longleftrightarrow> a < b"
+  unfolding vector_less_def by auto
+
+
+subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *}
+
+lemma has_derivative_within_vec1_dest_vec1: fixes f::"real\<Rightarrow>real" shows
+  "((vec1 \<circ> f \<circ> dest_vec1) has_derivative (vec1 \<circ> f' \<circ> dest_vec1)) (at (vec1 x) within vec1 ` s)
+  = (f has_derivative f') (at x within s)"
+  unfolding has_derivative_within unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear]
+  unfolding o_def Lim_within Ball_def unfolding forall_vec1 
+  unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto  
+
+lemma has_derivative_at_vec1_dest_vec1: fixes f::"real\<Rightarrow>real" shows
+  "((vec1 \<circ> f \<circ> dest_vec1) has_derivative (vec1 \<circ> f' \<circ> dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)"
+  using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV] by auto
+
+lemma bounded_linear_vec1': fixes f::"'a::real_normed_vector\<Rightarrow>real"
+  shows "bounded_linear f = bounded_linear (vec1 \<circ> f)"
+  unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def
+  unfolding vec1_dest_vec1_simps by auto
+
+lemma bounded_linear_dest_vec1: fixes f::"real\<Rightarrow>'a::real_normed_vector"
+  shows "bounded_linear f = bounded_linear (f \<circ> dest_vec1)"
+  unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def
+  unfolding vec1_dest_vec1_simps by auto
+
+lemma has_derivative_at_vec1: fixes f::"'a::real_normed_vector\<Rightarrow>real" shows
+  "(f has_derivative f') (at x) = ((vec1 \<circ> f) has_derivative (vec1 \<circ> f')) (at x)"
+  unfolding has_derivative_at unfolding bounded_linear_vec1'[unfolded linear_conv_bounded_linear]
+  unfolding o_def Lim_at unfolding vec1_dest_vec1_simps dist_vec1_0 by auto
+
+lemma has_derivative_within_dest_vec1:fixes f::"real\<Rightarrow>'a::real_normed_vector" shows
+  "((f \<circ> dest_vec1) has_derivative (f' \<circ> dest_vec1)) (at (vec1 x) within vec1 ` s) = (f has_derivative f') (at x within s)"
+  unfolding has_derivative_within bounded_linear_dest_vec1 unfolding o_def Lim_within Ball_def
+  unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto
+
+lemma has_derivative_at_dest_vec1:fixes f::"real\<Rightarrow>'a::real_normed_vector" shows
+  "((f \<circ> dest_vec1) has_derivative (f' \<circ> dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)"
+  using has_derivative_within_dest_vec1[where s=UNIV] by(auto simp add:within_UNIV)
+
+subsection {* In particular if we have a mapping into @{typ "real^1"}. *}
+
+lemma onorm_vec1: fixes f::"real \<Rightarrow> real"
+  shows "onorm (\<lambda>x. vec1 (f (dest_vec1 x))) = onorm f" proof-
+  have "\<forall>x::real^1. norm x = 1 \<longleftrightarrow> x\<in>{vec1 -1, vec1 (1::real)}" unfolding forall_vec1 by(auto simp add:Cart_eq)
+  hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto
+  have 2:"{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = (\<lambda>x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" by auto
+  have "\<forall>x::real. norm x = 1 \<longleftrightarrow> x\<in>{-1, 1}" by auto hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto
+  have 4:"{norm (f x) |x. norm x = 1} = (\<lambda>x. norm (f x)) ` {x. norm x=1}" by auto
+  show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max) qed
+
+lemma convex_vec1:"convex (vec1 ` s) = convex (s::real set)"
+  unfolding convex_def Ball_def forall_vec1 unfolding vec1_dest_vec1_simps image_iff by auto
+
+lemma bounded_linear_component_cart[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)"
+  apply(rule bounded_linearI[where K=1]) 
+  using component_le_norm_cart[of _ k] unfolding real_norm_def by auto
+
+lemma bounded_vec1[intro]:  "bounded s \<Longrightarrow> bounded (vec1 ` (s::real set))"
+  unfolding bounded_def apply safe apply(rule_tac x="vec1 x" in exI,rule_tac x=e in exI)
+  by(auto simp add: dist_real dist_real_def)
+
+(*lemma content_closed_interval_cases_cart:
+  "content {a..b::real^'n} =
+  (if {a..b} = {} then 0 else setprod (\<lambda>i. b$i - a$i) UNIV)" 
+proof(cases "{a..b} = {}")
+  case True thus ?thesis unfolding content_def by auto
+next case Falsethus ?thesis unfolding content_def unfolding if_not_P[OF False]
+  proof(cases "\<forall>i. a $ i \<le> b $ i")
+    case False thus ?thesis unfolding content_def using t by auto
+  next case True note interval_eq_empty
+   apply auto 
+  
+  sorry*)
+
+lemma integral_component_eq_cart[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real^'m"
+  assumes "f integrable_on s" shows "integral s (\<lambda>x. f x $ k) = integral s f $ k"
+  using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] .
+
+lemma interval_split_cart:
+  "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
+  "{a..b} \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
+  apply(rule_tac[!] set_ext) unfolding Int_iff mem_interval_cart mem_Collect_eq
+  unfolding Cart_lambda_beta by auto
+
+(*lemma content_split_cart:
+  "content {a..b::real^'n} = content({a..b} \<inter> {x. x$k \<le> c}) + content({a..b} \<inter> {x. x$k >= c})"
+proof- note simps = interval_split_cart content_closed_interval_cases_cart Cart_lambda_beta vector_le_def
+  { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps by auto }
+  have *:"UNIV = insert k (UNIV - {k})" "\<And>x. finite (UNIV-{x::'n})" "\<And>x. x\<notin>UNIV-{x}" by auto
+  have *:"\<And>X Y Z. (\<Prod>i\<in>UNIV. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>UNIV-{k}. Z i (Y i))"
+    "(\<Prod>i\<in>UNIV. b$i - a$i) = (\<Prod>i\<in>UNIV-{k}. b$i - a$i) * (b$k - a$k)" 
+    apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto
+  assume as:"a\<le>b" moreover have "\<And>x. min (b $ k) c = max (a $ k) c
+    \<Longrightarrow> x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)"
+    by  (auto simp add:field_simps)
+  moreover have "\<not> a $ k \<le> c \<Longrightarrow> \<not> c \<le> b $ k \<Longrightarrow> False"
+    unfolding not_le using as[unfolded vector_le_def,rule_format,of k] by auto
+  ultimately show ?thesis 
+    unfolding simps unfolding *(1)[of "\<lambda>i x. b$i - x"] *(1)[of "\<lambda>i x. x - a$i"] *(2) by(auto)
+qed*)
+
+lemma has_integral_vec1: assumes "(f has_integral k) {a..b}"
+  shows "((\<lambda>x. vec1 (f x)) has_integral (vec1 k)) {a..b}"
+proof- have *:"\<And>p. (\<Sum>(x, k)\<in>p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k)"
+    unfolding vec_sub Cart_eq by(auto simp add: split_beta)
+  show ?thesis using assms unfolding has_integral apply safe
+    apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe)
+    apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed
+
+end
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -13,14 +13,18 @@
 (* 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]
+lemma basis_0[simp]:"(basis i::'a::euclidean_space) = 0 \<longleftrightarrow> i\<ge>DIM('a)"
+  using norm_basis[of i, where 'a='a] unfolding norm_eq_zero[where 'a='a,THEN sym] by auto
 
-(*lemma dim1in[intro]:"Suc 0 \<in> {1::nat .. CARD(1)}" by auto*)
+lemma scaleR_2:
+  fixes x :: "'a::real_vector"
+  shows "scaleR 2 x = x + x"
+unfolding one_add_one_is_two [symmetric] scaleR_left_distrib by simp
 
-lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_le_def Cart_lambda_beta basis_component vector_uminus_component
+declare euclidean_simps[simp]
 
-lemma norm_not_0:"(x::real^'n)\<noteq>0 \<Longrightarrow> norm x \<noteq> 0" by auto
+lemma vector_choose_size: "0 <= c ==> \<exists>(x::'a::euclidean_space). norm x = c"
+  apply (rule exI[where x="c *\<^sub>R basis 0 ::'a"]) using DIM_positive[where 'a='a] 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"
@@ -37,23 +41,20 @@
   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 image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n)) ` {a..b} =
+lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::'a::ordered_euclidean_space)) ` {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 dist_triangle_eq:
-  fixes x y z :: "real ^ _"
+  fixes x y z :: "'a::euclidean_space"
   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 *]
+  show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *]
     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) = - y \<Longrightarrow> norm x = norm y" by auto
+lemma norm_minus_eqI:"x = - 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
@@ -61,10 +62,13 @@
 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 real_eq_of_nat real_of_nat_1 real_of_nat_le_iff) 
+lemma norm_lt: "norm x < norm y \<longleftrightarrow> inner x x < inner y y"
+  unfolding norm_eq_sqrt_inner by simp
 
-lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto
+lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
+  unfolding norm_eq_sqrt_inner by simp
+
+
 
 subsection {* Affine set and affine hull.*}
 
@@ -149,10 +153,12 @@
         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
+      hence "u x + (1 - u x) = 1 \<Longrightarrow> u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
+        apply-apply(rule as(3)[rule_format]) 
+        unfolding  RealVector.scaleR_right.setsum using x(1) as(6) by auto
       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
+         using `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
@@ -182,7 +188,7 @@
       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 scaleR_scaleR[THEN sym] RealVector.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:
@@ -272,9 +278,9 @@
 subsection {* Some relations between affine hull and subspaces. *}
 
 lemma affine_hull_insert_subset_span:
-  fixes a :: "real ^ _"
+  fixes a :: "'a::euclidean_space"
   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
+  unfolding subset_eq Ball_def unfolding affine_hull_explicit span_explicit mem_Collect_eq
   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
@@ -290,14 +296,14 @@
     unfolding as by simp qed
 
 lemma affine_hull_insert_span:
-  fixes a :: "real ^ _"
+  fixes a :: "'a::euclidean_space"
   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
+  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 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])
@@ -310,7 +316,7 @@
     by (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps *) qed
 
 lemma affine_hull_span:
-  fixes a :: "real ^ _"
+  fixes a :: "'a::euclidean_space"
   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
@@ -427,12 +433,12 @@
 
 subsection {* Balls, being convex, are connected. *}
 
-lemma convex_box:
-  assumes "\<And>i. convex {x. P i x}"
-  shows "convex {x. \<forall>i. P i (x$i)}"
-using assms unfolding convex_def by auto
+lemma convex_box: fixes a::"'a::euclidean_space"
+  assumes "\<And>i. i<DIM('a) \<Longrightarrow> convex {x. P i x}"
+  shows "convex {x. \<forall>i<DIM('a). P i (x$$i)}"
+  using assms unfolding convex_def by(auto simp add:euclidean_simps)
 
-lemma convex_positive_orthant: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
+lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\<forall>i<DIM('a). 0 \<le> x$$i)}"
   by (rule convex_box) (simp add: atLeast_def[symmetric] convex_real_interval)
 
 lemma convex_local_global_minimum:
@@ -755,7 +761,7 @@
 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^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: field_simps)
+         "\<And>x y z ::_::euclidean_space. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: field_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
@@ -773,22 +779,22 @@
 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
+  fixes s :: "(_::euclidean_space) set" shows "subspace s \<Longrightarrow> affine s"
+  unfolding subspace_def affine_def 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"
+  fixes s :: "(_::euclidean_space) 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)"
+  fixes s :: "(_::euclidean_space) set" shows "(affine hull s) \<subseteq> (span s)"
 by (metis hull_minimal mem_def span_inc subspace_imp_affine subspace_span)
 
 lemma convex_hull_subset_span:
-  fixes s :: "(real ^ _) set" shows "(convex hull s) \<subseteq> (span s)"
+  fixes s :: "(_::euclidean_space) set" shows "(convex hull s) \<subseteq> (span s)"
 by (metis hull_minimal mem_def span_inc subspace_imp_convex subspace_span)
 
 lemma convex_hull_subset_affine_hull: "(convex hull s) \<subseteq> (affine hull s)"
@@ -796,16 +802,16 @@
 
 
 lemma affine_dependent_imp_dependent:
-  fixes s :: "(real ^ _) set" shows "affine_dependent s \<Longrightarrow> dependent s"
+  fixes s :: "(_::euclidean_space) 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"
+  fixes s :: "(_::euclidean_space) 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 
+  from assms(1)[unfolded dependent_explicit] 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"
 
@@ -826,7 +832,7 @@
     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) 
+    unfolding setsum_clauses(2)[OF fin] using `a\<notin>s` `t\<subseteq>s` by (auto simp add: *)
   ultimately show ?thesis unfolding affine_dependent_explicit
     apply(rule_tac x="insert a t" in exI) by auto 
 qed
@@ -842,22 +848,22 @@
   thus ?thesis unfolding convex_def cone_def by blast
 qed
 
-lemma affine_dependent_biggerset: fixes s::"(real^'n) set"
-  assumes "finite s" "card s \<ge> CARD('n) + 2"
+lemma affine_dependent_biggerset: fixes s::"('a::euclidean_space) set"
+  assumes "finite s" "card s \<ge> DIM('a) + 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)
+  also have "\<dots> > DIM('a)" 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) set)" "card s \<ge> dim s + 2"
+  assumes "finite (s::('a::euclidean_space) set)" "card s \<ge> dim s + 2"
   shows "affine_dependent s"
 proof-
   from assms(2) have "s \<noteq> {}" by auto
@@ -876,8 +882,8 @@
 
 subsection {* Caratheodory's theorem. *}
 
-lemma convex_hull_caratheodory: fixes p::"(real^'n) set"
-  shows "convex hull p = {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> CARD('n) + 1 \<and>
+lemma convex_hull_caratheodory: fixes p::"('a::euclidean_space) set"
+  shows "convex hull p = {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 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)
@@ -888,8 +894,8 @@
   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"
+  have "card s \<le> DIM('a) + 1" proof(rule ccontr, simp only: not_le)
+    assume "DIM('a) + 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
@@ -918,33 +924,34 @@
     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 *:"\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
+      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)
+      using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
     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)
+      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
+  thus "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 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. \<exists>s. finite s \<and> s \<subseteq> p \<and>
-      card s \<le> CARD('n) + 1 \<and> x \<in> convex hull s}"
+ "convex hull p = {x::'a::euclidean_space. \<exists>s. finite s \<and> s \<subseteq> p \<and>
+      card s \<le> DIM('a) + 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"
+  then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 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"
+  thus "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 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
+  fix x assume "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s"
+  then obtain s where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1" "x \<in> convex hull s" by auto
   thus "x \<in> convex hull p" using hull_mono[OF `s\<subseteq>p`] by auto
 qed
 
@@ -988,16 +995,6 @@
   qed
 qed
 
-(* TODO: move *)
-lemma compact_real_interval:
-  fixes a b :: real shows "compact {a..b}"
-proof (rule bounded_closed_imp_compact)
-  have "\<forall>y\<in>{a..b}. dist a y \<le> dist a b"
-    unfolding dist_real_def by auto
-  thus "bounded {a..b}" unfolding bounded_def by fast
-  show "closed {a..b}" by (rule closed_real_atLeastAtMost)
-qed
-
 lemma compact_convex_combinations:
   fixes s t :: "'a::real_normed_vector set"
   assumes "compact s" "compact t"
@@ -1016,18 +1013,18 @@
     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)
+    apply (intro compact_Times compact_interval assms)
     done
 qed
 
-lemma compact_convex_hull: fixes s::"(real^'n) set"
+lemma compact_convex_hull: fixes s::"('a::euclidean_space) 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"))
+  proof(induct ("DIM('a) + 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
     case 0 thus ?case unfolding * by simp
@@ -1095,7 +1092,7 @@
 qed
 
 lemma finite_imp_compact_convex_hull:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) set"
   shows "finite s \<Longrightarrow> compact(convex hull s)"
 by (metis compact_convex_hull finite_imp_compact)
 
@@ -1178,7 +1175,7 @@
 qed (auto simp add: assms)
 
 lemma simplex_furthest_le:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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-
@@ -1194,12 +1191,12 @@
 qed
 
 lemma simplex_furthest_le_exists:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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"
+  fixes s :: "('a::euclidean_space) 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-
@@ -1220,7 +1217,7 @@
 qed 
 
 lemma simplex_extremal_le_exists:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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
@@ -1254,14 +1251,6 @@
  "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:
   assumes "inner y z > 0"
   shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y"
@@ -1386,24 +1375,26 @@
 qed
 
 lemma separating_hyperplane_closed_0:
-  assumes "convex (s::(real^'n) set)" "closed s" "0 \<notin> s"
+  assumes "convex (s::('a::euclidean_space) 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) = 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
+  proof(cases "s={}")
+  case True have "norm ((basis 0)::'a) = 1" by auto
+  hence "norm ((basis 0)::'a) = 1" "basis 0 \<noteq> (0::'a)" defer
+    apply(subst norm_le_zero_iff[THEN sym]) by auto
+  thus ?thesis apply(rule_tac x="basis 0" in exI, rule_tac x=1 in exI)
+    using True using DIM_positive[where 'a='a] by auto
 next case False thus ?thesis using False using separating_hyperplane_closed_point[OF assms]
     apply - apply(erule exE)+ unfolding inner.zero_right 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) set)" "closed s" "convex t" "compact t" "t \<noteq> {}" "s \<inter> t = {}"
+  assumes "convex (s::('a::euclidean_space) 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
+  obtain z::"'a" 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
@@ -1430,7 +1421,7 @@
 qed
 
 lemma separating_hyperplane_compact_closed:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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)"
@@ -1440,9 +1431,9 @@
 subsection {* General case without assuming closure and getting non-strict separation. *}
 
 lemma separating_hyperplane_set_0:
-  assumes "convex s" "(0::real^'n) \<notin> s"
+  assumes "convex s" "(0::'a::euclidean_space) \<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}"
+proof- let ?k = "\<lambda>c. {x::'a. 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-
@@ -1462,7 +1453,7 @@
   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) set)" "s \<noteq> {}" "t \<noteq> {}" "s \<inter> t = {}"
+  assumes "convex s" "convex (t::('a::euclidean_space) 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" 
@@ -1540,7 +1531,7 @@
 subsection {* Convex set as intersection of halfspaces. *}
 
 lemma convex_halfspace_intersection:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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- 
@@ -1567,7 +1558,7 @@
     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^_)"
+  assumes "finite s" "setsum f s = 0" "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
   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 
@@ -1602,13 +1593,13 @@
   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])
+    by(auto simp add: setsum_negf 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])
+    by(auto simp add: setsum_negf 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
 
@@ -1621,16 +1612,16 @@
 
 subsection {* Helly's theorem. *}
 
-lemma helly_induct: fixes f::"(real^'n) set set"
-  assumes "card f = 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> {}"
+lemma helly_induct: fixes f::"('a::euclidean_space) set set"
+  assumes "card f = n" "n \<ge> DIM('a) + 1"
+  "\<forall>s\<in>f. convex s" "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter> t \<noteq> {}"
   shows "\<Inter> f \<noteq> {}"
 using assms proof(induct n arbitrary: f)
 case (Suc n)
 have "finite f" using `card f = Suc n` by (auto intro: card_ge_0_finite)
-show "\<Inter> f \<noteq> {}" apply(cases "n = CARD('n)") apply(rule Suc(5)[rule_format])
+show "\<Inter> f \<noteq> {}" apply(cases "n = DIM('a)") apply(rule Suc(5)[rule_format])
   unfolding `card f = Suc n` 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
+  assume ng:"n \<noteq> DIM('a)" 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 `finite f`] `card f = Suc n`
     defer defer apply(rule Suc(4)[rule_format]) defer apply(rule Suc(5)[rule_format]) using Suc(3) `finite f` by auto
   then obtain X where X:"\<forall>s\<in>f. X s \<in> \<Inter>(f - {s})" by auto
@@ -1659,9 +1650,9 @@
     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) set set"
-  assumes "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> {}"
+lemma helly: fixes f::"('a::euclidean_space) set set"
+  assumes "card f \<ge> DIM('a) + 1" "\<forall>s\<in>f. convex s"
+          "\<forall>t\<subseteq>f. card t = DIM('a) + 1 \<longrightarrow> \<Inter> t \<noteq> {}"
   shows "\<Inter> f \<noteq>{}"
   apply(rule helly_induct) using assms by auto
 
@@ -1690,7 +1681,7 @@
 subsection {* Homeomorphism of all convex compact sets with nonempty interior. *}
 
 lemma compact_frontier_line_lemma:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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-
@@ -1698,10 +1689,11 @@
   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. u *\<^sub>R x) ` {0 .. b / norm x}"
     by auto
+  have *:"\<And>x A B. x\<in>A \<Longrightarrow> x\<in>B \<Longrightarrow> A\<inter>B \<noteq> {}" by blast
   have "compact ?A" unfolding A apply(rule compact_continuous_image, rule continuous_at_imp_continuous_on)
     apply(rule, rule continuous_vmul)
-    apply(rule continuous_at_id) by(rule compact_real_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)])
+    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 *[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
@@ -1726,12 +1718,12 @@
 qed
 
 lemma starlike_compact_projective:
-  assumes "compact s" "cball (0::real^'n) 1 \<subseteq> s "
+  assumes "compact s" "cball (0::'a::euclidean_space) 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) 1)"
+  shows "s homeomorphic (cball (0::'a::euclidean_space) 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"
+  def pi \<equiv> "\<lambda>x::'a. 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
@@ -1744,7 +1736,7 @@
     apply simp
     apply (rule continuous_at_id)
     done
-  def sphere \<equiv> "{x::real^'n. norm x = 1}"
+  def sphere \<equiv> "{x::'a. 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
@@ -1790,7 +1782,7 @@
   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"
+  { fix x assume as:"x \<in> cball (0::'a) 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])
@@ -1826,18 +1818,19 @@
   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"
+    fix x::"'a" 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])
+    next 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 0" in ballE) defer 
+        apply(erule_tac x="basis 0" in ballE)
+        unfolding Ball_def mem_cball dist_norm using DIM_positive[where 'a='a]
+        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_zero 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"
+        fix e and x::"'a" 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
@@ -1864,9 +1857,9 @@
         ultimately show ?thesis using injpi by auto qed qed
   qed auto qed
 
-lemma homeomorphic_convex_compact_lemma: fixes s::"(real^'n) set"
+lemma homeomorphic_convex_compact_lemma: fixes s::"('a::euclidean_space) set"
   assumes "convex s" "compact s" "cball 0 1 \<subseteq> s"
-  shows "s homeomorphic (cball (0::real^'n) 1)"
+  shows "s homeomorphic (cball (0::'a) 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
@@ -1882,12 +1875,12 @@
       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) set"
+lemma homeomorphic_convex_compact_cball: fixes e::real and s::"('a::euclidean_space) set"
   assumes "convex s" "compact s" "interior s \<noteq> {}" "0 < e"
-  shows "s homeomorphic (cball (b::real^'n) e)"
+  shows "s homeomorphic (cball (b::'a) 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"
+  let ?d = "inverse d" and ?n = "0::'a"
   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
@@ -1899,7 +1892,7 @@
     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) set" and t::"(real^'n) set"
+lemma homeomorphic_convex_compact: fixes s::"('a::euclidean_space) set" and t::"('a) set"
   assumes "convex s" "compact s" "interior s \<noteq> {}"
           "convex t" "compact t" "interior t \<noteq> {}"
   shows "s homeomorphic t"
@@ -1951,7 +1944,7 @@
 subsection {* Convexity of general and special intervals. *}
 
 lemma is_interval_convex:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::ordered_euclidean_space) 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)"
@@ -1966,14 +1959,14 @@
     hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: field_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 by auto qed
+    using as(3-) DIM_positive[where 'a='a] by(auto simp add:euclidean_simps) qed
 
 lemma is_interval_connected:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::ordered_euclidean_space) 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}"
+lemma convex_interval: "convex {a .. b}" "convex {a<..<b::'a::ordered_euclidean_space}"
   apply(rule_tac[!] is_interval_convex) using is_interval_interval by auto
 
 (* FIXME: rewrite these lemmas without using vec1
@@ -2008,31 +2001,31 @@
 *)
 subsection {* Another intermediate value theorem formulation. *}
 
-lemma ivt_increasing_component_on_1: fixes f::"real \<Rightarrow> real^'n"
-  assumes "a \<le> 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"
+lemma ivt_increasing_component_on_1: fixes f::"real \<Rightarrow> 'a::ordered_euclidean_space"
+  assumes "a \<le> 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_le_def)
+    using assms(1) by auto
   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_real_interval(5)]]
     using assms by(auto intro!: imageI) qed
 
-lemma ivt_increasing_component_1: fixes f::"real \<Rightarrow> real^'n"
+lemma ivt_increasing_component_1: fixes f::"real \<Rightarrow> 'a::ordered_euclidean_space"
   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a .. b}. continuous (at x) f
-   \<Longrightarrow> f a$k \<le> y \<Longrightarrow> y \<le> f b$k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)$k = y"
+   \<Longrightarrow> f a$$k \<le> y \<Longrightarrow> y \<le> f b$$k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)$$k = y"
 by(rule ivt_increasing_component_on_1)
   (auto simp add: continuous_at_imp_continuous_on)
 
-lemma ivt_decreasing_component_on_1: fixes f::"real \<Rightarrow> real^'n"
-  assumes "a \<le> 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
+lemma ivt_decreasing_component_on_1: fixes f::"real \<Rightarrow> 'a::ordered_euclidean_space"
+  assumes "a \<le> 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])
+  using ivt_increasing_component_on_1[of a b "\<lambda>x. - f x" k "- y"] using assms using continuous_on_neg
+  by (auto simp add:euclidean_simps)
 
-lemma ivt_decreasing_component_1: fixes f::"real \<Rightarrow> real^'n"
+lemma ivt_decreasing_component_1: fixes f::"real \<Rightarrow> 'a::ordered_euclidean_space"
   shows "a \<le> b \<Longrightarrow> \<forall>x\<in>{a .. b}. continuous (at x) f
-    \<Longrightarrow> f b$k \<le> y \<Longrightarrow> y \<le> f a$k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)$k = y"
+    \<Longrightarrow> f b$$k \<le> y \<Longrightarrow> y \<le> f a$$k \<Longrightarrow> \<exists>x\<in>{a..b}. (f x)$$k = y"
 by(rule ivt_decreasing_component_on_1)
   (auto simp: continuous_at_imp_continuous_on)
 
@@ -2051,91 +2044,103 @@
     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 .. 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" 
+  "{0::'a::ordered_euclidean_space .. (\<chi>\<chi> i. 1)} = convex hull {x. \<forall>i<DIM('a). (x$$i = 0) \<or> (x$$i = 1)}"
+  (is "?int = convex hull ?points")
+proof- have 01:"{0,(\<chi>\<chi> i. 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::'a::ordered_euclidean_space .. \<chi>\<chi> i. 1}" "n \<le> DIM('a)" "card {i. i<DIM('a) \<and> 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
+    case 0 hence "x = 0" apply(subst euclidean_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
+    case (Suc n) show "x\<in>convex hull ?points" proof(cases "{i. i<DIM('a) \<and> x$$i \<noteq> 0} = {}")
+      case True hence "x = 0" apply(subst euclidean_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"
+      case False def xi \<equiv> "Min ((\<lambda>i. x$$i) ` {i. i<DIM('a) \<and> x$$i \<noteq> 0})"
+      have "xi \<in> (\<lambda>i. x$$i) ` {i. i<DIM('a) \<and> 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" "i<DIM('a)" by auto
+      have i:"\<And>j. j<DIM('a) \<Longrightarrow> 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
-      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_le_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_le_def elim!:ballE[where x=j]) qed
+      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
+      show ?thesis proof(cases "x$$i=1")
+        case True have "\<forall>j\<in>{i. i<DIM('a) \<and> x$$i \<noteq> 0}. x$$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq
+        proof(erule conjE) fix j assume as:"x $$ j \<noteq> 0" "x $$ j \<noteq> 1" "j<DIM('a)"
+          hence j:"x$$j \<in> {0<..<1}" using Suc(2) by(auto simp add: eucl_le[where 'a='a] elim!:allE[where x=j])
+          hence "x$$j \<in> op $$ x ` {i. i<DIM('a) \<and> x $$ i \<noteq> 0}" using as(3) 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: elim!:ballE[where x=j]) qed
         thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
           by auto
-      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: 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"
+      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>\<chi> j. if x$$j = 0 then 0 else 1) + (1 - x$$i) *\<^sub>R (\<chi>\<chi> j. ?y j)"
+          apply(subst euclidean_eq) by(auto simp add: field_simps euclidean_simps)
+        { fix j assume j:"j<DIM('a)"
+          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)
+            using Suc(2)[unfolded mem_interval, rule_format, of j] using j
+            by(auto simp add:field_simps euclidean_simps)
           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
-        hence "{j. x$j \<noteq> 0} \<noteq> {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" by auto
-        hence **:"{j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by auto
-        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
+        moreover have "i\<in>{j. j<DIM('a) \<and> x$$j \<noteq> 0} - {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}"
+          using i01 using i'(3) by auto
+        hence "{j. j<DIM('a) \<and> x$$j \<noteq> 0} \<noteq> {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0}" using i'(3) by blast
+        hence **:"{j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ j \<noteq> 0} \<subset> {j. j<DIM('a) \<and> x$$j \<noteq> 0}" apply - apply rule 
+          by( auto simp add:euclidean_simps)
+        have "card {j. j<DIM('a) \<and> ((\<chi>\<chi> j. ?y j)::'a) $$ 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
       qed qed qed } note * = this
+  have **:"DIM('a) = card {..<DIM('a)}" by auto
   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
+    apply(rule_tac n2="DIM('a)" in *) prefer 3 apply(subst(2) **) 
+    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_le_def mem_def[of _ convex]) qed
+    by(auto simp add: 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} = convex hull s"
-  apply(rule that[of "{x::real^'n. \<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) ` UNIV"])
+lemma unit_cube_convex_hull: obtains s where "finite s" "{0 .. (\<chi>\<chi> i. 1)::'a::ordered_euclidean_space} = convex hull s"
+  apply(rule that[of "{x::'a. \<forall>i<DIM('a). x$$i=0 \<or> x$$i=1}"])
+  apply(rule finite_subset[of _ "(\<lambda>s. (\<chi>\<chi> i. if i\<in>s then 1::real else 0)::'a) ` Pow {..<DIM('a)}"])
   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 qed auto
+  fix x::"'a" assume as:"\<forall>i<DIM('a). x $$ i = 0 \<or> x $$ i = 1"
+  show "x \<in> (\<lambda>s. \<chi>\<chi> i. if i \<in> s then 1 else 0) ` Pow {..<DIM('a)}"
+    apply(rule image_eqI[where x="{i. i<DIM('a) \<and> x$$i = 1}"])
+    using as apply(subst euclidean_eq) by auto qed auto
 
 subsection {* Hence any cube (could do any nonempty interval). *}
 
 lemma cube_convex_hull:
-  assumes "0 < d" obtains s::"(real^'n) 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)
+  assumes "0 < d" obtains s::"('a::ordered_euclidean_space) set" where
+  "finite s" "{x - (\<chi>\<chi> i. d) .. x + (\<chi>\<chi> i. d)} = convex hull s" proof-
+  let ?d = "(\<chi>\<chi> i. d)::'a"
+  have *:"{x - ?d .. x + ?d} = (\<lambda>y. x - ?d + (2 * d) *\<^sub>R y) ` {0 .. \<chi>\<chi> i. 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
-      hence "1 \<ge> inverse d * (x $ i - y $ i)" "1 \<ge> inverse d * (y $ i - x $ i)"
+    { fix i assume i:"i<DIM('a)" have "x $$ i \<le> d + y $$ i" "y $$ i \<le> d + x $$ i"
+        using as[unfolded mem_interval, THEN spec[where x=i]] i
+        by(auto simp add:euclidean_simps)
+      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)
-      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 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_le_def)
+      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..\<chi>\<chi> i.1}" unfolding mem_interval using assms
+      by(auto simp add: euclidean_simps field_simps)
+    thus "\<exists>z\<in>{0..\<chi>\<chi> i.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
   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)
+    fix y z assume as:"z\<in>{0..\<chi>\<chi> i.1}" "y = x - ?d + (2*d) *\<^sub>R z" 
+    have "\<And>i. i<DIM('a) \<Longrightarrow> 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: Cart_eq)
+      using assms by auto
     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: Cart_eq) qed
-  obtain s where "finite s" "{0..1::real^'n} = convex hull s" using unit_cube_convex_hull by auto
+      apply(erule_tac x=i in allE) using assms by(auto simp add: euclidean_simps) qed
+  obtain s where "finite s" "{0::'a..\<chi>\<chi> i.1} = 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. *}
@@ -2191,11 +2196,6 @@
 
 subsection {* Upper bound on a ball implies upper and lower bounds. *}
 
-lemma scaleR_2:
-  fixes x :: "'a::real_vector"
-  shows "scaleR 2 x = x + x"
-unfolding one_add_one_is_two [symmetric] scaleR_left_distrib by simp
-
 lemma convex_bounds_lemma:
   fixes x :: "'a::real_normed_vector"
   assumes "convex_on (cball x e) f"  "\<forall>y \<in> cball x e. f y \<le> b"
@@ -2214,26 +2214,28 @@
 subsection {* Hence a convex function on an open set is continuous. *}
 
 lemma convex_on_continuous:
-  assumes "open (s::(real^'n) set)" "convex_on s f" 
+  assumes "open (s::('a::ordered_euclidean_space) 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]
+  note dimge1 = DIM_positive[where 'a='a]
   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)"
+  def d \<equiv> "e / real DIM('a)"
   have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) 
-  let ?d = "(\<chi> i. d)::real^'n"
+  let ?d = "(\<chi>\<chi> i. d)::'a"
   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
+  have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:euclidean_simps)
   hence "c\<noteq>{}" using c by auto
   def k \<equiv> "Max (f ` c)"
+  have real_dimindex_ge_1:"real (CARD('n::finite)) \<ge> 1" 
+    by(metis dimindex_ge_1 real_eq_of_nat real_of_nat_1 real_of_nat_le_iff) 
   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 eq_divide_imp times_divide_eq_left real_dimindex_gt_0 real_eq_of_nat less_le mult_commute)
+    have e:"e = setsum (\<lambda>i. d) {..<DIM('a)}" unfolding setsum_constant d_def using dimge1
+      unfolding real_eq_of_nat by auto
     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 qed
+      using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:euclidean_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
@@ -2241,10 +2243,10 @@
   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  }
+    { fix i assume "i<DIM('a)" hence "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:euclidean_simps)  }
     thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm 
-      by auto qed
+      by(auto simp add:euclidean_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)
     apply force
@@ -2338,14 +2340,14 @@
   unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto
 
 lemma segment_furthest_le:
-  fixes a b x y :: "real ^ 'n"
+  fixes a b x y :: "'a::euclidean_space"
   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:
-  fixes x a b :: "real ^ 'n"
+  fixes x a b :: "'a::euclidean_space"
   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]
@@ -2357,7 +2359,7 @@
 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) \<longleftrightarrow> dist a b = (dist a x) + (dist x b)"
+lemma between:"between (a,b) (x::'a::euclidean_space) \<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
@@ -2369,27 +2371,31 @@
       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)
+        unfolding norm_minus_commute[of x a] * using as(2,3)
         by(auto simp add: 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 
+      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: 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)
-          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
+        unfolding dist_norm apply(subst euclidean_eq) apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4
+      proof(rule,rule) fix i assume i:"i<DIM('a)"
+          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: field_simps euclidean_simps)
+          also have "\<dots> = x$$i" apply(rule divide_eq_imp[OF Fal])
+            unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq] apply-
+            apply(subst (asm) euclidean_eq) using i apply(erule_tac x=i in allE) by(auto simp add:field_simps euclidean_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" shows
+lemma between_midpoint: fixes a::"'a::euclidean_space" 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) qed
+    unfolding euclidean_eq[where 'a='a]
+    by(auto simp add:field_simps euclidean_simps) qed
 
 lemma between_mem_convex_hull:
   "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
@@ -2398,7 +2404,7 @@
 subsection {* Shrinking towards the interior of a convex set. *}
 
 lemma mem_interior_convex_shrink:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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
@@ -2407,9 +2413,9 @@
     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: 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)
+      unfolding dist_norm unfolding norm_scaleR[THEN sym] apply(rule arg_cong[where f=norm]) using `e>0`
+      by(auto simp add: euclidean_simps euclidean_eq[where 'a='a] field_simps) 
+    also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:arg_cong[where f=norm] 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`] mult_commute)
     finally show "y \<in> s" apply(subst *) apply(rule assms(1)[unfolded convex_alt,rule_format])
@@ -2417,7 +2423,7 @@
   qed(rule mult_pos_pos, insert `e>0` `d>0`, auto) qed
 
 lemma mem_interior_closure_convex_shrink:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) 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
@@ -2453,76 +2459,96 @@
   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 . (\<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
+  "convex hull (insert 0 { basis i | i. i<DIM('a)}) =
+        {x::'a::euclidean_space . (\<forall>i<DIM('a). 0 \<le> x$$i) \<and> setsum (\<lambda>i. x$$i) {..<DIM('a)} \<le> 1 }"
+  (is "convex hull (insert 0 ?p) = ?s")
+proof- let ?D = "{..<DIM('a)}"
+  have *:"finite ?p" "0\<notin>?p" by auto
+  have "{(basis i)::'a |i. i<DIM('a)} = 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
+  show ?thesis unfolding simplex[OF *] 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
+    fix x::"'a" and u assume as: "\<forall>x\<in>{basis i |i. i<DIM('a)}. 0 \<le> u x"
+      "setsum u {basis i |i. i<DIM('a)} \<le> 1" "(\<Sum>x\<in>{basis i |i. i<DIM('a)}. u x *\<^sub>R x) = x"
+    have *:"\<forall>i<DIM('a). u (basis i) = x$$i" 
+    proof safe case goal1
+      have "x$$i = (\<Sum>j<DIM('a). (if j = i then u (basis j) else 0))"
+        unfolding as(3)[THEN sym] euclidean_component.setsum unfolding  sumbas
+        apply(rule setsum_cong2) by(auto simp add: basis_component)
+      also have "... = u (basis i)" apply(subst setsum_delta) using goal1 by auto
+      finally show ?case by auto
+    qed
+    hence **:"setsum u {basis i |i. i<DIM('a)} = setsum (op $$ x) ?D" unfolding sumbas 
+      apply-apply(rule setsum_cong2) by auto
+    show "(\<forall>i<DIM('a). 0 \<le> x $$ i) \<and> setsum (op $$ x) ?D \<le> 1" apply - proof(rule,rule,rule)
+      fix i assume i:"i<DIM('a)" show "0 \<le> x$$i"  unfolding *[rule_format,OF i,THEN sym]
+        apply(rule_tac as(1)[rule_format]) using i 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 
+  next fix x::"'a" assume as:"\<forall>i<DIM('a). 0 \<le> x $$ i" "setsum (op $$ x) ?D \<le> 1"
+    show "\<exists>u. (\<forall>x\<in>{basis i |i. i<DIM('a)}. 0 \<le> u x) \<and>
+      setsum u {basis i |i. i<DIM('a)} \<le> 1 \<and> (\<Sum>x\<in>{basis i |i. i<DIM('a)}. u x *\<^sub>R x) = x"
+      apply(rule_tac x="\<lambda>y. inner y x" in exI) apply safe using as(1)
+      proof- show "(\<Sum>y\<in>{basis i |i. i < DIM('a)}. y \<bullet> x) \<le> 1" unfolding sumbas
+          using as(2) unfolding euclidean_component_def[THEN sym] .
+        show "(\<Sum>xa\<in>{basis i |i. i < DIM('a)}. (xa \<bullet> x) *\<^sub>R xa) = x" unfolding sumbas
+          apply(subst (7) euclidean_representation) apply(rule setsum_cong2)
+          unfolding euclidean_component_def by auto
+      qed (auto simp add:euclidean_component_def)
+    qed qed
 
 lemma interior_std_simplex:
-  "interior (convex hull (insert 0 { basis i| i. i\<in>UNIV})) =
-  {x::real^'n. (\<forall>i. 0 < x$i) \<and> setsum (\<lambda>i. x$i) UNIV < 1 }"
+  "interior (convex hull (insert 0 { basis i| i. i<DIM('a)})) =
+  {x::'a::euclidean_space. (\<forall>i<DIM('a). 0 < x$$i) \<and> setsum (\<lambda>i. x$$i) {..<DIM('a)} < 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 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]
+  fix x::"'a" and e assume "0<e" and as:"\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x<DIM('a). 0 \<le> xa $$ x) \<and> setsum (op $$ xa) {..<DIM('a)} \<le> 1"
+  show "(\<forall>xa<DIM('a). 0 < x $$ xa) \<and> setsum (op $$ x) {..<DIM('a)} < 1" apply(safe) proof-
+    fix i assume i:"i<DIM('a)" thus "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: inner_simps euclidean_component_def dot_basis elim!:allE[where x=i])
+  next have **:"dist x (x + (e / 2) *\<^sub>R basis 0) < e" using  `e>0`
       unfolding dist_norm by(auto 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
-    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')
+    have "\<And>i. i<DIM('a) \<Longrightarrow> (x + (e / 2) *\<^sub>R basis 0) $$ i = x$$i + (if i = 0 then e/2 else 0)"
+      unfolding euclidean_component_def by(auto simp add:inner_simps dot_basis)
+    hence *:"setsum (op $$ (x + (e / 2) *\<^sub>R basis 0)) {..<DIM('a)} = setsum (\<lambda>i. x$$i + (if 0 = i then e/2 else 0)) {..<DIM('a)}"
+      apply(rule_tac setsum_cong) by auto
+    have "setsum (op $$ x) {..<DIM('a)} < setsum (op $$ (x + (e / 2) *\<^sub>R basis 0)) {..<DIM('a)}" unfolding * setsum_addf
+      using `0<e` DIM_positive[where 'a='a] apply(subst setsum_delta') by auto
     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" assume as:"\<forall>i. 0 < x $ i" "setsum (op $ x) UNIV < 1"
+    finally show "setsum (op $$ x) {..<DIM('a)} < 1" by auto qed
+next fix x::"'a" assume as:"\<forall>i<DIM('a). 0 < x $$ i" "setsum (op $$ x) {..<DIM('a)} < 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
+  let ?d = "(1 - setsum (op $$ x) {..<DIM('a)}) / real (DIM('a))"
+  have "Min ((op $$ x) ` {..<DIM('a)}) > 0" apply(rule Min_grI) using as(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]
+  ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i<DIM('a). 0 \<le> y $$ i) \<and> setsum (op $$ y) {..<DIM('a)} \<le> 1"
+    apply(rule_tac x="min (Min ((op $$ x) ` {..<DIM('a)})) ?D" in exI) apply rule defer apply(rule,rule) proof-
+    fix y assume y:"dist x y < min (Min (op $$ x ` {..<DIM('a)})) ?d"
+    have "setsum (op $$ y) {..<DIM('a)} \<le> setsum (\<lambda>i. x$$i + ?d) {..<DIM('a)}" proof(rule setsum_mono)
+      fix i assume "i\<in>{..<DIM('a)}" hence "abs (y$$i - x$$i) < ?d" apply-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: norm_minus_commute)
-      thus "y $ i \<le> x $ i + ?d" by auto qed
+      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]
-        by auto
-      thus "0 \<le> y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto
-    qed auto qed auto qed
+    finally show "(\<forall>i<DIM('a). 0 \<le> y $$ i) \<and> setsum (op $$ y) {..<DIM('a)} \<le> 1" 
+    proof safe fix i assume i:"i<DIM('a)"
+      have "norm (x - y) < x$$i" apply(rule less_le_trans) 
+        apply(rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) using i by auto
+      thus "0 \<le> y$$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto
+    qed qed auto qed
 
-lemma interior_std_simplex_nonempty: obtains a::"real^'n" 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]) }
+lemma interior_std_simplex_nonempty: obtains a::"'a::euclidean_space" where
+  "a \<in> interior(convex hull (insert 0 {basis i | i . i<DIM('a)}))" proof-
+  let ?D = "{..<DIM('a)}" let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) {(basis i) | i. i<DIM('a)}"
+  have *:"{basis i :: 'a | i. i<DIM('a)} = basis ` ?D" by auto
+  { fix i assume i:"i<DIM('a)" have "?a $$ i = inverse (2 * real DIM('a))"
+      unfolding euclidean_component.setsum * and setsum_reindex[OF basis_inj] and o_def
+      apply(rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) apply(rule setsum_cong2)
+      defer apply(subst setsum_delta') unfolding euclidean_component_def using i by(auto simp add:dot_basis) }
   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 **) 
+  show ?thesis apply(rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe
+    fix i assume i:"i<DIM('a)" show "0 < ?a $$ i" unfolding **[OF i] by(auto simp add: Suc_le_eq)
+  next have "setsum (op $$ ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D" apply(rule setsum_cong2, rule **) by auto
     also have "\<dots> < 1" unfolding setsum_constant card_enum real_eq_of_nat divide_inverse[THEN sym] by (auto simp add:field_simps)
-    finally show "setsum (op $ ?a) ?D < 1" by auto qed qed
+    finally show "setsum (op $$ ?a) ?D < 1" by auto qed qed
 
 end
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -6,10 +6,9 @@
 header {* Multivariate calculus in Euclidean space. *}
 
 theory Derivative
-imports Brouwer_Fixpoint Vec1 RealVector Operator_Norm
+imports Brouwer_Fixpoint Operator_Norm
 begin
 
-
 (* Because I do not want to type this all the time *)
 lemmas linear_linear = linear_conv_bounded_linear[THEN sym]
 
@@ -74,7 +73,7 @@
         (\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm(x' - x) \<and> norm(x' - x) < d
         \<longrightarrow> norm(f x' - f x - f'(x' - x)) / norm(x' - x) < e)"
   unfolding has_derivative_within Lim_within dist_norm
-  unfolding diff_0_right norm_mul by (simp add: diff_diff_eq)
+  unfolding diff_0_right by (simp add: diff_diff_eq)
 
 lemma has_derivative_at':
  "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -89,48 +88,10 @@
   "a \<in> s \<Longrightarrow> open s \<Longrightarrow> ((f has_derivative f') (at a within s) \<longleftrightarrow> (f has_derivative f') (at a))"
   unfolding has_derivative_within has_derivative_at using Lim_within_open by auto
 
-subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *}
-
-lemma has_derivative_within_vec1_dest_vec1: fixes f::"real\<Rightarrow>real" shows
-  "((vec1 \<circ> f \<circ> dest_vec1) has_derivative (vec1 \<circ> f' \<circ> dest_vec1)) (at (vec1 x) within vec1 ` s)
-  = (f has_derivative f') (at x within s)"
-  unfolding has_derivative_within unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear]
-  unfolding o_def Lim_within Ball_def unfolding forall_vec1 
-  unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto  
-
-lemma has_derivative_at_vec1_dest_vec1: fixes f::"real\<Rightarrow>real" shows
-  "((vec1 \<circ> f \<circ> dest_vec1) has_derivative (vec1 \<circ> f' \<circ> dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)"
-  using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV] by auto
-
-lemma bounded_linear_vec1: fixes f::"'a::real_normed_vector\<Rightarrow>real"
-  shows "bounded_linear f = bounded_linear (vec1 \<circ> f)"
-  unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def
-  unfolding vec1_dest_vec1_simps by auto
-
-lemma bounded_linear_dest_vec1: fixes f::"real\<Rightarrow>'a::real_normed_vector"
-  shows "bounded_linear f = bounded_linear (f \<circ> dest_vec1)"
-  unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def
-  unfolding vec1_dest_vec1_simps by auto
-
-lemma has_derivative_at_vec1: fixes f::"'a::real_normed_vector\<Rightarrow>real" shows
-  "(f has_derivative f') (at x) = ((vec1 \<circ> f) has_derivative (vec1 \<circ> f')) (at x)"
-  unfolding has_derivative_at unfolding bounded_linear_vec1[unfolded linear_conv_bounded_linear]
-  unfolding o_def Lim_at unfolding vec1_dest_vec1_simps dist_vec1_0 by auto
-
-lemma has_derivative_within_dest_vec1:fixes f::"real\<Rightarrow>'a::real_normed_vector" shows
-  "((f \<circ> dest_vec1) has_derivative (f' \<circ> dest_vec1)) (at (vec1 x) within vec1 ` s) = (f has_derivative f') (at x within s)"
-  unfolding has_derivative_within bounded_linear_dest_vec1 unfolding o_def Lim_within Ball_def
-  unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto
-
-lemma has_derivative_at_dest_vec1:fixes f::"real\<Rightarrow>'a::real_normed_vector" shows
-  "((f \<circ> dest_vec1) has_derivative (f' \<circ> dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)"
-  using has_derivative_within_dest_vec1[where s=UNIV] by(auto simp add:within_UNIV)
-
-lemma derivative_is_linear: fixes f::"real^'a \<Rightarrow> real^'b" shows
+lemma derivative_is_linear: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space" shows
   "(f has_derivative f') net \<Longrightarrow> linear f'"
   unfolding has_derivative_def and linear_conv_bounded_linear by auto
 
-
 subsection {* Combining theorems. *}
 
 lemma (in bounded_linear) has_derivative: "(f has_derivative f) net"
@@ -198,37 +159,47 @@
 
 subsection {* somewhat different results for derivative of scalar multiplier. *}
 
-lemma has_derivative_vmul_component: fixes c::"real^'a \<Rightarrow> real^'b" and v::"real^'c"
+(** move **)
+lemma linear_vmul_component:
+  assumes lf: "linear f"
+  shows "linear (\<lambda>x. f x $$ k *\<^sub>R v)"
+  using lf
+  by (auto simp add: linear_def algebra_simps)
+
+lemma has_derivative_vmul_component: fixes c::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space" and v::"'c::euclidean_space"
   assumes "(c has_derivative c') net"
-  shows "((\<lambda>x. c(x)$k *\<^sub>R v) has_derivative (\<lambda>x. (c' x)$k *\<^sub>R v)) net" proof-
-  have *:"\<And>y. (c y $ k *\<^sub>R v - (c (netlimit net) $ k *\<^sub>R v + c' (y - netlimit net) $ k *\<^sub>R v)) = 
-        (c y $ k - (c (netlimit net) $ k + c' (y - netlimit net) $ k)) *\<^sub>R v" 
+  shows "((\<lambda>x. c(x)$$k *\<^sub>R v) has_derivative (\<lambda>x. (c' x)$$k *\<^sub>R v)) net" proof-
+  have *:"\<And>y. (c y $$ k *\<^sub>R v - (c (netlimit net) $$ k *\<^sub>R v + c' (y - netlimit net) $$ k *\<^sub>R v)) = 
+        (c y $$ k - (c (netlimit net) $$ k + c' (y - netlimit net) $$ k)) *\<^sub>R v" 
     unfolding scaleR_left_diff_distrib scaleR_left_distrib by auto
   show ?thesis unfolding has_derivative_def and * and linear_conv_bounded_linear[symmetric]
-    apply(rule,rule linear_vmul_component[of c' k v, unfolded smult_conv_scaleR]) defer 
-    apply(subst vector_smult_lzero[THEN sym, of v]) unfolding scaleR_scaleR smult_conv_scaleR apply(rule Lim_vmul)
+    apply(rule,rule linear_vmul_component[of c' k v]) defer 
+    apply(subst scaleR_zero_left[THEN sym, of v]) unfolding scaleR_scaleR apply(rule Lim_vmul)
     using assms[unfolded has_derivative_def] unfolding Lim o_def apply- apply(cases "trivial_limit net")
     apply(rule,assumption,rule disjI2,rule,rule) proof-
-    have *:"\<And>x. x - vec 0 = (x::real^'n)" by auto 
-    have **:"\<And>d x. d * (c x $ k - (c (netlimit net) $ k + c' (x - netlimit net) $ k)) = (d *\<^sub>R (c x - (c (netlimit net) + c' (x - netlimit net) ))) $k" by(auto simp add:field_simps)
+    have *:"\<And>x. x - 0 = (x::'a)" by auto 
+    have **:"\<And>d x. d * (c x $$ k - (c (netlimit net) $$ k + c' (x - netlimit net) $$ k)) =
+      (d *\<^sub>R (c x - (c (netlimit net) + c' (x - netlimit net) ))) $$k" by(auto simp add:field_simps)
     fix e assume "\<not> trivial_limit net" "0 < (e::real)"
-    then have "eventually (\<lambda>x. dist ((1 / norm (x - netlimit net)) *\<^sub>R (c x - (c (netlimit net) + c' (x - netlimit net)))) 0 < e) net"
+    then have "eventually (\<lambda>x. dist ((1 / norm (x - netlimit net)) *\<^sub>R
+      (c x - (c (netlimit net) + c' (x - netlimit net)))) 0 < e) net"
       using assms[unfolded has_derivative_def Lim] by auto
-    thus "eventually (\<lambda>x. dist (1 / norm (x - netlimit net) * (c x $ k - (c (netlimit net) $ k + c' (x - netlimit net) $ k))) 0 < e) net"
+    thus "eventually (\<lambda>x. dist (1 / norm (x - netlimit net) *
+      (c x $$ k - (c (netlimit net) $$ k + c' (x - netlimit net) $$ k))) 0 < e) net"
       proof (rule eventually_elim1)
-      case goal1 thus ?case apply - unfolding dist_norm  apply(rule le_less_trans) prefer 2 apply assumption unfolding * ** and norm_vec1
-        using component_le_norm[of "(1 / norm (x - netlimit net)) *\<^sub>R (c x - (c (netlimit net) + c' (x - netlimit net))) - 0" k] by auto
-      qed
-      qed(insert assms[unfolded has_derivative_def], auto simp add:linear_conv_bounded_linear) qed 
+      case goal1 thus ?case apply - unfolding dist_norm  apply(rule le_less_trans)
+        prefer 2 apply assumption unfolding * **
+        using component_le_norm[of "(1 / norm (x - netlimit net)) *\<^sub>R
+          (c x - (c (netlimit net) + c' (x - netlimit net))) - 0" k] by auto
+    qed
+  qed(insert assms[unfolded has_derivative_def], auto simp add:linear_conv_bounded_linear) qed 
 
-lemma has_derivative_vmul_within: fixes c::"real \<Rightarrow> real" and v::"real^'a"
+lemma has_derivative_vmul_within: fixes c::"real \<Rightarrow> real" and v::"'a::euclidean_space"
   assumes "(c has_derivative c') (at x within s)"
-  shows "((\<lambda>x. (c x) *\<^sub>R v) has_derivative (\<lambda>x. (c' x) *\<^sub>R v)) (at x within s)" proof-
-  have *:"\<And>c. (\<lambda>x. (vec1 \<circ> c \<circ> dest_vec1) x $ 1 *\<^sub>R v) = (\<lambda>x. (c x) *\<^sub>R v) \<circ> dest_vec1" unfolding o_def by auto
-  show ?thesis using has_derivative_vmul_component[of "vec1 \<circ> c \<circ> dest_vec1" "vec1 \<circ> c' \<circ> dest_vec1" "at (vec1 x) within vec1 ` s" 1 v]
-  unfolding * and has_derivative_within_vec1_dest_vec1 unfolding has_derivative_within_dest_vec1 using assms by auto qed
+  shows "((\<lambda>x. (c x) *\<^sub>R v) has_derivative (\<lambda>x. (c' x) *\<^sub>R v)) (at x within s)"
+  using has_derivative_vmul_component[OF assms, of 0 v] by auto
 
-lemma has_derivative_vmul_at: fixes c::"real \<Rightarrow> real" and v::"real^'a"
+lemma has_derivative_vmul_at: fixes c::"real \<Rightarrow> real" and v::"'a::euclidean_space"
   assumes "(c has_derivative c') (at x)"
   shows "((\<lambda>x. (c x) *\<^sub>R v) has_derivative (\<lambda>x. (c' x) *\<^sub>R v)) (at x)"
   using has_derivative_vmul_within[where s=UNIV] and assms by(auto simp add: within_UNIV)
@@ -287,9 +258,6 @@
   "f differentiable (at a within s) \<longleftrightarrow> (f differentiable (at a))"
   unfolding differentiable_def has_derivative_within_open[OF assms] by auto
 
-lemma differentiable_at_imp_differentiable_on: "(\<forall>x\<in>(s::(real^'n) set). f differentiable at x) \<Longrightarrow> f differentiable_on s"
-  unfolding differentiable_on_def by(auto intro!: differentiable_at_withinI)
-
 lemma differentiable_on_eq_differentiable_at: "open s \<Longrightarrow> (f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable at x))"
   unfolding differentiable_on_def by(auto simp add: differentiable_within_open)
 
@@ -311,22 +279,16 @@
  "f differentiable net \<longleftrightarrow> (f has_derivative (frechet_derivative f net)) net"
   unfolding frechet_derivative_def differentiable_def and some_eq_ex[of "\<lambda> f' . (f has_derivative f') net"] ..
 
-lemma linear_frechet_derivative: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma linear_frechet_derivative: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "f differentiable net \<Longrightarrow> linear(frechet_derivative f net)"
   unfolding frechet_derivative_works has_derivative_def unfolding linear_conv_bounded_linear by auto
 
-definition "jacobian f net = matrix(frechet_derivative f net)"
-
-lemma jacobian_works: "(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow> (f has_derivative (\<lambda>h. (jacobian f net) *v h)) net"
-  apply rule unfolding jacobian_def apply(simp only: matrix_works[OF linear_frechet_derivative]) defer
-  apply(rule differentiableI) apply assumption unfolding frechet_derivative_works by assumption
-
 subsection {* Differentiability implies continuity. *}
 
 lemma Lim_mul_norm_within: fixes f::"'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
   shows "(f ---> 0) (at a within s) \<Longrightarrow> ((\<lambda>x. norm(x - a) *\<^sub>R f(x)) ---> 0) (at a within s)"
   unfolding Lim_within apply(rule,rule) apply(erule_tac x=e in allE,erule impE,assumption,erule exE,erule conjE)
-  apply(rule_tac x="min d 1" in exI) apply rule defer apply(rule,erule_tac x=x in ballE) unfolding dist_norm diff_0_right norm_mul
+  apply(rule_tac x="min d 1" in exI) apply rule defer apply(rule,erule_tac x=x in ballE) unfolding dist_norm diff_0_right
   by(auto intro!: mult_strict_mono[of _ "1::real", unfolded mult_1_left])
 
 lemma differentiable_imp_continuous_within: assumes "f differentiable (at x within s)" 
@@ -379,7 +341,7 @@
     show "norm (f y - f x - f' (y - x)) \<le> e * norm (y - x)" proof(cases "y=x")
       case True thus ?thesis using `bounded_linear f'` by(auto simp add: zero) next
       case False hence "norm (f y - (f x + f' (y - x))) < e * norm (y - x)" using as(4)[rule_format, OF `y\<in>s`]
-	unfolding dist_norm diff_0_right norm_mul using as(3)
+	unfolding dist_norm diff_0_right using as(3)
 	using pos_divide_less_eq[OF False[unfolded dist_nz], unfolded dist_norm]
 	by (auto simp add: linear_0 linear_sub)
       thus ?thesis by(auto simp add:algebra_simps) qed qed next
@@ -480,13 +442,13 @@
   unfolding differentiable_def apply(erule exE)+ apply(rule_tac x="\<lambda>z. f' z - f'a z" in exI)
     apply(rule has_derivative_sub) by auto 
 
-lemma differentiable_setsum: fixes f::"'a \<Rightarrow> (real^'n \<Rightarrow>real^'n)"
+lemma differentiable_setsum: fixes f::"'a \<Rightarrow> ('n::euclidean_space \<Rightarrow> 'n::euclidean_space)"
   assumes "finite s" "\<forall>a\<in>s. (f a) differentiable net"
   shows "(\<lambda>x. setsum (\<lambda>a. f a x) s) differentiable net" proof-
   guess f' using bchoice[OF assms(2)[unfolded differentiable_def]] ..
   thus ?thesis unfolding differentiable_def apply- apply(rule,rule has_derivative_setsum[where f'=f'],rule assms(1)) by auto qed
 
-lemma differentiable_setsum_numseg: fixes f::"_ \<Rightarrow> (real^'n \<Rightarrow>real^'n)"
+lemma differentiable_setsum_numseg: fixes f::"_ \<Rightarrow> ('n::euclidean_space \<Rightarrow> 'n::euclidean_space)"
   shows "\<forall>i. m \<le> i \<and> i \<le> n \<longrightarrow> (f i) differentiable net \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) {m::nat..n}) differentiable net"
   apply(rule differentiable_setsum) using finite_atLeastAtMost[of n m] by auto
 
@@ -504,135 +466,178 @@
 (* The general result is a bit messy because we need approachability of the  *)
 (* limit point from any direction. But OK for nontrivial intervals etc. *}
     
-lemma frechet_derivative_unique_within: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma frechet_derivative_unique_within: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "(f has_derivative f') (at x within s)" "(f has_derivative f'') (at x within s)"
-  "(\<forall>i::'a::finite. \<forall>e>0. \<exists>d. 0 < abs(d) \<and> abs(d) < e \<and> (x + d *\<^sub>R basis i) \<in> s)" shows "f' = f''" proof-
+  "(\<forall>i<DIM('a). \<forall>e>0. \<exists>d. 0 < abs(d) \<and> abs(d) < e \<and> (x + d *\<^sub>R basis i) \<in> s)" shows "f' = f''" proof-
   note as = assms(1,2)[unfolded has_derivative_def]
   then interpret f': bounded_linear f' by auto from as interpret f'': bounded_linear f'' by auto
   have "x islimpt s" unfolding islimpt_approachable proof(rule,rule)
-    guess a using UNIV_witness[where 'a='a] ..
-    fix e::real assume "0<e" guess d using assms(3)[rule_format,OF`e>0`,of a] ..
-    thus "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e" apply(rule_tac x="x + d*\<^sub>R basis a" in bexI)
-      using basis_nonzero[of a] norm_basis[of a] unfolding dist_norm by auto qed
+    fix e::real assume "0<e" guess d using assms(3)[rule_format,OF DIM_positive `e>0`] ..
+    thus "\<exists>x'\<in>s. x' \<noteq> x \<and> dist x' x < e" apply(rule_tac x="x + d *\<^sub>R basis 0" in bexI)
+      unfolding dist_norm by auto qed
   hence *:"netlimit (at x within s) = x" apply-apply(rule netlimit_within) unfolding trivial_limit_within by simp
   show ?thesis  apply(rule linear_eq_stdbasis) unfolding linear_conv_bounded_linear
-    apply(rule as(1,2)[THEN conjunct1])+ proof(rule,rule ccontr)
-    fix i::'a def e \<equiv> "norm (f' (basis i) - f'' (basis i))"
+    apply(rule as(1,2)[THEN conjunct1])+ proof(rule,rule,rule ccontr)
+    fix i assume i:"i<DIM('a)" def e \<equiv> "norm (f' (basis i) - f'' (basis i))"
     assume "f' (basis i) \<noteq> f'' (basis i)" hence "e>0" unfolding e_def by auto
     guess d using Lim_sub[OF as(1,2)[THEN conjunct2], unfolded * Lim_within,rule_format,OF `e>0`] .. note d=this
-    guess c using assms(3)[rule_format,OF d[THEN conjunct1],of i] .. note c=this
+    guess c using assms(3)[rule_format,OF i d[THEN conjunct1]] .. note c=this
     have *:"norm (- ((1 / \<bar>c\<bar>) *\<^sub>R f' (c *\<^sub>R basis i)) + (1 / \<bar>c\<bar>) *\<^sub>R f'' (c *\<^sub>R basis i)) = norm ((1 / abs c) *\<^sub>R (- (f' (c *\<^sub>R basis i)) + f'' (c *\<^sub>R basis i)))"
       unfolding scaleR_right_distrib by auto
     also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' (basis i)) + f'' (basis i))))"  
       unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto
-    also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by (auto simp add: add.commute ab_diff_minus)
-    finally show False using c using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] using norm_basis[of i] unfolding dist_norm 
-      unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib by auto qed qed
+    also have "\<dots> = e" unfolding e_def using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by (auto simp add: add.commute ab_diff_minus)
+    finally show False using c using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] 
+      unfolding dist_norm unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff
+        scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib using i by auto qed qed
 
-lemma frechet_derivative_unique_at: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma frechet_derivative_unique_at: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   shows "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f'') (at x) \<Longrightarrow> f' = f''"
   apply(rule frechet_derivative_unique_within[of f f' x UNIV f'']) unfolding within_UNIV apply(assumption)+
-  apply(rule,rule,rule) apply(rule_tac x="e/2" in exI) by auto
+  apply(rule,rule,rule,rule) apply(rule_tac x="e/2" in exI) by auto
  
 lemma "isCont f x = continuous (at x) f" unfolding isCont_def LIM_def
   unfolding continuous_at Lim_at unfolding dist_nz by auto
 
-lemma frechet_derivative_unique_within_closed_interval: fixes f::"real^'a \<Rightarrow> real^'b"
-  assumes "\<forall>i. a$i < b$i" "x \<in> {a..b}" (is "x\<in>?I") and
+lemma frechet_derivative_unique_within_closed_interval: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
+  assumes "\<forall>i<DIM('a). a$$i < b$$i" "x \<in> {a..b}" (is "x\<in>?I") and
   "(f has_derivative f' ) (at x within {a..b})" and
   "(f has_derivative f'') (at x within {a..b})"
-  shows "f' = f''" apply(rule frechet_derivative_unique_within) apply(rule assms(3,4))+ proof(rule,rule,rule)
-  fix e::real and i::'a assume "e>0"
-  thus "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R basis i \<in> {a..b}" proof(cases "x$i=a$i")
-    case True thus ?thesis apply(rule_tac x="(min (b$i - a$i)  e) / 2" in exI)
+  shows "f' = f''" apply(rule frechet_derivative_unique_within) apply(rule assms(3,4))+ proof(rule,rule,rule,rule)
+  fix e::real and i assume "e>0" and i:"i<DIM('a)"
+  thus "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R basis i \<in> {a..b}" proof(cases "x$$i=a$$i")
+    case True thus ?thesis apply(rule_tac x="(min (b$$i - a$$i)  e) / 2" in exI)
       using assms(1)[THEN spec[where x=i]] and `e>0` and assms(2)
-      unfolding mem_interval by(auto simp add:field_simps) next
-    note * = assms(2)[unfolded mem_interval,THEN spec[where x=i]]
-    case False moreover have "a $ i < x $ i" using False * by auto
-    moreover { have "a $ i * 2 + min (x $ i - a $ i) e \<le> a$i *2 + x$i - a$i" by auto
-    also have "\<dots> = a$i + x$i" by auto also have "\<dots> \<le> 2 * x$i" using * by auto 
-    finally have "a $ i * 2 + min (x $ i - a $ i) e \<le> x $ i * 2" by auto }
-    moreover have "min (x $ i - a $ i) e \<ge> 0" using * and `e>0` by auto
-    hence "x $ i * 2 \<le> b $ i * 2 + min (x $ i - a $ i) e" using * by auto
-    ultimately show ?thesis apply(rule_tac x="- (min (x$i - a$i) e) / 2" in exI)
+      unfolding mem_interval euclidean_simps basis_component using i by(auto simp add:field_simps)
+  next note * = assms(2)[unfolded mem_interval,THEN spec[where x=i]]
+    case False moreover have "a $$ i < x $$ i" using False * by auto
+    moreover { have "a $$ i * 2 + min (x $$ i - a $$ i) e \<le> a$$i *2 + x$$i - a$$i" by auto
+    also have "\<dots> = a$$i + x$$i" by auto also have "\<dots> \<le> 2 * x$$i" using * by auto 
+    finally have "a $$ i * 2 + min (x $$ i - a $$ i) e \<le> x $$ i * 2" by auto }
+    moreover have "min (x $$ i - a $$ i) e \<ge> 0" using * and `e>0` by auto
+    hence "x $$ i * 2 \<le> b $$ i * 2 + min (x $$ i - a $$ i) e" using * by auto
+    ultimately show ?thesis apply(rule_tac x="- (min (x$$i - a$$i) e) / 2" in exI)
       using assms(1)[THEN spec[where x=i]] and `e>0` and assms(2)
-      unfolding mem_interval by(auto simp add:field_simps) qed qed
+      unfolding mem_interval euclidean_simps basis_component using i by(auto simp add:field_simps) qed qed
 
-lemma frechet_derivative_unique_within_open_interval: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma frechet_derivative_unique_within_open_interval: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
   assumes "x \<in> {a<..<b}" "(f has_derivative f' ) (at x within {a<..<b})"
                          "(f has_derivative f'') (at x within {a<..<b})"
-  shows "f' = f''" apply(rule frechet_derivative_unique_within) apply(rule assms(2-3))+ proof(rule,rule,rule)
-  fix e::real and i::'a assume "e>0"
-  note * = assms(1)[unfolded mem_interval,THEN spec[where x=i]]
-  have "a $ i < x $ i" using  * by auto
-  moreover { have "a $ i * 2 + min (x $ i - a $ i) e \<le> a$i *2 + x$i - a$i" by auto
-  also have "\<dots> = a$i + x$i" by auto also have "\<dots> < 2 * x$i" using * by auto 
-  finally have "a $ i * 2 + min (x $ i - a $ i) e < x $ i * 2" by auto }
-  moreover have "min (x $ i - a $ i) e \<ge> 0" using * and `e>0` by auto
-  hence "x $ i * 2 < b $ i * 2 + min (x $ i - a $ i) e" using * by auto
+  shows "f' = f''" apply(rule frechet_derivative_unique_within) apply(rule assms(2-3))+ proof(rule,rule,rule,rule)
+  fix e::real and i assume "e>0" and i:"i<DIM('a)"
+  note * = assms(1)[unfolded mem_interval,rule_format,OF i]
+  have "a $$ i < x $$ i" using  * by auto
+  moreover { have "a $$ i * 2 + min (x $$ i - a $$ i) e \<le> a$$i *2 + x$$i - a$$i" by auto
+  also have "\<dots> = a$$i + x$$i" by auto also have "\<dots> < 2 * x$$i" using * by auto 
+  finally have "a $$ i * 2 + min (x $$ i - a $$ i) e < x $$ i * 2" by auto }
+  moreover have "min (x $$ i - a $$ i) e \<ge> 0" using * and `e>0` by auto
+  hence "x $$ i * 2 < b $$ i * 2 + min (x $$ i - a $$ i) e" using * by auto
   ultimately show "\<exists>d. 0 < \<bar>d\<bar> \<and> \<bar>d\<bar> < e \<and> x + d *\<^sub>R basis i \<in> {a<..<b}"
-    apply(rule_tac x="- (min (x$i - a$i) e) / 2" in exI)
-    using `e>0` and assms(1) unfolding mem_interval by(auto simp add:field_simps) qed
+    apply(rule_tac x="- (min (x$$i - a$$i) e) / 2" in exI)
+    using `e>0` and assms(1) unfolding mem_interval euclidean_simps apply safe unfolding basis_component
+    by(auto simp add:field_simps) qed
 
-lemma frechet_derivative_at: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma frechet_derivative_at: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
   shows "(f has_derivative f') (at x) \<Longrightarrow> (f' = frechet_derivative f (at x))"
   apply(rule frechet_derivative_unique_at[of f],assumption)
   unfolding frechet_derivative_works[THEN sym] using differentiable_def by auto
 
-lemma frechet_derivative_within_closed_interval: fixes f::"real^'a \<Rightarrow> real^'b"
-  assumes "\<forall>i. a$i < b$i" "x \<in> {a..b}" "(f has_derivative f') (at x within {a.. b})"
+lemma frechet_derivative_within_closed_interval: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
+  assumes "\<forall>i<DIM('a). a$$i < b$$i" "x \<in> {a..b}" "(f has_derivative f') (at x within {a.. b})"
   shows "frechet_derivative f (at x within {a.. b}) = f'"
   apply(rule frechet_derivative_unique_within_closed_interval[where f=f]) 
   apply(rule assms(1,2))+ unfolding frechet_derivative_works[THEN sym]
   unfolding differentiable_def using assms(3) by auto 
 
-subsection {* Component of the differential must be zero if it exists at a local        *)
-(* maximum or minimum for that corresponding component. *}
+subsection {* The traditional Rolle theorem in one dimension. *}
+
+lemma linear_componentwise:
+  fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes lf: "linear f"
+  shows "(f x) $$ j = (\<Sum>i<DIM('a). (x$$i) * (f (basis i)$$j))" (is "?lhs = ?rhs")
+proof -
+  have fA: "finite {..<DIM('a)}" by simp
+  have "?rhs = (\<Sum>i<DIM('a). x$$i *\<^sub>R f (basis i))$$j"
+    by (simp add: euclidean_simps)
+  then show ?thesis
+    unfolding linear_setsum_mul[OF lf fA, symmetric]
+    unfolding euclidean_representation[symmetric] ..
+qed
+
+text {* We do not introduce @{text jacobian}, which is defined on matrices, instead we use
+  the unfolding of it. *}
 
-lemma differential_zero_maxmin_component: fixes f::"real^'a \<Rightarrow> real^'b"
-  assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))"
-  "f differentiable (at x)" shows "jacobian f (at x) $ k = 0" proof(rule ccontr)
-  def D \<equiv> "jacobian f (at x)" assume "jacobian f (at x) $ k \<noteq> 0"
-  then obtain j where j:"D$k$j \<noteq> 0" unfolding Cart_eq D_def by auto
-  hence *:"abs (jacobian f (at x) $ k $ j) / 2 > 0" unfolding D_def by auto
-  note as = assms(3)[unfolded jacobian_works has_derivative_at_alt]
-  guess e' using as[THEN conjunct2,rule_format,OF *] .. note e' = this
-  guess d using real_lbound_gt_zero[OF assms(1) e'[THEN conjunct1]] .. note d = this
-  { fix c assume "abs c \<le> d" 
+lemma jacobian_works:
+  "(f::('a::euclidean_space) \<Rightarrow> ('b::euclidean_space)) differentiable net \<longleftrightarrow>
+   (f has_derivative (\<lambda>h. \<chi>\<chi> i.
+      \<Sum>j<DIM('a). frechet_derivative f net (basis j) $$ i * h $$ j)) net"
+  (is "?differentiable \<longleftrightarrow> (f has_derivative (\<lambda>h. \<chi>\<chi> i. ?SUM h i)) net")
+proof
+  assume *: ?differentiable
+  { fix h i
+    have "?SUM h i = frechet_derivative f net h $$ i" using *
+      by (auto intro!: setsum_cong
+               simp: linear_componentwise[of _ h i] linear_frechet_derivative) }
+  thus "(f has_derivative (\<lambda>h. \<chi>\<chi> i. ?SUM h i)) net"
+    using * by (simp add: frechet_derivative_works)
+qed (auto intro!: differentiableI)
+
+lemma differential_zero_maxmin_component:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes k: "k < DIM('b)"
+    and ball: "0 < e" "((\<forall>y \<in> ball x e. (f y)$$k \<le> (f x)$$k) \<or> (\<forall>y\<in>ball x e. (f x)$$k \<le> (f y)$$k))"
+    and diff: "f differentiable (at x)"
+  shows "(\<chi>\<chi> j. frechet_derivative f (at x) (basis j) $$ k) = (0::'a)" (is "?D k = 0")
+proof (rule ccontr)
+  assume "?D k \<noteq> 0"
+  then obtain j where j: "?D k $$ j \<noteq> 0" "j < DIM('a)"
+    unfolding euclidean_lambda_beta euclidean_eq[of _ "0::'a"] by auto
+  hence *: "\<bar>?D k $$ j\<bar> / 2 > 0" by auto
+  note as = diff[unfolded jacobian_works has_derivative_at_alt]
+  guess e' using as[THEN conjunct2, rule_format, OF *] .. note e' = this
+  guess d using real_lbound_gt_zero[OF ball(1) e'[THEN conjunct1]] .. note d = this
+  { fix c assume "abs c \<le> d"
     hence *:"norm (x + c *\<^sub>R basis j - x) < e'" using norm_basis[of j] d by auto
-    have "\<bar>(f (x + c *\<^sub>R basis j) - f x - D *v (c *\<^sub>R basis j)) $ k\<bar> \<le> norm (f (x + c *\<^sub>R basis j) - f x - D *v (c *\<^sub>R basis j))" by(rule component_le_norm)
-    also have "\<dots> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>" using e'[THEN conjunct2,rule_format,OF *] and norm_basis[of j] unfolding D_def[symmetric] by auto
-    finally have "\<bar>(f (x + c *\<^sub>R basis j) - f x - D *v (c *\<^sub>R basis j)) $ k\<bar> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>" by simp
-    hence "\<bar>f (x + c *\<^sub>R basis j) $ k - f x $ k - c * D $ k $ j\<bar> \<le> \<bar>D $ k $ j\<bar> / 2 * \<bar>c\<bar>"
-      unfolding vector_component_simps matrix_vector_mul_component unfolding smult_conv_scaleR[symmetric] 
-      unfolding inner_simps dot_basis smult_conv_scaleR by simp  } note * = this
+    let ?v = "(\<chi>\<chi> i. \<Sum>l<DIM('a). ?D i $$ l * (c *\<^sub>R basis j :: 'a) $$ l)"
+    have if_dist: "\<And> P a b c. a * (if P then b else c) = (if P then a * b else a * c)" by auto
+    have "\<bar>(f (x + c *\<^sub>R basis j) - f x - ?v) $$ k\<bar> \<le>
+        norm (f (x + c *\<^sub>R basis j) - f x - ?v)" by (rule component_le_norm)
+    also have "\<dots> \<le> \<bar>?D k $$ j\<bar> / 2 * \<bar>c\<bar>"
+      using e'[THEN conjunct2, rule_format, OF *] and norm_basis[of j] by fastsimp
+    finally have "\<bar>(f (x + c *\<^sub>R basis j) - f x - ?v) $$ k\<bar> \<le> \<bar>?D k $$ j\<bar> / 2 * \<bar>c\<bar>" by simp
+    hence "\<bar>f (x + c *\<^sub>R basis j) $$ k - f x $$ k - c * ?D k $$ j\<bar> \<le> \<bar>?D k $$ j\<bar> / 2 * \<bar>c\<bar>"
+      unfolding euclidean_simps euclidean_lambda_beta using j k
+      by (simp add: if_dist setsum_cases field_simps) } note * = this
   have "x + d *\<^sub>R basis j \<in> ball x e" "x - d *\<^sub>R basis j \<in> ball x e"
     unfolding mem_ball dist_norm using norm_basis[of j] d by auto
-  hence **:"((f (x - d *\<^sub>R basis j))$k \<le> (f x)$k \<and> (f (x + d *\<^sub>R basis j))$k \<le> (f x)$k) \<or>
-         ((f (x - d *\<^sub>R basis j))$k \<ge> (f x)$k \<and> (f (x + d *\<^sub>R basis j))$k \<ge> (f x)$k)" using assms(2) by auto
-  have ***:"\<And>y y1 y2 d dx::real. (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
-  show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\<bar>D $ k $ j\<bar> / 2 * \<bar>d\<bar>"]) 
+  hence **:"((f (x - d *\<^sub>R basis j))$$k \<le> (f x)$$k \<and> (f (x + d *\<^sub>R basis j))$$k \<le> (f x)$$k) \<or>
+         ((f (x - d *\<^sub>R basis j))$$k \<ge> (f x)$$k \<and> (f (x + d *\<^sub>R basis j))$$k \<ge> (f x)$$k)" using ball by auto
+  have ***: "\<And>y y1 y2 d dx::real.
+    (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
+  show False apply(rule ***[OF **, where dx="d * ?D k $$ j" and d="\<bar>?D k $$ j\<bar> / 2 * \<bar>d\<bar>"])
     using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left
     unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding algebra_simps by (auto intro: mult_pos_pos)
 qed
 
-subsection {* In particular if we have a mapping into @{typ "real^1"}. *}
+subsection {* In particular if we have a mapping into @{typ "real"}. *}
 
-lemma differential_zero_maxmin: fixes f::"real^'a \<Rightarrow> real"
-  assumes "x \<in> s" "open s" "(f has_derivative f') (at x)"
-  "(\<forall>y\<in>s. f y \<le> f x) \<or> (\<forall>y\<in>s. f x \<le> f y)"
-  shows "f' = (\<lambda>v. 0)" proof-
-  note deriv = assms(3)[unfolded has_derivative_at_vec1]
-  obtain e where e:"e>0" "ball x e \<subseteq> s" using assms(2)[unfolded open_contains_ball] and assms(1) by auto
-  hence **:"(jacobian (vec1 \<circ> f) (at x)) $ 1 = 0" using differential_zero_maxmin_component[of e x "\<lambda>x. vec1 (f x)" 1]
-    using assms(4) and assms(3)[unfolded has_derivative_at_vec1 o_def]
-    unfolding differentiable_def o_def by auto 
-  have *:"jacobian (vec1 \<circ> f) (at x) = matrix (vec1 \<circ> f')" unfolding jacobian_def and frechet_derivative_at[OF deriv] ..
-  have "vec1 \<circ> f' = (\<lambda>x. 0)" apply(rule) unfolding matrix_works[OF derivative_is_linear[OF deriv],THEN sym]
-    unfolding Cart_eq matrix_vector_mul_component using **[unfolded *] by auto
-  thus ?thesis apply-apply(rule,subst vec1_eq[THEN sym]) unfolding o_def apply(drule fun_cong) by auto qed
-
-subsection {* The traditional Rolle theorem in one dimension. *}
+lemma differential_zero_maxmin:
+  fixes f::"'a\<Colon>ordered_euclidean_space \<Rightarrow> real"
+  assumes "x \<in> s" "open s"
+  and deriv: "(f has_derivative f') (at x)"
+  and mono: "(\<forall>y\<in>s. f y \<le> f x) \<or> (\<forall>y\<in>s. f x \<le> f y)"
+  shows "f' = (\<lambda>v. 0)"
+proof -
+  obtain e where e:"e>0" "ball x e \<subseteq> s" using `open s`[unfolded open_contains_ball] and `x \<in> s` by auto
+  with differential_zero_maxmin_component[where 'b=real, of 0 e x f, simplified]
+  have "(\<chi>\<chi> j. frechet_derivative f (at x) (basis j)) = (0::'a)"
+    unfolding differentiable_def using mono deriv by auto
+  with frechet_derivative_at[OF deriv, symmetric]
+  have "\<forall>i<DIM('a). f' (basis i) = 0"
+    by (simp add: euclidean_eq[of _ "0::'a"])
+  with derivative_is_linear[OF deriv, THEN linear_componentwise, of _ 0]
+  show ?thesis by (simp add: expand_fun_eq)
+qed
 
 lemma rolle: fixes f::"real\<Rightarrow>real"
   assumes "a < b" "f a = f b" "continuous_on {a..b} f"
@@ -640,8 +645,8 @@
   shows "\<exists>x\<in>{a<..<b}. f' x = (\<lambda>v. 0)" proof-
   have "\<exists>x\<in>{a<..<b}. ((\<forall>y\<in>{a<..<b}. f x \<le> f y) \<or> (\<forall>y\<in>{a<..<b}. f y \<le> f x))" proof-
     have "(a + b) / 2 \<in> {a .. b}" using assms(1) by auto hence *:"{a .. b}\<noteq>{}" by auto
-    guess d using continuous_attains_sup[OF compact_real_interval * assms(3)] .. note d=this
-    guess c using continuous_attains_inf[OF compact_real_interval * assms(3)] .. note c=this
+    guess d using continuous_attains_sup[OF compact_interval * assms(3)] .. note d=this
+    guess c using continuous_attains_inf[OF compact_interval * assms(3)] .. note c=this
     show ?thesis proof(cases "d\<in>{a<..<b} \<or> c\<in>{a<..<b}")
       case True thus ?thesis apply(erule_tac disjE) apply(rule_tac x=d in bexI)
 	apply(rule_tac[3] x=c in bexI) using d c by auto next def e \<equiv> "(a + b) /2"
@@ -649,12 +654,12 @@
       hence "\<And>x. x\<in>{a..b} \<Longrightarrow> f x = f d" using c d apply- apply(erule_tac x=x in ballE)+ by auto
       thus ?thesis apply(rule_tac x=e in bexI) unfolding e_def using assms(1) by auto qed qed
   then guess x .. note x=this
-  hence "f' x \<circ> dest_vec1 = (\<lambda>v. 0)" apply(rule_tac differential_zero_maxmin[of "vec1 x" "vec1 ` {a<..<b}" "f \<circ> dest_vec1" "(f' x) \<circ> dest_vec1"]) 
-    unfolding vec1_interval defer apply(rule open_interval) 
-    apply(rule assms(4)[unfolded has_derivative_at_dest_vec1[THEN sym],THEN bspec[where x=x]],assumption)
-    unfolding o_def apply(erule disjE,rule disjI2) by(auto simp add: vector_less_def) 
-  thus ?thesis apply(rule_tac x=x in bexI) unfolding o_def apply rule 
-    apply(drule_tac x="vec1 v" in fun_cong) unfolding vec1_dest_vec1 using x(1) by auto qed
+  hence "f' x = (\<lambda>v. 0)" apply(rule_tac differential_zero_maxmin[of x "{a<..<b}" f "f' x"])
+    defer apply(rule open_interval)
+    apply(rule assms(4)[unfolded has_derivative_at[THEN sym],THEN bspec[where x=x]],assumption)
+    unfolding o_def apply(erule disjE,rule disjI2) by auto
+  thus ?thesis apply(rule_tac x=x in bexI) unfolding o_def apply rule
+    apply(drule_tac x=v in fun_cong) using x(1) by auto qed
 
 subsection {* One-dimensional mean value theorem. *}
 
@@ -676,7 +681,8 @@
   shows "\<exists>x\<in>{a<..<b}. f b - f a = f' x (b - a)"
   apply(rule mvt) apply(rule assms(1), rule differentiable_imp_continuous_on)
   unfolding differentiable_on_def differentiable_def defer proof 
-  fix x assume x:"x \<in> {a<..<b}" show "(f has_derivative f' x) (at x)" unfolding has_derivative_within_open[OF x open_interval_real,THEN sym] 
+  fix x assume x:"x \<in> {a<..<b}" show "(f has_derivative f' x) (at x)"
+    unfolding has_derivative_within_open[OF x open_interval,THEN sym] 
     apply(rule has_derivative_within_subset) apply(rule assms(2)[rule_format]) using x by auto qed(insert assms(2), auto)
 
 lemma mvt_very_simple: fixes f::"real \<Rightarrow> real"
@@ -690,7 +696,7 @@
 
 subsection {* A nice generalization (see Havin's proof of 5.19 from Rudin's book). *}
 
-lemma mvt_general: fixes f::"real\<Rightarrow>real^'n"
+lemma mvt_general: fixes f::"real\<Rightarrow>'a::euclidean_space"
   assumes "a<b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_derivative f'(x)) (at x)"
   shows "\<exists>x\<in>{a<..<b}. norm(f b - f a) \<le> norm(f'(x) (b - a))" proof-
   have "\<exists>x\<in>{a<..<b}. (op \<bullet> (f b - f a) \<circ> f) b - (op \<bullet> (f b - f a) \<circ> f) a = (f b - f a) \<bullet> f' x (b - a)"
@@ -708,7 +714,7 @@
 
 subsection {* Still more general bound theorem. *}
 
-lemma differentiable_bound: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma differentiable_bound: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "convex s" "\<forall>x\<in>s. (f has_derivative f'(x)) (at x within s)" "\<forall>x\<in>s. onorm(f' x) \<le> B" and x:"x\<in>s" and y:"y\<in>s"
   shows "norm(f x - f y) \<le> B * norm(x - y)" proof-
   let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)"
@@ -723,7 +729,7 @@
     have "(f \<circ> ?p has_derivative (f' ?u) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u within {0<..<1})" 
       apply(rule diff_chain_within) apply(rule has_derivative_intros)+ 
       apply(rule has_derivative_within_subset) apply(rule assms(2)[rule_format]) using goal1 * by auto
-    thus ?case unfolding has_derivative_within_open[OF goal1 open_interval_real] by auto qed
+    thus ?case unfolding has_derivative_within_open[OF goal1 open_interval] by auto qed
   guess u using mvt_general[OF zero_less_one 1 2] .. note u = this
   have **:"\<And>x y. x\<in>s \<Longrightarrow> norm (f' x y) \<le> B * norm y" proof- case goal1
     have "norm (f' x y) \<le> onorm (f' x) * norm y"
@@ -737,25 +743,12 @@
   also have "\<dots> \<le> B * norm(y - x)" apply(rule **) using * and u by auto
   finally show ?thesis by(auto simp add:norm_minus_commute) qed 
 
-lemma onorm_vec1: fixes f::"real \<Rightarrow> real"
-  shows "onorm (\<lambda>x. vec1 (f (dest_vec1 x))) = onorm f" proof-
-  have "\<forall>x::real^1. norm x = 1 \<longleftrightarrow> x\<in>{vec1 -1, vec1 (1::real)}" unfolding forall_vec1 by(auto simp add:Cart_eq)
-  hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto
-  have 2:"{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = (\<lambda>x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" by auto
-  have "\<forall>x::real. norm x = 1 \<longleftrightarrow> x\<in>{-1, 1}" by auto hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto
-  have 4:"{norm (f x) |x. norm x = 1} = (\<lambda>x. norm (f x)) ` {x. norm x=1}" by auto
-  show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max) qed
-
-lemma convex_vec1:"convex (vec1 ` s) = convex (s::real set)"
-  unfolding convex_def Ball_def forall_vec1 unfolding vec1_dest_vec1_simps image_iff by auto
-
 lemma differentiable_bound_real: fixes f::"real \<Rightarrow> real"
   assumes "convex s" "\<forall>x\<in>s. (f has_derivative f' x) (at x within s)" "\<forall>x\<in>s. onorm(f' x) \<le> B" and x:"x\<in>s" and y:"y\<in>s"
-  shows "norm(f x - f y) \<le> B * norm(x - y)" 
-  using differentiable_bound[of "vec1 ` s" "vec1 \<circ> f \<circ> dest_vec1" "\<lambda>x. vec1 \<circ> (f' (dest_vec1 x)) \<circ> dest_vec1" B "vec1 x" "vec1 y"]
-  unfolding Ball_def forall_vec1 unfolding has_derivative_within_vec1_dest_vec1 image_iff 
-  unfolding convex_vec1 unfolding o_def vec1_dest_vec1_simps onorm_vec1 using assms by auto
- 
+  shows "norm(f x - f y) \<le> B * norm(x - y)"
+  using differentiable_bound[of s f f' B x y]
+  unfolding Ball_def image_iff o_def using assms by auto
+
 subsection {* In particular. *}
 
 lemma has_derivative_zero_constant: fixes f::"real\<Rightarrow>real"
@@ -764,7 +757,7 @@
   case False then obtain x where "x\<in>s" by auto
   have "\<And>y. y\<in>s \<Longrightarrow> f x = f y" proof- case goal1
     thus ?case using differentiable_bound_real[OF assms(1-2), of 0 x y] and `x\<in>s`
-    unfolding onorm_vec1[of "\<lambda>x. 0", THEN sym] onorm_const norm_vec1 by auto qed
+    unfolding onorm_const by auto qed
   thus ?thesis apply(rule_tac x="f x" in exI) by auto qed auto
 
 lemma has_derivative_zero_unique: fixes f::"real\<Rightarrow>real"
@@ -773,7 +766,7 @@
 
 subsection {* Differentiability of inverse function (most basic form). *}
 
-lemma has_derivative_inverse_basic: fixes f::"real^'b \<Rightarrow> real^'c"
+lemma has_derivative_inverse_basic: fixes f::"'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
   assumes "(f has_derivative f') (at (g y))" "bounded_linear g'" "g' \<circ> f' = id" "continuous (at y) g"
   "open t" "y \<in> t" "\<forall>z\<in>t. f(g z) = z"
   shows "(g has_derivative g') (at y)" proof-
@@ -817,7 +810,7 @@
 
 subsection {* Simply rewrite that based on the domain point x. *}
 
-lemma has_derivative_inverse_basic_x: fixes f::"real^'b \<Rightarrow> real^'c"
+lemma has_derivative_inverse_basic_x: fixes f::"'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
   assumes "(f has_derivative f') (at x)" "bounded_linear g'" "g' o f' = id"
   "continuous (at (f x)) g" "g(f x) = x" "open t" "f x \<in> t" "\<forall>y\<in>t. f(g y) = y"
   shows "(g has_derivative g') (at (f(x)))"
@@ -825,7 +818,7 @@
 
 subsection {* This is the version in Dieudonne', assuming continuity of f and g. *}
 
-lemma has_derivative_inverse_dieudonne: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma has_derivative_inverse_dieudonne: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "open s" "open (f ` s)" "continuous_on s f" "continuous_on (f ` s) g" "\<forall>x\<in>s. g(f x) = x"
   (**) "x\<in>s" "(f has_derivative f') (at x)"  "bounded_linear g'" "g' o f' = id"
   shows "(g has_derivative g') (at (f x))"
@@ -834,7 +827,7 @@
 
 subsection {* Here's the simplest way of not assuming much about g. *}
 
-lemma has_derivative_inverse: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma has_derivative_inverse: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "compact s" "x \<in> s" "f x \<in> interior(f ` s)" "continuous_on s f"
   "\<forall>y\<in>s. g(f y) = y" "(f has_derivative f') (at x)" "bounded_linear g'" "g' \<circ> f' = id"
   shows "(g has_derivative g') (at (f x))" proof-
@@ -847,7 +840,7 @@
 
 subsection {* Proving surjectivity via Brouwer fixpoint theorem. *}
 
-lemma brouwer_surjective: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma brouwer_surjective: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'n"
   assumes "compact t" "convex t"  "t \<noteq> {}" "continuous_on t f"
   "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" "x\<in>s"
   shows "\<exists>y\<in>t. f y = x" proof-
@@ -855,7 +848,7 @@
   show ?thesis  unfolding * apply(rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"])
     apply(rule continuous_on_intros assms)+ using assms(4-6) by auto qed
 
-lemma brouwer_surjective_cball: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma brouwer_surjective_cball: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'n"
   assumes "0 < e" "continuous_on (cball a e) f"
   "\<forall>x\<in>s. \<forall>y\<in>cball a e. x + (y - f y) \<in> cball a e" "x\<in>s"
   shows "\<exists>y\<in>cball a e. f y = x" apply(rule brouwer_surjective) apply(rule compact_cball convex_cball)+
@@ -863,10 +856,10 @@
 
 text {* See Sussmann: "Multidifferential calculus", Theorem 2.1.1 *}
 
-lemma sussmann_open_mapping: fixes f::"real^'a \<Rightarrow> real^'b"
+lemma sussmann_open_mapping: fixes f::"'a::euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
   assumes "open s" "continuous_on s f" "x \<in> s" 
   "(f has_derivative f') (at x)" "bounded_linear g'" "f' \<circ> g' = id"
-  (**) "t \<subseteq> s" "x \<in> interior t"
+  "t \<subseteq> s" "x \<in> interior t"
   shows "f x \<in> interior (f ` t)" proof- 
   interpret f':bounded_linear f' using assms unfolding has_derivative_def by auto
   interpret g':bounded_linear g' using assms by auto
@@ -920,7 +913,24 @@
 (* We could put f' o g = I but this happens to fit with the minimal linear   *)
 (* algebra theory I've set up so far. *}
 
-lemma has_derivative_inverse_strong: fixes f::"real^'n \<Rightarrow> real^'n"
+(* move  before left_inverse_linear in Euclidean_Space*)
+
+ lemma right_inverse_linear: fixes f::"'a::euclidean_space => 'a"
+   assumes lf: "linear f" 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 expand_fun_eq)
+     by metis 
+   from linear_surjective_isomorphism[OF lf fi]
+   obtain h:: "'a => 'a" 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 expand_fun_eq)
+     by metis
+   with h(1) show ?thesis by blast
+ qed
+ 
+lemma has_derivative_inverse_strong: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'n"
   assumes "open s" "x \<in> s" "continuous_on s f"
   "\<forall>x\<in>s. g(f x) = x" "(f has_derivative f') (at x)" "f' o g' = id"
   shows "(g has_derivative g') (at (f x))" proof-
@@ -951,7 +961,7 @@
 
 subsection {* A rewrite based on the other domain. *}
 
-lemma has_derivative_inverse_strong_x: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma has_derivative_inverse_strong_x: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'a"
   assumes "open s" "g y \<in> s" "continuous_on s f"
   "\<forall>x\<in>s. g(f x) = x" "(f has_derivative f') (at (g y))" "f' o g' = id" "f(g y) = y"
   shows "(g has_derivative g') (at y)"
@@ -959,7 +969,7 @@
 
 subsection {* On a region. *}
 
-lemma has_derivative_inverse_on: fixes f::"real^'n \<Rightarrow> real^'n"
+lemma has_derivative_inverse_on: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'n"
   assumes "open s" "\<forall>x\<in>s. (f has_derivative f'(x)) (at x)" "\<forall>x\<in>s. g(f x) = x" "f'(x) o g'(x) = id" "x\<in>s"
   shows "(g has_derivative g'(x)) (at (f x))"
   apply(rule has_derivative_inverse_strong[where g'="g' x" and f=f]) apply(rule assms)+
@@ -974,14 +984,15 @@
 lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g ==> bounded_linear (\<lambda>x. f x - g x)"
   using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:algebra_simps)
 
-lemma has_derivative_locally_injective: fixes f::"real^'n \<Rightarrow> real^'m"
+lemma has_derivative_locally_injective: fixes f::"'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
   assumes "a \<in> s" "open s" "bounded_linear g'" "g' o f'(a) = id"
   "\<forall>x\<in>s. (f has_derivative f'(x)) (at x)"
   "\<forall>e>0. \<exists>d>0. \<forall>x. dist a x < d \<longrightarrow> onorm(\<lambda>v. f' x v - f' a v) < e"
   obtains t where "a \<in> t" "open t" "\<forall>x\<in>t. \<forall>x'\<in>t. (f x' = f x) \<longrightarrow> (x' = x)" proof-
   interpret bounded_linear g' using assms by auto
   note f'g' = assms(4)[unfolded id_def o_def,THEN cong]
-  have "g' (f' a 1) = 1" using f'g' by auto
+  have "g' (f' a (\<chi>\<chi> i.1)) = (\<chi>\<chi> i.1)" "(\<chi>\<chi> i.1) \<noteq> (0::'n)" defer 
+    apply(subst euclidean_eq) using f'g' by auto
   hence *:"0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)[unfolded linear_linear]] by fastsimp
   def k \<equiv> "1 / onorm g' / 2" have *:"k>0" unfolding k_def using * by auto
   guess d1 using assms(6)[rule_format,OF *] .. note d1=this
@@ -1018,7 +1029,7 @@
 
 subsection {* Uniformly convergent sequence of derivatives. *}
 
-lemma has_derivative_sequence_lipschitz_lemma: fixes f::"nat \<Rightarrow> real^'m \<Rightarrow> real^'n"
+lemma has_derivative_sequence_lipschitz_lemma: fixes f::"nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes "convex s" "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
   "\<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm(f' n x h - g' x h) \<le> e * norm(h)"
   shows "\<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm(x - y)" proof(default)+ 
@@ -1035,7 +1046,7 @@
     thus "onorm (\<lambda>h. f' m x h - f' n x h) \<le> 2 * e" apply-apply(rule onorm(2)) apply(rule linear_compose_sub)
       unfolding linear_conv_bounded_linear using assms(2)[rule_format,OF `x\<in>s`, THEN derivative_linear] by auto qed qed
 
-lemma has_derivative_sequence_lipschitz: fixes f::"nat \<Rightarrow> real^'m \<Rightarrow> real^'n"
+lemma has_derivative_sequence_lipschitz: fixes f::"nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes "convex s" "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
   "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm(f' n x h - g' x h) \<le> e * norm(h)" "0 < e"
   shows "\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm((f m x - f n x) - (f m y - f n y)) \<le> e * norm(x - y)" proof(rule,rule)
@@ -1043,7 +1054,7 @@
   guess N using assms(3)[rule_format,OF *(2)] ..
   thus ?case apply(rule_tac x=N in exI) apply(rule has_derivative_sequence_lipschitz_lemma[where e="1/2 *e", unfolded *]) using assms by auto qed
 
-lemma has_derivative_sequence: fixes f::"nat\<Rightarrow>real^'m\<Rightarrow>real^'n"
+lemma has_derivative_sequence: fixes f::"nat\<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes "convex s" "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
   "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm(f' n x h - g' x h) \<le> e * norm(h)"
   "x0 \<in> s"  "((\<lambda>n. f n x0) ---> l) sequentially"
@@ -1089,10 +1100,10 @@
 	  show ?case unfolding dist_norm using N[rule_format,OF goal1 `x\<in>s`, of u] False `e>0`
 	    by (auto simp add:field_simps) qed qed qed
     show "bounded_linear (g' x)" unfolding linear_linear linear_def apply(rule,rule,rule) defer proof(rule,rule)
-      fix x' y z::"real^'m" and c::real
+      fix x' y z::"'m" and c::real
       note lin = assms(2)[rule_format,OF `x\<in>s`,THEN derivative_linear]
       show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply(rule Lim_unique[OF trivial_limit_sequentially])
-	apply(rule lem3[rule_format]) unfolding smult_conv_scaleR 
+	apply(rule lem3[rule_format])
         unfolding lin[unfolded bounded_linear_def bounded_linear_axioms_def,THEN conjunct2,THEN conjunct1,rule_format]
 	apply(rule Lim_cmul) by(rule lem3[rule_format])
       show "g' x (y + z) = g' x y + g' x z" apply(rule Lim_unique[OF trivial_limit_sequentially])
@@ -1117,7 +1128,7 @@
 
 subsection {* Can choose to line up antiderivatives if we want. *}
 
-lemma has_antiderivative_sequence: fixes f::"nat\<Rightarrow> real^'m \<Rightarrow> real^'n"
+lemma has_antiderivative_sequence: fixes f::"nat\<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes "convex s" "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
   "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm(f' n x h - g' x h) \<le> e * norm h"
   shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g'(x)) (at x within s)" proof(cases "s={}")
@@ -1126,7 +1137,7 @@
     apply(rule,rule) apply(rule has_derivative_add_const, rule assms(2)[rule_format], assumption)  
     apply(rule `a\<in>s`) by(auto intro!: Lim_const) qed auto
 
-lemma has_antiderivative_limit: fixes g'::"real^'m \<Rightarrow> real^'m \<Rightarrow> real^'n"
+lemma has_antiderivative_limit: fixes g'::"'m::euclidean_space \<Rightarrow> 'm \<Rightarrow> 'n::euclidean_space"
   assumes "convex s" "\<forall>e>0. \<exists>f f'. \<forall>x\<in>s. (f has_derivative (f' x)) (at x within s) \<and> (\<forall>h. norm(f' x h - g' x h) \<le> e * norm(h))"
   shows "\<exists>g. \<forall>x\<in>s. (g has_derivative g'(x)) (at x within s)" proof-
   have *:"\<forall>n. \<exists>f f'. \<forall>x\<in>s. (f has_derivative (f' x)) (at x within s) \<and> (\<forall>h. norm(f' x h - g' x h) \<le> inverse (real (Suc n)) * norm(h))"
@@ -1145,7 +1156,7 @@
 definition sums_seq :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> (nat set) \<Rightarrow> bool"
 (infixl "sums'_seq" 12) where "(f sums_seq l) s \<equiv> ((\<lambda>n. setsum f (s \<inter> {0..n})) ---> l) sequentially"
 
-lemma has_derivative_series: fixes f::"nat \<Rightarrow> real^'m \<Rightarrow> real^'n"
+lemma has_derivative_series: fixes f::"nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space"
   assumes "convex s" "\<forall>n. \<forall>x\<in>s. ((f n) has_derivative (f' n x)) (at x within s)"
   "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm(setsum (\<lambda>i. f' i x h) (k \<inter> {0..n}) - g' x h) \<le> e * norm(h)"
   "x\<in>s" "((\<lambda>n. f n x) sums_seq l) k"
@@ -1156,7 +1167,8 @@
 
 subsection {* Derivative with composed bilinear function. *}
 
-lemma has_derivative_bilinear_within: fixes h::"real^'m \<Rightarrow> real^'n \<Rightarrow> real^'p" and f::"real^'q \<Rightarrow> real^'m"
+lemma has_derivative_bilinear_within: fixes h::"'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
+  and f::"'q::euclidean_space \<Rightarrow> 'm::euclidean_space"
   assumes "(f has_derivative f') (at x within s)" "(g has_derivative g') (at x within s)" "bounded_bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x within s)" proof-
   have "(g ---> g x) (at x within s)" apply(rule differentiable_imp_continuous_within[unfolded continuous_within])
@@ -1189,21 +1201,21 @@
       finally show "dist ((1 / norm (y - x)) *\<^sub>R h (f' (y - x)) (g' (y - x))) 0 < e"
 	unfolding dist_norm apply-apply(cases "y = x") by(auto simp add:field_simps) qed qed
   have "bounded_linear (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))" unfolding linear_linear linear_def
-    unfolding smult_conv_scaleR unfolding g'.add f'.scaleR f'.add g'.scaleR 
+    unfolding g'.add f'.scaleR f'.add g'.scaleR 
     unfolding h.add_right h.add_left scaleR_right_distrib h.scaleR_left h.scaleR_right by auto
   thus ?thesis using Lim_add[OF * **] unfolding has_derivative_within 
-    unfolding smult_conv_scaleR unfolding g'.add f'.scaleR f'.add g'.scaleR f'.diff g'.diff
+    unfolding g'.add f'.scaleR f'.add g'.scaleR f'.diff g'.diff
      h.add_right h.add_left scaleR_right_distrib h.scaleR_left h.scaleR_right h.diff_right h.diff_left
     scaleR_right_diff_distrib h.zero_right h.zero_left by(auto simp add:field_simps) qed
 
-lemma has_derivative_bilinear_at: fixes h::"real^'m \<Rightarrow> real^'n \<Rightarrow> real^'p" and f::"real^'p \<Rightarrow> real^'m"
+lemma has_derivative_bilinear_at: fixes h::"'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'p::euclidean_space" and f::"'p::euclidean_space \<Rightarrow> 'm::euclidean_space"
   assumes "(f has_derivative f') (at x)" "(g has_derivative g') (at x)" "bounded_bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x)"
   using has_derivative_bilinear_within[of f f' x UNIV g g' h] unfolding within_UNIV using assms by auto
 
-subsection {* Considering derivative @{typ "real^1 \<Rightarrow> real^'n"} as a vector. *}
+subsection {* Considering derivative @{typ "real \<Rightarrow> 'b\<Colon>real_normed_vector"} as a vector. *}
 
-definition has_vector_derivative :: "(real \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> ('b) \<Rightarrow> (real net \<Rightarrow> bool)"
+definition has_vector_derivative :: "(real \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> (real net \<Rightarrow> bool)"
 (infixl "has'_vector'_derivative" 12) where
  "(f has_vector_derivative f') net \<equiv> (f has_derivative (\<lambda>x. x *\<^sub>R f')) net"
 
@@ -1218,33 +1230,32 @@
     using f' unfolding scaleR[THEN sym] by auto
 next assume ?r thus ?l  unfolding vector_derivative_def has_vector_derivative_def differentiable_def by auto qed
 
-lemma vector_derivative_unique_at: fixes f::"real\<Rightarrow>real^'n"
+lemma vector_derivative_unique_at: fixes f::"real\<Rightarrow> 'n::euclidean_space"
   assumes "(f has_vector_derivative f') (at x)" "(f has_vector_derivative f'') (at x)" shows "f' = f''" proof-
-  have *:"(\<lambda>x. x *\<^sub>R f') \<circ> dest_vec1 = (\<lambda>x. x *\<^sub>R f'') \<circ> dest_vec1" apply(rule frechet_derivative_unique_at)
-    using assms[unfolded has_vector_derivative_def] unfolding has_derivative_at_dest_vec1[THEN sym] by auto
+  have *:"(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')" apply(rule frechet_derivative_unique_at)
+    using assms[unfolded has_vector_derivative_def] by auto
   show ?thesis proof(rule ccontr) assume "f' \<noteq> f''" moreover
-    hence "((\<lambda>x. x *\<^sub>R f') \<circ> dest_vec1) (vec1 1) = ((\<lambda>x. x *\<^sub>R f'') \<circ> dest_vec1) (vec1 1)" using * by auto
-    ultimately show False unfolding o_def vec1_dest_vec1 by auto qed qed
+    hence "(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')" using * by auto
+    ultimately show False unfolding expand_fun_eq by auto qed qed
 
-lemma vector_derivative_unique_within_closed_interval: fixes f::"real \<Rightarrow> real^'n"
+lemma vector_derivative_unique_within_closed_interval: fixes f::"real \<Rightarrow> 'n::ordered_euclidean_space"
   assumes "a < b" "x \<in> {a..b}"
   "(f has_vector_derivative f') (at x within {a..b})"
   "(f has_vector_derivative f'') (at x within {a..b})" shows "f' = f''" proof-
-  have *:"(\<lambda>x. x *\<^sub>R f') \<circ> dest_vec1 = (\<lambda>x. x *\<^sub>R f'') \<circ> dest_vec1"
-    apply(rule frechet_derivative_unique_within_closed_interval[of "vec1 a" "vec1 b"])
-    using assms(3-)[unfolded has_vector_derivative_def]
-    unfolding has_derivative_within_dest_vec1[THEN sym] vec1_interval using assms(1-2) by auto
+  have *:"(\<lambda>x. x *\<^sub>R f') = (\<lambda>x. x *\<^sub>R f'')"
+    apply(rule frechet_derivative_unique_within_closed_interval[of "a" "b"])
+    using assms(3-)[unfolded has_vector_derivative_def] using assms(1-2) by auto
   show ?thesis proof(rule ccontr) assume "f' \<noteq> f''" moreover
-    hence "((\<lambda>x. x *\<^sub>R f') \<circ> dest_vec1) (vec1 1) = ((\<lambda>x. x *\<^sub>R f'') \<circ> dest_vec1) (vec1 1)" using * by auto
-    ultimately show False unfolding o_def vec1_dest_vec1 by auto qed qed
+    hence "(\<lambda>x. x *\<^sub>R f') 1 = (\<lambda>x. x *\<^sub>R f'') 1" using * by (auto simp: expand_fun_eq)
+    ultimately show False unfolding o_def by auto qed qed
 
-lemma vector_derivative_at: fixes f::"real \<Rightarrow> real^'a" shows
+lemma vector_derivative_at: fixes f::"real \<Rightarrow> 'a::euclidean_space" shows
  "(f has_vector_derivative f') (at x) \<Longrightarrow> vector_derivative f (at x) = f'"
   apply(rule vector_derivative_unique_at) defer apply assumption
   unfolding vector_derivative_works[THEN sym] differentiable_def
   unfolding has_vector_derivative_def by auto
 
-lemma vector_derivative_within_closed_interval: fixes f::"real \<Rightarrow> real^'a"
+lemma vector_derivative_within_closed_interval: fixes f::"real \<Rightarrow> 'a::ordered_euclidean_space"
   assumes "a < b" "x \<in> {a..b}" "(f has_vector_derivative f') (at x within {a..b})"
   shows "vector_derivative f (at x within {a..b}) = f'"
   apply(rule vector_derivative_unique_within_closed_interval)
@@ -1286,16 +1297,15 @@
   using has_derivative_sub[OF assms[unfolded has_vector_derivative_def]]
   unfolding has_vector_derivative_def scaleR_right_diff_distrib by auto
 
-lemma has_vector_derivative_bilinear_within: fixes h::"real^'m \<Rightarrow> real^'n \<Rightarrow> real^'p"
+lemma has_vector_derivative_bilinear_within: fixes h::"'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
   assumes "(f has_vector_derivative f') (at x within s)" "(g has_vector_derivative g') (at x within s)" "bounded_bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) has_vector_derivative (h (f x) g' + h f' (g x))) (at x within s)" proof-
   interpret bounded_bilinear h using assms by auto 
-  show ?thesis using has_derivative_bilinear_within[OF assms(1-2)[unfolded has_vector_derivative_def has_derivative_within_dest_vec1[THEN sym]], where h=h]
-    unfolding o_def vec1_dest_vec1 has_vector_derivative_def
-    unfolding has_derivative_within_dest_vec1[unfolded o_def, where f="\<lambda>x. h (f x) (g x)" and f'="\<lambda>d. h (f x) (d *\<^sub>R g') + h (d *\<^sub>R f') (g x)"]
+  show ?thesis using has_derivative_bilinear_within[OF assms(1-2)[unfolded has_vector_derivative_def], of h]
+    unfolding o_def has_vector_derivative_def
     using assms(3) unfolding scaleR_right scaleR_left scaleR_right_distrib by auto qed
 
-lemma has_vector_derivative_bilinear_at: fixes h::"real^'m \<Rightarrow> real^'n \<Rightarrow> real^'p"
+lemma has_vector_derivative_bilinear_at: fixes h::"'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
   assumes "(f has_vector_derivative f') (at x)" "(g has_vector_derivative g') (at x)" "bounded_bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) has_vector_derivative (h (f x) g' + h f' (g x))) (at x)"
   apply(rule has_vector_derivative_bilinear_within[where s=UNIV, unfolded within_UNIV]) using assms by auto
--- a/src/HOL/Multivariate_Analysis/Determinants.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -5,7 +5,7 @@
 header {* Traces, Determinant of square matrices and some properties *}
 
 theory Determinants
-imports Euclidean_Space Permutations Vec1
+imports Euclidean_Space Permutations
 begin
 
 subsection{* First some facts about products*}
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -7,193 +7,13 @@
 theory Euclidean_Space
 imports
   Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
-  Finite_Cartesian_Product Infinite_Set Numeral_Type
+  Infinite_Set Numeral_Type
   Inner_Product L2_Norm Convex
 uses "positivstellensatz.ML" ("normarith.ML")
 begin
 
-subsection{* Basic componentwise operations on vectors. *}
-
-instantiation cart :: (times,finite) times
-begin
-  definition vector_mult_def : "op * \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))"
-  instance ..
-end
-
-instantiation cart :: (one,finite) one
-begin
-  definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"
-  instance ..
-end
-
-instantiation cart :: (ord,finite) ord
-begin
-  definition vector_le_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
-
-text{* The ordering on one-dimensional vectors is linear. *}
-
-class cart_one = assumes UNIV_one: "card (UNIV \<Colon> 'a set) = Suc 0"
-begin
-  subclass finite
-  proof from UNIV_one show "finite (UNIV :: 'a set)"
-      by (auto intro!: card_ge_0_finite) qed
-end
-
-instantiation cart :: (linorder,cart_one) linorder begin
-instance proof
-  guess a B using UNIV_one[where 'a='b] unfolding card_Suc_eq apply- by(erule exE)+
-  hence *:"UNIV = {a}" by auto
-  have "\<And>P. (\<forall>i\<in>UNIV. P i) \<longleftrightarrow> P a" unfolding * by auto hence all:"\<And>P. (\<forall>i. P i) \<longleftrightarrow> P a" by auto
-  fix x y z::"'a^'b::cart_one" note * = vector_le_def vector_less_def all Cart_eq
-  show "x\<le>x" "(x < y) = (x \<le> y \<and> \<not> y \<le> x)" "x\<le>y \<or> y\<le>x" unfolding * by(auto simp only:field_simps)
-  { assume "x\<le>y" "y\<le>z" thus "x\<le>z" unfolding * by(auto simp only:field_simps) }
-  { assume "x\<le>y" "y\<le>x" thus "x=y" unfolding * by(auto simp only:field_simps) }
-qed 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)"
-
-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 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 $ i = x"
-  by (vector vec_def)
-
-lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i"
-  by vector
-
-lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$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 cart :: (semigroup_mult,finite) semigroup_mult
-  apply (intro_classes) by (vector mult_assoc)
-
-instance cart :: (monoid_mult,finite) monoid_mult
-  apply (intro_classes) by vector+
-
-instance cart :: (ab_semigroup_mult,finite) ab_semigroup_mult
-  apply (intro_classes) by (vector mult_commute)
-
-instance cart :: (ab_semigroup_idem_mult,finite) ab_semigroup_idem_mult
-  apply (intro_classes) by (vector mult_idem)
-
-instance cart :: (comm_monoid_mult,finite) comm_monoid_mult
-  apply (intro_classes) by vector
-
-instance cart :: (semiring,finite) semiring
-  apply (intro_classes) by (vector field_simps)+
-
-instance cart :: (semiring_0,finite) semiring_0
-  apply (intro_classes) by (vector field_simps)+
-instance cart :: (semiring_1,finite) semiring_1
-  apply (intro_classes) by vector
-instance cart :: (comm_semiring,finite) comm_semiring
-  apply (intro_classes) by (vector field_simps)+
-
-instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes)
-instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
-instance cart :: (semiring_0_cancel,finite) semiring_0_cancel by (intro_classes)
-instance cart :: (comm_semiring_0_cancel,finite) comm_semiring_0_cancel by (intro_classes)
-instance cart :: (ring,finite) ring by (intro_classes)
-instance cart :: (semiring_1_cancel,finite) semiring_1_cancel by (intro_classes)
-instance cart :: (comm_semiring_1,finite) comm_semiring_1 by (intro_classes)
-
-instance cart :: (ring_1,finite) ring_1 ..
-
-instance cart :: (real_algebra,finite) real_algebra
-  apply intro_classes
-  apply (simp_all add: vector_scaleR_def field_simps)
-  apply vector
-  apply vector
-  done
-
-instance cart :: (real_algebra_1,finite) 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 one_index[simp]:
-  "(1 :: 'a::one ^'n)$i = 1" by vector
-
-instance cart :: (semiring_char_0,finite) 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 cart :: (comm_ring_1,finite) comm_ring_1 by intro_classes
-instance cart :: (ring_char_0,finite) 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 field_simps)
-lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
-  by (vector field_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 field_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 field_simps)
-
-lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
-  by (simp add: Cart_eq)
+lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
+  by auto
 
 abbreviation inner_bullet (infix "\<bullet>" 70)  where "x \<bullet> y \<equiv> inner x y"
 
@@ -299,22 +119,13 @@
 
 text{* Hence derive more interesting properties of the norm. *}
 
-lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x"
+(* FIXME: same as norm_scaleR
+lemma norm_mul[simp]: "norm(a *\<^sub>R x) = abs(a) * norm x"
   by (simp add: norm_vector_def setL2_right_distrib abs_mult)
+*)
 
 lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (inner x x = (0::real))"
-  by (simp add: norm_vector_def setL2_def power2_eq_square)
-lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" 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)
+  by (simp add: setL2_def power2_eq_square)
 
 lemma norm_cauchy_schwarz:
   shows "inner x y <= norm x * norm y"
@@ -329,20 +140,6 @@
   shows "norm x \<le> norm y  + norm (x - y)"
   using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps)
 
-lemma component_le_norm: "\<bar>x$i\<bar> <= norm x"
-  apply (simp add: norm_vector_def)
-  apply (rule member_le_setL2, simp_all)
-  done
-
-lemma norm_bound_component_le: "norm x <= e ==> \<bar>x$i\<bar> <= e"
-  by (metis component_le_norm order_trans)
-
-lemma norm_bound_component_lt: "norm x < e ==> \<bar>x$i\<bar> < e"
-  by (metis component_le_norm basic_trans_rules(21))
-
-lemma norm_le_l1: "norm x <= 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"
   by (rule abs_norm_cancel)
 lemma real_abs_sub_norm: "\<bar>norm x - norm y\<bar> <= norm(x - y)"
@@ -577,33 +374,17 @@
   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"
@@ -638,72 +419,12 @@
   using setsum_norm_le[OF fS K] setsum_constant[symmetric]
   by simp
 
-lemma setsum_vmul:
-  fixes f :: "'a \<Rightarrow> 'b::semiring_0"
-  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
-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"
-  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
+  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 dot_lsum: "finite S \<Longrightarrow> setsum f S \<bullet> y = setsum (\<lambda>x. f x \<bullet> y) S "
   apply(induct rule: finite_induct) by(auto simp add: inner_simps)
@@ -711,80 +432,6 @@
 lemma dot_rsum: "finite S \<Longrightarrow> y \<bullet> setsum f S = setsum (\<lambda>x. y \<bullet> f x) S "
   apply(induct rule: finite_induct) by(auto simp add: inner_simps)
 
-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) = 1"
-  apply (simp add: basis_def norm_eq_sqrt_inner) unfolding inner_vector_def
-  apply (vector delta_mult_idempotent)
-  using setsum_delta[of "UNIV :: 'n set" "k" "\<lambda>k. 1::real"] by auto
-
-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). 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). 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)"
-  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)" (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 smult_conv_scaleR: "c *s x = scaleR c x"
-  unfolding vector_scalar_mult_def vector_scaleR_def by simp
-
-lemma basis_expansion':
-  "setsum (\<lambda>i. (x$i) *\<^sub>R basis i) UNIV = x"
-  by (rule basis_expansion [where 'a=real, unfolded smult_conv_scaleR])
-
-lemma basis_expansion_unique:
-  "setsum (\<lambda>i. f i *s basis i) UNIV = (x::('a::comm_ring_1) ^'n) \<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) = (x$i)"
-  unfolding inner_vector_def by (auto simp add: 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"
-  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"
 proof
   assume "\<forall>x. x \<bullet> y = x \<bullet> z"
@@ -803,29 +450,26 @@
 
 subsection{* Orthogonality. *}
 
+context real_inner
+begin
+
 definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"
 
-lemma orthogonal_basis:
-  shows "orthogonal (basis i) x \<longleftrightarrow> x$i = (0::real)"
-  by (auto simp add: orthogonal_def inner_vector_def basis_def cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
-
-lemma orthogonal_basis_basis:
-  shows "orthogonal (basis i :: real^'n) (basis j) \<longleftrightarrow> i \<noteq> j"
-  unfolding orthogonal_basis[of i] basis_component[of j] by simp
-
 lemma orthogonal_clauses:
   "orthogonal a 0"
-  "orthogonal a x ==> orthogonal a (c *\<^sub>R 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 a x \<Longrightarrow> orthogonal a (c *\<^sub>R x)"
+  "orthogonal a x \<Longrightarrow> orthogonal a (-x)"
+  "orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x + y)"
+  "orthogonal a x \<Longrightarrow> orthogonal a y \<Longrightarrow> orthogonal a (x - y)"
   "orthogonal 0 a"
-  "orthogonal x a ==> orthogonal (c *\<^sub>R 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"
+  "orthogonal x a \<Longrightarrow> orthogonal (c *\<^sub>R x) a"
+  "orthogonal x a \<Longrightarrow> orthogonal (-x) a"
+  "orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x + y) a"
+  "orthogonal x a \<Longrightarrow> orthogonal y a \<Longrightarrow> orthogonal (x - y) a"
   unfolding orthogonal_def inner_simps by auto
 
+end
+
 lemma orthogonal_commute: "orthogonal x y \<longleftrightarrow> orthogonal y x"
   by (simp add: orthogonal_def inner_commute)
 
@@ -864,13 +508,7 @@
   apply (induct rule: finite_induct[OF fS])
   by (auto simp add: linear_zero intro: linear_compose_add)
 
-lemma linear_vmul_component:
-  assumes lf: "linear f"
-  shows "linear (\<lambda>x. f x $ k *\<^sub>R v)"
-  using lf
-  by (auto simp add: linear_def algebra_simps)
-
-lemma linear_0: "linear f ==> f 0 = 0"
+lemma linear_0: "linear f \<Longrightarrow> f 0 = 0"
   unfolding linear_def
   apply clarsimp
   apply (erule allE[where x="0::'a"])
@@ -919,93 +557,6 @@
   finally show ?thesis .
 qed
 
-lemma linear_bounded:
-  fixes f:: "real ^'m \<Rightarrow> real ^'n"
-  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) *\<^sub>R (basis i) :: real ^ 'm)"
-    have "norm (f x) = norm (f (setsum (\<lambda>i. (x$i) *\<^sub>R (basis i)) ?S))"
-      by (simp add: basis_expansion')
-    also have "\<dots> = norm (setsum (\<lambda>i. (x$i) *\<^sub>R 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) *\<^sub>R f (basis i))?S)" .
-    {fix i assume i: "i \<in> ?S"
-      from component_le_norm[of x i]
-      have "norm ((x$i) *\<^sub>R f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x"
-      unfolding norm_scaleR
-      apply (simp only: mult_commute)
-      apply (rule mult_mono)
-      by (auto simp add: field_simps) }
-    then have th: "\<forall>i\<in> ?S. norm ((x$i) *\<^sub>R f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis
-    from setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *\<^sub>R (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 \<Rightarrow> real ^'m"
-  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
-      with C have "B * norm (1:: real ^ 'n) < 0"
-        by (simp add: mult_less_0_iff)
-      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: field_simps split add: abs_split)
-      apply (erule order_trans, simp)
-      done
-  }
-  then show ?thesis using Kp by blast
-qed
-
-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
-
-lemma bounded_linearI': fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "\<And>x y. f (x + y) = f x + f y" "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"
-  shows "bounded_linear f" unfolding linear_conv_bounded_linear[THEN sym]
-  by(rule linearI[OF assms])
-
 subsection{* Bilinear functions. *}
 
 definition "bilinear f \<longleftrightarrow> (\<forall>x. linear(\<lambda>y. f x y)) \<and> (\<forall>y. linear(\<lambda>x. f x y))"
@@ -1060,89 +611,6 @@
   finally show ?thesis unfolding setsum_cartesian_product .
 qed
 
-lemma bilinear_bounded:
-  fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^'k"
-  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) *\<^sub>R basis i) ?M) (setsum (\<lambda>i. (y$i) *\<^sub>R basis i) ?N))" unfolding basis_expansion' ..
-    also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$i) *\<^sub>R basis i) ((y$j) *\<^sub>R 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 setsum_norm_le)
-      using fN fM
-      apply simp
-      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR)
-      apply (rule mult_mono)
-      apply (auto simp add: zero_le_mult_iff component_le_norm)
-      apply (rule mult_mono)
-      apply (auto simp add: zero_le_mult_iff component_le_norm)
-      done}
-  then show ?thesis by metis
-qed
-
-lemma bilinear_bounded_pos:
-  fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^'k"
-  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
-    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)"
@@ -1164,242 +632,6 @@
 
 lemma choice_iff: "(\<forall>x. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x. P x (f x))" by metis
 
-text {* TODO: The following lemmas about adjoints should hold for any
-Hilbert space (i.e. complete inner product space).
-(see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint})
-*}
-
-lemma adjoint_works_lemma:
-  fixes f:: "real ^'n \<Rightarrow> real ^'m"
-  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:: "real ^ 'm"
-    let ?w = "(\<chi> i. (f (basis i) \<bullet> y)) :: real ^ 'n"
-    {fix x
-      have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *\<^sub>R basis i) ?N) \<bullet> y"
-        by (simp only: basis_expansion')
-      also have "\<dots> = (setsum (\<lambda>i. (x$i) *\<^sub>R 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: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_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:: "real ^'n \<Rightarrow> real ^'m"
-  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:: "real ^'n \<Rightarrow> real ^'m"
-  assumes lf: "linear f"
-  shows "linear (adjoint f)"
-  unfolding linear_def vector_eq_ldot[where 'a="real^'n", symmetric] apply safe
-  unfolding inner_simps smult_conv_scaleR adjoint_works[OF lf] by auto
-
-lemma adjoint_clauses:
-  fixes f:: "real ^'n \<Rightarrow> real ^'m"
-  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] inner_commute)
-
-lemma adjoint_adjoint:
-  fixes f:: "real ^'n \<Rightarrow> real ^'m"
-  assumes lf: "linear f"
-  shows "adjoint (adjoint f) = f"
-  by (rule adjoint_unique, simp add: adjoint_clauses [OF lf])
-
-subsection {* Matrix operations *}
-
-text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *}
-
-definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
-  where "m ** m' == (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
-
-definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
-  where "m *v x \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
-
-definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
-  where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
-
-definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
-definition transpose where 
-  "(transpose::'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 ** B) + (A ** C)"
-  by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps)
-
-lemma matrix_mul_lid:
-  fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
-  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 ^ '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)"
-  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_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"
-  by (simp add: matrix_matrix_mult_def transpose_def Cart_eq mult_commute)
-
-lemma matrix_eq:
-  fixes A B :: "'a::semiring_1 ^ 'n ^ '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::real^_^_) *v x)$k = (A$k) \<bullet> x"
-  by (simp add: matrix_vector_mult_def inner_vector_def)
-
-lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
-  apply (simp add: inner_vector_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)
-  apply (subst setsum_commute)
-  by simp
-
-lemma transpose_mat: "transpose (mat n) = mat n"
-  by (vector transpose_def mat_def)
-
-lemma transpose_transpose: "transpose(transpose A) = A"
-  by (vector transpose_def)
-
-lemma row_transpose:
-  fixes A:: "'a::semiring_1^_^_"
-  shows "row i (transpose A) = column i A"
-  by (simp add: row_def column_def transpose_def Cart_eq)
-
-lemma column_transpose:
-  fixes A:: "'a::semiring_1^_^_"
-  shows "column i (transpose A) = row i A"
-  by (simp add: row_def column_def transpose_def Cart_eq)
-
-lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A"
-by (auto simp add: rows_def columns_def row_transpose intro: set_ext)
-
-lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" by (metis transpose_transpose rows_transpose)
-
-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 inner_vector_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) = (\<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:: "real ^'m \<Rightarrow> real ^ _"
-  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) *\<^sub>R f (basis i) ) ?M)$j"
-    unfolding vector_smult_component[symmetric] smult_conv_scaleR
-    unfolding setsum_component[of "(\<lambda>i.(x$i) *\<^sub>R f (basis i :: real^'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::real ^ _))"
-  by (simp add: linear_def matrix_vector_mult_def Cart_eq field_simps setsum_right_distrib setsum_addf)
-
-lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::real ^ 'n)"
-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::real ^ 'n))" by (simp add: ext matrix_works)
-
-lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: real ^ 'n)) = A"
-  by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
-
-lemma matrix_compose:
-  assumes lf: "linear (f::real^'n \<Rightarrow> real^'m)"
-  and lg: "linear (g::real^'m \<Rightarrow> real^_)"
-  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^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
-  by (simp add: matrix_vector_mult_def transpose_def Cart_eq mult_commute)
-
-lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
-  apply (rule adjoint_unique)
-  apply (simp add: transpose_def inner_vector_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 :: real^'n \<Rightarrow> real ^'m)"
-  shows "matrix(adjoint f) = transpose(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"
@@ -1433,53 +665,6 @@
   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
-
-lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
-
-lemma vec_add: "vec(x + y) = vec x + vec y"  by (vector vec_def)
-lemma vec_sub: "vec(x - y) = vec x - vec y" by (vector vec_def)
-lemma vec_cmul: "vec(c* x) = c *s vec x " by (vector vec_def)
-lemma vec_neg: "vec(- x) = - vec x " by (vector vec_def)
-
-lemma vec_setsum: assumes fS: "finite S"
-  shows "vec(setsum f S) = setsum (vec o f) S"
-  apply (induct rule: finite_induct[OF fS])
-  apply (simp)
-  apply (auto simp add: vec_add)
-  done
-
-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
-
 text {* TODO: move to NthRoot *}
 lemma sqrt_add_le_add_sqrt:
   assumes x: "0 \<le> x" and y: "0 \<le> y"
@@ -1693,24 +878,24 @@
 
 subsection{* A bit of linear algebra. *}
 
-definition
-  subspace :: "'a::real_vector set \<Rightarrow> bool" where
+definition (in real_vector)
+  subspace :: "'a set \<Rightarrow> bool" where
   "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 *\<^sub>R 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)"
+definition (in real_vector) "span S = (subspace hull S)"
+definition (in real_vector) "dependent S \<longleftrightarrow> (\<exists>a \<in> S. a \<in> span(S - {a}))"
+abbreviation (in real_vector) "independent s == ~(dependent s)"
 
 text {* 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"
+lemma (in real_vector) subspace_0: "subspace S ==> 0 \<in> S" by (metis subspace_def)
+
+lemma (in real_vector) 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 *\<^sub>R x \<in> S"
+lemma (in real_vector) subspace_mul: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> c *\<^sub>R x \<in> S"
   by (metis subspace_def)
 
 lemma subspace_neg: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> - x \<in> S"
@@ -1719,7 +904,7 @@
 lemma subspace_sub: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
   by (metis diff_def subspace_add subspace_neg)
 
-lemma subspace_setsum:
+lemma (in real_vector) 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"
@@ -1743,14 +928,13 @@
 lemma subspace_trivial: "subspace {0}"
   by (simp add: subspace_def)
 
-lemma subspace_inter: "subspace A \<Longrightarrow> subspace B ==> subspace (A \<inter> B)"
+lemma (in real_vector) 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"
+lemma (in real_vector) span_mono: "A \<subseteq> B ==> span A \<subseteq> span B"
   by (metis span_def hull_mono)
 
-lemma subspace_span: "subspace(span S)"
+lemma (in real_vector) 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)
@@ -1772,7 +956,7 @@
   apply simp
   done
 
-lemma span_clauses:
+lemma (in real_vector) 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"
@@ -1780,7 +964,7 @@
   by (metis span_def hull_subset subset_eq)
      (metis subspace_span subspace_def)+
 
-lemma span_induct: assumes SP: "\<And>x. x \<in> S ==> P x"
+lemma (in real_vector) 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)
@@ -1789,7 +973,7 @@
   show "P x" by (metis mem_def subset_eq)
 qed
 
-lemma span_empty: "span {} = {0}"
+lemma span_empty[simp]: "span {} = {0}"
   apply (simp add: span_def)
   apply (rule hull_unique)
   apply (auto simp add: mem_def subspace_def)
@@ -1797,10 +981,14 @@
   apply simp
   done
 
-lemma independent_empty: "independent {}"
+lemma (in real_vector) independent_empty[intro]: "independent {}"
   by (simp add: dependent_def)
 
-lemma independent_mono: "independent A \<Longrightarrow> B \<subseteq> A ==> independent B"
+lemma dependent_single[simp]:
+  "dependent {x} \<longleftrightarrow> x = 0"
+  unfolding dependent_def by auto
+
+lemma (in real_vector) 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
@@ -1808,14 +996,14 @@
   apply auto
   done
 
-lemma span_subspace: "A \<subseteq> B \<Longrightarrow> B \<le> span A \<Longrightarrow>  subspace B \<Longrightarrow> span A = B"
+lemma (in real_vector) 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"
+lemma (in real_vector) 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::real_vector \<Rightarrow> bool"
+inductive (in real_vector) span_induct_alt_help for S:: "'a \<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 *\<^sub>R x + z)"
@@ -1876,14 +1064,18 @@
 
 text {* Individual closure properties. *}
 
-lemma span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses(1))
-
-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"
+lemma (in real_vector) span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses(1))
+
+lemma (in real_vector) span_0: "0 \<in> span S" by (metis subspace_span subspace_0)
+
+lemma (in real_vector) dependent_0: assumes "0\<in>A" shows "dependent A"
+  unfolding dependent_def apply(rule_tac x=0 in bexI)
+  using assms span_0 by auto
+
+lemma (in real_vector) 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 *\<^sub>R x) \<in> span S"
+lemma (in real_vector) span_mul: "x \<in> span S ==> (c *\<^sub>R x) \<in> span S"
   by (metis subspace_span subspace_mul)
 
 lemma span_neg: "x \<in> span S ==> - x \<in> span S"
@@ -1892,7 +1084,7 @@
 lemma span_sub: "x \<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"
+lemma (in real_vector) span_setsum: "finite A \<Longrightarrow> \<forall>x \<in> A. f x \<in> span S ==> setsum f A \<in> span S"
   by (rule subspace_setsum, rule subspace_span)
 
 lemma span_add_eq: "x \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"
@@ -1990,7 +1182,7 @@
       done}
   moreover
   { fix k assume k: "x - k *\<^sub>R a \<in> span S"
-    have eq: "x = (x - k *\<^sub>R a) + k *\<^sub>R a" by vector
+    have eq: "x = (x - k *\<^sub>R a) + k *\<^sub>R a" by simp
     have "(x - k *\<^sub>R a) + k *\<^sub>R a \<in> span (insert a S)"
       apply (rule span_add)
       apply (rule set_rev_mp[of _ "span S" _])
@@ -2024,7 +1216,7 @@
     with na  have ?thesis by blast}
   moreover
   {assume k0: "k \<noteq> 0"
-    have eq: "b = (1/k) *\<^sub>R a - ((1/k) *\<^sub>R a - b)" by vector
+    have eq: "b = (1/k) *\<^sub>R a - ((1/k) *\<^sub>R a - b)" by simp
     from k0 have eq': "(1/k) *\<^sub>R (a - k*\<^sub>R b) = (1/k) *\<^sub>R a - b"
       by (simp add: algebra_simps)
     from k have "(1/k) *\<^sub>R (a - k*\<^sub>R b) \<in> span (S - {b})"
@@ -2066,7 +1258,7 @@
 proof-
   from span_breakdown[of x "insert x S" y, OF insertI1 y]
   obtain k where k: "y -k*\<^sub>R x \<in> span (S - {x})" by auto
-  have eq: "y = (y - k *\<^sub>R x) + k *\<^sub>R x" by vector
+  have eq: "y = (y - k *\<^sub>R x) + k *\<^sub>R x" by simp
   show ?thesis
     apply (subst eq)
     apply (rule span_add)
@@ -2078,6 +1270,9 @@
     by (rule x)
 qed
 
+lemma span_insert_0[simp]: "span (insert 0 S) = span S"
+  using span_mono[of S "insert 0 S"] by (auto intro: span_trans span_0)
+
 text {* An explicit expansion is sometimes needed. *}
 
 lemma span_explicit:
@@ -2153,7 +1348,7 @@
     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 *\<^sub>R v) ?S = 0"
       using fS aS
-      apply (simp add: vector_smult_lneg setsum_clauses field_simps)
+      apply (simp add: setsum_clauses field_simps)
       apply (subst (2) ua[symmetric])
       apply (rule setsum_cong2)
       by auto
@@ -2172,8 +1367,7 @@
     have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"       using fS SP vS by auto
     have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *\<^sub>R (u v *\<^sub>R v)) S - ?u v *\<^sub>R v"
       using fS vS uv
-      by (simp add: setsum_diff1 vector_smult_lneg divide_inverse
-        vector_smult_assoc field_simps)
+      by (simp add: setsum_diff1 divide_inverse field_simps)
     also have "\<dots> = ?a"
       unfolding scaleR_right.setsum [symmetric] u
       using uv by simp
@@ -2198,10 +1392,8 @@
     from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and
       u: "setsum (\<lambda>v. u v *\<^sub>R 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 *\<^sub>R v" S', symmetric] SS'
     have "setsum (\<lambda>v. ?u v *\<^sub>R v) S = setsum (\<lambda>v. u v *\<^sub>R v) S'"
-      unfolding cond_value_iff cond_application_beta
-      by (simp add: cond_value_iff inf_absorb2 cong del: if_weak_cong)
+      using SS' fS by (auto intro!: setsum_mono_zero_cong_right)
     hence "setsum (\<lambda>v. ?u v *\<^sub>R v) S = y" by (metis u)
     hence "y \<in> ?rhs" by auto}
   moreover
@@ -2211,71 +1403,6 @@
 qed
 
 
-text {* Standard bases are a spanning set, and obviously finite. *}
-
-lemma span_stdbasis:"span {basis i :: real^'n | 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 finite_stdbasis: "finite {basis i ::real^'n |i. i\<in> (UNIV:: 'n set)}" (is "finite ?S")
-proof-
-  have eq: "?S = basis ` UNIV" by blast
-  show ?thesis unfolding eq by auto
-qed
-
-lemma card_stdbasis: "card {basis i ::real^'n |i. i\<in> (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _")
-proof-
-  have eq: "?S = basis ` UNIV" by blast
-  show ?thesis unfolding eq using card_image[OF basis_inj] by simp
-qed
-
-
-lemma independent_stdbasis_lemma:
-  assumes x: "(x::real ^ '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::real^_). \<forall>i\<in> ?U. i \<notin> S \<longrightarrow> x$i =0"
- {fix x::"real^_" 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 |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
-
 text {* This is useful for building a basis step-by-step. *}
 
 lemma independent_insert:
@@ -2445,33 +1572,526 @@
     done
 qed
 
+subsection{* Euclidean Spaces as Typeclass*}
+
+class real_basis = real_vector +
+  fixes basis :: "nat \<Rightarrow> 'a"
+  assumes span_basis: "span (range basis) = UNIV"
+  assumes dimension_exists: "\<exists>d>0.
+    basis ` {d..} = {0} \<and>
+    independent (basis ` {..<d}) \<and>
+    inj_on basis {..<d}"
+
+definition (in real_basis) dimension :: "'a set \<Rightarrow> nat" where
+  "dimension x =
+    (THE d. basis ` {d..} = {0} \<and> independent (basis ` {..<d}) \<and> inj_on basis {..<d})"
+
+syntax "_type_dimension" :: "type => nat" ("(1DIM/(1'(_')))")
+
+translations "DIM('t)" => "CONST dimension (CONST UNIV \<Colon> 't set)"
+
+typed_print_translation {*
+let
+  fun card_univ_tr' show_sorts _ [Const (@{const_syntax UNIV}, Type(_, [T, _]))] =
+    Syntax.const @{syntax_const "_type_dimension"} $ Syntax.term_of_typ show_sorts T;
+in [(@{const_syntax dimension}, card_univ_tr')]
+end
+*}
+
+lemma (in real_basis) dimensionI:
+  assumes "\<And>d. \<lbrakk> 0 < d; basis ` {d..} = {0}; independent (basis ` {..<d});
+    inj_on basis {..<d} \<rbrakk> \<Longrightarrow> P d"
+  shows "P DIM('a)"
+proof -
+  obtain d where "0 < d" and d: "basis ` {d..} = {0} \<and>
+      independent (basis ` {..<d}) \<and> inj_on basis {..<d}" (is "?P d")
+    using dimension_exists by auto
+  show ?thesis unfolding dimension_def
+  proof (rule theI2)
+    fix d' assume "?P d'"
+    with d have "basis d' = 0" "basis d = 0" and
+      "d < d' \<Longrightarrow> basis d \<noteq> 0"
+      "d' < d \<Longrightarrow> basis d' \<noteq> 0"
+      using dependent_0 by auto
+    thus "d' = d" by (cases rule: linorder_cases) auto
+    moreover have "P d" using assms[of d] `0 < d` d by auto
+    ultimately show "P d'" by simp
+  next show "?P d" using `?P d` .
+  qed
+qed
+
+lemma (in real_basis) dimension_eq:
+  assumes "\<And>i. i < d \<Longrightarrow> basis i \<noteq> 0"
+  assumes "\<And>i. d \<le> i \<Longrightarrow> basis i = 0"
+  shows "DIM('a) = d"
+proof (rule dimensionI)
+  let ?b = "basis :: nat \<Rightarrow> 'a"
+  fix d' assume *: "?b ` {d'..} = {0}" "independent (?b ` {..<d'})"
+  show "d' = d"
+  proof (cases rule: linorder_cases)
+    assume "d' < d" hence "basis d' \<noteq> 0" using assms by auto
+    with * show ?thesis by auto
+  next
+    assume "d < d'" hence "basis d \<noteq> 0" using * dependent_0 by auto
+    with assms(2) `d < d'` show ?thesis by auto
+  qed
+qed
+
+lemma (in real_basis)
+  shows basis_finite: "basis ` {DIM('a)..} = {0}"
+  and independent_basis: "independent (basis ` {..<DIM('a)})"
+  and DIM_positive[intro]: "(DIM('a) :: nat) > 0"
+  and basis_inj[simp, intro]: "inj_on basis {..<DIM('a)}"
+  by (auto intro!: dimensionI)
+
+lemma (in real_basis) basis_eq_0_iff: "basis j = 0 \<longleftrightarrow> DIM('a) \<le> j"
+proof
+  show "DIM('a) \<le> j \<Longrightarrow> basis j = 0" using basis_finite by auto
+next
+  have "j < DIM('a) \<Longrightarrow> basis j \<noteq> 0"
+    using independent_basis by (auto intro!: dependent_0)
+  thus "basis j = 0 \<Longrightarrow> DIM('a) \<le> j" unfolding not_le[symmetric] by blast
+qed
+
+lemma (in real_basis) basis_range:
+    "range (basis) = {0} \<union> basis ` {..<DIM('a)}"
+  using basis_finite by (fastsimp simp: image_def)
+
+lemma (in real_basis) range_basis:
+    "range basis = insert 0 (basis ` {..<DIM('a)})"
+proof -
+  have *: "UNIV = {..<DIM('a)} \<union> {DIM('a)..}" by auto
+  show ?thesis unfolding * image_Un basis_finite by auto
+qed
+
+lemma (in real_basis) range_basis_finite[intro]:
+    "finite (range basis)"
+  unfolding range_basis by auto
+
+lemma (in real_basis) basis_neq_0[intro]:
+  assumes "i<DIM('a)" shows "(basis i) \<noteq> 0"
+proof(rule ccontr) assume "\<not> basis i \<noteq> (0::'a)"
+  hence "(0::'a) \<in> basis ` {..<DIM('a)}" using assms by auto
+  from dependent_0[OF this] show False using independent_basis by auto
+qed
+
+lemma (in real_basis) basis_zero[simp]:
+  assumes"i \<ge> DIM('a)" shows "basis i = 0"
+proof-
+  have "(basis i::'a) \<in> basis ` {DIM('a)..}" using assms by auto
+  thus ?thesis unfolding basis_finite by fastsimp
+qed
+
+lemma basis_representation:
+  "\<exists>u. x = (\<Sum>v\<in>basis ` {..<DIM('a)}. u v *\<^sub>R (v\<Colon>'a\<Colon>real_basis))"
+proof -
+  have "x\<in>UNIV" by auto from this[unfolded span_basis[THEN sym]]
+  have "\<exists>u. (\<Sum>v\<in>basis ` {..<DIM('a)}. u v *\<^sub>R v) = x"
+    unfolding range_basis span_insert_0 apply(subst (asm) span_finite) by auto
+  thus ?thesis by fastsimp
+qed
+
+class real_basis_with_inner = real_inner + real_basis
+begin
+
+definition euclidean_component (infixl "$$" 90) where
+  "x $$ i = inner (basis i) x"
+
+definition Chi (binder "\<chi>\<chi> " 10) where
+  "(\<chi>\<chi> i. f i) = (\<Sum>i<DIM('a). f i *\<^sub>R basis i)"
+
+lemma basis_at_neq_0[intro]:
+  "i < DIM('a) \<Longrightarrow> basis i $$ i \<noteq> 0"
+  unfolding euclidean_component_def by (auto intro!: basis_neq_0)
+
+lemma euclidean_component_ge[simp]:
+  assumes "i \<ge> DIM('a)" shows "x $$ i = 0"
+  unfolding euclidean_component_def basis_zero[OF assms] by auto
+
+lemma euclidean_scaleR:
+  shows "(a *\<^sub>R x) $$ i = a * (x$$i)"
+  unfolding euclidean_component_def by auto
+
+end
+
+interpretation euclidean_component: additive "\<lambda>x. euclidean_component x i"
+proof qed (simp add: euclidean_component_def inner_right.add)
+
+subsection{* Euclidean Spaces as Typeclass *}
+
+class euclidean_space = real_basis_with_inner +
+  assumes basis_orthonormal: "\<forall>i<DIM('a). \<forall>j<DIM('a).
+    inner (basis i) (basis j) = (if i = j then 1 else 0)"
+
+lemma (in euclidean_space) dot_basis:
+  "inner (basis i) (basis j) = (if i = j \<and> i<DIM('a) then 1 else 0)"
+proof (cases "(i < DIM('a) \<and> j < DIM('a))")
+  case False
+  hence "basis i = 0 \<or> basis j = 0"
+    using basis_finite by fastsimp
+  hence "inner (basis i) (basis j) = 0" by (rule disjE) simp_all
+  thus ?thesis using False by auto
+next
+  case True thus ?thesis using basis_orthonormal by auto
+qed
+
+lemma (in euclidean_space) basis_component[simp]:
+  "basis i $$ j = (if i = j \<and> i < DIM('a) then 1 else 0)"
+  unfolding euclidean_component_def dot_basis by auto
+
+lemmas euclidean_simps =
+  euclidean_component.add
+  euclidean_component.diff
+  euclidean_scaleR
+  euclidean_component.minus
+  euclidean_component.setsum
+  basis_component
+
+lemma euclidean_representation:
+  "(x\<Colon>'a\<Colon>euclidean_space) = (\<Sum>i\<in>{..<DIM('a)}. (x$$i) *\<^sub>R basis i)"
+proof-
+  from basis_representation[of x] guess u ..
+  hence *:"x = (\<Sum>i\<in>{..<DIM('a)}. u (basis i) *\<^sub>R (basis i))"
+    by (simp add: setsum_reindex)
+  show ?thesis apply(subst *)
+  proof(safe intro!: setsum_cong2)
+    fix i assume i: "i < DIM('a)"
+    hence "x$$i = (\<Sum>x<DIM('a). (if i = x then u (basis x) else 0))"
+      by (auto simp: euclidean_simps * intro!: setsum_cong)
+    also have "... = u (basis i)" using i by (auto simp: setsum_cases)
+    finally show "u (basis i) *\<^sub>R basis i = x $$ i *\<^sub>R basis i" by simp
+  qed
+qed
+
+lemma euclidean_eq:
+  fixes x y :: "'a\<Colon>euclidean_space"
+  shows "x = y \<longleftrightarrow> (\<forall>i<DIM('a). x$$i = y$$i)" (is "?l = ?r")
+proof safe
+  assume "\<forall>i<DIM('a). x $$ i = y $$ i"
+  thus "x = y"
+    by (subst (1 2) euclidean_representation) auto
+qed
+
+lemma euclidean_lambda_beta[simp]:
+  "((\<chi>\<chi> i. f i)::'a::euclidean_space) $$ j = (if j < DIM('a) then f j else 0)"
+  by (auto simp: euclidean_simps Chi_def if_distrib setsum_cases
+           intro!: setsum_cong)
+
+lemma euclidean_lambda_beta':
+  "j < DIM('a) \<Longrightarrow> ((\<chi>\<chi> i. f i)::'a::euclidean_space) $$ j = f j"
+  by simp
+
+lemma euclidean_lambda_beta'':"(\<forall>j < DIM('a::euclidean_space). P j (((\<chi>\<chi> i. f i)::'a) $$ j)) \<longleftrightarrow>
+  (\<forall>j < DIM('a::euclidean_space). P j (f j))" by auto
+
+lemma euclidean_beta_reduce[simp]:
+  "(\<chi>\<chi> i. x $$ i) = (x::'a::euclidean_space)"
+  by (subst euclidean_eq) (simp add: euclidean_lambda_beta)
+
+lemma euclidean_lambda_beta_0[simp]:
+    "((\<chi>\<chi> i. f i)::'a::euclidean_space) $$ 0 = f 0"
+  by (simp add: DIM_positive)
+
+lemma euclidean_inner:
+  "x \<bullet> (y::'a) = (\<Sum>i<DIM('a::euclidean_space). (x $$ i) \<bullet> (y $$ i))"
+proof -
+  have "x \<bullet> y = (\<Sum>i<DIM('a). x $$ i *\<^sub>R basis i) \<bullet>
+                (\<Sum>i<DIM('a). y $$ i *\<^sub>R (basis i :: 'a))"
+    by (subst (1 2) euclidean_representation) simp
+  also have "\<dots> = (\<Sum>i<DIM('a::euclidean_space). (x $$ i) \<bullet> (y $$ i))"
+    unfolding inner_left.setsum inner_right.setsum
+    by (auto simp add: dot_basis if_distrib setsum_cases intro!: setsum_cong)
+  finally show ?thesis .
+qed
+
+lemma norm_basis[simp]:"norm (basis i::'a::euclidean_space) = (if i<DIM('a) then 1 else 0)"
+  unfolding norm_eq_sqrt_inner dot_basis by auto
+
+lemma component_le_norm: "\<bar>x$$i\<bar> \<le> norm (x::'a::euclidean_space)"
+  unfolding euclidean_component_def
+  apply(rule order_trans[OF real_inner_class.Cauchy_Schwarz_ineq2]) by auto
+
+lemma norm_bound_component_le: "norm (x::'a::euclidean_space) \<le> e \<Longrightarrow> \<bar>x$$i\<bar> <= e"
+  by (metis component_le_norm order_trans)
+
+lemma norm_bound_component_lt: "norm (x::'a::euclidean_space) < e \<Longrightarrow> \<bar>x$$i\<bar> < e"
+  by (metis component_le_norm basic_trans_rules(21))
+
+lemma norm_le_l1: "norm (x::'a::euclidean_space) \<le> (\<Sum>i<DIM('a). \<bar>x $$ i\<bar>)"
+  apply (subst euclidean_representation[of x])
+  apply (rule order_trans[OF setsum_norm])
+  by (auto intro!: setsum_mono)
+
+lemma setsum_norm_allsubsets_bound:
+  fixes f:: "'a \<Rightarrow> 'n::euclidean_space"
+  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 DIM('n) *  e"
+proof-
+  let ?d = "real DIM('n)"
+  let ?nf = "\<lambda>x. norm (f x)"
+  let ?U = "{..<DIM('n)}"
+  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]
+      unfolding euclidean_component.setsum 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]
+      unfolding euclidean_component.setsum euclidean_component.minus
+      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 choice_iff': "(\<forall>x<d. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x<d. P x (f x))" by metis
+
+lemma lambda_skolem': "(\<forall>i<DIM('a::euclidean_space). \<exists>x. P i x) \<longleftrightarrow>
+   (\<exists>x::'a. \<forall>i<DIM('a). P i (x$$i))" (is "?lhs \<longleftrightarrow> ?rhs")
+proof-
+  let ?S = "{..<DIM('a)}"
+  {assume H: "?rhs"
+    then have ?lhs by auto}
+  moreover
+  {assume H: "?lhs"
+    then obtain f where f:"\<forall>i<DIM('a). P i (f i)" unfolding choice_iff' by metis
+    let ?x = "(\<chi>\<chi> i. (f i)) :: 'a"
+    {fix i assume i:"i<DIM('a)"
+      with f have "P i (f i)" by metis
+      then have "P i (?x$$i)" using i by auto
+    }
+    hence "\<forall>i<DIM('a). P i (?x$$i)" by metis
+    hence ?rhs by metis }
+  ultimately show ?thesis by metis
+qed
+
+subsection {* An ordering on euclidean spaces that will allow us to talk about intervals *}
+
+class ordered_euclidean_space = ord + euclidean_space +
+  assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i \<le> y $$ i)"
+  and eucl_less: "x < y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i < y $$ i)"
+
+subsection {* Linearity and Bilinearity continued *}
+
+lemma linear_bounded:
+  fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes lf: "linear f"
+  shows "\<exists>B. \<forall>x. norm (f x) \<le> B * norm x"
+proof-
+  let ?S = "{..<DIM('a)}"
+  let ?B = "setsum (\<lambda>i. norm(f(basis i))) ?S"
+  have fS: "finite ?S" by simp
+  {fix x:: "'a"
+    let ?g = "(\<lambda> i. (x$$i) *\<^sub>R (basis i) :: 'a)"
+    have "norm (f x) = norm (f (setsum (\<lambda>i. (x$$i) *\<^sub>R (basis i)) ?S))"
+      apply(subst euclidean_representation[of x]) ..
+    also have "\<dots> = norm (setsum (\<lambda> i. (x$$i) *\<^sub>R 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) *\<^sub>R f (basis i))?S)" .
+    {fix i assume i: "i \<in> ?S"
+      from component_le_norm[of x i]
+      have "norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \<le> norm (f (basis i)) * norm x"
+      unfolding norm_scaleR
+      apply (simp only: mult_commute)
+      apply (rule mult_mono)
+      by (auto simp add: field_simps) }
+    then have th: "\<forall>i\<in> ?S. norm ((x$$i) *\<^sub>R f (basis i :: 'a)) \<le> norm (f (basis i)) * norm x" by metis
+    from setsum_norm_le[OF fS, of "\<lambda>i. (x$$i) *\<^sub>R (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:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  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 "((\<chi>\<chi> i. 1)::'a) \<noteq> 0" unfolding euclidean_eq[where 'a='a]
+        by(auto intro!:exI[where x=0] simp add:euclidean_component.zero)
+      hence "norm ((\<chi>\<chi> i. 1)::'a) > 0" by auto
+      with C have "B * norm ((\<chi>\<chi> i. 1)::'a) < 0"
+        by (simp add: mult_less_0_iff)
+      with B[rule_format, of "(\<chi>\<chi> i. 1)::'a"] norm_ge_zero[of "f ((\<chi>\<chi> i. 1)::'a)"] have False by simp
+    }
+    then have Bp: "B \<ge> 0" by ferrack
+    {fix x::"'a"
+      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: field_simps split add: abs_split)
+      apply (erule order_trans, simp)
+      done
+  }
+  then show ?thesis using Kp by blast
+qed
+
+lemma linear_conv_bounded_linear:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  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
+  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"
+    by (simp add: f.add f.scaleR linear_def)
+qed
+
+lemma bounded_linearI': fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "\<And>x y. f (x + y) = f x + f y" "\<And>c x. f (c *\<^sub>R x) = c *\<^sub>R f x"
+  shows "bounded_linear f" unfolding linear_conv_bounded_linear[THEN sym]
+  by(rule linearI[OF assms])
+
+
+lemma bilinear_bounded:
+  fixes h:: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'k::euclidean_space"
+  assumes bh: "bilinear h"
+  shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
+proof-
+  let ?M = "{..<DIM('m)}"
+  let ?N = "{..<DIM('n)}"
+  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:: "'m" and  y :: "'n"
+    have "norm (h x y) = norm (h (setsum (\<lambda>i. (x$$i) *\<^sub>R basis i) ?M) (setsum (\<lambda>i. (y$$i) *\<^sub>R basis i) ?N))" 
+      apply(subst euclidean_representation[where 'a='m])
+      apply(subst euclidean_representation[where 'a='n]) ..
+    also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x$$i) *\<^sub>R basis i) ((y$$j) *\<^sub>R 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 setsum_norm_le)
+      using fN fM
+      apply simp
+      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps simp del: scaleR_scaleR)
+      apply (rule mult_mono)
+      apply (auto simp add: zero_le_mult_iff component_le_norm)
+      apply (rule mult_mono)
+      apply (auto simp add: zero_le_mult_iff component_le_norm)
+      done}
+  then show ?thesis by metis
+qed
+
+lemma bilinear_bounded_pos:
+  fixes h:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
+  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::'a and y::'b
+    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
+    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 :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::euclidean_space"
+  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
+  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
+  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 {* We continue. *}
+
+(** move **)
+lemma span_basis'[simp]:"span ((basis::nat=>'a) ` {..<DIM('a::real_basis)}) = UNIV"
+  apply(subst(2) span_basis[THEN sym]) unfolding basis_range by auto
+
+lemma card_basis[simp]:"card ((basis::nat=>'a) ` {..<DIM('a::real_basis)}) = DIM('a)"
+  apply(subst card_image) using basis_inj by auto
 
 lemma independent_bound:
-  fixes S:: "(real^'n) 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) set) ==> card S > CARD('n)) ==> dependent S"
+  fixes S:: "('a::euclidean_space) set"
+  shows "independent S \<Longrightarrow> finite S \<and> card S <= DIM('a::euclidean_space)"
+  using independent_span_bound[of "(basis::nat=>'a) ` {..<DIM('a)}" S] by auto
+
+lemma dependent_biggerset: "(finite (S::('a::euclidean_space) set) ==> card S > DIM('a)) ==> dependent S"
   by (metis independent_bound not_less)
 
 text {* Hence we can create a maximal independent subset. *}
 
 lemma maximal_independent_subset_extend:
-  assumes sv: "(S::(real^'n) set) \<subseteq> V" and iS: "independent S"
+  assumes sv: "(S::('a::euclidean_space) 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 "CARD('n) - card S" arbitrary: S rule: less_induct)
+proof(induct "DIM('a) - card S" arbitrary: S rule: less_induct)
   case less
   note sv = `S \<subseteq> V` and i = `independent 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)"
+  let ?d = "DIM('a)"
   {assume "V \<subseteq> span S"
     then have ?ths  using sv i by blast }
   moreover
@@ -2494,14 +2114,15 @@
 qed
 
 lemma maximal_independent_subset:
-  "\<exists>(B:: (real ^'n) 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)
+  "\<exists>(B:: ('a::euclidean_space) set). B\<subseteq> V \<and> independent B \<and> V \<subseteq> span B"
+  by (metis maximal_independent_subset_extend[of "{}:: ('a::euclidean_space) set"] empty_subsetI independent_empty)
+
 
 text {* Notion of dimension. *}
 
 definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = n))"
 
-lemma basis_exists:  "\<exists>B. (B :: (real ^'n) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = dim V)"
+lemma basis_exists:  "\<exists>B. (B :: ('a::euclidean_space) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (card B = 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> (card B = n)"]
 using maximal_independent_subset[of V] independent_bound
 by auto
@@ -2509,7 +2130,7 @@
 text {* Consequences of independence or spanning for cardinality. *}
 
 lemma independent_card_le_dim: 
-  assumes "(B::(real ^'n) set) \<subseteq> V" and "independent B" shows "card B \<le> dim V"
+  assumes "(B::('a::euclidean_space) set) \<subseteq> V" and "independent B" shows "card B \<le> dim V"
 proof -
   from basis_exists[of V] `B \<subseteq> V`
   obtain B' where "independent B'" and "B \<subseteq> span B'" and "card B' = dim V" by blast
@@ -2517,34 +2138,34 @@
   show ?thesis by auto
 qed
 
-lemma span_card_ge_dim:  "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> finite B \<Longrightarrow> dim V \<le> card B"
+lemma span_card_ge_dim:  "(B::('a::euclidean_space) 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 subset_trans)
 
 lemma basis_card_eq_dim:
-  "B \<subseteq> (V:: (real ^'n) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
+  "B \<subseteq> (V:: ('a::euclidean_space) 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_bound)
 
-lemma dim_unique: "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"
+lemma dim_unique: "(B::('a::euclidean_space) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"
   by (metis basis_card_eq_dim)
 
 text {* More lemmas about dimension. *}
 
-lemma dim_univ: "dim (UNIV :: (real^'n) set) = CARD('n)"
-  apply (rule dim_unique[of "{basis i |i. i\<in> (UNIV :: 'n set)}"])
-  by (auto simp only: span_stdbasis card_stdbasis finite_stdbasis independent_stdbasis)
+lemma dim_UNIV: "dim (UNIV :: ('a::euclidean_space) set) = DIM('a)"
+  apply (rule dim_unique[of "(basis::nat=>'a) ` {..<DIM('a)}"])
+  using independent_basis by auto
 
 lemma dim_subset:
-  "(S:: (real ^'n) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
+  "(S:: ('a::euclidean_space) set) \<subseteq> T \<Longrightarrow> dim S \<le> dim T"
   using basis_exists[of T] basis_exists[of S]
   by (metis independent_card_le_dim subset_trans)
 
-lemma dim_subset_univ: "dim (S:: (real^'n) set) \<le> CARD('n)"
-  by (metis dim_subset subset_UNIV dim_univ)
+lemma dim_subset_UNIV: "dim (S:: ('a::euclidean_space) set) \<le> DIM('a)"
+  by (metis dim_subset subset_UNIV dim_UNIV)
 
 text {* Converses to those. *}
 
 lemma card_ge_dim_independent:
-  assumes BV:"(B::(real ^'n) set) \<subseteq> V" and iB:"independent B" and dVB:"dim V \<le> card B"
+  assumes BV:"(B::('a::euclidean_space) 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"
@@ -2558,7 +2179,7 @@
 qed
 
 lemma card_le_dim_spanning:
-  assumes BV: "(B:: (real ^'n) set) \<subseteq> V" and VB: "V \<subseteq> span B"
+  assumes BV: "(B:: ('a::euclidean_space) set) \<subseteq> V" and VB: "V \<subseteq> span B"
   and fB: "finite B" and dVB: "dim V \<ge> card B"
   shows "independent B"
 proof-
@@ -2579,20 +2200,20 @@
   then show ?thesis unfolding dependent_def by blast
 qed
 
-lemma card_eq_dim: "(B:: (real ^'n) set) \<subseteq> V \<Longrightarrow> card B = dim V \<Longrightarrow> finite B \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
+lemma card_eq_dim: "(B:: ('a::euclidean_space) set) \<subseteq> V \<Longrightarrow> card B = dim V \<Longrightarrow> finite B \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
   by (metis order_eq_iff card_le_dim_spanning
     card_ge_dim_independent)
 
 text {* More general size bound lemmas. *}
 
 lemma independent_bound_general:
-  "independent (S:: (real^'n) set) \<Longrightarrow> finite S \<and> card S \<le> dim S"
+  "independent (S:: ('a::euclidean_space) 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) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
+lemma dependent_biggerset_general: "(finite (S:: ('a::euclidean_space) 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) set)) = dim S"
+lemma dim_span: "dim (span (S:: ('a::euclidean_space) set)) = dim S"
 proof-
   have th0: "dim S \<le> dim (span S)"
     by (auto simp add: subset_eq intro: dim_subset span_superset)
@@ -2605,10 +2226,10 @@
     using fB(2)  by arith
 qed
 
-lemma subset_le_dim: "(S:: (real ^'n) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
+lemma subset_le_dim: "(S:: ('a::euclidean_space) set) \<subseteq> span T \<Longrightarrow> dim S \<le> dim T"
   by (metis dim_span dim_subset)
 
-lemma span_eq_dim: "span (S:: (real ^'n) set) = span T ==> dim S = dim T"
+lemma span_eq_dim: "span (S:: ('a::euclidean_space) set) = span T ==> dim S = dim T"
   by (metis dim_span)
 
 lemma spans_image:
@@ -2618,8 +2239,8 @@
   by (metis VB image_mono)
 
 lemma dim_image_le:
-  fixes f :: "real^'n \<Rightarrow> real^'m"
-  assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n) set)"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S)"
 proof-
   from basis_exists[of S] obtain B where
     B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "card B = dim S" by blast
@@ -2662,11 +2283,11 @@
     (* 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::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
-  unfolding inner_simps smult_conv_scaleR by auto
+lemma vector_sub_project_orthogonal: "(b::'a::euclidean_space) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *\<^sub>R b) = 0"
+  unfolding inner_simps by auto
 
 lemma basis_orthogonal:
-  fixes B :: "(real ^'n) set"
+  fixes B :: "('a::euclidean_space) 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")
@@ -2683,7 +2304,7 @@
   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: field_simps)
+    have th0: "\<And>(a::'a) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps)
     have "x - k *\<^sub>R (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x)) \<in> span C \<longleftrightarrow> x - k *\<^sub>R a \<in> span C"
       apply (simp only: scaleR_right_diff_distrib th0)
       apply (rule span_add_eq)
@@ -2708,7 +2329,7 @@
         apply simp
         apply (subst Cy)
         using C(1) fth
-        apply (simp only: setsum_clauses) unfolding smult_conv_scaleR
+        apply (simp only: setsum_clauses)
         apply (auto simp add: inner_simps inner_commute[of y a] dot_lsum[OF fth])
         apply (rule setsum_0')
         apply clarsimp
@@ -2724,7 +2345,7 @@
         apply simp
         apply (subst Cx)
         using C(1) fth
-        apply (simp only: setsum_clauses) unfolding smult_conv_scaleR
+        apply (simp only: setsum_clauses)
         apply (subst inner_commute[of x])
         apply (auto simp add: inner_simps inner_commute[of x a] dot_rsum[OF fth])
         apply (rule setsum_0')
@@ -2741,7 +2362,7 @@
 qed
 
 lemma orthogonal_basis_exists:
-  fixes V :: "(real ^'n) set"
+  fixes V :: "('a::euclidean_space) set"
   shows "\<exists>B. independent B \<and> B \<subseteq> span V \<and> V \<subseteq> span B \<and> (card B = 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" "card B = dim V" by blast
@@ -2766,9 +2387,9 @@
 
 text {* Low-dimensional subset is in a hyperplane (weak orthogonal complement). *}
 
-lemma span_not_univ_orthogonal:
+lemma span_not_univ_orthogonal: fixes S::"('a::euclidean_space) set"
   assumes sU: "span S \<noteq> UNIV"
-  shows "\<exists>(a:: real ^'n). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"
+  shows "\<exists>(a::'a). 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
@@ -2787,8 +2408,7 @@
   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 inner_simps smult_conv_scaleR)
-  
+    show "subspace (\<lambda>x. ?a \<bullet> x = 0)" by (auto simp add: subspace_def mem_def inner_simps)
 next
     {fix x assume x: "x \<in> B"
       from x have B': "B = insert x (B - {x})" by blast
@@ -2796,8 +2416,8 @@
       have "?a \<bullet> x = 0"
         apply (subst B') using fB fth
         unfolding setsum_clauses(2)[OF fth]
-        apply simp unfolding inner_simps smult_conv_scaleR
-        apply (clarsimp simp add: inner_simps smult_conv_scaleR dot_lsum)
+        apply simp unfolding inner_simps
+        apply (clarsimp simp add: inner_simps dot_lsum)
         apply (rule setsum_0', rule ballI)
         unfolding inner_commute
         by (auto simp add: x field_simps intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
@@ -2807,17 +2427,17 @@
 qed
 
 lemma span_not_univ_subset_hyperplane:
-  assumes SU: "span S \<noteq> (UNIV ::(real^'n) set)"
+  assumes SU: "span S \<noteq> (UNIV ::('a::euclidean_space) 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). a  \<noteq> 0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
+lemma lowdim_subset_hyperplane: fixes S::"('a::euclidean_space) set"
+  assumes d: "dim S < DIM('a)"
+  shows "\<exists>(a::'a). 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)
+    hence "dim (span S) = dim (UNIV :: ('a) set)" by simp
+    hence "dim S = DIM('a)" 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 .
@@ -2853,7 +2473,7 @@
   hence "f x - k*\<^sub>R f a \<in> span (f ` b)"
     by (simp add: linear_sub[OF lf] linear_cmul[OF lf])
   hence th: "-k *\<^sub>R f a \<in> span (f ` b)"
-    using "2.prems"(5) by (simp add: vector_smult_lneg)
+    using "2.prems"(5) by simp
   {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]
@@ -2862,7 +2482,7 @@
   {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)
+      by auto
     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"]
@@ -2907,7 +2527,7 @@
       have khz: "(k - ?h z) *\<^sub>R 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)
+        have "a \<in> span b" by simp
         with "2.prems"(1) "2.hyps"(2) have False
           by (auto simp add: dependent_def)}
       then have "k = ?h z" by blast}
@@ -2965,7 +2585,7 @@
 qed
 
 lemma linear_independent_extend:
-  assumes iB: "independent (B:: (real ^'n) set)"
+  assumes iB: "independent (B:: ('a::euclidean_space) 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
@@ -3018,8 +2638,8 @@
 qed
 
 lemma subspace_isomorphism:
-  assumes s: "subspace (S:: (real ^'n) set)"
-  and t: "subspace (T :: (real ^'m) set)"
+  assumes s: "subspace (S:: ('a::euclidean_space) set)"
+  and t: "subspace (T :: ('b::euclidean_space) set)"
   and d: "dim S = dim T"
   shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"
 proof-
@@ -3092,17 +2712,16 @@
 qed
 
 lemma linear_eq_stdbasis:
-  assumes lf: "linear (f::real^'m \<Rightarrow> _)" and lg: "linear g"
-  and fg: "\<forall>i. f (basis i) = g(basis i)"
+  assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> _)" and lg: "linear g"
+  and fg: "\<forall>i<DIM('a::euclidean_space). f (basis i) = g(basis i)"
   shows "f = g"
 proof-
-  let ?U = "UNIV :: 'm set"
-  let ?I = "{basis i:: real^'m|i. i \<in> ?U}"
-  {fix x assume x: "x \<in> (UNIV :: (real^'m) set)"
-    from equalityD2[OF span_stdbasis]
-    have IU: " (UNIV :: (real^'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}
+  let ?U = "{..<DIM('a)}"
+  let ?I = "(basis::nat=>'a) ` {..<DIM('a)}"
+  {fix x assume x: "x \<in> (UNIV :: 'a set)"
+    from equalityD2[OF span_basis'[where 'a='a]]
+    have IU: " (UNIV :: 'a set) \<subseteq> span ?I" by blast
+    have "f x = g x" apply(rule linear_eq[OF lf lg IU,rule_format]) using fg x by auto }
   then show ?thesis by (auto intro: ext)
 qed
 
@@ -3136,35 +2755,29 @@
   then show ?thesis using SB TC by (auto intro: ext)
 qed
 
-lemma bilinear_eq_stdbasis:
-  assumes bf: "bilinear (f:: real^'m \<Rightarrow> real^'n \<Rightarrow> _)"
+lemma bilinear_eq_stdbasis: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> _"
+  assumes bf: "bilinear f"
   and bg: "bilinear g"
-  and fg: "\<forall>i j. f (basis i) (basis j) = g (basis i) (basis j)"
+  and fg: "\<forall>i<DIM('a). \<forall>j<DIM('b). 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)
+  from fg have th: "\<forall>x \<in> (basis ` {..<DIM('a)}). \<forall>y\<in> (basis ` {..<DIM('b)}). f x y = g x y" by blast
+  from bilinear_eq[OF bf bg equalityD2[OF span_basis'] equalityD2[OF span_basis'] th]
+  show ?thesis by (blast intro: ext)
 qed
 
 text {* Detailed theorems about left and right invertibility in general case. *}
 
-lemma left_invertible_transpose:
-  "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
-  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
-
-lemma right_invertible_transpose:
-  "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
-  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
-
-lemma linear_injective_left_inverse:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and fi: "inj f"
+lemma linear_injective_left_inverse: fixes f::"'a::euclidean_space => 'b::euclidean_space"
+  assumes lf: "linear f" 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 linear_independent_extend[OF independent_injective_image, OF independent_basis, OF lf fi]
+  obtain h:: "'b => 'a" where h: "linear h"
+    " \<forall>x \<in> f ` basis ` {..<DIM('a)}. 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]
+  have th: "\<forall>i<DIM('a). (h \<circ> f) (basis i) = id (basis i)"
+    using inv_o_cancel[OF fi, unfolded expand_fun_eq id_def o_def]
     by auto
 
   from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]
@@ -3172,183 +2785,28 @@
   then show ?thesis using h(1) by blast
 qed
 
-lemma linear_surjective_right_inverse:
-  assumes lf: "linear (f:: real ^'m \<Rightarrow> real ^'n)" and sf: "surj f"
+lemma linear_surjective_right_inverse: fixes f::"'a::euclidean_space => 'b::euclidean_space"
+  assumes lf: "linear f" 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 linear_independent_extend[OF independent_basis[where 'a='b],of "inv f"]
+  obtain h:: "'b \<Rightarrow> 'a" where
+    h: "linear h" "\<forall> x\<in> basis ` {..<DIM('b)}. 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
-
+  have th: "\<forall>i<DIM('b). (f o h) (basis i) = id (basis i)"
+    using sf by(auto simp add: surj_iff o_def expand_fun_eq)
   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^'m) = 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^'n) ** (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^'m) ** (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^'m"
-  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^'m"
-  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_transpose[symmetric]
-    matrix_left_invertible_independent_columns
-  by (simp add: column_transpose)
-
-lemma matrix_right_invertible_span_columns:
-  "(\<exists>(B::real ^'n^'m). (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
-          unfolding smult_conv_scaleR
-          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", folded smult_conv_scaleR])
-        show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
-          by (rule exI[where x=0], simp)
-      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: field_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^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
-  unfolding right_invertible_transpose[symmetric]
-  unfolding columns_transpose[symmetric]
-  unfolding matrix_right_invertible_span_columns
- ..
-
-text {* An injective map @{typ "real^'n \<Rightarrow> real^'n"} is also surjective. *}
-
-lemma linear_injective_imp_surjective:
-  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and fi: "inj f"
+text {* An injective map @{typ "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"} is also surjective. *}
+
+lemma linear_injective_imp_surjective:  fixes f::"'a::euclidean_space => 'a::euclidean_space"
+  assumes lf: "linear f" and fi: "inj f"
   shows "surj f"
 proof-
-  let ?U = "UNIV :: (real ^'n) set"
+  let ?U = "UNIV :: 'a set"
   from basis_exists[of ?U] obtain B
     where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "card B = dim ?U"
     by blast
@@ -3406,11 +2864,11 @@
   ultimately show ?thesis by blast
 qed
 
-lemma linear_surjective_imp_injective:
-  assumes lf: "linear (f::real ^'n => real ^'n)" and sf: "surj f"
+lemma linear_surjective_imp_injective: fixes f::"'a::euclidean_space => 'a::euclidean_space"
+  assumes lf: "linear f" and sf: "surj f"
   shows "inj f"
 proof-
-  let ?U = "UNIV :: (real ^'n) set"
+  let ?U = "UNIV :: 'a set"
   from basis_exists[of ?U] obtain B
     where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" and d: "card B = dim ?U"
     by blast
@@ -3465,172 +2923,103 @@
   "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 \<Rightarrow> real ^'n)" and fi: "inj f"
+lemma linear_injective_isomorphism: fixes f::"'a::euclidean_space => 'a::euclidean_space"
+  assumes lf: "linear f" 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 \<Rightarrow> real ^'n)" and sf: "surj f"
+lemma linear_surjective_isomorphism: fixes f::"'a::euclidean_space => 'a::euclidean_space"
+  assumes lf: "linear f" 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)
 
-text {* Left and right inverses are the same for @{typ "real^'n \<Rightarrow> real^'n"}. *}
-
-lemma linear_inverse_left:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and lf': "linear f'"
+text {* Left and right inverses are the same for @{typ "'a::euclidean_space => 'a::euclidean_space"}. *}
+
+lemma linear_inverse_left: fixes f::"'a::euclidean_space => 'a::euclidean_space"
+  assumes lf: "linear f" and lf': "linear f'"
   shows "f o f' = id \<longleftrightarrow> f' o f = id"
 proof-
-  {fix f f':: "real ^'n \<Rightarrow> real ^'n"
+  {fix f f':: "'a => 'a"
     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)
+      apply (auto simp add: o_def expand_fun_eq 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
+    have "f' o f = id" unfolding expand_fun_eq o_def id_def
       by metis}
   then show ?thesis using lf lf' by metis
 qed
 
 text {* Moreover, a one-sided inverse is automatically linear. *}
 
-lemma left_inverse_linear:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and gf: "g o f = id"
+lemma left_inverse_linear: fixes f::"'a::euclidean_space => 'a::euclidean_space"
+  assumes lf: "linear f" 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])
+  from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def expand_fun_eq)
     by metis
   from linear_injective_isomorphism[OF lf fi]
-  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
+  obtain h:: "'a \<Rightarrow> 'a" 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 \<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])
+    apply (simp add: o_def id_def expand_fun_eq)
     by metis
   with h(1) show ?thesis by blast
 qed
 
-text {* The same result in terms of square matrices. *}
-
-lemma matrix_left_right_inverse:
-  fixes A A' :: "real ^'n^'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
-
-text {* 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 transpose_columnvector:
- "transpose(columnvector v) = rowvector v"
-  by (simp add: transpose_def rowvector_def columnvector_def Cart_eq)
-
-lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"
-  by (simp add: transpose_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::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1"
-  by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vector_def)
-
-lemma dot_matrix_vector_mul:
-  fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
-  shows "(A *v x) \<bullet> (B *v y) =
-      (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
-unfolding dot_matrix_product transpose_columnvector[symmetric]
-  dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
-
 subsection {* Infinity norm *}
 
-definition "infnorm (x::real^'n) = Sup {abs(x$i) |i. i\<in> (UNIV :: 'n set)}"
+definition "infnorm (x::'a::euclidean_space) = Sup {abs(x$$i) |i. i<DIM('a)}"
 
 lemma numseg_dimindex_nonempty: "\<exists>i. i \<in> (UNIV :: 'n set)"
   by auto
 
 lemma infnorm_set_image:
-  "{abs(x$i) |i. i\<in> (UNIV :: _ set)} =
-  (\<lambda>i. abs(x$i)) ` (UNIV)" by blast
+  "{abs((x::'a::euclidean_space)$$i) |i. i<DIM('a)} =
+  (\<lambda>i. abs(x$$i)) ` {..<DIM('a)}" by blast
 
 lemma infnorm_set_lemma:
-  shows "finite {abs((x::'a::abs ^'n)$i) |i. i\<in> (UNIV :: 'n set)}"
-  and "{abs(x$i) |i. i\<in> (UNIV :: 'n::finite set)} \<noteq> {}"
+  shows "finite {abs((x::'a::euclidean_space)$$i) |i. i<DIM('a)}"
+  and "{abs(x$$i) |i. i<DIM('a::euclidean_space)} \<noteq> {}"
   unfolding infnorm_set_image
   by (auto intro: finite_imageI)
 
-lemma infnorm_pos_le: "0 \<le> infnorm (x::real^'n)"
+lemma infnorm_pos_le: "0 \<le> infnorm (x::'a::euclidean_space)"
   unfolding infnorm_def
   unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]
   unfolding infnorm_set_image
   by auto
 
-lemma infnorm_triangle: "infnorm ((x::real^'n) + y) \<le> infnorm x + infnorm y"
+lemma infnorm_triangle: "infnorm ((x::'a::euclidean_space) + 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
+  have *:"\<And>i. i \<in> {..<DIM('a)} \<longleftrightarrow> i <DIM('a)" by auto
   show ?thesis
-  unfolding infnorm_def
-  unfolding Sup_finite_le_iff[ OF infnorm_set_lemma]
+  unfolding infnorm_def unfolding  Sup_finite_le_iff[ OF infnorm_set_lemma[where 'a='a]]
   apply (subst diff_le_eq[symmetric])
   unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]
   unfolding infnorm_set_image bex_simps
   apply (subst th)
-  unfolding th1
-  unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma]
-
+  unfolding th1 *
+  unfolding Sup_finite_ge_iff[ OF infnorm_set_lemma[where 'a='a]]
   unfolding infnorm_set_image ball_simps bex_simps
-  apply simp
-  apply (metis th2)
-  done
+  unfolding euclidean_simps by (metis th2)
 qed
 
-lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::real ^'n) = 0"
+lemma infnorm_eq_0: "infnorm x = 0 \<longleftrightarrow> (x::_::euclidean_space) = 0"
 proof-
   have "infnorm x <= 0 \<longleftrightarrow> x = 0"
     unfolding infnorm_def
     unfolding Sup_finite_le_iff[OF infnorm_set_lemma]
     unfolding infnorm_set_image ball_simps
-    by vector
+    apply(subst (1) euclidean_eq) unfolding euclidean_component.zero
+    by auto
   then show ?thesis using infnorm_pos_le[of x] by simp
 qed
 
@@ -3640,10 +3029,7 @@
 lemma infnorm_neg: "infnorm (- x) = infnorm x"
   unfolding infnorm_def
   apply (rule cong[of "Sup" "Sup"])
-  apply blast
-  apply (rule set_ext)
-  apply auto
-  done
+  apply blast by(auto simp add: euclidean_simps)
 
 lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"
 proof-
@@ -3666,44 +3052,39 @@
   using infnorm_pos_le[of x] by arith
 
 lemma component_le_infnorm:
-  shows "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
-proof-
-  let ?U = "UNIV :: 'n set"
-  let ?S = "{\<bar>x$i\<bar> |i. i\<in> ?U}"
+  shows "\<bar>x$$i\<bar> \<le> infnorm (x::'a::euclidean_space)"
+proof(cases "i<DIM('a)")
+  case False thus ?thesis using infnorm_pos_le by auto
+next case True
+  let ?U = "{..<DIM('a)}"
+  let ?S = "{\<bar>x$$i\<bar> |i. i<DIM('a)}"
   have fS: "finite ?S" unfolding image_Collect[symmetric]
-    apply (rule finite_imageI) unfolding Collect_def mem_def by simp
+    apply (rule finite_imageI) by simp
   have S0: "?S \<noteq> {}" by blast
   have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
-  from Sup_finite_in[OF fS S0] 
-  show ?thesis unfolding infnorm_def infnorm_set_image 
-    by (metis Sup_finite_ge_iff finite finite_imageI UNIV_not_empty image_is_empty 
-              rangeI order_refl)
+  show ?thesis unfolding infnorm_def  
+    apply(subst Sup_finite_ge_iff) using Sup_finite_in[OF fS S0]
+    using infnorm_set_image using True by auto
 qed
 
-lemma infnorm_mul_lemma: "infnorm(a *s x) <= \<bar>a\<bar> * infnorm x"
+lemma infnorm_mul_lemma: "infnorm(a *\<^sub>R x) <= \<bar>a\<bar> * infnorm x"
   apply (subst infnorm_def)
   unfolding Sup_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"
+  unfolding infnorm_set_image ball_simps euclidean_scaleR abs_mult
+  using component_le_infnorm[of x] by(auto intro: mult_mono) 
+
+lemma infnorm_mul: "infnorm(a *\<^sub>R 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 th: "(1/a) *\<^sub>R (a *\<^sub>R x) = x" by simp
     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)"
+    from infnorm_mul_lemma[of "1/a" "a *\<^sub>R x"]
+    have "infnorm x \<le> 1/\<bar>a\<bar> * infnorm (a*\<^sub>R 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)"
+    with ap have "\<bar>a\<bar> * infnorm x \<le> \<bar>a\<bar> * (1/\<bar>a\<bar> * infnorm (a *\<^sub>R x))" by (simp add: field_simps)
+    then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*\<^sub>R x)"
       using ap by (simp add: field_simps)
     with infnorm_mul_lemma[of a x] have ?thesis by arith }
   ultimately show ?thesis by blast
@@ -3718,10 +3099,12 @@
   unfolding infnorm_def Sup_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)"
+
+lemma norm_le_infnorm: "norm(x) <= sqrt(real DIM('a)) * infnorm(x::'a::euclidean_space)"
 proof-
-  let ?d = "CARD('n)"
+  let ?d = "DIM('a)"
   have "real ?d \<ge> 0" by simp
   hence d2: "(sqrt (real ?d))^2 = real ?d"
     by (auto intro: real_sqrt_pow2)
@@ -3729,14 +3112,14 @@
     by (simp add: zero_le_mult_iff infnorm_pos_le)
   have th1: "x \<bullet> x \<le> (sqrt (real ?d) * infnorm x)^2"
     unfolding power_mult_distrib d2
-    unfolding real_of_nat_def inner_vector_def
-    apply (subst power2_abs[symmetric]) 
-    apply (rule setsum_bounded)
+    unfolding real_of_nat_def apply(subst euclidean_inner)
+    apply (subst power2_abs[symmetric])
+    apply(rule order_trans[OF setsum_bounded[where K="\<bar>infnorm x\<bar>\<twosuperior>"]])
     apply(auto simp add: power2_eq_square[symmetric])
     apply (subst power2_abs[symmetric])
     apply (rule power_mono)
     unfolding infnorm_def  Sup_finite_ge_iff[OF infnorm_set_lemma]
-    unfolding infnorm_set_image bex_simps apply(rule_tac x=i in exI) by auto
+    unfolding infnorm_set_image bex_simps apply(rule_tac x=i in bexI) by auto
   from real_le_lsqrt[OF inner_ge_zero th th1]
   show ?thesis unfolding norm_eq_sqrt_inner id_def .
 qed
@@ -3755,7 +3138,7 @@
     from inner_eq_zero_iff[of "norm y *\<^sub>R x - norm x *\<^sub>R 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 inner_simps smult_conv_scaleR
+      unfolding inner_simps
       unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: inner_commute)
       apply (simp add: field_simps) by metis
     also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
@@ -3843,7 +3226,7 @@
       from cy y have cy0: "cy \<noteq> 0" by auto
       let ?d = "cy / cx"
       from cx cy cx0 have "y = ?d *\<^sub>R x"
-        by (simp add: vector_smult_assoc)
+        by simp
       hence ?rhs using x y by blast}
     moreover
     {assume h: "?rhs"
@@ -3875,12 +3258,12 @@
 apply (rule exI[where x="(1/norm x) * norm y"])
 apply (drule sym)
 unfolding scaleR_scaleR[symmetric]
-apply (simp add: vector_smult_assoc field_simps)
+apply (simp add: field_simps)
 apply (rule exI[where x="(1/norm x) * - norm y"])
 apply clarify
 apply (drule sym)
 unfolding scaleR_scaleR[symmetric]
-apply (simp add: vector_smult_assoc field_simps)
+apply (simp add: field_simps)
 apply (erule exE)
 apply (erule ssubst)
 unfolding scaleR_scaleR
@@ -3894,4 +3277,78 @@
 apply simp
 done
 
+print_antiquotations
+
+section "Instanciate @{typ real} and @{typ complex} as typeclass @{text ordered_euclidean_space}."
+
+instantiation real :: real_basis_with_inner
+begin
+definition [simp]: "basis i = (if i = 0 then (1::real) else 0)"
+
+lemma basis_real_range: "basis ` {..<1} = {1::real}" by auto
+
+instance proof
+  let ?b = "basis::nat \<Rightarrow> real"
+
+  from basis_real_range have "independent (?b ` {..<1})" by auto
+  thus "\<exists>d>0. ?b ` {d..} = {0} \<and> independent (?b ` {..<d}) \<and> inj_on ?b {..<d}"
+    by (auto intro!: exI[of _ 1] inj_onI)
+
+  { fix x::real
+    have "x \<in> span (range ?b)"
+      using span_mul[of 1 "range ?b" x] span_clauses(1)[of 1 "range ?b"]
+      by auto }
+  thus "span (range ?b) = UNIV" by auto
+qed
 end
+
+lemma DIM_real[simp]: "DIM(real) = 1"
+  by (rule dimension_eq) (auto simp: basis_real_def)
+
+instance real::ordered_euclidean_space proof qed(auto simp add:euclidean_component_def)
+
+lemma Eucl_real_simps[simp]:
+  "(x::real) $$ 0 = x"
+  "(\<chi>\<chi> i. f i) = ((f 0)::real)"
+  "\<And>i. i > 0 \<Longrightarrow> x $$ i = 0"
+  defer apply(subst euclidean_eq) apply safe
+  unfolding euclidean_lambda_beta'
+  unfolding euclidean_component_def by auto
+
+instantiation complex :: real_basis_with_inner
+begin
+definition "basis i = (if i = 0 then 1 else if i = 1 then ii else 0)"
+
+lemma complex_basis[simp]:"basis 0 = (1::complex)" "basis 1 = ii" "basis (Suc 0) = ii"
+  unfolding basis_complex_def by auto
+
+instance
+proof
+  let ?b = "basis::nat \<Rightarrow> complex"
+  have [simp]: "(range ?b) = {0, basis 0, basis 1}"
+    by (auto simp: basis_complex_def split: split_if_asm)
+  { fix z::complex
+    have "z \<in> span (range ?b)"
+      by (auto simp: span_finite complex_equality
+        intro!: exI[of _ "\<lambda>i. if i = 1 then Re z else if i = ii then Im z else 0"]) }
+  thus "span (range ?b) = UNIV" by auto
+
+  have "{..<2} = {0, 1::nat}" by auto
+  hence *: "?b ` {..<2} = {1, ii}"
+    by (auto simp add: basis_complex_def)
+  moreover have "1 \<notin> span {\<i>}"
+    by (simp add: span_finite complex_equality complex_scaleR_def)
+  hence "independent (?b ` {..<2})"
+    by (simp add: * basis_complex_def independent_empty independent_insert)
+  ultimately show "\<exists>d>0. ?b ` {d..} = {0} \<and> independent (?b ` {..<d}) \<and> inj_on ?b {..<d}"
+    by (auto intro!: exI[of _ 2] inj_onI simp: basis_complex_def split: split_if_asm)
+qed
+end
+
+lemma DIM_complex[simp]: "DIM(complex) = 2"
+  by (rule dimension_eq) (auto simp: basis_complex_def)
+
+instance complex :: euclidean_space
+  proof qed (auto simp add: basis_complex_def inner_complex_def)
+
+end
--- a/src/HOL/Multivariate_Analysis/Fashoda.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Fashoda.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -4,13 +4,13 @@
 header {* Fashoda meet theorem. *}
 
 theory Fashoda
-imports Brouwer_Fixpoint Vec1 Path_Connected
+imports Brouwer_Fixpoint Path_Connected
 begin
 
 subsection {*Fashoda meet theorem. *}
 
 lemma infnorm_2: "infnorm (x::real^2) = max (abs(x$1)) (abs(x$2))"
-  unfolding infnorm_def UNIV_2 apply(rule Sup_eq) by auto
+  unfolding infnorm_cart UNIV_2 apply(rule Sup_eq) by auto
 
 lemma infnorm_eq_1_2: "infnorm (x::real^2) = 1 \<longleftrightarrow>
         (abs(x$1) \<le> 1 \<and> abs(x$2) \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1))"
@@ -31,16 +31,16 @@
     unfolding negatex_def infnorm_2 vector_2 by auto
   have lem2:"\<forall>z. z\<noteq>0 \<longrightarrow> infnorm(sqprojection z) = 1" unfolding sqprojection_def
     unfolding infnorm_mul[unfolded smult_conv_scaleR] unfolding abs_inverse real_abs_infnorm
-    unfolding infnorm_eq_0[THEN sym] by auto
+    apply(subst infnorm_eq_0[THEN sym]) by auto
   let ?F = "(\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w)"
   have *:"\<And>i. (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real}"
-    apply(rule set_ext) unfolding image_iff Bex_def mem_interval apply rule defer 
+    apply(rule set_ext) unfolding image_iff Bex_def mem_interval_cart apply rule defer 
     apply(rule_tac x="vec x" in exI) by auto
   { fix x assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` {- 1..1::real^2}"
     then guess w unfolding image_iff .. note w = this
-    hence "x \<noteq> 0" using as[of "w$1" "w$2"] unfolding mem_interval by auto} note x0=this
+    hence "x \<noteq> 0" using as[of "w$1" "w$2"] unfolding mem_interval_cart by auto} note x0=this
   have 21:"\<And>i::2. i\<noteq>1 \<Longrightarrow> i=2" using UNIV_2 by auto
-  have 1:"{- 1<..<1::real^2} \<noteq> {}" unfolding interval_eq_empty by auto
+  have 1:"{- 1<..<1::real^2} \<noteq> {}" unfolding interval_eq_empty_cart by auto
   have 2:"continuous_on {- 1..1} (negatex \<circ> sqprojection \<circ> ?F)" apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+
     prefer 2 apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+ unfolding *
     apply(rule assms)+ apply(rule continuous_on_compose,subst sqprojection_def)
@@ -54,7 +54,7 @@
     case goal1 then guess y unfolding image_iff .. note y=this have "?F y \<noteq> 0" apply(rule x0) using y(1) by auto
     hence *:"infnorm (sqprojection (?F y)) = 1" unfolding y o_def apply- by(rule lem2[rule_format])
     have "infnorm x = 1" unfolding *[THEN sym] y o_def by(rule lem1[rule_format])
-    thus "x\<in>{- 1..1}" unfolding mem_interval infnorm_2 apply- apply rule
+    thus "x\<in>{- 1..1}" unfolding mem_interval_cart infnorm_2 apply- apply rule
     proof-case goal1 thus ?case apply(cases "i=1") defer apply(drule 21) by auto qed qed
   guess x apply(rule brouwer_weak[of "{- 1..1::real^2}" "negatex \<circ> sqprojection \<circ> ?F"])
     apply(rule compact_interval convex_interval)+ unfolding interior_closed_interval
@@ -69,7 +69,7 @@
       unfolding sqprojection_def vector_component_simps Cart_nth.scaleR real_scaleR_def
       unfolding zero_less_mult_iff mult_less_0_iff by(auto simp add:field_simps) qed
   note lem3 = this[rule_format]
-  have x1:"x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}" using x(1) unfolding mem_interval by auto
+  have x1:"x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}" using x(1) unfolding mem_interval_cart by auto
   hence nz:"f (x $ 1) - g (x $ 2) \<noteq> 0" unfolding right_minus_eq apply-apply(rule as) by auto
   have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1" using nx unfolding infnorm_eq_1_2 by auto 
   thus False proof- fix P Q R S 
@@ -80,7 +80,7 @@
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]]
       unfolding as negatex_def vector_2 by auto moreover
     from x1 have "g (x $ 2) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
-    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
+    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart 
       apply(erule_tac x=1 in allE) by auto 
   next assume as:"x$1 = -1"
     hence *:"f (x $ 1) $ 1 = - 1" using assms(5) by auto
@@ -88,7 +88,7 @@
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]]
       unfolding as negatex_def vector_2 by auto moreover
     from x1 have "g (x $ 2) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
-    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
+    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart 
       apply(erule_tac x=1 in allE) by auto
   next assume as:"x$2 = 1"
     hence *:"g (x $ 2) $ 2 = 1" using assms(8) by auto
@@ -96,7 +96,7 @@
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]]
       unfolding as negatex_def vector_2 by auto moreover
     from x1 have "f (x $ 1) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
-    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
+    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart 
      apply(erule_tac x=2 in allE) by auto
  next assume as:"x$2 = -1"
     hence *:"g (x $ 2) $ 2 = - 1" using assms(7) by auto
@@ -104,7 +104,7 @@
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]]
       unfolding as negatex_def vector_2 by auto moreover
     from x1 have "f (x $ 1) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
-    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
+    ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval_cart 
       apply(erule_tac x=2 in allE) by auto qed(auto) qed
 
 lemma fashoda_unit_path: fixes f ::"real \<Rightarrow> real^2" and g ::"real \<Rightarrow> real^2"
@@ -129,6 +129,12 @@
     apply(rule_tac x="iscale s" in bexI) prefer 3 apply(rule_tac x="iscale t" in bexI)
     using isc[unfolded subset_eq, rule_format] by auto qed
 
+(* move *)
+lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" 
+  shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
+  unfolding interval_bij_cart split_conv Cart_eq Cart_lambda_beta
+  apply(rule,insert assms,erule_tac x=i in allE) by auto
+
 lemma fashoda: fixes b::"real^2"
   assumes "path f" "path g" "path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}"
   "(pathstart f)$1 = a$1" "(pathfinish f)$1 = b$1"
@@ -136,28 +142,28 @@
   obtains z where "z \<in> path_image f" "z \<in> path_image g" proof-
   fix P Q S presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" "Q \<Longrightarrow> thesis" "S \<Longrightarrow> thesis" thus thesis by auto
 next have "{a..b} \<noteq> {}" using assms(3) using path_image_nonempty by auto
-  hence "a \<le> b" unfolding interval_eq_empty vector_le_def by(auto simp add: not_less)
+  hence "a \<le> b" unfolding interval_eq_empty_cart vector_le_def by(auto simp add: not_less)
   thus "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)" unfolding vector_le_def forall_2 by auto
-next assume as:"a$1 = b$1" have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" apply(rule connected_ivt_component)
+next assume as:"a$1 = b$1" have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2" apply(rule connected_ivt_component_cart)
     apply(rule connected_path_image assms)+apply(rule pathstart_in_path_image,rule pathfinish_in_path_image)
     unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"]
     unfolding pathstart_def by(auto simp add: vector_le_def) then guess z .. note z=this
   have "z \<in> {a..b}" using z(1) assms(4) unfolding path_image_def by blast 
   hence "z = f 0" unfolding Cart_eq forall_2 unfolding z(2) pathstart_def
-    using assms(3)[unfolded path_image_def subset_eq mem_interval,rule_format,of "f 0" 1]
-    unfolding mem_interval apply(erule_tac x=1 in allE) using as by auto
+    using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1]
+    unfolding mem_interval_cart apply(erule_tac x=1 in allE) using as by auto
   thus thesis apply-apply(rule that[OF _ z(1)]) unfolding path_image_def by auto
-next assume as:"a$2 = b$2" have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" apply(rule connected_ivt_component)
+next assume as:"a$2 = b$2" have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1" apply(rule connected_ivt_component_cart)
     apply(rule connected_path_image assms)+apply(rule pathstart_in_path_image,rule pathfinish_in_path_image)
     unfolding assms using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"]
     unfolding pathstart_def by(auto simp add: vector_le_def) then guess z .. note z=this
   have "z \<in> {a..b}" using z(1) assms(3) unfolding path_image_def by blast 
   hence "z = g 0" unfolding Cart_eq forall_2 unfolding z(2) pathstart_def
-    using assms(4)[unfolded path_image_def subset_eq mem_interval,rule_format,of "g 0" 2]
-    unfolding mem_interval apply(erule_tac x=2 in allE) using as by auto
+    using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2]
+    unfolding mem_interval_cart apply(erule_tac x=2 in allE) using as by auto
   thus thesis apply-apply(rule that[OF z(1)]) unfolding path_image_def by auto
 next assume as:"a $ 1 < b $ 1 \<and> a $ 2 < b $ 2"
-  have int_nem:"{- 1..1::real^2} \<noteq> {}" unfolding interval_eq_empty by auto
+  have int_nem:"{- 1..1::real^2} \<noteq> {}" unfolding interval_eq_empty_cart by auto
   guess z apply(rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"]) 
     unfolding path_def path_image_def pathstart_def pathfinish_def
     apply(rule_tac[1-2] continuous_on_compose) apply(rule assms[unfolded path_def] continuous_on_interval_bij)+
@@ -173,13 +179,14 @@
   next show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1"
       "(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1"
       "(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1"
-      "(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1" unfolding interval_bij_def Cart_lambda_beta vector_component_simps o_def split_conv
+      "(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1"
+      unfolding interval_bij_cart Cart_lambda_beta vector_component_simps o_def split_conv
       unfolding assms[unfolded pathstart_def pathfinish_def] using as by auto qed note z=this
   from z(1) guess zf unfolding image_iff .. note zf=this
   from z(2) guess zg unfolding image_iff .. note zg=this
   have *:"\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i" unfolding forall_2 using as by auto
   show thesis apply(rule_tac z="interval_bij (- 1,1) (a,b) z" in that)
-    apply(subst zf) defer apply(subst zg) unfolding o_def interval_bij_bij[OF *] path_image_def
+    apply(subst zf) defer apply(subst zg) unfolding o_def interval_bij_bij_cart[OF *] path_image_def
     using zf(1) zg(1) by auto qed
 
 subsection {*Some slightly ad hoc lemmas I use below*}
@@ -244,10 +251,10 @@
   obtains z where "z \<in> path_image f" "z \<in> path_image g"
 proof-
   have "{a..b} \<noteq> {}" using path_image_nonempty using assms(3) by auto
-  note ab=this[unfolded interval_eq_empty not_ex forall_2 not_less]
+  note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less]
   have "pathstart f \<in> {a..b}" "pathfinish f \<in> {a..b}" "pathstart g \<in> {a..b}" "pathfinish g \<in> {a..b}"
     using pathstart_in_path_image pathfinish_in_path_image using assms(3-4) by auto
-  note startfin = this[unfolded mem_interval forall_2]
+  note startfin = this[unfolded mem_interval_cart forall_2]
   let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++
      linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++
      linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++
@@ -273,12 +280,12 @@
     show "path ?P1" "path ?P2" using assms by auto
     have "path_image ?P1 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 3
       apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format])
-      unfolding mem_interval forall_2 vector_2 using ab startfin abab assms(3)
+      unfolding mem_interval_cart forall_2 vector_2 using ab startfin abab assms(3)
       using assms(9-) unfolding assms by(auto simp add:field_simps)
     thus "path_image ?P1  \<subseteq> {?a .. ?b}" .
     have "path_image ?P2 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 2
       apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format])
-      unfolding mem_interval forall_2 vector_2 using ab startfin abab assms(4)
+      unfolding mem_interval_cart forall_2 vector_2 using ab startfin abab assms(4)
       using assms(9-) unfolding assms  by(auto simp add:field_simps)
     thus "path_image ?P2  \<subseteq> {?a .. ?b}" . 
     show "a $ 1 - 2 = a $ 1 - 2"  "b $ 1 + 2 = b $ 1 + 2" "pathstart g $ 2 - 3 = a $ 2 - 3"  "b $ 2 + 3 = b $ 2 + 3"
@@ -295,15 +302,15 @@
   z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False"
       apply(simp only: segment_vertical segment_horizontal vector_2) proof- case goal1 note as=this
       have "pathfinish f \<in> {a..b}" using assms(3) pathfinish_in_path_image[of f] by auto 
-      hence "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" unfolding mem_interval forall_2 by auto
+      hence "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False" unfolding mem_interval_cart forall_2 by auto
       hence "z$1 \<noteq> pathfinish f$1" using as(2) using assms ab by(auto simp add:field_simps)
       moreover have "pathstart f \<in> {a..b}" using assms(3) pathstart_in_path_image[of f] by auto 
-      hence "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" unfolding mem_interval forall_2 by auto
+      hence "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False" unfolding mem_interval_cart forall_2 by auto
       hence "z$1 \<noteq> pathstart f$1" using as(2) using assms ab by(auto simp add:field_simps)
       ultimately have *:"z$2 = a$2 - 2" using goal1(1) by auto
       have "z$1 \<noteq> pathfinish g$1" using as(2) using assms ab by(auto simp add:field_simps *)
       moreover have "pathstart g \<in> {a..b}" using assms(4) pathstart_in_path_image[of g] by auto 
-      note this[unfolded mem_interval forall_2]
+      note this[unfolded mem_interval_cart forall_2]
       hence "z$1 \<noteq> pathstart g$1" using as(1) using assms ab by(auto simp add:field_simps *)
       ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1"
         using as(2) unfolding * assms by(auto simp add:field_simps)
@@ -312,11 +319,11 @@
     hence z':"z\<in>{a..b}" using assms(3-4) by auto
     have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow> (z = pathstart f \<or> z = pathfinish f)"
       unfolding Cart_eq forall_2 assms by auto
-    with z' show "z\<in>path_image f" using z(1) unfolding Un_iff mem_interval forall_2 apply-
+    with z' show "z\<in>path_image f" using z(1) unfolding Un_iff mem_interval_cart forall_2 apply-
       apply(simp only: segment_vertical segment_horizontal vector_2) unfolding assms by auto
     have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow> (z = pathstart g \<or> z = pathfinish g)"
       unfolding Cart_eq forall_2 assms by auto
-    with z' show "z\<in>path_image g" using z(2) unfolding Un_iff mem_interval forall_2 apply-
+    with z' show "z\<in>path_image g" using z(2) unfolding Un_iff mem_interval_cart forall_2 apply-
       apply(simp only: segment_vertical segment_horizontal vector_2) unfolding assms by auto
   qed qed
 
@@ -535,7 +542,7 @@
                (!x. x IN path_image p ==> f x = x)`,
   REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN
   MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN
-  ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);;
+  ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);;
 
 let PATH_CONNECTED_ARC_COMPLEMENT = prove
  (`!p. 2 <= dimindex(:N) /\ arc p
--- a/src/HOL/Multivariate_Analysis/Integration.certs	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.certs	Mon Jun 21 19:33:51 2010 +0200
@@ -1,3 +1,1337 @@
+2e3a7f41999e849037c0bc39fd6c0ffa4f5b7eb6 910 0
+#2 := false
+#369 := 0::real
+decl f8 :: (-> S5 S2 real)
+decl f11 :: (-> S4 S2)
+decl f7 :: S4
+#13 := f7
+#20 := (f11 f7)
+decl f21 :: (-> S3 S5)
+decl f4 :: S3
+#8 := f4
+#101 := (f21 f4)
+#1205 := (f8 #101 #20)
+#367 := -1::real
+#1507 := (* -1::real #1205)
+decl f19 :: S3
+#60 := f19
+#88 := (f21 f19)
+#1037 := (f8 #88 #20)
+#1698 := (+ #1037 #1507)
+#1765 := (>= #1698 0::real)
+#1697 := (= #1037 #1205)
+decl f10 :: S5
+#19 := f10
+#21 := (f8 f10 #20)
+#1208 := (= #21 #1205)
+decl f12 :: S5
+#22 := f12
+#1129 := (f8 f12 #20)
+#1207 := (= #1129 #1205)
+decl f6 :: (-> S2 S4)
+#84 := (f6 #20)
+#351 := (= f7 #84)
+#1211 := (ite #351 #1208 #1207)
+decl f9 :: S5
+#16 := f9
+#31 := (f8 f9 #20)
+#1206 := (= #31 #1205)
+#70 := 0::int
+decl f5 :: (-> S4 int)
+#1041 := (f5 #84)
+#155 := -1::int
+#1051 := (* -1::int #1041)
+#14 := (f5 f7)
+#1097 := (+ #14 #1051)
+#1098 := (<= #1097 0::int)
+#1214 := (ite #1098 #1211 #1206)
+#9 := (:var 0 S2)
+#17 := (f8 f9 #9)
+#697 := (pattern #17)
+#23 := (f8 f12 #9)
+#696 := (pattern #23)
+#102 := (f8 #101 #9)
+#695 := (pattern #102)
+#11 := (f6 #9)
+#694 := (pattern #11)
+#580 := (= #17 #102)
+#578 := (= #23 #102)
+#577 := (= #21 #102)
+#18 := (= #11 f7)
+#579 := (ite #18 #577 #578)
+#158 := (* -1::int #14)
+#12 := (f5 #11)
+#159 := (+ #12 #158)
+#157 := (>= #159 0::int)
+#581 := (ite #157 #579 #580)
+#698 := (forall (vars (?v0 S2)) (:pat #694 #695 #696 #697) #581)
+#584 := (forall (vars (?v0 S2)) #581)
+#701 := (iff #584 #698)
+#699 := (iff #581 #581)
+#700 := [refl]: #699
+#702 := [quant-intro #700]: #701
+#24 := (ite #18 #21 #23)
+#165 := (ite #157 #24 #17)
+#486 := (= #102 #165)
+#487 := (forall (vars (?v0 S2)) #486)
+#585 := (iff #487 #584)
+#582 := (iff #486 #581)
+#583 := [rewrite]: #582
+#586 := [quant-intro #583]: #585
+#480 := (~ #487 #487)
+#482 := (~ #486 #486)
+#483 := [refl]: #482
+#481 := [nnf-pos #483]: #480
+decl f3 :: (-> S3 S2 real)
+#10 := (f3 f4 #9)
+#170 := (= #10 #165)
+#173 := (forall (vars (?v0 S2)) #170)
+#488 := (iff #173 #487)
+#130 := (:var 1 S3)
+#133 := (f3 #130 #9)
+#131 := (f21 #130)
+#132 := (f8 #131 #9)
+#134 := (= #132 #133)
+#135 := (forall (vars (?v0 S3) (?v1 S2)) #134)
+#441 := [asserted]: #135
+#489 := [rewrite* #441]: #488
+#15 := (< #12 #14)
+#25 := (ite #15 #17 #24)
+#26 := (= #10 #25)
+#27 := (forall (vars (?v0 S2)) #26)
+#174 := (iff #27 #173)
+#171 := (iff #26 #170)
+#168 := (= #25 #165)
+#156 := (not #157)
+#162 := (ite #156 #17 #24)
+#166 := (= #162 #165)
+#167 := [rewrite]: #166
+#163 := (= #25 #162)
+#160 := (iff #15 #156)
+#161 := [rewrite]: #160
+#164 := [monotonicity #161]: #163
+#169 := [trans #164 #167]: #168
+#172 := [monotonicity #169]: #171
+#175 := [quant-intro #172]: #174
+#152 := [asserted]: #27
+#176 := [mp #152 #175]: #173
+#490 := [mp #176 #489]: #487
+#478 := [mp~ #490 #481]: #487
+#587 := [mp #478 #586]: #584
+#703 := [mp #587 #702]: #698
+#962 := (not #698)
+#1217 := (or #962 #1214)
+#85 := (= #84 f7)
+#1209 := (ite #85 #1208 #1207)
+#1088 := (+ #1041 #158)
+#1089 := (>= #1088 0::int)
+#1210 := (ite #1089 #1209 #1206)
+#1218 := (or #962 #1210)
+#1220 := (iff #1218 #1217)
+#1222 := (iff #1217 #1217)
+#1223 := [rewrite]: #1222
+#1215 := (iff #1210 #1214)
+#1212 := (iff #1209 #1211)
+#353 := (iff #85 #351)
+#354 := [rewrite]: #353
+#1213 := [monotonicity #354]: #1212
+#1101 := (iff #1089 #1098)
+#1091 := (+ #158 #1041)
+#1094 := (>= #1091 0::int)
+#1099 := (iff #1094 #1098)
+#1100 := [rewrite]: #1099
+#1095 := (iff #1089 #1094)
+#1092 := (= #1088 #1091)
+#1093 := [rewrite]: #1092
+#1096 := [monotonicity #1093]: #1095
+#1102 := [trans #1096 #1100]: #1101
+#1216 := [monotonicity #1102 #1213]: #1215
+#1221 := [monotonicity #1216]: #1220
+#1224 := [trans #1221 #1223]: #1220
+#1219 := [quant-inst]: #1218
+#1225 := [mp #1219 #1224]: #1217
+#1767 := [unit-resolution #1225 #703]: #1214
+#1396 := (= #14 #1041)
+#1718 := (= #1041 #14)
+#350 := [asserted]: #85
+#357 := [mp #350 #354]: #351
+#1717 := [symm #357]: #85
+#1719 := [monotonicity #1717]: #1718
+#1749 := [symm #1719]: #1396
+#1750 := (not #1396)
+#1768 := (or #1750 #1098)
+#1769 := [th-lemma]: #1768
+#1770 := [unit-resolution #1769 #1749]: #1098
+#1116 := (not #1098)
+#1238 := (not #1214)
+#1239 := (or #1238 #1116 #1211)
+#1240 := [def-axiom]: #1239
+#1771 := [unit-resolution #1240 #1770 #1767]: #1211
+#1226 := (not #1211)
+#1772 := (or #1226 #1208)
+#1227 := (not #351)
+#1228 := (or #1226 #1227 #1208)
+#1229 := [def-axiom]: #1228
+#1773 := [unit-resolution #1229 #357]: #1772
+#1774 := [unit-resolution #1773 #1771]: #1208
+#1811 := (= #1037 #21)
+#1038 := (= #21 #1037)
+decl f16 :: S4
+#40 := f16
+#41 := (f5 f16)
+#1052 := (+ #41 #1051)
+#1053 := (<= #1052 0::int)
+#1074 := (not #1053)
+#196 := (* -1::int #41)
+#1665 := (+ #14 #196)
+#1666 := (>= #1665 0::int)
+#1754 := (not #1666)
+#1679 := (<= #1665 0::int)
+#44 := (f11 f16)
+#82 := (f6 #44)
+#778 := (f5 #82)
+#788 := (* -1::int #778)
+#835 := (+ #14 #788)
+#836 := (<= #835 0::int)
+decl f18 :: S3
+#55 := f18
+#98 := (f21 f18)
+#823 := (f8 #98 #44)
+#45 := (f8 f10 #44)
+#824 := (= #45 #823)
+#863 := (not #824)
+decl f15 :: S3
+#38 := f15
+#93 := (f21 f15)
+#902 := (f8 #93 #44)
+#1699 := (= #823 #902)
+#1759 := (not #1699)
+#1793 := (iff #1759 #863)
+#1791 := (iff #1699 #824)
+#1786 := (= #823 #45)
+#1789 := (iff #1786 #824)
+#1790 := [commutativity]: #1789
+#1787 := (iff #1699 #1786)
+#1784 := (= #902 #45)
+#905 := (= #45 #902)
+#869 := (f8 f12 #44)
+#904 := (= #869 #902)
+#347 := (= f16 #82)
+#908 := (ite #347 #905 #904)
+#867 := (f8 f9 #44)
+#903 := (= #867 #902)
+#789 := (+ #41 #788)
+#790 := (<= #789 0::int)
+#911 := (ite #790 #908 #903)
+#94 := (f8 #93 #9)
+#713 := (pattern #94)
+#602 := (= #17 #94)
+#600 := (= #23 #94)
+#599 := (= #45 #94)
+#43 := (= #11 f16)
+#601 := (ite #43 #599 #600)
+#197 := (+ #12 #196)
+#195 := (>= #197 0::int)
+#603 := (ite #195 #601 #602)
+#714 := (forall (vars (?v0 S2)) (:pat #694 #713 #696 #697) #603)
+#606 := (forall (vars (?v0 S2)) #603)
+#717 := (iff #606 #714)
+#715 := (iff #603 #603)
+#716 := [refl]: #715
+#718 := [quant-intro #716]: #717
+#46 := (ite #43 #45 #23)
+#203 := (ite #195 #46 #17)
+#497 := (= #94 #203)
+#498 := (forall (vars (?v0 S2)) #497)
+#607 := (iff #498 #606)
+#604 := (iff #497 #603)
+#605 := [rewrite]: #604
+#608 := [quant-intro #605]: #607
+#470 := (~ #498 #498)
+#472 := (~ #497 #497)
+#473 := [refl]: #472
+#471 := [nnf-pos #473]: #470
+#39 := (f3 f15 #9)
+#208 := (= #39 #203)
+#211 := (forall (vars (?v0 S2)) #208)
+#499 := (iff #211 #498)
+#500 := [rewrite* #441]: #499
+#42 := (< #12 #41)
+#47 := (ite #42 #17 #46)
+#48 := (= #39 #47)
+#49 := (forall (vars (?v0 S2)) #48)
+#212 := (iff #49 #211)
+#209 := (iff #48 #208)
+#206 := (= #47 #203)
+#194 := (not #195)
+#200 := (ite #194 #17 #46)
+#204 := (= #200 #203)
+#205 := [rewrite]: #204
+#201 := (= #47 #200)
+#198 := (iff #42 #194)
+#199 := [rewrite]: #198
+#202 := [monotonicity #199]: #201
+#207 := [trans #202 #205]: #206
+#210 := [monotonicity #207]: #209
+#213 := [quant-intro #210]: #212
+#154 := [asserted]: #49
+#214 := [mp #154 #213]: #211
+#501 := [mp #214 #500]: #498
+#468 := [mp~ #501 #471]: #498
+#609 := [mp #468 #608]: #606
+#719 := [mp #609 #718]: #714
+#914 := (not #714)
+#915 := (or #914 #911)
+#83 := (= #82 f16)
+#906 := (ite #83 #905 #904)
+#779 := (+ #778 #196)
+#780 := (>= #779 0::int)
+#907 := (ite #780 #906 #903)
+#916 := (or #914 #907)
+#918 := (iff #916 #915)
+#920 := (iff #915 #915)
+#921 := [rewrite]: #920
+#912 := (iff #907 #911)
+#909 := (iff #906 #908)
+#348 := (iff #83 #347)
+#349 := [rewrite]: #348
+#910 := [monotonicity #349]: #909
+#793 := (iff #780 #790)
+#782 := (+ #196 #778)
+#785 := (>= #782 0::int)
+#791 := (iff #785 #790)
+#792 := [rewrite]: #791
+#786 := (iff #780 #785)
+#783 := (= #779 #782)
+#784 := [rewrite]: #783
+#787 := [monotonicity #784]: #786
+#794 := [trans #787 #792]: #793
+#913 := [monotonicity #794 #910]: #912
+#919 := [monotonicity #913]: #918
+#922 := [trans #919 #921]: #918
+#917 := [quant-inst]: #916
+#923 := [mp #917 #922]: #915
+#1776 := [unit-resolution #923 #719]: #911
+#1394 := (= #41 #778)
+#1674 := (= #778 #41)
+#346 := [asserted]: #83
+#352 := [mp #346 #349]: #347
+#1673 := [symm #352]: #83
+#1675 := [monotonicity #1673]: #1674
+#1676 := [symm #1675]: #1394
+#1739 := (not #1394)
+#1777 := (or #1739 #790)
+#1778 := [th-lemma]: #1777
+#1779 := [unit-resolution #1778 #1676]: #790
+#812 := (not #790)
+#936 := (not #911)
+#937 := (or #936 #812 #908)
+#938 := [def-axiom]: #937
+#1780 := [unit-resolution #938 #1779 #1776]: #908
+#924 := (not #908)
+#1781 := (or #924 #905)
+#925 := (not #347)
+#926 := (or #924 #925 #905)
+#927 := [def-axiom]: #926
+#1782 := [unit-resolution #927 #352]: #1781
+#1783 := [unit-resolution #1782 #1780]: #905
+#1785 := [symm #1783]: #1784
+#1788 := [monotonicity #1785]: #1787
+#1792 := [trans #1788 #1790]: #1791
+#1794 := [monotonicity #1792]: #1793
+#1462 := (* -1::real #902)
+#1709 := (+ #823 #1462)
+#1711 := (>= #1709 0::real)
+#1708 := (not #1711)
+decl f22 :: S5
+#90 := f22
+#1463 := (f8 f22 #44)
+#1466 := (* -1::real #1463)
+#1477 := (+ #902 #1466)
+#1478 := (<= #1477 0::real)
+#1502 := (not #1478)
+#774 := (f8 #88 #44)
+#1467 := (+ #774 #1466)
+#1468 := (>= #1467 0::real)
+#1483 := (or #1468 #1478)
+#1486 := (not #1483)
+#91 := (f8 f22 #9)
+#761 := (pattern #91)
+#89 := (f8 #88 #9)
+#734 := (pattern #89)
+#376 := (* -1::real #94)
+#377 := (+ #91 #376)
+#375 := (>= #377 0::real)
+#371 := (* -1::real #91)
+#372 := (+ #89 #371)
+#370 := (>= #372 0::real)
+#561 := (or #370 #375)
+#562 := (not #561)
+#762 := (forall (vars (?v0 S2)) (:pat #734 #761 #713) #562)
+#565 := (forall (vars (?v0 S2)) #562)
+#765 := (iff #565 #762)
+#763 := (iff #562 #562)
+#764 := [refl]: #763
+#766 := [quant-intro #764]: #765
+#378 := (not #375)
+#368 := (not #370)
+#381 := (and #368 #378)
+#384 := (forall (vars (?v0 S2)) #381)
+#566 := (iff #384 #565)
+#563 := (iff #381 #562)
+#564 := [rewrite]: #563
+#567 := [quant-intro #564]: #566
+#553 := (~ #384 #384)
+#551 := (~ #381 #381)
+#552 := [refl]: #551
+#554 := [nnf-pos #552]: #553
+#394 := (* -1::real #102)
+#395 := (+ #91 #394)
+#393 := (>= #395 0::real)
+#396 := (not #393)
+#99 := (f8 #98 #9)
+#387 := (* -1::real #99)
+#388 := (+ #91 #387)
+#389 := (<= #388 0::real)
+#390 := (not #389)
+#399 := (and #390 #396)
+#402 := (forall (vars (?v0 S2)) #399)
+#405 := (and #384 #402)
+#103 := (< #91 #102)
+#100 := (< #99 #91)
+#104 := (and #100 #103)
+#105 := (forall (vars (?v0 S2)) #104)
+#95 := (< #91 #94)
+#92 := (< #89 #91)
+#96 := (and #92 #95)
+#97 := (forall (vars (?v0 S2)) #96)
+#106 := (and #97 #105)
+#406 := (iff #106 #405)
+#403 := (iff #105 #402)
+#400 := (iff #104 #399)
+#397 := (iff #103 #396)
+#398 := [rewrite]: #397
+#391 := (iff #100 #390)
+#392 := [rewrite]: #391
+#401 := [monotonicity #392 #398]: #400
+#404 := [quant-intro #401]: #403
+#385 := (iff #97 #384)
+#382 := (iff #96 #381)
+#379 := (iff #95 #378)
+#380 := [rewrite]: #379
+#373 := (iff #92 #368)
+#374 := [rewrite]: #373
+#383 := [monotonicity #374 #380]: #382
+#386 := [quant-intro #383]: #385
+#407 := [monotonicity #386 #404]: #406
+#363 := [asserted]: #106
+#408 := [mp #363 #407]: #405
+#409 := [and-elim #408]: #384
+#555 := [mp~ #409 #554]: #384
+#568 := [mp #555 #567]: #565
+#767 := [mp #568 #766]: #762
+#1489 := (not #762)
+#1490 := (or #1489 #1486)
+#1464 := (+ #1463 #1462)
+#1465 := (>= #1464 0::real)
+#1469 := (or #1468 #1465)
+#1470 := (not #1469)
+#1491 := (or #1489 #1470)
+#1493 := (iff #1491 #1490)
+#1495 := (iff #1490 #1490)
+#1496 := [rewrite]: #1495
+#1487 := (iff #1470 #1486)
+#1484 := (iff #1469 #1483)
+#1481 := (iff #1465 #1478)
+#1471 := (+ #1462 #1463)
+#1474 := (>= #1471 0::real)
+#1479 := (iff #1474 #1478)
+#1480 := [rewrite]: #1479
+#1475 := (iff #1465 #1474)
+#1472 := (= #1464 #1471)
+#1473 := [rewrite]: #1472
+#1476 := [monotonicity #1473]: #1475
+#1482 := [trans #1476 #1480]: #1481
+#1485 := [monotonicity #1482]: #1484
+#1488 := [monotonicity #1485]: #1487
+#1494 := [monotonicity #1488]: #1493
+#1497 := [trans #1494 #1496]: #1493
+#1492 := [quant-inst]: #1491
+#1498 := [mp #1492 #1497]: #1490
+#1701 := [unit-resolution #1498 #767]: #1486
+#1503 := (or #1483 #1502)
+#1504 := [def-axiom]: #1503
+#1702 := [unit-resolution #1504 #1701]: #1502
+#1579 := (+ #823 #1466)
+#1580 := (>= #1579 0::real)
+#1612 := (not #1580)
+#946 := (f8 #101 #44)
+#1591 := (+ #946 #1466)
+#1592 := (<= #1591 0::real)
+#1597 := (or #1580 #1592)
+#1600 := (not #1597)
+#727 := (pattern #99)
+#569 := (or #389 #393)
+#570 := (not #569)
+#768 := (forall (vars (?v0 S2)) (:pat #761 #727 #695) #570)
+#573 := (forall (vars (?v0 S2)) #570)
+#771 := (iff #573 #768)
+#769 := (iff #570 #570)
+#770 := [refl]: #769
+#772 := [quant-intro #770]: #771
+#574 := (iff #402 #573)
+#571 := (iff #399 #570)
+#572 := [rewrite]: #571
+#575 := [quant-intro #572]: #574
+#558 := (~ #402 #402)
+#556 := (~ #399 #399)
+#557 := [refl]: #556
+#559 := [nnf-pos #557]: #558
+#410 := [and-elim #408]: #402
+#560 := [mp~ #410 #559]: #402
+#576 := [mp #560 #575]: #573
+#773 := [mp #576 #772]: #768
+#1547 := (not #768)
+#1603 := (or #1547 #1600)
+#1565 := (* -1::real #946)
+#1566 := (+ #1463 #1565)
+#1567 := (>= #1566 0::real)
+#1568 := (* -1::real #823)
+#1569 := (+ #1463 #1568)
+#1570 := (<= #1569 0::real)
+#1571 := (or #1570 #1567)
+#1572 := (not #1571)
+#1604 := (or #1547 #1572)
+#1606 := (iff #1604 #1603)
+#1608 := (iff #1603 #1603)
+#1609 := [rewrite]: #1608
+#1601 := (iff #1572 #1600)
+#1598 := (iff #1571 #1597)
+#1595 := (iff #1567 #1592)
+#1585 := (+ #1565 #1463)
+#1588 := (>= #1585 0::real)
+#1593 := (iff #1588 #1592)
+#1594 := [rewrite]: #1593
+#1589 := (iff #1567 #1588)
+#1586 := (= #1566 #1585)
+#1587 := [rewrite]: #1586
+#1590 := [monotonicity #1587]: #1589
+#1596 := [trans #1590 #1594]: #1595
+#1583 := (iff #1570 #1580)
+#1573 := (+ #1568 #1463)
+#1576 := (<= #1573 0::real)
+#1581 := (iff #1576 #1580)
+#1582 := [rewrite]: #1581
+#1577 := (iff #1570 #1576)
+#1574 := (= #1569 #1573)
+#1575 := [rewrite]: #1574
+#1578 := [monotonicity #1575]: #1577
+#1584 := [trans #1578 #1582]: #1583
+#1599 := [monotonicity #1584 #1596]: #1598
+#1602 := [monotonicity #1599]: #1601
+#1607 := [monotonicity #1602]: #1606
+#1610 := [trans #1607 #1609]: #1606
+#1605 := [quant-inst]: #1604
+#1611 := [mp #1605 #1610]: #1603
+#1703 := [unit-resolution #1611 #773]: #1600
+#1613 := (or #1597 #1612)
+#1614 := [def-axiom]: #1613
+#1704 := [unit-resolution #1614 #1703]: #1612
+#1713 := (or #1708 #1580 #1478)
+#1714 := [th-lemma]: #1713
+#1715 := [unit-resolution #1714 #1704 #1702]: #1708
+#1760 := (or #1759 #1711)
+#1761 := [th-lemma]: #1760
+#1775 := [unit-resolution #1761 #1715]: #1759
+#1795 := [mp #1775 #1794]: #863
+#1797 := (or #836 #824)
+decl f14 :: S5
+#32 := f14
+#776 := (f8 f14 #44)
+#825 := (= #776 #823)
+#841 := (ite #836 #825 #824)
+#30 := (f8 f10 #9)
+#706 := (pattern #30)
+#33 := (f8 f14 #9)
+#705 := (pattern #33)
+#620 := (= #30 #99)
+#619 := (= #33 #99)
+#621 := (ite #157 #619 #620)
+#728 := (forall (vars (?v0 S2)) (:pat #694 #705 #727 #706) #621)
+#624 := (forall (vars (?v0 S2)) #621)
+#731 := (iff #624 #728)
+#729 := (iff #621 #621)
+#730 := [refl]: #729
+#732 := [quant-intro #730]: #731
+#235 := (ite #157 #33 #30)
+#508 := (= #99 #235)
+#509 := (forall (vars (?v0 S2)) #508)
+#625 := (iff #509 #624)
+#622 := (iff #508 #621)
+#623 := [rewrite]: #622
+#626 := [quant-intro #623]: #625
+#460 := (~ #509 #509)
+#462 := (~ #508 #508)
+#463 := [refl]: #462
+#461 := [nnf-pos #463]: #460
+#56 := (f3 f18 #9)
+#240 := (= #56 #235)
+#243 := (forall (vars (?v0 S2)) #240)
+#510 := (iff #243 #509)
+#511 := [rewrite* #441]: #510
+#57 := (ite #15 #30 #33)
+#58 := (= #56 #57)
+#59 := (forall (vars (?v0 S2)) #58)
+#244 := (iff #59 #243)
+#241 := (iff #58 #240)
+#238 := (= #57 #235)
+#232 := (ite #156 #30 #33)
+#236 := (= #232 #235)
+#237 := [rewrite]: #236
+#233 := (= #57 #232)
+#234 := [monotonicity #161]: #233
+#239 := [trans #234 #237]: #238
+#242 := [monotonicity #239]: #241
+#245 := [quant-intro #242]: #244
+#193 := [asserted]: #59
+#246 := [mp #193 #245]: #243
+#512 := [mp #246 #511]: #509
+#530 := [mp~ #512 #461]: #509
+#627 := [mp #530 #626]: #624
+#733 := [mp #627 #732]: #728
+#844 := (not #728)
+#845 := (or #844 #841)
+#826 := (+ #778 #158)
+#827 := (>= #826 0::int)
+#828 := (ite #827 #825 #824)
+#846 := (or #844 #828)
+#848 := (iff #846 #845)
+#850 := (iff #845 #845)
+#851 := [rewrite]: #850
+#842 := (iff #828 #841)
+#839 := (iff #827 #836)
+#829 := (+ #158 #778)
+#832 := (>= #829 0::int)
+#837 := (iff #832 #836)
+#838 := [rewrite]: #837
+#833 := (iff #827 #832)
+#830 := (= #826 #829)
+#831 := [rewrite]: #830
+#834 := [monotonicity #831]: #833
+#840 := [trans #834 #838]: #839
+#843 := [monotonicity #840]: #842
+#849 := [monotonicity #843]: #848
+#852 := [trans #849 #851]: #848
+#847 := [quant-inst]: #846
+#853 := [mp #847 #852]: #845
+#1796 := [unit-resolution #853 #733]: #841
+#854 := (not #841)
+#858 := (or #854 #836 #824)
+#859 := [def-axiom]: #858
+#1798 := [unit-resolution #859 #1796]: #1797
+#1799 := [unit-resolution #1798 #1795]: #836
+#855 := (not #836)
+#1746 := (or #1679 #855)
+#1738 := [hypothesis]: #836
+#1395 := (>= #789 0::int)
+#1740 := (or #1739 #1395)
+#1741 := [th-lemma]: #1740
+#1742 := [unit-resolution #1741 #1676]: #1395
+#1743 := (not #1679)
+#1744 := [hypothesis]: #1743
+#1745 := [th-lemma #1744 #1742 #1738]: false
+#1747 := [lemma #1745]: #1746
+#1800 := [unit-resolution #1747 #1799]: #1679
+#1803 := (or #1743 #1754)
+#1712 := (= #14 #41)
+#1736 := (not #1712)
+#356 := (= f7 f16)
+#953 := (= f7 #82)
+decl f20 :: (-> int S4)
+#1398 := (f20 #778)
+#1727 := (= #1398 #82)
+#1399 := (= #82 #1398)
+#65 := (:var 0 S4)
+#66 := (f5 #65)
+#741 := (pattern #66)
+#67 := (f20 #66)
+#247 := (= #65 #67)
+#742 := (forall (vars (?v0 S4)) (:pat #741) #247)
+#265 := (forall (vars (?v0 S4)) #247)
+#745 := (iff #265 #742)
+#743 := (iff #247 #247)
+#744 := [refl]: #743
+#746 := [quant-intro #744]: #745
+#538 := (~ #265 #265)
+#536 := (~ #247 #247)
+#537 := [refl]: #536
+#539 := [nnf-pos #537]: #538
+#68 := (= #67 #65)
+#69 := (forall (vars (?v0 S4)) #68)
+#266 := (iff #69 #265)
+#263 := (iff #68 #247)
+#264 := [rewrite]: #263
+#267 := [quant-intro #264]: #266
+#231 := [asserted]: #69
+#270 := [mp #231 #267]: #265
+#540 := [mp~ #270 #539]: #265
+#747 := [mp #540 #746]: #742
+#1294 := (not #742)
+#1402 := (or #1294 #1399)
+#1403 := [quant-inst]: #1402
+#1672 := [unit-resolution #1403 #747]: #1399
+#1728 := [symm #1672]: #1727
+#1731 := (= f7 #1398)
+#1400 := (f20 #1041)
+#1725 := (= #1400 #1398)
+#1722 := (= #1041 #778)
+#1720 := (= #1041 #41)
+#1716 := [hypothesis]: #1712
+#1721 := [trans #1719 #1716]: #1720
+#1723 := [trans #1721 #1676]: #1722
+#1726 := [monotonicity #1723]: #1725
+#1729 := (= f7 #1400)
+#1401 := (= #84 #1400)
+#1406 := (or #1294 #1401)
+#1407 := [quant-inst]: #1406
+#1724 := [unit-resolution #1407 #747]: #1401
+#1730 := [trans #357 #1724]: #1729
+#1732 := [trans #1730 #1726]: #1731
+#1733 := [trans #1732 #1728]: #953
+#1734 := [trans #1733 #1673]: #356
+#360 := (not #356)
+#86 := (= f16 f7)
+#87 := (not #86)
+#361 := (iff #87 #360)
+#358 := (iff #86 #356)
+#359 := [rewrite]: #358
+#362 := [monotonicity #359]: #361
+#355 := [asserted]: #87
+#365 := [mp #355 #362]: #360
+#1735 := [unit-resolution #365 #1734]: false
+#1737 := [lemma #1735]: #1736
+#1801 := (or #1712 #1743 #1754)
+#1802 := [th-lemma]: #1801
+#1804 := [unit-resolution #1802 #1737]: #1803
+#1805 := [unit-resolution #1804 #1800]: #1754
+#1757 := (or #1666 #1074)
+#1748 := [hypothesis]: #1053
+#1397 := (>= #1097 0::int)
+#1751 := (or #1750 #1397)
+#1752 := [th-lemma]: #1751
+#1753 := [unit-resolution #1752 #1749]: #1397
+#1755 := [hypothesis]: #1754
+#1756 := [th-lemma #1755 #1753 #1748]: false
+#1758 := [lemma #1756]: #1757
+#1806 := [unit-resolution #1758 #1805]: #1074
+#1808 := (or #1053 #1038)
+#1039 := (f8 f14 #20)
+#1058 := (= #1037 #1039)
+#1061 := (ite #1053 #1058 #1038)
+#629 := (= #30 #89)
+#628 := (= #33 #89)
+#630 := (ite #195 #628 #629)
+#735 := (forall (vars (?v0 S2)) (:pat #694 #705 #734 #706) #630)
+#633 := (forall (vars (?v0 S2)) #630)
+#738 := (iff #633 #735)
+#736 := (iff #630 #630)
+#737 := [refl]: #736
+#739 := [quant-intro #737]: #738
+#251 := (ite #195 #33 #30)
+#513 := (= #89 #251)
+#514 := (forall (vars (?v0 S2)) #513)
+#634 := (iff #514 #633)
+#631 := (iff #513 #630)
+#632 := [rewrite]: #631
+#635 := [quant-intro #632]: #634
+#533 := (~ #514 #514)
+#531 := (~ #513 #513)
+#532 := [refl]: #531
+#534 := [nnf-pos #532]: #533
+#61 := (f3 f19 #9)
+#256 := (= #61 #251)
+#259 := (forall (vars (?v0 S2)) #256)
+#515 := (iff #259 #514)
+#516 := [rewrite* #441]: #515
+#62 := (ite #42 #30 #33)
+#63 := (= #61 #62)
+#64 := (forall (vars (?v0 S2)) #63)
+#260 := (iff #64 #259)
+#257 := (iff #63 #256)
+#254 := (= #62 #251)
+#248 := (ite #194 #30 #33)
+#252 := (= #248 #251)
+#253 := [rewrite]: #252
+#249 := (= #62 #248)
+#250 := [monotonicity #199]: #249
+#255 := [trans #250 #253]: #254
+#258 := [monotonicity #255]: #257
+#261 := [quant-intro #258]: #260
+#215 := [asserted]: #64
+#262 := [mp #215 #261]: #259
+#517 := [mp #262 #516]: #514
+#535 := [mp~ #517 #534]: #514
+#636 := [mp #535 #635]: #633
+#740 := [mp #636 #739]: #735
+#801 := (not #735)
+#1064 := (or #801 #1061)
+#1040 := (= #1039 #1037)
+#1042 := (+ #1041 #196)
+#1043 := (>= #1042 0::int)
+#1044 := (ite #1043 #1040 #1038)
+#1065 := (or #801 #1044)
+#1067 := (iff #1065 #1064)
+#1069 := (iff #1064 #1064)
+#1070 := [rewrite]: #1069
+#1062 := (iff #1044 #1061)
+#1059 := (iff #1040 #1058)
+#1060 := [rewrite]: #1059
+#1056 := (iff #1043 #1053)
+#1045 := (+ #196 #1041)
+#1048 := (>= #1045 0::int)
+#1054 := (iff #1048 #1053)
+#1055 := [rewrite]: #1054
+#1049 := (iff #1043 #1048)
+#1046 := (= #1042 #1045)
+#1047 := [rewrite]: #1046
+#1050 := [monotonicity #1047]: #1049
+#1057 := [trans #1050 #1055]: #1056
+#1063 := [monotonicity #1057 #1060]: #1062
+#1068 := [monotonicity #1063]: #1067
+#1071 := [trans #1068 #1070]: #1067
+#1066 := [quant-inst]: #1065
+#1072 := [mp #1066 #1071]: #1064
+#1807 := [unit-resolution #1072 #740]: #1061
+#1073 := (not #1061)
+#1077 := (or #1073 #1053 #1038)
+#1078 := [def-axiom]: #1077
+#1809 := [unit-resolution #1078 #1807]: #1808
+#1810 := [unit-resolution #1809 #1806]: #1038
+#1812 := [symm #1810]: #1811
+#1813 := [trans #1812 #1774]: #1697
+#1814 := (not #1697)
+#1815 := (or #1814 #1765)
+#1816 := [th-lemma]: #1815
+#1817 := [unit-resolution #1816 #1813]: #1765
+#1508 := (f8 f22 #20)
+#1522 := (* -1::real #1508)
+#1535 := (+ #1205 #1522)
+#1536 := (<= #1535 0::real)
+#1560 := (not #1536)
+#1085 := (f8 #98 #20)
+#1523 := (+ #1085 #1522)
+#1524 := (>= #1523 0::real)
+#1541 := (or #1524 #1536)
+#1544 := (not #1541)
+#1548 := (or #1547 #1544)
+#1509 := (+ #1508 #1507)
+#1510 := (>= #1509 0::real)
+#1511 := (* -1::real #1085)
+#1512 := (+ #1508 #1511)
+#1513 := (<= #1512 0::real)
+#1514 := (or #1513 #1510)
+#1515 := (not #1514)
+#1549 := (or #1547 #1515)
+#1551 := (iff #1549 #1548)
+#1553 := (iff #1548 #1548)
+#1554 := [rewrite]: #1553
+#1545 := (iff #1515 #1544)
+#1542 := (iff #1514 #1541)
+#1539 := (iff #1510 #1536)
+#1529 := (+ #1507 #1508)
+#1532 := (>= #1529 0::real)
+#1537 := (iff #1532 #1536)
+#1538 := [rewrite]: #1537
+#1533 := (iff #1510 #1532)
+#1530 := (= #1509 #1529)
+#1531 := [rewrite]: #1530
+#1534 := [monotonicity #1531]: #1533
+#1540 := [trans #1534 #1538]: #1539
+#1527 := (iff #1513 #1524)
+#1516 := (+ #1511 #1508)
+#1519 := (<= #1516 0::real)
+#1525 := (iff #1519 #1524)
+#1526 := [rewrite]: #1525
+#1520 := (iff #1513 #1519)
+#1517 := (= #1512 #1516)
+#1518 := [rewrite]: #1517
+#1521 := [monotonicity #1518]: #1520
+#1528 := [trans #1521 #1526]: #1527
+#1543 := [monotonicity #1528 #1540]: #1542
+#1546 := [monotonicity #1543]: #1545
+#1552 := [monotonicity #1546]: #1551
+#1555 := [trans #1552 #1554]: #1551
+#1550 := [quant-inst]: #1549
+#1556 := [mp #1550 #1555]: #1548
+#1818 := [unit-resolution #1556 #773]: #1544
+#1561 := (or #1541 #1560)
+#1562 := [def-axiom]: #1561
+#1819 := [unit-resolution #1562 #1818]: #1560
+#1623 := (+ #1037 #1522)
+#1624 := (>= #1623 0::real)
+#1654 := (not #1624)
+#1158 := (f8 #93 #20)
+#1633 := (+ #1158 #1522)
+#1634 := (<= #1633 0::real)
+#1639 := (or #1624 #1634)
+#1642 := (not #1639)
+#1645 := (or #1489 #1642)
+#1620 := (* -1::real #1158)
+#1621 := (+ #1508 #1620)
+#1622 := (>= #1621 0::real)
+#1625 := (or #1624 #1622)
+#1626 := (not #1625)
+#1646 := (or #1489 #1626)
+#1648 := (iff #1646 #1645)
+#1650 := (iff #1645 #1645)
+#1651 := [rewrite]: #1650
+#1643 := (iff #1626 #1642)
+#1640 := (iff #1625 #1639)
+#1637 := (iff #1622 #1634)
+#1627 := (+ #1620 #1508)
+#1630 := (>= #1627 0::real)
+#1635 := (iff #1630 #1634)
+#1636 := [rewrite]: #1635
+#1631 := (iff #1622 #1630)
+#1628 := (= #1621 #1627)
+#1629 := [rewrite]: #1628
+#1632 := [monotonicity #1629]: #1631
+#1638 := [trans #1632 #1636]: #1637
+#1641 := [monotonicity #1638]: #1640
+#1644 := [monotonicity #1641]: #1643
+#1649 := [monotonicity #1644]: #1648
+#1652 := [trans #1649 #1651]: #1648
+#1647 := [quant-inst]: #1646
+#1653 := [mp #1647 #1652]: #1645
+#1820 := [unit-resolution #1653 #767]: #1642
+#1655 := (or #1639 #1654)
+#1656 := [def-axiom]: #1655
+#1821 := [unit-resolution #1656 #1820]: #1654
+[th-lemma #1821 #1819 #1817]: false
+unsat
+6c73093b27236ef09bc4a53162dee78b6dc31895 422 0
+#2 := false
+decl f12 :: S2
+#42 := f12
+decl f5 :: S2
+#25 := f5
+#45 := (= f5 f12)
+decl f3 :: (-> int S2)
+decl f4 :: (-> S2 int)
+#43 := (f4 f12)
+#598 := (f3 #43)
+#696 := (= #598 f12)
+#599 := (= f12 #598)
+#8 := (:var 0 S2)
+#9 := (f4 #8)
+#551 := (pattern #9)
+#10 := (f3 #9)
+#98 := (= #8 #10)
+#552 := (forall (vars (?v0 S2)) (:pat #551) #98)
+#101 := (forall (vars (?v0 S2)) #98)
+#555 := (iff #101 #552)
+#553 := (iff #98 #98)
+#554 := [refl]: #553
+#556 := [quant-intro #554]: #555
+#455 := (~ #101 #101)
+#457 := (~ #98 #98)
+#458 := [refl]: #457
+#456 := [nnf-pos #458]: #455
+#11 := (= #10 #8)
+#12 := (forall (vars (?v0 S2)) #11)
+#102 := (iff #12 #101)
+#99 := (iff #11 #98)
+#100 := [rewrite]: #99
+#103 := [quant-intro #100]: #102
+#97 := [asserted]: #12
+#106 := [mp #97 #103]: #101
+#453 := [mp~ #106 #456]: #101
+#557 := [mp #453 #556]: #552
+#600 := (not #552)
+#605 := (or #600 #599)
+#606 := [quant-inst]: #605
+#690 := [unit-resolution #606 #557]: #599
+#697 := [symm #690]: #696
+#698 := (= f5 #598)
+#26 := (f4 f5)
+#596 := (f3 #26)
+#694 := (= #596 #598)
+#692 := (= #598 #596)
+#688 := (= #43 #26)
+#686 := (= #26 #43)
+#13 := 0::int
+#231 := -1::int
+#234 := (* -1::int #43)
+#235 := (+ #26 #234)
+#295 := (<= #235 0::int)
+#74 := (<= #26 #43)
+#393 := (iff #74 #295)
+#394 := [rewrite]: #393
+#346 := [asserted]: #74
+#395 := [mp #346 #394]: #295
+#233 := (>= #235 0::int)
+decl f6 :: (-> S3 S4 real)
+decl f8 :: (-> S2 S4)
+#29 := (f8 f5)
+decl f7 :: S3
+#28 := f7
+#30 := (f6 f7 #29)
+decl f9 :: S3
+#31 := f9
+#32 := (f6 f9 #29)
+#46 := (f8 f12)
+decl f11 :: S3
+#37 := f11
+#47 := (f6 f11 #46)
+#48 := (ite #45 #47 #32)
+#241 := (ite #233 #48 #30)
+#572 := (= #30 #241)
+#658 := (not #572)
+#199 := 0::real
+#197 := -1::real
+#249 := (* -1::real #241)
+#647 := (+ #30 #249)
+#648 := (<= #647 0::real)
+#652 := (not #648)
+#650 := [hypothesis]: #648
+decl f10 :: S3
+#34 := f10
+#35 := (f6 f10 #29)
+#250 := (+ #35 #249)
+#251 := (<= #250 0::real)
+#252 := (not #251)
+#38 := (f6 f11 #29)
+decl f13 :: S3
+#51 := f13
+#52 := (f6 f13 #29)
+#260 := (ite #233 #52 #38)
+#269 := (* -1::real #260)
+#270 := (+ #35 #269)
+#268 := (>= #270 0::real)
+#271 := (not #268)
+#276 := (and #252 #271)
+#44 := (< #26 #43)
+#53 := (ite #44 #38 #52)
+#54 := (< #35 #53)
+#49 := (ite #44 #30 #48)
+#50 := (< #49 #35)
+#55 := (and #50 #54)
+#277 := (iff #55 #276)
+#274 := (iff #54 #271)
+#265 := (< #35 #260)
+#272 := (iff #265 #271)
+#273 := [rewrite]: #272
+#266 := (iff #54 #265)
+#263 := (= #53 #260)
+#232 := (not #233)
+#257 := (ite #232 #38 #52)
+#261 := (= #257 #260)
+#262 := [rewrite]: #261
+#258 := (= #53 #257)
+#236 := (iff #44 #232)
+#237 := [rewrite]: #236
+#259 := [monotonicity #237]: #258
+#264 := [trans #259 #262]: #263
+#267 := [monotonicity #264]: #266
+#275 := [trans #267 #273]: #274
+#255 := (iff #50 #252)
+#246 := (< #241 #35)
+#253 := (iff #246 #252)
+#254 := [rewrite]: #253
+#247 := (iff #50 #246)
+#244 := (= #49 #241)
+#238 := (ite #232 #30 #48)
+#242 := (= #238 #241)
+#243 := [rewrite]: #242
+#239 := (= #49 #238)
+#240 := [monotonicity #237]: #239
+#245 := [trans #240 #243]: #244
+#248 := [monotonicity #245]: #247
+#256 := [trans #248 #254]: #255
+#278 := [monotonicity #256 #275]: #277
+#183 := [asserted]: #55
+#279 := [mp #183 #278]: #276
+#280 := [and-elim #279]: #252
+#201 := (* -1::real #35)
+#217 := (+ #30 #201)
+#218 := (<= #217 0::real)
+#219 := (not #218)
+#202 := (+ #32 #201)
+#200 := (>= #202 0::real)
+#198 := (not #200)
+#224 := (and #198 #219)
+#27 := (< #26 #26)
+#39 := (ite #27 #38 #30)
+#40 := (< #35 #39)
+#33 := (ite #27 #30 #32)
+#36 := (< #33 #35)
+#41 := (and #36 #40)
+#225 := (iff #41 #224)
+#222 := (iff #40 #219)
+#214 := (< #35 #30)
+#220 := (iff #214 #219)
+#221 := [rewrite]: #220
+#215 := (iff #40 #214)
+#212 := (= #39 #30)
+#207 := (ite false #38 #30)
+#210 := (= #207 #30)
+#211 := [rewrite]: #210
+#208 := (= #39 #207)
+#185 := (iff #27 false)
+#186 := [rewrite]: #185
+#209 := [monotonicity #186]: #208
+#213 := [trans #209 #211]: #212
+#216 := [monotonicity #213]: #215
+#223 := [trans #216 #221]: #222
+#205 := (iff #36 #198)
+#194 := (< #32 #35)
+#203 := (iff #194 #198)
+#204 := [rewrite]: #203
+#195 := (iff #36 #194)
+#192 := (= #33 #32)
+#187 := (ite false #30 #32)
+#190 := (= #187 #32)
+#191 := [rewrite]: #190
+#188 := (= #33 #187)
+#189 := [monotonicity #186]: #188
+#193 := [trans #189 #191]: #192
+#196 := [monotonicity #193]: #195
+#206 := [trans #196 #204]: #205
+#226 := [monotonicity #206 #223]: #225
+#182 := [asserted]: #41
+#227 := [mp #182 #226]: #224
+#229 := [and-elim #227]: #219
+#651 := [th-lemma #229 #280 #650]: false
+#653 := [lemma #651]: #652
+#657 := [hypothesis]: #572
+#659 := (or #658 #648)
+#660 := [th-lemma]: #659
+#661 := [unit-resolution #660 #657 #653]: false
+#662 := [lemma #661]: #658
+#582 := (or #233 #572)
+#583 := [def-axiom]: #582
+#685 := [unit-resolution #583 #662]: #233
+#687 := [th-lemma #685 #395]: #686
+#689 := [symm #687]: #688
+#693 := [monotonicity #689]: #692
+#695 := [symm #693]: #694
+#597 := (= f5 #596)
+#601 := (or #600 #597)
+#602 := [quant-inst]: #601
+#691 := [unit-resolution #602 #557]: #597
+#699 := [trans #691 #695]: #698
+#700 := [trans #699 #697]: #45
+#575 := (not #45)
+#63 := (f6 f13 #46)
+#283 := (ite #45 #30 #63)
+#466 := (* -1::real #283)
+#642 := (+ #30 #466)
+#644 := (>= #642 0::real)
+#590 := (= #30 #283)
+#666 := [hypothesis]: #45
+#592 := (or #575 #590)
+#593 := [def-axiom]: #592
+#667 := [unit-resolution #593 #666]: #590
+#668 := (not #590)
+#669 := (or #668 #644)
+#670 := [th-lemma]: #669
+#671 := [unit-resolution #670 #667]: #644
+#60 := (f6 f10 #46)
+#362 := (* -1::real #60)
+#363 := (+ #47 #362)
+#361 := (>= #363 0::real)
+#360 := (not #361)
+#379 := (* -1::real #63)
+#380 := (+ #60 #379)
+#378 := (>= #380 0::real)
+#381 := (not #378)
+#386 := (and #360 #381)
+#68 := (< #43 #43)
+#71 := (ite #68 #47 #63)
+#72 := (< #60 #71)
+#57 := (f6 f7 #46)
+#69 := (ite #68 #57 #47)
+#70 := (< #69 #60)
+#73 := (and #70 #72)
+#387 := (iff #73 #386)
+#384 := (iff #72 #381)
+#375 := (< #60 #63)
+#382 := (iff #375 #381)
+#383 := [rewrite]: #382
+#376 := (iff #72 #375)
+#373 := (= #71 #63)
+#368 := (ite false #47 #63)
+#371 := (= #368 #63)
+#372 := [rewrite]: #371
+#369 := (= #71 #368)
+#348 := (iff #68 false)
+#349 := [rewrite]: #348
+#370 := [monotonicity #349]: #369
+#374 := [trans #370 #372]: #373
+#377 := [monotonicity #374]: #376
+#385 := [trans #377 #383]: #384
+#366 := (iff #70 #360)
+#357 := (< #47 #60)
+#364 := (iff #357 #360)
+#365 := [rewrite]: #364
+#358 := (iff #70 #357)
+#355 := (= #69 #47)
+#350 := (ite false #57 #47)
+#353 := (= #350 #47)
+#354 := [rewrite]: #353
+#351 := (= #69 #350)
+#352 := [monotonicity #349]: #351
+#356 := [trans #352 #354]: #355
+#359 := [monotonicity #356]: #358
+#367 := [trans #359 #365]: #366
+#388 := [monotonicity #367 #385]: #387
+#345 := [asserted]: #73
+#389 := [mp #345 #388]: #386
+#390 := [and-elim #389]: #360
+#399 := (* -1::real #57)
+#400 := (+ #47 #399)
+#398 := (>= #400 0::real)
+#58 := (f6 f9 #46)
+#407 := (* -1::real #58)
+#408 := (+ #57 #407)
+#406 := (>= #408 0::real)
+#402 := (+ #47 #379)
+#403 := (<= #402 0::real)
+#417 := (and #398 #403 #406)
+#77 := (<= #47 #63)
+#76 := (<= #57 #47)
+#78 := (and #76 #77)
+#75 := (<= #58 #57)
+#79 := (and #75 #78)
+#420 := (iff #79 #417)
+#411 := (and #398 #403)
+#414 := (and #406 #411)
+#418 := (iff #414 #417)
+#419 := [rewrite]: #418
+#415 := (iff #79 #414)
+#412 := (iff #78 #411)
+#404 := (iff #77 #403)
+#405 := [rewrite]: #404
+#397 := (iff #76 #398)
+#401 := [rewrite]: #397
+#413 := [monotonicity #401 #405]: #412
+#409 := (iff #75 #406)
+#410 := [rewrite]: #409
+#416 := [monotonicity #410 #413]: #415
+#421 := [trans #416 #419]: #420
+#347 := [asserted]: #79
+#422 := [mp #347 #421]: #417
+#423 := [and-elim #422]: #398
+#655 := (+ #30 #399)
+#656 := (<= #655 0::real)
+#654 := (= #30 #57)
+#676 := (= #57 #30)
+#674 := (= #46 #29)
+#672 := (= #29 #46)
+#673 := [monotonicity #666]: #672
+#675 := [symm #673]: #674
+#677 := [monotonicity #675]: #676
+#678 := [symm #677]: #654
+#679 := (not #654)
+#680 := (or #679 #656)
+#681 := [th-lemma]: #680
+#682 := [unit-resolution #681 #678]: #656
+#469 := (+ #60 #466)
+#472 := (>= #469 0::real)
+#445 := (not #472)
+#321 := (ite #295 #283 #47)
+#331 := (* -1::real #321)
+#332 := (+ #60 #331)
+#330 := (>= #332 0::real)
+#329 := (not #330)
+#446 := (iff #329 #445)
+#473 := (iff #330 #472)
+#470 := (= #332 #469)
+#467 := (= #331 #466)
+#464 := (= #321 #283)
+#1 := true
+#459 := (ite true #283 #47)
+#462 := (= #459 #283)
+#463 := [rewrite]: #462
+#460 := (= #321 #459)
+#451 := (iff #295 true)
+#452 := [iff-true #395]: #451
+#461 := [monotonicity #452]: #460
+#465 := [trans #461 #463]: #464
+#468 := [monotonicity #465]: #467
+#471 := [monotonicity #468]: #470
+#474 := [monotonicity #471]: #473
+#475 := [monotonicity #474]: #446
+#302 := (ite #295 #58 #57)
+#310 := (* -1::real #302)
+#311 := (+ #60 #310)
+#312 := (<= #311 0::real)
+#313 := (not #312)
+#337 := (and #313 #329)
+#62 := (= f12 f5)
+#64 := (ite #62 #30 #63)
+#56 := (< #43 #26)
+#65 := (ite #56 #47 #64)
+#66 := (< #60 #65)
+#59 := (ite #56 #57 #58)
+#61 := (< #59 #60)
+#67 := (and #61 #66)
+#340 := (iff #67 #337)
+#286 := (ite #56 #47 #283)
+#289 := (< #60 #286)
+#292 := (and #61 #289)
+#338 := (iff #292 #337)
+#335 := (iff #289 #329)
+#326 := (< #60 #321)
+#333 := (iff #326 #329)
+#334 := [rewrite]: #333
+#327 := (iff #289 #326)
+#324 := (= #286 #321)
+#296 := (not #295)
+#318 := (ite #296 #47 #283)
+#322 := (= #318 #321)
+#323 := [rewrite]: #322
+#319 := (= #286 #318)
+#297 := (iff #56 #296)
+#298 := [rewrite]: #297
+#320 := [monotonicity #298]: #319
+#325 := [trans #320 #323]: #324
+#328 := [monotonicity #325]: #327
+#336 := [trans #328 #334]: #335
+#316 := (iff #61 #313)
+#307 := (< #302 #60)
+#314 := (iff #307 #313)
+#315 := [rewrite]: #314
+#308 := (iff #61 #307)
+#305 := (= #59 #302)
+#299 := (ite #296 #57 #58)
+#303 := (= #299 #302)
+#304 := [rewrite]: #303
+#300 := (= #59 #299)
+#301 := [monotonicity #298]: #300
+#306 := [trans #301 #304]: #305
+#309 := [monotonicity #306]: #308
+#317 := [trans #309 #315]: #316
+#339 := [monotonicity #317 #336]: #338
+#293 := (iff #67 #292)
+#290 := (iff #66 #289)
+#287 := (= #65 #286)
+#284 := (= #64 #283)
+#230 := (iff #62 #45)
+#282 := [rewrite]: #230
+#285 := [monotonicity #282]: #284
+#288 := [monotonicity #285]: #287
+#291 := [monotonicity #288]: #290
+#294 := [monotonicity #291]: #293
+#341 := [trans #294 #339]: #340
+#184 := [asserted]: #67
+#342 := [mp #184 #341]: #337
+#344 := [and-elim #342]: #329
+#476 := [mp #344 #475]: #445
+#683 := [th-lemma #476 #682 #423 #390 #671]: false
+#684 := [lemma #683]: #575
+[unit-resolution #684 #700]: false
+unsat
 5ee060971856d2def7cc6d40549073dace7efe45 428 0
 #2 := false
 decl f12 :: S2
@@ -427,1351 +1761,6 @@
 #679 := [lemma #678]: #576
 [unit-resolution #679 #695]: false
 unsat
-6c73093b27236ef09bc4a53162dee78b6dc31895 422 0
-#2 := false
-decl f12 :: S2
-#42 := f12
-decl f5 :: S2
-#25 := f5
-#45 := (= f5 f12)
-decl f3 :: (-> int S2)
-decl f4 :: (-> S2 int)
-#43 := (f4 f12)
-#598 := (f3 #43)
-#696 := (= #598 f12)
-#599 := (= f12 #598)
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#551 := (pattern #9)
-#10 := (f3 #9)
-#98 := (= #8 #10)
-#552 := (forall (vars (?v0 S2)) (:pat #551) #98)
-#101 := (forall (vars (?v0 S2)) #98)
-#555 := (iff #101 #552)
-#553 := (iff #98 #98)
-#554 := [refl]: #553
-#556 := [quant-intro #554]: #555
-#455 := (~ #101 #101)
-#457 := (~ #98 #98)
-#458 := [refl]: #457
-#456 := [nnf-pos #458]: #455
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#102 := (iff #12 #101)
-#99 := (iff #11 #98)
-#100 := [rewrite]: #99
-#103 := [quant-intro #100]: #102
-#97 := [asserted]: #12
-#106 := [mp #97 #103]: #101
-#453 := [mp~ #106 #456]: #101
-#557 := [mp #453 #556]: #552
-#600 := (not #552)
-#605 := (or #600 #599)
-#606 := [quant-inst]: #605
-#690 := [unit-resolution #606 #557]: #599
-#697 := [symm #690]: #696
-#698 := (= f5 #598)
-#26 := (f4 f5)
-#596 := (f3 #26)
-#694 := (= #596 #598)
-#692 := (= #598 #596)
-#688 := (= #43 #26)
-#686 := (= #26 #43)
-#13 := 0::int
-#231 := -1::int
-#234 := (* -1::int #43)
-#235 := (+ #26 #234)
-#295 := (<= #235 0::int)
-#74 := (<= #26 #43)
-#393 := (iff #74 #295)
-#394 := [rewrite]: #393
-#346 := [asserted]: #74
-#395 := [mp #346 #394]: #295
-#233 := (>= #235 0::int)
-decl f6 :: (-> S3 S4 real)
-decl f8 :: (-> S2 S4)
-#29 := (f8 f5)
-decl f7 :: S3
-#28 := f7
-#30 := (f6 f7 #29)
-decl f9 :: S3
-#31 := f9
-#32 := (f6 f9 #29)
-#46 := (f8 f12)
-decl f11 :: S3
-#37 := f11
-#47 := (f6 f11 #46)
-#48 := (ite #45 #47 #32)
-#241 := (ite #233 #48 #30)
-#572 := (= #30 #241)
-#658 := (not #572)
-#199 := 0::real
-#197 := -1::real
-#249 := (* -1::real #241)
-#647 := (+ #30 #249)
-#648 := (<= #647 0::real)
-#652 := (not #648)
-#650 := [hypothesis]: #648
-decl f10 :: S3
-#34 := f10
-#35 := (f6 f10 #29)
-#250 := (+ #35 #249)
-#251 := (<= #250 0::real)
-#252 := (not #251)
-#38 := (f6 f11 #29)
-decl f13 :: S3
-#51 := f13
-#52 := (f6 f13 #29)
-#260 := (ite #233 #52 #38)
-#269 := (* -1::real #260)
-#270 := (+ #35 #269)
-#268 := (>= #270 0::real)
-#271 := (not #268)
-#276 := (and #252 #271)
-#44 := (< #26 #43)
-#53 := (ite #44 #38 #52)
-#54 := (< #35 #53)
-#49 := (ite #44 #30 #48)
-#50 := (< #49 #35)
-#55 := (and #50 #54)
-#277 := (iff #55 #276)
-#274 := (iff #54 #271)
-#265 := (< #35 #260)
-#272 := (iff #265 #271)
-#273 := [rewrite]: #272
-#266 := (iff #54 #265)
-#263 := (= #53 #260)
-#232 := (not #233)
-#257 := (ite #232 #38 #52)
-#261 := (= #257 #260)
-#262 := [rewrite]: #261
-#258 := (= #53 #257)
-#236 := (iff #44 #232)
-#237 := [rewrite]: #236
-#259 := [monotonicity #237]: #258
-#264 := [trans #259 #262]: #263
-#267 := [monotonicity #264]: #266
-#275 := [trans #267 #273]: #274
-#255 := (iff #50 #252)
-#246 := (< #241 #35)
-#253 := (iff #246 #252)
-#254 := [rewrite]: #253
-#247 := (iff #50 #246)
-#244 := (= #49 #241)
-#238 := (ite #232 #30 #48)
-#242 := (= #238 #241)
-#243 := [rewrite]: #242
-#239 := (= #49 #238)
-#240 := [monotonicity #237]: #239
-#245 := [trans #240 #243]: #244
-#248 := [monotonicity #245]: #247
-#256 := [trans #248 #254]: #255
-#278 := [monotonicity #256 #275]: #277
-#183 := [asserted]: #55
-#279 := [mp #183 #278]: #276
-#280 := [and-elim #279]: #252
-#201 := (* -1::real #35)
-#217 := (+ #30 #201)
-#218 := (<= #217 0::real)
-#219 := (not #218)
-#202 := (+ #32 #201)
-#200 := (>= #202 0::real)
-#198 := (not #200)
-#224 := (and #198 #219)
-#27 := (< #26 #26)
-#39 := (ite #27 #38 #30)
-#40 := (< #35 #39)
-#33 := (ite #27 #30 #32)
-#36 := (< #33 #35)
-#41 := (and #36 #40)
-#225 := (iff #41 #224)
-#222 := (iff #40 #219)
-#214 := (< #35 #30)
-#220 := (iff #214 #219)
-#221 := [rewrite]: #220
-#215 := (iff #40 #214)
-#212 := (= #39 #30)
-#207 := (ite false #38 #30)
-#210 := (= #207 #30)
-#211 := [rewrite]: #210
-#208 := (= #39 #207)
-#185 := (iff #27 false)
-#186 := [rewrite]: #185
-#209 := [monotonicity #186]: #208
-#213 := [trans #209 #211]: #212
-#216 := [monotonicity #213]: #215
-#223 := [trans #216 #221]: #222
-#205 := (iff #36 #198)
-#194 := (< #32 #35)
-#203 := (iff #194 #198)
-#204 := [rewrite]: #203
-#195 := (iff #36 #194)
-#192 := (= #33 #32)
-#187 := (ite false #30 #32)
-#190 := (= #187 #32)
-#191 := [rewrite]: #190
-#188 := (= #33 #187)
-#189 := [monotonicity #186]: #188
-#193 := [trans #189 #191]: #192
-#196 := [monotonicity #193]: #195
-#206 := [trans #196 #204]: #205
-#226 := [monotonicity #206 #223]: #225
-#182 := [asserted]: #41
-#227 := [mp #182 #226]: #224
-#229 := [and-elim #227]: #219
-#651 := [th-lemma #229 #280 #650]: false
-#653 := [lemma #651]: #652
-#657 := [hypothesis]: #572
-#659 := (or #658 #648)
-#660 := [th-lemma]: #659
-#661 := [unit-resolution #660 #657 #653]: false
-#662 := [lemma #661]: #658
-#582 := (or #233 #572)
-#583 := [def-axiom]: #582
-#685 := [unit-resolution #583 #662]: #233
-#687 := [th-lemma #685 #395]: #686
-#689 := [symm #687]: #688
-#693 := [monotonicity #689]: #692
-#695 := [symm #693]: #694
-#597 := (= f5 #596)
-#601 := (or #600 #597)
-#602 := [quant-inst]: #601
-#691 := [unit-resolution #602 #557]: #597
-#699 := [trans #691 #695]: #698
-#700 := [trans #699 #697]: #45
-#575 := (not #45)
-#63 := (f6 f13 #46)
-#283 := (ite #45 #30 #63)
-#466 := (* -1::real #283)
-#642 := (+ #30 #466)
-#644 := (>= #642 0::real)
-#590 := (= #30 #283)
-#666 := [hypothesis]: #45
-#592 := (or #575 #590)
-#593 := [def-axiom]: #592
-#667 := [unit-resolution #593 #666]: #590
-#668 := (not #590)
-#669 := (or #668 #644)
-#670 := [th-lemma]: #669
-#671 := [unit-resolution #670 #667]: #644
-#60 := (f6 f10 #46)
-#362 := (* -1::real #60)
-#363 := (+ #47 #362)
-#361 := (>= #363 0::real)
-#360 := (not #361)
-#379 := (* -1::real #63)
-#380 := (+ #60 #379)
-#378 := (>= #380 0::real)
-#381 := (not #378)
-#386 := (and #360 #381)
-#68 := (< #43 #43)
-#71 := (ite #68 #47 #63)
-#72 := (< #60 #71)
-#57 := (f6 f7 #46)
-#69 := (ite #68 #57 #47)
-#70 := (< #69 #60)
-#73 := (and #70 #72)
-#387 := (iff #73 #386)
-#384 := (iff #72 #381)
-#375 := (< #60 #63)
-#382 := (iff #375 #381)
-#383 := [rewrite]: #382
-#376 := (iff #72 #375)
-#373 := (= #71 #63)
-#368 := (ite false #47 #63)
-#371 := (= #368 #63)
-#372 := [rewrite]: #371
-#369 := (= #71 #368)
-#348 := (iff #68 false)
-#349 := [rewrite]: #348
-#370 := [monotonicity #349]: #369
-#374 := [trans #370 #372]: #373
-#377 := [monotonicity #374]: #376
-#385 := [trans #377 #383]: #384
-#366 := (iff #70 #360)
-#357 := (< #47 #60)
-#364 := (iff #357 #360)
-#365 := [rewrite]: #364
-#358 := (iff #70 #357)
-#355 := (= #69 #47)
-#350 := (ite false #57 #47)
-#353 := (= #350 #47)
-#354 := [rewrite]: #353
-#351 := (= #69 #350)
-#352 := [monotonicity #349]: #351
-#356 := [trans #352 #354]: #355
-#359 := [monotonicity #356]: #358
-#367 := [trans #359 #365]: #366
-#388 := [monotonicity #367 #385]: #387
-#345 := [asserted]: #73
-#389 := [mp #345 #388]: #386
-#390 := [and-elim #389]: #360
-#399 := (* -1::real #57)
-#400 := (+ #47 #399)
-#398 := (>= #400 0::real)
-#58 := (f6 f9 #46)
-#407 := (* -1::real #58)
-#408 := (+ #57 #407)
-#406 := (>= #408 0::real)
-#402 := (+ #47 #379)
-#403 := (<= #402 0::real)
-#417 := (and #398 #403 #406)
-#77 := (<= #47 #63)
-#76 := (<= #57 #47)
-#78 := (and #76 #77)
-#75 := (<= #58 #57)
-#79 := (and #75 #78)
-#420 := (iff #79 #417)
-#411 := (and #398 #403)
-#414 := (and #406 #411)
-#418 := (iff #414 #417)
-#419 := [rewrite]: #418
-#415 := (iff #79 #414)
-#412 := (iff #78 #411)
-#404 := (iff #77 #403)
-#405 := [rewrite]: #404
-#397 := (iff #76 #398)
-#401 := [rewrite]: #397
-#413 := [monotonicity #401 #405]: #412
-#409 := (iff #75 #406)
-#410 := [rewrite]: #409
-#416 := [monotonicity #410 #413]: #415
-#421 := [trans #416 #419]: #420
-#347 := [asserted]: #79
-#422 := [mp #347 #421]: #417
-#423 := [and-elim #422]: #398
-#655 := (+ #30 #399)
-#656 := (<= #655 0::real)
-#654 := (= #30 #57)
-#676 := (= #57 #30)
-#674 := (= #46 #29)
-#672 := (= #29 #46)
-#673 := [monotonicity #666]: #672
-#675 := [symm #673]: #674
-#677 := [monotonicity #675]: #676
-#678 := [symm #677]: #654
-#679 := (not #654)
-#680 := (or #679 #656)
-#681 := [th-lemma]: #680
-#682 := [unit-resolution #681 #678]: #656
-#469 := (+ #60 #466)
-#472 := (>= #469 0::real)
-#445 := (not #472)
-#321 := (ite #295 #283 #47)
-#331 := (* -1::real #321)
-#332 := (+ #60 #331)
-#330 := (>= #332 0::real)
-#329 := (not #330)
-#446 := (iff #329 #445)
-#473 := (iff #330 #472)
-#470 := (= #332 #469)
-#467 := (= #331 #466)
-#464 := (= #321 #283)
-#1 := true
-#459 := (ite true #283 #47)
-#462 := (= #459 #283)
-#463 := [rewrite]: #462
-#460 := (= #321 #459)
-#451 := (iff #295 true)
-#452 := [iff-true #395]: #451
-#461 := [monotonicity #452]: #460
-#465 := [trans #461 #463]: #464
-#468 := [monotonicity #465]: #467
-#471 := [monotonicity #468]: #470
-#474 := [monotonicity #471]: #473
-#475 := [monotonicity #474]: #446
-#302 := (ite #295 #58 #57)
-#310 := (* -1::real #302)
-#311 := (+ #60 #310)
-#312 := (<= #311 0::real)
-#313 := (not #312)
-#337 := (and #313 #329)
-#62 := (= f12 f5)
-#64 := (ite #62 #30 #63)
-#56 := (< #43 #26)
-#65 := (ite #56 #47 #64)
-#66 := (< #60 #65)
-#59 := (ite #56 #57 #58)
-#61 := (< #59 #60)
-#67 := (and #61 #66)
-#340 := (iff #67 #337)
-#286 := (ite #56 #47 #283)
-#289 := (< #60 #286)
-#292 := (and #61 #289)
-#338 := (iff #292 #337)
-#335 := (iff #289 #329)
-#326 := (< #60 #321)
-#333 := (iff #326 #329)
-#334 := [rewrite]: #333
-#327 := (iff #289 #326)
-#324 := (= #286 #321)
-#296 := (not #295)
-#318 := (ite #296 #47 #283)
-#322 := (= #318 #321)
-#323 := [rewrite]: #322
-#319 := (= #286 #318)
-#297 := (iff #56 #296)
-#298 := [rewrite]: #297
-#320 := [monotonicity #298]: #319
-#325 := [trans #320 #323]: #324
-#328 := [monotonicity #325]: #327
-#336 := [trans #328 #334]: #335
-#316 := (iff #61 #313)
-#307 := (< #302 #60)
-#314 := (iff #307 #313)
-#315 := [rewrite]: #314
-#308 := (iff #61 #307)
-#305 := (= #59 #302)
-#299 := (ite #296 #57 #58)
-#303 := (= #299 #302)
-#304 := [rewrite]: #303
-#300 := (= #59 #299)
-#301 := [monotonicity #298]: #300
-#306 := [trans #301 #304]: #305
-#309 := [monotonicity #306]: #308
-#317 := [trans #309 #315]: #316
-#339 := [monotonicity #317 #336]: #338
-#293 := (iff #67 #292)
-#290 := (iff #66 #289)
-#287 := (= #65 #286)
-#284 := (= #64 #283)
-#230 := (iff #62 #45)
-#282 := [rewrite]: #230
-#285 := [monotonicity #282]: #284
-#288 := [monotonicity #285]: #287
-#291 := [monotonicity #288]: #290
-#294 := [monotonicity #291]: #293
-#341 := [trans #294 #339]: #340
-#184 := [asserted]: #67
-#342 := [mp #184 #341]: #337
-#344 := [and-elim #342]: #329
-#476 := [mp #344 #475]: #445
-#683 := [th-lemma #476 #682 #423 #390 #671]: false
-#684 := [lemma #683]: #575
-[unit-resolution #684 #700]: false
-unsat
-c220892677421b557c184d2f3de28c1bae1b8341 921 0
-#2 := false
-#58 := 0::int
-decl f5 :: (-> S4 int)
-decl f6 :: (-> S2 S4)
-decl f11 :: (-> S4 S2)
-decl f14 :: S4
-#30 := f14
-#34 := (f11 f14)
-#70 := (f6 #34)
-#605 := (f5 #70)
-#121 := -1::int
-#615 := (* -1::int #605)
-decl f7 :: S4
-#13 := f7
-#14 := (f5 f7)
-#662 := (+ #14 #615)
-#663 := (<= #662 0::int)
-decl f8 :: (-> S5 S2 real)
-decl f19 :: (-> S3 S5)
-decl f15 :: S3
-#40 := f15
-#86 := (f19 f15)
-#650 := (f8 #86 #34)
-decl f10 :: S5
-#19 := f10
-#35 := (f8 f10 #34)
-#651 := (= #35 #650)
-#690 := (not #651)
-decl f13 :: S3
-#28 := f13
-#81 := (f19 f13)
-#749 := (f8 #81 #34)
-#1233 := (= #650 #749)
-#1246 := (not #1233)
-#1335 := (iff #1246 #690)
-#1333 := (iff #1233 #651)
-#1328 := (= #650 #35)
-#1331 := (iff #1328 #651)
-#1332 := [commutativity]: #1331
-#1329 := (iff #1233 #1328)
-#1326 := (= #749 #35)
-#752 := (= #35 #749)
-decl f12 :: S5
-#22 := f12
-#696 := (f8 f12 #34)
-#751 := (= #696 #749)
-#281 := (= f14 #70)
-#755 := (ite #281 #752 #751)
-decl f9 :: S5
-#16 := f9
-#694 := (f8 f9 #34)
-#750 := (= #694 #749)
-#31 := (f5 f14)
-#616 := (+ #31 #615)
-#617 := (<= #616 0::int)
-#758 := (ite #617 #755 #750)
-#9 := (:var 0 S2)
-#17 := (f8 f9 #9)
-#538 := (pattern #17)
-#23 := (f8 f12 #9)
-#537 := (pattern #23)
-#82 := (f8 #81 #9)
-#545 := (pattern #82)
-#11 := (f6 #9)
-#535 := (pattern #11)
-#452 := (= #17 #82)
-#450 := (= #23 #82)
-#449 := (= #35 #82)
-#33 := (= #11 f14)
-#451 := (ite #33 #449 #450)
-#146 := (* -1::int #31)
-#12 := (f5 #11)
-#147 := (+ #12 #146)
-#145 := (>= #147 0::int)
-#453 := (ite #145 #451 #452)
-#546 := (forall (vars (?v0 S2)) (:pat #535 #545 #537 #538) #453)
-#456 := (forall (vars (?v0 S2)) #453)
-#549 := (iff #456 #546)
-#547 := (iff #453 #453)
-#548 := [refl]: #547
-#550 := [quant-intro #548]: #549
-#36 := (ite #33 #35 #23)
-#153 := (ite #145 #36 #17)
-#380 := (= #82 #153)
-#381 := (forall (vars (?v0 S2)) #380)
-#457 := (iff #381 #456)
-#454 := (iff #380 #453)
-#455 := [rewrite]: #454
-#458 := [quant-intro #455]: #457
-#366 := (~ #381 #381)
-#368 := (~ #380 #380)
-#365 := [refl]: #368
-#363 := [nnf-pos #365]: #366
-decl f3 :: (-> S3 S2 real)
-#29 := (f3 f13 #9)
-#158 := (= #29 #153)
-#161 := (forall (vars (?v0 S2)) #158)
-#382 := (iff #161 #381)
-#96 := (:var 1 S3)
-#99 := (f3 #96 #9)
-#97 := (f19 #96)
-#98 := (f8 #97 #9)
-#100 := (= #98 #99)
-#101 := (forall (vars (?v0 S3) (?v1 S2)) #100)
-#298 := [asserted]: #101
-#383 := [rewrite* #298]: #382
-#32 := (< #12 #31)
-#37 := (ite #32 #17 #36)
-#38 := (= #29 #37)
-#39 := (forall (vars (?v0 S2)) #38)
-#162 := (iff #39 #161)
-#159 := (iff #38 #158)
-#156 := (= #37 #153)
-#144 := (not #145)
-#150 := (ite #144 #17 #36)
-#154 := (= #150 #153)
-#155 := [rewrite]: #154
-#151 := (= #37 #150)
-#148 := (iff #32 #144)
-#149 := [rewrite]: #148
-#152 := [monotonicity #149]: #151
-#157 := [trans #152 #155]: #156
-#160 := [monotonicity #157]: #159
-#163 := [quant-intro #160]: #162
-#119 := [asserted]: #39
-#164 := [mp #119 #163]: #161
-#384 := [mp #164 #383]: #381
-#364 := [mp~ #384 #363]: #381
-#459 := [mp #364 #458]: #456
-#551 := [mp #459 #550]: #546
-#761 := (not #546)
-#762 := (or #761 #758)
-#71 := (= #70 f14)
-#753 := (ite #71 #752 #751)
-#606 := (+ #605 #146)
-#607 := (>= #606 0::int)
-#754 := (ite #607 #753 #750)
-#763 := (or #761 #754)
-#765 := (iff #763 #762)
-#767 := (iff #762 #762)
-#768 := [rewrite]: #767
-#759 := (iff #754 #758)
-#756 := (iff #753 #755)
-#282 := (iff #71 #281)
-#283 := [rewrite]: #282
-#757 := [monotonicity #283]: #756
-#620 := (iff #607 #617)
-#609 := (+ #146 #605)
-#612 := (>= #609 0::int)
-#618 := (iff #612 #617)
-#619 := [rewrite]: #618
-#613 := (iff #607 #612)
-#610 := (= #606 #609)
-#611 := [rewrite]: #610
-#614 := [monotonicity #611]: #613
-#621 := [trans #614 #619]: #620
-#760 := [monotonicity #621 #757]: #759
-#766 := [monotonicity #760]: #765
-#769 := [trans #766 #768]: #765
-#764 := [quant-inst]: #763
-#770 := [mp #764 #769]: #762
-#1317 := [unit-resolution #770 #551]: #758
-#981 := (= #31 #605)
-#1295 := (= #605 #31)
-#280 := [asserted]: #71
-#286 := [mp #280 #283]: #281
-#1290 := [symm #286]: #71
-#1296 := [monotonicity #1290]: #1295
-#1297 := [symm #1296]: #981
-#1318 := (not #981)
-#1319 := (or #1318 #617)
-#1320 := [th-lemma]: #1319
-#1321 := [unit-resolution #1320 #1297]: #617
-#639 := (not #617)
-#783 := (not #758)
-#784 := (or #783 #639 #755)
-#785 := [def-axiom]: #784
-#1322 := [unit-resolution #785 #1321 #1317]: #755
-#771 := (not #755)
-#1323 := (or #771 #752)
-#772 := (not #281)
-#773 := (or #771 #772 #752)
-#774 := [def-axiom]: #773
-#1324 := [unit-resolution #774 #286]: #1323
-#1325 := [unit-resolution #1324 #1322]: #752
-#1327 := [symm #1325]: #1326
-#1330 := [monotonicity #1327]: #1329
-#1334 := [trans #1330 #1332]: #1333
-#1336 := [monotonicity #1334]: #1335
-#303 := 0::real
-#301 := -1::real
-#1033 := (* -1::real #749)
-#1234 := (+ #650 #1033)
-#1236 := (>= #1234 0::real)
-#1243 := (not #1236)
-#1237 := [hypothesis]: #1236
-decl f20 :: S5
-#78 := f20
-#1034 := (f8 f20 #34)
-#1037 := (* -1::real #1034)
-#1048 := (+ #749 #1037)
-#1049 := (<= #1048 0::real)
-#1073 := (not #1049)
-decl f17 :: S3
-#48 := f17
-#76 := (f19 f17)
-#601 := (f8 #76 #34)
-#1038 := (+ #601 #1037)
-#1039 := (>= #1038 0::real)
-#1054 := (or #1039 #1049)
-#1057 := (not #1054)
-#79 := (f8 f20 #9)
-#588 := (pattern #79)
-#77 := (f8 #76 #9)
-#561 := (pattern #77)
-#310 := (* -1::real #82)
-#311 := (+ #79 #310)
-#309 := (>= #311 0::real)
-#305 := (* -1::real #79)
-#306 := (+ #77 #305)
-#304 := (>= #306 0::real)
-#422 := (or #304 #309)
-#423 := (not #422)
-#589 := (forall (vars (?v0 S2)) (:pat #561 #588 #545) #423)
-#426 := (forall (vars (?v0 S2)) #423)
-#592 := (iff #426 #589)
-#590 := (iff #423 #423)
-#591 := [refl]: #590
-#593 := [quant-intro #591]: #592
-#312 := (not #309)
-#302 := (not #304)
-#315 := (and #302 #312)
-#318 := (forall (vars (?v0 S2)) #315)
-#427 := (iff #318 #426)
-#424 := (iff #315 #423)
-#425 := [rewrite]: #424
-#428 := [quant-intro #425]: #427
-#414 := (~ #318 #318)
-#412 := (~ #315 #315)
-#413 := [refl]: #412
-#415 := [nnf-pos #413]: #414
-decl f4 :: S3
-#8 := f4
-#89 := (f19 f4)
-#90 := (f8 #89 #9)
-#328 := (* -1::real #90)
-#329 := (+ #79 #328)
-#327 := (>= #329 0::real)
-#330 := (not #327)
-#87 := (f8 #86 #9)
-#321 := (* -1::real #87)
-#322 := (+ #79 #321)
-#323 := (<= #322 0::real)
-#324 := (not #323)
-#333 := (and #324 #330)
-#336 := (forall (vars (?v0 S2)) #333)
-#339 := (and #318 #336)
-#91 := (< #79 #90)
-#88 := (< #87 #79)
-#92 := (and #88 #91)
-#93 := (forall (vars (?v0 S2)) #92)
-#83 := (< #79 #82)
-#80 := (< #77 #79)
-#84 := (and #80 #83)
-#85 := (forall (vars (?v0 S2)) #84)
-#94 := (and #85 #93)
-#340 := (iff #94 #339)
-#337 := (iff #93 #336)
-#334 := (iff #92 #333)
-#331 := (iff #91 #330)
-#332 := [rewrite]: #331
-#325 := (iff #88 #324)
-#326 := [rewrite]: #325
-#335 := [monotonicity #326 #332]: #334
-#338 := [quant-intro #335]: #337
-#319 := (iff #85 #318)
-#316 := (iff #84 #315)
-#313 := (iff #83 #312)
-#314 := [rewrite]: #313
-#307 := (iff #80 #302)
-#308 := [rewrite]: #307
-#317 := [monotonicity #308 #314]: #316
-#320 := [quant-intro #317]: #319
-#341 := [monotonicity #320 #338]: #340
-#297 := [asserted]: #94
-#342 := [mp #297 #341]: #339
-#343 := [and-elim #342]: #318
-#416 := [mp~ #343 #415]: #318
-#429 := [mp #416 #428]: #426
-#594 := [mp #429 #593]: #589
-#1060 := (not #589)
-#1061 := (or #1060 #1057)
-#1035 := (+ #1034 #1033)
-#1036 := (>= #1035 0::real)
-#1040 := (or #1039 #1036)
-#1041 := (not #1040)
-#1062 := (or #1060 #1041)
-#1064 := (iff #1062 #1061)
-#1066 := (iff #1061 #1061)
-#1067 := [rewrite]: #1066
-#1058 := (iff #1041 #1057)
-#1055 := (iff #1040 #1054)
-#1052 := (iff #1036 #1049)
-#1042 := (+ #1033 #1034)
-#1045 := (>= #1042 0::real)
-#1050 := (iff #1045 #1049)
-#1051 := [rewrite]: #1050
-#1046 := (iff #1036 #1045)
-#1043 := (= #1035 #1042)
-#1044 := [rewrite]: #1043
-#1047 := [monotonicity #1044]: #1046
-#1053 := [trans #1047 #1051]: #1052
-#1056 := [monotonicity #1053]: #1055
-#1059 := [monotonicity #1056]: #1058
-#1065 := [monotonicity #1059]: #1064
-#1068 := [trans #1065 #1067]: #1064
-#1063 := [quant-inst]: #1062
-#1069 := [mp #1063 #1068]: #1061
-#1238 := [unit-resolution #1069 #594]: #1057
-#1074 := (or #1054 #1073)
-#1075 := [def-axiom]: #1074
-#1239 := [unit-resolution #1075 #1238]: #1073
-#1150 := (+ #650 #1037)
-#1151 := (>= #1150 0::real)
-#1183 := (not #1151)
-#693 := (f8 #89 #34)
-#1162 := (+ #693 #1037)
-#1163 := (<= #1162 0::real)
-#1168 := (or #1151 #1163)
-#1171 := (not #1168)
-#536 := (pattern #90)
-#553 := (pattern #87)
-#430 := (or #323 #327)
-#431 := (not #430)
-#595 := (forall (vars (?v0 S2)) (:pat #588 #553 #536) #431)
-#434 := (forall (vars (?v0 S2)) #431)
-#598 := (iff #434 #595)
-#596 := (iff #431 #431)
-#597 := [refl]: #596
-#599 := [quant-intro #597]: #598
-#435 := (iff #336 #434)
-#432 := (iff #333 #431)
-#433 := [rewrite]: #432
-#436 := [quant-intro #433]: #435
-#419 := (~ #336 #336)
-#417 := (~ #333 #333)
-#418 := [refl]: #417
-#420 := [nnf-pos #418]: #419
-#344 := [and-elim #342]: #336
-#421 := [mp~ #344 #420]: #336
-#437 := [mp #421 #436]: #434
-#600 := [mp #437 #599]: #595
-#1118 := (not #595)
-#1174 := (or #1118 #1171)
-#1136 := (* -1::real #693)
-#1137 := (+ #1034 #1136)
-#1138 := (>= #1137 0::real)
-#1139 := (* -1::real #650)
-#1140 := (+ #1034 #1139)
-#1141 := (<= #1140 0::real)
-#1142 := (or #1141 #1138)
-#1143 := (not #1142)
-#1175 := (or #1118 #1143)
-#1177 := (iff #1175 #1174)
-#1179 := (iff #1174 #1174)
-#1180 := [rewrite]: #1179
-#1172 := (iff #1143 #1171)
-#1169 := (iff #1142 #1168)
-#1166 := (iff #1138 #1163)
-#1156 := (+ #1136 #1034)
-#1159 := (>= #1156 0::real)
-#1164 := (iff #1159 #1163)
-#1165 := [rewrite]: #1164
-#1160 := (iff #1138 #1159)
-#1157 := (= #1137 #1156)
-#1158 := [rewrite]: #1157
-#1161 := [monotonicity #1158]: #1160
-#1167 := [trans #1161 #1165]: #1166
-#1154 := (iff #1141 #1151)
-#1144 := (+ #1139 #1034)
-#1147 := (<= #1144 0::real)
-#1152 := (iff #1147 #1151)
-#1153 := [rewrite]: #1152
-#1148 := (iff #1141 #1147)
-#1145 := (= #1140 #1144)
-#1146 := [rewrite]: #1145
-#1149 := [monotonicity #1146]: #1148
-#1155 := [trans #1149 #1153]: #1154
-#1170 := [monotonicity #1155 #1167]: #1169
-#1173 := [monotonicity #1170]: #1172
-#1178 := [monotonicity #1173]: #1177
-#1181 := [trans #1178 #1180]: #1177
-#1176 := [quant-inst]: #1175
-#1182 := [mp #1176 #1181]: #1174
-#1240 := [unit-resolution #1182 #600]: #1171
-#1184 := (or #1168 #1183)
-#1185 := [def-axiom]: #1184
-#1241 := [unit-resolution #1185 #1240]: #1183
-#1242 := [th-lemma #1241 #1239 #1237]: false
-#1244 := [lemma #1242]: #1243
-#1247 := (or #1246 #1236)
-#1248 := [th-lemma]: #1247
-#1316 := [unit-resolution #1248 #1244]: #1246
-#1337 := [mp #1316 #1336]: #690
-#1339 := (or #663 #651)
-decl f16 :: S5
-#43 := f16
-#603 := (f8 f16 #34)
-#652 := (= #603 #650)
-#668 := (ite #663 #652 #651)
-#42 := (f8 f10 #9)
-#554 := (pattern #42)
-#44 := (f8 f16 #9)
-#552 := (pattern #44)
-#461 := (= #42 #87)
-#460 := (= #44 #87)
-#124 := (* -1::int #14)
-#125 := (+ #12 #124)
-#123 := (>= #125 0::int)
-#462 := (ite #123 #460 #461)
-#555 := (forall (vars (?v0 S2)) (:pat #535 #552 #553 #554) #462)
-#465 := (forall (vars (?v0 S2)) #462)
-#558 := (iff #465 #555)
-#556 := (iff #462 #462)
-#557 := [refl]: #556
-#559 := [quant-intro #557]: #558
-#169 := (ite #123 #44 #42)
-#385 := (= #87 #169)
-#386 := (forall (vars (?v0 S2)) #385)
-#466 := (iff #386 #465)
-#463 := (iff #385 #462)
-#464 := [rewrite]: #463
-#467 := [quant-intro #464]: #466
-#359 := (~ #386 #386)
-#361 := (~ #385 #385)
-#362 := [refl]: #361
-#360 := [nnf-pos #362]: #359
-#41 := (f3 f15 #9)
-#174 := (= #41 #169)
-#177 := (forall (vars (?v0 S2)) #174)
-#387 := (iff #177 #386)
-#388 := [rewrite* #298]: #387
-#15 := (< #12 #14)
-#45 := (ite #15 #42 #44)
-#46 := (= #41 #45)
-#47 := (forall (vars (?v0 S2)) #46)
-#178 := (iff #47 #177)
-#175 := (iff #46 #174)
-#172 := (= #45 #169)
-#122 := (not #123)
-#166 := (ite #122 #42 #44)
-#170 := (= #166 #169)
-#171 := [rewrite]: #170
-#167 := (= #45 #166)
-#126 := (iff #15 #122)
-#127 := [rewrite]: #126
-#168 := [monotonicity #127]: #167
-#173 := [trans #168 #171]: #172
-#176 := [monotonicity #173]: #175
-#179 := [quant-intro #176]: #178
-#120 := [asserted]: #47
-#180 := [mp #120 #179]: #177
-#389 := [mp #180 #388]: #386
-#357 := [mp~ #389 #360]: #386
-#468 := [mp #357 #467]: #465
-#560 := [mp #468 #559]: #555
-#671 := (not #555)
-#672 := (or #671 #668)
-#653 := (+ #605 #124)
-#654 := (>= #653 0::int)
-#655 := (ite #654 #652 #651)
-#673 := (or #671 #655)
-#675 := (iff #673 #672)
-#677 := (iff #672 #672)
-#678 := [rewrite]: #677
-#669 := (iff #655 #668)
-#666 := (iff #654 #663)
-#656 := (+ #124 #605)
-#659 := (>= #656 0::int)
-#664 := (iff #659 #663)
-#665 := [rewrite]: #664
-#660 := (iff #654 #659)
-#657 := (= #653 #656)
-#658 := [rewrite]: #657
-#661 := [monotonicity #658]: #660
-#667 := [trans #661 #665]: #666
-#670 := [monotonicity #667]: #669
-#676 := [monotonicity #670]: #675
-#679 := [trans #676 #678]: #675
-#674 := [quant-inst]: #673
-#680 := [mp #674 #679]: #672
-#1338 := [unit-resolution #680 #560]: #668
-#681 := (not #668)
-#685 := (or #681 #663 #651)
-#686 := [def-axiom]: #685
-#1340 := [unit-resolution #686 #1338]: #1339
-#1341 := [unit-resolution #1340 #1337]: #663
-#1286 := (+ #14 #146)
-#1287 := (<= #1286 0::int)
-#1376 := (not #1287)
-#1285 := (= #14 #31)
-#1314 := (not #1285)
-#290 := (= f7 f14)
-#702 := (= f7 #70)
-decl f18 :: (-> int S4)
-#985 := (f18 #605)
-#1305 := (= #985 #70)
-#986 := (= #70 #985)
-#53 := (:var 0 S4)
-#54 := (f5 #53)
-#568 := (pattern #54)
-#55 := (f18 #54)
-#181 := (= #53 #55)
-#569 := (forall (vars (?v0 S4)) (:pat #568) #181)
-#199 := (forall (vars (?v0 S4)) #181)
-#572 := (iff #199 #569)
-#570 := (iff #181 #181)
-#571 := [refl]: #570
-#573 := [quant-intro #571]: #572
-#399 := (~ #199 #199)
-#397 := (~ #181 #181)
-#398 := [refl]: #397
-#400 := [nnf-pos #398]: #399
-#56 := (= #55 #53)
-#57 := (forall (vars (?v0 S4)) #56)
-#200 := (iff #57 #199)
-#197 := (iff #56 #181)
-#198 := [rewrite]: #197
-#201 := [quant-intro #198]: #200
-#165 := [asserted]: #57
-#204 := [mp #165 #201]: #199
-#401 := [mp~ #204 #400]: #199
-#574 := [mp #401 #573]: #569
-#989 := (not #569)
-#990 := (or #989 #986)
-#991 := [quant-inst]: #990
-#1289 := [unit-resolution #991 #574]: #986
-#1306 := [symm #1289]: #1305
-#1309 := (= f7 #985)
-#20 := (f11 f7)
-#72 := (f6 #20)
-#797 := (f5 #72)
-#987 := (f18 #797)
-#1303 := (= #987 #985)
-#1300 := (= #797 #605)
-#1298 := (= #797 #31)
-#1291 := [hypothesis]: #1285
-#1293 := (= #797 #14)
-#73 := (= #72 f7)
-#285 := (= f7 #72)
-#287 := (iff #73 #285)
-#288 := [rewrite]: #287
-#284 := [asserted]: #73
-#291 := [mp #284 #288]: #285
-#1292 := [symm #291]: #73
-#1294 := [monotonicity #1292]: #1293
-#1299 := [trans #1294 #1291]: #1298
-#1301 := [trans #1299 #1297]: #1300
-#1304 := [monotonicity #1301]: #1303
-#1307 := (= f7 #987)
-#988 := (= #72 #987)
-#994 := (or #989 #988)
-#995 := [quant-inst]: #994
-#1302 := [unit-resolution #995 #574]: #988
-#1308 := [trans #291 #1302]: #1307
-#1310 := [trans #1308 #1304]: #1309
-#1311 := [trans #1310 #1306]: #702
-#1312 := [trans #1311 #1290]: #290
-#294 := (not #290)
-#74 := (= f14 f7)
-#75 := (not #74)
-#295 := (iff #75 #294)
-#292 := (iff #74 #290)
-#293 := [rewrite]: #292
-#296 := [monotonicity #293]: #295
-#289 := [asserted]: #75
-#299 := [mp #289 #296]: #294
-#1313 := [unit-resolution #299 #1312]: false
-#1315 := [lemma #1313]: #1314
-#1380 := (or #1285 #1376)
-#1288 := (>= #1286 0::int)
-#807 := (* -1::int #797)
-#808 := (+ #31 #807)
-#809 := (<= #808 0::int)
-#793 := (f8 #76 #20)
-#21 := (f8 f10 #20)
-#794 := (= #21 #793)
-#838 := (not #794)
-#883 := (f8 #89 #20)
-#1259 := (= #793 #883)
-#1272 := (not #1259)
-#1362 := (iff #1272 #838)
-#1360 := (iff #1259 #794)
-#1355 := (= #793 #21)
-#1358 := (iff #1355 #794)
-#1359 := [commutativity]: #1358
-#1356 := (iff #1259 #1355)
-#1353 := (= #883 #21)
-#888 := (= #21 #883)
-#886 := (f8 f12 #20)
-#891 := (= #883 #886)
-#894 := (ite #285 #888 #891)
-#884 := (f8 f9 #20)
-#897 := (= #883 #884)
-#853 := (+ #14 #807)
-#854 := (<= #853 0::int)
-#900 := (ite #854 #894 #897)
-#441 := (= #17 #90)
-#439 := (= #23 #90)
-#438 := (= #21 #90)
-#18 := (= #11 f7)
-#440 := (ite #18 #438 #439)
-#442 := (ite #123 #440 #441)
-#539 := (forall (vars (?v0 S2)) (:pat #535 #536 #537 #538) #442)
-#445 := (forall (vars (?v0 S2)) #442)
-#542 := (iff #445 #539)
-#540 := (iff #442 #442)
-#541 := [refl]: #540
-#543 := [quant-intro #541]: #542
-#24 := (ite #18 #21 #23)
-#131 := (ite #123 #24 #17)
-#375 := (= #90 #131)
-#376 := (forall (vars (?v0 S2)) #375)
-#446 := (iff #376 #445)
-#443 := (iff #375 #442)
-#444 := [rewrite]: #443
-#447 := [quant-intro #444]: #446
-#369 := (~ #376 #376)
-#371 := (~ #375 #375)
-#372 := [refl]: #371
-#370 := [nnf-pos #372]: #369
-#10 := (f3 f4 #9)
-#136 := (= #10 #131)
-#139 := (forall (vars (?v0 S2)) #136)
-#377 := (iff #139 #376)
-#378 := [rewrite* #298]: #377
-#25 := (ite #15 #17 #24)
-#26 := (= #10 #25)
-#27 := (forall (vars (?v0 S2)) #26)
-#140 := (iff #27 #139)
-#137 := (iff #26 #136)
-#134 := (= #25 #131)
-#128 := (ite #122 #17 #24)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #25 #128)
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#138 := [monotonicity #135]: #137
-#141 := [quant-intro #138]: #140
-#118 := [asserted]: #27
-#142 := [mp #118 #141]: #139
-#379 := [mp #142 #378]: #376
-#367 := [mp~ #379 #370]: #376
-#448 := [mp #367 #447]: #445
-#544 := [mp #448 #543]: #539
-#717 := (not #539)
-#903 := (or #717 #900)
-#885 := (= #884 #883)
-#887 := (= #886 #883)
-#889 := (ite #73 #888 #887)
-#844 := (+ #797 #124)
-#845 := (>= #844 0::int)
-#890 := (ite #845 #889 #885)
-#904 := (or #717 #890)
-#906 := (iff #904 #903)
-#908 := (iff #903 #903)
-#909 := [rewrite]: #908
-#901 := (iff #890 #900)
-#898 := (iff #885 #897)
-#899 := [rewrite]: #898
-#895 := (iff #889 #894)
-#892 := (iff #887 #891)
-#893 := [rewrite]: #892
-#896 := [monotonicity #288 #893]: #895
-#857 := (iff #845 #854)
-#847 := (+ #124 #797)
-#850 := (>= #847 0::int)
-#855 := (iff #850 #854)
-#856 := [rewrite]: #855
-#851 := (iff #845 #850)
-#848 := (= #844 #847)
-#849 := [rewrite]: #848
-#852 := [monotonicity #849]: #851
-#858 := [trans #852 #856]: #857
-#902 := [monotonicity #858 #896 #899]: #901
-#907 := [monotonicity #902]: #906
-#910 := [trans #907 #909]: #906
-#905 := [quant-inst]: #904
-#911 := [mp #905 #910]: #903
-#1343 := [unit-resolution #911 #544]: #900
-#983 := (= #14 #797)
-#1344 := [symm #1294]: #983
-#1345 := (not #983)
-#1346 := (or #1345 #854)
-#1347 := [th-lemma]: #1346
-#1348 := [unit-resolution #1347 #1344]: #854
-#872 := (not #854)
-#924 := (not #900)
-#925 := (or #924 #872 #894)
-#926 := [def-axiom]: #925
-#1349 := [unit-resolution #926 #1348 #1343]: #894
-#912 := (not #894)
-#1350 := (or #912 #888)
-#913 := (not #285)
-#914 := (or #912 #913 #888)
-#915 := [def-axiom]: #914
-#1351 := [unit-resolution #915 #291]: #1350
-#1352 := [unit-resolution #1351 #1349]: #888
-#1354 := [symm #1352]: #1353
-#1357 := [monotonicity #1354]: #1356
-#1361 := [trans #1357 #1359]: #1360
-#1363 := [monotonicity #1361]: #1362
-#1078 := (* -1::real #883)
-#1260 := (+ #793 #1078)
-#1262 := (>= #1260 0::real)
-#1269 := (not #1262)
-#1263 := [hypothesis]: #1262
-#1079 := (f8 f20 #20)
-#1093 := (* -1::real #1079)
-#1106 := (+ #883 #1093)
-#1107 := (<= #1106 0::real)
-#1131 := (not #1107)
-#841 := (f8 #86 #20)
-#1094 := (+ #841 #1093)
-#1095 := (>= #1094 0::real)
-#1112 := (or #1095 #1107)
-#1115 := (not #1112)
-#1119 := (or #1118 #1115)
-#1080 := (+ #1079 #1078)
-#1081 := (>= #1080 0::real)
-#1082 := (* -1::real #841)
-#1083 := (+ #1079 #1082)
-#1084 := (<= #1083 0::real)
-#1085 := (or #1084 #1081)
-#1086 := (not #1085)
-#1120 := (or #1118 #1086)
-#1122 := (iff #1120 #1119)
-#1124 := (iff #1119 #1119)
-#1125 := [rewrite]: #1124
-#1116 := (iff #1086 #1115)
-#1113 := (iff #1085 #1112)
-#1110 := (iff #1081 #1107)
-#1100 := (+ #1078 #1079)
-#1103 := (>= #1100 0::real)
-#1108 := (iff #1103 #1107)
-#1109 := [rewrite]: #1108
-#1104 := (iff #1081 #1103)
-#1101 := (= #1080 #1100)
-#1102 := [rewrite]: #1101
-#1105 := [monotonicity #1102]: #1104
-#1111 := [trans #1105 #1109]: #1110
-#1098 := (iff #1084 #1095)
-#1087 := (+ #1082 #1079)
-#1090 := (<= #1087 0::real)
-#1096 := (iff #1090 #1095)
-#1097 := [rewrite]: #1096
-#1091 := (iff #1084 #1090)
-#1088 := (= #1083 #1087)
-#1089 := [rewrite]: #1088
-#1092 := [monotonicity #1089]: #1091
-#1099 := [trans #1092 #1097]: #1098
-#1114 := [monotonicity #1099 #1111]: #1113
-#1117 := [monotonicity #1114]: #1116
-#1123 := [monotonicity #1117]: #1122
-#1126 := [trans #1123 #1125]: #1122
-#1121 := [quant-inst]: #1120
-#1127 := [mp #1121 #1126]: #1119
-#1264 := [unit-resolution #1127 #600]: #1115
-#1132 := (or #1112 #1131)
-#1133 := [def-axiom]: #1132
-#1265 := [unit-resolution #1133 #1264]: #1131
-#1194 := (+ #793 #1093)
-#1195 := (>= #1194 0::real)
-#1225 := (not #1195)
-#934 := (f8 #81 #20)
-#1204 := (+ #934 #1093)
-#1205 := (<= #1204 0::real)
-#1210 := (or #1195 #1205)
-#1213 := (not #1210)
-#1216 := (or #1060 #1213)
-#1191 := (* -1::real #934)
-#1192 := (+ #1079 #1191)
-#1193 := (>= #1192 0::real)
-#1196 := (or #1195 #1193)
-#1197 := (not #1196)
-#1217 := (or #1060 #1197)
-#1219 := (iff #1217 #1216)
-#1221 := (iff #1216 #1216)
-#1222 := [rewrite]: #1221
-#1214 := (iff #1197 #1213)
-#1211 := (iff #1196 #1210)
-#1208 := (iff #1193 #1205)
-#1198 := (+ #1191 #1079)
-#1201 := (>= #1198 0::real)
-#1206 := (iff #1201 #1205)
-#1207 := [rewrite]: #1206
-#1202 := (iff #1193 #1201)
-#1199 := (= #1192 #1198)
-#1200 := [rewrite]: #1199
-#1203 := [monotonicity #1200]: #1202
-#1209 := [trans #1203 #1207]: #1208
-#1212 := [monotonicity #1209]: #1211
-#1215 := [monotonicity #1212]: #1214
-#1220 := [monotonicity #1215]: #1219
-#1223 := [trans #1220 #1222]: #1219
-#1218 := [quant-inst]: #1217
-#1224 := [mp #1218 #1223]: #1216
-#1266 := [unit-resolution #1224 #594]: #1213
-#1226 := (or #1210 #1225)
-#1227 := [def-axiom]: #1226
-#1267 := [unit-resolution #1227 #1266]: #1225
-#1268 := [th-lemma #1267 #1265 #1263]: false
-#1270 := [lemma #1268]: #1269
-#1273 := (or #1272 #1262)
-#1274 := [th-lemma]: #1273
-#1342 := [unit-resolution #1274 #1270]: #1272
-#1364 := [mp #1342 #1363]: #838
-#1366 := (or #809 #794)
-#795 := (f8 f16 #20)
-#814 := (= #793 #795)
-#817 := (ite #809 #814 #794)
-#470 := (= #42 #77)
-#469 := (= #44 #77)
-#471 := (ite #145 #469 #470)
-#562 := (forall (vars (?v0 S2)) (:pat #535 #552 #561 #554) #471)
-#474 := (forall (vars (?v0 S2)) #471)
-#565 := (iff #474 #562)
-#563 := (iff #471 #471)
-#564 := [refl]: #563
-#566 := [quant-intro #564]: #565
-#185 := (ite #145 #44 #42)
-#390 := (= #77 #185)
-#391 := (forall (vars (?v0 S2)) #390)
-#475 := (iff #391 #474)
-#472 := (iff #390 #471)
-#473 := [rewrite]: #472
-#476 := [quant-intro #473]: #475
-#348 := (~ #391 #391)
-#358 := (~ #390 #390)
-#347 := [refl]: #358
-#395 := [nnf-pos #347]: #348
-#49 := (f3 f17 #9)
-#190 := (= #49 #185)
-#193 := (forall (vars (?v0 S2)) #190)
-#392 := (iff #193 #391)
-#393 := [rewrite* #298]: #392
-#50 := (ite #32 #42 #44)
-#51 := (= #49 #50)
-#52 := (forall (vars (?v0 S2)) #51)
-#194 := (iff #52 #193)
-#191 := (iff #51 #190)
-#188 := (= #50 #185)
-#182 := (ite #144 #42 #44)
-#186 := (= #182 #185)
-#187 := [rewrite]: #186
-#183 := (= #50 #182)
-#184 := [monotonicity #149]: #183
-#189 := [trans #184 #187]: #188
-#192 := [monotonicity #189]: #191
-#195 := [quant-intro #192]: #194
-#143 := [asserted]: #52
-#196 := [mp #143 #195]: #193
-#394 := [mp #196 #393]: #391
-#396 := [mp~ #394 #395]: #391
-#477 := [mp #396 #476]: #474
-#567 := [mp #477 #566]: #562
-#628 := (not #562)
-#820 := (or #628 #817)
-#796 := (= #795 #793)
-#798 := (+ #797 #146)
-#799 := (>= #798 0::int)
-#800 := (ite #799 #796 #794)
-#821 := (or #628 #800)
-#823 := (iff #821 #820)
-#825 := (iff #820 #820)
-#826 := [rewrite]: #825
-#818 := (iff #800 #817)
-#815 := (iff #796 #814)
-#816 := [rewrite]: #815
-#812 := (iff #799 #809)
-#801 := (+ #146 #797)
-#804 := (>= #801 0::int)
-#810 := (iff #804 #809)
-#811 := [rewrite]: #810
-#805 := (iff #799 #804)
-#802 := (= #798 #801)
-#803 := [rewrite]: #802
-#806 := [monotonicity #803]: #805
-#813 := [trans #806 #811]: #812
-#819 := [monotonicity #813 #816]: #818
-#824 := [monotonicity #819]: #823
-#827 := [trans #824 #826]: #823
-#822 := [quant-inst]: #821
-#828 := [mp #822 #827]: #820
-#1365 := [unit-resolution #828 #567]: #817
-#829 := (not #817)
-#833 := (or #829 #809 #794)
-#834 := [def-axiom]: #833
-#1367 := [unit-resolution #834 #1365]: #1366
-#1368 := [unit-resolution #1367 #1364]: #809
-#984 := (>= #853 0::int)
-#1369 := (or #1345 #984)
-#1370 := [th-lemma]: #1369
-#1371 := [unit-resolution #1370 #1344]: #984
-#830 := (not #809)
-#1372 := (not #984)
-#1373 := (or #1288 #1372 #830)
-#1374 := [th-lemma]: #1373
-#1375 := [unit-resolution #1374 #1371 #1368]: #1288
-#1377 := (not #1288)
-#1378 := (or #1285 #1376 #1377)
-#1379 := [th-lemma]: #1378
-#1381 := [unit-resolution #1379 #1375]: #1380
-#1382 := [unit-resolution #1381 #1315]: #1376
-#982 := (>= #616 0::int)
-#1383 := (or #1318 #982)
-#1384 := [th-lemma]: #1383
-#1385 := [unit-resolution #1384 #1297]: #982
-[th-lemma #1385 #1382 #1341]: false
-unsat
 ca942f6174c1f53254d5ef1b69b0e75f0d4027d4 208 0
 #2 := false
 #37 := 0::real
@@ -2631,6 +2620,156 @@
 #158 := [mp #152 #157]: #150
 [unit-resolution #158 #134 #165]: false
 unsat
+9ecd5f8eb0c8f78bd68a366175093e04632f1f73 149 0
+#2 := false
+#23 := 0::real
+decl f3 :: (-> S2 S3 real)
+decl f5 :: S3
+#9 := f5
+decl f6 :: S2
+#11 := f6
+#12 := (f3 f6 f5)
+#49 := -1::real
+#161 := (* -1::real #12)
+decl f4 :: S2
+#8 := f4
+#10 := (f3 f4 f5)
+#208 := (+ #10 #161)
+#210 := (>= #208 0::real)
+#13 := (= #10 #12)
+#45 := [asserted]: #13
+#213 := (not #13)
+#214 := (or #213 #210)
+#215 := [th-lemma]: #214
+#216 := [unit-resolution #215 #45]: #210
+decl f7 :: S2
+#16 := f7
+#26 := (f3 f7 f5)
+#165 := (* -1::real #26)
+#166 := (+ #10 #165)
+#212 := (>= #166 0::real)
+#227 := (not #212)
+#211 := (= #10 #26)
+#221 := (not #211)
+#67 := (= #12 #26)
+#75 := (not #67)
+#222 := (iff #75 #221)
+#219 := (iff #67 #211)
+#217 := (iff #211 #67)
+#218 := [monotonicity #45]: #217
+#220 := [symm #218]: #219
+#223 := [monotonicity #220]: #222
+#27 := (= #26 #12)
+#28 := (not #27)
+#76 := (iff #28 #75)
+#73 := (iff #27 #67)
+#74 := [rewrite]: #73
+#77 := [monotonicity #74]: #76
+#48 := [asserted]: #28
+#80 := [mp #48 #77]: #75
+#224 := [mp #80 #223]: #221
+#230 := (or #211 #227)
+#167 := (<= #166 0::real)
+#177 := (+ #12 #165)
+#178 := (>= #177 0::real)
+#183 := (not #178)
+#168 := (not #167)
+#186 := (or #168 #183)
+#189 := (not #186)
+#14 := (:var 0 S3)
+#19 := (f3 f6 #14)
+#154 := (pattern #19)
+#17 := (f3 f7 #14)
+#153 := (pattern #17)
+#15 := (f3 f4 #14)
+#152 := (pattern #15)
+#55 := (* -1::real #19)
+#56 := (+ #17 #55)
+#57 := (<= #56 0::real)
+#82 := (not #57)
+#50 := (* -1::real #17)
+#51 := (+ #15 #50)
+#52 := (<= #51 0::real)
+#85 := (not #52)
+#83 := (or #85 #82)
+#81 := (not #83)
+#155 := (forall (vars (?v0 S3)) (:pat #152 #153 #154) #81)
+#91 := (forall (vars (?v0 S3)) #81)
+#158 := (iff #91 #155)
+#156 := (iff #81 #81)
+#157 := [refl]: #156
+#159 := [quant-intro #157]: #158
+#60 := (and #52 #57)
+#63 := (forall (vars (?v0 S3)) #60)
+#92 := (iff #63 #91)
+#78 := (iff #60 #81)
+#90 := [rewrite]: #78
+#93 := [quant-intro #90]: #92
+#86 := (~ #63 #63)
+#88 := (~ #60 #60)
+#89 := [refl]: #88
+#87 := [nnf-pos #89]: #86
+#20 := (<= #17 #19)
+#18 := (<= #15 #17)
+#21 := (and #18 #20)
+#22 := (forall (vars (?v0 S3)) #21)
+#64 := (iff #22 #63)
+#61 := (iff #21 #60)
+#58 := (iff #20 #57)
+#59 := [rewrite]: #58
+#53 := (iff #18 #52)
+#54 := [rewrite]: #53
+#62 := [monotonicity #54 #59]: #61
+#65 := [quant-intro #62]: #64
+#46 := [asserted]: #22
+#66 := [mp #46 #65]: #63
+#84 := [mp~ #66 #87]: #63
+#94 := [mp #84 #93]: #91
+#160 := [mp #94 #159]: #155
+#192 := (not #155)
+#193 := (or #192 #189)
+#162 := (+ #26 #161)
+#163 := (<= #162 0::real)
+#164 := (not #163)
+#169 := (or #168 #164)
+#170 := (not #169)
+#194 := (or #192 #170)
+#196 := (iff #194 #193)
+#198 := (iff #193 #193)
+#199 := [rewrite]: #198
+#190 := (iff #170 #189)
+#187 := (iff #169 #186)
+#184 := (iff #164 #183)
+#181 := (iff #163 #178)
+#171 := (+ #161 #26)
+#174 := (<= #171 0::real)
+#179 := (iff #174 #178)
+#180 := [rewrite]: #179
+#175 := (iff #163 #174)
+#172 := (= #162 #171)
+#173 := [rewrite]: #172
+#176 := [monotonicity #173]: #175
+#182 := [trans #176 #180]: #181
+#185 := [monotonicity #182]: #184
+#188 := [monotonicity #185]: #187
+#191 := [monotonicity #188]: #190
+#197 := [monotonicity #191]: #196
+#200 := [trans #197 #199]: #196
+#195 := [quant-inst]: #194
+#201 := [mp #195 #200]: #193
+#225 := [unit-resolution #201 #160]: #189
+#202 := (or #186 #167)
+#203 := [def-axiom]: #202
+#226 := [unit-resolution #203 #225]: #167
+#228 := (or #211 #168 #227)
+#229 := [th-lemma]: #228
+#231 := [unit-resolution #229 #226]: #230
+#232 := [unit-resolution #231 #224]: #227
+#204 := (or #186 #178)
+#205 := [def-axiom]: #204
+#233 := [unit-resolution #205 #225]: #178
+[th-lemma #233 #232 #216]: false
+unsat
 ada412db5ba79d588ff49226c319d0dae76f5f87 271 0
 #2 := false
 #8 := 0::real
@@ -3192,156 +3331,6 @@
 #425 := [th-lemma]: #424
 [unit-resolution #425 #422 #408]: false
 unsat
-9ecd5f8eb0c8f78bd68a366175093e04632f1f73 149 0
-#2 := false
-#23 := 0::real
-decl f3 :: (-> S2 S3 real)
-decl f5 :: S3
-#9 := f5
-decl f6 :: S2
-#11 := f6
-#12 := (f3 f6 f5)
-#49 := -1::real
-#161 := (* -1::real #12)
-decl f4 :: S2
-#8 := f4
-#10 := (f3 f4 f5)
-#208 := (+ #10 #161)
-#210 := (>= #208 0::real)
-#13 := (= #10 #12)
-#45 := [asserted]: #13
-#213 := (not #13)
-#214 := (or #213 #210)
-#215 := [th-lemma]: #214
-#216 := [unit-resolution #215 #45]: #210
-decl f7 :: S2
-#16 := f7
-#26 := (f3 f7 f5)
-#165 := (* -1::real #26)
-#166 := (+ #10 #165)
-#212 := (>= #166 0::real)
-#227 := (not #212)
-#211 := (= #10 #26)
-#221 := (not #211)
-#67 := (= #12 #26)
-#75 := (not #67)
-#222 := (iff #75 #221)
-#219 := (iff #67 #211)
-#217 := (iff #211 #67)
-#218 := [monotonicity #45]: #217
-#220 := [symm #218]: #219
-#223 := [monotonicity #220]: #222
-#27 := (= #26 #12)
-#28 := (not #27)
-#76 := (iff #28 #75)
-#73 := (iff #27 #67)
-#74 := [rewrite]: #73
-#77 := [monotonicity #74]: #76
-#48 := [asserted]: #28
-#80 := [mp #48 #77]: #75
-#224 := [mp #80 #223]: #221
-#230 := (or #211 #227)
-#167 := (<= #166 0::real)
-#177 := (+ #12 #165)
-#178 := (>= #177 0::real)
-#183 := (not #178)
-#168 := (not #167)
-#186 := (or #168 #183)
-#189 := (not #186)
-#14 := (:var 0 S3)
-#19 := (f3 f6 #14)
-#154 := (pattern #19)
-#17 := (f3 f7 #14)
-#153 := (pattern #17)
-#15 := (f3 f4 #14)
-#152 := (pattern #15)
-#55 := (* -1::real #19)
-#56 := (+ #17 #55)
-#57 := (<= #56 0::real)
-#82 := (not #57)
-#50 := (* -1::real #17)
-#51 := (+ #15 #50)
-#52 := (<= #51 0::real)
-#85 := (not #52)
-#83 := (or #85 #82)
-#81 := (not #83)
-#155 := (forall (vars (?v0 S3)) (:pat #152 #153 #154) #81)
-#91 := (forall (vars (?v0 S3)) #81)
-#158 := (iff #91 #155)
-#156 := (iff #81 #81)
-#157 := [refl]: #156
-#159 := [quant-intro #157]: #158
-#60 := (and #52 #57)
-#63 := (forall (vars (?v0 S3)) #60)
-#92 := (iff #63 #91)
-#78 := (iff #60 #81)
-#90 := [rewrite]: #78
-#93 := [quant-intro #90]: #92
-#86 := (~ #63 #63)
-#88 := (~ #60 #60)
-#89 := [refl]: #88
-#87 := [nnf-pos #89]: #86
-#20 := (<= #17 #19)
-#18 := (<= #15 #17)
-#21 := (and #18 #20)
-#22 := (forall (vars (?v0 S3)) #21)
-#64 := (iff #22 #63)
-#61 := (iff #21 #60)
-#58 := (iff #20 #57)
-#59 := [rewrite]: #58
-#53 := (iff #18 #52)
-#54 := [rewrite]: #53
-#62 := [monotonicity #54 #59]: #61
-#65 := [quant-intro #62]: #64
-#46 := [asserted]: #22
-#66 := [mp #46 #65]: #63
-#84 := [mp~ #66 #87]: #63
-#94 := [mp #84 #93]: #91
-#160 := [mp #94 #159]: #155
-#192 := (not #155)
-#193 := (or #192 #189)
-#162 := (+ #26 #161)
-#163 := (<= #162 0::real)
-#164 := (not #163)
-#169 := (or #168 #164)
-#170 := (not #169)
-#194 := (or #192 #170)
-#196 := (iff #194 #193)
-#198 := (iff #193 #193)
-#199 := [rewrite]: #198
-#190 := (iff #170 #189)
-#187 := (iff #169 #186)
-#184 := (iff #164 #183)
-#181 := (iff #163 #178)
-#171 := (+ #161 #26)
-#174 := (<= #171 0::real)
-#179 := (iff #174 #178)
-#180 := [rewrite]: #179
-#175 := (iff #163 #174)
-#172 := (= #162 #171)
-#173 := [rewrite]: #172
-#176 := [monotonicity #173]: #175
-#182 := [trans #176 #180]: #181
-#185 := [monotonicity #182]: #184
-#188 := [monotonicity #185]: #187
-#191 := [monotonicity #188]: #190
-#197 := [monotonicity #191]: #196
-#200 := [trans #197 #199]: #196
-#195 := [quant-inst]: #194
-#201 := [mp #195 #200]: #193
-#225 := [unit-resolution #201 #160]: #189
-#202 := (or #186 #167)
-#203 := [def-axiom]: #202
-#226 := [unit-resolution #203 #225]: #167
-#228 := (or #211 #168 #227)
-#229 := [th-lemma]: #228
-#231 := [unit-resolution #229 #226]: #230
-#232 := [unit-resolution #231 #224]: #227
-#204 := (or #186 #178)
-#205 := [def-axiom]: #204
-#233 := [unit-resolution #205 #225]: #178
-[th-lemma #233 #232 #216]: false
-unsat
 2dea73fd0603d00ddaec5e14116c465addb0b89e 870 0
 #2 := false
 #11 := 0::real
@@ -4213,3 +4202,5165 @@
 #949 := [unit-resolution #948 #945]: #819
 [th-lemma #926 #949 #646 #648 #899 #651 #938 #944 #644 #642]: false
 unsat
+fda4738b9d427b4c846961908e5b41bb384b40b6 89 0
+f1 -> val!2
+f2 -> val!0
+f23 -> val!1
+f24 -> val!12
+f7 -> val!20
+f25 -> val!18
+f21 -> val!3
+f10 -> val!5
+f9 -> val!6
+f14 -> val!7
+f16 -> val!8
+f4 -> val!10
+f13 -> val!13
+f15 -> val!14
+f31 -> val!17
+f22 -> {
+  val!1 val!12 -> val!2
+  else -> val!2
+}
+f6 -> {
+  val!18 -> val!21
+  else -> val!21
+}
+f5 -> {
+  val!20 -> 7720
+  val!21 -> 7719
+  val!17 -> 8365
+  val!18 -> 1796
+  val!19 -> 1797
+  val!15 -> 1
+  val!22 -> 8366
+  else -> 8366
+}
+f20 -> {
+  val!3 -> val!19
+  else -> val!19
+}
+f26 -> {
+  val!18 val!19 -> val!4
+  else -> val!4
+}
+f8 -> {
+  val!5 val!18 -> 0
+  val!6 val!18 -> 0
+  val!7 val!18 -> 0
+  val!1 val!18 -> -1
+  else -> -1
+}
+f28 -> {
+  val!8 -> val!9
+  val!10 -> val!11
+  else -> val!11
+}
+f27 -> {
+  val!9 val!11 -> val!12
+  else -> val!12
+}
+f17 -> {
+  1 -> val!15
+  8366 -> val!22
+  8365 -> val!17
+  1796 -> val!18
+  1797 -> val!19
+  7720 -> val!20
+  7719 -> val!21
+  else -> val!21
+}
+f30 -> {
+  val!15 val!22 -> val!16
+  else -> val!16
+}
+f29 -> {
+  val!20 val!16 -> val!2
+  else -> val!2
+}
+f3 -> {
+  val!13 val!18 -> 0
+  val!8 val!18 -> 0
+  val!14 val!18 -> 0
+  val!10 val!18 -> 0
+  else -> 0
+}
+f18 -> (ite (forall (?v2 S2)
+              (or (not (= f1 (f19 ?v2 (f20 f21))))
+                  (<= (+ (f8 #1 ?v2) (* -1 (f8 #2 ?v2))) 0))
+              :qid {k!45})
+            f1
+            (f18!0 #0 #1))
+unknown
+7f975502da925ceb9dd4add3271a5a407e743846 262 0
+f1 -> val!7
+f2 -> val!0
+f16 -> val!2
+f7 -> val!4
+f27 -> val!25
+f28 -> val!12
+f18 -> val!8
+f4 -> val!10
+f13 -> val!13
+f17 -> val!14
+f12 -> val!16
+f9 -> val!17
+f14 -> val!18
+f10 -> val!19
+f15 -> val!20
+f19 -> val!21
+f24 -> val!23
+f11 -> {
+  val!2 -> val!1
+  val!4 -> val!3
+  else -> val!3
+}
+f6 -> {
+  val!1 -> val!2
+  val!3 -> val!4
+  val!26 -> val!29
+  val!4 -> val!28
+  val!25 -> val!29
+  val!2 -> val!27
+  val!5 -> val!29
+  val!27 -> val!30
+  val!28 -> val!4
+  val!29 -> val!31
+  val!30 -> val!4
+  val!31 -> val!2
+  else -> val!2
+}
+f20 -> {
+  1 -> val!5
+  7720 -> val!26
+  7719 -> val!25
+  14272 -> val!2
+  14270 -> val!4
+  8365 -> val!27
+  14269 -> val!28
+  14271 -> val!29
+  8457 -> val!30
+  0 -> val!31
+  else -> val!31
+}
+f26 -> {
+  val!5 val!25 -> val!6
+  val!5 val!26 -> val!15
+  else -> val!15
+}
+f25 -> {
+  val!2 val!6 -> val!7
+  val!4 val!6 -> val!7
+  val!4 val!15 -> val!7
+  else -> val!7
+}
+f23 -> {
+  val!8 -> val!9
+  val!10 -> val!11
+  val!21 -> val!22
+  val!20 -> val!24
+  else -> val!24
+}
+f29 -> {
+  val!9 val!11 -> val!12
+  else -> val!12
+}
+f5 -> {
+  val!25 -> 7719
+  val!2 -> 14272
+  val!4 -> 14270
+  val!5 -> 1
+  val!26 -> 7720
+  val!29 -> 14271
+  val!28 -> 14269
+  val!27 -> 8365
+  val!30 -> 8457
+  val!31 -> 0
+  else -> 0
+}
+f3 -> {
+  val!21 val!1 -> 2
+  val!8 val!1 -> 2
+  val!14 val!1 -> 3
+  val!20 val!1 -> 4
+  val!10 val!1 -> 3
+  val!13 val!1 -> 2
+  val!21 val!3 -> 7
+  val!8 val!3 -> 5
+  val!14 val!3 -> 6
+  val!20 val!3 -> 8
+  val!10 val!3 -> 7
+  val!13 val!3 -> 8
+  val!13 val!4 -> 29
+  val!10 val!2 -> 30
+  val!10 val!5 -> 36
+  val!10 val!4 -> 31
+  val!10 val!25 -> 37
+  val!10 val!26 -> 38
+  val!13 val!2 -> 32
+  val!13 val!5 -> 39
+  val!13 val!25 -> 40
+  val!13 val!26 -> 41
+  val!20 val!2 -> 30
+  val!20 val!5 -> 33
+  val!20 val!4 -> 31
+  val!20 val!25 -> 34
+  val!20 val!26 -> 35
+  val!14 val!2 -> 30
+  val!14 val!5 -> 36
+  val!14 val!4 -> 31
+  val!14 val!25 -> 37
+  val!14 val!26 -> 38
+  val!8 val!2 -> 32
+  val!8 val!5 -> 39
+  val!8 val!4 -> 29
+  val!8 val!25 -> 40
+  val!8 val!26 -> 41
+  val!21 val!2 -> 32
+  val!21 val!5 -> 42
+  val!21 val!4 -> 29
+  val!21 val!25 -> 43
+  val!21 val!26 -> 44
+  val!10 val!29 -> 73
+  val!10 val!28 -> 7
+  val!10 val!27 -> 75
+  val!13 val!29 -> 78
+  val!13 val!28 -> 8
+  val!13 val!27 -> 80
+  val!20 val!29 -> 73
+  val!20 val!28 -> 74
+  val!20 val!27 -> 75
+  val!14 val!29 -> 73
+  val!14 val!28 -> 76
+  val!14 val!27 -> 75
+  val!8 val!29 -> 78
+  val!8 val!28 -> 77
+  val!8 val!27 -> 80
+  val!21 val!29 -> 78
+  val!21 val!28 -> 79
+  val!21 val!27 -> 80
+  val!10 val!30 -> 7
+  val!13 val!30 -> 8
+  val!20 val!30 -> 93
+  val!14 val!30 -> 94
+  val!8 val!30 -> 95
+  val!21 val!30 -> 96
+  val!10 val!31 -> 101
+  val!13 val!31 -> 102
+  val!20 val!31 -> 4
+  val!14 val!31 -> 101
+  val!8 val!31 -> 102
+  val!21 val!31 -> 102
+  else -> 102
+}
+f8 -> {
+  val!18 val!1 -> 2
+  val!19 val!1 -> 4
+  val!16 val!1 -> 3
+  val!19 val!3 -> 7
+  val!17 val!3 -> 8
+  val!18 val!3 -> 5
+  val!16 val!3 -> 6
+  val!23 val!25 -> 0
+  val!9 val!25 -> -1
+  val!11 val!25 -> 1
+  val!24 val!25 -> 1
+  val!22 val!25 -> -1
+  val!23 val!2 -> 0
+  val!9 val!2 -> -1
+  val!11 val!2 -> 1
+  val!24 val!2 -> 1
+  val!22 val!2 -> -1
+  val!23 val!4 -> 0
+  val!9 val!4 -> -1
+  val!11 val!4 -> 1
+  val!24 val!4 -> 1
+  val!22 val!4 -> -1
+  val!23 val!5 -> 0
+  val!9 val!5 -> -1
+  val!11 val!5 -> 1
+  val!24 val!5 -> 1
+  val!22 val!5 -> -1
+  val!23 val!26 -> 0
+  val!9 val!26 -> -1
+  val!11 val!26 -> 1
+  val!24 val!26 -> 1
+  val!22 val!26 -> -1
+  val!17 val!2 -> 30
+  val!16 val!5 -> 36
+  val!17 val!5 -> 33
+  val!17 val!4 -> 31
+  val!16 val!25 -> 37
+  val!17 val!25 -> 34
+  val!16 val!26 -> 38
+  val!17 val!26 -> 35
+  val!19 val!2 -> 32
+  val!18 val!5 -> 39
+  val!19 val!5 -> 42
+  val!19 val!4 -> 29
+  val!18 val!25 -> 40
+  val!19 val!25 -> 43
+  val!18 val!26 -> 41
+  val!19 val!26 -> 44
+  val!23 val!3 -> 0
+  val!22 val!3 -> -1
+  val!24 val!3 -> 1
+  val!23 val!1 -> 0
+  val!22 val!1 -> -1
+  val!24 val!1 -> 1
+  val!11 val!3 -> 1
+  val!9 val!3 -> -1
+  val!11 val!1 -> 1
+  val!9 val!1 -> -1
+  val!23 val!27 -> 0
+  val!9 val!27 -> -1
+  val!11 val!27 -> 1
+  val!24 val!27 -> 1
+  val!22 val!27 -> -1
+  val!23 val!28 -> 0
+  val!9 val!28 -> -1
+  val!11 val!28 -> 1
+  val!24 val!28 -> 1
+  val!22 val!28 -> -1
+  val!23 val!29 -> 0
+  val!9 val!29 -> -1
+  val!11 val!29 -> 1
+  val!24 val!29 -> 1
+  val!22 val!29 -> -1
+  val!17 val!29 -> 73
+  val!16 val!28 -> 76
+  val!17 val!28 -> 74
+  val!17 val!27 -> 75
+  val!19 val!29 -> 78
+  val!18 val!28 -> 77
+  val!19 val!28 -> 79
+  val!19 val!27 -> 80
+  val!23 val!31 -> 0
+  val!9 val!31 -> -1
+  val!11 val!31 -> 1
+  val!24 val!31 -> 1
+  val!22 val!31 -> -1
+  val!23 val!30 -> 0
+  val!9 val!30 -> -1
+  val!11 val!30 -> 1
+  val!24 val!30 -> 1
+  val!22 val!30 -> -1
+  val!16 val!30 -> 94
+  val!17 val!30 -> 93
+  val!18 val!30 -> 95
+  val!19 val!30 -> 96
+  val!16 val!31 -> 101
+  val!18 val!31 -> 102
+  else -> 102
+}
+f32 -> (f30 (f31 #0) #1)
+unknown
+bd6da22de14f35502495633a6d03f6d719a5ebda 538 0
+f1 -> val!7
+f2 -> val!0
+f16 -> val!2
+f7 -> val!4
+f27 -> val!27
+f28 -> val!12
+f18 -> val!8
+f4 -> val!10
+f13 -> val!13
+f17 -> val!14
+f12 -> val!16
+f9 -> val!17
+f14 -> val!18
+f10 -> val!19
+f15 -> val!20
+f19 -> val!21
+f24 -> val!23
+f22 -> val!25
+f11 -> {
+  val!2 -> val!1
+  val!4 -> val!3
+  else -> val!3
+}
+f6 -> {
+  val!1 -> val!2
+  val!3 -> val!4
+  val!30 -> val!44
+  val!2 -> val!31
+  val!27 -> val!32
+  val!5 -> val!33
+  val!4 -> val!34
+  val!31 -> val!37
+  val!32 -> val!35
+  val!34 -> val!36
+  val!33 -> val!38
+  val!36 -> val!44
+  val!37 -> val!41
+  val!38 -> val!39
+  val!35 -> val!40
+  val!39 -> val!46
+  val!40 -> val!43
+  val!41 -> val!42
+  val!42 -> val!50
+  val!43 -> val!44
+  val!44 -> val!45
+  val!45 -> val!48
+  val!46 -> val!47
+  val!47 -> val!49
+  val!48 -> val!48
+  val!49 -> val!44
+  val!50 -> val!44
+  else -> val!44
+}
+f20 -> {
+  1 -> val!5
+  7720 -> val!30
+  7719 -> val!27
+  13211 -> val!2
+  13214 -> val!4
+  2944 -> val!31
+  8366 -> val!32
+  5724 -> val!33
+  1597 -> val!34
+  1176 -> val!35
+  1143 -> val!36
+  8753 -> val!37
+  10625 -> val!38
+  4877 -> val!39
+  2854 -> val!40
+  13213 -> val!41
+  5166 -> val!42
+  13216 -> val!43
+  6737 -> val!46
+  2853 -> val!45
+  1276 -> val!47
+  13210 -> val!48
+  2140 -> val!49
+  13212 -> val!44
+  10932 -> val!50
+  else -> val!50
+}
+f26 -> {
+  val!5 val!27 -> val!6
+  val!5 val!30 -> val!15
+  else -> val!15
+}
+f25 -> {
+  val!2 val!6 -> val!7
+  val!4 val!6 -> val!7
+  val!4 val!15 -> val!7
+  else -> val!7
+}
+f23 -> {
+  val!8 -> val!9
+  val!10 -> val!11
+  val!21 -> val!22
+  val!20 -> val!24
+  else -> val!24
+}
+f29 -> {
+  val!9 val!11 -> val!12
+  else -> val!12
+}
+f5 -> {
+  val!27 -> 7719
+  val!2 -> 13211
+  val!4 -> 13214
+  val!5 -> 1
+  val!30 -> 7720
+  val!31 -> 2944
+  val!32 -> 8366
+  val!33 -> 5724
+  val!34 -> 1597
+  val!37 -> 8753
+  val!35 -> 1176
+  val!36 -> 1143
+  val!38 -> 10625
+  val!41 -> 13213
+  val!39 -> 4877
+  val!40 -> 2854
+  val!46 -> 6737
+  val!43 -> 13216
+  val!48 -> 13210
+  val!42 -> 5166
+  val!50 -> 10932
+  val!44 -> 13212
+  val!45 -> 2853
+  val!47 -> 1276
+  val!49 -> 2140
+  else -> 2140
+}
+f21 -> {
+  val!25 -> val!26
+  else -> val!26
+}
+f3 -> {
+  val!21 val!1 -> 2
+  val!8 val!1 -> 4
+  val!14 val!1 -> 3
+  val!20 val!1 -> 4
+  val!10 val!1 -> 3
+  val!13 val!1 -> 4
+  val!21 val!3 -> 5
+  val!8 val!3 -> 5
+  val!14 val!3 -> 6
+  val!20 val!3 -> 6
+  val!10 val!3 -> 7
+  val!13 val!3 -> 8
+  val!10 val!30 -> 37
+  val!13 val!36 -> 30
+  val!10 val!36 -> 96
+  val!10 val!40 -> 32
+  val!21 val!27 -> 40
+  val!20 val!27 -> 35
+  val!21 val!2 -> 38
+  val!20 val!2 -> 33
+  val!10 val!27 -> 35
+  val!8 val!27 -> 40
+  val!10 val!4 -> 36
+  val!8 val!4 -> 41
+  val!10 val!2 -> 33
+  val!8 val!2 -> 38
+  val!21 val!4 -> 41
+  val!20 val!4 -> 36
+  val!10 val!5 -> 34
+  val!8 val!5 -> 39
+  val!8 val!30 -> 42
+  val!21 val!30 -> 43
+  val!20 val!30 -> 29
+  val!21 val!5 -> 39
+  val!20 val!5 -> 34
+  val!13 val!2 -> 38
+  val!13 val!5 -> 39
+  val!13 val!27 -> 40
+  val!13 val!4 -> 41
+  val!13 val!30 -> 42
+  val!14 val!2 -> 33
+  val!14 val!5 -> 34
+  val!14 val!27 -> 35
+  val!14 val!4 -> 36
+  val!14 val!30 -> 37
+  val!10 val!34 -> 70
+  val!8 val!34 -> 74
+  val!10 val!33 -> 68
+  val!8 val!33 -> 72
+  val!21 val!32 -> 73
+  val!20 val!32 -> 69
+  val!21 val!34 -> 74
+  val!20 val!34 -> 70
+  val!21 val!33 -> 72
+  val!20 val!33 -> 68
+  val!10 val!31 -> 67
+  val!8 val!31 -> 71
+  val!21 val!31 -> 71
+  val!20 val!31 -> 67
+  val!10 val!32 -> 69
+  val!8 val!32 -> 73
+  val!13 val!31 -> 71
+  val!13 val!33 -> 72
+  val!13 val!32 -> 73
+  val!13 val!34 -> 74
+  val!14 val!31 -> 67
+  val!14 val!33 -> 68
+  val!14 val!32 -> 69
+  val!14 val!34 -> 70
+  val!10 val!37 -> 98
+  val!8 val!37 -> 102
+  val!21 val!37 -> 103
+  val!20 val!37 -> 97
+  val!21 val!35 -> 104
+  val!20 val!35 -> 99
+  val!10 val!35 -> 99
+  val!8 val!35 -> 104
+  val!21 val!38 -> 100
+  val!20 val!38 -> 95
+  val!21 val!36 -> 101
+  val!20 val!36 -> 31
+  val!8 val!36 -> 30
+  val!10 val!38 -> 95
+  val!8 val!38 -> 100
+  val!13 val!38 -> 100
+  val!13 val!37 -> 102
+  val!13 val!35 -> 104
+  val!14 val!38 -> 95
+  val!14 val!36 -> 96
+  val!14 val!37 -> 98
+  val!14 val!35 -> 99
+  val!10 val!39 -> 125
+  val!8 val!39 -> 128
+  val!10 val!41 -> 127
+  val!8 val!41 -> 130
+  val!21 val!39 -> 128
+  val!20 val!39 -> 125
+  val!21 val!41 -> 130
+  val!20 val!41 -> 127
+  val!21 val!40 -> 131
+  val!20 val!40 -> 32
+  val!8 val!40 -> 131
+  val!13 val!39 -> 128
+  val!13 val!41 -> 130
+  val!13 val!40 -> 131
+  val!14 val!39 -> 125
+  val!14 val!41 -> 127
+  val!14 val!40 -> 32
+  val!10 val!48 -> 153
+  val!8 val!48 -> 159
+  val!8 val!42 -> 158
+  val!10 val!42 -> 152
+  val!21 val!43 -> 161
+  val!20 val!43 -> 154
+  val!21 val!48 -> 159
+  val!20 val!48 -> 153
+  val!10 val!43 -> 155
+  val!8 val!43 -> 160
+  val!10 val!46 -> 157
+  val!8 val!46 -> 163
+  val!21 val!46 -> 163
+  val!20 val!46 -> 157
+  val!20 val!42 -> 152
+  val!21 val!42 -> 158
+  val!13 val!42 -> 158
+  val!13 val!48 -> 159
+  val!13 val!43 -> 160
+  val!13 val!46 -> 163
+  val!14 val!42 -> 152
+  val!14 val!48 -> 153
+  val!14 val!43 -> 155
+  val!14 val!46 -> 157
+  val!21 val!44 -> 129
+  val!20 val!44 -> 126
+  val!10 val!44 -> 126
+  val!8 val!44 -> 129
+  val!20 val!47 -> 186
+  val!21 val!47 -> 189
+  val!10 val!47 -> 186
+  val!8 val!47 -> 189
+  val!21 val!50 -> 188
+  val!20 val!50 -> 184
+  val!10 val!50 -> 185
+  val!8 val!50 -> 187
+  val!13 val!44 -> 129
+  val!13 val!50 -> 187
+  val!13 val!47 -> 189
+  val!14 val!44 -> 126
+  val!14 val!50 -> 185
+  val!14 val!47 -> 186
+  val!10 val!45 -> 156
+  val!8 val!45 -> 162
+  val!10 val!49 -> 207
+  val!8 val!49 -> 208
+  val!21 val!45 -> 162
+  val!20 val!45 -> 156
+  val!21 val!49 -> 209
+  val!20 val!49 -> 206
+  val!13 val!45 -> 162
+  val!13 val!49 -> 208
+  val!14 val!45 -> 156
+  val!14 val!49 -> 207
+  else -> 207
+}
+f8 -> {
+  val!18 val!1 -> 2
+  val!19 val!1 -> 4
+  val!17 val!1 -> 3
+  val!19 val!3 -> 7
+  val!17 val!3 -> 8
+  val!18 val!3 -> 5
+  val!16 val!3 -> 6
+  val!23 val!27 -> 0
+  val!9 val!27 -> -1
+  val!11 val!27 -> 1
+  val!24 val!27 -> 1
+  val!22 val!27 -> -1
+  val!23 val!2 -> 0
+  val!9 val!2 -> -1
+  val!11 val!2 -> 1
+  val!24 val!2 -> 1
+  val!22 val!2 -> -1
+  val!23 val!4 -> 0
+  val!9 val!4 -> -1
+  val!11 val!4 -> 1
+  val!24 val!4 -> 1
+  val!22 val!4 -> -1
+  val!23 val!5 -> 0
+  val!9 val!5 -> -1
+  val!11 val!5 -> 1
+  val!24 val!5 -> 1
+  val!22 val!5 -> -1
+  val!22 val!1 -> -1
+  val!9 val!3 -> -1
+  val!24 val!1 -> 1
+  val!11 val!3 -> 1
+  val!23 val!30 -> 0
+  val!9 val!30 -> -1
+  val!11 val!30 -> 1
+  val!24 val!30 -> 1
+  val!22 val!30 -> -1
+  val!16 val!30 -> 29
+  val!16 val!36 -> 31
+  val!16 val!40 -> 32
+  val!9 val!1 -> -1
+  val!11 val!1 -> 1
+  val!22 val!3 -> -1
+  val!24 val!3 -> 1
+  val!17 val!2 -> 33
+  val!17 val!5 -> 34
+  val!17 val!27 -> 35
+  val!17 val!4 -> 36
+  val!17 val!30 -> 37
+  val!19 val!2 -> 38
+  val!19 val!5 -> 39
+  val!19 val!27 -> 40
+  val!19 val!4 -> 41
+  val!18 val!30 -> 43
+  val!19 val!30 -> 42
+  val!23 val!3 -> 0
+  val!23 val!1 -> 0
+  val!23 val!31 -> 0
+  val!9 val!31 -> -1
+  val!11 val!31 -> 1
+  val!24 val!31 -> 1
+  val!22 val!31 -> -1
+  val!23 val!33 -> 0
+  val!9 val!33 -> -1
+  val!11 val!33 -> 1
+  val!24 val!33 -> 1
+  val!22 val!33 -> -1
+  val!23 val!32 -> 0
+  val!9 val!32 -> -1
+  val!11 val!32 -> 1
+  val!24 val!32 -> 1
+  val!22 val!32 -> -1
+  val!23 val!34 -> 0
+  val!9 val!34 -> -1
+  val!11 val!34 -> 1
+  val!24 val!34 -> 1
+  val!22 val!34 -> -1
+  val!17 val!31 -> 67
+  val!17 val!33 -> 68
+  val!17 val!32 -> 69
+  val!17 val!34 -> 70
+  val!19 val!31 -> 71
+  val!19 val!33 -> 72
+  val!19 val!32 -> 73
+  val!19 val!34 -> 74
+  val!23 val!37 -> 0
+  val!9 val!37 -> -1
+  val!11 val!37 -> 1
+  val!24 val!37 -> 1
+  val!22 val!37 -> -1
+  val!23 val!38 -> 0
+  val!9 val!38 -> -1
+  val!11 val!38 -> 1
+  val!24 val!38 -> 1
+  val!22 val!38 -> -1
+  val!23 val!35 -> 0
+  val!9 val!35 -> -1
+  val!11 val!35 -> 1
+  val!24 val!35 -> 1
+  val!22 val!35 -> -1
+  val!23 val!36 -> 0
+  val!9 val!36 -> -1
+  val!11 val!36 -> 1
+  val!24 val!36 -> 1
+  val!22 val!36 -> -1
+  val!17 val!38 -> 95
+  val!17 val!36 -> 96
+  val!16 val!37 -> 97
+  val!17 val!37 -> 98
+  val!17 val!35 -> 99
+  val!19 val!38 -> 100
+  val!18 val!36 -> 101
+  val!19 val!36 -> 30
+  val!18 val!37 -> 103
+  val!19 val!37 -> 102
+  val!19 val!35 -> 104
+  val!23 val!39 -> 0
+  val!9 val!39 -> -1
+  val!11 val!39 -> 1
+  val!24 val!39 -> 1
+  val!22 val!39 -> -1
+  val!23 val!41 -> 0
+  val!9 val!41 -> -1
+  val!11 val!41 -> 1
+  val!24 val!41 -> 1
+  val!22 val!41 -> -1
+  val!23 val!40 -> 0
+  val!9 val!40 -> -1
+  val!11 val!40 -> 1
+  val!24 val!40 -> 1
+  val!22 val!40 -> -1
+  val!17 val!39 -> 125
+  val!17 val!41 -> 127
+  val!19 val!39 -> 128
+  val!19 val!41 -> 130
+  val!18 val!40 -> 131
+  val!23 val!48 -> 0
+  val!9 val!48 -> -1
+  val!11 val!48 -> 1
+  val!24 val!48 -> 1
+  val!22 val!48 -> -1
+  val!23 val!46 -> 0
+  val!9 val!46 -> -1
+  val!11 val!46 -> 1
+  val!24 val!46 -> 1
+  val!22 val!46 -> -1
+  val!23 val!42 -> 0
+  val!11 val!42 -> 1
+  val!9 val!42 -> -1
+  val!22 val!42 -> -1
+  val!24 val!42 -> 1
+  val!23 val!43 -> 0
+  val!9 val!43 -> -1
+  val!11 val!43 -> 1
+  val!24 val!43 -> 1
+  val!22 val!43 -> -1
+  val!17 val!42 -> 152
+  val!17 val!48 -> 153
+  val!16 val!43 -> 154
+  val!17 val!43 -> 155
+  val!17 val!46 -> 157
+  val!19 val!42 -> 158
+  val!19 val!48 -> 159
+  val!18 val!43 -> 161
+  val!19 val!43 -> 160
+  val!19 val!46 -> 163
+  val!23 val!50 -> 0
+  val!9 val!50 -> -1
+  val!11 val!50 -> 1
+  val!24 val!50 -> 1
+  val!22 val!50 -> -1
+  val!23 val!44 -> 0
+  val!9 val!44 -> -1
+  val!11 val!44 -> 1
+  val!24 val!44 -> 1
+  val!22 val!44 -> -1
+  val!23 val!47 -> 0
+  val!9 val!47 -> -1
+  val!11 val!47 -> 1
+  val!22 val!47 -> -1
+  val!24 val!47 -> 1
+  val!17 val!44 -> 126
+  val!16 val!50 -> 184
+  val!17 val!50 -> 185
+  val!17 val!47 -> 186
+  val!19 val!44 -> 129
+  val!18 val!50 -> 188
+  val!19 val!50 -> 187
+  val!19 val!47 -> 189
+  val!23 val!45 -> 0
+  val!9 val!45 -> -1
+  val!11 val!45 -> 1
+  val!24 val!45 -> 1
+  val!22 val!45 -> -1
+  val!23 val!49 -> 0
+  val!9 val!49 -> -1
+  val!11 val!49 -> 1
+  val!24 val!49 -> 1
+  val!22 val!49 -> -1
+  val!17 val!45 -> 156
+  val!16 val!49 -> 206
+  val!17 val!49 -> 207
+  val!19 val!45 -> 162
+  val!18 val!49 -> 209
+  val!19 val!49 -> 208
+  else -> 208
+}
+f30 -> {
+  val!1 val!26 -> val!28
+  val!3 val!26 -> val!29
+  val!27 val!26 -> val!51
+  val!2 val!26 -> val!52
+  val!4 val!26 -> val!53
+  val!5 val!26 -> val!54
+  val!30 val!26 -> val!55
+  val!34 val!26 -> val!56
+  val!33 val!26 -> val!57
+  val!32 val!26 -> val!58
+  val!31 val!26 -> val!59
+  val!37 val!26 -> val!60
+  val!35 val!26 -> val!61
+  val!38 val!26 -> val!62
+  val!36 val!26 -> val!63
+  val!39 val!26 -> val!64
+  val!41 val!26 -> val!65
+  val!40 val!26 -> val!67
+  val!48 val!26 -> val!68
+  val!42 val!26 -> val!69
+  val!43 val!26 -> val!71
+  val!46 val!26 -> val!72
+  val!44 val!26 -> val!66
+  val!47 val!26 -> val!73
+  val!50 val!26 -> val!74
+  val!45 val!26 -> val!70
+  val!49 val!26 -> val!75
+  else -> val!75
+}
+unknown
+68b48c983404c9f923b5c2c801b81a1ebfa54f09 60 0
+f1 -> val!0
+f2 -> val!1
+f7 -> val!3
+f8 -> val!5
+f12 -> val!13
+f11 -> val!6
+f17 -> val!8
+f18 -> val!9
+f16 -> val!10
+f20 -> val!11
+f14 -> val!12
+f6 -> {
+  val!3 -> val!2
+  val!5 -> val!4
+  else -> val!4
+}
+f5 -> {
+  val!2 -> val!3
+  val!4 -> val!5
+  else -> val!5
+}
+f4 -> {
+  val!5 -> 2
+  val!13 -> 40
+  val!14 -> 1
+  val!3 -> 3
+  val!17 -> 0
+  val!4 -> 1236
+  else -> 1236
+}
+f3 -> {
+  1 -> val!14
+  2 -> val!5
+  40 -> val!13
+  3 -> val!3
+  0 -> val!17
+  1236 -> val!4
+  else -> val!4
+}
+f10 -> {
+  val!6 val!13 -> val!7
+  else -> val!7
+}
+f13 -> {
+  val!12 -> val!17
+  else -> val!17
+}
+f9 -> {
+  val!3 val!7 -> val!15
+  val!5 val!7 -> val!16
+  else -> val!16
+}
+f15 -> {
+  val!9 val!2 -> 0
+  val!8 val!2 -> -1
+  val!10 val!2 -> 1
+  val!11 val!2 -> 1
+  else -> 1
+}
+unknown
+2cf0827cf5826be278dbca8a4dccba886e2b01ef 611 0
+f1 -> val!7
+f2 -> val!0
+f16 -> val!2
+f7 -> val!4
+f24 -> val!31
+f29 -> val!16
+f18 -> val!8
+f4 -> val!10
+f13 -> val!12
+f17 -> val!14
+f12 -> val!18
+f9 -> val!19
+f14 -> val!20
+f10 -> val!21
+f15 -> val!22
+f19 -> val!23
+f23 -> val!24
+f28 -> val!27
+f26 -> val!29
+f11 -> {
+  val!2 -> val!1
+  val!4 -> val!3
+  else -> val!3
+}
+f6 -> {
+  val!1 -> val!2
+  val!3 -> val!4
+  val!36 -> val!40
+  val!31 -> val!41
+  val!5 -> val!37
+  val!2 -> val!38
+  val!4 -> val!39
+  val!37 -> val!43
+  val!38 -> val!54
+  val!39 -> val!44
+  val!40 -> val!45
+  val!41 -> val!42
+  val!43 -> val!47
+  val!44 -> val!46
+  val!45 -> val!56
+  val!42 -> val!51
+  val!46 -> val!53
+  val!47 -> val!48
+  val!48 -> val!49
+  val!49 -> val!50
+  val!50 -> val!49
+  val!51 -> val!52
+  val!54 -> val!59
+  val!53 -> val!55
+  val!57 -> val!54
+  val!55 -> val!2
+  val!56 -> val!58
+  val!52 -> val!57
+  val!58 -> val!54
+  val!59 -> val!60
+  val!60 -> val!86
+  val!86 -> val!2
+  else -> val!2
+}
+f20 -> {
+  1 -> val!5
+  7720 -> val!36
+  7719 -> val!31
+  31267 -> val!2
+  36564 -> val!4
+  8366 -> val!37
+  9246 -> val!38
+  8852 -> val!39
+  0 -> val!40
+  2944 -> val!41
+  26221 -> val!42
+  29498 -> val!43
+  15353 -> val!44
+  6996 -> val!45
+  2997 -> val!46
+  9245 -> val!47
+  19679 -> val!49
+  45859 -> val!50
+  9531 -> val!51
+  24738 -> val!52
+  30204 -> val!57
+  9247 -> val!53
+  23570 -> val!55
+  26809 -> val!56
+  23086 -> val!58
+  5529 -> val!59
+  36563 -> val!60
+  36565 -> val!48
+  27246 -> val!86
+  31266 -> val!54
+  else -> val!54
+}
+f22 -> {
+  val!5 val!31 -> val!6
+  val!5 val!36 -> val!17
+  val!24 val!31 -> val!25
+  else -> val!25
+}
+f21 -> {
+  val!2 val!6 -> val!7
+  val!4 val!6 -> val!7
+  val!4 val!17 -> val!7
+  val!2 val!25 -> val!32
+  val!4 val!25 -> val!33
+  else -> val!33
+}
+f27 -> {
+  val!8 -> val!9
+  val!10 -> val!11
+  val!12 -> val!13
+  val!14 -> val!15
+  val!23 -> val!26
+  val!22 -> val!28
+  else -> val!28
+}
+f30 -> {
+  val!13 val!15 -> val!16
+  else -> val!16
+}
+f5 -> {
+  val!31 -> 7719
+  val!2 -> 31267
+  val!4 -> 36564
+  val!5 -> 1
+  val!36 -> 7720
+  val!40 -> 0
+  val!41 -> 2944
+  val!37 -> 8366
+  val!38 -> 9246
+  val!39 -> 8852
+  val!44 -> 15353
+  val!43 -> 29498
+  val!42 -> 26221
+  val!45 -> 6996
+  val!46 -> 2997
+  val!56 -> 26809
+  val!51 -> 9531
+  val!53 -> 9247
+  val!47 -> 9245
+  val!48 -> 36565
+  val!49 -> 19679
+  val!50 -> 45859
+  val!52 -> 24738
+  val!59 -> 5529
+  val!55 -> 23570
+  val!54 -> 31266
+  val!58 -> 23086
+  val!57 -> 30204
+  val!60 -> 36563
+  val!86 -> 27246
+  else -> 27246
+}
+f25 -> {
+  val!29 -> val!30
+  else -> val!30
+}
+f3 -> {
+  val!23 val!1 -> 2
+  val!8 val!1 -> 4
+  val!14 val!1 -> 3
+  val!22 val!1 -> 4
+  val!10 val!1 -> 3
+  val!12 val!1 -> 4
+  val!23 val!3 -> 5
+  val!8 val!3 -> 5
+  val!14 val!3 -> 6
+  val!22 val!3 -> 6
+  val!10 val!3 -> 7
+  val!12 val!3 -> 8
+  val!22 val!31 -> 32
+  val!22 val!37 -> 62
+  val!22 val!40 -> 60
+  val!22 val!43 -> 88
+  val!10 val!5 -> 30
+  val!8 val!5 -> 35
+  val!22 val!60 -> 150
+  val!23 val!36 -> 37
+  val!22 val!36 -> 31
+  val!23 val!31 -> 36
+  val!10 val!2 -> 28
+  val!8 val!2 -> 33
+  val!23 val!5 -> 35
+  val!22 val!5 -> 30
+  val!10 val!31 -> 32
+  val!8 val!31 -> 36
+  val!23 val!2 -> 33
+  val!22 val!2 -> 28
+  val!10 val!4 -> 29
+  val!8 val!4 -> 34
+  val!23 val!4 -> 34
+  val!22 val!4 -> 29
+  val!10 val!36 -> 31
+  val!8 val!36 -> 37
+  val!12 val!2 -> 33
+  val!12 val!4 -> 34
+  val!12 val!5 -> 35
+  val!12 val!31 -> 36
+  val!12 val!36 -> 37
+  val!14 val!2 -> 28
+  val!14 val!4 -> 29
+  val!14 val!5 -> 30
+  val!14 val!31 -> 32
+  val!14 val!36 -> 31
+  val!10 val!38 -> 59
+  val!8 val!38 -> 64
+  val!23 val!38 -> 64
+  val!22 val!38 -> 59
+  val!10 val!39 -> 58
+  val!8 val!39 -> 63
+  val!10 val!37 -> 62
+  val!8 val!37 -> 67
+  val!23 val!39 -> 63
+  val!22 val!39 -> 58
+  val!10 val!40 -> 60
+  val!8 val!40 -> 65
+  val!23 val!40 -> 65
+  val!23 val!41 -> 66
+  val!22 val!41 -> 61
+  val!23 val!37 -> 67
+  val!10 val!41 -> 61
+  val!8 val!41 -> 66
+  val!12 val!39 -> 63
+  val!12 val!38 -> 64
+  val!12 val!40 -> 65
+  val!12 val!41 -> 66
+  val!12 val!37 -> 67
+  val!14 val!39 -> 58
+  val!14 val!38 -> 59
+  val!14 val!40 -> 60
+  val!14 val!41 -> 61
+  val!14 val!37 -> 62
+  val!10 val!42 -> 89
+  val!8 val!42 -> 93
+  val!23 val!54 -> 94
+  val!22 val!54 -> 90
+  val!10 val!43 -> 88
+  val!8 val!43 -> 92
+  val!23 val!43 -> 92
+  val!23 val!42 -> 93
+  val!22 val!42 -> 89
+  val!23 val!44 -> 95
+  val!22 val!44 -> 91
+  val!10 val!44 -> 91
+  val!8 val!44 -> 95
+  val!10 val!54 -> 90
+  val!8 val!54 -> 94
+  val!23 val!45 -> 107
+  val!22 val!45 -> 106
+  val!10 val!45 -> 106
+  val!8 val!45 -> 107
+  val!12 val!43 -> 92
+  val!12 val!42 -> 93
+  val!12 val!54 -> 94
+  val!12 val!44 -> 95
+  val!14 val!43 -> 88
+  val!14 val!42 -> 89
+  val!14 val!54 -> 90
+  val!14 val!44 -> 91
+  val!12 val!45 -> 107
+  val!14 val!45 -> 106
+  val!10 val!46 -> 118
+  val!8 val!46 -> 122
+  val!23 val!46 -> 122
+  val!22 val!46 -> 118
+  val!23 val!47 -> 123
+  val!22 val!47 -> 127
+  val!8 val!59 -> 126
+  val!10 val!59 -> 120
+  val!23 val!59 -> 125
+  val!22 val!59 -> 119
+  val!10 val!47 -> 127
+  val!8 val!47 -> 123
+  val!10 val!51 -> 116
+  val!8 val!51 -> 124
+  val!23 val!51 -> 124
+  val!22 val!51 -> 116
+  val!23 val!56 -> 121
+  val!22 val!56 -> 117
+  val!10 val!56 -> 117
+  val!8 val!56 -> 121
+  val!12 val!56 -> 121
+  val!12 val!46 -> 122
+  val!12 val!47 -> 123
+  val!12 val!51 -> 124
+  val!12 val!59 -> 126
+  val!14 val!56 -> 117
+  val!14 val!46 -> 118
+  val!14 val!47 -> 127
+  val!14 val!51 -> 116
+  val!14 val!59 -> 120
+  val!23 val!58 -> 153
+  val!22 val!58 -> 148
+  val!10 val!53 -> 151
+  val!8 val!53 -> 156
+  val!23 val!53 -> 156
+  val!22 val!53 -> 151
+  val!10 val!48 -> 149
+  val!8 val!48 -> 154
+  val!23 val!48 -> 154
+  val!22 val!48 -> 149
+  val!23 val!52 -> 157
+  val!22 val!52 -> 152
+  val!10 val!58 -> 148
+  val!8 val!58 -> 153
+  val!10 val!60 -> 150
+  val!8 val!60 -> 155
+  val!10 val!52 -> 152
+  val!8 val!52 -> 157
+  val!23 val!60 -> 155
+  val!12 val!58 -> 153
+  val!12 val!48 -> 154
+  val!12 val!60 -> 155
+  val!12 val!53 -> 156
+  val!12 val!52 -> 157
+  val!14 val!58 -> 148
+  val!14 val!48 -> 149
+  val!14 val!60 -> 150
+  val!14 val!53 -> 151
+  val!14 val!52 -> 152
+  val!23 val!49 -> 186
+  val!22 val!49 -> 187
+  val!10 val!55 -> 179
+  val!8 val!55 -> 182
+  val!23 val!55 -> 178
+  val!22 val!55 -> 4
+  val!10 val!86 -> 181
+  val!8 val!86 -> 184
+  val!23 val!86 -> 185
+  val!22 val!86 -> 4
+  val!23 val!57 -> 183
+  val!22 val!57 -> 180
+  val!10 val!49 -> 187
+  val!8 val!49 -> 186
+  val!8 val!57 -> 183
+  val!10 val!57 -> 180
+  val!12 val!55 -> 182
+  val!12 val!57 -> 183
+  val!12 val!86 -> 184
+  val!12 val!49 -> 186
+  val!14 val!55 -> 179
+  val!14 val!57 -> 180
+  val!14 val!86 -> 181
+  val!14 val!49 -> 187
+  val!23 val!50 -> 205
+  val!22 val!50 -> 204
+  val!10 val!50 -> 204
+  val!8 val!50 -> 205
+  val!12 val!50 -> 205
+  val!14 val!50 -> 204
+  else -> 204
+}
+f8 -> {
+  val!20 val!1 -> 2
+  val!21 val!1 -> 4
+  val!19 val!1 -> 3
+  val!21 val!3 -> 7
+  val!19 val!3 -> 8
+  val!20 val!3 -> 5
+  val!18 val!3 -> 6
+  val!27 val!31 -> 0
+  val!9 val!31 -> -1
+  val!11 val!31 -> 1
+  val!28 val!31 -> 1
+  val!26 val!31 -> -1
+  val!27 val!2 -> 0
+  val!9 val!2 -> -1
+  val!11 val!2 -> 1
+  val!28 val!2 -> 1
+  val!26 val!2 -> -1
+  val!27 val!4 -> 0
+  val!9 val!4 -> -1
+  val!11 val!4 -> 1
+  val!28 val!4 -> 1
+  val!26 val!4 -> -1
+  val!27 val!5 -> 0
+  val!9 val!5 -> -1
+  val!11 val!5 -> 1
+  val!28 val!5 -> 1
+  val!26 val!5 -> -1
+  val!9 val!36 -> -1
+  val!27 val!36 -> 0
+  val!11 val!36 -> 1
+  val!28 val!36 -> 1
+  val!26 val!36 -> -1
+  val!19 val!2 -> 28
+  val!19 val!4 -> 29
+  val!19 val!5 -> 30
+  val!19 val!36 -> 31
+  val!19 val!31 -> 32
+  val!21 val!2 -> 33
+  val!21 val!4 -> 34
+  val!21 val!5 -> 35
+  val!21 val!31 -> 36
+  val!21 val!36 -> 37
+  val!27 val!38 -> 0
+  val!9 val!38 -> -1
+  val!11 val!38 -> 1
+  val!28 val!38 -> 1
+  val!26 val!38 -> -1
+  val!27 val!39 -> 0
+  val!9 val!39 -> -1
+  val!11 val!39 -> 1
+  val!28 val!39 -> 1
+  val!26 val!39 -> -1
+  val!27 val!37 -> 0
+  val!9 val!37 -> -1
+  val!11 val!37 -> 1
+  val!28 val!37 -> 1
+  val!26 val!37 -> -1
+  val!27 val!40 -> 0
+  val!9 val!40 -> -1
+  val!11 val!40 -> 1
+  val!28 val!40 -> 1
+  val!26 val!40 -> -1
+  val!27 val!41 -> 0
+  val!9 val!41 -> -1
+  val!11 val!41 -> 1
+  val!28 val!41 -> 1
+  val!26 val!41 -> -1
+  val!19 val!39 -> 58
+  val!19 val!38 -> 59
+  val!19 val!40 -> 60
+  val!19 val!41 -> 61
+  val!19 val!37 -> 62
+  val!21 val!39 -> 63
+  val!21 val!38 -> 64
+  val!21 val!40 -> 65
+  val!21 val!41 -> 66
+  val!21 val!37 -> 67
+  val!27 val!44 -> 0
+  val!9 val!44 -> -1
+  val!11 val!44 -> 1
+  val!28 val!44 -> 1
+  val!26 val!44 -> -1
+  val!27 val!45 -> 0
+  val!9 val!45 -> -1
+  val!11 val!45 -> 1
+  val!28 val!45 -> 1
+  val!26 val!45 -> -1
+  val!27 val!42 -> 0
+  val!9 val!42 -> -1
+  val!11 val!42 -> 1
+  val!28 val!42 -> 1
+  val!26 val!42 -> -1
+  val!27 val!43 -> 0
+  val!9 val!43 -> -1
+  val!11 val!43 -> 1
+  val!28 val!43 -> 1
+  val!26 val!43 -> -1
+  val!19 val!43 -> 88
+  val!19 val!42 -> 89
+  val!19 val!54 -> 90
+  val!19 val!44 -> 91
+  val!21 val!43 -> 92
+  val!21 val!42 -> 93
+  val!21 val!54 -> 94
+  val!21 val!44 -> 95
+  val!27 val!47 -> 0
+  val!9 val!47 -> -1
+  val!11 val!47 -> 1
+  val!28 val!47 -> 1
+  val!26 val!47 -> -1
+  val!27 val!59 -> 0
+  val!11 val!59 -> 1
+  val!9 val!59 -> -1
+  val!28 val!59 -> 1
+  val!26 val!59 -> -1
+  val!27 val!46 -> 0
+  val!9 val!46 -> -1
+  val!11 val!46 -> 1
+  val!28 val!46 -> 1
+  val!26 val!46 -> -1
+  val!19 val!45 -> 106
+  val!21 val!45 -> 107
+  val!27 val!51 -> 0
+  val!9 val!51 -> -1
+  val!11 val!51 -> 1
+  val!28 val!51 -> 1
+  val!26 val!51 -> -1
+  val!27 val!56 -> 0
+  val!9 val!56 -> -1
+  val!11 val!56 -> 1
+  val!28 val!56 -> 1
+  val!26 val!56 -> -1
+  val!19 val!56 -> 117
+  val!19 val!46 -> 118
+  val!18 val!47 -> 127
+  val!19 val!51 -> 116
+  val!18 val!59 -> 119
+  val!19 val!59 -> 120
+  val!21 val!56 -> 121
+  val!21 val!46 -> 122
+  val!20 val!47 -> 123
+  val!21 val!51 -> 124
+  val!20 val!59 -> 125
+  val!21 val!59 -> 126
+  val!27 val!58 -> 0
+  val!9 val!58 -> -1
+  val!11 val!58 -> 1
+  val!28 val!58 -> 1
+  val!26 val!58 -> -1
+  val!27 val!53 -> 0
+  val!9 val!53 -> -1
+  val!11 val!53 -> 1
+  val!28 val!53 -> 1
+  val!26 val!53 -> -1
+  val!27 val!48 -> 0
+  val!9 val!48 -> -1
+  val!11 val!48 -> 1
+  val!28 val!48 -> 1
+  val!26 val!48 -> -1
+  val!27 val!52 -> 0
+  val!9 val!52 -> -1
+  val!11 val!52 -> 1
+  val!28 val!52 -> 1
+  val!26 val!52 -> -1
+  val!27 val!60 -> 0
+  val!9 val!60 -> -1
+  val!11 val!60 -> 1
+  val!26 val!60 -> -1
+  val!28 val!60 -> 1
+  val!19 val!58 -> 148
+  val!19 val!48 -> 149
+  val!19 val!60 -> 150
+  val!19 val!53 -> 151
+  val!19 val!52 -> 152
+  val!21 val!58 -> 153
+  val!21 val!48 -> 154
+  val!21 val!60 -> 155
+  val!21 val!53 -> 156
+  val!21 val!52 -> 157
+  val!27 val!49 -> 0
+  val!9 val!49 -> -1
+  val!11 val!49 -> 1
+  val!28 val!49 -> 1
+  val!26 val!49 -> -1
+  val!27 val!86 -> 0
+  val!9 val!86 -> -1
+  val!11 val!86 -> 1
+  val!28 val!86 -> 1
+  val!26 val!86 -> -1
+  val!27 val!55 -> 0
+  val!9 val!55 -> -1
+  val!11 val!55 -> 1
+  val!28 val!55 -> 1
+  val!26 val!55 -> -1
+  val!27 val!57 -> 0
+  val!11 val!57 -> 1
+  val!9 val!57 -> -1
+  val!28 val!57 -> 1
+  val!26 val!57 -> -1
+  val!19 val!55 -> 179
+  val!19 val!57 -> 180
+  val!19 val!86 -> 181
+  val!18 val!49 -> 187
+  val!20 val!55 -> 178
+  val!21 val!55 -> 182
+  val!21 val!57 -> 183
+  val!20 val!86 -> 185
+  val!21 val!86 -> 184
+  val!20 val!49 -> 186
+  val!27 val!50 -> 0
+  val!9 val!50 -> -1
+  val!11 val!50 -> 1
+  val!28 val!50 -> 1
+  val!26 val!50 -> -1
+  val!27 val!54 -> 0
+  val!9 val!54 -> -1
+  val!11 val!54 -> 1
+  val!28 val!54 -> 1
+  val!26 val!54 -> -1
+  val!19 val!50 -> 204
+  val!21 val!50 -> 205
+  else -> 205
+}
+f31 -> {
+  val!1 val!30 -> val!34
+  val!3 val!30 -> val!35
+  val!5 val!30 -> val!61
+  val!36 val!30 -> val!62
+  val!31 val!30 -> val!63
+  val!2 val!30 -> val!64
+  val!4 val!30 -> val!65
+  val!38 val!30 -> val!66
+  val!39 val!30 -> val!67
+  val!37 val!30 -> val!68
+  val!40 val!30 -> val!69
+  val!41 val!30 -> val!70
+  val!42 val!30 -> val!71
+  val!54 val!30 -> val!72
+  val!43 val!30 -> val!73
+  val!44 val!30 -> val!74
+  val!45 val!30 -> val!75
+  val!46 val!30 -> val!76
+  val!47 val!30 -> val!77
+  val!59 val!30 -> val!78
+  val!51 val!30 -> val!79
+  val!56 val!30 -> val!80
+  val!58 val!30 -> val!81
+  val!53 val!30 -> val!82
+  val!48 val!30 -> val!83
+  val!52 val!30 -> val!84
+  val!60 val!30 -> val!85
+  val!49 val!30 -> val!87
+  val!55 val!30 -> val!88
+  val!86 val!30 -> val!89
+  val!57 val!30 -> val!90
+  val!50 val!30 -> val!91
+  else -> val!91
+}
+unknown
+a1ff1dee861d393a5412b6a4273eb86deab7593f 77 0
+f1 -> val!0
+f2 -> val!1
+f7 -> val!3
+f8 -> val!5
+f14 -> val!6
+f11 -> val!7
+f12 -> val!8
+f17 -> val!10
+f18 -> val!11
+f16 -> val!12
+f20 -> val!13
+f19 -> val!14
+f6 -> {
+  val!3 -> val!2
+  val!5 -> val!4
+  else -> val!4
+}
+f5 -> {
+  val!2 -> val!3
+  val!4 -> val!5
+  val!16 -> val!19
+  val!5 -> val!16
+  val!19 -> val!16
+  val!3 -> val!16
+  else -> val!16
+}
+f13 -> {
+  val!6 -> val!15
+  else -> val!15
+}
+f4 -> {
+  val!15 -> 40
+  val!5 -> 2
+  val!16 -> 1
+  val!3 -> 3
+  val!4 -> 1276
+  val!19 -> 0
+  else -> 0
+}
+f3 -> {
+  1 -> val!16
+  40 -> val!15
+  2 -> val!5
+  3 -> val!3
+  0 -> val!19
+  1276 -> val!4
+  else -> val!4
+}
+f10 -> {
+  val!7 val!8 -> val!9
+  else -> val!9
+}
+f9 -> {
+  val!3 val!9 -> val!17
+  val!5 val!9 -> val!18
+  else -> val!18
+}
+f15 -> {
+  val!11 val!2 -> 0
+  val!10 val!2 -> -1
+  val!12 val!2 -> 1
+  val!13 val!2 -> 1
+  val!11 val!16 -> 0
+  val!12 val!16 -> -1
+  val!14 val!16 -> 1
+  val!11 val!3 -> 0
+  val!12 val!3 -> -1
+  val!14 val!3 -> 1
+  val!11 val!5 -> 0
+  val!12 val!5 -> -1
+  val!14 val!5 -> 1
+  val!11 val!19 -> 0
+  val!12 val!19 -> -1
+  val!14 val!19 -> 1
+  else -> 1
+}
+unknown
+26a56d9f61c23c26da64882e31e65a36923764cb 72 0
+f1 -> val!0
+f2 -> val!1
+f7 -> val!3
+f8 -> val!5
+f12 -> val!13
+f11 -> val!6
+f17 -> val!8
+f18 -> val!9
+f16 -> val!10
+f20 -> val!11
+f19 -> val!12
+f6 -> {
+  val!3 -> val!2
+  val!5 -> val!4
+  else -> val!4
+}
+f5 -> {
+  val!2 -> val!3
+  val!4 -> val!5
+  val!3 -> val!17
+  val!14 -> val!14
+  val!5 -> val!14
+  val!17 -> val!14
+  else -> val!14
+}
+f4 -> {
+  val!13 -> 40
+  val!5 -> 2
+  val!14 -> 1
+  val!3 -> 3
+  val!4 -> 1276
+  val!17 -> 0
+  else -> 0
+}
+f3 -> {
+  1 -> val!14
+  40 -> val!13
+  2 -> val!5
+  3 -> val!3
+  1276 -> val!4
+  0 -> val!17
+  else -> val!17
+}
+f10 -> {
+  val!6 val!13 -> val!7
+  else -> val!7
+}
+f9 -> {
+  val!3 val!7 -> val!15
+  val!5 val!7 -> val!16
+  else -> val!16
+}
+f15 -> {
+  val!9 val!2 -> 0
+  val!8 val!2 -> -1
+  val!10 val!2 -> 1
+  val!11 val!2 -> 1
+  val!9 val!3 -> 0
+  val!10 val!3 -> -1
+  val!12 val!3 -> 1
+  val!9 val!14 -> 0
+  val!10 val!14 -> -1
+  val!12 val!14 -> 1
+  val!9 val!5 -> 0
+  val!10 val!5 -> -1
+  val!12 val!5 -> 1
+  val!9 val!17 -> 0
+  val!10 val!17 -> -1
+  val!12 val!17 -> 1
+  else -> 1
+}
+unknown
+a4a52206bfedbc1fa95069df86dc718ab11f5b3a 59 0
+f1 -> val!0
+f2 -> val!1
+f7 -> val!3
+f8 -> val!5
+f17 -> val!6
+f20 -> val!7
+f19 -> val!8
+f18 -> val!9
+f16 -> val!10
+f14 -> val!11
+f11 -> val!12
+f12 -> val!13
+f6 -> {
+  val!3 -> val!2
+  val!5 -> val!4
+  else -> val!4
+}
+f5 -> {
+  val!2 -> val!3
+  val!4 -> val!5
+  else -> val!5
+}
+f15 -> {
+  val!6 val!2 -> 0
+  val!7 val!2 -> 1
+  val!8 val!2 -> -1
+  val!6 val!4 -> 0
+  val!9 val!4 -> 1
+  val!10 val!4 -> -1
+  else -> -1
+}
+f4 -> {
+  val!3 -> 1237
+  val!5 -> 1238
+  val!15 -> 8957
+  val!16 -> 1
+  else -> 1
+}
+f13 -> {
+  val!11 -> val!15
+  else -> val!15
+}
+f3 -> {
+  1 -> val!16
+  1237 -> val!3
+  1238 -> val!5
+  8957 -> val!15
+  else -> val!15
+}
+f10 -> {
+  val!12 val!13 -> val!14
+  else -> val!14
+}
+f9 -> {
+  val!3 val!14 -> val!17
+  val!5 val!14 -> val!18
+  else -> val!18
+}
+unknown
+14333af854a5b65017dd8bd92066cf80fe83d74b 43 0
+f1 -> val!0
+f2 -> val!1
+f15 -> val!2
+f6 -> val!9
+f12 -> val!3
+f16 -> val!4
+f5 -> val!9
+f13 -> val!5
+f10 -> val!7
+f8 -> val!8
+f11 -> {
+  val!9 -> val!6
+  else -> val!6
+}
+f9 -> {
+  val!5 val!6 -> 1
+  val!3 val!6 -> 0
+  val!4 val!6 -> 1
+  val!2 val!6 -> -1
+  val!7 val!6 -> -1
+  else -> -1
+}
+f4 -> {
+  val!9 -> 7720
+  val!10 -> 7758
+  val!11 -> 1
+  else -> 1
+}
+f14 -> {
+  val!6 -> val!9
+  else -> val!9
+}
+f7 -> {
+  val!8 -> val!10
+  else -> val!10
+}
+f3 -> {
+  1 -> val!11
+  7720 -> val!9
+  7758 -> val!10
+  else -> val!10
+}
+unknown
+5f2c36ac6c49043bec7b255aa0d7a0c690b930a4 21 0
+f1 -> val!0
+f2 -> val!1
+f17 -> val!2
+f4 -> val!4
+f14 -> val!7
+f16 -> val!9
+f7 -> val!12
+f10 -> val!12
+f20 -> {
+  val!2 -> val!3
+  val!4 -> val!5
+  val!7 -> val!8
+  val!9 -> val!10
+  else -> val!10
+}
+f19 -> {
+  val!3 val!5 -> val!6
+  val!8 val!10 -> val!11
+  else -> val!11
+}
+unknown
+4bd0e90cc9ab46b702c811a3ea671a168ec22aba 65 0
+f1 -> val!0
+f2 -> val!1
+f14 -> val!19
+f21 -> val!3
+f13 -> val!6
+f23 -> val!8
+f17 -> val!9
+f16 -> val!11
+f4 -> val!13
+f7 -> val!19
+f10 -> val!15
+f9 -> val!16
+f12 -> val!17
+f15 -> val!18
+f11 -> {
+  val!19 -> val!2
+  else -> val!2
+}
+f20 -> {
+  val!3 -> val!4
+  else -> val!4
+}
+f19 -> {
+  val!2 val!4 -> val!5
+  else -> val!5
+}
+f22 -> {
+  val!6 -> val!7
+  val!9 -> val!10
+  val!11 -> val!12
+  val!13 -> val!14
+  else -> val!14
+}
+f8 -> {
+  val!7 val!2 -> -1
+  val!8 val!2 -> 0
+  val!10 val!2 -> 1
+  val!12 val!2 -> -1
+  val!14 val!2 -> 1
+  val!15 val!2 -> 0
+  val!16 val!2 -> -1
+  val!17 val!2 -> 1
+  val!18 val!2 -> 0
+  else -> 0
+}
+f5 -> {
+  val!19 -> 0
+  else -> 0
+}
+f18 -> {
+  0 -> val!19
+  else -> val!19
+}
+f6 -> {
+  val!2 -> val!19
+  else -> val!19
+}
+f3 -> {
+  val!13 val!2 -> 0
+  val!6 val!2 -> -1
+  val!11 val!2 -> 0
+  val!9 val!2 -> 1
+  else -> 1
+}
+unknown
+cdc3b5c1b7947a0fe95f5d2f65f941e2774116a3 63 0
+f1 -> val!0
+f2 -> val!1
+f14 -> val!17
+f21 -> val!3
+f13 -> val!6
+f23 -> val!8
+f17 -> val!9
+f16 -> val!11
+f4 -> val!13
+f7 -> val!17
+f10 -> val!15
+f9 -> val!16
+f11 -> {
+  val!17 -> val!2
+  else -> val!2
+}
+f20 -> {
+  val!3 -> val!4
+  else -> val!4
+}
+f19 -> {
+  val!2 val!4 -> val!5
+  else -> val!5
+}
+f22 -> {
+  val!6 -> val!7
+  val!9 -> val!10
+  val!11 -> val!12
+  val!13 -> val!14
+  else -> val!14
+}
+f8 -> {
+  val!7 val!2 -> -1
+  val!8 val!2 -> 0
+  val!10 val!2 -> 1
+  val!12 val!2 -> -1
+  val!14 val!2 -> 1
+  val!15 val!2 -> 0
+  val!16 val!2 -> 2
+  else -> 2
+}
+f5 -> {
+  val!17 -> 1
+  val!18 -> 0
+  else -> 0
+}
+f18 -> {
+  1 -> val!17
+  0 -> val!18
+  else -> val!18
+}
+f3 -> {
+  val!13 val!2 -> 2
+  val!11 val!2 -> 0
+  val!6 val!2 -> 0
+  val!9 val!2 -> 2
+  else -> 2
+}
+f6 -> {
+  val!2 -> val!18
+  else -> val!18
+}
+unknown
+e7b427c80202d9c4c29f2a9278c2bcdecdabe460 42 0
+f1 -> val!0
+f2 -> val!1
+f3 -> val!2
+f4 -> val!2
+f5 -> val!3
+f6 -> val!3
+f18 -> val!4
+f19 -> val!5
+f20 -> val!6
+f17 -> val!8
+f23 -> val!14
+f13 -> {
+  val!4 val!5 -> 0
+  val!6 val!5 -> -3
+  val!10 val!5 -> 0
+  val!13 val!5 -> 0
+  else -> 0
+}
+f16 -> {
+  val!3 -> val!7
+  val!2 -> val!11
+  else -> val!11
+}
+f15 -> {
+  val!7 val!8 -> val!9
+  val!11 val!8 -> val!12
+  else -> val!12
+}
+f14 -> {
+  val!9 val!4 -> val!10
+  val!12 val!6 -> val!13
+  else -> val!13
+}
+f22 -> {
+  val!14 -> val!15
+  else -> val!15
+}
+f21 -> {
+  val!5 val!15 -> val!16
+  else -> val!16
+}
+unknown
+fc367e02b7e503644f5c70cc9716512639619a51 52 0
+f1 -> val!0
+f2 -> val!1
+f3 -> val!2
+f4 -> val!2
+f5 -> val!3
+f6 -> val!3
+f20 -> val!4
+f21 -> val!15
+f22 -> val!5
+f19 -> val!7
+f25 -> val!13
+f15 -> {
+  val!4 val!15 -> 6
+  val!5 val!15 -> 0
+  val!9 val!15 -> 1
+  val!12 val!15 -> -1
+  else -> -1
+}
+f18 -> {
+  val!3 -> val!6
+  val!2 -> val!10
+  else -> val!10
+}
+f17 -> {
+  val!6 val!7 -> val!8
+  val!10 val!7 -> val!11
+  else -> val!11
+}
+f16 -> {
+  val!8 val!4 -> val!9
+  val!11 val!5 -> val!12
+  else -> val!12
+}
+f24 -> {
+  val!13 -> val!16
+  else -> val!16
+}
+f23 -> {
+  val!15 val!16 -> val!14
+  else -> val!14
+}
+f14 -> {
+  val!15 -> 0
+  val!16 -> 1
+  else -> 1
+}
+f13 -> {
+  0 -> val!15
+  1 -> val!16
+  else -> val!16
+}
+unknown
+2b6a5d632b112638788d36706b106b4e4d10b262 35 0
+f1 -> val!0
+f2 -> val!1
+f3 -> val!2
+f4 -> val!2
+f5 -> val!3
+f6 -> val!3
+f18 -> val!4
+f19 -> val!5
+f20 -> val!6
+f17 -> val!8
+f13 -> {
+  val!4 val!5 -> 0
+  val!6 val!5 -> -3
+  val!10 val!5 -> 0
+  val!13 val!5 -> 0
+  val!9 val!5 -> 0
+  val!12 val!5 -> 0
+  else -> 0
+}
+f16 -> {
+  val!3 -> val!7
+  val!2 -> val!11
+  else -> val!11
+}
+f15 -> {
+  val!7 val!8 -> val!9
+  val!11 val!8 -> val!12
+  else -> val!12
+}
+f14 -> {
+  val!9 val!4 -> val!10
+  val!12 val!6 -> val!13
+  else -> val!13
+}
+unknown
+6220e161feeb3d35604e0055563cdf27e5108197 35 0
+f1 -> val!0
+f2 -> val!1
+f8 -> val!2
+f12 -> val!2
+f11 -> val!3
+f13 -> val!3
+f22 -> val!4
+f9 -> val!5
+f23 -> val!6
+f10 -> val!7
+f21 -> val!8
+f4 -> val!9
+f5 -> {
+  val!4 val!5 -> 3
+  val!6 val!5 -> 0
+  val!11 val!5 -> 0
+  val!13 val!5 -> 0
+  else -> 0
+}
+f20 -> {
+  val!7 val!8 -> 3
+  val!9 val!8 -> 0
+  else -> 0
+}
+f7 -> {
+  val!3 -> val!10
+  val!2 -> val!12
+  else -> val!12
+}
+f24 -> {
+  val!10 val!8 -> val!11
+  val!12 val!8 -> val!13
+  else -> val!13
+}
+unknown
+c29b6c64de39317156c532c91451be225c15adb2 238 0
+#2 := false
+#48 := 0::real
+decl f19 :: (-> S3 S10 real)
+decl f20 :: S10
+#43 := f20
+decl f4 :: S3
+#8 := f4
+#58 := (f19 f4 f20)
+#109 := -1::real
+#153 := (* -1::real #58)
+decl f5 :: (-> S4 S5 real)
+decl f8 :: S5
+#13 := f8
+decl f22 :: S4
+#54 := f22
+#55 := (f5 f22 f8)
+#154 := (+ #55 #153)
+#137 := (* -1::real #55)
+#144 := (+ #137 #58)
+#188 := (<= #154 0::real)
+#195 := (ite #188 #144 #154)
+#451 := (* -1::real #195)
+#452 := (+ #144 #451)
+#453 := (<= #452 0::real)
+#435 := (= #144 #195)
+decl f21 :: S4
+#45 := f21
+#46 := (f5 f21 f8)
+decl f9 :: S3
+#17 := f9
+#44 := (f19 f9 f20)
+#120 := (* -1::real #44)
+#121 := (+ #120 #46)
+#110 := (* -1::real #46)
+#111 := (+ #44 #110)
+#216 := (>= #111 0::real)
+#223 := (ite #216 #111 #121)
+#447 := (* -1::real #223)
+#450 := (+ #121 #447)
+#454 := (<= #450 0::real)
+#442 := (= #121 #223)
+#217 := (not #216)
+#455 := [hypothesis]: #216
+#184 := (+ #44 #153)
+#185 := (<= #184 0::real)
+#206 := -3::real
+#234 := (* -3::real #223)
+#235 := (+ #137 #234)
+#236 := (+ #46 #235)
+#237 := (<= #236 0::real)
+#238 := (not #237)
+#207 := (* -3::real #195)
+#208 := (+ #137 #207)
+#209 := (+ #46 #208)
+#210 := (<= #209 0::real)
+#211 := (not #210)
+#249 := (and #185 #211 #238)
+#65 := (<= #44 #58)
+#56 := (- #46 #55)
+#52 := 3::real
+#59 := (- #58 #55)
+#61 := (- #59)
+#60 := (< #59 0::real)
+#62 := (ite #60 #61 #59)
+#63 := (* #62 3::real)
+#64 := (< #63 #56)
+#66 := (and #64 #65)
+#47 := (- #44 #46)
+#50 := (- #47)
+#49 := (< #47 0::real)
+#51 := (ite #49 #50 #47)
+#53 := (* #51 3::real)
+#57 := (< #53 #56)
+#67 := (and #57 #66)
+#254 := (iff #67 #249)
+#138 := (+ #46 #137)
+#147 := (< #144 0::real)
+#159 := (ite #147 #154 #144)
+#165 := (* 3::real #159)
+#170 := (< #165 #138)
+#176 := (and #65 #170)
+#114 := (< #111 0::real)
+#126 := (ite #114 #121 #111)
+#132 := (* 3::real #126)
+#141 := (< #132 #138)
+#181 := (and #141 #176)
+#252 := (iff #181 #249)
+#243 := (and #185 #211)
+#246 := (and #238 #243)
+#250 := (iff #246 #249)
+#251 := [rewrite]: #250
+#247 := (iff #181 #246)
+#244 := (iff #176 #243)
+#214 := (iff #170 #211)
+#200 := (* 3::real #195)
+#203 := (< #200 #138)
+#212 := (iff #203 #211)
+#213 := [rewrite]: #212
+#204 := (iff #170 #203)
+#201 := (= #165 #200)
+#198 := (= #159 #195)
+#189 := (not #188)
+#192 := (ite #189 #154 #144)
+#196 := (= #192 #195)
+#197 := [rewrite]: #196
+#193 := (= #159 #192)
+#190 := (iff #147 #189)
+#191 := [rewrite]: #190
+#194 := [monotonicity #191]: #193
+#199 := [trans #194 #197]: #198
+#202 := [monotonicity #199]: #201
+#205 := [monotonicity #202]: #204
+#215 := [trans #205 #213]: #214
+#186 := (iff #65 #185)
+#187 := [rewrite]: #186
+#245 := [monotonicity #187 #215]: #244
+#241 := (iff #141 #238)
+#228 := (* 3::real #223)
+#231 := (< #228 #138)
+#239 := (iff #231 #238)
+#240 := [rewrite]: #239
+#232 := (iff #141 #231)
+#229 := (= #132 #228)
+#226 := (= #126 #223)
+#220 := (ite #217 #121 #111)
+#224 := (= #220 #223)
+#225 := [rewrite]: #224
+#221 := (= #126 #220)
+#218 := (iff #114 #217)
+#219 := [rewrite]: #218
+#222 := [monotonicity #219]: #221
+#227 := [trans #222 #225]: #226
+#230 := [monotonicity #227]: #229
+#233 := [monotonicity #230]: #232
+#242 := [trans #233 #240]: #241
+#248 := [monotonicity #242 #245]: #247
+#253 := [trans #248 #251]: #252
+#182 := (iff #67 #181)
+#179 := (iff #66 #176)
+#173 := (and #170 #65)
+#177 := (iff #173 #176)
+#178 := [rewrite]: #177
+#174 := (iff #66 #173)
+#171 := (iff #64 #170)
+#139 := (= #56 #138)
+#140 := [rewrite]: #139
+#168 := (= #63 #165)
+#162 := (* #159 3::real)
+#166 := (= #162 #165)
+#167 := [rewrite]: #166
+#163 := (= #63 #162)
+#160 := (= #62 #159)
+#145 := (= #59 #144)
+#146 := [rewrite]: #145
+#157 := (= #61 #154)
+#150 := (- #144)
+#155 := (= #150 #154)
+#156 := [rewrite]: #155
+#151 := (= #61 #150)
+#152 := [monotonicity #146]: #151
+#158 := [trans #152 #156]: #157
+#148 := (iff #60 #147)
+#149 := [monotonicity #146]: #148
+#161 := [monotonicity #149 #158 #146]: #160
+#164 := [monotonicity #161]: #163
+#169 := [trans #164 #167]: #168
+#172 := [monotonicity #169 #140]: #171
+#175 := [monotonicity #172]: #174
+#180 := [trans #175 #178]: #179
+#142 := (iff #57 #141)
+#135 := (= #53 #132)
+#129 := (* #126 3::real)
+#133 := (= #129 #132)
+#134 := [rewrite]: #133
+#130 := (= #53 #129)
+#127 := (= #51 #126)
+#112 := (= #47 #111)
+#113 := [rewrite]: #112
+#124 := (= #50 #121)
+#117 := (- #111)
+#122 := (= #117 #121)
+#123 := [rewrite]: #122
+#118 := (= #50 #117)
+#119 := [monotonicity #113]: #118
+#125 := [trans #119 #123]: #124
+#115 := (iff #49 #114)
+#116 := [monotonicity #113]: #115
+#128 := [monotonicity #116 #125 #113]: #127
+#131 := [monotonicity #128]: #130
+#136 := [trans #131 #134]: #135
+#143 := [monotonicity #136 #140]: #142
+#183 := [monotonicity #143 #180]: #182
+#255 := [trans #183 #253]: #254
+#108 := [asserted]: #67
+#256 := [mp #108 #255]: #249
+#257 := [and-elim #256]: #185
+#259 := [and-elim #256]: #238
+#448 := (+ #111 #447)
+#449 := (<= #448 0::real)
+#441 := (= #111 #223)
+#443 := (or #217 #441)
+#444 := [def-axiom]: #443
+#467 := [unit-resolution #444 #455]: #441
+#468 := (not #441)
+#469 := (or #468 #449)
+#470 := [th-lemma]: #469
+#471 := [unit-resolution #470 #467]: #449
+#463 := (or #189 #217)
+#456 := [hypothesis]: #188
+#258 := [and-elim #256]: #211
+#437 := (or #189 #435)
+#438 := [def-axiom]: #437
+#457 := [unit-resolution #438 #456]: #435
+#458 := (not #435)
+#459 := (or #458 #453)
+#460 := [th-lemma]: #459
+#461 := [unit-resolution #460 #457]: #453
+#462 := [th-lemma #257 #461 #258 #456 #455]: false
+#464 := [lemma #462]: #463
+#472 := [unit-resolution #464 #455]: #189
+#473 := [th-lemma #472 #471 #259 #257 #455]: false
+#474 := [lemma #473]: #217
+#445 := (or #216 #442)
+#446 := [def-axiom]: #445
+#475 := [unit-resolution #446 #474]: #442
+#476 := (not #442)
+#477 := (or #476 #454)
+#478 := [th-lemma]: #477
+#479 := [unit-resolution #478 #475]: #454
+#481 := (not #185)
+#480 := (not #454)
+#482 := (or #188 #480 #237 #481 #216)
+#483 := [th-lemma]: #482
+#484 := [unit-resolution #483 #257 #474 #259 #479]: #188
+#485 := [unit-resolution #438 #484]: #435
+#486 := [unit-resolution #460 #485]: #453
+[th-lemma #479 #259 #257 #474 #258 #486]: false
+unsat
+e14093e1f049d6c695e00b050a39f30b5ee75c84 41 0
+f1 -> val!6
+f2 -> val!0
+?v0!0 -> val!11
+f12 -> val!1
+f14 -> val!2
+f15 -> val!3
+f9 -> val!7
+f8 -> val!8
+f6 -> val!9
+f13 -> {
+  val!2 val!3 -> val!4
+  else -> val!4
+}
+f11 -> {
+  val!1 val!4 -> val!5
+  else -> val!5
+}
+f10 -> {
+  val!11 val!5 -> val!6
+  else -> val!6
+}
+f7 -> {
+  val!7 val!11 -> 0
+  val!8 val!11 -> 0
+  else -> 0
+}
+f5 -> {
+  val!9 -> val!10
+  else -> val!10
+}
+f4 -> {
+  val!10 -> 0
+  val!11 -> 38
+  else -> 38
+}
+f3 -> {
+  0 -> val!10
+  38 -> val!11
+  else -> val!11
+}
+unknown
+b39ccddf1d4f138215d1b195d383905c9937a82f 44 0
+f1 -> val!7
+f2 -> val!0
+?v0!0 -> val!11
+f6 -> val!1
+f14 -> val!3
+f15 -> val!4
+f9 -> val!8
+f8 -> val!9
+f5 -> {
+  val!1 -> val!10
+  else -> val!10
+}
+f12 -> {
+  val!10 -> val!2
+  else -> val!2
+}
+f13 -> {
+  val!3 val!4 -> val!5
+  else -> val!5
+}
+f11 -> {
+  val!2 val!5 -> val!6
+  else -> val!6
+}
+f10 -> {
+  val!11 val!6 -> val!7
+  else -> val!7
+}
+f7 -> {
+  val!8 val!11 -> 0
+  val!9 val!11 -> 0
+  else -> 0
+}
+f4 -> {
+  val!10 -> 0
+  val!11 -> 38
+  else -> 38
+}
+f3 -> {
+  0 -> val!10
+  38 -> val!11
+  else -> val!11
+}
+unknown
+e03a7d5d1030a490d5669c1c764d61d605302d39 48 0
+f1 -> val!6
+f2 -> val!0
+f5 -> val!9
+f7 -> val!1
+?v0!0 -> val!11
+f15 -> val!3
+f10 -> val!7
+f9 -> val!8
+f4 -> {
+  val!9 -> 0
+  val!10 -> 1
+  val!11 -> 1237
+  else -> 1237
+}
+f6 -> {
+  val!1 -> val!10
+  else -> val!10
+}
+f13 -> {
+  val!10 -> val!2
+  else -> val!2
+}
+f14 -> {
+  val!9 val!3 -> val!4
+  else -> val!4
+}
+f12 -> {
+  val!2 val!4 -> val!5
+  else -> val!5
+}
+f11 -> {
+  val!11 val!5 -> val!6
+  else -> val!6
+}
+f8 -> {
+  val!7 val!11 -> 0
+  val!8 val!11 -> 0
+  val!7 val!9 -> -1
+  val!8 val!9 -> 0
+  else -> 0
+}
+f3 -> {
+  0 -> val!9
+  1 -> val!10
+  1237 -> val!11
+  else -> val!11
+}
+unknown
+e7836464956771492f5454e164bcd43446470f5a 48 0
+f1 -> val!6
+f2 -> val!0
+f5 -> val!9
+f7 -> val!1
+?v0!0 -> val!11
+f15 -> val!3
+f10 -> val!7
+f9 -> val!8
+f4 -> {
+  val!9 -> 0
+  val!10 -> 1
+  val!11 -> 1237
+  else -> 1237
+}
+f6 -> {
+  val!1 -> val!10
+  else -> val!10
+}
+f13 -> {
+  val!10 -> val!2
+  else -> val!2
+}
+f14 -> {
+  val!9 val!3 -> val!4
+  else -> val!4
+}
+f12 -> {
+  val!2 val!4 -> val!5
+  else -> val!5
+}
+f11 -> {
+  val!11 val!5 -> val!6
+  else -> val!6
+}
+f8 -> {
+  val!7 val!11 -> 0
+  val!8 val!11 -> 0
+  val!7 val!9 -> 1
+  val!8 val!9 -> 0
+  else -> 0
+}
+f3 -> {
+  0 -> val!9
+  1 -> val!10
+  1237 -> val!11
+  else -> val!11
+}
+unknown
+94aef9051873492c2e350c3ee8f6b777964eebe4 147 0
+#2 := false
+#159 := 0::real
+decl f8 :: (-> S4 S2 real)
+decl f5 :: S2
+#25 := f5
+decl f10 :: S4
+#33 := f10
+#34 := (f8 f10 f5)
+#156 := -1::real
+#157 := (* -1::real #34)
+decl f9 :: S4
+#31 := f9
+#32 := (f8 f9 f5)
+#158 := (+ #32 #157)
+#315 := (>= #158 0::real)
+#392 := (not #315)
+#202 := (= #32 #34)
+#205 := (not #202)
+#43 := (= #34 #32)
+#44 := (not #43)
+#206 := (iff #44 #205)
+#203 := (iff #43 #202)
+#204 := [rewrite]: #203
+#207 := [monotonicity #204]: #206
+#201 := [asserted]: #44
+#210 := [mp #201 #207]: #205
+#395 := (or #202 #392)
+#160 := (<= #158 0::real)
+#13 := 0::int
+decl f4 :: (-> S2 int)
+decl f6 :: (-> S3 S2)
+decl f7 :: S3
+#27 := f7
+#28 := (f6 f7)
+#29 := (f4 #28)
+#149 := -1::int
+#152 := (* -1::int #29)
+#26 := (f4 f5)
+#153 := (+ #26 #152)
+#151 := (>= #153 0::int)
+#150 := (not #151)
+#163 := (and #150 #160)
+#35 := (<= #32 #34)
+#30 := (< #26 #29)
+#36 := (and #30 #35)
+#164 := (iff #36 #163)
+#161 := (iff #35 #160)
+#162 := [rewrite]: #161
+#154 := (iff #30 #150)
+#155 := [rewrite]: #154
+#165 := [monotonicity #155 #162]: #164
+#146 := [asserted]: #36
+#166 := [mp #146 #165]: #163
+#168 := [and-elim #166]: #160
+#391 := (not #160)
+#393 := (or #202 #391 #392)
+#394 := [th-lemma]: #393
+#396 := [unit-resolution #394 #168]: #395
+#397 := [unit-resolution #396 #210]: #392
+#8 := (:var 0 S2)
+#39 := (f8 f9 #8)
+#308 := (pattern #39)
+#38 := (f8 f10 #8)
+#307 := (pattern #38)
+#9 := (f4 #8)
+#287 := (pattern #9)
+#187 := (* -1::real #39)
+#188 := (+ #38 #187)
+#189 := (<= #188 0::real)
+#177 := (+ #9 #152)
+#176 := (>= #177 0::int)
+#192 := (or #176 #189)
+#309 := (forall (vars (?v0 S2)) (:pat #287 #307 #308) #192)
+#195 := (forall (vars (?v0 S2)) #192)
+#312 := (iff #195 #309)
+#310 := (iff #192 #192)
+#311 := [refl]: #310
+#313 := [quant-intro #311]: #312
+#227 := (~ #195 #195)
+#208 := (~ #192 #192)
+#226 := [refl]: #208
+#228 := [nnf-pos #226]: #227
+#40 := (<= #38 #39)
+#37 := (< #9 #29)
+#41 := (implies #37 #40)
+#42 := (forall (vars (?v0 S2)) #41)
+#198 := (iff #42 #195)
+#148 := (not #37)
+#169 := (or #148 #40)
+#172 := (forall (vars (?v0 S2)) #169)
+#196 := (iff #172 #195)
+#193 := (iff #169 #192)
+#190 := (iff #40 #189)
+#191 := [rewrite]: #190
+#185 := (iff #148 #176)
+#175 := (not #176)
+#180 := (not #175)
+#183 := (iff #180 #176)
+#184 := [rewrite]: #183
+#181 := (iff #148 #180)
+#178 := (iff #37 #175)
+#179 := [rewrite]: #178
+#182 := [monotonicity #179]: #181
+#186 := [trans #182 #184]: #185
+#194 := [monotonicity #186 #191]: #193
+#197 := [quant-intro #194]: #196
+#173 := (iff #42 #172)
+#170 := (iff #41 #169)
+#171 := [rewrite]: #170
+#174 := [quant-intro #171]: #173
+#199 := [trans #174 #197]: #198
+#147 := [asserted]: #42
+#200 := [mp #147 #199]: #195
+#229 := [mp~ #200 #228]: #195
+#314 := [mp #229 #313]: #309
+#167 := [and-elim #166]: #150
+#338 := (not #309)
+#339 := (or #338 #151 #315)
+#318 := (* -1::real #32)
+#319 := (+ #34 #318)
+#323 := (<= #319 0::real)
+#324 := (or #151 #323)
+#340 := (or #338 #324)
+#347 := (iff #340 #339)
+#335 := (or #151 #315)
+#342 := (or #338 #335)
+#345 := (iff #342 #339)
+#346 := [rewrite]: #345
+#343 := (iff #340 #342)
+#336 := (iff #324 #335)
+#333 := (iff #323 #315)
+#325 := (+ #318 #34)
+#328 := (<= #325 0::real)
+#331 := (iff #328 #315)
+#332 := [rewrite]: #331
+#329 := (iff #323 #328)
+#326 := (= #319 #325)
+#327 := [rewrite]: #326
+#330 := [monotonicity #327]: #329
+#334 := [trans #330 #332]: #333
+#337 := [monotonicity #334]: #336
+#344 := [monotonicity #337]: #343
+#348 := [trans #344 #346]: #347
+#341 := [quant-inst]: #340
+#349 := [mp #341 #348]: #339
+[unit-resolution #349 #167 #314 #397]: false
+unsat
+a776c8d80c92e36c87766b3cd2a3d1707a6a4f17 29 0
+f1 -> val!0
+f2 -> val!1
+f6 -> val!2
+f7 -> val!7
+f8 -> val!3
+f12 -> 1
+f11 -> val!4
+f10 -> val!5
+f5 -> {
+  val!2 val!7 -> 0
+  val!3 val!7 -> 0
+  val!4 val!7 -> 1
+  else -> 1
+}
+f9 -> {
+  val!5 -> val!6
+  else -> val!6
+}
+f4 -> {
+  val!6 -> 0
+  val!7 -> 7719
+  else -> 7719
+}
+f3 -> {
+  0 -> val!6
+  7719 -> val!7
+  else -> val!7
+}
+unknown
+2b1630352ec036f9060c61ea0f08be80276b84d2 29 0
+f1 -> val!0
+f2 -> val!1
+f6 -> val!2
+f7 -> val!7
+f8 -> val!3
+f12 -> 1
+f11 -> val!4
+f10 -> val!5
+f5 -> {
+  val!2 val!7 -> 0
+  val!3 val!7 -> 0
+  val!4 val!7 -> 1
+  else -> 1
+}
+f9 -> {
+  val!5 -> val!6
+  else -> val!6
+}
+f4 -> {
+  val!6 -> 0
+  val!7 -> 7719
+  else -> 7719
+}
+f3 -> {
+  0 -> val!6
+  7719 -> val!7
+  else -> val!7
+}
+unknown
+478564f7f355dca34baaecde71bc3afeb25b293b 204 0
+#2 := false
+#46 := 0::real
+decl f5 :: (-> S3 S2 real)
+decl f7 :: S2
+#26 := f7
+decl f11 :: S3
+#38 := f11
+#49 := (f5 f11 f7)
+#183 := -1::real
+#349 := (* -1::real #49)
+decl f8 :: S3
+#28 := f8
+#29 := (f5 f8 f7)
+#362 := (+ #29 #349)
+#363 := (>= #362 0::real)
+#368 := (not #363)
+decl f6 :: S3
+#25 := f6
+#27 := (f5 f6 f7)
+#350 := (+ #27 #349)
+#351 := (<= #350 0::real)
+#352 := (not #351)
+#371 := (or #352 #368)
+#374 := (not #371)
+#8 := (:var 0 S2)
+#41 := (f5 f8 #8)
+#333 := (pattern #41)
+#39 := (f5 f11 #8)
+#332 := (pattern #39)
+#37 := (f5 f6 #8)
+#331 := (pattern #37)
+decl f4 :: (-> S2 int)
+#9 := (f4 #8)
+#311 := (pattern #9)
+#189 := (* -1::real #41)
+#190 := (+ #39 #189)
+#191 := (<= #190 0::real)
+#242 := (not #191)
+#184 := (* -1::real #39)
+#185 := (+ #37 #184)
+#186 := (<= #185 0::real)
+#241 := (not #186)
+#243 := (or #241 #242)
+#244 := (not #243)
+#13 := 0::int
+decl f9 :: (-> S4 S2)
+decl f10 :: S4
+#32 := f10
+#33 := (f9 f10)
+#34 := (f4 #33)
+#157 := -1::int
+#160 := (* -1::int #34)
+#173 := (+ #9 #160)
+#172 := (>= #173 0::int)
+#247 := (or #172 #244)
+#334 := (forall (vars (?v0 S2)) (:pat #311 #331 #332 #333) #247)
+#250 := (forall (vars (?v0 S2)) #247)
+#337 := (iff #250 #334)
+#335 := (iff #247 #247)
+#336 := [refl]: #335
+#338 := [quant-intro #336]: #337
+#194 := (and #186 #191)
+#197 := (or #172 #194)
+#200 := (forall (vars (?v0 S2)) #197)
+#251 := (iff #200 #250)
+#248 := (iff #197 #247)
+#245 := (iff #194 #244)
+#246 := [rewrite]: #245
+#249 := [monotonicity #246]: #248
+#252 := [quant-intro #249]: #251
+#219 := (~ #200 #200)
+#224 := (~ #197 #197)
+#222 := [refl]: #224
+#239 := [nnf-pos #222]: #219
+#42 := (<= #39 #41)
+#40 := (<= #37 #39)
+#43 := (and #40 #42)
+#36 := (< #9 #34)
+#44 := (implies #36 #43)
+#45 := (forall (vars (?v0 S2)) #44)
+#203 := (iff #45 #200)
+#156 := (not #36)
+#165 := (or #156 #43)
+#168 := (forall (vars (?v0 S2)) #165)
+#201 := (iff #168 #200)
+#198 := (iff #165 #197)
+#195 := (iff #43 #194)
+#192 := (iff #42 #191)
+#193 := [rewrite]: #192
+#187 := (iff #40 #186)
+#188 := [rewrite]: #187
+#196 := [monotonicity #188 #193]: #195
+#181 := (iff #156 #172)
+#171 := (not #172)
+#176 := (not #171)
+#179 := (iff #176 #172)
+#180 := [rewrite]: #179
+#177 := (iff #156 #176)
+#174 := (iff #36 #171)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#182 := [trans #178 #180]: #181
+#199 := [monotonicity #182 #196]: #198
+#202 := [quant-intro #199]: #201
+#169 := (iff #45 #168)
+#166 := (iff #44 #165)
+#167 := [rewrite]: #166
+#170 := [quant-intro #167]: #169
+#204 := [trans #170 #202]: #203
+#155 := [asserted]: #45
+#205 := [mp #155 #204]: #200
+#240 := [mp~ #205 #239]: #200
+#253 := [mp #240 #252]: #250
+#339 := [mp #253 #338]: #334
+#31 := (f4 f7)
+#161 := (+ #31 #160)
+#159 := (>= #161 0::int)
+#158 := (not #159)
+#35 := (< #31 #34)
+#162 := (iff #35 #158)
+#163 := [rewrite]: #162
+#154 := [asserted]: #35
+#164 := [mp #154 #163]: #158
+#380 := (not #334)
+#381 := (or #380 #159 #374)
+#342 := (* -1::real #29)
+#343 := (+ #49 #342)
+#347 := (<= #343 0::real)
+#348 := (not #347)
+#353 := (or #352 #348)
+#354 := (not #353)
+#355 := (or #159 #354)
+#382 := (or #380 #355)
+#389 := (iff #382 #381)
+#377 := (or #159 #374)
+#384 := (or #380 #377)
+#387 := (iff #384 #381)
+#388 := [rewrite]: #387
+#385 := (iff #382 #384)
+#378 := (iff #355 #377)
+#375 := (iff #354 #374)
+#372 := (iff #353 #371)
+#369 := (iff #348 #368)
+#366 := (iff #347 #363)
+#356 := (+ #342 #49)
+#359 := (<= #356 0::real)
+#364 := (iff #359 #363)
+#365 := [rewrite]: #364
+#360 := (iff #347 #359)
+#357 := (= #343 #356)
+#358 := [rewrite]: #357
+#361 := [monotonicity #358]: #360
+#367 := [trans #361 #365]: #366
+#370 := [monotonicity #367]: #369
+#373 := [monotonicity #370]: #372
+#376 := [monotonicity #373]: #375
+#379 := [monotonicity #376]: #378
+#386 := [monotonicity #379]: #385
+#390 := [trans #386 #388]: #389
+#383 := [quant-inst]: #382
+#391 := [mp #383 #390]: #381
+#513 := [unit-resolution #391 #164 #339]: #374
+#394 := (or #371 #363)
+#395 := [def-axiom]: #394
+#514 := [unit-resolution #395 #513]: #363
+#475 := (>= #350 0::real)
+#524 := (not #475)
+#474 := (= #27 #49)
+#519 := (not #474)
+#208 := (= #29 #49)
+#216 := (not #208)
+#520 := (iff #216 #519)
+#517 := (iff #208 #474)
+#515 := (iff #474 #208)
+#30 := (= #27 #29)
+#153 := [asserted]: #30
+#516 := [monotonicity #153]: #515
+#518 := [symm #516]: #517
+#521 := [monotonicity #518]: #520
+#50 := (= #49 #29)
+#51 := (not #50)
+#217 := (iff #51 #216)
+#214 := (iff #50 #208)
+#215 := [rewrite]: #214
+#218 := [monotonicity #215]: #217
+#207 := [asserted]: #51
+#221 := [mp #207 #218]: #216
+#522 := [mp #221 #521]: #519
+#527 := (or #474 #524)
+#392 := (or #371 #351)
+#393 := [def-axiom]: #392
+#523 := [unit-resolution #393 #513]: #351
+#525 := (or #474 #352 #524)
+#526 := [th-lemma]: #525
+#528 := [unit-resolution #526 #523]: #527
+#529 := [unit-resolution #528 #522]: #524
+#471 := (+ #27 #342)
+#473 := (>= #471 0::real)
+#530 := (not #30)
+#531 := (or #530 #473)
+#532 := [th-lemma]: #531
+#533 := [unit-resolution #532 #153]: #473
+[th-lemma #533 #529 #514]: false
+unsat
+07499e3240b5b9ff9168a39bf4f0848339ecb137 29 0
+f5 -> 1
+f12 -> val!0
+f11 -> val!7
+f9 -> val!1
+f1 -> val!2
+f2 -> val!3
+f10 -> val!4
+f7 -> val!5
+f8 -> {
+  val!0 val!7 -> 0
+  val!1 val!7 -> 0
+  val!4 val!7 -> 0
+  else -> 0
+}
+f6 -> {
+  val!5 -> val!6
+  else -> val!6
+}
+f4 -> {
+  val!6 -> 0
+  val!7 -> 38
+  else -> 38
+}
+f3 -> {
+  0 -> val!6
+  38 -> val!7
+  else -> val!7
+}
+unknown
+3521e5c8a799f779cef25ded86c422c58ebfc7fd 318 0
+#2 := false
+#25 := 0::real
+decl f8 :: (-> S4 S2 real)
+decl f11 :: S2
+#40 := f11
+decl f9 :: S4
+#32 := f9
+#47 := (f8 f9 f11)
+decl f12 :: S4
+#44 := f12
+#45 := (f8 f12 f11)
+#192 := -1::real
+#232 := (* -1::real #45)
+#233 := (+ #232 #47)
+decl f5 :: real
+#26 := f5
+#268 := (* -1::real #47)
+#271 := (+ #45 #268)
+#274 := (+ f5 #271)
+#275 := (<= #274 0::real)
+#278 := (ite #275 f5 #233)
+#593 := (* -1::real #278)
+#594 := (+ f5 #593)
+#595 := (<= #594 0::real)
+#603 := (not #595)
+#223 := 1/2::real
+#281 := (* 1/2::real #278)
+#457 := (<= #281 0::real)
+#292 := (= #281 0::real)
+#306 := (<= #271 0::real)
+decl f10 :: S4
+#34 := f10
+#43 := (f8 f10 f11)
+#302 := (+ #43 #232)
+#303 := (<= #302 0::real)
+#309 := (and #303 #306)
+#13 := 0::int
+decl f4 :: (-> S2 int)
+#41 := (f4 f11)
+#178 := -1::int
+#213 := (* -1::int #41)
+decl f6 :: (-> S3 S2)
+decl f7 :: S3
+#28 := f7
+#29 := (f6 f7)
+#30 := (f4 #29)
+#214 := (+ #30 #213)
+#215 := (<= #214 0::int)
+#312 := (or #215 #309)
+#225 := (* 1/2::real #47)
+#272 := (+ #232 #225)
+#224 := (* 1/2::real #43)
+#273 := (+ #224 #272)
+#270 := (>= #273 0::real)
+#321 := (and #270 #292 #312)
+#52 := 2::real
+#55 := (- #47 #45)
+#56 := (<= f5 #55)
+#57 := (ite #56 f5 #55)
+#58 := (/ #57 2::real)
+#59 := (+ #45 #58)
+#60 := (= #59 #45)
+#51 := (+ #43 #47)
+#53 := (/ #51 2::real)
+#54 := (<= #45 #53)
+#61 := (and #54 #60)
+#48 := (<= #45 #47)
+#46 := (<= #43 #45)
+#49 := (and #46 #48)
+#42 := (< #41 #30)
+#50 := (implies #42 #49)
+#62 := (and #50 #61)
+#326 := (iff #62 #321)
+#236 := (<= f5 #233)
+#239 := (ite #236 f5 #233)
+#245 := (* 1/2::real #239)
+#250 := (+ #45 #245)
+#256 := (= #45 #250)
+#226 := (+ #224 #225)
+#229 := (<= #45 #226)
+#261 := (and #229 #256)
+#212 := (not #42)
+#220 := (or #212 #49)
+#264 := (and #220 #261)
+#324 := (iff #264 #321)
+#315 := (and #270 #292)
+#318 := (and #312 #315)
+#322 := (iff #318 #321)
+#323 := [rewrite]: #322
+#319 := (iff #264 #318)
+#316 := (iff #261 #315)
+#293 := (iff #256 #292)
+#284 := (+ #45 #281)
+#287 := (= #45 #284)
+#290 := (iff #287 #292)
+#291 := [rewrite]: #290
+#288 := (iff #256 #287)
+#285 := (= #250 #284)
+#282 := (= #245 #281)
+#279 := (= #239 #278)
+#276 := (iff #236 #275)
+#277 := [rewrite]: #276
+#280 := [monotonicity #277]: #279
+#283 := [monotonicity #280]: #282
+#286 := [monotonicity #283]: #285
+#289 := [monotonicity #286]: #288
+#294 := [trans #289 #291]: #293
+#267 := (iff #229 #270)
+#269 := [rewrite]: #267
+#317 := [monotonicity #269 #294]: #316
+#313 := (iff #220 #312)
+#310 := (iff #49 #309)
+#307 := (iff #48 #306)
+#308 := [rewrite]: #307
+#304 := (iff #46 #303)
+#305 := [rewrite]: #304
+#311 := [monotonicity #305 #308]: #310
+#300 := (iff #212 #215)
+#216 := (not #215)
+#295 := (not #216)
+#298 := (iff #295 #215)
+#299 := [rewrite]: #298
+#296 := (iff #212 #295)
+#217 := (iff #42 #216)
+#218 := [rewrite]: #217
+#297 := [monotonicity #218]: #296
+#301 := [trans #297 #299]: #300
+#314 := [monotonicity #301 #311]: #313
+#320 := [monotonicity #314 #317]: #319
+#325 := [trans #320 #323]: #324
+#265 := (iff #62 #264)
+#262 := (iff #61 #261)
+#259 := (iff #60 #256)
+#253 := (= #250 #45)
+#257 := (iff #253 #256)
+#258 := [rewrite]: #257
+#254 := (iff #60 #253)
+#251 := (= #59 #250)
+#248 := (= #58 #245)
+#242 := (/ #239 2::real)
+#246 := (= #242 #245)
+#247 := [rewrite]: #246
+#243 := (= #58 #242)
+#240 := (= #57 #239)
+#234 := (= #55 #233)
+#235 := [rewrite]: #234
+#237 := (iff #56 #236)
+#238 := [monotonicity #235]: #237
+#241 := [monotonicity #238 #235]: #240
+#244 := [monotonicity #241]: #243
+#249 := [trans #244 #247]: #248
+#252 := [monotonicity #249]: #251
+#255 := [monotonicity #252]: #254
+#260 := [trans #255 #258]: #259
+#230 := (iff #54 #229)
+#227 := (= #53 #226)
+#228 := [rewrite]: #227
+#231 := [monotonicity #228]: #230
+#263 := [monotonicity #231 #260]: #262
+#221 := (iff #50 #220)
+#222 := [rewrite]: #221
+#266 := [monotonicity #222 #263]: #265
+#327 := [trans #266 #325]: #326
+#211 := [asserted]: #62
+#328 := [mp #211 #327]: #321
+#330 := [and-elim #328]: #292
+#597 := (not #292)
+#598 := (or #597 #457)
+#599 := [th-lemma]: #598
+#600 := [unit-resolution #599 #330]: #457
+#601 := [hypothesis]: #595
+#167 := (<= f5 0::real)
+#168 := (not #167)
+#27 := (< 0::real f5)
+#169 := (iff #27 #168)
+#170 := [rewrite]: #169
+#164 := [asserted]: #27
+#171 := [mp #164 #170]: #168
+#602 := [th-lemma #171 #601 #600]: false
+#604 := [lemma #602]: #603
+#450 := (= f5 #278)
+#451 := (= #233 #278)
+#613 := (not #451)
+#596 := (+ #233 #593)
+#605 := (<= #596 0::real)
+#610 := (not #605)
+#532 := (+ #43 #268)
+#533 := (>= #532 0::real)
+#538 := (not #533)
+#210 := [asserted]: #42
+#219 := [mp #210 #218]: #216
+#8 := (:var 0 S2)
+#35 := (f8 f10 #8)
+#443 := (pattern #35)
+#33 := (f8 f9 #8)
+#442 := (pattern #33)
+#9 := (f4 #8)
+#422 := (pattern #9)
+#193 := (* -1::real #35)
+#194 := (+ #33 #193)
+#195 := (<= #194 0::real)
+#198 := (not #195)
+#181 := (* -1::int #30)
+#182 := (+ #9 #181)
+#180 := (>= #182 0::int)
+#201 := (or #180 #198)
+#444 := (forall (vars (?v0 S2)) (:pat #422 #442 #443) #201)
+#204 := (forall (vars (?v0 S2)) #201)
+#447 := (iff #204 #444)
+#445 := (iff #201 #201)
+#446 := [refl]: #445
+#448 := [quant-intro #446]: #447
+#362 := (~ #204 #204)
+#332 := (~ #201 #201)
+#361 := [refl]: #332
+#363 := [nnf-pos #361]: #362
+#36 := (<= #33 #35)
+#37 := (not #36)
+#31 := (< #9 #30)
+#38 := (implies #31 #37)
+#39 := (forall (vars (?v0 S2)) #38)
+#207 := (iff #39 #204)
+#166 := (not #31)
+#172 := (or #166 #37)
+#175 := (forall (vars (?v0 S2)) #172)
+#205 := (iff #175 #204)
+#202 := (iff #172 #201)
+#199 := (iff #37 #198)
+#196 := (iff #36 #195)
+#197 := [rewrite]: #196
+#200 := [monotonicity #197]: #199
+#190 := (iff #166 #180)
+#179 := (not #180)
+#185 := (not #179)
+#188 := (iff #185 #180)
+#189 := [rewrite]: #188
+#186 := (iff #166 #185)
+#183 := (iff #31 #179)
+#184 := [rewrite]: #183
+#187 := [monotonicity #184]: #186
+#191 := [trans #187 #189]: #190
+#203 := [monotonicity #191 #200]: #202
+#206 := [quant-intro #203]: #205
+#176 := (iff #39 #175)
+#173 := (iff #38 #172)
+#174 := [rewrite]: #173
+#177 := [quant-intro #174]: #176
+#208 := [trans #177 #206]: #207
+#165 := [asserted]: #39
+#209 := [mp #165 #208]: #204
+#364 := [mp~ #209 #363]: #204
+#449 := [mp #364 #448]: #444
+#544 := (not #444)
+#545 := (or #544 #215 #538)
+#507 := (* -1::real #43)
+#508 := (+ #47 #507)
+#511 := (<= #508 0::real)
+#512 := (not #511)
+#513 := (+ #41 #181)
+#514 := (>= #513 0::int)
+#515 := (or #514 #512)
+#546 := (or #544 #515)
+#553 := (iff #546 #545)
+#541 := (or #215 #538)
+#548 := (or #544 #541)
+#551 := (iff #548 #545)
+#552 := [rewrite]: #551
+#549 := (iff #546 #548)
+#542 := (iff #515 #541)
+#539 := (iff #512 #538)
+#536 := (iff #511 #533)
+#526 := (+ #507 #47)
+#529 := (<= #526 0::real)
+#534 := (iff #529 #533)
+#535 := [rewrite]: #534
+#530 := (iff #511 #529)
+#527 := (= #508 #526)
+#528 := [rewrite]: #527
+#531 := [monotonicity #528]: #530
+#537 := [trans #531 #535]: #536
+#540 := [monotonicity #537]: #539
+#524 := (iff #514 #215)
+#516 := (+ #181 #41)
+#519 := (>= #516 0::int)
+#522 := (iff #519 #215)
+#523 := [rewrite]: #522
+#520 := (iff #514 #519)
+#517 := (= #513 #516)
+#518 := [rewrite]: #517
+#521 := [monotonicity #518]: #520
+#525 := [trans #521 #523]: #524
+#543 := [monotonicity #525 #540]: #542
+#550 := [monotonicity #543]: #549
+#554 := [trans #550 #552]: #553
+#547 := [quant-inst]: #546
+#555 := [mp #547 #554]: #545
+#607 := [unit-resolution #555 #449 #219]: #538
+#329 := [and-elim #328]: #270
+#608 := [hypothesis]: #605
+#609 := [th-lemma #608 #329 #607 #600]: false
+#611 := [lemma #609]: #610
+#612 := [hypothesis]: #451
+#614 := (or #613 #605)
+#615 := [th-lemma]: #614
+#616 := [unit-resolution #615 #612 #611]: false
+#617 := [lemma #616]: #613
+#455 := (or #275 #451)
+#456 := [def-axiom]: #455
+#618 := [unit-resolution #456 #617]: #275
+#452 := (not #275)
+#453 := (or #452 #450)
+#454 := [def-axiom]: #453
+#619 := [unit-resolution #454 #618]: #450
+#620 := (not #450)
+#621 := (or #620 #595)
+#622 := [th-lemma]: #621
+[unit-resolution #622 #619 #604]: false
+unsat
+f7a332c4ab50576b47f6154d9c204a030b7f3346 295 0
+#2 := false
+#25 := 0::real
+decl f8 :: (-> S4 S2 real)
+decl f11 :: S2
+#40 := f11
+decl f12 :: S4
+#44 := f12
+#45 := (f8 f12 f11)
+decl f10 :: S4
+#34 := f10
+#43 := (f8 f10 f11)
+#199 := -1::real
+#263 := (* -1::real #43)
+#264 := (+ #263 #45)
+decl f5 :: real
+#26 := f5
+#242 := (* -1::real #45)
+#331 := (+ #43 #242)
+#332 := (+ f5 #331)
+#333 := (<= #332 0::real)
+#336 := (ite #333 f5 #264)
+#666 := (* -1::real #336)
+#667 := (+ f5 #666)
+#668 := (<= #667 0::real)
+#676 := (not #668)
+#230 := 1/2::real
+#414 := (* 1/2::real #336)
+#531 := (<= #414 0::real)
+#415 := (= #414 0::real)
+#284 := -1/2::real
+#339 := (* -1/2::real #336)
+#342 := (+ #45 #339)
+decl f9 :: S4
+#32 := f9
+#47 := (f8 f9 f11)
+#243 := (+ #242 #47)
+#316 := (* -1::real #47)
+#317 := (+ #45 #316)
+#318 := (+ f5 #317)
+#319 := (<= #318 0::real)
+#322 := (ite #319 f5 #243)
+#325 := (* 1/2::real #322)
+#328 := (+ #45 #325)
+#232 := (* 1/2::real #47)
+#312 := (+ #242 #232)
+#231 := (* 1/2::real #43)
+#313 := (+ #231 #312)
+#310 := (>= #313 0::real)
+#345 := (ite #310 #328 #342)
+#348 := (= #45 #345)
+#418 := (iff #348 #415)
+#411 := (= #45 #342)
+#416 := (iff #411 #415)
+#417 := [rewrite]: #416
+#412 := (iff #348 #411)
+#409 := (= #345 #342)
+#404 := (ite false #328 #342)
+#407 := (= #404 #342)
+#408 := [rewrite]: #407
+#405 := (= #345 #404)
+#402 := (iff #310 false)
+#309 := (not #310)
+#361 := (<= #317 0::real)
+#358 := (<= #331 0::real)
+#364 := (and #358 #361)
+#13 := 0::int
+decl f4 :: (-> S2 int)
+#41 := (f4 f11)
+#185 := -1::int
+#220 := (* -1::int #41)
+decl f6 :: (-> S3 S2)
+decl f7 :: S3
+#28 := f7
+#29 := (f6 f7)
+#30 := (f4 #29)
+#221 := (+ #30 #220)
+#222 := (<= #221 0::int)
+#367 := (or #222 #364)
+#376 := (and #309 #348 #367)
+#52 := 2::real
+#61 := (- #45 #43)
+#62 := (<= f5 #61)
+#63 := (ite #62 f5 #61)
+#64 := (/ #63 2::real)
+#65 := (- #45 #64)
+#56 := (- #47 #45)
+#57 := (<= f5 #56)
+#58 := (ite #57 f5 #56)
+#59 := (/ #58 2::real)
+#60 := (+ #45 #59)
+#51 := (+ #43 #47)
+#53 := (/ #51 2::real)
+#55 := (<= #45 #53)
+#66 := (ite #55 #60 #65)
+#67 := (= #66 #45)
+#54 := (< #53 #45)
+#68 := (and #54 #67)
+#48 := (<= #45 #47)
+#46 := (<= #43 #45)
+#49 := (and #46 #48)
+#42 := (< #41 #30)
+#50 := (implies #42 #49)
+#69 := (and #50 #68)
+#381 := (iff #69 #376)
+#267 := (<= f5 #264)
+#270 := (ite #267 f5 #264)
+#285 := (* -1/2::real #270)
+#286 := (+ #45 #285)
+#246 := (<= f5 #243)
+#249 := (ite #246 f5 #243)
+#255 := (* 1/2::real #249)
+#260 := (+ #45 #255)
+#233 := (+ #231 #232)
+#239 := (<= #45 #233)
+#291 := (ite #239 #260 #286)
+#297 := (= #45 #291)
+#236 := (< #233 #45)
+#302 := (and #236 #297)
+#219 := (not #42)
+#227 := (or #219 #49)
+#305 := (and #227 #302)
+#379 := (iff #305 #376)
+#370 := (and #309 #348)
+#373 := (and #367 #370)
+#377 := (iff #373 #376)
+#378 := [rewrite]: #377
+#374 := (iff #305 #373)
+#371 := (iff #302 #370)
+#349 := (iff #297 #348)
+#346 := (= #291 #345)
+#343 := (= #286 #342)
+#340 := (= #285 #339)
+#337 := (= #270 #336)
+#334 := (iff #267 #333)
+#335 := [rewrite]: #334
+#338 := [monotonicity #335]: #337
+#341 := [monotonicity #338]: #340
+#344 := [monotonicity #341]: #343
+#329 := (= #260 #328)
+#326 := (= #255 #325)
+#323 := (= #249 #322)
+#320 := (iff #246 #319)
+#321 := [rewrite]: #320
+#324 := [monotonicity #321]: #323
+#327 := [monotonicity #324]: #326
+#330 := [monotonicity #327]: #329
+#315 := (iff #239 #310)
+#314 := [rewrite]: #315
+#347 := [monotonicity #314 #330 #344]: #346
+#350 := [monotonicity #347]: #349
+#308 := (iff #236 #309)
+#311 := [rewrite]: #308
+#372 := [monotonicity #311 #350]: #371
+#368 := (iff #227 #367)
+#365 := (iff #49 #364)
+#362 := (iff #48 #361)
+#363 := [rewrite]: #362
+#359 := (iff #46 #358)
+#360 := [rewrite]: #359
+#366 := [monotonicity #360 #363]: #365
+#356 := (iff #219 #222)
+#223 := (not #222)
+#351 := (not #223)
+#354 := (iff #351 #222)
+#355 := [rewrite]: #354
+#352 := (iff #219 #351)
+#224 := (iff #42 #223)
+#225 := [rewrite]: #224
+#353 := [monotonicity #225]: #352
+#357 := [trans #353 #355]: #356
+#369 := [monotonicity #357 #366]: #368
+#375 := [monotonicity #369 #372]: #374
+#380 := [trans #375 #378]: #379
+#306 := (iff #69 #305)
+#303 := (iff #68 #302)
+#300 := (iff #67 #297)
+#294 := (= #291 #45)
+#298 := (iff #294 #297)
+#299 := [rewrite]: #298
+#295 := (iff #67 #294)
+#292 := (= #66 #291)
+#289 := (= #65 #286)
+#276 := (* 1/2::real #270)
+#281 := (- #45 #276)
+#287 := (= #281 #286)
+#288 := [rewrite]: #287
+#282 := (= #65 #281)
+#279 := (= #64 #276)
+#273 := (/ #270 2::real)
+#277 := (= #273 #276)
+#278 := [rewrite]: #277
+#274 := (= #64 #273)
+#271 := (= #63 #270)
+#265 := (= #61 #264)
+#266 := [rewrite]: #265
+#268 := (iff #62 #267)
+#269 := [monotonicity #266]: #268
+#272 := [monotonicity #269 #266]: #271
+#275 := [monotonicity #272]: #274
+#280 := [trans #275 #278]: #279
+#283 := [monotonicity #280]: #282
+#290 := [trans #283 #288]: #289
+#261 := (= #60 #260)
+#258 := (= #59 #255)
+#252 := (/ #249 2::real)
+#256 := (= #252 #255)
+#257 := [rewrite]: #256
+#253 := (= #59 #252)
+#250 := (= #58 #249)
+#244 := (= #56 #243)
+#245 := [rewrite]: #244
+#247 := (iff #57 #246)
+#248 := [monotonicity #245]: #247
+#251 := [monotonicity #248 #245]: #250
+#254 := [monotonicity #251]: #253
+#259 := [trans #254 #257]: #258
+#262 := [monotonicity #259]: #261
+#240 := (iff #55 #239)
+#234 := (= #53 #233)
+#235 := [rewrite]: #234
+#241 := [monotonicity #235]: #240
+#293 := [monotonicity #241 #262 #290]: #292
+#296 := [monotonicity #293]: #295
+#301 := [trans #296 #299]: #300
+#237 := (iff #54 #236)
+#238 := [monotonicity #235]: #237
+#304 := [monotonicity #238 #301]: #303
+#228 := (iff #50 #227)
+#229 := [rewrite]: #228
+#307 := [monotonicity #229 #304]: #306
+#382 := [trans #307 #380]: #381
+#218 := [asserted]: #69
+#383 := [mp #218 #382]: #376
+#384 := [and-elim #383]: #309
+#403 := [iff-false #384]: #402
+#406 := [monotonicity #403]: #405
+#410 := [trans #406 #408]: #409
+#413 := [monotonicity #410]: #412
+#419 := [trans #413 #417]: #418
+#385 := [and-elim #383]: #348
+#420 := [mp #385 #419]: #415
+#670 := (not #415)
+#671 := (or #670 #531)
+#672 := [th-lemma]: #671
+#673 := [unit-resolution #672 #420]: #531
+#674 := [hypothesis]: #668
+#174 := (<= f5 0::real)
+#175 := (not #174)
+#27 := (< 0::real f5)
+#176 := (iff #27 #175)
+#177 := [rewrite]: #176
+#171 := [asserted]: #27
+#178 := [mp #171 #177]: #175
+#675 := [th-lemma #178 #674 #673]: false
+#677 := [lemma #675]: #676
+#524 := (= f5 #336)
+#525 := (= #264 #336)
+#685 := (not #525)
+#669 := (+ #264 #666)
+#678 := (<= #669 0::real)
+#682 := (not #678)
+#428 := (iff #367 #364)
+#423 := (or false #364)
+#426 := (iff #423 #364)
+#427 := [rewrite]: #426
+#424 := (iff #367 #423)
+#400 := (iff #222 false)
+#217 := [asserted]: #42
+#226 := [mp #217 #225]: #223
+#401 := [iff-false #226]: #400
+#425 := [monotonicity #401]: #424
+#429 := [trans #425 #427]: #428
+#386 := [and-elim #383]: #367
+#430 := [mp #386 #429]: #364
+#422 := [and-elim #430]: #361
+#680 := [hypothesis]: #678
+#681 := [th-lemma #680 #422 #384 #673]: false
+#683 := [lemma #681]: #682
+#684 := [hypothesis]: #525
+#686 := (or #685 #678)
+#687 := [th-lemma]: #686
+#688 := [unit-resolution #687 #684 #683]: false
+#689 := [lemma #688]: #685
+#529 := (or #333 #525)
+#530 := [def-axiom]: #529
+#690 := [unit-resolution #530 #689]: #333
+#526 := (not #333)
+#527 := (or #526 #524)
+#528 := [def-axiom]: #527
+#691 := [unit-resolution #528 #690]: #524
+#692 := (not #524)
+#693 := (or #692 #668)
+#694 := [th-lemma]: #693
+[unit-resolution #694 #691 #677]: false
+unsat
+5c9bfd1097eaec646745ccbc19c98d1ba847f0f5 308 0
+#2 := false
+#10 := 0::real
+decl f6 :: real
+#15 := f6
+#143 := (<= f6 0::real)
+#144 := (not #143)
+decl f5 :: real
+#13 := f5
+#53 := -1::real
+#203 := (* -1::real f5)
+decl f4 :: real
+#8 := f4
+#204 := (+ f4 #203)
+#202 := (>= #204 0::real)
+decl f3 :: (-> real real)
+#9 := (f3 f4)
+#54 := (* -1::real #9)
+#154 := (>= #9 0::real)
+#161 := (ite #154 #9 #54)
+#166 := (* f6 #161)
+#187 := (* f5 #166)
+#193 := (* -1::real #187)
+#66 := (* #9 f6)
+#78 := (* f4 #66)
+#72 := (* f5 f6)
+#73 := (* #9 #72)
+#96 := (* -1::real #73)
+#97 := (+ #96 #78)
+#84 := (* -1::real #78)
+#85 := (+ #73 #84)
+#172 := (>= #85 0::real)
+#179 := (ite #172 #85 #97)
+#194 := (+ #179 #193)
+#169 := (* f4 #166)
+#195 := (+ #169 #194)
+#196 := (<= #195 0::real)
+#11 := (= #9 0::real)
+#223 := (or #11 #143 #196 #202)
+#228 := (not #223)
+#18 := (- #9)
+#17 := (< #9 0::real)
+#19 := (ite #17 #18 #9)
+#20 := (* f6 #19)
+#30 := (* f5 #20)
+#22 := (* f6 #9)
+#24 := (* f4 #22)
+#23 := (* f5 #22)
+#25 := (- #23 #24)
+#27 := (- #25)
+#26 := (< #25 0::real)
+#28 := (ite #26 #27 #25)
+#21 := (* f4 #20)
+#29 := (+ #21 #28)
+#31 := (<= #29 #30)
+#16 := (< 0::real f6)
+#32 := (implies #16 #31)
+#14 := (< f4 f5)
+#33 := (implies #14 #32)
+#12 := (not #11)
+#34 := (implies #12 #33)
+#35 := (not #34)
+#231 := (iff #35 #228)
+#57 := (ite #17 #54 #9)
+#60 := (* f6 #57)
+#108 := (* f5 #60)
+#90 := (< #85 0::real)
+#102 := (ite #90 #97 #85)
+#63 := (* f4 #60)
+#105 := (+ #63 #102)
+#111 := (<= #105 #108)
+#117 := (not #16)
+#118 := (or #117 #111)
+#126 := (not #14)
+#127 := (or #126 #118)
+#135 := (or #11 #127)
+#140 := (not #135)
+#229 := (iff #140 #228)
+#226 := (iff #135 #223)
+#214 := (or #143 #196)
+#217 := (or #202 #214)
+#220 := (or #11 #217)
+#224 := (iff #220 #223)
+#225 := [rewrite]: #224
+#221 := (iff #135 #220)
+#218 := (iff #127 #217)
+#215 := (iff #118 #214)
+#199 := (iff #111 #196)
+#184 := (+ #169 #179)
+#190 := (<= #184 #187)
+#197 := (iff #190 #196)
+#198 := [rewrite]: #197
+#191 := (iff #111 #190)
+#188 := (= #108 #187)
+#167 := (= #60 #166)
+#164 := (= #57 #161)
+#155 := (not #154)
+#158 := (ite #155 #54 #9)
+#162 := (= #158 #161)
+#163 := [rewrite]: #162
+#159 := (= #57 #158)
+#156 := (iff #17 #155)
+#157 := [rewrite]: #156
+#160 := [monotonicity #157]: #159
+#165 := [trans #160 #163]: #164
+#168 := [monotonicity #165]: #167
+#189 := [monotonicity #168]: #188
+#185 := (= #105 #184)
+#182 := (= #102 #179)
+#173 := (not #172)
+#176 := (ite #173 #97 #85)
+#180 := (= #176 #179)
+#181 := [rewrite]: #180
+#177 := (= #102 #176)
+#174 := (iff #90 #173)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#183 := [trans #178 #181]: #182
+#170 := (= #63 #169)
+#171 := [monotonicity #168]: #170
+#186 := [monotonicity #171 #183]: #185
+#192 := [monotonicity #186 #189]: #191
+#200 := [trans #192 #198]: #199
+#152 := (iff #117 #143)
+#147 := (not #144)
+#150 := (iff #147 #143)
+#151 := [rewrite]: #150
+#148 := (iff #117 #147)
+#145 := (iff #16 #144)
+#146 := [rewrite]: #145
+#149 := [monotonicity #146]: #148
+#153 := [trans #149 #151]: #152
+#216 := [monotonicity #153 #200]: #215
+#212 := (iff #126 #202)
+#201 := (not #202)
+#207 := (not #201)
+#210 := (iff #207 #202)
+#211 := [rewrite]: #210
+#208 := (iff #126 #207)
+#205 := (iff #14 #201)
+#206 := [rewrite]: #205
+#209 := [monotonicity #206]: #208
+#213 := [trans #209 #211]: #212
+#219 := [monotonicity #213 #216]: #218
+#222 := [monotonicity #219]: #221
+#227 := [trans #222 #225]: #226
+#230 := [monotonicity #227]: #229
+#141 := (iff #35 #140)
+#138 := (iff #34 #135)
+#132 := (implies #12 #127)
+#136 := (iff #132 #135)
+#137 := [rewrite]: #136
+#133 := (iff #34 #132)
+#130 := (iff #33 #127)
+#123 := (implies #14 #118)
+#128 := (iff #123 #127)
+#129 := [rewrite]: #128
+#124 := (iff #33 #123)
+#121 := (iff #32 #118)
+#114 := (implies #16 #111)
+#119 := (iff #114 #118)
+#120 := [rewrite]: #119
+#115 := (iff #32 #114)
+#112 := (iff #31 #111)
+#109 := (= #30 #108)
+#61 := (= #20 #60)
+#58 := (= #19 #57)
+#55 := (= #18 #54)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#62 := [monotonicity #59]: #61
+#110 := [monotonicity #62]: #109
+#106 := (= #29 #105)
+#103 := (= #28 #102)
+#88 := (= #25 #85)
+#81 := (- #73 #78)
+#86 := (= #81 #85)
+#87 := [rewrite]: #86
+#82 := (= #25 #81)
+#79 := (= #24 #78)
+#67 := (= #22 #66)
+#68 := [rewrite]: #67
+#80 := [monotonicity #68]: #79
+#76 := (= #23 #73)
+#69 := (* f5 #66)
+#74 := (= #69 #73)
+#75 := [rewrite]: #74
+#70 := (= #23 #69)
+#71 := [monotonicity #68]: #70
+#77 := [trans #71 #75]: #76
+#83 := [monotonicity #77 #80]: #82
+#89 := [trans #83 #87]: #88
+#100 := (= #27 #97)
+#93 := (- #85)
+#98 := (= #93 #97)
+#99 := [rewrite]: #98
+#94 := (= #27 #93)
+#95 := [monotonicity #89]: #94
+#101 := [trans #95 #99]: #100
+#91 := (iff #26 #90)
+#92 := [monotonicity #89]: #91
+#104 := [monotonicity #92 #101 #89]: #103
+#64 := (= #21 #63)
+#65 := [monotonicity #62]: #64
+#107 := [monotonicity #65 #104]: #106
+#113 := [monotonicity #107 #110]: #112
+#116 := [monotonicity #113]: #115
+#122 := [trans #116 #120]: #121
+#125 := [monotonicity #122]: #124
+#131 := [trans #125 #129]: #130
+#134 := [monotonicity #131]: #133
+#139 := [trans #134 #137]: #138
+#142 := [monotonicity #139]: #141
+#232 := [trans #142 #230]: #231
+#52 := [asserted]: #35
+#233 := [mp #52 #232]: #228
+#235 := [not-or-elim #233]: #144
+#241 := (<= #9 0::real)
+#317 := (not #241)
+#251 := (= #54 #161)
+#290 := (not #251)
+#236 := (not #196)
+#237 := [not-or-elim #233]: #236
+#239 := (* -1::real #179)
+#299 := (+ #97 #239)
+#301 := (>= #299 0::real)
+#247 := (= #97 #179)
+#246 := (= #85 #179)
+#280 := (not #246)
+#250 := (= #9 #161)
+#264 := (not #250)
+#254 := (+ #85 #239)
+#256 := (>= #254 0::real)
+#279 := [hypothesis]: #246
+#281 := (or #280 #256)
+#282 := [th-lemma]: #281
+#283 := [unit-resolution #282 #279]: #256
+#255 := (<= #254 0::real)
+#284 := (or #280 #255)
+#285 := [th-lemma]: #284
+#286 := [unit-resolution #285 #279]: #255
+#273 := (not #256)
+#272 := (not #255)
+#274 := (or #264 #272 #273)
+#261 := [hypothesis]: #256
+#262 := [hypothesis]: #255
+#257 := (* -1::real #161)
+#258 := (+ #9 #257)
+#260 := (>= #258 0::real)
+#263 := [hypothesis]: #250
+#265 := (or #264 #260)
+#266 := [th-lemma]: #265
+#267 := [unit-resolution #266 #263]: #260
+#259 := (<= #258 0::real)
+#268 := (or #264 #259)
+#269 := [th-lemma]: #268
+#270 := [unit-resolution #269 #263]: #259
+#271 := [th-lemma #270 #267 #262 #261 #237]: false
+#275 := [lemma #271]: #274
+#287 := [unit-resolution #275 #286 #283]: #264
+#252 := (or #155 #250)
+#253 := [def-axiom]: #252
+#288 := [unit-resolution #253 #287]: #155
+#238 := [not-or-elim #233]: #201
+#276 := (+ #54 #257)
+#278 := (>= #276 0::real)
+#248 := (or #154 #251)
+#249 := [def-axiom]: #248
+#289 := [unit-resolution #249 #288]: #251
+#291 := (or #290 #278)
+#292 := [th-lemma]: #291
+#293 := [unit-resolution #292 #289]: #278
+#277 := (<= #276 0::real)
+#294 := (or #290 #277)
+#295 := [th-lemma]: #294
+#296 := [unit-resolution #295 #289]: #277
+#297 := [th-lemma #296 #293 #286 #283 #237 #238 #288 #235]: false
+#298 := [lemma #297]: #280
+#244 := (or #173 #246)
+#245 := [def-axiom]: #244
+#302 := [unit-resolution #245 #298]: #173
+#242 := (or #172 #247)
+#243 := [def-axiom]: #242
+#303 := [unit-resolution #243 #302]: #247
+#304 := (not #247)
+#305 := (or #304 #301)
+#306 := [th-lemma]: #305
+#307 := [unit-resolution #306 #303]: #301
+#300 := (<= #299 0::real)
+#308 := (or #304 #300)
+#309 := [th-lemma]: #308
+#310 := [unit-resolution #309 #303]: #300
+#311 := [hypothesis]: #251
+#312 := [unit-resolution #292 #311]: #278
+#313 := [unit-resolution #295 #311]: #277
+#314 := [th-lemma #313 #312 #310 #307 #237]: false
+#315 := [lemma #314]: #290
+#316 := [unit-resolution #249 #315]: #154
+#320 := (or #317 #155)
+#234 := [not-or-elim #233]: #12
+#318 := (or #11 #317 #155)
+#319 := [th-lemma]: #318
+#321 := [unit-resolution #319 #234]: #320
+#322 := [unit-resolution #321 #316]: #317
+#323 := [unit-resolution #253 #316]: #250
+#324 := [unit-resolution #266 #323]: #260
+#325 := [unit-resolution #269 #323]: #259
+[th-lemma #325 #324 #310 #307 #237 #238 #322 #235]: false
+unsat
+9171919cd704be8e7c103a2053e7cd7353c1a910 19 0
+f1 -> val!0
+f2 -> val!1
+f3 -> 0
+f4 -> 1
+f5 -> 1
+f6 -> {
+  0 -> 3
+  else -> 3
+}
+f7 -> {
+  3 -> 1
+  8 -> 1
+  else -> 1
+}
+f8 -> {
+  1236 3 -> 8
+  else -> 8
+}
+unknown
+2890cc0ea13f6a8b550bc60d1dd56915e31e1ba5 386 0
+#2 := false
+#11 := 0::real
+decl f5 :: real
+#12 := f5
+#62 := (<= f5 0::real)
+#63 := (not #62)
+#13 := (< 0::real f5)
+#64 := (iff #13 #63)
+#65 := [rewrite]: #64
+#51 := [asserted]: #13
+#66 := [mp #51 #65]: #63
+decl f4 :: real
+#9 := f4
+#53 := -1::real
+#56 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#57 := (+ f3 #56)
+#55 := (>= #57 0::real)
+#54 := (not #55)
+#10 := (< f3 f4)
+#58 := (iff #10 #54)
+#59 := [rewrite]: #58
+#50 := [asserted]: #10
+#60 := [mp #50 #59]: #54
+decl f6 :: (-> real real)
+#14 := (f6 f3)
+#99 := (* -1::real #14)
+#152 := (>= #14 0::real)
+#159 := (ite #152 #14 #99)
+#271 := (<= #159 0::real)
+#336 := (not #271)
+#252 := (<= #14 0::real)
+#397 := (not #252)
+#153 := (not #152)
+#384 := [hypothesis]: #153
+#267 := (* -1::real #159)
+#305 := (+ #99 #267)
+#306 := (<= #305 0::real)
+#241 := (= #99 #159)
+#244 := (or #152 #241)
+#245 := [def-axiom]: #244
+#385 := [unit-resolution #245 #384]: #241
+#314 := (not #241)
+#318 := (or #314 #306)
+#319 := [th-lemma]: #318
+#386 := [unit-resolution #319 #385]: #306
+#337 := (not #306)
+#338 := (or #336 #152 #337)
+#339 := [th-lemma]: #338
+#387 := [unit-resolution #339 #386 #384]: #336
+#81 := (* f4 f5)
+#90 := 1/8::real
+#93 := (* 1/8::real #81)
+#79 := (* f3 f5)
+#91 := -1/8::real
+#92 := (* -1/8::real #79)
+#94 := (+ #92 #93)
+#164 := (/ #94 #159)
+#170 := (<= #164 0::real)
+#176 := (* #14 #164)
+#188 := (* -1::real #176)
+#183 := (>= #176 0::real)
+#194 := (ite #183 #176 #188)
+#203 := (* -1/8::real #81)
+#204 := (+ #203 #194)
+#202 := (* 1/8::real #79)
+#205 := (+ #202 #204)
+#206 := (<= #205 0::real)
+#382 := (or #206 #314)
+#237 := (not #206)
+#273 := [hypothesis]: #237
+#263 := (* -1::real #194)
+#348 := (+ #188 #263)
+#350 := (>= #348 0::real)
+#247 := (= #188 #194)
+#182 := (not #183)
+#246 := (= #176 #194)
+#325 := (not #246)
+#346 := (or #325 #206)
+#226 := (= #14 #159)
+#293 := (not #226)
+#264 := (+ #176 #263)
+#266 := (>= #264 0::real)
+#324 := [hypothesis]: #246
+#326 := (or #325 #266)
+#327 := [th-lemma]: #326
+#328 := [unit-resolution #327 #324]: #266
+#265 := (<= #264 0::real)
+#329 := (or #325 #265)
+#330 := [th-lemma]: #329
+#331 := [unit-resolution #330 #324]: #265
+#302 := (not #266)
+#301 := (not #265)
+#303 := (or #293 #301 #302 #206)
+#254 := (* #159 #164)
+#258 := (+ #203 #254)
+#259 := (+ #202 #258)
+#257 := (>= #259 0::real)
+#260 := (= #259 0::real)
+#253 := (= #159 0::real)
+#277 := (not #253)
+#15 := (= #14 0::real)
+#16 := (not #15)
+#278 := (iff #16 #277)
+#275 := (iff #15 #253)
+#274 := [hypothesis]: #226
+#276 := [monotonicity #274]: #275
+#279 := [monotonicity #276]: #278
+#171 := (not #170)
+#211 := (and #171 #206)
+#214 := (or #15 #211)
+#217 := (not #214)
+#19 := 8::real
+#17 := (- f4 f3)
+#18 := (* f5 #17)
+#20 := (/ #18 8::real)
+#22 := (- #14)
+#21 := (< #14 0::real)
+#23 := (ite #21 #22 #14)
+#24 := (/ #20 #23)
+#26 := (* #24 #14)
+#28 := (- #26)
+#27 := (< #26 0::real)
+#29 := (ite #27 #28 #26)
+#30 := (<= #29 #20)
+#25 := (< 0::real #24)
+#31 := (and #25 #30)
+#32 := (implies #16 #31)
+#33 := (not #32)
+#220 := (iff #33 #217)
+#102 := (ite #21 #99 #14)
+#105 := (/ #94 #102)
+#114 := (* #14 #105)
+#125 := (* -1::real #114)
+#119 := (< #114 0::real)
+#130 := (ite #119 #125 #114)
+#133 := (<= #130 #94)
+#108 := (< 0::real #105)
+#136 := (and #108 #133)
+#61 := (= 0::real #14)
+#142 := (or #61 #136)
+#147 := (not #142)
+#218 := (iff #147 #217)
+#215 := (iff #142 #214)
+#212 := (iff #136 #211)
+#209 := (iff #133 #206)
+#199 := (<= #194 #94)
+#207 := (iff #199 #206)
+#208 := [rewrite]: #207
+#200 := (iff #133 #199)
+#197 := (= #130 #194)
+#191 := (ite #182 #188 #176)
+#195 := (= #191 #194)
+#196 := [rewrite]: #195
+#192 := (= #130 #191)
+#177 := (= #114 #176)
+#165 := (= #105 #164)
+#162 := (= #102 #159)
+#156 := (ite #153 #99 #14)
+#160 := (= #156 #159)
+#161 := [rewrite]: #160
+#157 := (= #102 #156)
+#154 := (iff #21 #153)
+#155 := [rewrite]: #154
+#158 := [monotonicity #155]: #157
+#163 := [trans #158 #161]: #162
+#166 := [monotonicity #163]: #165
+#178 := [monotonicity #166]: #177
+#189 := (= #125 #188)
+#190 := [monotonicity #178]: #189
+#186 := (iff #119 #182)
+#179 := (< #176 0::real)
+#184 := (iff #179 #182)
+#185 := [rewrite]: #184
+#180 := (iff #119 #179)
+#181 := [monotonicity #178]: #180
+#187 := [trans #181 #185]: #186
+#193 := [monotonicity #187 #190 #178]: #192
+#198 := [trans #193 #196]: #197
+#201 := [monotonicity #198]: #200
+#210 := [trans #201 #208]: #209
+#174 := (iff #108 #171)
+#167 := (< 0::real #164)
+#172 := (iff #167 #171)
+#173 := [rewrite]: #172
+#168 := (iff #108 #167)
+#169 := [monotonicity #166]: #168
+#175 := [trans #169 #173]: #174
+#213 := [monotonicity #175 #210]: #212
+#150 := (iff #61 #15)
+#151 := [rewrite]: #150
+#216 := [monotonicity #151 #213]: #215
+#219 := [monotonicity #216]: #218
+#148 := (iff #33 #147)
+#145 := (iff #32 #142)
+#69 := (not #61)
+#139 := (implies #69 #136)
+#143 := (iff #139 #142)
+#144 := [rewrite]: #143
+#140 := (iff #32 #139)
+#137 := (iff #31 #136)
+#134 := (iff #30 #133)
+#97 := (= #20 #94)
+#80 := (* -1::real #79)
+#82 := (+ #80 #81)
+#87 := (/ #82 8::real)
+#95 := (= #87 #94)
+#96 := [rewrite]: #95
+#88 := (= #20 #87)
+#85 := (= #18 #82)
+#72 := (* -1::real f3)
+#73 := (+ #72 f4)
+#76 := (* f5 #73)
+#83 := (= #76 #82)
+#84 := [rewrite]: #83
+#77 := (= #18 #76)
+#74 := (= #17 #73)
+#75 := [rewrite]: #74
+#78 := [monotonicity #75]: #77
+#86 := [trans #78 #84]: #85
+#89 := [monotonicity #86]: #88
+#98 := [trans #89 #96]: #97
+#131 := (= #29 #130)
+#117 := (= #26 #114)
+#111 := (* #105 #14)
+#115 := (= #111 #114)
+#116 := [rewrite]: #115
+#112 := (= #26 #111)
+#106 := (= #24 #105)
+#103 := (= #23 #102)
+#100 := (= #22 #99)
+#101 := [rewrite]: #100
+#104 := [monotonicity #101]: #103
+#107 := [monotonicity #98 #104]: #106
+#113 := [monotonicity #107]: #112
+#118 := [trans #113 #116]: #117
+#128 := (= #28 #125)
+#122 := (- #114)
+#126 := (= #122 #125)
+#127 := [rewrite]: #126
+#123 := (= #28 #122)
+#124 := [monotonicity #118]: #123
+#129 := [trans #124 #127]: #128
+#120 := (iff #27 #119)
+#121 := [monotonicity #118]: #120
+#132 := [monotonicity #121 #129 #118]: #131
+#135 := [monotonicity #132 #98]: #134
+#109 := (iff #25 #108)
+#110 := [monotonicity #107]: #109
+#138 := [monotonicity #110 #135]: #137
+#70 := (iff #16 #69)
+#67 := (iff #15 #61)
+#68 := [rewrite]: #67
+#71 := [monotonicity #68]: #70
+#141 := [monotonicity #71 #138]: #140
+#146 := [trans #141 #144]: #145
+#149 := [monotonicity #146]: #148
+#221 := [trans #149 #219]: #220
+#52 := [asserted]: #33
+#222 := [mp #52 #221]: #217
+#223 := [not-or-elim #222]: #16
+#280 := [mp #223 #279]: #277
+#281 := (or #253 #260)
+#282 := [th-lemma]: #281
+#283 := [unit-resolution #282 #280]: #260
+#284 := (not #260)
+#285 := (or #284 #257)
+#286 := [th-lemma]: #285
+#287 := [unit-resolution #286 #283]: #257
+#256 := (<= #259 0::real)
+#288 := (or #284 #256)
+#289 := [th-lemma]: #288
+#290 := [unit-resolution #289 #283]: #256
+#291 := [hypothesis]: #266
+#292 := [hypothesis]: #265
+#268 := (+ #14 #267)
+#270 := (>= #268 0::real)
+#294 := (or #293 #270)
+#295 := [th-lemma]: #294
+#296 := [unit-resolution #295 #274]: #270
+#269 := (<= #268 0::real)
+#297 := (or #293 #269)
+#298 := [th-lemma]: #297
+#299 := [unit-resolution #298 #274]: #269
+#300 := [th-lemma #299 #296 #292 #291 #290 #287 #273]: false
+#304 := [lemma #300]: #303
+#332 := [unit-resolution #304 #331 #328 #273]: #293
+#242 := (or #153 #226)
+#243 := [def-axiom]: #242
+#333 := [unit-resolution #243 #332]: #153
+#334 := [unit-resolution #245 #333]: #241
+#335 := [unit-resolution #319 #334]: #306
+#340 := [unit-resolution #339 #333 #335]: #336
+#322 := (or #284 #301 #302 #206)
+#308 := [hypothesis]: #260
+#309 := [unit-resolution #286 #308]: #257
+#310 := [unit-resolution #289 #308]: #256
+#307 := (>= #305 0::real)
+#311 := [unit-resolution #304 #292 #291 #273]: #293
+#312 := [unit-resolution #243 #311]: #153
+#313 := [unit-resolution #245 #312]: #241
+#315 := (or #314 #307)
+#316 := [th-lemma]: #315
+#317 := [unit-resolution #316 #313]: #307
+#320 := [unit-resolution #319 #313]: #306
+#321 := [th-lemma #320 #317 #292 #291 #310 #309 #273 #60 #66]: false
+#323 := [lemma #321]: #322
+#341 := [unit-resolution #323 #331 #328 #273]: #284
+#342 := [unit-resolution #282 #341]: #253
+#343 := (or #277 #271)
+#344 := [th-lemma]: #343
+#345 := [unit-resolution #344 #342 #340]: false
+#347 := [lemma #345]: #346
+#365 := [unit-resolution #347 #273]: #325
+#248 := (or #182 #246)
+#249 := [def-axiom]: #248
+#366 := [unit-resolution #249 #365]: #182
+#250 := (or #183 #247)
+#251 := [def-axiom]: #250
+#367 := [unit-resolution #251 #366]: #247
+#368 := (not #247)
+#369 := (or #368 #350)
+#370 := [th-lemma]: #369
+#371 := [unit-resolution #370 #367]: #350
+#349 := (<= #348 0::real)
+#372 := (or #368 #349)
+#373 := [th-lemma]: #372
+#374 := [unit-resolution #373 #367]: #349
+#351 := [hypothesis]: #253
+#352 := [unit-resolution #344 #351]: #271
+#357 := (iff #16 #293)
+#355 := (iff #15 #226)
+#353 := (iff #226 #15)
+#354 := [monotonicity #351]: #353
+#356 := [symm #354]: #355
+#358 := [monotonicity #356]: #357
+#359 := [mp #223 #358]: #293
+#360 := [unit-resolution #243 #359]: #153
+#361 := [unit-resolution #339 #360 #352]: #337
+#362 := [unit-resolution #245 #360]: #241
+#363 := [unit-resolution #319 #362 #361]: false
+#364 := [lemma #363]: #277
+#375 := [unit-resolution #282 #364]: #260
+#376 := [unit-resolution #286 #375]: #257
+#377 := [unit-resolution #289 #375]: #256
+#378 := [hypothesis]: #241
+#379 := [unit-resolution #316 #378]: #307
+#380 := [unit-resolution #319 #378]: #306
+#381 := [th-lemma #380 #379 #377 #376 #374 #371 #273]: false
+#383 := [lemma #381]: #382
+#388 := [unit-resolution #383 #385]: #206
+#238 := (or #170 #237)
+#224 := (not #211)
+#229 := (iff #224 #238)
+#239 := (not #238)
+#236 := (not #239)
+#231 := (iff #236 #238)
+#232 := [rewrite]: #231
+#233 := (iff #224 #236)
+#240 := (iff #211 #239)
+#235 := [rewrite]: #240
+#234 := [monotonicity #235]: #233
+#230 := [trans #234 #232]: #229
+#225 := [not-or-elim #222]: #224
+#228 := [mp #225 #230]: #238
+#389 := [unit-resolution #228 #388]: #170
+#390 := [th-lemma #377 #376 #389 #387 #60 #66]: false
+#391 := [lemma #390]: #152
+#400 := (or #397 #153)
+#398 := (or #15 #397 #153)
+#399 := [th-lemma]: #398
+#401 := [unit-resolution #399 #223]: #400
+#402 := [unit-resolution #401 #391]: #397
+#392 := [unit-resolution #243 #391]: #226
+#394 := [unit-resolution #298 #392]: #269
+#403 := (not #269)
+#404 := (or #336 #252 #403)
+#405 := [th-lemma]: #404
+#406 := [unit-resolution #405 #394 #402]: #336
+#393 := [unit-resolution #295 #392]: #270
+#395 := [th-lemma #394 #393 #377 #376 #374 #371 #273 #60 #66]: false
+#396 := [lemma #395]: #206
+#407 := [unit-resolution #228 #396]: #170
+[th-lemma #377 #376 #407 #406 #60 #66]: false
+unsat
+726e0e73e0888d06cfe039319b4774bfabee92ef 421 0
+#2 := false
+#10 := 0::real
+decl f6 :: real
+#15 := f6
+#150 := (<= f6 0::real)
+#151 := (not #150)
+decl f5 :: real
+#13 := f5
+#53 := -1::real
+#223 := (* -1::real f5)
+decl f4 :: real
+#8 := f4
+#224 := (+ f4 #223)
+#225 := (>= #224 0::real)
+decl f3 :: (-> real real)
+#9 := (f3 f4)
+#81 := (* -1::real #9)
+#161 := (>= #9 0::real)
+#168 := (ite #161 #9 #81)
+#63 := (* f5 f6)
+#72 := 1/8::real
+#75 := (* 1/8::real #63)
+#61 := (* f4 f6)
+#73 := -1/8::real
+#74 := (* -1/8::real #61)
+#76 := (+ #74 #75)
+#173 := (/ #76 #168)
+#185 := (* #9 #173)
+#197 := (* -1::real #185)
+#192 := (>= #185 0::real)
+#203 := (ite #192 #185 #197)
+#212 := (* -1/8::real #63)
+#213 := (+ #212 #203)
+#211 := (* 1/8::real #61)
+#214 := (+ #211 #213)
+#215 := (<= #214 0::real)
+#179 := (<= #173 0::real)
+#180 := (not #179)
+#220 := (and #180 #215)
+#11 := (= #9 0::real)
+#245 := (or #11 #150 #220 #225)
+#250 := (not #245)
+#19 := 8::real
+#17 := (- f5 f4)
+#18 := (* f6 #17)
+#20 := (/ #18 8::real)
+#22 := (- #9)
+#21 := (< #9 0::real)
+#23 := (ite #21 #22 #9)
+#24 := (/ #20 #23)
+#26 := (* #24 #9)
+#28 := (- #26)
+#27 := (< #26 0::real)
+#29 := (ite #27 #28 #26)
+#30 := (<= #29 #20)
+#25 := (< 0::real #24)
+#31 := (and #25 #30)
+#16 := (< 0::real f6)
+#32 := (implies #16 #31)
+#14 := (< f4 f5)
+#33 := (implies #14 #32)
+#12 := (not #11)
+#34 := (implies #12 #33)
+#35 := (not #34)
+#253 := (iff #35 #250)
+#84 := (ite #21 #81 #9)
+#87 := (/ #76 #84)
+#96 := (* #9 #87)
+#107 := (* -1::real #96)
+#101 := (< #96 0::real)
+#112 := (ite #101 #107 #96)
+#115 := (<= #112 #76)
+#90 := (< 0::real #87)
+#118 := (and #90 #115)
+#124 := (not #16)
+#125 := (or #124 #118)
+#133 := (not #14)
+#134 := (or #133 #125)
+#142 := (or #11 #134)
+#147 := (not #142)
+#251 := (iff #147 #250)
+#248 := (iff #142 #245)
+#236 := (or #150 #220)
+#239 := (or #225 #236)
+#242 := (or #11 #239)
+#246 := (iff #242 #245)
+#247 := [rewrite]: #246
+#243 := (iff #142 #242)
+#240 := (iff #134 #239)
+#237 := (iff #125 #236)
+#221 := (iff #118 #220)
+#218 := (iff #115 #215)
+#208 := (<= #203 #76)
+#216 := (iff #208 #215)
+#217 := [rewrite]: #216
+#209 := (iff #115 #208)
+#206 := (= #112 #203)
+#191 := (not #192)
+#200 := (ite #191 #197 #185)
+#204 := (= #200 #203)
+#205 := [rewrite]: #204
+#201 := (= #112 #200)
+#186 := (= #96 #185)
+#174 := (= #87 #173)
+#171 := (= #84 #168)
+#162 := (not #161)
+#165 := (ite #162 #81 #9)
+#169 := (= #165 #168)
+#170 := [rewrite]: #169
+#166 := (= #84 #165)
+#163 := (iff #21 #162)
+#164 := [rewrite]: #163
+#167 := [monotonicity #164]: #166
+#172 := [trans #167 #170]: #171
+#175 := [monotonicity #172]: #174
+#187 := [monotonicity #175]: #186
+#198 := (= #107 #197)
+#199 := [monotonicity #187]: #198
+#195 := (iff #101 #191)
+#188 := (< #185 0::real)
+#193 := (iff #188 #191)
+#194 := [rewrite]: #193
+#189 := (iff #101 #188)
+#190 := [monotonicity #187]: #189
+#196 := [trans #190 #194]: #195
+#202 := [monotonicity #196 #199 #187]: #201
+#207 := [trans #202 #205]: #206
+#210 := [monotonicity #207]: #209
+#219 := [trans #210 #217]: #218
+#183 := (iff #90 #180)
+#176 := (< 0::real #173)
+#181 := (iff #176 #180)
+#182 := [rewrite]: #181
+#177 := (iff #90 #176)
+#178 := [monotonicity #175]: #177
+#184 := [trans #178 #182]: #183
+#222 := [monotonicity #184 #219]: #221
+#159 := (iff #124 #150)
+#154 := (not #151)
+#157 := (iff #154 #150)
+#158 := [rewrite]: #157
+#155 := (iff #124 #154)
+#152 := (iff #16 #151)
+#153 := [rewrite]: #152
+#156 := [monotonicity #153]: #155
+#160 := [trans #156 #158]: #159
+#238 := [monotonicity #160 #222]: #237
+#234 := (iff #133 #225)
+#226 := (not #225)
+#229 := (not #226)
+#232 := (iff #229 #225)
+#233 := [rewrite]: #232
+#230 := (iff #133 #229)
+#227 := (iff #14 #226)
+#228 := [rewrite]: #227
+#231 := [monotonicity #228]: #230
+#235 := [trans #231 #233]: #234
+#241 := [monotonicity #235 #238]: #240
+#244 := [monotonicity #241]: #243
+#249 := [trans #244 #247]: #248
+#252 := [monotonicity #249]: #251
+#148 := (iff #35 #147)
+#145 := (iff #34 #142)
+#139 := (implies #12 #134)
+#143 := (iff #139 #142)
+#144 := [rewrite]: #143
+#140 := (iff #34 #139)
+#137 := (iff #33 #134)
+#130 := (implies #14 #125)
+#135 := (iff #130 #134)
+#136 := [rewrite]: #135
+#131 := (iff #33 #130)
+#128 := (iff #32 #125)
+#121 := (implies #16 #118)
+#126 := (iff #121 #125)
+#127 := [rewrite]: #126
+#122 := (iff #32 #121)
+#119 := (iff #31 #118)
+#116 := (iff #30 #115)
+#79 := (= #20 #76)
+#62 := (* -1::real #61)
+#64 := (+ #62 #63)
+#69 := (/ #64 8::real)
+#77 := (= #69 #76)
+#78 := [rewrite]: #77
+#70 := (= #20 #69)
+#67 := (= #18 #64)
+#54 := (* -1::real f4)
+#55 := (+ #54 f5)
+#58 := (* f6 #55)
+#65 := (= #58 #64)
+#66 := [rewrite]: #65
+#59 := (= #18 #58)
+#56 := (= #17 #55)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#68 := [trans #60 #66]: #67
+#71 := [monotonicity #68]: #70
+#80 := [trans #71 #78]: #79
+#113 := (= #29 #112)
+#99 := (= #26 #96)
+#93 := (* #87 #9)
+#97 := (= #93 #96)
+#98 := [rewrite]: #97
+#94 := (= #26 #93)
+#88 := (= #24 #87)
+#85 := (= #23 #84)
+#82 := (= #22 #81)
+#83 := [rewrite]: #82
+#86 := [monotonicity #83]: #85
+#89 := [monotonicity #80 #86]: #88
+#95 := [monotonicity #89]: #94
+#100 := [trans #95 #98]: #99
+#110 := (= #28 #107)
+#104 := (- #96)
+#108 := (= #104 #107)
+#109 := [rewrite]: #108
+#105 := (= #28 #104)
+#106 := [monotonicity #100]: #105
+#111 := [trans #106 #109]: #110
+#102 := (iff #27 #101)
+#103 := [monotonicity #100]: #102
+#114 := [monotonicity #103 #111 #100]: #113
+#117 := [monotonicity #114 #80]: #116
+#91 := (iff #25 #90)
+#92 := [monotonicity #89]: #91
+#120 := [monotonicity #92 #117]: #119
+#123 := [monotonicity #120]: #122
+#129 := [trans #123 #127]: #128
+#132 := [monotonicity #129]: #131
+#138 := [trans #132 #136]: #137
+#141 := [monotonicity #138]: #140
+#146 := [trans #141 #144]: #145
+#149 := [monotonicity #146]: #148
+#254 := [trans #149 #252]: #253
+#52 := [asserted]: #35
+#255 := [mp #52 #254]: #250
+#257 := [not-or-elim #255]: #151
+#260 := [not-or-elim #255]: #226
+#306 := (<= #168 0::real)
+#371 := (not #306)
+#287 := (<= #9 0::real)
+#432 := (not #287)
+#419 := [hypothesis]: #162
+#302 := (* -1::real #168)
+#340 := (+ #81 #302)
+#341 := (<= #340 0::real)
+#276 := (= #81 #168)
+#279 := (or #161 #276)
+#280 := [def-axiom]: #279
+#420 := [unit-resolution #280 #419]: #276
+#349 := (not #276)
+#353 := (or #349 #341)
+#354 := [th-lemma]: #353
+#421 := [unit-resolution #354 #420]: #341
+#372 := (not #341)
+#373 := (or #371 #161 #372)
+#374 := [th-lemma]: #373
+#422 := [unit-resolution #374 #421 #419]: #371
+#417 := (or #215 #349)
+#272 := (not #215)
+#308 := [hypothesis]: #272
+#298 := (* -1::real #203)
+#383 := (+ #197 #298)
+#385 := (>= #383 0::real)
+#282 := (= #197 #203)
+#281 := (= #185 #203)
+#360 := (not #281)
+#381 := (or #360 #215)
+#261 := (= #9 #168)
+#328 := (not #261)
+#299 := (+ #185 #298)
+#301 := (>= #299 0::real)
+#359 := [hypothesis]: #281
+#361 := (or #360 #301)
+#362 := [th-lemma]: #361
+#363 := [unit-resolution #362 #359]: #301
+#300 := (<= #299 0::real)
+#364 := (or #360 #300)
+#365 := [th-lemma]: #364
+#366 := [unit-resolution #365 #359]: #300
+#337 := (not #301)
+#336 := (not #300)
+#338 := (or #328 #336 #337 #215)
+#289 := (* #168 #173)
+#293 := (+ #212 #289)
+#294 := (+ #211 #293)
+#292 := (>= #294 0::real)
+#295 := (= #294 0::real)
+#288 := (= #168 0::real)
+#312 := (not #288)
+#313 := (iff #12 #312)
+#310 := (iff #11 #288)
+#309 := [hypothesis]: #261
+#311 := [monotonicity #309]: #310
+#314 := [monotonicity #311]: #313
+#256 := [not-or-elim #255]: #12
+#315 := [mp #256 #314]: #312
+#316 := (or #288 #295)
+#317 := [th-lemma]: #316
+#318 := [unit-resolution #317 #315]: #295
+#319 := (not #295)
+#320 := (or #319 #292)
+#321 := [th-lemma]: #320
+#322 := [unit-resolution #321 #318]: #292
+#291 := (<= #294 0::real)
+#323 := (or #319 #291)
+#324 := [th-lemma]: #323
+#325 := [unit-resolution #324 #318]: #291
+#326 := [hypothesis]: #301
+#327 := [hypothesis]: #300
+#303 := (+ #9 #302)
+#305 := (>= #303 0::real)
+#329 := (or #328 #305)
+#330 := [th-lemma]: #329
+#331 := [unit-resolution #330 #309]: #305
+#304 := (<= #303 0::real)
+#332 := (or #328 #304)
+#333 := [th-lemma]: #332
+#334 := [unit-resolution #333 #309]: #304
+#335 := [th-lemma #334 #331 #327 #326 #325 #322 #308]: false
+#339 := [lemma #335]: #338
+#367 := [unit-resolution #339 #366 #363 #308]: #328
+#277 := (or #162 #261)
+#278 := [def-axiom]: #277
+#368 := [unit-resolution #278 #367]: #162
+#369 := [unit-resolution #280 #368]: #276
+#370 := [unit-resolution #354 #369]: #341
+#375 := [unit-resolution #374 #368 #370]: #371
+#357 := (or #319 #336 #337 #215)
+#343 := [hypothesis]: #295
+#344 := [unit-resolution #321 #343]: #292
+#345 := [unit-resolution #324 #343]: #291
+#342 := (>= #340 0::real)
+#346 := [unit-resolution #339 #327 #326 #308]: #328
+#347 := [unit-resolution #278 #346]: #162
+#348 := [unit-resolution #280 #347]: #276
+#350 := (or #349 #342)
+#351 := [th-lemma]: #350
+#352 := [unit-resolution #351 #348]: #342
+#355 := [unit-resolution #354 #348]: #341
+#356 := [th-lemma #355 #352 #327 #326 #345 #344 #308 #260 #257]: false
+#358 := [lemma #356]: #357
+#376 := [unit-resolution #358 #366 #363 #308]: #319
+#377 := [unit-resolution #317 #376]: #288
+#378 := (or #312 #306)
+#379 := [th-lemma]: #378
+#380 := [unit-resolution #379 #377 #375]: false
+#382 := [lemma #380]: #381
+#400 := [unit-resolution #382 #308]: #360
+#283 := (or #191 #281)
+#284 := [def-axiom]: #283
+#401 := [unit-resolution #284 #400]: #191
+#285 := (or #192 #282)
+#286 := [def-axiom]: #285
+#402 := [unit-resolution #286 #401]: #282
+#403 := (not #282)
+#404 := (or #403 #385)
+#405 := [th-lemma]: #404
+#406 := [unit-resolution #405 #402]: #385
+#384 := (<= #383 0::real)
+#407 := (or #403 #384)
+#408 := [th-lemma]: #407
+#409 := [unit-resolution #408 #402]: #384
+#386 := [hypothesis]: #288
+#387 := [unit-resolution #379 #386]: #306
+#392 := (iff #12 #328)
+#390 := (iff #11 #261)
+#388 := (iff #261 #11)
+#389 := [monotonicity #386]: #388
+#391 := [symm #389]: #390
+#393 := [monotonicity #391]: #392
+#394 := [mp #256 #393]: #328
+#395 := [unit-resolution #278 #394]: #162
+#396 := [unit-resolution #374 #395 #387]: #372
+#397 := [unit-resolution #280 #395]: #276
+#398 := [unit-resolution #354 #397 #396]: false
+#399 := [lemma #398]: #312
+#410 := [unit-resolution #317 #399]: #295
+#411 := [unit-resolution #321 #410]: #292
+#412 := [unit-resolution #324 #410]: #291
+#413 := [hypothesis]: #276
+#414 := [unit-resolution #351 #413]: #342
+#415 := [unit-resolution #354 #413]: #341
+#416 := [th-lemma #415 #414 #412 #411 #409 #406 #308]: false
+#418 := [lemma #416]: #417
+#423 := [unit-resolution #418 #420]: #215
+#273 := (or #179 #272)
+#258 := (not #220)
+#264 := (iff #258 #273)
+#274 := (not #273)
+#271 := (not #274)
+#266 := (iff #271 #273)
+#267 := [rewrite]: #266
+#268 := (iff #258 #271)
+#275 := (iff #220 #274)
+#270 := [rewrite]: #275
+#269 := [monotonicity #270]: #268
+#265 := [trans #269 #267]: #264
+#259 := [not-or-elim #255]: #258
+#263 := [mp #259 #265]: #273
+#424 := [unit-resolution #263 #423]: #179
+#425 := [th-lemma #412 #411 #424 #422 #260 #257]: false
+#426 := [lemma #425]: #161
+#435 := (or #432 #162)
+#433 := (or #11 #432 #162)
+#434 := [th-lemma]: #433
+#436 := [unit-resolution #434 #256]: #435
+#437 := [unit-resolution #436 #426]: #432
+#427 := [unit-resolution #278 #426]: #261
+#429 := [unit-resolution #333 #427]: #304
+#438 := (not #304)
+#439 := (or #371 #287 #438)
+#440 := [th-lemma]: #439
+#441 := [unit-resolution #440 #429 #437]: #371
+#428 := [unit-resolution #330 #427]: #305
+#430 := [th-lemma #429 #428 #412 #411 #409 #406 #308 #260 #257]: false
+#431 := [lemma #430]: #215
+#442 := [unit-resolution #263 #431]: #179
+[th-lemma #412 #411 #442 #441 #260 #257]: false
+unsat
+bae4db9bec568394a074211dd1bdf1addbc19bd3 401 0
+#2 := false
+#16 := 0::real
+decl f3 :: real
+#8 := f3
+#165 := (<= f3 0::real)
+#166 := (not #165)
+decl f5 :: real
+#10 := f5
+#54 := -1::real
+#55 := (* -1::real f5)
+decl f4 :: real
+#9 := f4
+#56 := (+ f4 #55)
+#237 := (<= #56 0::real)
+decl f6 :: (-> real real)
+#15 := (f6 f5)
+#82 := (* -1::real #15)
+#176 := (>= #15 0::real)
+#183 := (ite #176 #15 #82)
+#63 := (* f3 f5)
+#75 := -1/8::real
+#76 := (* -1/8::real #63)
+#62 := (* f3 f4)
+#73 := 1/8::real
+#74 := (* 1/8::real #62)
+#77 := (+ #74 #76)
+#188 := (/ #77 #183)
+#228 := (<= #188 0::real)
+#229 := (not #228)
+#191 := (* #15 #188)
+#203 := (* -1::real #191)
+#198 := (>= #191 0::real)
+#209 := (ite #198 #191 #203)
+#221 := (* -1::real #209)
+#222 := (+ #76 #221)
+#223 := (+ #74 #222)
+#219 := (>= #223 0::real)
+#234 := (and #219 #229)
+#248 := (not #219)
+#26 := (= #15 0::real)
+#263 := (or #26 #165 #248 #234 #237)
+#268 := (not #263)
+#13 := 8::real
+#11 := (- f4 f5)
+#12 := (* f3 #11)
+#14 := (/ #12 8::real)
+#18 := (- #15)
+#17 := (< #15 0::real)
+#19 := (ite #17 #18 #15)
+#20 := (/ #14 #19)
+#21 := (* #20 #15)
+#23 := (- #21)
+#22 := (< #21 0::real)
+#24 := (ite #22 #23 #21)
+#25 := (<= #24 #14)
+#30 := (< 0::real #20)
+#31 := (and #30 #25)
+#29 := (< 0::real f3)
+#32 := (implies #29 #31)
+#28 := (< f5 f4)
+#33 := (implies #28 #32)
+#27 := (not #26)
+#34 := (implies #27 #33)
+#35 := (implies #25 #34)
+#36 := (not #35)
+#271 := (iff #36 #268)
+#85 := (ite #17 #82 #15)
+#88 := (/ #77 #85)
+#116 := (< 0::real #88)
+#94 := (* #15 #88)
+#105 := (* -1::real #94)
+#99 := (< #94 0::real)
+#110 := (ite #99 #105 #94)
+#113 := (<= #110 #77)
+#122 := (and #113 #116)
+#130 := (not #29)
+#131 := (or #130 #122)
+#139 := (not #28)
+#140 := (or #139 #131)
+#148 := (or #26 #140)
+#156 := (not #113)
+#157 := (or #156 #148)
+#162 := (not #157)
+#269 := (iff #162 #268)
+#266 := (iff #157 #263)
+#251 := (or #165 #234)
+#254 := (or #237 #251)
+#257 := (or #26 #254)
+#260 := (or #248 #257)
+#264 := (iff #260 #263)
+#265 := [rewrite]: #264
+#261 := (iff #157 #260)
+#258 := (iff #148 #257)
+#255 := (iff #140 #254)
+#252 := (iff #131 #251)
+#235 := (iff #122 #234)
+#232 := (iff #116 #229)
+#225 := (< 0::real #188)
+#230 := (iff #225 #229)
+#231 := [rewrite]: #230
+#226 := (iff #116 #225)
+#189 := (= #88 #188)
+#186 := (= #85 #183)
+#177 := (not #176)
+#180 := (ite #177 #82 #15)
+#184 := (= #180 #183)
+#185 := [rewrite]: #184
+#181 := (= #85 #180)
+#178 := (iff #17 #177)
+#179 := [rewrite]: #178
+#182 := [monotonicity #179]: #181
+#187 := [trans #182 #185]: #186
+#190 := [monotonicity #187]: #189
+#227 := [monotonicity #190]: #226
+#233 := [trans #227 #231]: #232
+#220 := (iff #113 #219)
+#214 := (<= #209 #77)
+#218 := (iff #214 #219)
+#217 := [rewrite]: #218
+#215 := (iff #113 #214)
+#212 := (= #110 #209)
+#197 := (not #198)
+#206 := (ite #197 #203 #191)
+#210 := (= #206 #209)
+#211 := [rewrite]: #210
+#207 := (= #110 #206)
+#192 := (= #94 #191)
+#193 := [monotonicity #190]: #192
+#204 := (= #105 #203)
+#205 := [monotonicity #193]: #204
+#201 := (iff #99 #197)
+#194 := (< #191 0::real)
+#199 := (iff #194 #197)
+#200 := [rewrite]: #199
+#195 := (iff #99 #194)
+#196 := [monotonicity #193]: #195
+#202 := [trans #196 #200]: #201
+#208 := [monotonicity #202 #205 #193]: #207
+#213 := [trans #208 #211]: #212
+#216 := [monotonicity #213]: #215
+#224 := [trans #216 #217]: #220
+#236 := [monotonicity #224 #233]: #235
+#174 := (iff #130 #165)
+#169 := (not #166)
+#172 := (iff #169 #165)
+#173 := [rewrite]: #172
+#170 := (iff #130 #169)
+#167 := (iff #29 #166)
+#168 := [rewrite]: #167
+#171 := [monotonicity #168]: #170
+#175 := [trans #171 #173]: #174
+#253 := [monotonicity #175 #236]: #252
+#246 := (iff #139 #237)
+#238 := (not #237)
+#241 := (not #238)
+#244 := (iff #241 #237)
+#245 := [rewrite]: #244
+#242 := (iff #139 #241)
+#239 := (iff #28 #238)
+#240 := [rewrite]: #239
+#243 := [monotonicity #240]: #242
+#247 := [trans #243 #245]: #246
+#256 := [monotonicity #247 #253]: #255
+#259 := [monotonicity #256]: #258
+#249 := (iff #156 #248)
+#250 := [monotonicity #224]: #249
+#262 := [monotonicity #250 #259]: #261
+#267 := [trans #262 #265]: #266
+#270 := [monotonicity #267]: #269
+#163 := (iff #36 #162)
+#160 := (iff #35 #157)
+#153 := (implies #113 #148)
+#158 := (iff #153 #157)
+#159 := [rewrite]: #158
+#154 := (iff #35 #153)
+#151 := (iff #34 #148)
+#145 := (implies #27 #140)
+#149 := (iff #145 #148)
+#150 := [rewrite]: #149
+#146 := (iff #34 #145)
+#143 := (iff #33 #140)
+#136 := (implies #28 #131)
+#141 := (iff #136 #140)
+#142 := [rewrite]: #141
+#137 := (iff #33 #136)
+#134 := (iff #32 #131)
+#127 := (implies #29 #122)
+#132 := (iff #127 #131)
+#133 := [rewrite]: #132
+#128 := (iff #32 #127)
+#125 := (iff #31 #122)
+#119 := (and #116 #113)
+#123 := (iff #119 #122)
+#124 := [rewrite]: #123
+#120 := (iff #31 #119)
+#114 := (iff #25 #113)
+#80 := (= #14 #77)
+#64 := (* -1::real #63)
+#65 := (+ #62 #64)
+#70 := (/ #65 8::real)
+#78 := (= #70 #77)
+#79 := [rewrite]: #78
+#71 := (= #14 #70)
+#68 := (= #12 #65)
+#59 := (* f3 #56)
+#66 := (= #59 #65)
+#67 := [rewrite]: #66
+#60 := (= #12 #59)
+#57 := (= #11 #56)
+#58 := [rewrite]: #57
+#61 := [monotonicity #58]: #60
+#69 := [trans #61 #67]: #68
+#72 := [monotonicity #69]: #71
+#81 := [trans #72 #79]: #80
+#111 := (= #24 #110)
+#97 := (= #21 #94)
+#91 := (* #88 #15)
+#95 := (= #91 #94)
+#96 := [rewrite]: #95
+#92 := (= #21 #91)
+#89 := (= #20 #88)
+#86 := (= #19 #85)
+#83 := (= #18 #82)
+#84 := [rewrite]: #83
+#87 := [monotonicity #84]: #86
+#90 := [monotonicity #81 #87]: #89
+#93 := [monotonicity #90]: #92
+#98 := [trans #93 #96]: #97
+#108 := (= #23 #105)
+#102 := (- #94)
+#106 := (= #102 #105)
+#107 := [rewrite]: #106
+#103 := (= #23 #102)
+#104 := [monotonicity #98]: #103
+#109 := [trans #104 #107]: #108
+#100 := (iff #22 #99)
+#101 := [monotonicity #98]: #100
+#112 := [monotonicity #101 #109 #98]: #111
+#115 := [monotonicity #112 #81]: #114
+#117 := (iff #30 #116)
+#118 := [monotonicity #90]: #117
+#121 := [monotonicity #118 #115]: #120
+#126 := [trans #121 #124]: #125
+#129 := [monotonicity #126]: #128
+#135 := [trans #129 #133]: #134
+#138 := [monotonicity #135]: #137
+#144 := [trans #138 #142]: #143
+#147 := [monotonicity #144]: #146
+#152 := [trans #147 #150]: #151
+#155 := [monotonicity #115 #152]: #154
+#161 := [trans #155 #159]: #160
+#164 := [monotonicity #161]: #163
+#272 := [trans #164 #270]: #271
+#53 := [asserted]: #36
+#273 := [mp #53 #272]: #268
+#275 := [not-or-elim #273]: #166
+#279 := [not-or-elim #273]: #238
+#322 := (<= #183 0::real)
+#370 := (not #322)
+#318 := (<= #15 0::real)
+#404 := (not #318)
+#276 := [not-or-elim #273]: #219
+#352 := (+ #191 #221)
+#353 := (<= #352 0::real)
+#306 := (= #191 #209)
+#386 := [hypothesis]: #177
+#389 := (or #198 #176)
+#277 := (not #234)
+#301 := (iff #277 #228)
+#296 := (not #229)
+#299 := (iff #296 #228)
+#300 := [rewrite]: #299
+#297 := (iff #277 #296)
+#294 := (iff #234 #229)
+#1 := true
+#289 := (and true #229)
+#292 := (iff #289 #229)
+#293 := [rewrite]: #292
+#290 := (iff #234 #289)
+#287 := (iff #219 true)
+#288 := [iff-true #276]: #287
+#291 := [monotonicity #288]: #290
+#295 := [trans #291 #293]: #294
+#298 := [monotonicity #295]: #297
+#302 := [trans #298 #300]: #301
+#278 := [not-or-elim #273]: #277
+#303 := [mp #278 #302]: #228
+#387 := [hypothesis]: #197
+#388 := [th-lemma #387 #303 #386]: false
+#390 := [lemma #388]: #389
+#376 := [unit-resolution #390 #386]: #198
+#282 := (or #197 #306)
+#280 := [def-axiom]: #282
+#377 := [unit-resolution #280 #376]: #306
+#366 := (not #306)
+#367 := (or #366 #353)
+#368 := [th-lemma]: #367
+#391 := [unit-resolution #368 #377]: #353
+#324 := (* -1::real #183)
+#325 := (+ #82 #324)
+#326 := (<= #325 0::real)
+#305 := (= #82 #183)
+#283 := (or #176 #305)
+#284 := [def-axiom]: #283
+#392 := [unit-resolution #284 #386]: #305
+#342 := (not #305)
+#343 := (or #342 #326)
+#344 := [th-lemma]: #343
+#393 := [unit-resolution #344 #392]: #326
+#394 := (not #326)
+#395 := (or #370 #176 #394)
+#396 := [th-lemma]: #395
+#397 := [unit-resolution #396 #386 #393]: #370
+#311 := (* #183 #188)
+#319 := (* -1::real #311)
+#320 := (+ #76 #319)
+#321 := (+ #74 #320)
+#314 := (>= #321 0::real)
+#317 := (= #321 0::real)
+#310 := (= #183 0::real)
+#346 := (not #310)
+#304 := (= #15 #183)
+#336 := (not #304)
+#337 := (iff #27 #336)
+#334 := (iff #26 #304)
+#332 := (iff #304 #26)
+#331 := [hypothesis]: #310
+#333 := [monotonicity #331]: #332
+#335 := [symm #333]: #334
+#338 := [monotonicity #335]: #337
+#274 := [not-or-elim #273]: #27
+#339 := [mp #274 #338]: #336
+#285 := (or #177 #304)
+#286 := [def-axiom]: #285
+#340 := [unit-resolution #286 #339]: #177
+#341 := [unit-resolution #284 #340]: #305
+#345 := [unit-resolution #344 #341]: #326
+#347 := (or #346 #322)
+#348 := [th-lemma]: #347
+#349 := [unit-resolution #348 #331]: #322
+#350 := [th-lemma #340 #349 #345]: false
+#351 := [lemma #350]: #346
+#357 := (or #310 #317)
+#358 := [th-lemma]: #357
+#359 := [unit-resolution #358 #351]: #317
+#360 := (not #317)
+#398 := (or #360 #314)
+#399 := [th-lemma]: #398
+#400 := [unit-resolution #399 #359]: #314
+#313 := (<= #321 0::real)
+#361 := (or #360 #313)
+#362 := [th-lemma]: #361
+#363 := [unit-resolution #362 #359]: #313
+#401 := [th-lemma #363 #400 #303 #397 #376 #391 #276 #279 #275]: false
+#402 := [lemma #401]: #176
+#407 := (or #404 #177)
+#405 := (or #26 #404 #177)
+#406 := [th-lemma]: #405
+#408 := [unit-resolution #406 #274]: #407
+#409 := [unit-resolution #408 #402]: #404
+#328 := (+ #15 #324)
+#329 := (<= #328 0::real)
+#410 := [unit-resolution #286 #402]: #304
+#411 := (or #336 #329)
+#412 := [th-lemma]: #411
+#413 := [unit-resolution #412 #410]: #329
+#414 := (not #329)
+#415 := (or #370 #318 #414)
+#416 := [th-lemma]: #415
+#417 := [unit-resolution #416 #413 #409]: #370
+#327 := (+ #203 #221)
+#354 := (<= #327 0::real)
+#307 := (= #203 #209)
+#384 := (or #307 #322)
+#355 := (not #307)
+#356 := [hypothesis]: #355
+#308 := (or #198 #307)
+#309 := [def-axiom]: #308
+#364 := [unit-resolution #309 #356]: #198
+#365 := [unit-resolution #280 #364]: #306
+#380 := (= #203 #191)
+#378 := (= 0::real #191)
+#372 := (= #191 0::real)
+#369 := [unit-resolution #368 #365]: #353
+#371 := [hypothesis]: #370
+#373 := [th-lemma #364 #371 #364 #369 #276 #363 #303]: #372
+#379 := [symm #373]: #378
+#374 := (= #203 0::real)
+#375 := [th-lemma #364 #371 #364 #369 #276 #363 #303]: #374
+#381 := [trans #375 #379]: #380
+#382 := [trans #381 #365]: #307
+#383 := [unit-resolution #356 #382]: false
+#385 := [lemma #383]: #384
+#418 := [unit-resolution #385 #417]: #307
+#419 := (or #355 #354)
+#420 := [th-lemma]: #419
+#421 := [unit-resolution #420 #418]: #354
+#422 := [th-lemma #387 #421 #276 #363 #303 #417]: false
+#423 := [lemma #422]: #198
+[th-lemma #363 #400 #303 #409 #423 #279 #275]: false
+unsat
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -1,5 +1,5 @@
 
-header {* Kurzweil-Henstock gauge integration in many dimensions. *}
+header {* Kurzweil-Henstock Gauge Integration in many dimensions. *}
 (*  Author:                     John Harrison
     Translation from HOL light: Robert Himmelmann, TU Muenchen *)
 
@@ -8,20 +8,30 @@
 begin
 
 declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration.certs"]]
-declare [[smt_fixed=true]]
+declare [[smt_fixed=false]]
 declare [[z3_proofs=true]]
 
 setup {* Arith_Data.add_tactic "Ferrante-Rackoff" (K FerranteRackoff.dlo_tac) *}
 
-
+(*declare not_less[simp] not_le[simp]*)
+
+lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
+  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
+  scaleR_cancel_left scaleR_cancel_right scaleR.add_right scaleR.add_left real_vector_class.scaleR_one
+
+lemma real_arch_invD:
+  "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
+  by(subst(asm) real_arch_inv)
 subsection {* Sundries *}
 
+(*declare basis_component[simp]*)
+
 lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
 lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto
 lemma conjunctD4: assumes "a \<and> b \<and> c \<and> d" shows a b c d using assms by auto
 lemma conjunctD5: assumes "a \<and> b \<and> c \<and> d \<and> e" shows a b c d e using assms by auto
 
-declare smult_conv_scaleR[simp]
+declare norm_triangle_ineq4[intro] 
 
 lemma simple_image: "{f x |x . x \<in> s} = f ` s" by blast
 
@@ -47,20 +57,10 @@
   using assms apply-apply(erule exE) apply(rule_tac x=b in exI)
   unfolding isUb_def setle_def by auto
 
-lemma dist_trans[simp]:"dist (vec1 x) (vec1 y) = dist x (y::real)"
-  unfolding dist_real_def dist_vec1 ..
-
-lemma Lim_trans[simp]: fixes f::"'a \<Rightarrow> real"
-  shows "((\<lambda>x. vec1 (f x)) ---> vec1 l) net \<longleftrightarrow> (f ---> l) net"
-  using assms unfolding Lim dist_trans ..
-
-lemma bounded_linear_component[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)"
+lemma bounded_linear_component[intro]: "bounded_linear (\<lambda>x::'a::euclidean_space. x $$ k)"
   apply(rule bounded_linearI[where K=1]) 
   using component_le_norm[of _ k] unfolding real_norm_def by auto
 
-lemma bounded_vec1[intro]:  "bounded s \<Longrightarrow> bounded (vec1 ` (s::real set))"
-  unfolding bounded_def apply safe apply(rule_tac x="vec1 x" in exI,rule_tac x=e in exI) by auto
-
 lemma transitive_stepwise_lt_eq:
   assumes "(\<And>x y z::nat. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z)"
   shows "((\<forall>m. \<forall>n>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n)))" (is "?l = ?r")
@@ -97,15 +97,16 @@
     apply(rule assms) apply(rule assms,assumption,assumption) using assms(3) by auto
   thus ?thesis by auto qed
 
-
 subsection {* Some useful lemmas about intervals. *}
 
-lemma empty_as_interval: "{} = {1..0::real^'n}"
-  apply(rule set_ext,rule) defer unfolding vector_le_def mem_interval
+abbreviation One  where "One \<equiv> ((\<chi>\<chi> i. 1)::_::ordered_euclidean_space)"
+
+lemma empty_as_interval: "{} = {One..0}"
+  apply(rule set_ext,rule) defer unfolding mem_interval
   using UNIV_witness[where 'a='n] apply(erule_tac exE,rule_tac x=x in allE) by auto
 
 lemma interior_subset_union_intervals: 
-  assumes "i = {a..b::real^'n}" "j = {c..d}" "interior j \<noteq> {}" "i \<subseteq> j \<union> s" "interior(i) \<inter> interior(j) = {}"
+  assumes "i = {a..b::'a::ordered_euclidean_space}" "j = {c..d}" "interior j \<noteq> {}" "i \<subseteq> j \<union> s" "interior(i) \<inter> interior(j) = {}"
   shows "interior i \<subseteq> interior s" proof-
   have "{a<..<b} \<inter> {c..d} = {}" using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
     unfolding assms(1,2) interior_closed_interval by auto
@@ -114,7 +115,7 @@
   ultimately show ?thesis apply-apply(rule interior_maximal) defer apply(rule open_interior)
     unfolding assms(1,2) interior_closed_interval by auto qed
 
-lemma inter_interior_unions_intervals: fixes f::"(real^'n) set set"
+lemma inter_interior_unions_intervals: fixes f::"('a::ordered_euclidean_space) set set"
   assumes "finite f" "open s" "\<forall>t\<in>f. \<exists>a b. t = {a..b}" "\<forall>t\<in>f. s \<inter> (interior t) = {}"
   shows "s \<inter> interior(\<Union>f) = {}" proof(rule ccontr,unfold ex_in_conv[THEN sym]) case goal1
   have lem1:"\<And>x e s U. ball x e \<subseteq> s \<inter> interior U \<longleftrightarrow> ball x e \<subseteq> s \<inter> U" apply rule  defer apply(rule_tac Int_greatest)
@@ -128,21 +129,22 @@
     guess a using insert(4)[rule_format,OF insertI1] .. then guess b .. note ab = this
     show ?case proof(cases "x\<in>i") case False hence "x \<in> UNIV - {a..b}" unfolding ab by auto
       then guess d unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_interval],rule_format] ..
-      hence "0 < d" "ball x (min d e) \<subseteq> UNIV - i" using e unfolding ab by auto
-      hence "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)" using e unfolding lem1 by auto hence "x \<in> s \<inter> interior (\<Union>f)" using `d>0` e by auto
+      hence "0 < d" "ball x (min d e) \<subseteq> UNIV - i" unfolding ab ball_min_Int by auto
+      hence "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)" using e unfolding lem1 unfolding  ball_min_Int by auto
+      hence "x \<in> s \<inter> interior (\<Union>f)" using `d>0` e by auto
       hence "\<exists>t\<in>f. \<exists>x e. 0 < e \<and> ball x e \<subseteq> s \<inter> t" apply-apply(rule insert(3)) using insert(4) by auto thus ?thesis by auto next
     case True show ?thesis proof(cases "x\<in>{a<..<b}")
       case True then guess d unfolding open_contains_ball_eq[OF open_interval,rule_format] ..
       thus ?thesis apply(rule_tac x=i in bexI,rule_tac x=x in exI,rule_tac x="min d e" in exI)
 	unfolding ab using interval_open_subset_closed[of a b] and e by fastsimp+ next
-    case False then obtain k where "x$k \<le> a$k \<or> x$k \<ge> b$k" unfolding mem_interval by(auto simp add:not_less) 
-    hence "x$k = a$k \<or> x$k = b$k" using True unfolding ab and mem_interval apply(erule_tac x=k in allE) by auto
+    case False then obtain k where "x$$k \<le> a$$k \<or> x$$k \<ge> b$$k" and k:"k<DIM('a)" unfolding mem_interval by(auto simp add:not_less) 
+    hence "x$$k = a$$k \<or> x$$k = b$$k" using True unfolding ab and mem_interval apply(erule_tac x=k in allE) by auto
     hence "\<exists>x. ball x (e/2) \<subseteq> s \<inter> (\<Union>f)" proof(erule_tac disjE)
-      let ?z = "x - (e/2) *\<^sub>R basis k" assume as:"x$k = a$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
+      let ?z = "x - (e/2) *\<^sub>R basis k" assume as:"x$$k = a$$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
 	fix y assume "y \<in> ball ?z (e / 2) \<inter> i" hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
-	hence "\<bar>(?z - y) $ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding dist_norm by auto
-	hence "y$k < a$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps)
-	hence "y \<notin> i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed
+	hence "\<bar>(?z - y) $$ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding dist_norm by auto
+	hence "y$$k < a$$k" using e[THEN conjunct1] k by(auto simp add:field_simps basis_component as)
+	hence "y \<notin> i" unfolding ab mem_interval not_all apply(rule_tac x=k in exI) using k by auto thus False using yi by auto qed
       moreover have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof
 	fix y assume as:"y\<in> ball ?z (e/2)" have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R basis k)"
 	   apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R basis k"])
@@ -150,15 +152,15 @@
 	also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball dist_norm using e by(auto simp add:field_simps)
 	finally show "y\<in>ball x e" unfolding mem_ball dist_norm using e by(auto simp add:field_simps) qed
       ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto
-    next let ?z = "x + (e/2) *\<^sub>R basis k" assume as:"x$k = b$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
+    next let ?z = "x + (e/2) *\<^sub>R basis k" assume as:"x$$k = b$$k" have "ball ?z (e / 2) \<inter> i = {}" apply(rule ccontr) unfolding ex_in_conv[THEN sym] proof(erule exE)
 	fix y assume "y \<in> ball ?z (e / 2) \<inter> i" hence "dist ?z y < e/2" and yi:"y\<in>i" by auto
-	hence "\<bar>(?z - y) $ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding dist_norm by auto
-	hence "y$k > b$k" unfolding vector_component_simps vector_scaleR_component as using e[THEN conjunct1] by(auto simp add:field_simps)
-	hence "y \<notin> i" unfolding ab mem_interval not_all by(rule_tac x=k in exI,auto) thus False using yi by auto qed
+	hence "\<bar>(?z - y) $$ k\<bar> < e/2" using component_le_norm[of "?z - y" k] unfolding dist_norm by auto
+	hence "y$$k > b$$k" using e[THEN conjunct1] k by(auto simp add:field_simps as)
+	hence "y \<notin> i" unfolding ab mem_interval not_all using k by(rule_tac x=k in exI,auto) thus False using yi by auto qed
       moreover have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)" apply(rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]]) proof
 	fix y assume as:"y\<in> ball ?z (e/2)" have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R basis k)"
 	   apply-apply(rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R basis k"])
-	  unfolding norm_scaleR norm_basis by auto
+	  unfolding norm_scaleR by auto
 	also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2" apply(rule add_strict_left_mono) using as unfolding mem_ball dist_norm using e by(auto simp add:field_simps)
 	finally show "y\<in>ball x e" unfolding mem_ball dist_norm using e by(auto simp add:field_simps) qed
       ultimately show ?thesis apply(rule_tac x="?z" in exI) unfolding Union_insert by auto qed 
@@ -167,117 +169,111 @@
   guess t using *[OF assms(1,3) goal1]  .. from this(2) guess x .. then guess e ..
   hence "x \<in> s" "x\<in>interior t" defer using open_subset_interior[OF open_ball, of x e t] by auto
   thus False using `t\<in>f` assms(4) by auto qed
+
 subsection {* Bounds on intervals where they exist. *}
 
-definition "interval_upperbound (s::(real^'n) set) = (\<chi> i. Sup {a. \<exists>x\<in>s. x$i = a})"
-
-definition "interval_lowerbound (s::(real^'n) set) = (\<chi> i. Inf {a. \<exists>x\<in>s. x$i = a})"
-
-lemma interval_upperbound[simp]: assumes "\<forall>i. a$i \<le> b$i" shows "interval_upperbound {a..b} = b"
-  using assms unfolding interval_upperbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE)
+definition "interval_upperbound (s::('a::ordered_euclidean_space) set) = ((\<chi>\<chi> i. Sup {a. \<exists>x\<in>s. x$$i = a})::'a)"
+
+definition "interval_lowerbound (s::('a::ordered_euclidean_space) set) = ((\<chi>\<chi> i. Inf {a. \<exists>x\<in>s. x$$i = a})::'a)"
+
+lemma interval_upperbound[simp]: assumes "\<forall>i<DIM('a::ordered_euclidean_space). a$$i \<le> (b::'a)$$i" shows "interval_upperbound {a..b} = b"
+  using assms unfolding interval_upperbound_def apply(subst euclidean_eq[where 'a='a]) apply safe
+  unfolding euclidean_lambda_beta' apply(erule_tac x=i in allE)
   apply(rule Sup_unique) unfolding setle_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer
-  apply(rule,rule) apply(rule_tac x="b$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=b in bexI)
+  apply(rule,rule) apply(rule_tac x="b$$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=b in bexI)
   unfolding mem_interval using assms by auto
 
-lemma interval_lowerbound[simp]: assumes "\<forall>i. a$i \<le> b$i" shows "interval_lowerbound {a..b} = a"
-  using assms unfolding interval_lowerbound_def Cart_eq Cart_lambda_beta apply-apply(rule,erule_tac x=i in allE)
+lemma interval_lowerbound[simp]: assumes "\<forall>i<DIM('a::ordered_euclidean_space). a$$i \<le> (b::'a)$$i" shows "interval_lowerbound {a..b} = a"
+  using assms unfolding interval_lowerbound_def apply(subst euclidean_eq[where 'a='a]) apply safe
+  unfolding euclidean_lambda_beta' apply(erule_tac x=i in allE)
   apply(rule Inf_unique) unfolding setge_def apply rule unfolding mem_Collect_eq apply(erule bexE) unfolding mem_interval defer
-  apply(rule,rule) apply(rule_tac x="a$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=a in bexI)
-  unfolding mem_interval using assms by auto
+  apply(rule,rule) apply(rule_tac x="a$$i" in bexI) defer unfolding mem_Collect_eq apply(rule_tac x=a in bexI)
+  unfolding mem_interval using assms by auto 
 
 lemmas interval_bounds = interval_upperbound interval_lowerbound
 
 lemma interval_bounds'[simp]: assumes "{a..b}\<noteq>{}" shows "interval_upperbound {a..b} = b" "interval_lowerbound {a..b} = a"
   using assms unfolding interval_ne_empty by auto
 
-lemma interval_upperbound_1[simp]: "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> interval_upperbound {a..b} = (b::real^1)"
-  apply(rule interval_upperbound) by auto
-
-lemma interval_lowerbound_1[simp]: "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> interval_lowerbound {a..b} = (a::real^1)"
-  apply(rule interval_lowerbound) by auto
-
-lemmas interval_bound_1 = interval_upperbound_1 interval_lowerbound_1
-
 subsection {* Content (length, area, volume...) of an interval. *}
 
-definition "content (s::(real^'n) set) =
-       (if s = {} then 0 else (\<Prod>i\<in>UNIV. (interval_upperbound s)$i - (interval_lowerbound s)$i))"
-
-lemma interval_not_empty:"\<forall>i. a$i \<le> b$i \<Longrightarrow> {a..b::real^'n} \<noteq> {}"
-  unfolding interval_eq_empty unfolding not_ex not_less by assumption
-
-lemma content_closed_interval: assumes "\<forall>i. a$i \<le> b$i"
-  shows "content {a..b} = (\<Prod>i\<in>UNIV. b$i - a$i)"
+definition "content (s::('a::ordered_euclidean_space) set) =
+       (if s = {} then 0 else (\<Prod>i<DIM('a). (interval_upperbound s)$$i - (interval_lowerbound s)$$i))"
+
+lemma interval_not_empty:"\<forall>i<DIM('a). a$$i \<le> b$$i \<Longrightarrow> {a..b::'a::ordered_euclidean_space} \<noteq> {}"
+  unfolding interval_eq_empty unfolding not_ex not_less by auto
+
+lemma content_closed_interval: fixes a::"'a::ordered_euclidean_space" assumes "\<forall>i<DIM('a). a$$i \<le> b$$i"
+  shows "content {a..b} = (\<Prod>i<DIM('a). b$$i - a$$i)"
   using interval_not_empty[OF assms] unfolding content_def interval_upperbound[OF assms] interval_lowerbound[OF assms] by auto
 
-lemma content_closed_interval': assumes "{a..b}\<noteq>{}" shows "content {a..b} = (\<Prod>i\<in>UNIV. b$i - a$i)"
+lemma content_closed_interval': fixes a::"'a::ordered_euclidean_space" assumes "{a..b}\<noteq>{}" shows "content {a..b} = (\<Prod>i<DIM('a). b$$i - a$$i)"
   apply(rule content_closed_interval) using assms unfolding interval_ne_empty .
 
-lemma content_1:"dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> content {a..b} = dest_vec1 b - dest_vec1 a"
-  using content_closed_interval[of a b] by auto
-
-lemma content_1':"a \<le> b \<Longrightarrow> content {vec1 a..vec1 b} = b - a" using content_1[of "vec a" "vec b"] by auto
-
-lemma content_unit[intro]: "content{0..1::real^'n} = 1" proof-
-  have *:"\<forall>i. 0$i \<le> (1::real^'n::finite)$i" by auto
-  have "0 \<in> {0..1::real^'n::finite}" unfolding mem_interval by auto
+lemma content_real:assumes "a\<le>b" shows "content {a..b} = b-a"
+proof- have *:"{..<Suc 0} = {0}" by auto
+  show ?thesis unfolding content_def using assms by(auto simp: *)
+qed
+
+lemma content_unit[intro]: "content{0..One::'a::ordered_euclidean_space} = 1" proof-
+  have *:"\<forall>i<DIM('a). (0::'a)$$i \<le> (One::'a)$$i" by auto
+  have "0 \<in> {0..One::'a}" unfolding mem_interval by auto
   thus ?thesis unfolding content_def interval_bounds[OF *] using setprod_1 by auto qed
 
-lemma content_pos_le[intro]: "0 \<le> content {a..b}" proof(cases "{a..b}={}")
-  case False hence *:"\<forall>i. a $ i \<le> b $ i" unfolding interval_ne_empty by assumption
-  have "(\<Prod>i\<in>UNIV. interval_upperbound {a..b} $ i - interval_lowerbound {a..b} $ i) \<ge> 0"
+lemma content_pos_le[intro]: fixes a::"'a::ordered_euclidean_space" shows "0 \<le> content {a..b}" proof(cases "{a..b}={}")
+  case False hence *:"\<forall>i<DIM('a). a $$ i \<le> b $$ i" unfolding interval_ne_empty by assumption
+  have "(\<Prod>i<DIM('a). interval_upperbound {a..b} $$ i - interval_lowerbound {a..b} $$ i) \<ge> 0"
     apply(rule setprod_nonneg) unfolding interval_bounds[OF *] using * apply(erule_tac x=x in allE) by auto
   thus ?thesis unfolding content_def by(auto simp del:interval_bounds') qed(unfold content_def, auto)
 
-lemma content_pos_lt: assumes "\<forall>i. a$i < b$i" shows "0 < content {a..b}"
-proof- have help_lemma1: "\<forall>i. a$i < b$i \<Longrightarrow> \<forall>i. a$i \<le> ((b$i)::real)" apply(rule,erule_tac x=i in allE) by auto
+lemma content_pos_lt: fixes a::"'a::ordered_euclidean_space" assumes "\<forall>i<DIM('a). a$$i < b$$i" shows "0 < content {a..b}"
+proof- have help_lemma1: "\<forall>i<DIM('a). a$$i < b$$i \<Longrightarrow> \<forall>i<DIM('a). a$$i \<le> ((b$$i)::real)" apply(rule,erule_tac x=i in allE) by auto
   show ?thesis unfolding content_closed_interval[OF help_lemma1[OF assms]] apply(rule setprod_pos)
     using assms apply(erule_tac x=x in allE) by auto qed
 
-lemma content_pos_lt_1: "dest_vec1 a < dest_vec1 b \<Longrightarrow> 0 < content({a..b})"
-  apply(rule content_pos_lt) by auto
-
-lemma content_eq_0: "content({a..b::real^'n}) = 0 \<longleftrightarrow> (\<exists>i. b$i \<le> a$i)" proof(cases "{a..b} = {}")
+lemma content_eq_0: "content{a..b::'a::ordered_euclidean_space} = 0 \<longleftrightarrow> (\<exists>i<DIM('a). b$$i \<le> a$$i)" proof(cases "{a..b} = {}")
   case True thus ?thesis unfolding content_def if_P[OF True] unfolding interval_eq_empty apply-
     apply(rule,erule exE) apply(rule_tac x=i in exI) by auto next
-  guess a using UNIV_witness[where 'a='n] .. case False note as=this[unfolded interval_eq_empty not_ex not_less]
-  show ?thesis unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_UNIV]
+  case False note this[unfolded interval_eq_empty not_ex not_less]
+  hence as:"\<forall>i<DIM('a). b $$ i \<ge> a $$ i" by fastsimp
+  show ?thesis unfolding content_def if_not_P[OF False] setprod_zero_iff[OF finite_lessThan]
     apply(rule) apply(erule_tac[!] exE bexE) unfolding interval_bounds[OF as] apply(rule_tac x=x in exI) defer
     apply(rule_tac x=i in bexI) using as apply(erule_tac x=i in allE) by auto qed
 
 lemma cond_cases:"(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)" by auto
 
 lemma content_closed_interval_cases:
-  "content {a..b} = (if \<forall>i. a$i \<le> b$i then setprod (\<lambda>i. b$i - a$i) UNIV else 0)" apply(rule cond_cases) 
+  "content {a..b::'a::ordered_euclidean_space} = (if \<forall>i<DIM('a). a$$i \<le> b$$i then setprod (\<lambda>i. b$$i - a$$i) {..<DIM('a)} else 0)" apply(rule cond_cases) 
   apply(rule content_closed_interval) unfolding content_eq_0 not_all not_le defer apply(erule exE,rule_tac x=x in exI) by auto
 
 lemma content_eq_0_interior: "content {a..b} = 0 \<longleftrightarrow> interior({a..b}) = {}"
   unfolding content_eq_0 interior_closed_interval interval_eq_empty by auto
 
-lemma content_eq_0_1: "content {a..b::real^1} = 0 \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
-  unfolding content_eq_0 by auto
-
-lemma content_pos_lt_eq: "0 < content {a..b} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
+(*lemma content_eq_0_1: "content {a..b::real^1} = 0 \<longleftrightarrow> dest_vec1 b \<le> dest_vec1 a"
+  unfolding content_eq_0 by auto*)
+
+lemma content_pos_lt_eq: "0 < content {a..b::'a::ordered_euclidean_space} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i < b$$i)"
   apply(rule) defer apply(rule content_pos_lt,assumption) proof- assume "0 < content {a..b}"
-  hence "content {a..b} \<noteq> 0" by auto thus "\<forall>i. a$i < b$i" unfolding content_eq_0 not_ex not_le by auto qed
+  hence "content {a..b} \<noteq> 0" by auto thus "\<forall>i<DIM('a). a$$i < b$$i" unfolding content_eq_0 not_ex not_le by fastsimp qed
 
 lemma content_empty[simp]: "content {} = 0" unfolding content_def by auto
 
-lemma content_subset: assumes "{a..b} \<subseteq> {c..d}" shows "content {a..b::real^'n} \<le> content {c..d}" proof(cases "{a..b}={}")
+lemma content_subset: assumes "{a..b} \<subseteq> {c..d}" shows "content {a..b::'a::ordered_euclidean_space} \<le> content {c..d}" proof(cases "{a..b}={}")
   case True thus ?thesis using content_pos_le[of c d] by auto next
-  case False hence ab_ne:"\<forall>i. a $ i \<le> b $ i" unfolding interval_ne_empty by auto
+  case False hence ab_ne:"\<forall>i<DIM('a). a $$ i \<le> b $$ i" unfolding interval_ne_empty by auto
   hence ab_ab:"a\<in>{a..b}" "b\<in>{a..b}" unfolding mem_interval by auto
   have "{c..d} \<noteq> {}" using assms False by auto
-  hence cd_ne:"\<forall>i. c $ i \<le> d $ i" using assms unfolding interval_ne_empty by auto
+  hence cd_ne:"\<forall>i<DIM('a). c $$ i \<le> d $$ i" using assms unfolding interval_ne_empty by auto
   show ?thesis unfolding content_def unfolding interval_bounds[OF ab_ne] interval_bounds[OF cd_ne]
-    unfolding if_not_P[OF False] if_not_P[OF `{c..d} \<noteq> {}`] apply(rule setprod_mono,rule) proof fix i::'n
-    show "0 \<le> b $ i - a $ i" using ab_ne[THEN spec[where x=i]] by auto
-    show "b $ i - a $ i \<le> d $ i - c $ i"
+    unfolding if_not_P[OF False] if_not_P[OF `{c..d} \<noteq> {}`] apply(rule setprod_mono,rule) proof
+    fix i assume i:"i\<in>{..<DIM('a)}"
+    show "0 \<le> b $$ i - a $$ i" using ab_ne[THEN spec[where x=i]] i by auto
+    show "b $$ i - a $$ i \<le> d $$ i - c $$ i"
       using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(2),of i]
-      using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i] by auto qed qed
+      using assms[unfolded subset_eq mem_interval,rule_format,OF ab_ab(1),of i] using i by auto qed qed
 
 lemma content_lt_nz: "0 < content {a..b} \<longleftrightarrow> content {a..b} \<noteq> 0"
-  unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by auto
+  unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastsimp
 
 subsection {* The notion of a gauge --- simply an open set containing the point. *}
 
@@ -331,13 +327,13 @@
 
 lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}" unfolding division_of_def by auto 
 
-lemma division_of_sing[simp]: "s division_of {a..a::real^'n} \<longleftrightarrow> s = {{a..a}}" (is "?l = ?r") proof
+lemma division_of_sing[simp]: "s division_of {a..a::'a::ordered_euclidean_space} \<longleftrightarrow> s = {{a..a}}" (is "?l = ?r") proof
   assume ?r moreover { assume "s = {{a}}" moreover fix k assume "k\<in>s" 
-    ultimately have"\<exists>x y. k = {x..y}" apply(rule_tac x=a in exI)+ unfolding interval_sing[THEN conjunct1] by auto }
-  ultimately show ?l unfolding division_of_def interval_sing[THEN conjunct1] by auto next
-  assume ?l note as=conjunctD4[OF this[unfolded division_of_def interval_sing[THEN conjunct1]]]
+    ultimately have"\<exists>x y. k = {x..y}" apply(rule_tac x=a in exI)+ unfolding interval_sing by auto }
+  ultimately show ?l unfolding division_of_def interval_sing by auto next
+  assume ?l note as=conjunctD4[OF this[unfolded division_of_def interval_sing]]
   { fix x assume x:"x\<in>s" have "x={a}" using as(2)[rule_format,OF x] by auto }
-  moreover have "s \<noteq> {}" using as(4) by auto ultimately show ?r unfolding interval_sing[THEN conjunct1] by auto qed
+  moreover have "s \<noteq> {}" using as(4) by auto ultimately show ?r unfolding interval_sing by auto qed
 
 lemma elementary_empty: obtains p where "p division_of {}"
   unfolding division_of_trivial by auto
@@ -366,7 +362,7 @@
   unfolding forall_in_division[OF assms(2)] apply(rule,rule,rule) apply(drule division_ofD(2)[OF assms(2)])
   apply(drule content_subset) unfolding assms(1) proof- case goal1 thus ?case using content_pos_le[of a b] by auto qed
 
-lemma division_inter: assumes "p1 division_of s1" "p2 division_of (s2::(real^'a) set)"
+lemma division_inter: assumes "p1 division_of s1" "p2 division_of (s2::('a::ordered_euclidean_space) set)"
   shows "{k1 \<inter> k2 | k1 k2 .k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)" (is "?A' division_of _") proof-
 let ?A = "{s. s \<in>  (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}" have *:"?A' = ?A" by auto
 show ?thesis unfolding * proof(rule division_ofI) have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)" by auto
@@ -389,17 +385,17 @@
     using division_ofD(5)[OF assms(1) k1(2) k2(2)]
     using division_ofD(5)[OF assms(2) k1(3) k2(3)] using th by auto qed qed
 
-lemma division_inter_1: assumes "d division_of i" "{a..b::real^'n} \<subseteq> i"
+lemma division_inter_1: assumes "d division_of i" "{a..b::'a::ordered_euclidean_space} \<subseteq> i"
   shows "{ {a..b} \<inter> k |k. k \<in> d \<and> {a..b} \<inter> k \<noteq> {} } division_of {a..b}" proof(cases "{a..b} = {}")
   case True show ?thesis unfolding True and division_of_trivial by auto next
   have *:"{a..b} \<inter> i = {a..b}" using assms(2) by auto 
   case False show ?thesis using division_inter[OF division_of_self[OF False] assms(1)] unfolding * by auto qed
 
-lemma elementary_inter: assumes "p1 division_of s" "p2 division_of (t::(real^'n) set)"
+lemma elementary_inter: assumes "p1 division_of s" "p2 division_of (t::('a::ordered_euclidean_space) set)"
   shows "\<exists>p. p division_of (s \<inter> t)"
   by(rule,rule division_inter[OF assms])
 
-lemma elementary_inters: assumes "finite f" "f\<noteq>{}" "\<forall>s\<in>f. \<exists>p. p division_of (s::(real^'n) set)"
+lemma elementary_inters: assumes "finite f" "f\<noteq>{}" "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::ordered_euclidean_space) set)"
   shows "\<exists>p. p division_of (\<Inter> f)" using assms apply-proof(induct f rule:finite_induct)
 case (insert x f) show ?case proof(cases "f={}")
   case True thus ?thesis unfolding True using insert by auto next
@@ -422,57 +418,71 @@
   fix k assume k:"k \<in> p1 \<union> p2"  show "k \<subseteq> s1 \<union> s2" using k d1(2) d2(2) by auto
   show "k \<noteq> {}" using k d1(3) d2(3) by auto show "\<exists>a b. k = {a..b}" using k d1(4) d2(4) by auto qed
 
+(* move *)
+lemma Eucl_nth_inverse[simp]: fixes x::"'a::euclidean_space" shows "(\<chi>\<chi> i. x $$ i) = x"
+  apply(subst euclidean_eq) by auto
+
 lemma partial_division_extend_1:
-  assumes "{c..d} \<subseteq> {a..b::real^'n}" "{c..d} \<noteq> {}"
+  assumes "{c..d} \<subseteq> {a..b::'a::ordered_euclidean_space}" "{c..d} \<noteq> {}"
   obtains p where "p division_of {a..b}" "{c..d} \<in> p"
-proof- def n \<equiv> "CARD('n)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def by auto
-  guess \<pi> using ex_bij_betw_nat_finite_1[OF finite_UNIV[where 'a='n]] .. note \<pi>=this
+proof- def n \<equiv> "DIM('a)" have n:"1 \<le> n" "0 < n" "n \<noteq> 0" unfolding n_def using DIM_positive[where 'a='a] by auto
+  guess \<pi> using ex_bij_betw_nat_finite_1[OF finite_lessThan[of "DIM('a)"]] .. note \<pi>=this
   def \<pi>' \<equiv> "inv_into {1..n} \<pi>"
-  have \<pi>':"bij_betw \<pi>' UNIV {1..n}" using bij_betw_inv_into[OF \<pi>] unfolding \<pi>'_def n_def by auto
-  hence \<pi>'i:"\<And>i. \<pi>' i \<in> {1..n}" unfolding bij_betw_def by auto 
-  have \<pi>\<pi>'[simp]:"\<And>i. \<pi> (\<pi>' i) = i" unfolding \<pi>'_def apply(rule f_inv_into_f) unfolding n_def using \<pi> unfolding bij_betw_def by auto
-  have \<pi>'\<pi>[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi>' (\<pi> i) = i" unfolding \<pi>'_def apply(rule inv_into_f_eq) using \<pi> unfolding n_def bij_betw_def by auto
+  have \<pi>':"bij_betw \<pi>' {..<DIM('a)} {1..n}" using bij_betw_inv_into[OF \<pi>] unfolding \<pi>'_def n_def by auto
+  hence \<pi>'i:"\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i \<in> {1..n}" unfolding bij_betw_def by auto 
+  have \<pi>i:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi> i <DIM('a)" using \<pi> unfolding bij_betw_def n_def by auto 
+  have \<pi>\<pi>'[simp]:"\<And>i. i<DIM('a) \<Longrightarrow> \<pi> (\<pi>' i) = i" unfolding \<pi>'_def
+    apply(rule f_inv_into_f) unfolding n_def using \<pi> unfolding bij_betw_def by auto
+  have \<pi>'\<pi>[simp]:"\<And>i. i\<in>{1..n} \<Longrightarrow> \<pi>' (\<pi> i) = i" unfolding \<pi>'_def apply(rule inv_into_f_eq)
+    using \<pi> unfolding n_def bij_betw_def by auto
   have "{c..d} \<noteq> {}" using assms by auto
-  let ?p1 = "\<lambda>l. {(\<chi> i. if \<pi>' i < l then c$i else a$i) .. (\<chi> i. if \<pi>' i < l then d$i else if \<pi>' i = l then c$\<pi> l else b$i)}"
-  let ?p2 = "\<lambda>l. {(\<chi> i. if \<pi>' i < l then c$i else if \<pi>' i = l then d$\<pi> l else a$i) .. (\<chi> i. if \<pi>' i < l then d$i else b$i)}"
+  let ?p1 = "\<lambda>l. {(\<chi>\<chi> i. if \<pi>' i < l then c$$i else a$$i)::'a .. (\<chi>\<chi> i. if \<pi>' i < l then d$$i else if \<pi>' i = l then c$$\<pi> l else b$$i)}"
+  let ?p2 = "\<lambda>l. {(\<chi>\<chi> i. if \<pi>' i < l then c$$i else if \<pi>' i = l then d$$\<pi> l else a$$i)::'a .. (\<chi>\<chi> i. if \<pi>' i < l then d$$i else b$$i)}"
   let ?p =  "{?p1 l |l. l \<in> {1..n+1}} \<union> {?p2 l |l. l \<in> {1..n+1}}"
-  have abcd:"\<And>i. a $ i \<le> c $ i \<and> c$i \<le> d$i \<and> d $ i \<le> b $ i" using assms unfolding subset_interval interval_eq_empty by(auto simp add:not_le not_less)
+  have abcd:"\<And>i. i<DIM('a) \<Longrightarrow> a $$ i \<le> c $$ i \<and> c$$i \<le> d$$i \<and> d $$ i \<le> b $$ i" using assms
+    unfolding subset_interval interval_eq_empty by auto
   show ?thesis apply(rule that[of ?p]) apply(rule division_ofI)
-  proof- have "\<And>i. \<pi>' i < Suc n"
-    proof(rule ccontr,unfold not_less) fix i assume "Suc n \<le> \<pi>' i"
-      hence "\<pi>' i \<notin> {1..n}" by auto thus False using \<pi>' unfolding bij_betw_def by auto
-    qed hence "c = (\<chi> i. if \<pi>' i < Suc n then c $ i else a $ i)"
-        "d = (\<chi> i. if \<pi>' i < Suc n then d $ i else if \<pi>' i = n + 1 then c $ \<pi> (n + 1) else b $ i)"
-      unfolding Cart_eq Cart_lambda_beta using \<pi>' unfolding bij_betw_def by auto
+  proof- have "\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i < Suc n"
+    proof(rule ccontr,unfold not_less) fix i assume i:"i<DIM('a)" and "Suc n \<le> \<pi>' i"
+      hence "\<pi>' i \<notin> {1..n}" by auto thus False using \<pi>' i unfolding bij_betw_def by auto
+    qed hence "c = (\<chi>\<chi> i. if \<pi>' i < Suc n then c $$ i else a $$ i)"
+        "d = (\<chi>\<chi> i. if \<pi>' i < Suc n then d $$ i else if \<pi>' i = n + 1 then c $$ \<pi> (n + 1) else b $$ i)"
+      unfolding euclidean_eq[where 'a='a] using \<pi>' unfolding bij_betw_def by auto
     thus cdp:"{c..d} \<in> ?p" apply-apply(rule UnI1) unfolding mem_Collect_eq apply(rule_tac x="n + 1" in exI) by auto
     have "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p1 l \<subseteq> {a..b}"  "\<And>l. l\<in>{1..n+1} \<Longrightarrow> ?p2 l \<subseteq> {a..b}"
       unfolding subset_eq apply(rule_tac[!] ballI,rule_tac[!] ccontr)
     proof- fix l assume l:"l\<in>{1..n+1}" fix x assume "x\<notin>{a..b}"
-      then guess i unfolding mem_interval not_all .. note i=this
+      then guess i unfolding mem_interval not_all not_imp .. note i=conjunctD2[OF this]
       show "x \<in> ?p1 l \<Longrightarrow> False" "x \<in> ?p2 l \<Longrightarrow> False" unfolding mem_interval apply(erule_tac[!] x=i in allE)
         apply(case_tac[!] "\<pi>' i < l", case_tac[!] "\<pi>' i = l") using abcd[of i] i by auto 
     qed moreover have "\<And>x. x \<in> {a..b} \<Longrightarrow> x \<in> \<Union>?p"
     proof- fix x assume x:"x\<in>{a..b}"
       { presume "x\<notin>{c..d} \<Longrightarrow> x \<in> \<Union>?p" thus "x \<in> \<Union>?p" using cdp by blast }
-      let ?M = "{i. i\<in>{1..n+1} \<and> \<not> (c $ \<pi> i \<le> x $ \<pi> i \<and> x $ \<pi> i \<le> d $ \<pi> i)}"
-      assume "x\<notin>{c..d}" then guess i0 unfolding mem_interval not_all ..
+      let ?M = "{i. i\<in>{1..n+1} \<and> \<not> (c $$ \<pi> i \<le> x $$ \<pi> i \<and> x $$ \<pi> i \<le> d $$ \<pi> i)}"
+      assume "x\<notin>{c..d}" then guess i0 unfolding mem_interval not_all not_imp ..
       hence "\<pi>' i0 \<in> ?M" using \<pi>' unfolding bij_betw_def by(auto intro!:le_SucI)
       hence M:"finite ?M" "?M \<noteq> {}" by auto
       def l \<equiv> "Min ?M" note l = Min_less_iff[OF M,unfolded l_def[symmetric]] Min_in[OF M,unfolded mem_Collect_eq l_def[symmetric]]
         Min_gr_iff[OF M,unfolded l_def[symmetric]]
       have "x\<in>?p1 l \<or> x\<in>?p2 l" using l(2)[THEN conjunct2] unfolding de_Morgan_conj not_le
         apply- apply(erule disjE) apply(rule disjI1) defer apply(rule disjI2)
-      proof- assume as:"x $ \<pi> l < c $ \<pi> l"
-        show "x \<in> ?p1 l" unfolding mem_interval Cart_lambda_beta
-        proof case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le by auto
+      proof- assume as:"x $$ \<pi> l < c $$ \<pi> l"
+        show "x \<in> ?p1 l" unfolding mem_interval apply safe unfolding euclidean_lambda_beta'
+        proof- case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le using goal1 by auto
           thus ?case using as x[unfolded mem_interval,rule_format,of i]
-            apply auto using l(3)[of "\<pi>' i"] by(auto elim!:ballE[where x="\<pi>' i"])
+            apply auto using l(3)[of "\<pi>' i"] using goal1 by(auto elim!:ballE[where x="\<pi>' i"])
+        next case goal2 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le using goal2 by auto
+          thus ?case using as x[unfolded mem_interval,rule_format,of i]
+            apply auto using l(3)[of "\<pi>' i"] using goal2 by(auto elim!:ballE[where x="\<pi>' i"])
         qed
-      next assume as:"x $ \<pi> l > d $ \<pi> l"
-        show "x \<in> ?p2 l" unfolding mem_interval Cart_lambda_beta
-        proof case goal1 have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le by auto
-          thus ?case using as x[unfolded mem_interval,rule_format,of i]
-            apply auto using l(3)[of "\<pi>' i"] by(auto elim!:ballE[where x="\<pi>' i"])
+      next assume as:"x $$ \<pi> l > d $$ \<pi> l"
+        show "x \<in> ?p2 l" unfolding mem_interval apply safe unfolding euclidean_lambda_beta'
+        proof- fix i assume i:"i<DIM('a)"
+          have "\<pi>' i \<in> {1..n}" using \<pi>' unfolding bij_betw_def not_le using i by auto
+          thus "(if \<pi>' i < l then c $$ i else if \<pi>' i = l then d $$ \<pi> l else a $$ i) \<le> x $$ i"
+            "x $$ i \<le> (if \<pi>' i < l then d $$ i else b $$ i)"
+            using as x[unfolded mem_interval,rule_format,of i]
+            apply auto using l(3)[of "\<pi>' i"] i by(auto elim!:ballE[where x="\<pi>' i"])
         qed qed
       thus "x \<in> \<Union>?p" using l(2) by blast 
     qed ultimately show "\<Union>?p = {a..b}" apply-apply(rule) defer apply(rule) by(assumption,blast)
@@ -480,8 +490,9 @@
     show "finite ?p" by auto
     fix k assume k:"k\<in>?p" then obtain l where l:"k = ?p1 l \<or> k = ?p2 l" "l \<in> {1..n + 1}" by auto
     show "k\<subseteq>{a..b}" apply(rule,unfold mem_interval,rule,rule) 
-    proof- fix i::'n and x assume "x \<in> k" moreover have "\<pi>' i < l \<or> \<pi>' i = l \<or> \<pi>' i > l" by auto
-      ultimately show "a$i \<le> x$i" "x$i \<le> b$i" using abcd[of i] using l by(auto elim:disjE elim!:allE[where x=i] simp add:vector_le_def)
+    proof fix i x assume i:"i<DIM('a)" assume "x \<in> k" moreover have "\<pi>' i < l \<or> \<pi>' i = l \<or> \<pi>' i > l" by auto
+      ultimately show "a$$i \<le> x$$i" "x$$i \<le> b$$i" using abcd[of i] using l using i
+        by(auto elim:disjE elim!:allE[where x=i] simp add:eucl_le[where 'a='a])
     qed have "\<And>l. ?p1 l \<noteq> {}" "\<And>l. ?p2 l \<noteq> {}" unfolding interval_eq_empty not_ex apply(rule_tac[!] allI)
     proof- case goal1 thus ?case using abcd[of x] by auto
     next   case goal2 thus ?case using abcd[of x] by auto
@@ -494,20 +505,22 @@
       assume "l \<le> l'" fix x
       have "x \<notin> interior k \<inter> interior k'" 
       proof(rule,cases "l' = n+1") assume x:"x \<in> interior k \<inter> interior k'"
-        case True hence "\<And>i. \<pi>' i < l'" using \<pi>'i by(auto simp add:less_Suc_eq_le)
-        hence k':"k' = {c..d}" using l'(1) \<pi>'i by(auto simp add:Cart_nth_inverse)
+        case True hence "\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i < l'" using \<pi>'i using l' by(auto simp add:less_Suc_eq_le)
+        hence *:"\<And> P Q. (\<chi>\<chi> i. if \<pi>' i < l' then P i else Q i) = ((\<chi>\<chi> i. P i)::'a)" apply-apply(subst euclidean_eq) by auto
+        hence k':"k' = {c..d}" using l'(1) unfolding * by auto
         have ln:"l < n + 1" 
         proof(rule ccontr) case goal1 hence l2:"l = n+1" using l by auto
-          hence "\<And>i. \<pi>' i < l" using \<pi>'i by(auto simp add:less_Suc_eq_le)
-          hence "k = {c..d}" using l(1) \<pi>'i by(auto simp add:Cart_nth_inverse)
+          hence "\<And>i. i<DIM('a) \<Longrightarrow> \<pi>' i < l" using \<pi>'i by(auto simp add:less_Suc_eq_le)
+          hence *:"\<And> P Q. (\<chi>\<chi> i. if \<pi>' i < l then P i else Q i) = ((\<chi>\<chi> i. P i)::'a)" apply-apply(subst euclidean_eq) by auto
+          hence "k = {c..d}" using l(1) \<pi>'i unfolding * by(auto)
           thus False using `k\<noteq>k'` k' by auto
         qed have **:"\<pi>' (\<pi> l) = l" using \<pi>'\<pi>[of l] using l ln by auto
-        have "x $ \<pi> l < c $ \<pi> l \<or> d $ \<pi> l < x $ \<pi> l" using l(1) apply-
+        have "x $$ \<pi> l < c $$ \<pi> l \<or> d $$ \<pi> l < x $$ \<pi> l" using l(1) apply-
         proof(erule disjE)
           assume as:"k = ?p1 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show ?thesis using *[of "\<pi> l"] using ln unfolding Cart_lambda_beta ** by auto
+          show ?thesis using *[of "\<pi> l"] using ln l(2) using \<pi>i[of l] by(auto simp add:** not_less)
         next assume as:"k = ?p2 l" note * = conjunct1[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show ?thesis using *[of "\<pi> l"] using ln unfolding Cart_lambda_beta ** by auto
+          show ?thesis using *[of "\<pi> l"] using ln l(2) using \<pi>i[of l] unfolding ** by auto
         qed thus False using x unfolding k' unfolding Int_iff interior_closed_interval mem_interval
           by(auto elim!:allE[where x="\<pi> l"])
       next case False hence "l < n + 1" using l'(2) using `l\<le>l'` by auto
@@ -517,22 +530,22 @@
         show False using l(1) l'(1) apply-
         proof(erule_tac[!] disjE)+
           assume as:"k = ?p1 l" "k' = ?p1 l'"
-          note * = x[unfolded as Int_iff interior_closed_interval mem_interval]
+          note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
           have "l \<noteq> l'" using k'(2)[unfolded as] by auto
-          thus False using * by(smt Cart_lambda_beta \<pi>l)
+          thus False using *[of "\<pi> l'"] *[of "\<pi> l"] ln using \<pi>i[OF ln(1)] \<pi>i[OF ln(2)] apply(cases "l<l'")
+            by(auto simp add:euclidean_lambda_beta' \<pi>l \<pi>i n_def)
         next assume as:"k = ?p2 l" "k' = ?p2 l'"
           note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
           have "l \<noteq> l'" apply(rule) using k'(2)[unfolded as] by auto
-          thus False using *[of "\<pi> l"] *[of "\<pi> l'"]
-            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` by auto
+          thus False using *[of "\<pi> l"] *[of "\<pi> l'"]  `l \<le> l'` ln by(auto simp add:euclidean_lambda_beta' \<pi>l \<pi>i n_def)
         next assume as:"k = ?p1 l" "k' = ?p2 l'"
           note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show False using *[of "\<pi> l"] *[of "\<pi> l'"]
-            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` using abcd[of "\<pi> l'"] by smt 
+          show False using abcd[of "\<pi> l'"] using *[of "\<pi> l"] *[of "\<pi> l'"]  `l \<le> l'` ln apply(cases "l=l'")
+            by(auto simp add:euclidean_lambda_beta' \<pi>l \<pi>i n_def)
         next assume as:"k = ?p2 l" "k' = ?p1 l'"
           note * = conjunctD2[OF x[unfolded as Int_iff interior_closed_interval mem_interval],rule_format]
-          show False using *[of "\<pi> l"] *[of "\<pi> l'"]
-            unfolding Cart_lambda_beta \<pi>l using `l \<le> l'` using abcd[of "\<pi> l'"] by smt
+          show False using *[of "\<pi> l"] *[of "\<pi> l'"] ln `l \<le> l'` apply(cases "l=l'") using abcd[of "\<pi> l'"] 
+            by(auto simp add:euclidean_lambda_beta' \<pi>l \<pi>i n_def)
         qed qed } 
     from this[OF k l k' l'] this[OF k'(1) l' k _ l] have "\<And>x. x \<notin> interior k \<inter> interior k'"
       apply - apply(cases "l' \<le> l") using k'(2) by auto            
@@ -540,7 +553,7 @@
 qed qed
 
 lemma partial_division_extend_interval: assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> {a..b}"
-  obtains q where "p \<subseteq> q" "q division_of {a..b::real^'n}" proof(cases "p = {}")
+  obtains q where "p \<subseteq> q" "q division_of {a..b::'a::ordered_euclidean_space}" proof(cases "p = {}")
   case True guess q apply(rule elementary_interval[of a b]) .
   thus ?thesis apply- apply(rule that[of q]) unfolding True by auto next
   case False note p = division_ofD[OF assms(1)]
@@ -569,13 +582,13 @@
 	have *:"\<And>x s. x \<in> s \<Longrightarrow> \<Inter>s \<subseteq> x" by auto show "interior (\<Inter>(\<lambda>i. \<Union>(q i - {i})) ` p) \<subseteq> interior (\<Union>(q k - {k}))"
 	  apply(rule subset_interior *)+ using k by auto qed qed qed auto qed
 
-lemma elementary_bounded[dest]: "p division_of s \<Longrightarrow> bounded (s::(real^'n) set)"
+lemma elementary_bounded[dest]: "p division_of s \<Longrightarrow> bounded (s::('a::ordered_euclidean_space) set)"
   unfolding division_of_def by(metis bounded_Union bounded_interval) 
 
-lemma elementary_subset_interval: "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> {a..b::real^'n}"
+lemma elementary_subset_interval: "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> {a..b::'a::ordered_euclidean_space}"
   by(meson elementary_bounded bounded_subset_closed_interval)
 
-lemma division_union_intervals_exists: assumes "{a..b::real^'n} \<noteq> {}"
+lemma division_union_intervals_exists: assumes "{a..b::'a::ordered_euclidean_space} \<noteq> {}"
   obtains p where "(insert {a..b} p) division_of ({a..b} \<union> {c..d})" proof(cases "{c..d} = {}")
   case True show ?thesis apply(rule that[of "{}"]) unfolding True using assms by auto next
   case False note false=this show ?thesis proof(cases "{a..b} \<inter> {c..d} = {}")
@@ -602,7 +615,7 @@
   using division_ofD[OF assms(2)] by auto
   
 lemma elementary_union_interval: assumes "p division_of \<Union>p"
-  obtains q where "q division_of ({a..b::real^'n} \<union> \<Union>p)" proof-
+  obtains q where "q division_of ({a..b::'a::ordered_euclidean_space} \<union> \<Union>p)" proof-
   note assm=division_ofD[OF assms]
   have lem1:"\<And>f s. \<Union>\<Union> (f ` s) = \<Union>(\<lambda>x.\<Union>(f x)) ` s" by auto
   have lem2:"\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f" by auto
@@ -653,7 +666,7 @@
     qed qed } qed
 
 lemma elementary_unions_intervals:
-  assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = {a..b::real^'n}"
+  assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = {a..b::'a::ordered_euclidean_space}"
   obtains p where "p division_of (\<Union>f)" proof-
   have "\<exists>p. p division_of (\<Union>f)" proof(induct_tac f rule:finite_subset_induct) 
     show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
@@ -665,7 +678,7 @@
       unfolding Union_insert ab * by auto
   qed(insert assms,auto) thus ?thesis apply-apply(erule exE,rule that) by auto qed
 
-lemma elementary_union: assumes "ps division_of s" "pt division_of (t::(real^'n) set)"
+lemma elementary_union: assumes "ps division_of s" "pt division_of (t::('a::ordered_euclidean_space) set)"
   obtains p where "p division_of (s \<union> t)"
 proof- have "s \<union> t = \<Union>ps \<union> \<Union>pt" using assms unfolding division_of_def by auto
   hence *:"\<Union>(ps \<union> pt) = s \<union> t" by auto
@@ -673,7 +686,7 @@
     unfolding * prefer 3 apply(rule_tac p=p in that)
     using assms[unfolded division_of_def] by auto qed
 
-lemma partial_division_extend: fixes t::"(real^'n) set"
+lemma partial_division_extend: fixes t::"('a::ordered_euclidean_space) set"
   assumes "p division_of s" "q division_of t" "s \<subseteq> t"
   obtains r where "p \<subseteq> r" "r division_of t" proof-
   note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
@@ -776,7 +789,7 @@
   shows "t tagged_partial_division_of i"
   using assms unfolding tagged_partial_division_of_def using finite_subset[OF assms(2)] by blast
 
-lemma setsum_over_tagged_division_lemma: fixes d::"(real^'m) set \<Rightarrow> 'a::real_normed_vector"
+lemma setsum_over_tagged_division_lemma: fixes d::"('m::ordered_euclidean_space) set \<Rightarrow> 'a::real_normed_vector"
   assumes "p tagged_division_of i" "\<And>u v. {u..v} \<noteq> {} \<Longrightarrow> content {u..v} = 0 \<Longrightarrow> d {u..v} = 0"
   shows "setsum (\<lambda>(x,k). d k) p = setsum d (snd ` p)"
 proof- note assm=tagged_division_ofD[OF assms(1)]
@@ -886,7 +899,7 @@
                         \<longrightarrow> norm(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))"
 
 definition has_integral (infixr "has'_integral" 46) where 
-"((f::(real^'n \<Rightarrow> 'b::real_normed_vector)) has_integral y) i \<equiv>
+"((f::('n::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector)) has_integral y) i \<equiv>
         if (\<exists>a b. i = {a..b}) then (f has_integral_compact_interval y) i
         else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b}
               \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral_compact_interval z) {a..b} \<and>
@@ -933,14 +946,6 @@
 lemma has_integral_integral:"f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s"
   by auto
 
-lemma has_integral_vec1: assumes "(f has_integral k) {a..b}"
-  shows "((\<lambda>x. vec1 (f x)) has_integral (vec1 k)) {a..b}"
-proof- have *:"\<And>p. (\<Sum>(x, k)\<in>p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k)"
-    unfolding vec_sub Cart_eq by(auto simp add: split_beta)
-  show ?thesis using assms unfolding has_integral apply safe
-    apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe)
-    apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed
-
 lemma setsum_content_null:
   assumes "content({a..b}) = 0" "p tagged_division_of {a..b}"
   shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
@@ -968,15 +973,15 @@
 
 subsection {* The set we're concerned with must be closed. *}
 
-lemma division_of_closed: "s division_of i \<Longrightarrow> closed (i::(real^'n) set)"
+lemma division_of_closed: "s division_of i \<Longrightarrow> closed (i::('n::ordered_euclidean_space) set)"
   unfolding division_of_def by(fastsimp intro!: closed_Union closed_interval)
 
 subsection {* General bisection principle for intervals; might be useful elsewhere. *}
 
-lemma interval_bisection_step:
-  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "~(P {a..b::real^'n})"
+lemma interval_bisection_step:  fixes type::"'a::ordered_euclidean_space"
+  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "~(P {a..b::'a})"
   obtains c d where "~(P{c..d})"
-  "\<forall>i. a$i \<le> c$i \<and> c$i \<le> d$i \<and> d$i \<le> b$i \<and> 2 * (d$i - c$i) \<le> b$i - a$i"
+  "\<forall>i<DIM('a). a$$i \<le> c$$i \<and> c$$i \<le> d$$i \<and> d$$i \<le> b$$i \<and> 2 * (d$$i - c$$i) \<le> b$$i - a$$i"
 proof- have "{a..b} \<noteq> {}" using assms(1,3) by auto
   note ab=this[unfolded interval_eq_empty not_ex not_less]
   { fix f have "finite f \<Longrightarrow>
@@ -989,68 +994,70 @@
         apply rule defer apply rule defer apply(rule inter_interior_unions_intervals)
         using insert by auto
     qed } note * = this
-  let ?A = "{{c..d} | c d. \<forall>i. (c$i = a$i) \<and> (d$i = (a$i + b$i) / 2) \<or> (c$i = (a$i + b$i) / 2) \<and> (d$i = b$i)}"
-  let ?PP = "\<lambda>c d. \<forall>i. a$i \<le> c$i \<and> c$i \<le> d$i \<and> d$i \<le> b$i \<and> 2 * (d$i - c$i) \<le> b$i - a$i"
+  let ?A = "{{c..d} | c d::'a. \<forall>i<DIM('a). (c$$i = a$$i) \<and> (d$$i = (a$$i + b$$i) / 2) \<or> (c$$i = (a$$i + b$$i) / 2) \<and> (d$$i = b$$i)}"
+  let ?PP = "\<lambda>c d. \<forall>i<DIM('a). a$$i \<le> c$$i \<and> c$$i \<le> d$$i \<and> d$$i \<le> b$$i \<and> 2 * (d$$i - c$$i) \<le> b$$i - a$$i"
   { presume "\<forall>c d. ?PP c d \<longrightarrow> P {c..d} \<Longrightarrow> False"
     thus thesis unfolding atomize_not not_all apply-apply(erule exE)+ apply(rule_tac c=x and d=xa in that) by auto }
   assume as:"\<forall>c d. ?PP c d \<longrightarrow> P {c..d}"
   have "P (\<Union> ?A)" proof(rule *, rule_tac[2-] ballI, rule_tac[4] ballI, rule_tac[4] impI) 
-    let ?B = "(\<lambda>s.{(\<chi> i. if i \<in> s then a$i else (a$i + b$i) / 2) ..
-      (\<chi> i. if i \<in> s then (a$i + b$i) / 2 else b$i)}) ` {s. s \<subseteq> UNIV}"
+    let ?B = "(\<lambda>s.{(\<chi>\<chi> i. if i \<in> s then a$$i else (a$$i + b$$i) / 2)::'a ..
+      (\<chi>\<chi> i. if i \<in> s then (a$$i + b$$i) / 2 else b$$i)}) ` {s. s \<subseteq> {..<DIM('a)}}"
     have "?A \<subseteq> ?B" proof case goal1
       then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+) note c_d=this[rule_format]
       have *:"\<And>a b c d. a = c \<Longrightarrow> b = d \<Longrightarrow> {a..b} = {c..d}" by auto
-      show "x\<in>?B" unfolding image_iff apply(rule_tac x="{i. c$i = a$i}" in bexI)
-        unfolding c_d apply(rule * ) unfolding Cart_eq cond_component Cart_lambda_beta
-      proof(rule_tac[1-2] allI) fix i show "c $ i = (if i \<in> {i. c $ i = a $ i} then a $ i else (a $ i + b $ i) / 2)"
-          "d $ i = (if i \<in> {i. c $ i = a $ i} then (a $ i + b $ i) / 2 else b $ i)"
+      show "x\<in>?B" unfolding image_iff apply(rule_tac x="{i. i<DIM('a) \<and> c$$i = a$$i}" in bexI)
+        unfolding c_d apply(rule * ) unfolding euclidean_eq[where 'a='a] apply safe unfolding euclidean_lambda_beta' mem_Collect_eq
+      proof- fix i assume "i<DIM('a)" thus " c $$ i = (if i < DIM('a) \<and> c $$ i = a $$ i then a $$ i else (a $$ i + b $$ i) / 2)"
+          "d $$ i = (if i < DIM('a) \<and> c $$ i = a $$ i then (a $$ i + b $$ i) / 2 else b $$ i)"
           using c_d(2)[of i] ab[THEN spec[where x=i]] by(auto simp add:field_simps)
-      qed auto qed
-    thus "finite ?A" apply(rule finite_subset[of _ ?B]) by auto
+      qed qed
+    thus "finite ?A" apply(rule finite_subset) by auto
     fix s assume "s\<in>?A" then guess c unfolding mem_Collect_eq .. then guess d apply- by(erule exE,(erule conjE)+)
     note c_d=this[rule_format]
-    show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 show ?case 
+    show "P s" unfolding c_d apply(rule as[rule_format]) proof- case goal1 thus ?case 
         using c_d(2)[of i] using ab[THEN spec[where x=i]] by auto qed
     show "\<exists>a b. s = {a..b}" unfolding c_d by auto
     fix t assume "t\<in>?A" then guess e unfolding mem_Collect_eq .. then guess f apply- by(erule exE,(erule conjE)+)
     note e_f=this[rule_format]
     assume "s \<noteq> t" hence "\<not> (c = e \<and> d = f)" unfolding c_d e_f by auto
-    then obtain i where "c$i \<noteq> e$i \<or> d$i \<noteq> f$i" unfolding de_Morgan_conj Cart_eq by auto
-    hence i:"c$i \<noteq> e$i" "d$i \<noteq> f$i" apply- apply(erule_tac[!] disjE)
-    proof- assume "c$i \<noteq> e$i" thus "d$i \<noteq> f$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
-    next   assume "d$i \<noteq> f$i" thus "c$i \<noteq> e$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
+    then obtain i where "c$$i \<noteq> e$$i \<or> d$$i \<noteq> f$$i" and i':"i<DIM('a)" unfolding de_Morgan_conj euclidean_eq[where 'a='a] by auto
+    hence i:"c$$i \<noteq> e$$i" "d$$i \<noteq> f$$i" apply- apply(erule_tac[!] disjE)
+    proof- assume "c$$i \<noteq> e$$i" thus "d$$i \<noteq> f$$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
+    next   assume "d$$i \<noteq> f$$i" thus "c$$i \<noteq> e$$i" using c_d(2)[of i] e_f(2)[of i] by fastsimp
     qed have *:"\<And>s t. (\<And>a. a\<in>s \<Longrightarrow> a\<in>t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}" by auto
     show "interior s \<inter> interior t = {}" unfolding e_f c_d interior_closed_interval proof(rule *)
       fix x assume "x\<in>{c<..<d}" "x\<in>{e<..<f}"
-      hence x:"c$i < d$i" "e$i < f$i" "c$i < f$i" "e$i < d$i" unfolding mem_interval apply-apply(erule_tac[!] x=i in allE)+ by auto
-      show False using c_d(2)[of i] apply- apply(erule_tac disjE)
-      proof(erule_tac[!] conjE) assume as:"c $ i = a $ i" "d $ i = (a $ i + b $ i) / 2"
+      hence x:"c$$i < d$$i" "e$$i < f$$i" "c$$i < f$$i" "e$$i < d$$i" unfolding mem_interval using i'
+        apply-apply(erule_tac[!] x=i in allE)+ by auto
+      show False using c_d(2)[OF i'] apply- apply(erule_tac disjE)
+      proof(erule_tac[!] conjE) assume as:"c $$ i = a $$ i" "d $$ i = (a $$ i + b $$ i) / 2"
         show False using e_f(2)[of i] and i x unfolding as by(fastsimp simp add:field_simps)
-      next assume as:"c $ i = (a $ i + b $ i) / 2" "d $ i = b $ i"
+      next assume as:"c $$ i = (a $$ i + b $$ i) / 2" "d $$ i = b $$ i"
         show False using e_f(2)[of i] and i x unfolding as by(fastsimp simp add:field_simps)
       qed qed qed
   also have "\<Union> ?A = {a..b}" proof(rule set_ext,rule)
     fix x assume "x\<in>\<Union>?A" then guess Y unfolding Union_iff ..
     from this(1) guess c unfolding mem_Collect_eq .. then guess d ..
     note c_d = this[THEN conjunct2,rule_format] `x\<in>Y`[unfolded this[THEN conjunct1]]
-    show "x\<in>{a..b}" unfolding mem_interval proof 
-      fix i show "a $ i \<le> x $ i \<and> x $ i \<le> b $ i"
+    show "x\<in>{a..b}" unfolding mem_interval proof safe
+      fix i assume "i<DIM('a)" thus "a $$ i \<le> x $$ i" "x $$ i \<le> b $$ i"
         using c_d(1)[of i] c_d(2)[unfolded mem_interval,THEN spec[where x=i]] by auto qed
   next fix x assume x:"x\<in>{a..b}"
-    have "\<forall>i. \<exists>c d. (c = a$i \<and> d = (a$i + b$i) / 2 \<or> c = (a$i + b$i) / 2 \<and> d = b$i) \<and> c\<le>x$i \<and> x$i \<le> d"
-      (is "\<forall>i. \<exists>c d. ?P i c d") unfolding mem_interval proof fix i
-      have "?P i (a$i) ((a $ i + b $ i) / 2) \<or> ?P i ((a $ i + b $ i) / 2) (b$i)"
+    have "\<forall>i<DIM('a). \<exists>c d. (c = a$$i \<and> d = (a$$i + b$$i) / 2 \<or> c = (a$$i + b$$i) / 2 \<and> d = b$$i) \<and> c\<le>x$$i \<and> x$$i \<le> d"
+      (is "\<forall>i<DIM('a). \<exists>c d. ?P i c d") unfolding mem_interval proof(rule,rule) fix i
+      have "?P i (a$$i) ((a $$ i + b $$ i) / 2) \<or> ?P i ((a $$ i + b $$ i) / 2) (b$$i)"
         using x[unfolded mem_interval,THEN spec[where x=i]] by auto thus "\<exists>c d. ?P i c d" by blast
-    qed thus "x\<in>\<Union>?A" unfolding Union_iff lambda_skolem unfolding Bex_def mem_Collect_eq
+    qed thus "x\<in>\<Union>?A" unfolding Union_iff unfolding lambda_skolem' unfolding Bex_def mem_Collect_eq
       apply-apply(erule exE)+ apply(rule_tac x="{xa..xaa}" in exI) unfolding mem_interval by auto
   qed finally show False using assms by auto qed
 
-lemma interval_bisection:
-  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "\<not> P {a..b::real^'n}"
+lemma interval_bisection: fixes type::"'a::ordered_euclidean_space"
+  assumes "P {}" "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))" "\<not> P {a..b::'a}"
   obtains x where "x \<in> {a..b}" "\<forall>e>0. \<exists>c d. x \<in> {c..d} \<and> {c..d} \<subseteq> ball x e \<and> {c..d} \<subseteq> {a..b} \<and> ~P({c..d})"
 proof-
-  have "\<forall>x. \<exists>y. \<not> P {fst x..snd x} \<longrightarrow> (\<not> P {fst y..snd y} \<and> (\<forall>i. fst x$i \<le> fst y$i \<and> fst y$i \<le> snd y$i \<and> snd y$i \<le> snd x$i \<and>
-                           2 * (snd y$i - fst y$i) \<le> snd x$i - fst x$i))" proof case goal1 thus ?case proof-
+  have "\<forall>x. \<exists>y. \<not> P {fst x..snd x} \<longrightarrow> (\<not> P {fst y..snd y} \<and>
+    (\<forall>i<DIM('a). fst x$$i \<le> fst y$$i \<and> fst y$$i \<le> snd y$$i \<and> snd y$$i \<le> snd x$$i \<and>
+                           2 * (snd y$$i - fst y$$i) \<le> snd x$$i - fst x$$i))" proof case goal1 thus ?case proof-
       presume "\<not> P {fst x..snd x} \<Longrightarrow> ?thesis"
       thus ?thesis apply(cases "P {fst x..snd x}") by auto
     next assume as:"\<not> P {fst x..snd x}" from interval_bisection_step[of P, OF assms(1-2) as] guess c d . 
@@ -1058,8 +1065,8 @@
     qed qed then guess f apply-apply(drule choice) by(erule exE) note f=this
   def AB \<equiv> "\<lambda>n. (f ^^ n) (a,b)" def A \<equiv> "\<lambda>n. fst(AB n)" and B \<equiv> "\<lambda>n. snd(AB n)" note ab_def = this AB_def
   have "A 0 = a" "B 0 = b" "\<And>n. \<not> P {A(Suc n)..B(Suc n)} \<and>
-    (\<forall>i. A(n)$i \<le> A(Suc n)$i \<and> A(Suc n)$i \<le> B(Suc n)$i \<and> B(Suc n)$i \<le> B(n)$i \<and> 
-    2 * (B(Suc n)$i - A(Suc n)$i) \<le> B(n)$i - A(n)$i)" (is "\<And>n. ?P n")
+    (\<forall>i<DIM('a). A(n)$$i \<le> A(Suc n)$$i \<and> A(Suc n)$$i \<le> B(Suc n)$$i \<and> B(Suc n)$$i \<le> B(n)$$i \<and> 
+    2 * (B(Suc n)$$i - A(Suc n)$$i) \<le> B(n)$$i - A(n)$$i)" (is "\<And>n. ?P n")
   proof- show "A 0 = a" "B 0 = b" unfolding ab_def by auto
     case goal3 note S = ab_def funpow.simps o_def id_apply show ?case
     proof(induct n) case 0 thus ?case unfolding S apply(rule f[rule_format]) using assms(3) by auto
@@ -1067,19 +1074,19 @@
     qed qed note AB = this(1-2) conjunctD2[OF this(3),rule_format]
 
   have interv:"\<And>e. 0 < e \<Longrightarrow> \<exists>n. \<forall>x\<in>{A n..B n}. \<forall>y\<in>{A n..B n}. dist x y < e"
-  proof- case goal1 guess n using real_arch_pow2[of "(setsum (\<lambda>i. b$i - a$i) UNIV) / e"] .. note n=this
+  proof- case goal1 guess n using real_arch_pow2[of "(setsum (\<lambda>i. b$$i - a$$i) {..<DIM('a)}) / e"] .. note n=this
     show ?case apply(rule_tac x=n in exI) proof(rule,rule)
       fix x y assume xy:"x\<in>{A n..B n}" "y\<in>{A n..B n}"
-      have "dist x y \<le> setsum (\<lambda>i. abs((x - y)$i)) UNIV" unfolding dist_norm by(rule norm_le_l1)
-      also have "\<dots> \<le> setsum (\<lambda>i. B n$i - A n$i) UNIV"
-      proof(rule setsum_mono) fix i show "\<bar>(x - y) $ i\<bar> \<le> B n $ i - A n $ i"
-          using xy[unfolded mem_interval,THEN spec[where x=i]]
-          unfolding vector_minus_component by auto qed
-      also have "\<dots> \<le> setsum (\<lambda>i. b$i - a$i) UNIV / 2^n" unfolding setsum_divide_distrib
+      have "dist x y \<le> setsum (\<lambda>i. abs((x - y)$$i)) {..<DIM('a)}" unfolding dist_norm by(rule norm_le_l1)
+      also have "\<dots> \<le> setsum (\<lambda>i. B n$$i - A n$$i) {..<DIM('a)}"
+      proof(rule setsum_mono) fix i show "\<bar>(x - y) $$ i\<bar> \<le> B n $$ i - A n $$ i"
+          using xy[unfolded mem_interval,THEN spec[where x=i]] by auto qed
+      also have "\<dots> \<le> setsum (\<lambda>i. b$$i - a$$i) {..<DIM('a)} / 2^n" unfolding setsum_divide_distrib
       proof(rule setsum_mono) case goal1 thus ?case
         proof(induct n) case 0 thus ?case unfolding AB by auto
-        next case (Suc n) have "B (Suc n) $ i - A (Suc n) $ i \<le> (B n $ i - A n $ i) / 2" using AB(4)[of n i] by auto
-          also have "\<dots> \<le> (b $ i - a $ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case .
+        next case (Suc n) have "B (Suc n) $$ i - A (Suc n) $$ i \<le> (B n $$ i - A n $$ i) / 2"
+            using AB(4)[of i n] using goal1 by auto
+          also have "\<dots> \<le> (b $$ i - a $$ i) / 2 ^ Suc n" using Suc by(auto simp add:field_simps) finally show ?case .
         qed qed
       also have "\<dots> < e" using n using goal1 by(auto simp add:field_simps) finally show "dist x y < e" .
     qed qed
@@ -1088,7 +1095,7 @@
     proof(induct d) case 0 thus ?case by auto
     next case (Suc d) show ?case apply(rule subset_trans[OF _ Suc])
         apply(rule) unfolding mem_interval apply(rule,erule_tac x=i in allE)
-      proof- case goal1 thus ?case using AB(4)[of "m + d" i] by(auto simp add:field_simps)
+      proof- case goal1 thus ?case using AB(4)[of i "m + d"] by(auto simp add:field_simps)
       qed qed } note ABsubset = this 
   have "\<exists>a. \<forall>n. a\<in>{A n..B n}" apply(rule decreasing_closed_nest[rule_format,OF closed_interval _ ABsubset interv])
   proof- fix n show "{A n..B n} \<noteq> {}" apply(cases "0<n") using AB(3)[of "n - 1"] assms(1,3) AB(1-2) by auto qed auto
@@ -1106,7 +1113,7 @@
 subsection {* Cousin's lemma. *}
 
 lemma fine_division_exists: assumes "gauge g" 
-  obtains p where "p tagged_division_of {a..b::real^'n}" "g fine p"
+  obtains p where "p tagged_division_of {a..b::'a::ordered_euclidean_space}" "g fine p"
 proof- presume "\<not> (\<exists>p. p tagged_division_of {a..b} \<and> g fine p) \<Longrightarrow> False"
   then guess p unfolding atomize_not not_not .. thus thesis apply-apply(rule that[of p]) by auto
 next assume as:"\<not> (\<exists>p. p tagged_division_of {a..b} \<and> g fine p)"
@@ -1124,10 +1131,10 @@
 
 subsection {* Basic theorems about integrals. *}
 
-lemma has_integral_unique: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+lemma has_integral_unique: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::real_normed_vector"
   assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2"
 proof(rule ccontr) let ?e = "norm(k1 - k2) / 2" assume as:"k1 \<noteq> k2" hence e:"?e > 0" by auto
-  have lem:"\<And>f::real^'n \<Rightarrow> 'a.  \<And> a b k1 k2.
+  have lem:"\<And>f::'n \<Rightarrow> 'a.  \<And> a b k1 k2.
     (f has_integral k1) ({a..b}) \<Longrightarrow> (f has_integral k2) ({a..b}) \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> False"
   proof- case goal1 let ?e = "norm(k1 - k2) / 2" from goal1(3) have e:"?e > 0" by auto
     guess d1 by(rule has_integralD[OF goal1(1) e]) note d1=this
@@ -1144,7 +1151,7 @@
   assume as:"\<not> (\<exists>a b. i = {a..b})"
   guess B1 by(rule has_integral_altD[OF assms(1) as,OF e]) note B1=this[rule_format]
   guess B2 by(rule has_integral_altD[OF assms(2) as,OF e]) note B2=this[rule_format]
-  have "\<exists>a b::real^'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> {a..b}" apply(rule bounded_subset_closed_interval)
+  have "\<exists>a b::'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> {a..b}" apply(rule bounded_subset_closed_interval)
     using bounded_Un bounded_ball by auto then guess a b apply-by(erule exE)+
   note ab=conjunctD2[OF this[unfolded Un_subset_iff]]
   guess w using B1(2)[OF ab(1)] .. note w=conjunctD2[OF this]
@@ -1159,11 +1166,11 @@
   "(f has_integral y) k \<Longrightarrow> integral k f = y"
   unfolding integral_def apply(rule some_equality) by(auto intro: has_integral_unique) 
 
-lemma has_integral_is_0: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector" 
+lemma has_integral_is_0: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::real_normed_vector" 
   assumes "\<forall>x\<in>s. f x = 0" shows "(f has_integral 0) s"
-proof- have lem:"\<And>a b. \<And>f::real^'n \<Rightarrow> 'a.
+proof- have lem:"\<And>a b. \<And>f::'n \<Rightarrow> 'a.
     (\<forall>x\<in>{a..b}. f(x) = 0) \<Longrightarrow> (f has_integral 0) ({a..b})" unfolding has_integral
-  proof(rule,rule) fix a b e and f::"real^'n \<Rightarrow> 'a"
+  proof(rule,rule) fix a b e and f::"'n \<Rightarrow> 'a"
     assume as:"\<forall>x\<in>{a..b}. f x = 0" "0 < (e::real)"
     show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e)"
       apply(rule_tac x="\<lambda>x. ball x 1" in exI)  apply(rule,rule gaugeI) unfolding centre_in_ball defer apply(rule open_ball)
@@ -1179,20 +1186,20 @@
   assume "\<not> (\<exists>a b. s = {a..b})" thus ?thesis apply(subst has_integral_alt) unfolding if_not_P *
   apply(rule,rule,rule_tac x=1 in exI,rule) defer apply(rule,rule,rule)
   proof- fix e::real and a b assume "e>0"
-    thus "\<exists>z. ((\<lambda>x::real^'n. 0::'a) has_integral z) {a..b} \<and> norm (z - 0) < e"
+    thus "\<exists>z. ((\<lambda>x::'n. 0::'a) has_integral z) {a..b} \<and> norm (z - 0) < e"
       apply(rule_tac x=0 in exI) apply(rule,rule lem) by auto
   qed auto qed
 
-lemma has_integral_0[simp]: "((\<lambda>x::real^'n. 0) has_integral 0) s"
+lemma has_integral_0[simp]: "((\<lambda>x::'n::ordered_euclidean_space. 0) has_integral 0) s"
   apply(rule has_integral_is_0) by auto 
 
 lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) s \<longleftrightarrow> i = 0"
   using has_integral_unique[OF has_integral_0] by auto
 
-lemma has_integral_linear: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+lemma has_integral_linear: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::real_normed_vector"
   assumes "(f has_integral y) s" "bounded_linear h" shows "((h o f) has_integral ((h y))) s"
 proof- interpret bounded_linear h using assms(2) . from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format]
-  have lem:"\<And>f::real^'n \<Rightarrow> 'a. \<And> y a b.
+  have lem:"\<And>f::'n \<Rightarrow> 'a. \<And> y a b.
     (f has_integral y) ({a..b}) \<Longrightarrow> ((h o f) has_integral h(y)) ({a..b})"
   proof(subst has_integral,rule,rule) case goal1
     from pos_bounded guess B .. note B=conjunctD2[OF this,rule_format]
@@ -1231,10 +1238,10 @@
   shows "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral (-k)) s"
   apply(drule_tac c="-1" in has_integral_cmul) by auto
 
-lemma has_integral_add: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector" 
+lemma has_integral_add: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::real_normed_vector" 
   assumes "(f has_integral k) s" "(g has_integral l) s"
   shows "((\<lambda>x. f x + g x) has_integral (k + l)) s"
-proof- have lem:"\<And>f g::real^'n \<Rightarrow> 'a. \<And>a b k l.
+proof- have lem:"\<And>f g::'n \<Rightarrow> 'a. \<And>a b k l.
     (f has_integral k) ({a..b}) \<Longrightarrow> (g has_integral l) ({a..b}) \<Longrightarrow>
      ((\<lambda>x. f(x) + g(x)) has_integral (k + l)) ({a..b})" proof- case goal1
     show ?case unfolding has_integral proof(rule,rule) fix e::real assume e:"e>0" hence *:"e/2>0" by auto
@@ -1259,8 +1266,8 @@
     from has_integral_altD[OF assms(1) as *] guess B1 . note B1=this[rule_format]
     from has_integral_altD[OF assms(2) as *] guess B2 . note B2=this[rule_format]
     show ?case apply(rule_tac x="max B1 B2" in exI) apply(rule,rule min_max.less_supI1,rule B1)
-    proof(rule,rule,rule) fix a b assume "ball 0 (max B1 B2) \<subseteq> {a..b::real^'n}"
-      hence *:"ball 0 B1 \<subseteq> {a..b::real^'n}" "ball 0 B2 \<subseteq> {a..b::real^'n}" by auto
+    proof(rule,rule,rule) fix a b assume "ball 0 (max B1 B2) \<subseteq> {a..b::'n}"
+      hence *:"ball 0 B1 \<subseteq> {a..b::'n}" "ball 0 B2 \<subseteq> {a..b::'n}" by auto
       guess w using B1(2)[OF *(1)] .. note w=conjunctD2[OF this]
       guess z using B2(2)[OF *(2)] .. note z=conjunctD2[OF this]
       have *:"\<And>x. (if x \<in> s then f x + g x else 0) = (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)" by auto
@@ -1273,7 +1280,7 @@
   shows "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_integral (k - l)) s"
   using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding algebra_simps by auto
 
-lemma integral_0: "integral s (\<lambda>x::real^'n. 0::real^'m) = 0"
+lemma integral_0: "integral s (\<lambda>x::'n::ordered_euclidean_space. 0::'m::real_normed_vector) = 0"
   by(rule integral_unique has_integral_0)+
 
 lemma integral_add:
@@ -1326,9 +1333,9 @@
   apply(drule has_integral_linear,assumption,assumption) unfolding has_integral_integral[THEN sym]
   apply(rule integrable_linear) by assumption+
 
-lemma integral_component_eq[simp]: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "f integrable_on s" shows "integral s (\<lambda>x. f x $ k) = integral s f $ k"
-  using integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] .
+lemma integral_component_eq[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+  assumes "f integrable_on s" shows "integral s (\<lambda>x. f x $$ k) = integral s f $$ k"
+  unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] ..
 
 lemma has_integral_setsum:
   assumes "finite t" "\<forall>a\<in>t. ((f a) has_integral (i a)) s"
@@ -1394,8 +1401,8 @@
 lemma integral_empty[simp]: shows "integral {} f = 0"
   apply(rule integral_unique) using has_integral_empty .
 
-lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}" "(f has_integral 0) {a}"
-proof- have *:"{a} = {a..a}" apply(rule set_ext) unfolding mem_interval singleton_iff Cart_eq
+lemma has_integral_refl[intro]: shows "(f has_integral 0) {a..a}" "(f has_integral 0) {a::'a::ordered_euclidean_space}"
+proof- have *:"{a} = {a..a}" apply(rule set_ext) unfolding mem_interval singleton_iff euclidean_eq[where 'a='a]
     apply safe prefer 3 apply(erule_tac x=i in allE) by(auto simp add: field_simps)
   show "(f has_integral 0) {a..a}" "(f has_integral 0) {a}" unfolding *
     apply(rule_tac[!] has_integral_null) unfolding content_eq_0_interior
@@ -1407,7 +1414,8 @@
 
 subsection {* Cauchy-type criterion for integrability. *}
 
-lemma integrable_cauchy: fixes f::"real^'n \<Rightarrow> 'a::{real_normed_vector,complete_space}" 
+(* XXXXXXX *)
+lemma integrable_cauchy: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::{real_normed_vector,complete_space}" 
   shows "f integrable_on {a..b} \<longleftrightarrow>
   (\<forall>e>0.\<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<and> d fine p1 \<and>
                             p2 tagged_division_of {a..b} \<and> d fine p2
@@ -1455,118 +1463,127 @@
 
 subsection {* Additivity of integral on abutting intervals. *}
 
-lemma interval_split:
-  "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
-  "{a..b} \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
-  apply(rule_tac[!] set_ext) unfolding Int_iff mem_interval mem_Collect_eq
-  unfolding Cart_lambda_beta by auto
-
-lemma content_split:
-  "content {a..b::real^'n} = content({a..b} \<inter> {x. x$k \<le> c}) + content({a..b} \<inter> {x. x$k >= c})"
-proof- note simps = interval_split content_closed_interval_cases Cart_lambda_beta vector_le_def
-  { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps by auto }
-  have *:"UNIV = insert k (UNIV - {k})" "\<And>x. finite (UNIV-{x::'n})" "\<And>x. x\<notin>UNIV-{x}" by auto
-  have *:"\<And>X Y Z. (\<Prod>i\<in>UNIV. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>UNIV-{k}. Z i (Y i))"
-    "(\<Prod>i\<in>UNIV. b$i - a$i) = (\<Prod>i\<in>UNIV-{k}. b$i - a$i) * (b$k - a$k)" 
+lemma interval_split: fixes a::"'a::ordered_euclidean_space" assumes "k<DIM('a)" shows
+  "{a..b} \<inter> {x. x$$k \<le> c} = {a .. (\<chi>\<chi> i. if i = k then min (b$$k) c else b$$i)}"
+  "{a..b} \<inter> {x. x$$k \<ge> c} = {(\<chi>\<chi> i. if i = k then max (a$$k) c else a$$i) .. b}"
+  apply(rule_tac[!] set_ext) unfolding Int_iff mem_interval mem_Collect_eq using assms by auto
+
+lemma content_split: fixes a::"'a::ordered_euclidean_space" assumes "k<DIM('a)" shows
+  "content {a..b} = content({a..b} \<inter> {x. x$$k \<le> c}) + content({a..b} \<inter> {x. x$$k >= c})"
+proof- note simps = interval_split[OF assms] content_closed_interval_cases eucl_le[where 'a='a]
+  { presume "a\<le>b \<Longrightarrow> ?thesis" thus ?thesis apply(cases "a\<le>b") unfolding simps using assms by auto }
+  have *:"{..<DIM('a)} = insert k ({..<DIM('a)} - {k})" "\<And>x. finite ({..<DIM('a)}-{x})" "\<And>x. x\<notin>{..<DIM('a)}-{x}"
+    using assms by auto
+  have *:"\<And>X Y Z. (\<Prod>i\<in>{..<DIM('a)}. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>{..<DIM('a)}-{k}. Z i (Y i))"
+    "(\<Prod>i\<in>{..<DIM('a)}. b$$i - a$$i) = (\<Prod>i\<in>{..<DIM('a)}-{k}. b$$i - a$$i) * (b$$k - a$$k)" 
     apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto
-  assume as:"a\<le>b" moreover have "\<And>x. min (b $ k) c = max (a $ k) c
-    \<Longrightarrow> x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)"
+  assume as:"a\<le>b" moreover have "\<And>x. min (b $$ k) c = max (a $$ k) c
+    \<Longrightarrow> x* (b$$k - a$$k) = x*(max (a $$ k) c - a $$ k) + x*(b $$ k - max (a $$ k) c)"
     by  (auto simp add:field_simps)
-  moreover have "\<not> a $ k \<le> c \<Longrightarrow> \<not> c \<le> b $ k \<Longrightarrow> False"
-    unfolding not_le using as[unfolded vector_le_def,rule_format,of k] by auto
-  ultimately show ?thesis 
-    unfolding simps unfolding *(1)[of "\<lambda>i x. b$i - x"] *(1)[of "\<lambda>i x. x - a$i"] *(2) by(auto)
+  moreover have **:"(\<Prod>i<DIM('a). ((\<chi>\<chi> i. if i = k then min (b $$ k) c else b $$ i)::'a) $$ i - a $$ i) = 
+    (\<Prod>i<DIM('a). (if i = k then min (b $$ k) c else b $$ i) - a $$ i)"
+    "(\<Prod>i<DIM('a). b $$ i - ((\<chi>\<chi> i. if i = k then max (a $$ k) c else a $$ i)::'a) $$ i) =
+    (\<Prod>i<DIM('a). b $$ i - (if i = k then max (a $$ k) c else a $$ i))"
+    apply(rule_tac[!] setprod.cong) by auto
+  have "\<not> a $$ k \<le> c \<Longrightarrow> \<not> c \<le> b $$ k \<Longrightarrow> False"
+    unfolding not_le using as[unfolded eucl_le[where 'a='a],rule_format,of k] assms by auto
+  ultimately show ?thesis using assms unfolding simps **
+    unfolding *(1)[of "\<lambda>i x. b$$i - x"] *(1)[of "\<lambda>i x. x - a$$i"] unfolding  *(2) 
+    apply(subst(2) euclidean_lambda_beta''[where 'a='a])
+    apply(subst(3) euclidean_lambda_beta''[where 'a='a]) by auto
 qed
 
-lemma division_split_left_inj:
-  assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2"
-  "k1 \<inter> {x::real^'n. x$k \<le> c} = k2 \<inter> {x. x$k \<le> c}"
-  shows "content(k1 \<inter> {x. x$k \<le> c}) = 0"
+lemma division_split_left_inj: fixes type::"'a::ordered_euclidean_space"
+  assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2" 
+  "k1 \<inter> {x::'a. x$$k \<le> c} = k2 \<inter> {x. x$$k \<le> c}"and k:"k<DIM('a)"
+  shows "content(k1 \<inter> {x. x$$k \<le> c}) = 0"
 proof- note d=division_ofD[OF assms(1)]
-  have *:"\<And>a b::real^'n. \<And> c k. (content({a..b} \<inter> {x. x$k \<le> c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$k \<le> c}) = {})"
-    unfolding interval_split content_eq_0_interior by auto
+  have *:"\<And>a b::'a. \<And> c. (content({a..b} \<inter> {x. x$$k \<le> c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$$k \<le> c}) = {})"
+    unfolding  interval_split[OF k] content_eq_0_interior by auto
   guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
   guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this
   have **:"\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}" by auto
   show ?thesis unfolding uv1 uv2 * apply(rule **[OF d(5)[OF assms(2-4)]])
     defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
-
-lemma division_split_right_inj:
+ 
+lemma division_split_right_inj: fixes type::"'a::ordered_euclidean_space"
   assumes "d division_of i" "k1 \<in> d" "k2 \<in> d"  "k1 \<noteq> k2"
-  "k1 \<inter> {x::real^'n. x$k \<ge> c} = k2 \<inter> {x. x$k \<ge> c}"
-  shows "content(k1 \<inter> {x. x$k \<ge> c}) = 0"
+  "k1 \<inter> {x::'a. x$$k \<ge> c} = k2 \<inter> {x. x$$k \<ge> c}" and k:"k<DIM('a)"
+  shows "content(k1 \<inter> {x. x$$k \<ge> c}) = 0"
 proof- note d=division_ofD[OF assms(1)]
-  have *:"\<And>a b::real^'n. \<And> c k. (content({a..b} \<inter> {x. x$k >= c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$k >= c}) = {})"
-    unfolding interval_split content_eq_0_interior by auto
+  have *:"\<And>a b::'a. \<And> c. (content({a..b} \<inter> {x. x$$k >= c}) = 0 \<longleftrightarrow> interior({a..b} \<inter> {x. x$$k >= c}) = {})"
+    unfolding interval_split[OF k] content_eq_0_interior by auto
   guess u1 v1 using d(4)[OF assms(2)] apply-by(erule exE)+ note uv1=this
   guess u2 v2 using d(4)[OF assms(3)] apply-by(erule exE)+ note uv2=this
   have **:"\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}" by auto
   show ?thesis unfolding uv1 uv2 * apply(rule **[OF d(5)[OF assms(2-4)]])
     defer apply(subst assms(5)[unfolded uv1 uv2]) unfolding uv1 uv2 by auto qed
 
-lemma tagged_division_split_left_inj:
-  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$k \<le> c} = k2 \<inter> {x. x$k \<le> c}" 
-  shows "content(k1 \<inter> {x. x$k \<le> c}) = 0"
+lemma tagged_division_split_left_inj: fixes x1::"'a::ordered_euclidean_space"
+  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$$k \<le> c} = k2 \<inter> {x. x$$k \<le> c}" 
+  and k:"k<DIM('a)"
+  shows "content(k1 \<inter> {x. x$$k \<le> c}) = 0"
 proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
   show ?thesis apply(rule division_split_left_inj[OF division_of_tagged_division[OF assms(1)]])
     apply(rule_tac[1-2] *) using assms(2-) by auto qed
 
-lemma tagged_division_split_right_inj:
-  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$k \<ge> c} = k2 \<inter> {x. x$k \<ge> c}" 
-  shows "content(k1 \<inter> {x. x$k \<ge> c}) = 0"
+lemma tagged_division_split_right_inj: fixes x1::"'a::ordered_euclidean_space"
+  assumes "d tagged_division_of i" "(x1,k1) \<in> d" "(x2,k2) \<in> d" "k1 \<noteq> k2"  "k1 \<inter> {x. x$$k \<ge> c} = k2 \<inter> {x. x$$k \<ge> c}" 
+  and k:"k<DIM('a)"
+  shows "content(k1 \<inter> {x. x$$k \<ge> c}) = 0"
 proof- have *:"\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c" unfolding image_iff apply(rule_tac x="(a,b)" in bexI) by auto
   show ?thesis apply(rule division_split_right_inj[OF division_of_tagged_division[OF assms(1)]])
     apply(rule_tac[1-2] *) using assms(2-) by auto qed
 
-lemma division_split:
-  assumes "p division_of {a..b::real^'n}"
-  shows "{l \<inter> {x. x$k \<le> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$k \<le> c} = {})} division_of ({a..b} \<inter> {x. x$k \<le> c})" (is "?p1 division_of ?I1") and 
-        "{l \<inter> {x. x$k \<ge> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$k \<ge> c} = {})} division_of ({a..b} \<inter> {x. x$k \<ge> c})" (is "?p2 division_of ?I2")
-proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms]
+lemma division_split: fixes a::"'a::ordered_euclidean_space"
+  assumes "p division_of {a..b}" and k:"k<DIM('a)"
+  shows "{l \<inter> {x. x$$k \<le> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$$k \<le> c} = {})} division_of({a..b} \<inter> {x. x$$k \<le> c})" (is "?p1 division_of ?I1") and 
+        "{l \<inter> {x. x$$k \<ge> c} | l. l \<in> p \<and> ~(l \<inter> {x. x$$k \<ge> c} = {})} division_of ({a..b} \<inter> {x. x$$k \<ge> c})" (is "?p2 division_of ?I2")
+proof(rule_tac[!] division_ofI) note p=division_ofD[OF assms(1)]
   show "finite ?p1" "finite ?p2" using p(1) by auto show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2" unfolding p(6)[THEN sym] by auto
   { fix k assume "k\<in>?p1" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this
     guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this
     show "k\<subseteq>?I1" "k \<noteq> {}" "\<exists>a b. k = {a..b}" unfolding l
-      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto
+      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split[OF k]) by auto
     fix k' assume "k'\<in>?p1" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this
     assume "k\<noteq>k'" thus "interior k \<inter> interior k' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto }
   { fix k assume "k\<in>?p2" then guess l unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l=this
     guess u v using p(4)[OF l(2)] apply-by(erule exE)+ note uv=this
     show "k\<subseteq>?I2" "k \<noteq> {}" "\<exists>a b. k = {a..b}" unfolding l
-      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split) by auto
+      using p(2-3)[OF l(2)] l(3) unfolding uv apply- prefer 3 apply(subst interval_split[OF k]) by auto
     fix k' assume "k'\<in>?p2" then guess l' unfolding mem_Collect_eq apply-by(erule exE,(erule conjE)+) note l'=this
     assume "k\<noteq>k'" thus "interior k \<inter> interior k' = {}" unfolding l l' using p(5)[OF l(2) l'(2)] by auto }
 qed
 
-lemma has_integral_split: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral i) ({a..b} \<inter> {x. x$k \<le> c})"  "(f has_integral j) ({a..b} \<inter> {x. x$k \<ge> c})"
+lemma has_integral_split: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "(f has_integral i) ({a..b} \<inter> {x. x$$k \<le> c})"  "(f has_integral j) ({a..b} \<inter> {x. x$$k \<ge> c})" and k:"k<DIM('a)"
   shows "(f has_integral (i + j)) ({a..b})"
 proof(unfold has_integral,rule,rule) case goal1 hence e:"e/2>0" by auto
-  guess d1 using has_integralD[OF assms(1)[unfolded interval_split] e] . note d1=this[unfolded interval_split[THEN sym]]
-  guess d2 using has_integralD[OF assms(2)[unfolded interval_split] e] . note d2=this[unfolded interval_split[THEN sym]]
-  let ?d = "\<lambda>x. if x$k = c then (d1 x \<inter> d2 x) else ball x (abs(x$k - c)) \<inter> d1 x \<inter> d2 x"
+  guess d1 using has_integralD[OF assms(1)[unfolded interval_split[OF k]] e] . note d1=this[unfolded interval_split[THEN sym,OF k]]
+  guess d2 using has_integralD[OF assms(2)[unfolded interval_split[OF k]] e] . note d2=this[unfolded interval_split[THEN sym,OF k]]
+  let ?d = "\<lambda>x. if x$$k = c then (d1 x \<inter> d2 x) else ball x (abs(x$$k - c)) \<inter> d1 x \<inter> d2 x"
   show ?case apply(rule_tac x="?d" in exI,rule) defer apply(rule,rule,(erule conjE)+)
   proof- show "gauge ?d" using d1(1) d2(1) unfolding gauge_def by auto
     fix p assume "p tagged_division_of {a..b}" "?d fine p" note p = this tagged_division_ofD[OF this(1)]
-    have lem0:"\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$k \<le> c} = {}) \<Longrightarrow> x$k \<le> c"
-         "\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$k \<ge> c} = {}) \<Longrightarrow> x$k \<ge> c"
+    have lem0:"\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$$k \<le> c} = {}) \<Longrightarrow> x$$k \<le> c"
+         "\<And>x kk. (x,kk) \<in> p \<Longrightarrow> ~(kk \<inter> {x. x$$k \<ge> c} = {}) \<Longrightarrow> x$$k \<ge> c"
     proof- fix x kk assume as:"(x,kk)\<in>p"
-      show "~(kk \<inter> {x. x$k \<le> c} = {}) \<Longrightarrow> x$k \<le> c"
+      show "~(kk \<inter> {x. x$$k \<le> c} = {}) \<Longrightarrow> x$$k \<le> c"
       proof(rule ccontr) case goal1
-        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $ k - c\<bar>"
+        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $$ k - c\<bar>"
           using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
-        hence "\<exists>y. y \<in> ball x \<bar>x $ k - c\<bar> \<inter> {x. x $ k \<le> c}" using goal1(1) by blast 
-        then guess y .. hence "\<bar>x $ k - y $ k\<bar> < \<bar>x $ k - c\<bar>" "y$k \<le> c" apply-apply(rule le_less_trans)
-          using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:dist_norm)
+        hence "\<exists>y. y \<in> ball x \<bar>x $$ k - c\<bar> \<inter> {x. x $$ k \<le> c}" using goal1(1) by blast 
+        then guess y .. hence "\<bar>x $$ k - y $$ k\<bar> < \<bar>x $$ k - c\<bar>" "y$$k \<le> c" apply-apply(rule le_less_trans)
+          using component_le_norm[of "x - y" k] by(auto simp add:dist_norm)
         thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
       qed
-      show "~(kk \<inter> {x. x$k \<ge> c} = {}) \<Longrightarrow> x$k \<ge> c"
+      show "~(kk \<inter> {x. x$$k \<ge> c} = {}) \<Longrightarrow> x$$k \<ge> c"
       proof(rule ccontr) case goal1
-        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $ k - c\<bar>"
+        from this(2)[unfolded not_le] have "kk \<subseteq> ball x \<bar>x $$ k - c\<bar>"
           using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
-        hence "\<exists>y. y \<in> ball x \<bar>x $ k - c\<bar> \<inter> {x. x $ k \<ge> c}" using goal1(1) by blast 
-        then guess y .. hence "\<bar>x $ k - y $ k\<bar> < \<bar>x $ k - c\<bar>" "y$k \<ge> c" apply-apply(rule le_less_trans)
-          using component_le_norm[of "x - y" k,unfolded vector_minus_component] by(auto simp add:dist_norm)
+        hence "\<exists>y. y \<in> ball x \<bar>x $$ k - c\<bar> \<inter> {x. x $$ k \<ge> c}" using goal1(1) by blast 
+        then guess y .. hence "\<bar>x $$ k - y $$ k\<bar> < \<bar>x $$ k - c\<bar>" "y$$k \<ge> c" apply-apply(rule le_less_trans)
+          using component_le_norm[of "x - y" k] by(auto simp add:dist_norm)
         thus False using goal1(2)[unfolded not_le] by(auto simp add:field_simps)
       qed
     qed
@@ -1574,11 +1591,11 @@
     have lem1: "\<And>f P Q. (\<forall>x k. (x,k) \<in> {(x,f k) | x k. P x k} \<longrightarrow> Q x k) \<longleftrightarrow> (\<forall>x k. P x k \<longrightarrow> Q x (f k))" by auto
     have lem2: "\<And>f s P f. finite s \<Longrightarrow> finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}"
     proof- case goal1 thus ?case apply-apply(rule finite_subset[of _ "(\<lambda>(x,k). (x,f k)) ` s"]) by auto qed
-    have lem3: "\<And>g::(real ^ 'n \<Rightarrow> bool) \<Rightarrow> real ^ 'n \<Rightarrow> bool. finite p \<Longrightarrow>
+    have lem3: "\<And>g::('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool. finite p \<Longrightarrow>
       setsum (\<lambda>(x,k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> ~(g k = {})}
                = setsum (\<lambda>(x,k). content k *\<^sub>R f x) ((\<lambda>(x,k). (x,g k)) ` p)"
       apply(rule setsum_mono_zero_left) prefer 3
-    proof fix g::"(real ^ 'n \<Rightarrow> bool) \<Rightarrow> real ^ 'n \<Rightarrow> bool" and i::"(real^'n) \<times> ((real^'n) set)"
+    proof fix g::"('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" and i::"('a) \<times> (('a) set)"
       assume "i \<in> (\<lambda>(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
       then obtain x k where xk:"i=(x,g k)" "(x,k)\<in>p" "(x,g k) \<notin> {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}" by auto
       have "content (g k) = 0" using xk using content_empty by auto
@@ -1586,17 +1603,17 @@
     qed auto
     have lem4:"\<And>g. (\<lambda>(x,l). content (g l) *\<^sub>R f x) = (\<lambda>(x,l). content l *\<^sub>R f x) o (\<lambda>(x,l). (x,g l))" apply(rule ext) by auto
 
-    let ?M1 = "{(x,kk \<inter> {x. x$k \<le> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$k \<le> c} \<noteq> {}}"
+    let ?M1 = "{(x,kk \<inter> {x. x$$k \<le> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$$k \<le> c} \<noteq> {}}"
     have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2" apply(rule d1(2),rule tagged_division_ofI)
       apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
-    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = {a..b} \<inter> {x. x$k \<le> c}" unfolding p(8)[THEN sym] by auto
+    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = {a..b} \<inter> {x. x$$k \<le> c}" unfolding p(8)[THEN sym] by auto
       fix x l assume xl:"(x,l)\<in>?M1"
       then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note xl'=this
       have "l' \<subseteq> d1 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
       thus "l \<subseteq> d1 x" unfolding xl' by auto
-      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $ k \<le> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
+      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $$ k \<le> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
         using lem0(1)[OF xl'(3-4)] by auto
-      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k])
+      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[OF k,where c=c])
       fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M1"
       then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note yr'=this
       assume as:"(x,l) \<noteq> (y,r)" show "interior l \<inter> interior r = {}"
@@ -1606,17 +1623,17 @@
         thus ?thesis using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
       qed qed moreover
 
-    let ?M2 = "{(x,kk \<inter> {x. x$k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$k \<ge> c} \<noteq> {}}" 
+    let ?M2 = "{(x,kk \<inter> {x. x$$k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x$$k \<ge> c} \<noteq> {}}" 
     have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2" apply(rule d2(2),rule tagged_division_ofI)
       apply(rule lem2 p(3))+ prefer 6 apply(rule fineI)
-    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = {a..b} \<inter> {x. x$k \<ge> c}" unfolding p(8)[THEN sym] by auto
+    proof- show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = {a..b} \<inter> {x. x$$k \<ge> c}" unfolding p(8)[THEN sym] by auto
       fix x l assume xl:"(x,l)\<in>?M2"
       then guess x' l' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note xl'=this
       have "l' \<subseteq> d2 x'" apply(rule order_trans[OF fineD[OF p(2) xl'(3)]]) by auto
       thus "l \<subseteq> d2 x" unfolding xl' by auto
-      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $ k \<ge> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
+      show "x\<in>l" "l \<subseteq> {a..b} \<inter> {x. x $$ k \<ge> c}" unfolding xl' using p(4-6)[OF xl'(3)] using xl'(4)
         using lem0(2)[OF xl'(3-4)] by auto
-      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[where c=c and k=k])
+      show  "\<exists>a b. l = {a..b}" unfolding xl' using p(6)[OF xl'(3)] by(fastsimp simp add: interval_split[OF k, where c=c])
       fix y r let ?goal = "interior l \<inter> interior r = {}" assume yr:"(y,r)\<in>?M2"
       then guess y' r' unfolding mem_Collect_eq apply- unfolding Pair_eq apply((erule exE)+,(erule conjE)+) .  note yr'=this
       assume as:"(x,l) \<noteq> (y,r)" show "interior l \<inter> interior r = {}"
@@ -1628,108 +1645,121 @@
 
     have "norm (((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2"
       apply- apply(rule norm_triangle_lt) by auto
-    also { have *:"\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'a) = 0" using scaleR_zero_left by auto
+    also { have *:"\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0" using scaleR_zero_left by auto
       have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)
        = (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)" by auto
-      also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x $ k \<le> c}) *\<^sub>R f x) + (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x $ k}) *\<^sub>R f x) - (i + j)"
+      also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x $$ k \<le> c}) *\<^sub>R f x) +
+        (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x $$ k}) *\<^sub>R f x) - (i + j)"
         unfolding lem3[OF p(3)] apply(subst setsum_reindex_nonzero[OF p(3)]) defer apply(subst setsum_reindex_nonzero[OF p(3)])
         defer unfolding lem4[THEN sym] apply(rule refl) unfolding split_paired_all split_conv apply(rule_tac[!] *)
-      proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) by auto
-      next case   goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) by auto
+      proof- case goal1 thus ?case apply- apply(rule tagged_division_split_left_inj [OF p(1), of a b aa ba]) using k by auto
+      next case   goal2 thus ?case apply- apply(rule tagged_division_split_right_inj[OF p(1), of a b aa ba]) using k by auto
       qed also note setsum_addf[THEN sym]
-      also have *:"\<And>x. x\<in>p \<Longrightarrow> (\<lambda>(x, ka). content (ka \<inter> {x. x $ k \<le> c}) *\<^sub>R f x) x + (\<lambda>(x, ka). content (ka \<inter> {x. c \<le> x $ k}) *\<^sub>R f x) x
+      also have *:"\<And>x. x\<in>p \<Longrightarrow> (\<lambda>(x, ka). content (ka \<inter> {x. x $$ k \<le> c}) *\<^sub>R f x) x + (\<lambda>(x, ka). content (ka \<inter> {x. c \<le> x $$ k}) *\<^sub>R f x) x
         = (\<lambda>(x,ka). content ka *\<^sub>R f x) x" unfolding split_paired_all split_conv
       proof- fix a b assume "(a,b) \<in> p" from p(6)[OF this] guess u v apply-by(erule exE)+ note uv=this
-        thus "content (b \<inter> {x. x $ k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x $ k}) *\<^sub>R f a = content b *\<^sub>R f a"
-          unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[of u v k c] by auto
+        thus "content (b \<inter> {x. x $$ k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x $$ k}) *\<^sub>R f a = content b *\<^sub>R f a"
+          unfolding scaleR_left_distrib[THEN sym] unfolding uv content_split[OF k,of u v c] by auto
       qed note setsum_cong2[OF this]
-      finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x $ k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x $ k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
-        ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x $ k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x $ k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
+      finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x $$ k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x $$ k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
+        ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x $$ k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x $$ k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
         (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)" by auto }
     finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed
 
+(*lemma has_integral_split_cart: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral i) ({a..b} \<inter> {x. x$k \<le> c})"  "(f has_integral j) ({a..b} \<inter> {x. x$k \<ge> c})"
+  shows "(f has_integral (i + j)) ({a..b})" *)
+
 subsection {* A sort of converse, integrability on subintervals. *}
 
-lemma tagged_division_union_interval:
-  assumes "p1 tagged_division_of ({a..b} \<inter> {x::real^'n. x$k \<le> (c::real)})"  "p2 tagged_division_of ({a..b} \<inter> {x. x$k \<ge> c})"
+lemma tagged_division_union_interval: fixes a::"'a::ordered_euclidean_space"
+  assumes "p1 tagged_division_of ({a..b} \<inter> {x. x$$k \<le> (c::real)})"  "p2 tagged_division_of ({a..b} \<inter> {x. x$$k \<ge> c})"
+  and k:"k<DIM('a)"
   shows "(p1 \<union> p2) tagged_division_of ({a..b})"
-proof- have *:"{a..b} = ({a..b} \<inter> {x. x$k \<le> c}) \<union> ({a..b} \<inter> {x. x$k \<ge> c})" by auto
-  show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms])
-    unfolding interval_split interior_closed_interval
-    by(auto simp add: vector_less_def elim!:allE[where x=k]) qed
-
-lemma has_integral_separate_sides: fixes f::"real^'m \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral i) ({a..b})" "e>0"
-  obtains d where "gauge d" "(\<forall>p1 p2. p1 tagged_division_of ({a..b} \<inter> {x. x$k \<le> c}) \<and> d fine p1 \<and>
-                                p2 tagged_division_of ({a..b} \<inter> {x. x$k \<ge> c}) \<and> d fine p2
+proof- have *:"{a..b} = ({a..b} \<inter> {x. x$$k \<le> c}) \<union> ({a..b} \<inter> {x. x$$k \<ge> c})" by auto
+  show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms(1-2)])
+    unfolding interval_split[OF k] interior_closed_interval using k
+    by(auto simp add: eucl_less[where 'a='a] elim!:allE[where x=k]) qed
+
+lemma has_integral_separate_sides: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "(f has_integral i) ({a..b})" "e>0" and k:"k<DIM('a)"
+  obtains d where "gauge d" "(\<forall>p1 p2. p1 tagged_division_of ({a..b} \<inter> {x. x$$k \<le> c}) \<and> d fine p1 \<and>
+                                p2 tagged_division_of ({a..b} \<inter> {x. x$$k \<ge> c}) \<and> d fine p2
                                 \<longrightarrow> norm((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 +
                                           setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e)"
-proof- guess d using has_integralD[OF assms] . note d=this
+proof- guess d using has_integralD[OF assms(1-2)] . note d=this
   show ?thesis apply(rule that[of d]) apply(rule d) apply(rule,rule,rule,(erule conjE)+)
-  proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this
-                   assume "p2 tagged_division_of {a..b} \<inter> {x. c \<le> x $ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this
+  proof- fix p1 p2 assume "p1 tagged_division_of {a..b} \<inter> {x. x $$ k \<le> c}" "d fine p1" note p1=tagged_division_ofD[OF this(1)] this
+                   assume "p2 tagged_division_of {a..b} \<inter> {x. c \<le> x $$ k}" "d fine p2" note p2=tagged_division_ofD[OF this(1)] this
     note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this
     have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) = norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
       apply(subst setsum_Un_zero) apply(rule p1 p2)+ apply(rule) unfolding split_paired_all split_conv
     proof- fix a b assume ab:"(a,b) \<in> p1 \<inter> p2"
       have "(a,b) \<in> p1" using ab by auto from p1(4)[OF this] guess u v apply-by(erule exE)+ note uv =this
-      have "b \<subseteq> {x. x$k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastsimp
-      moreover have "interior {x. x $ k = c} = {}" 
-      proof(rule ccontr) case goal1 then obtain x where x:"x\<in>interior {x. x$k = c}" by auto
+      have "b \<subseteq> {x. x$$k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastsimp
+      moreover have "interior {x::'a. x $$ k = c} = {}" 
+      proof(rule ccontr) case goal1 then obtain x where x:"x\<in>interior {x::'a. x$$k = c}" by auto
         then guess e unfolding mem_interior .. note e=this
-        have x:"x$k = c" using x interior_subset by fastsimp
-        have *:"\<And>i. \<bar>(x - (x + (\<chi> i. if i = k then e / 2 else 0))) $ i\<bar> = (if i = k then e/2 else 0)" using e by auto
-        have "x + (\<chi> i. if i = k then e/2 else 0) \<in> ball x e" unfolding mem_ball dist_norm 
-          apply(rule le_less_trans[OF norm_le_l1]) unfolding * 
-          unfolding setsum_delta[OF finite_UNIV] using e by auto 
-        hence "x + (\<chi> i. if i = k then e/2 else 0) \<in> {x. x$k = c}" using e by auto
-        thus False unfolding mem_Collect_eq using e x by auto
+        have x:"x$$k = c" using x interior_subset by fastsimp
+        have *:"\<And>i. i<DIM('a) \<Longrightarrow> \<bar>(x - (x + (\<chi>\<chi> i. if i = k then e / 2 else 0))) $$ i\<bar>
+          = (if i = k then e/2 else 0)" using e by auto
+        have "(\<Sum>i<DIM('a). \<bar>(x - (x + (\<chi>\<chi> i. if i = k then e / 2 else 0))) $$ i\<bar>) =
+          (\<Sum>i<DIM('a). (if i = k then e / 2 else 0))" apply(rule setsum_cong2) apply(subst *) by auto
+        also have "... < e" apply(subst setsum_delta) using e by auto 
+        finally have "x + (\<chi>\<chi> i. if i = k then e/2 else 0) \<in> ball x e" unfolding mem_ball dist_norm
+          by(rule le_less_trans[OF norm_le_l1])
+        hence "x + (\<chi>\<chi> i. if i = k then e/2 else 0) \<in> {x. x$$k = c}" using e by auto
+        thus False unfolding mem_Collect_eq using e x k by auto
       qed ultimately have "content b = 0" unfolding uv content_eq_0_interior apply-apply(drule subset_interior) by auto
       thus "content b *\<^sub>R f a = 0" by auto
     qed auto
-    also have "\<dots> < e" by(rule d(2) p12 fine_union p1 p2)+
+    also have "\<dots> < e" by(rule k d(2) p12 fine_union p1 p2)+
     finally show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" . qed qed
 
-lemma integrable_split[intro]: fixes f::"real^'n \<Rightarrow> 'a::{real_normed_vector,complete_space}" assumes "f integrable_on {a..b}"
-  shows "f integrable_on ({a..b} \<inter> {x. x$k \<le> c})" (is ?t1) and "f integrable_on ({a..b} \<inter> {x. x$k \<ge> c})" (is ?t2) 
-proof- guess y using assms unfolding integrable_on_def .. note y=this
-  def b' \<equiv> "(\<chi> i. if i = k then min (b$k) c else b$i)::real^'n"
-  and a' \<equiv> "(\<chi> i. if i = k then max (a$k) c else a$i)::real^'n"
-  show ?t1 ?t2 unfolding interval_split integrable_cauchy unfolding interval_split[THEN sym]
+lemma integrable_split[intro]: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
+  assumes "f integrable_on {a..b}" and k:"k<DIM('a)"
+  shows "f integrable_on ({a..b} \<inter> {x. x$$k \<le> c})" (is ?t1) and "f integrable_on ({a..b} \<inter> {x. x$$k \<ge> c})" (is ?t2) 
+proof- guess y using assms(1) unfolding integrable_on_def .. note y=this
+  def b' \<equiv> "(\<chi>\<chi> i. if i = k then min (b$$k) c else b$$i)::'a"
+  and a' \<equiv> "(\<chi>\<chi> i. if i = k then max (a$$k) c else a$$i)::'a"
+  show ?t1 ?t2 unfolding interval_split[OF k] integrable_cauchy unfolding interval_split[THEN sym,OF k]
   proof(rule_tac[!] allI impI)+ fix e::real assume "e>0" hence "e/2>0" by auto
-    from has_integral_separate_sides[OF y this,of k c] guess d . note d=this[rule_format]
-    let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<inter> A \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> A \<and> d fine p2 \<longrightarrow>
-                              norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
-    show "?P {x. x $ k \<le> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
-    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<le> c} \<and> d fine p2"
+    from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format]
+    let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of {a..b} \<inter> A \<and> d fine p1
+      \<and> p2 tagged_division_of {a..b} \<inter> A \<and> d fine p2 \<longrightarrow>
+      norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
+    show "?P {x. x $$ k \<le> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
+    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $$ k \<le> c} \<and> d fine p1
+        \<and> p2 tagged_division_of {a..b} \<inter> {x. x $$ k \<le> c} \<and> d fine p2"
       show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
       proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
         show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
-          using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
+          using as unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
           using p using assms by(auto simp add:algebra_simps)
       qed qed  
-    show "?P {x. x $ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
-    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p2"
+    show "?P {x. x $$ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
+    proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $$ k \<ge> c} \<and> d fine p1
+        \<and> p2 tagged_division_of {a..b} \<inter> {x. x $$ k \<ge> c} \<and> d fine p2"
       show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
       proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
         show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
-          using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
+          using as unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
           using p using assms by(auto simp add:algebra_simps) qed qed qed qed
 
 subsection {* Generalized notion of additivity. *}
 
 definition "neutral opp = (SOME x. \<forall>y. opp x y = y \<and> opp y x = y)"
 
-definition operative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ((real^'n) set \<Rightarrow> 'a) \<Rightarrow> bool" where
+definition operative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (('b::ordered_euclidean_space) set \<Rightarrow> 'a) \<Rightarrow> bool" where
   "operative opp f \<equiv> 
     (\<forall>a b. content {a..b} = 0 \<longrightarrow> f {a..b} = neutral(opp)) \<and>
-    (\<forall>a b c k. f({a..b}) =
-                   opp (f({a..b} \<inter> {x. x$k \<le> c}))
-                       (f({a..b} \<inter> {x. x$k \<ge> c})))"
-
-lemma operativeD[dest]: assumes "operative opp f"
-  shows "\<And>a b. content {a..b} = 0 \<Longrightarrow> f {a..b} = neutral(opp)"
-  "\<And>a b c k. f({a..b}) = opp (f({a..b} \<inter> {x. x$k \<le> c})) (f({a..b} \<inter> {x. x$k \<ge> c}))"
+    (\<forall>a b c. \<forall>k<DIM('b). f({a..b}) =
+                   opp (f({a..b} \<inter> {x. x$$k \<le> c}))
+                       (f({a..b} \<inter> {x. x$$k \<ge> c})))"
+
+lemma operativeD[dest]: fixes type::"'a::ordered_euclidean_space"  assumes "operative opp f"
+  shows "\<And>a b. content {a..b} = 0 \<Longrightarrow> f {a..b::'a} = neutral(opp)"
+  "\<And>a b c k. k<DIM('a) \<Longrightarrow> f({a..b}) = opp (f({a..b} \<inter> {x. x$$k \<le> c})) (f({a..b} \<inter> {x. x$$k \<ge> c}))"
   using assms unfolding operative_def by auto
 
 lemma operative_trivial:
@@ -1835,15 +1865,15 @@
 proof(induct s) case empty thus ?case using assms by auto
 next case (insert x F) show ?case apply(subst iterate_insert) prefer 3 apply(subst if_not_P)
     defer unfolding insert(3) lifted.simps apply rule using assms insert by auto qed
-
 subsection {* Two key instances of additivity. *}
 
 lemma neutral_add[simp]:
   "neutral op + = (0::_::comm_monoid_add)" unfolding neutral_def 
   apply(rule some_equality) defer apply(erule_tac x=0 in allE) by auto
 
-lemma operative_content[intro]: "operative (op +) content"
-  unfolding operative_def content_split[THEN sym] neutral_add by auto
+lemma operative_content[intro]: "operative (op +) content" 
+  unfolding operative_def neutral_add apply safe 
+  unfolding content_split[THEN sym] ..
 
 lemma neutral_monoid: "neutral ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a) = 0"
   by (rule neutral_add) (* FIXME: duplicate *)
@@ -1852,110 +1882,119 @@
   shows "monoidal ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a)"
   unfolding monoidal_def neutral_monoid by(auto simp add: algebra_simps) 
 
-lemma operative_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma operative_integral: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
   shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
   unfolding operative_def unfolding neutral_lifted[OF monoidal_monoid] neutral_add
-  apply(rule,rule,rule,rule) defer apply(rule allI)+
-proof- fix a b c k show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) =
-              lifted op + (if f integrable_on {a..b} \<inter> {x. x $ k \<le> c} then Some (integral ({a..b} \<inter> {x. x $ k \<le> c}) f) else None)
-               (if f integrable_on {a..b} \<inter> {x. c \<le> x $ k} then Some (integral ({a..b} \<inter> {x. c \<le> x $ k}) f) else None)"
+  apply(rule,rule,rule,rule) defer apply(rule allI impI)+
+proof- fix a b c k assume k:"k<DIM('a)" show "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) =
+    lifted op + (if f integrable_on {a..b} \<inter> {x. x $$ k \<le> c} then Some (integral ({a..b} \<inter> {x. x $$ k \<le> c}) f) else None)
+    (if f integrable_on {a..b} \<inter> {x. c \<le> x $$ k} then Some (integral ({a..b} \<inter> {x. c \<le> x $$ k}) f) else None)"
   proof(cases "f integrable_on {a..b}") 
-    case True show ?thesis unfolding if_P[OF True]
-      unfolding if_P[OF integrable_split(1)[OF True]] if_P[OF integrable_split(2)[OF True]]
-      unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split) 
-      apply(rule_tac[!] integrable_integral integrable_split)+ using True by assumption+
-  next case False have "(\<not> (f integrable_on {a..b} \<inter> {x. x $ k \<le> c})) \<or> (\<not> ( f integrable_on {a..b} \<inter> {x. c \<le> x $ k}))"
+    case True show ?thesis unfolding if_P[OF True] using k apply-
+      unfolding if_P[OF integrable_split(1)[OF True]] unfolding if_P[OF integrable_split(2)[OF True]]
+      unfolding lifted.simps option.inject apply(rule integral_unique) apply(rule has_integral_split[OF _ _ k]) 
+      apply(rule_tac[!] integrable_integral integrable_split)+ using True k by auto
+  next case False have "(\<not> (f integrable_on {a..b} \<inter> {x. x $$ k \<le> c})) \<or> (\<not> ( f integrable_on {a..b} \<inter> {x. c \<le> x $$ k}))"
     proof(rule ccontr) case goal1 hence "f integrable_on {a..b}" apply- unfolding integrable_on_def
-        apply(rule_tac x="integral ({a..b} \<inter> {x. x $ k \<le> c}) f + integral ({a..b} \<inter> {x. x $ k \<ge> c}) f" in exI)
-        apply(rule has_integral_split) apply(rule_tac[!] integrable_integral) by auto
+        apply(rule_tac x="integral ({a..b} \<inter> {x. x $$ k \<le> c}) f + integral ({a..b} \<inter> {x. x $$ k \<ge> c}) f" in exI)
+        apply(rule has_integral_split[OF _ _ k]) apply(rule_tac[!] integrable_integral) by auto
       thus False using False by auto
     qed thus ?thesis using False by auto 
   qed next 
-  fix a b assume as:"content {a..b::real^'n} = 0"
+  fix a b assume as:"content {a..b::'a} = 0"
   thus "(if f integrable_on {a..b} then Some (integral {a..b} f) else None) = Some 0"
     unfolding if_P[OF integrable_on_null[OF as]] using has_integral_null_eq[OF as] by auto qed
 
 subsection {* Points of division of a partition. *}
 
-definition "division_points (k::(real^'n) set) d = 
-    {(j,x). (interval_lowerbound k)$j < x \<and> x < (interval_upperbound k)$j \<and>
-           (\<exists>i\<in>d. (interval_lowerbound i)$j = x \<or> (interval_upperbound i)$j = x)}"
-
-lemma division_points_finite: assumes "d division_of i"
-  shows "finite (division_points i d)"
+definition "division_points (k::('a::ordered_euclidean_space) set) d = 
+    {(j,x). j<DIM('a) \<and> (interval_lowerbound k)$$j < x \<and> x < (interval_upperbound k)$$j \<and>
+           (\<exists>i\<in>d. (interval_lowerbound i)$$j = x \<or> (interval_upperbound i)$$j = x)}"
+
+lemma division_points_finite: fixes i::"('a::ordered_euclidean_space) set"
+  assumes "d division_of i" shows "finite (division_points i d)"
 proof- note assm = division_ofD[OF assms]
-  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)$j < x \<and> x < (interval_upperbound i)$j \<and>
-           (\<exists>i\<in>d. (interval_lowerbound i)$j = x \<or> (interval_upperbound i)$j = x)}"
-  have *:"division_points i d = \<Union>(?M ` UNIV)"
+  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)$$j < x \<and> x < (interval_upperbound i)$$j \<and>
+           (\<exists>i\<in>d. (interval_lowerbound i)$$j = x \<or> (interval_upperbound i)$$j = x)}"
+  have *:"division_points i d = \<Union>(?M ` {..<DIM('a)})"
     unfolding division_points_def by auto
   show ?thesis unfolding * using assm by auto qed
 
-lemma division_points_subset:
-  assumes "d division_of {a..b}" "\<forall>i. a$i < b$i"  "a$k < c" "c < b$k"
-  shows "division_points ({a..b} \<inter> {x. x$k \<le> c}) {l \<inter> {x. x$k \<le> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$k \<le> c} = {})}
+lemma division_points_subset: fixes a::"'a::ordered_euclidean_space"
+  assumes "d division_of {a..b}" "\<forall>i<DIM('a). a$$i < b$$i"  "a$$k < c" "c < b$$k" and k:"k<DIM('a)"
+  shows "division_points ({a..b} \<inter> {x. x$$k \<le> c}) {l \<inter> {x. x$$k \<le> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$$k \<le> c} = {})}
                   \<subseteq> division_points ({a..b}) d" (is ?t1) and
-        "division_points ({a..b} \<inter> {x. x$k \<ge> c}) {l \<inter> {x. x$k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$k \<ge> c} = {})}
+        "division_points ({a..b} \<inter> {x. x$$k \<ge> c}) {l \<inter> {x. x$$k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x$$k \<ge> c} = {})}
                   \<subseteq> division_points ({a..b}) d" (is ?t2)
 proof- note assm = division_ofD[OF assms(1)]
-  have *:"\<forall>i. a$i \<le> b$i"   "\<forall>i. a$i \<le> (\<chi> i. if i = k then min (b $ k) c else b $ i) $ i"
-    "\<forall>i. (\<chi> i. if i = k then max (a $ k) c else a $ i) $ i \<le> b$i"  "min (b $ k) c = c" "max (a $ k) c = c"
+  have *:"\<forall>i<DIM('a). a$$i \<le> b$$i"   "\<forall>i<DIM('a). a$$i \<le> ((\<chi>\<chi> i. if i = k then min (b $$ k) c else b $$ i)::'a) $$ i"
+    "\<forall>i<DIM('a). ((\<chi>\<chi> i. if i = k then max (a $$ k) c else a $$ i)::'a) $$ i \<le> b$$i"  "min (b $$ k) c = c" "max (a $$ k) c = c"
     using assms using less_imp_le by auto
-  show ?t1 unfolding division_points_def interval_split[of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding *
-    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq apply(erule exE conjE)+
-  proof- fix i l x assume as:"a $ fst x < snd x" "snd x < (if fst x = k then c else b $ fst x)"
-      "interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x"  "i = l \<inter> {x. x $ k \<le> c}" "l \<in> d" "l \<inter> {x. x $ k \<le> c} \<noteq> {}"
+  show ?t1 unfolding division_points_def interval_split[OF k, of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding *
+    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+
+    unfolding mem_Collect_eq apply(erule exE conjE)+ unfolding euclidean_lambda_beta'
+  proof- fix i l x assume as:"a $$ fst x < snd x" "snd x < (if fst x = k then c else b $$ fst x)"
+      "interval_lowerbound i $$ fst x = snd x \<or> interval_upperbound i $$ fst x = snd x"
+      "i = l \<inter> {x. x $$ k \<le> c}" "l \<in> d" "l \<inter> {x. x $$ k \<le> c} \<noteq> {}" and fstx:"fst x <DIM('a)"
     from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
-    have *:"\<forall>i. u $ i \<le> (\<chi> i. if i = k then min (v $ k) c else v $ i) $ i" using as(6) unfolding l interval_split interval_ne_empty as .
-    have **:"\<forall>i. u$i \<le> v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
-    show "a $ fst x < snd x \<and> snd x < b $ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x)"
-      using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply-
+    have *:"\<forall>i<DIM('a). u $$ i \<le> ((\<chi>\<chi> i. if i = k then min (v $$ k) c else v $$ i)::'a) $$ i"
+      using as(6) unfolding l interval_split[OF k] interval_ne_empty as .
+    have **:"\<forall>i<DIM('a). u$$i \<le> v$$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
+    show "fst x <DIM('a) \<and> a $$ fst x < snd x \<and> snd x < b $$ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $$ fst x = snd x
+      \<or> interval_upperbound i $$ fst x = snd x)" apply(rule,rule fstx)
+      using as(1-3,5) unfolding l interval_split[OF k] interval_ne_empty as interval_bounds[OF *] apply-
       apply(rule,assumption,rule) defer apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
-      apply(case_tac[!] "fst x = k") using assms by auto
+      apply(case_tac[!] "fst x = k") using assms fstx apply- unfolding euclidean_lambda_beta by auto
   qed
-  show ?t2 unfolding division_points_def interval_split[of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] Cart_lambda_beta unfolding *
-    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+ unfolding mem_Collect_eq apply(erule exE conjE)+
-  proof- fix i l x assume as:"(if fst x = k then c else a $ fst x) < snd x" "snd x < b $ fst x" "interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x"
-      "i = l \<inter> {x. c \<le> x $ k}" "l \<in> d" "l \<inter> {x. c \<le> x $ k} \<noteq> {}"
+  show ?t2 unfolding division_points_def interval_split[OF k, of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)] unfolding *
+    unfolding subset_eq apply(rule) unfolding mem_Collect_eq split_beta apply(erule bexE conjE)+
+    unfolding mem_Collect_eq apply(erule exE conjE)+ unfolding euclidean_lambda_beta' apply(rule,assumption)
+  proof- fix i l x assume as:"(if fst x = k then c else a $$ fst x) < snd x" "snd x < b $$ fst x"
+      "interval_lowerbound i $$ fst x = snd x \<or> interval_upperbound i $$ fst x = snd x" 
+      "i = l \<inter> {x. c \<le> x $$ k}" "l \<in> d" "l \<inter> {x. c \<le> x $$ k} \<noteq> {}" and fstx:"fst x < DIM('a)"
     from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
-    have *:"\<forall>i. (\<chi> i. if i = k then max (u $ k) c else u $ i) $ i \<le> v $ i" using as(6) unfolding l interval_split interval_ne_empty as .
-    have **:"\<forall>i. u$i \<le> v$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
-    show "a $ fst x < snd x \<and> snd x < b $ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $ fst x = snd x \<or> interval_upperbound i $ fst x = snd x)"
-      using as(1-3,5) unfolding l interval_split interval_ne_empty as interval_bounds[OF *] Cart_lambda_beta apply-
+    have *:"\<forall>i<DIM('a). ((\<chi>\<chi> i. if i = k then max (u $$ k) c else u $$ i)::'a) $$ i \<le> v $$ i"
+      using as(6) unfolding l interval_split[OF k] interval_ne_empty as .
+    have **:"\<forall>i<DIM('a). u$$i \<le> v$$i" using l using as(6) unfolding interval_ne_empty[THEN sym] by auto
+    show "a $$ fst x < snd x \<and> snd x < b $$ fst x \<and> (\<exists>i\<in>d. interval_lowerbound i $$ fst x = snd x \<or>
+      interval_upperbound i $$ fst x = snd x)"
+      using as(1-3,5) unfolding l interval_split[OF k] interval_ne_empty as interval_bounds[OF *] apply-
       apply rule defer apply(rule,assumption) apply(rule_tac x="{u..v}" in bexI) unfolding interval_bounds[OF **]
-      apply(case_tac[!] "fst x = k") using assms by auto qed qed
-
-lemma division_points_psubset:
-  assumes "d division_of {a..b}"  "\<forall>i. a$i < b$i"  "a$k < c" "c < b$k"
-  "l \<in> d" "interval_lowerbound l$k = c \<or> interval_upperbound l$k = c"
-  shows "division_points ({a..b} \<inter> {x. x$k \<le> c}) {l \<inter> {x. x$k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x$k \<le> c} \<noteq> {}} \<subset> division_points ({a..b}) d" (is "?D1 \<subset> ?D") 
-        "division_points ({a..b} \<inter> {x. x$k \<ge> c}) {l \<inter> {x. x$k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x$k \<ge> c} \<noteq> {}} \<subset> division_points ({a..b}) d" (is "?D2 \<subset> ?D") 
-proof- have ab:"\<forall>i. a$i \<le> b$i" using assms(2) by(auto intro!:less_imp_le)
+      apply(case_tac[!] "fst x = k") using assms fstx apply-  by(auto simp add:euclidean_lambda_beta'[OF k]) qed qed
+
+lemma division_points_psubset: fixes a::"'a::ordered_euclidean_space"
+  assumes "d division_of {a..b}"  "\<forall>i<DIM('a). a$$i < b$$i"  "a$$k < c" "c < b$$k"
+  "l \<in> d" "interval_lowerbound l$$k = c \<or> interval_upperbound l$$k = c" and k:"k<DIM('a)"
+  shows "division_points ({a..b} \<inter> {x. x$$k \<le> c}) {l \<inter> {x. x$$k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x$$k \<le> c} \<noteq> {}}
+              \<subset> division_points ({a..b}) d" (is "?D1 \<subset> ?D") 
+        "division_points ({a..b} \<inter> {x. x$$k \<ge> c}) {l \<inter> {x. x$$k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x$$k \<ge> c} \<noteq> {}}
+              \<subset> division_points ({a..b}) d" (is "?D2 \<subset> ?D") 
+proof- have ab:"\<forall>i<DIM('a). a$$i \<le> b$$i" using assms(2) by(auto intro!:less_imp_le)
   guess u v using division_ofD(4)[OF assms(1,5)] apply-by(erule exE)+ note l=this
-  have uv:"\<forall>i. u$i \<le> v$i" "\<forall>i. a$i \<le> u$i \<and> v$i \<le> b$i" using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty
+  have uv:"\<forall>i<DIM('a). u$$i \<le> v$$i" "\<forall>i<DIM('a). a$$i \<le> u$$i \<and> v$$i \<le> b$$i"
+    using division_ofD(2,2,3)[OF assms(1,5)] unfolding l interval_ne_empty
     unfolding subset_eq apply- defer apply(erule_tac x=u in ballE, erule_tac x=v in ballE) unfolding mem_interval by auto
-  have *:"interval_upperbound ({a..b} \<inter> {x. x $ k \<le> interval_upperbound l $ k}) $ k = interval_upperbound l $ k"
-         "interval_upperbound ({a..b} \<inter> {x. x $ k \<le> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
-    unfolding interval_split apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
-    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab by auto
+  have *:"interval_upperbound ({a..b} \<inter> {x. x $$ k \<le> interval_upperbound l $$ k}) $$ k = interval_upperbound l $$ k"
+         "interval_upperbound ({a..b} \<inter> {x. x $$ k \<le> interval_lowerbound l $$ k}) $$ k = interval_lowerbound l $$ k"
+    unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
+    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto
   have "\<exists>x. x \<in> ?D - ?D1" using assms(2-) apply-apply(erule disjE)
-    apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
-    apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
-    unfolding division_points_def unfolding interval_bounds[OF ab]
-    apply auto unfolding * by auto
-  thus "?D1 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto
-
-  have *:"interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
-         "interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_upperbound l $ k}) $ k = interval_upperbound l $ k"
-    unfolding interval_split apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
-    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab by auto
+    apply(rule_tac x="(k,(interval_lowerbound l)$$k)" in exI) defer
+    apply(rule_tac x="(k,(interval_upperbound l)$$k)" in exI)
+    unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*) 
+  thus "?D1 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) using k by auto
+
+  have *:"interval_lowerbound ({a..b} \<inter> {x. x $$ k \<ge> interval_lowerbound l $$ k}) $$ k = interval_lowerbound l $$ k"
+         "interval_lowerbound ({a..b} \<inter> {x. x $$ k \<ge> interval_upperbound l $$ k}) $$ k = interval_upperbound l $$ k"
+    unfolding interval_split[OF k] apply(subst interval_bounds) prefer 3 apply(subst interval_bounds)
+    unfolding l interval_bounds[OF uv(1)] using uv[rule_format,of k] ab k by auto
   have "\<exists>x. x \<in> ?D - ?D2" using assms(2-) apply-apply(erule disjE)
-    apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
-    apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
-    unfolding division_points_def unfolding interval_bounds[OF ab]
-    apply auto unfolding * by auto
-  thus "?D2 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto qed
+    apply(rule_tac x="(k,(interval_lowerbound l)$$k)" in exI) defer
+    apply(rule_tac x="(k,(interval_upperbound l)$$k)" in exI)
+    unfolding division_points_def unfolding interval_bounds[OF ab] by(auto simp add:*) 
+  thus "?D2 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4) k]) by auto qed
 
 subsection {* Preservation by divisions and tagged divisions. *}
 
@@ -2036,7 +2075,7 @@
 
 lemma nonempty_witness: assumes "s \<noteq> {}" obtains x where "x \<in> s" using assms by auto
 
-lemma operative_division: fixes f::"(real^'n) set \<Rightarrow> 'a"
+lemma operative_division: fixes f::"('a::ordered_euclidean_space) set \<Rightarrow> 'b"
   assumes "monoidal opp" "operative opp f" "d division_of {a..b}"
   shows "iterate opp d f = f {a..b}"
 proof- def C \<equiv> "card (division_points {a..b} d)" thus ?thesis using assms
@@ -2052,82 +2091,82 @@
             using operativeD(1)[OF assms(2)] x by auto
         qed qed }
     assume "content {a..b} \<noteq> 0" note ab = this[unfolded content_lt_nz[THEN sym] content_pos_lt_eq]
-    hence ab':"\<forall>i. a$i \<le> b$i" by (auto intro!: less_imp_le) show ?case 
+    hence ab':"\<forall>i<DIM('a). a$$i \<le> b$$i" by (auto intro!: less_imp_le) show ?case 
     proof(cases "division_points {a..b} d = {}")
       case True have d':"\<forall>i\<in>d. \<exists>u v. i = {u..v} \<and>
-        (\<forall>j. u$j = a$j \<and> v$j = a$j \<or> u$j = b$j \<and> v$j = b$j \<or> u$j = a$j \<and> v$j = b$j)"
+        (\<forall>j<DIM('a). u$$j = a$$j \<and> v$$j = a$$j \<or> u$$j = b$$j \<and> v$$j = b$$j \<or> u$$j = a$$j \<and> v$$j = b$$j)"
         unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule)
-        apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) apply(rule)
-      proof- fix u v j assume as:"{u..v} \<in> d" note division_ofD(3)[OF goal1(4) this]
-        hence uv:"\<forall>i. u$i \<le> v$i" "u$j \<le> v$j" unfolding interval_ne_empty by auto
-        have *:"\<And>p r Q. p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> (Q {u..v})" using as by auto
-        have "(j, u$j) \<notin> division_points {a..b} d"
-          "(j, v$j) \<notin> division_points {a..b} d" using True by auto
+        apply(rule_tac x=a in exI,rule_tac x=b in exI) apply(rule,rule refl) apply(rule,rule)
+      proof- fix u v j assume j:"j<DIM('a)" assume as:"{u..v} \<in> d" note division_ofD(3)[OF goal1(4) this]
+        hence uv:"\<forall>i<DIM('a). u$$i \<le> v$$i" "u$$j \<le> v$$j" using j unfolding interval_ne_empty by auto
+        have *:"\<And>p r Q. \<not> j<DIM('a) \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> (Q {u..v})" using as j by auto
+        have "(j, u$$j) \<notin> division_points {a..b} d"
+          "(j, v$$j) \<notin> division_points {a..b} d" using True by auto
         note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps]
         note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
-        moreover have "a$j \<le> u$j" "v$j \<le> b$j" using division_ofD(2,2,3)[OF goal1(4) as] 
+        moreover have "a$$j \<le> u$$j" "v$$j \<le> b$$j" using division_ofD(2,2,3)[OF goal1(4) as] 
           unfolding subset_eq apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE)
-          unfolding interval_ne_empty mem_interval by auto
-        ultimately show "u$j = a$j \<and> v$j = a$j \<or> u$j = b$j \<and> v$j = b$j \<or> u$j = a$j \<and> v$j = b$j"
-          unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) by auto
+          unfolding interval_ne_empty mem_interval using j by auto
+        ultimately show "u$$j = a$$j \<and> v$$j = a$$j \<or> u$$j = b$$j \<and> v$$j = b$$j \<or> u$$j = a$$j \<and> v$$j = b$$j"
+          unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by auto
       qed have "(1/2) *\<^sub>R (a+b) \<in> {a..b}" unfolding mem_interval using ab by(auto intro!:less_imp_le)
       note this[unfolded division_ofD(6)[OF goal1(4),THEN sym] Union_iff]
       then guess i .. note i=this guess u v using d'[rule_format,OF i(1)] apply-by(erule exE conjE)+ note uv=this
       have "{a..b} \<in> d"
       proof- { presume "i = {a..b}" thus ?thesis using i by auto }
         { presume "u = a" "v = b" thus "i = {a..b}" using uv by auto }
-        show "u = a" "v = b" unfolding Cart_eq
-        proof(rule_tac[!] allI) fix j note i(2)[unfolded uv mem_interval,rule_format,of j]
-          thus "u $ j = a $ j" "v $ j = b $ j" using uv(2)[rule_format,of j] by auto
+        show "u = a" "v = b" unfolding euclidean_eq[where 'a='a]
+        proof(safe) fix j assume j:"j<DIM('a)" note i(2)[unfolded uv mem_interval,rule_format,of j]
+          thus "u $$ j = a $$ j" "v $$ j = b $$ j" using uv(2)[rule_format,of j] j by auto
         qed qed
       hence *:"d = insert {a..b} (d - {{a..b}})" by auto
       have "iterate opp (d - {{a..b}}) f = neutral opp" apply(rule iterate_eq_neutral[OF goal1(2)])
       proof fix x assume x:"x \<in> d - {{a..b}}" hence "x\<in>d" by auto note d'[rule_format,OF this]
         then guess u v apply-by(erule exE conjE)+ note uv=this
         have "u\<noteq>a \<or> v\<noteq>b" using x[unfolded uv] by auto  
-        then obtain j where "u$j \<noteq> a$j \<or> v$j \<noteq> b$j" unfolding Cart_eq by auto
-        hence "u$j = v$j" using uv(2)[rule_format,of j] by auto
-        hence "content {u..v} = 0"  unfolding content_eq_0 apply(rule_tac x=j in exI) by auto
+        then obtain j where "u$$j \<noteq> a$$j \<or> v$$j \<noteq> b$$j" and j:"j<DIM('a)" unfolding euclidean_eq[where 'a='a] by auto
+        hence "u$$j = v$$j" using uv(2)[rule_format,OF j] by auto
+        hence "content {u..v} = 0"  unfolding content_eq_0 apply(rule_tac x=j in exI) using j by auto
         thus "f x = neutral opp" unfolding uv(1) by(rule operativeD(1)[OF goal1(3)])
       qed thus "iterate opp d f = f {a..b}" apply-apply(subst *) 
         apply(subst iterate_insert[OF goal1(2)]) using goal1(2,4) by auto
     next case False hence "\<exists>x. x\<in>division_points {a..b} d" by auto
       then guess k c unfolding split_paired_Ex apply- unfolding division_points_def mem_Collect_eq split_conv
-        by(erule exE conjE)+ note kc=this[unfolded interval_bounds[OF ab']]
+        by(erule exE conjE)+ note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']]
       from this(3) guess j .. note j=this
-      def d1 \<equiv> "{l \<inter> {x. x$k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x$k \<le> c} \<noteq> {}}"
-      def d2 \<equiv> "{l \<inter> {x. x$k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x$k \<ge> c} \<noteq> {}}"
-      def cb \<equiv> "(\<chi> i. if i = k then c else b$i)" and ca \<equiv> "(\<chi> i. if i = k then c else a$i)"
+      def d1 \<equiv> "{l \<inter> {x. x$$k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x$$k \<le> c} \<noteq> {}}"
+      def d2 \<equiv> "{l \<inter> {x. x$$k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x$$k \<ge> c} \<noteq> {}}"
+      def cb \<equiv> "(\<chi>\<chi> i. if i = k then c else b$$i)::'a" and ca \<equiv> "(\<chi>\<chi> i. if i = k then c else a$$i)::'a"
       note division_points_psubset[OF goal1(4) ab kc(1-2) j]
       note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
-      hence *:"(iterate opp d1 f) = f ({a..b} \<inter> {x. x$k \<le> c})" "(iterate opp d2 f) = f ({a..b} \<inter> {x. x$k \<ge> c})"
-        apply- unfolding interval_split apply(rule_tac[!] goal1(1)[rule_format])
+      hence *:"(iterate opp d1 f) = f ({a..b} \<inter> {x. x$$k \<le> c})" "(iterate opp d2 f) = f ({a..b} \<inter> {x. x$$k \<ge> c})"
+        apply- unfolding interval_split[OF kc(4)] apply(rule_tac[!] goal1(1)[rule_format])
         using division_split[OF goal1(4), where k=k and c=c]
-        unfolding interval_split d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono
-        using goal1(2-3) using division_points_finite[OF goal1(4)] by auto
+        unfolding interval_split[OF kc(4)] d1_def[symmetric] d2_def[symmetric] unfolding goal1(2) Suc_le_mono
+        using goal1(2-3) using division_points_finite[OF goal1(4)] using kc(4) by auto
       have "f {a..b} = opp (iterate opp d1 f) (iterate opp d2 f)" (is "_ = ?prev")
-        unfolding * apply(rule operativeD(2)) using goal1(3) .
-      also have "iterate opp d1 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$k \<le> c}))"
+        unfolding * apply(rule operativeD(2)) using goal1(3) using kc(4) by auto 
+      also have "iterate opp d1 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$$k \<le> c}))"
         unfolding d1_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
         unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
         unfolding empty_as_interval[THEN sym] apply(rule content_empty)
-      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. x $ k \<le> c} = y \<inter> {x. x $ k \<le> c}" "l \<noteq> y" 
+      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. x $$ k \<le> c} = y \<inter> {x. x $$ k \<le> c}" "l \<noteq> y" 
         from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
-        show "f (l \<inter> {x. x $ k \<le> c}) = neutral opp" unfolding l interval_split
-          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_left_inj)
-          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+
-      qed also have "iterate opp d2 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$k \<ge> c}))"
+        show "f (l \<inter> {x. x $$ k \<le> c}) = neutral opp" unfolding l interval_split[OF kc(4)] 
+          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym,OF kc(4)] apply(rule division_split_left_inj)
+          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule kc(4) as)+
+      qed also have "iterate opp d2 f = iterate opp d (\<lambda>l. f(l \<inter> {x. x$$k \<ge> c}))"
         unfolding d2_def apply(rule iterate_nonzero_image_lemma[unfolded o_def])
         unfolding empty_as_interval apply(rule goal1 division_of_finite operativeD[OF goal1(3)])+
         unfolding empty_as_interval[THEN sym] apply(rule content_empty)
-      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x $ k} = y \<inter> {x. c \<le> x $ k}" "l \<noteq> y" 
+      proof(rule,rule,rule,erule conjE) fix l y assume as:"l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x $$ k} = y \<inter> {x. c \<le> x $$ k}" "l \<noteq> y" 
         from division_ofD(4)[OF goal1(4) this(1)] guess u v apply-by(erule exE)+ note l=this
-        show "f (l \<inter> {x. x $ k \<ge> c}) = neutral opp" unfolding l interval_split
-          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym] apply(rule division_split_right_inj)
-          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as)+
-      qed also have *:"\<forall>x\<in>d. f x = opp (f (x \<inter> {x. x $ k \<le> c})) (f (x \<inter> {x. c \<le> x $ k}))"
-        unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) .
-      have "opp (iterate opp d (\<lambda>l. f (l \<inter> {x. x $ k \<le> c}))) (iterate opp d (\<lambda>l. f (l \<inter> {x. c \<le> x $ k})))
+        show "f (l \<inter> {x. x $$ k \<ge> c}) = neutral opp" unfolding l interval_split[OF kc(4)] 
+          apply(rule operativeD(1) goal1)+ unfolding interval_split[THEN sym,OF kc(4)] apply(rule division_split_right_inj)
+          apply(rule goal1) unfolding l[THEN sym] apply(rule as(1),rule as(2)) by(rule as kc(4))+
+      qed also have *:"\<forall>x\<in>d. f x = opp (f (x \<inter> {x. x $$ k \<le> c})) (f (x \<inter> {x. c \<le> x $$ k}))"
+        unfolding forall_in_division[OF goal1(4)] apply(rule,rule,rule,rule operativeD(2)) using goal1(3) kc by auto 
+      have "opp (iterate opp d (\<lambda>l. f (l \<inter> {x. x $$ k \<le> c}))) (iterate opp d (\<lambda>l. f (l \<inter> {x. c \<le> x $$ k})))
         = iterate opp d f" apply(subst(3) iterate_eq[OF _ *[rule_format]]) prefer 3
         apply(rule iterate_op[THEN sym]) using goal1 by auto
       finally show ?thesis by auto
@@ -2186,7 +2225,7 @@
 subsection {* Finally, the integral of a constant *}
 
 lemma has_integral_const[intro]:
-  "((\<lambda>x. c) has_integral (content({a..b::real^'n}) *\<^sub>R c)) ({a..b})"
+  "((\<lambda>x. c) has_integral (content({a..b::'a::ordered_euclidean_space}) *\<^sub>R c)) ({a..b})"
   unfolding has_integral apply(rule,rule,rule_tac x="\<lambda>x. ball x 1" in exI)
   apply(rule,rule gauge_trivial)apply(rule,rule,erule conjE)
   unfolding split_def apply(subst scaleR_left.setsum[THEN sym, unfolded o_def])
@@ -2232,7 +2271,7 @@
   apply(rule order_trans[OF _ rsum_bound[OF assms]]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
   unfolding setsum_subtractf[THEN sym] apply(rule setsum_cong2) unfolding scaleR.diff_right by auto
 
-lemma has_integral_bound: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+lemma has_integral_bound: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::real_normed_vector"
   assumes "0 \<le> B" "(f has_integral i) ({a..b})" "\<forall>x\<in>{a..b}. norm(f x) \<le> B"
   shows "norm i \<le> B * content {a..b}"
 proof- let ?P = "content {a..b} > 0" { presume "?P \<Longrightarrow> ?thesis"
@@ -2252,53 +2291,52 @@
 
 subsection {* Similar theorems about relationship among components. *}
 
-lemma rsum_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. (f x)$i \<le> (g x)$i"
-  shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)$i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)$i"
-  unfolding setsum_component apply(rule setsum_mono)
-  apply(rule mp) defer apply assumption unfolding split_paired_all apply rule unfolding split_conv
+lemma rsum_component_le: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. (f x)$$i \<le> (g x)$$i"
+  shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)$$i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)$$i"
+  unfolding  euclidean_component.setsum apply(rule setsum_mono) apply safe
 proof- fix a b assume ab:"(a,b) \<in> p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab]
   from this(3) guess u v apply-by(erule exE)+ note b=this
-  show "(content b *\<^sub>R f a) $ i \<le> (content b *\<^sub>R g a) $ i" unfolding b
-    unfolding Cart_nth.scaleR real_scaleR_def apply(rule mult_left_mono)
+  show "(content b *\<^sub>R f a) $$ i \<le> (content b *\<^sub>R g a) $$ i" unfolding b
+    unfolding euclidean_simps real_scaleR_def apply(rule mult_left_mono)
     defer apply(rule content_pos_le,rule assms(2)[rule_format]) using assm by auto qed
 
-lemma has_integral_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "(f has_integral i) s" "(g has_integral j) s"  "\<forall>x\<in>s. (f x)$k \<le> (g x)$k"
-  shows "i$k \<le> j$k"
-proof- have lem:"\<And>a b g i j. \<And>f::real^'n \<Rightarrow> real^'m. (f has_integral i) ({a..b}) \<Longrightarrow> 
-    (g has_integral j) ({a..b}) \<Longrightarrow> \<forall>x\<in>{a..b}. (f x)$k \<le> (g x)$k \<Longrightarrow> i$k \<le> j$k"
-  proof(rule ccontr) case goal1 hence *:"0 < (i$k - j$k) / 3" by auto
+lemma has_integral_component_le: fixes f g::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "(f has_integral i) s" "(g has_integral j) s"  "\<forall>x\<in>s. (f x)$$k \<le> (g x)$$k"
+  shows "i$$k \<le> j$$k"
+proof- have lem:"\<And>a b i (j::'b). \<And>g f::'a \<Rightarrow> 'b. (f has_integral i) ({a..b}) \<Longrightarrow> 
+    (g has_integral j) ({a..b}) \<Longrightarrow> \<forall>x\<in>{a..b}. (f x)$$k \<le> (g x)$$k \<Longrightarrow> i$$k \<le> j$$k"
+  proof(rule ccontr) case goal1 hence *:"0 < (i$$k - j$$k) / 3" by auto
     guess d1 using goal1(1)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d1=this[rule_format]
     guess d2 using goal1(2)[unfolded has_integral,rule_format,OF *] apply-by(erule exE conjE)+ note d2=this[rule_format]
     guess p using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter .
-    note p = this(1) conjunctD2[OF this(2)]  note le_less_trans[OF component_le_norm, of _ _ k]
+    note p = this(1) conjunctD2[OF this(2)]  note le_less_trans[OF component_le_norm, of _ _ k] term g
     note this[OF d1(2)[OF conjI[OF p(1,2)]]] this[OF d2(2)[OF conjI[OF p(1,3)]]]
-    thus False unfolding Cart_nth.diff using rsum_component_le[OF p(1) goal1(3)] by smt
+    thus False unfolding euclidean_simps using rsum_component_le[OF p(1) goal1(3)] apply simp by smt
   qed let ?P = "\<exists>a b. s = {a..b}"
   { presume "\<not> ?P \<Longrightarrow> ?thesis" thus ?thesis proof(cases ?P)
       case True then guess a b apply-by(erule exE)+ note s=this
       show ?thesis apply(rule lem) using assms[unfolded s] by auto
     qed auto } assume as:"\<not> ?P"
   { presume "\<not> ?thesis \<Longrightarrow> False" thus ?thesis by auto }
-  assume "\<not> i$k \<le> j$k" hence ij:"(i$k - j$k) / 3 > 0" by auto
+  assume "\<not> i$$k \<le> j$$k" hence ij:"(i$$k - j$$k) / 3 > 0" by auto
   note has_integral_altD[OF _ as this] from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
-  have "bounded (ball 0 B1 \<union> ball (0::real^'n) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+
+  have "bounded (ball 0 B1 \<union> ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+
   from bounded_subset_closed_interval[OF this] guess a b apply- by(erule exE)+
   note ab = conjunctD2[OF this[unfolded Un_subset_iff]]
   guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this]
   guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this]
-  have *:"\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False" by smt(*SMTSMT*)
+  have *:"\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False" by smt
   note le_less_trans[OF component_le_norm[of _ k]] note this[OF w1(2)] this[OF w2(2)] moreover
-  have "w1$k \<le> w2$k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately
-  show False unfolding Cart_nth.diff by(rule *) qed
-
-lemma integral_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "f integrable_on s" "g integrable_on s"  "\<forall>x\<in>s. (f x)$k \<le> (g x)$k"
-  shows "(integral s f)$k \<le> (integral s g)$k"
+  have "w1$$k \<le> w2$$k" apply(rule lem[OF w1(1) w2(1)]) using assms by auto ultimately
+  show False unfolding euclidean_simps by(rule *) qed
+
+lemma integral_component_le: fixes g f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "f integrable_on s" "g integrable_on s"  "\<forall>x\<in>s. (f x)$$k \<le> (g x)$$k"
+  shows "(integral s f)$$k \<le> (integral s g)$$k"
   apply(rule has_integral_component_le) using integrable_integral assms by auto
 
-lemma has_integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
+(*lemma has_integral_dest_vec1_le: fixes f::"'a::ordered_euclidean_space \<Rightarrow> real^1"
   assumes "(f has_integral i) s"  "(g has_integral j) s" "\<forall>x\<in>s. f x \<le> g x"
   shows "dest_vec1 i \<le> dest_vec1 j" apply(rule has_integral_component_le[OF assms(1-2)])
   using assms(3) unfolding vector_le_def by auto
@@ -2306,61 +2344,57 @@
 lemma integral_dest_vec1_le: fixes f::"real^'n \<Rightarrow> real^1"
   assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. f x \<le> g x"
   shows "dest_vec1(integral s f) \<le> dest_vec1(integral s g)"
-  apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto
-
-lemma has_integral_component_nonneg: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> (f x)$k" shows "0 \<le> i$k"
-  using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2) by auto
-
-lemma integral_component_nonneg: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> (f x)$k" shows "0 \<le> (integral s f)$k"
+  apply(rule has_integral_dest_vec1_le) apply(rule_tac[1-2] integrable_integral) using assms by auto*)
+
+lemma has_integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> (f x)$$k" shows "0 \<le> i$$k" 
+  using has_integral_component_le[OF has_integral_0 assms(1)] using assms(2-) by auto
+
+lemma integral_component_nonneg: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> (f x)$$k" shows "0 \<le> (integral s f)$$k"
   apply(rule has_integral_component_nonneg) using assms by auto
 
-lemma has_integral_dest_vec1_nonneg: fixes f::"real^'n \<Rightarrow> real^1"
+(*lemma has_integral_dest_vec1_nonneg: fixes f::"real^'n \<Rightarrow> real^1"
   assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> i"
   using has_integral_component_nonneg[OF assms(1), of 1]
   using assms(2) unfolding vector_le_def by auto
 
 lemma integral_dest_vec1_nonneg: fixes f::"real^'n \<Rightarrow> real^1"
   assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> integral s f"
-  apply(rule has_integral_dest_vec1_nonneg) using assms by auto
-
-lemma has_integral_component_neg: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "(f has_integral i) s" "\<forall>x\<in>s. (f x)$k \<le> 0" shows "i$k \<le> 0"
-  using has_integral_component_le[OF assms(1) has_integral_0] assms(2) by auto
-
-lemma has_integral_dest_vec1_neg: fixes f::"real^'n \<Rightarrow> real^1"
+  apply(rule has_integral_dest_vec1_nonneg) using assms by auto*)
+
+lemma has_integral_component_neg: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
+  assumes "(f has_integral i) s" "\<forall>x\<in>s. (f x)$$k \<le> 0"shows "i$$k \<le> 0" 
+  using has_integral_component_le[OF assms(1) has_integral_0] assms(2-) by auto
+
+(*lemma has_integral_dest_vec1_neg: fixes f::"real^'n \<Rightarrow> real^1"
   assumes "(f has_integral i) s" "\<forall>x\<in>s. f x \<le> 0" shows "i \<le> 0"
-  using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto
-
-lemma has_integral_component_lbound:
-  assumes "(f has_integral i) {a..b}"  "\<forall>x\<in>{a..b}. B \<le> f(x)$k" shows "B * content {a..b} \<le> i$k"
-  using has_integral_component_le[OF has_integral_const assms(1),of "(\<chi> i. B)" k] assms(2)
-  unfolding Cart_lambda_beta vector_scaleR_component by(auto simp add:field_simps)
-
-lemma has_integral_component_ubound: 
-  assumes "(f has_integral i) {a..b}" "\<forall>x\<in>{a..b}. f x$k \<le> B"
-  shows "i$k \<le> B * content({a..b})"
-  using has_integral_component_le[OF assms(1) has_integral_const, of k "vec B"]
-  unfolding vec_component Cart_nth.scaleR using assms(2) by(auto simp add:field_simps)
-
-lemma integral_component_lbound:
-  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)$k"
-  shows "B * content({a..b}) \<le> (integral({a..b}) f)$k"
+  using has_integral_component_neg[OF assms(1),of 1] using assms(2) by auto*)
+
+lemma has_integral_component_lbound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
+  assumes "(f has_integral i) {a..b}"  "\<forall>x\<in>{a..b}. B \<le> f(x)$$k" "k<DIM('b)" shows "B * content {a..b} \<le> i$$k"
+  using has_integral_component_le[OF has_integral_const assms(1),of "(\<chi>\<chi> i. B)::'b" k] assms(2-)
+  unfolding euclidean_simps euclidean_lambda_beta'[OF assms(3)] by(auto simp add:field_simps)
+
+lemma has_integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
+  assumes "(f has_integral i) {a..b}" "\<forall>x\<in>{a..b}. f x$$k \<le> B" "k<DIM('b)"
+  shows "i$$k \<le> B * content({a..b})"
+  using has_integral_component_le[OF assms(1) has_integral_const, of k "\<chi>\<chi> i. B"]
+  unfolding euclidean_simps euclidean_lambda_beta'[OF assms(3)] using assms(2) by(auto simp add:field_simps)
+
+lemma integral_component_lbound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
+  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. B \<le> f(x)$$k" "k<DIM('b)"
+  shows "B * content({a..b}) \<le> (integral({a..b}) f)$$k"
   apply(rule has_integral_component_lbound) using assms unfolding has_integral_integral by auto
 
-lemma integral_component_ubound:
-  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. f(x)$k \<le> B"
-  shows "(integral({a..b}) f)$k \<le> B * content({a..b})"
+lemma integral_component_ubound: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
+  assumes "f integrable_on {a..b}" "\<forall>x\<in>{a..b}. f(x)$$k \<le> B" "k<DIM('b)" 
+  shows "(integral({a..b}) f)$$k \<le> B * content({a..b})"
   apply(rule has_integral_component_ubound) using assms unfolding has_integral_integral by auto
 
 subsection {* Uniform limit of integrable functions is integrable. *}
 
-lemma real_arch_invD:
-  "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
-  by(subst(asm) real_arch_inv)
-
-lemma integrable_uniform_limit: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_uniform_limit: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
   assumes "\<forall>e>0. \<exists>g. (\<forall>x\<in>{a..b}. norm(f x - g x) \<le> e) \<and> g integrable_on {a..b}"
   shows "f integrable_on {a..b}"
 proof- { presume *:"content {a..b} > 0 \<Longrightarrow> ?thesis"
@@ -2446,7 +2480,7 @@
 lemma dest_vec1_indicator_abs_le_1: "abs(indicator s x) \<le> 1"
   unfolding indicator_def by auto
 
-definition "negligible (s::(real^'n) set) \<equiv> (\<forall>a b. ((indicator s) has_integral 0) {a..b})"
+definition "negligible (s::(_::ordered_euclidean_space) set) \<equiv> (\<forall>a b. ((indicator s) has_integral 0) {a..b})"
 
 lemma indicator_simps[simp]:"x\<in>s \<Longrightarrow> indicator s x = 1" "x\<notin>s \<Longrightarrow> indicator s x = 0"
   unfolding indicator_def by auto
@@ -2461,98 +2495,104 @@
   apply(rule iterate_nonzero_image_lemma) apply(rule assms monoidal_monoid)+
   unfolding assms using neutral_add unfolding neutral_add using assms by auto 
 
-lemma interval_doublesplit: shows "{a..b} \<inter> {x . abs(x$k - c) \<le> (e::real)} =
-  {(\<chi> i. if i = k then max (a$k) (c - e) else a$i) .. (\<chi> i. if i = k then min (b$k) (c + e) else b$i)}"
+lemma interval_doublesplit:  fixes a::"'a::ordered_euclidean_space" assumes "k<DIM('a)"
+  shows "{a..b} \<inter> {x . abs(x$$k - c) \<le> (e::real)} = 
+  {(\<chi>\<chi> i. if i = k then max (a$$k) (c - e) else a$$i) .. (\<chi>\<chi> i. if i = k then min (b$$k) (c + e) else b$$i)}"
 proof- have *:"\<And>x c e::real. abs(x - c) \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e" by auto
   have **:"\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}" by blast
-  show ?thesis unfolding * ** interval_split by(rule refl) qed
-
-lemma division_doublesplit: assumes "p division_of {a..b::real^'n}" 
-  shows "{l \<inter> {x. abs(x$k - c) \<le> e} |l. l \<in> p \<and> l \<inter> {x. abs(x$k - c) \<le> e} \<noteq> {}} division_of ({a..b} \<inter> {x. abs(x$k - c) \<le> e})"
+  show ?thesis unfolding * ** interval_split[OF assms] by(rule refl) qed
+
+lemma division_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "p division_of {a..b}" and k:"k<DIM('a)"
+  shows "{l \<inter> {x. abs(x$$k - c) \<le> e} |l. l \<in> p \<and> l \<inter> {x. abs(x$$k - c) \<le> e} \<noteq> {}} division_of ({a..b} \<inter> {x. abs(x$$k - c) \<le> e})"
 proof- have *:"\<And>x c. abs(x - c) \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e" by auto
   have **:"\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'" by auto
-  note division_split(1)[OF assms, where c="c+e" and k=k,unfolded interval_split]
-  note division_split(2)[OF this, where c="c-e" and k=k] 
-  thus ?thesis apply(rule **) unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit
+  note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
+  note division_split(2)[OF this, where c="c-e" and k=k,OF k] 
+  thus ?thesis apply(rule **) using k apply- unfolding interval_doublesplit unfolding * unfolding interval_split interval_doublesplit
     apply(rule set_ext) unfolding mem_Collect_eq apply rule apply(erule conjE exE)+ apply(rule_tac x=la in exI) defer
-    apply(erule conjE exE)+ apply(rule_tac x="l \<inter> {x. c + e \<ge> x $ k}" in exI) apply rule defer apply rule
+    apply(erule conjE exE)+ apply(rule_tac x="l \<inter> {x. c + e \<ge> x $$ k}" in exI) apply rule defer apply rule
     apply(rule_tac x=l in exI) by blast+ qed
 
-lemma content_doublesplit: assumes "0 < e"
-  obtains d where "0 < d" "content({a..b} \<inter> {x. abs(x$k - c) \<le> d}) < e"
+lemma content_doublesplit: fixes a::"'a::ordered_euclidean_space" assumes "0 < e" and k:"k<DIM('a)"
+  obtains d where "0 < d" "content({a..b} \<inter> {x. abs(x$$k - c) \<le> d}) < e"
 proof(cases "content {a..b} = 0")
-  case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit
+  case True show ?thesis apply(rule that[of 1]) defer unfolding interval_doublesplit[OF k]
     apply(rule le_less_trans[OF content_subset]) defer apply(subst True)
-    unfolding interval_doublesplit[THEN sym] using assms by auto 
-next case False def d \<equiv> "e / 3 / setprod (\<lambda>i. b$i - a$i) (UNIV - {k})"
+    unfolding interval_doublesplit[THEN sym,OF k] using assms by auto 
+next case False def d \<equiv> "e / 3 / setprod (\<lambda>i. b$$i - a$$i) ({..<DIM('a)} - {k})"
   note False[unfolded content_eq_0 not_ex not_le, rule_format]
-  hence prod0:"0 < setprod (\<lambda>i. b$i - a$i) (UNIV - {k})" apply-apply(rule setprod_pos) by smt
+  hence "\<And>x. x<DIM('a) \<Longrightarrow> b$$x > a$$x" by(auto simp add:not_le)
+  hence prod0:"0 < setprod (\<lambda>i. b$$i - a$$i) ({..<DIM('a)} - {k})" apply-apply(rule setprod_pos) by(auto simp add:field_simps)
   hence "d > 0" unfolding d_def using assms by(auto simp add:field_simps) thus ?thesis
-  proof(rule that[of d]) have *:"UNIV = insert k (UNIV - {k})" by auto
-    have **:"{a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} \<noteq> {} \<Longrightarrow> 
-      (\<Prod>i\<in>UNIV - {k}. interval_upperbound ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) $ i - interval_lowerbound ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) $ i)
-      = (\<Prod>i\<in>UNIV - {k}. b$i - a$i)" apply(rule setprod_cong,rule refl)
-      unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds by auto
-    show "content ({a..b} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms)
+  proof(rule that[of d]) have *:"{..<DIM('a)} = insert k ({..<DIM('a)} - {k})" using k by auto
+    have **:"{a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} \<noteq> {} \<Longrightarrow> 
+      (\<Prod>i\<in>{..<DIM('a)} - {k}. interval_upperbound ({a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) $$ i
+      - interval_lowerbound ({a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) $$ i)
+      = (\<Prod>i\<in>{..<DIM('a)} - {k}. b$$i - a$$i)" apply(rule setprod_cong,rule refl) 
+      unfolding interval_doublesplit[OF k] apply(subst interval_bounds) defer apply(subst interval_bounds)
+      unfolding interval_eq_empty not_ex not_less by auto
+    show "content ({a..b} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) < e" apply(cases) unfolding content_def apply(subst if_P,assumption,rule assms)
       unfolding if_not_P apply(subst *) apply(subst setprod_insert) unfolding **
-      unfolding interval_doublesplit interval_eq_empty not_ex not_less unfolding interval_bounds unfolding Cart_lambda_beta if_P[OF refl]
-    proof- have "(min (b $ k) (c + d) - max (a $ k) (c - d)) \<le> 2 * d" by auto
-      also have "... < e / (\<Prod>i\<in>UNIV - {k}. b $ i - a $ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps)
-      finally show "(min (b $ k) (c + d) - max (a $ k) (c - d)) * (\<Prod>i\<in>UNIV - {k}. b $ i - a $ i) < e"
+      unfolding interval_doublesplit[OF k] interval_eq_empty not_ex not_less prefer 3
+      apply(subst interval_bounds) defer apply(subst interval_bounds) unfolding euclidean_lambda_beta'[OF k] if_P[OF refl]
+    proof- have "(min (b $$ k) (c + d) - max (a $$ k) (c - d)) \<le> 2 * d" by auto
+      also have "... < e / (\<Prod>i\<in>{..<DIM('a)} - {k}. b $$ i - a $$ i)" unfolding d_def using assms prod0 by(auto simp add:field_simps)
+      finally show "(min (b $$ k) (c + d) - max (a $$ k) (c - d)) * (\<Prod>i\<in>{..<DIM('a)} - {k}. b $$ i - a $$ i) < e"
         unfolding pos_less_divide_eq[OF prod0] . qed auto qed qed
 
-lemma negligible_standard_hyperplane[intro]: "negligible {x. x$k = (c::real)}" 
+lemma negligible_standard_hyperplane[intro]: fixes type::"'a::ordered_euclidean_space" assumes k:"k<DIM('a)"
+  shows "negligible {x::'a. x$$k = (c::real)}" 
   unfolding negligible_def has_integral apply(rule,rule,rule,rule)
-proof- case goal1 from content_doublesplit[OF this,of a b k c] guess d . note d=this let ?i = "indicator {x. x$k = c}"
+proof- case goal1 from content_doublesplit[OF this k,of a b c] guess d . note d=this let ?i = "indicator {x::'a. x$$k = c}"
   show ?case apply(rule_tac x="\<lambda>x. ball x d" in exI) apply(rule,rule gauge_ball,rule d)
   proof(rule,rule) fix p assume p:"p tagged_division_of {a..b} \<and> (\<lambda>x. ball x d) fine p"
-    have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x$k - c) \<le> d}) *\<^sub>R ?i x)"
+    have *:"(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x$$k - c) \<le> d}) *\<^sub>R ?i x)"
       apply(rule setsum_cong2) unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
       apply(cases,rule disjI1,assumption,rule disjI2)
-    proof- fix x l assume as:"(x,l)\<in>p" "?i x \<noteq> 0" hence xk:"x$k = c" unfolding indicator_def apply-by(rule ccontr,auto)
-      show "content l = content (l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d})" apply(rule arg_cong[where f=content])
+    proof- fix x l assume as:"(x,l)\<in>p" "?i x \<noteq> 0" hence xk:"x$$k = c" unfolding indicator_def apply-by(rule ccontr,auto)
+      show "content l = content (l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d})" apply(rule arg_cong[where f=content])
         apply(rule set_ext,rule,rule) unfolding mem_Collect_eq
       proof- fix y assume y:"y\<in>l" note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv]
         note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y] note le_less_trans[OF component_le_norm[of _ k] this]
-        thus "\<bar>y $ k - c\<bar> \<le> d" unfolding Cart_nth.diff xk by auto
+        thus "\<bar>y $$ k - c\<bar> \<le> d" unfolding euclidean_simps xk by auto
       qed auto qed
     note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]]
     show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e" unfolding diff_0_right * unfolding real_scaleR_def real_norm_def
       apply(subst abs_of_nonneg) apply(rule setsum_nonneg,rule) unfolding split_paired_all split_conv
       apply(rule mult_nonneg_nonneg) apply(drule p'(4)) apply(erule exE)+ apply(rule_tac b=b in back_subst)
       prefer 2 apply(subst(asm) eq_commute) apply assumption
-      apply(subst interval_doublesplit) apply(rule content_pos_le) apply(rule indicator_pos_le)
-    proof- have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}))"
+      apply(subst interval_doublesplit[OF k]) apply(rule content_pos_le) apply(rule indicator_pos_le)
+    proof- have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}))"
         apply(rule setsum_mono) unfolding split_paired_all split_conv 
-        apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit intro!:content_pos_le)
+        apply(rule mult_right_le_one_le) apply(drule p'(4)) by(auto simp add:interval_doublesplit[OF k] intro!:content_pos_le)
       also have "... < e" apply(subst setsum_over_tagged_division_lemma[OF p[THEN conjunct1]])
-      proof- case goal1 have "content ({u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<le> content {u..v}"
-          unfolding interval_doublesplit apply(rule content_subset) unfolding interval_doublesplit[THEN sym] by auto
-        thus ?case unfolding goal1 unfolding interval_doublesplit using content_pos_le by smt
-      next have *:"setsum content {l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} |l. l \<in> snd ` p \<and> l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} \<noteq> {}} \<ge> 0"
+      proof- case goal1 have "content ({u..v} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) \<le> content {u..v}"
+          unfolding interval_doublesplit[OF k] apply(rule content_subset) unfolding interval_doublesplit[THEN sym,OF k] by auto
+        thus ?case unfolding goal1 unfolding interval_doublesplit[OF k] using content_pos_le by smt
+      next have *:"setsum content {l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} |l. l \<in> snd ` p \<and> l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} \<noteq> {}} \<ge> 0"
           apply(rule setsum_nonneg,rule) unfolding mem_Collect_eq image_iff apply(erule exE bexE conjE)+ unfolding split_paired_all 
-        proof- fix x l a b assume as:"x = l \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}" "(a, b) \<in> p" "l = snd (a, b)"
+        proof- fix x l a b assume as:"x = l \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}" "(a, b) \<in> p" "l = snd (a, b)"
           guess u v using p'(4)[OF as(2)] apply-by(erule exE)+ note * = this
-          show "content x \<ge> 0" unfolding as snd_conv * interval_doublesplit by(rule content_pos_le)
-        qed have **:"norm (1::real) \<le> 1" by auto note division_doublesplit[OF p'',unfolded interval_doublesplit]
-        note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym]]
-        note this[unfolded real_scaleR_def real_norm_def mult_1_right mult_1, of k c d] note le_less_trans[OF this d(2)]
-        from this[unfolded abs_of_nonneg[OF *]] show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d})) < e"
+          show "content x \<ge> 0" unfolding as snd_conv * interval_doublesplit[OF k] by(rule content_pos_le)
+        qed have **:"norm (1::real) \<le> 1" by auto note division_doublesplit[OF p'' k,unfolded interval_doublesplit[OF k]]
+        note dsum_bound[OF this **,unfolded interval_doublesplit[THEN sym,OF k]]
+        note this[unfolded real_scaleR_def real_norm_def mult_1_right mult_1, of c d] note le_less_trans[OF this d(2)]
+        from this[unfolded abs_of_nonneg[OF *]] show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d})) < e"
           apply(subst vsum_nonzero_image_lemma[of "snd ` p" content "{}", unfolded o_def,THEN sym])
           apply(rule finite_imageI p' content_empty)+ unfolding forall_in_division[OF p'']
         proof(rule,rule,rule,rule,rule,rule,rule,erule conjE) fix m n u v
-          assume as:"{m..n} \<in> snd ` p" "{u..v} \<in> snd ` p" "{m..n} \<noteq> {u..v}"  "{m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d} = {u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}"
-          have "({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<inter> ({u..v} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) \<subseteq> {m..n} \<inter> {u..v}" by blast
+          assume as:"{m..n} \<in> snd ` p" "{u..v} \<in> snd ` p" "{m..n} \<noteq> {u..v}"  "{m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d} = {u..v} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}"
+          have "({m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) \<inter> ({u..v} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) \<subseteq> {m..n} \<inter> {u..v}" by blast
           note subset_interior[OF this, unfolded division_ofD(5)[OF p'' as(1-3)] interior_inter[of "{m..n}"]]
-          hence "interior ({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) = {}" unfolding as Int_absorb by auto
-          thus "content ({m..n} \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) = 0" unfolding interval_doublesplit content_eq_0_interior[THEN sym] .
+          hence "interior ({m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) = {}" unfolding as Int_absorb by auto
+          thus "content ({m..n} \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) = 0" unfolding interval_doublesplit[OF k] content_eq_0_interior[THEN sym] .
         qed qed
-      finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $ k - c\<bar> \<le> d}) * ?i x) < e" .
+      finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x $$ k - c\<bar> \<le> d}) * ?i x) < e" .
     qed qed qed
 
 subsection {* A technical lemma about "refinement" of division. *}
 
-lemma tagged_division_finer: fixes p::"((real^'n) \<times> ((real^'n) set)) set"
+lemma tagged_division_finer: fixes p::"(('a::ordered_euclidean_space) \<times> (('a::ordered_euclidean_space) set)) set"
   assumes "p tagged_division_of {a..b}" "gauge d"
   obtains q where "q tagged_division_of {a..b}" "d fine q" "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
 proof-
@@ -2562,7 +2602,7 @@
   { have *:"finite p" "p tagged_partial_division_of {a..b}" using assms(1) unfolding tagged_division_of_def by auto
     presume "\<And>p. finite p \<Longrightarrow> ?P p" from this[rule_format,OF * assms(2)] guess q .. note q=this
     thus ?thesis apply-apply(rule that[of q]) unfolding tagged_division_ofD[OF assms(1)] by auto
-  } fix p::"((real^'n) \<times> ((real^'n) set)) set" assume as:"finite p"
+  } fix p::"(('a::ordered_euclidean_space) \<times> (('a::ordered_euclidean_space) set)) set" assume as:"finite p"
   show "?P p" apply(rule,rule) using as proof(induct p) 
     case empty show ?case apply(rule_tac x="{}" in exI) unfolding fine_def by auto
   next case (insert xk p) guess x k using surj_pair[of xk] apply- by(erule exE)+ note xk=this
@@ -2614,10 +2654,10 @@
     show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}" apply(rule finite_product_dependent) using insert by auto
   qed(insert insert, auto) qed auto
 
-lemma has_integral_negligible: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+lemma has_integral_negligible: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::real_normed_vector"
   assumes "negligible s" "\<forall>x\<in>(t - s). f x = 0"
   shows "(f has_integral 0) t"
-proof- presume P:"\<And>f::real^'n \<Rightarrow> 'a. \<And>a b. (\<forall>x. ~(x \<in> s) \<longrightarrow> f x = 0) \<Longrightarrow> (f has_integral 0) ({a..b})"
+proof- presume P:"\<And>f::'b::ordered_euclidean_space \<Rightarrow> 'a. \<And>a b. (\<forall>x. ~(x \<in> s) \<longrightarrow> f x = 0) \<Longrightarrow> (f has_integral 0) ({a..b})"
   let ?f = "(\<lambda>x. if x \<in> t then f x else 0)"
   show ?thesis apply(rule_tac f="?f" in has_integral_eq) apply(rule) unfolding if_P apply(rule refl)
     apply(subst has_integral_alt) apply(cases,subst if_P,assumption) unfolding if_not_P
@@ -2627,7 +2667,7 @@
       apply(safe,rule_tac x=1 in exI,rule) apply(rule zero_less_one,safe) apply(rule_tac x=0 in exI)
       apply(rule,rule P) using assms(2) by auto
   qed
-next fix f::"real^'n \<Rightarrow> 'a" and a b::"real^'n" assume assm:"\<forall>x. x \<notin> s \<longrightarrow> f x = 0" 
+next fix f::"'b \<Rightarrow> 'a" and a b::"'b" assume assm:"\<forall>x. x \<notin> s \<longrightarrow> f x = 0" 
   show "(f has_integral 0) {a..b}" unfolding has_integral
   proof(safe) case goal1
     hence "\<And>n. e / 2 / ((real n+1) * (2 ^ n)) > 0" 
@@ -2680,10 +2720,10 @@
         apply(subst sumr_geometric) using goal1 by auto
       finally show "?goal" by auto qed qed qed
 
-lemma has_integral_spike: fixes f::"real^'n \<Rightarrow> 'a::real_normed_vector"
+lemma has_integral_spike: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::real_normed_vector"
   assumes "negligible s" "(\<forall>x\<in>(t - s). g x = f x)" "(f has_integral y) t"
   shows "(g has_integral y) t"
-proof- { fix a b::"real^'n" and f g ::"real^'n \<Rightarrow> 'a" and y::'a
+proof- { fix a b::"'b" and f g ::"'b \<Rightarrow> 'a" and y::'a
     assume as:"\<forall>x \<in> {a..b} - s. g x = f x" "(f has_integral y) {a..b}"
     have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) {a..b}" apply(rule has_integral_add[OF as(2)])
       apply(rule has_integral_negligible[OF assms(1)]) using as by auto
@@ -2726,9 +2766,8 @@
 lemma negligible_union_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> (negligible s \<and> negligible t)"
   using negligible_union by auto
 
-lemma negligible_sing[intro]: "negligible {a::real^'n}" 
-proof- guess x using UNIV_witness[where 'a='n] ..
-  show ?thesis using negligible_standard_hyperplane[of x "a$x"] by auto qed
+lemma negligible_sing[intro]: "negligible {a::_::ordered_euclidean_space}" 
+  using negligible_standard_hyperplane[of 0 "a$$0"] by auto 
 
 lemma negligible_insert[simp]: "negligible(insert a s) \<longleftrightarrow> negligible s"
   apply(subst insert_is_Un) unfolding negligible_union_eq by auto
@@ -2741,9 +2780,9 @@
 lemma negligible_unions[intro]: assumes "finite s" "\<forall>t\<in>s. negligible t" shows "negligible(\<Union>s)"
   using assms by(induct,auto) 
 
-lemma negligible:  "negligible s \<longleftrightarrow> (\<forall>t::(real^'n) set. (indicator s has_integral 0) t)"
+lemma negligible:  "negligible s \<longleftrightarrow> (\<forall>t::('a::ordered_euclidean_space) set. (indicator s has_integral 0) t)"
   apply safe defer apply(subst negligible_def)
-proof- fix t::"(real^'n) set" assume as:"negligible s"
+proof- fix t::"'a set" assume as:"negligible s"
   have *:"(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)" by(rule ext,auto)  
   show "(indicator s has_integral 0) t" apply(subst has_integral_alt)
     apply(cases,subst if_P,assumption) unfolding if_not_P apply(safe,rule as[unfolded negligible_def,rule_format])
@@ -2767,10 +2806,10 @@
 
 subsection {* In particular, the boundary of an interval is negligible. *}
 
-lemma negligible_frontier_interval: "negligible({a..b} - {a<..<b})"
-proof- let ?A = "\<Union>((\<lambda>k. {x. x$k = a$k} \<union> {x. x$k = b$k}) ` UNIV)"
+lemma negligible_frontier_interval: "negligible({a::'a::ordered_euclidean_space..b} - {a<..<b})"
+proof- let ?A = "\<Union>((\<lambda>k. {x. x$$k = a$$k} \<union> {x::'a. x$$k = b$$k}) ` {..<DIM('a)})"
   have "{a..b} - {a<..<b} \<subseteq> ?A" apply rule unfolding Diff_iff mem_interval not_all
-    apply(erule conjE exE)+ apply(rule_tac X="{x. x $ xa = a $ xa} \<union> {x. x $ xa = b $ xa}" in UnionI)
+    apply(erule conjE exE)+ apply(rule_tac X="{x. x $$ xa = a $$ xa} \<union> {x. x $$ xa = b $$ xa}" in UnionI)
     apply(erule_tac[!] x=xa in allE) by auto
   thus ?thesis apply-apply(rule negligible_subset[of ?A]) apply(rule negligible_unions[OF finite_imageI]) by auto qed
 
@@ -2799,34 +2838,35 @@
   shows "(\<forall>i\<in>d. P i) \<longleftrightarrow> P {a..b}"
   using operative_division[OF monoidal_and assms] division_of_finite[OF assms(2)] by auto
 
-lemma operative_approximable: assumes "0 \<le> e" fixes f::"real^'n \<Rightarrow> 'a::banach"
-  shows "operative op \<and> (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::real^'n)) \<le> e) \<and> g integrable_on i)" unfolding operative_def neutral_and
-proof safe fix a b::"real^'n" { assume "content {a..b} = 0"
+lemma operative_approximable: assumes "0 \<le> e" fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
+  shows "operative op \<and> (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)" unfolding operative_def neutral_and
+proof safe fix a b::"'b" { assume "content {a..b} = 0"
     thus "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" 
       apply(rule_tac x=f in exI) using assms by(auto intro!:integrable_on_null) }
-  { fix c k g assume as:"\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}"
-    show "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. x $ k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. x $ k \<le> c}"
-      "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $ k}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. c \<le> x $ k}"
-      apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2)] by auto }
-  fix c k g1 g2 assume as:"\<forall>x\<in>{a..b} \<inter> {x. x $ k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on {a..b} \<inter> {x. x $ k \<le> c}"
-                          "\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $ k}. norm (f x - g2 x) \<le> e" "g2 integrable_on {a..b} \<inter> {x. c \<le> x $ k}"
-  let ?g = "\<lambda>x. if x$k = c then f x else if x$k \<le> c then g1 x else g2 x"
+  { fix c k g assume as:"\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}" and k:"k<DIM('b)"
+    show "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. x $$ k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. x $$ k \<le> c}"
+      "\<exists>g. (\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $$ k}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b} \<inter> {x. c \<le> x $$ k}"
+      apply(rule_tac[!] x=g in exI) using as(1) integrable_split[OF as(2) k] by auto }
+  fix c k g1 g2 assume as:"\<forall>x\<in>{a..b} \<inter> {x. x $$ k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on {a..b} \<inter> {x. x $$ k \<le> c}"
+                          "\<forall>x\<in>{a..b} \<inter> {x. c \<le> x $$ k}. norm (f x - g2 x) \<le> e" "g2 integrable_on {a..b} \<inter> {x. c \<le> x $$ k}"
+  assume k:"k<DIM('b)"
+  let ?g = "\<lambda>x. if x$$k = c then f x else if x$$k \<le> c then g1 x else g2 x"
   show "\<exists>g. (\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e) \<and> g integrable_on {a..b}" apply(rule_tac x="?g" in exI)
-  proof safe case goal1 thus ?case apply- apply(cases "x$k=c", case_tac "x$k < c") using as assms by auto
-  next case goal2 presume "?g integrable_on {a..b} \<inter> {x. x $ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $ k \<ge> c}"
-    then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this]
+  proof safe case goal1 thus ?case apply- apply(cases "x$$k=c", case_tac "x$$k < c") using as assms by auto
+  next case goal2 presume "?g integrable_on {a..b} \<inter> {x. x $$ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $$ k \<ge> c}"
+    then guess h1 h2 unfolding integrable_on_def by auto from has_integral_split[OF this k] 
     show ?case unfolding integrable_on_def by auto
-  next show "?g integrable_on {a..b} \<inter> {x. x $ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $ k \<ge> c}"
-      apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using as(2,4) by auto qed qed
-
-lemma approximable_on_division: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  next show "?g integrable_on {a..b} \<inter> {x. x $$ k \<le> c}" "?g integrable_on {a..b} \<inter> {x. x $$ k \<ge> c}"
+      apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]]) using k as(2,4) by auto qed qed
+
+lemma approximable_on_division: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "0 \<le> e" "d division_of {a..b}" "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
   obtains g where "\<forall>x\<in>{a..b}. norm (f x - g x) \<le> e" "g integrable_on {a..b}"
 proof- note * = operative_division[OF monoidal_and operative_approximable[OF assms(1)] assms(2)]
   note this[unfolded iterate_and[OF division_of_finite[OF assms(2)]]] from assms(3)[unfolded this[of f]]
   guess g .. thus thesis apply-apply(rule that[of g]) by auto qed
 
-lemma integrable_continuous: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_continuous: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}"
 proof(rule integrable_uniform_limit,safe) fix e::real assume e:"0 < e"
   from compact_uniformly_continuous[OF assms compact_interval,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d ..
@@ -2847,48 +2887,52 @@
 subsection {* Specialization of additivity to one dimension. *}
 
 lemma operative_1_lt: assumes "monoidal opp"
-  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real^1} = neutral opp) \<and>
+  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real} = neutral opp) \<and>
                 (\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
-  unfolding operative_def content_eq_0_1 forall_1 vector_le_def vector_less_def
-proof safe fix a b c::"real^1" assume as:"\<forall>a b c. f {a..b} = opp (f ({a..b} \<inter> {x. x $ 1 \<le> c})) (f ({a..b} \<inter> {x. c \<le> x $ 1}))" "a $ 1 < c $ 1" "c $ 1 < b $ 1"
-    from this(2-) have "{a..b} \<inter> {x. x $ 1 \<le> c $ 1} = {a..c}" "{a..b} \<inter> {x. x $ 1 \<ge> c $ 1} = {c..b}" by auto
-    thus "opp (f {a..c}) (f {c..b}) = f {a..b}" unfolding as(1)[rule_format,of a b "c$1"] by auto
-next fix a b::"real^1" and c::real
-  assume as:"\<forall>a b. b $ 1 \<le> a $ 1 \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a $ 1 < c $ 1 \<and> c $ 1 < b $ 1 \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}"
-  show "f {a..b} = opp (f ({a..b} \<inter> {x. x $ 1 \<le> c})) (f ({a..b} \<inter> {x. c \<le> x $ 1}))"
-  proof(cases "c \<in> {a$1 .. b$1}")
-    case False hence "c<a$1 \<or> c>b$1" by auto
+  unfolding operative_def content_eq_0 DIM_real less_one dnf_simps(39,41) Eucl_real_simps
+    (* The dnf_simps simplify "\<exists> x. x= _ \<and> _" and "\<forall>k. k = _ \<longrightarrow> _" *)
+proof safe fix a b c::"real" assume as:"\<forall>a b c. f {a..b} = opp (f ({a..b} \<inter> {x. x \<le> c}))
+    (f ({a..b} \<inter> {x. c \<le> x}))" "a < c" "c < b"
+    from this(2-) have "{a..b} \<inter> {x. x \<le> c} = {a..c}" "{a..b} \<inter> {x. x \<ge> c} = {c..b}" by auto
+    thus "opp (f {a..c}) (f {c..b}) = f {a..b}" unfolding as(1)[rule_format,of a b "c"] by auto
+next fix a b c::real
+  assume as:"\<forall>a b. b \<le> a \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}"
+  show "f {a..b} = opp (f ({a..b} \<inter> {x. x \<le> c})) (f ({a..b} \<inter> {x. c \<le> x}))"
+  proof(cases "c \<in> {a .. b}")
+    case False hence "c<a \<or> c>b" by auto
     thus ?thesis apply-apply(erule disjE)
-    proof- assume "c<a$1" hence *:"{a..b} \<inter> {x. x $ 1 \<le> c} = {1..0}"  "{a..b} \<inter> {x. c \<le> x $ 1} = {a..b}" by auto
+    proof- assume "c<a" hence *:"{a..b} \<inter> {x. x \<le> c} = {1..0}"  "{a..b} \<inter> {x. c \<le> x} = {a..b}" by auto
       show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
-    next   assume "b$1<c" hence *:"{a..b} \<inter> {x. x $ 1 \<le> c} = {a..b}"  "{a..b} \<inter> {x. c \<le> x $ 1} = {1..0}" by auto
+    next   assume "b<c" hence *:"{a..b} \<inter> {x. x \<le> c} = {a..b}"  "{a..b} \<inter> {x. c \<le> x} = {1..0}" by auto
       show ?thesis unfolding * apply(subst as(1)[rule_format,of 0 1]) using assms by auto
     qed
-  next case True hence *:"min (b $ 1) c = c" "max (a $ 1) c = c" by auto
-    show ?thesis unfolding interval_split num1_eq_iff if_True * vec_def[THEN sym]
-    proof(cases "c = a$1 \<or> c = b$1")
-      case False thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})"
+  next case True hence *:"min (b) c = c" "max a c = c" by auto
+    have **:"0 < DIM(real)" by auto
+    have ***:"\<And>P Q. (\<chi>\<chi> i. if i = 0 then P i else Q i) = (P 0::real)" apply(subst euclidean_eq)
+      apply safe unfolding euclidean_lambda_beta' by auto
+    show ?thesis unfolding interval_split[OF **,unfolded Eucl_real_simps(1,3)] unfolding *** *
+    proof(cases "c = a \<or> c = b")
+      case False thus "f {a..b} = opp (f {a..c}) (f {c..b})"
         apply-apply(subst as(2)[rule_format]) using True by auto
-    next case True thus "f {a..b} = opp (f {a..vec1 c}) (f {vec1 c..b})" apply-
-      proof(erule disjE) assume "c=a$1" hence *:"a = vec1 c" unfolding Cart_eq by auto 
-        hence "f {a..vec1 c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
+    next case True thus "f {a..b} = opp (f {a..c}) (f {c..b})" apply-
+      proof(erule disjE) assume *:"c=a"
+        hence "f {a..c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
         thus ?thesis using assms unfolding * by auto
-      next assume "c=b$1" hence *:"b = vec1 c" unfolding Cart_eq by auto 
-        hence "f {vec1 c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
+      next assume *:"c=b" hence "f {c..b} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
         thus ?thesis using assms unfolding * by auto qed qed qed qed
 
 lemma operative_1_le: assumes "monoidal opp"
-  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real^1} = neutral opp) \<and>
+  shows "operative opp f \<longleftrightarrow> ((\<forall>a b. b \<le> a \<longrightarrow> f {a..b::real} = neutral opp) \<and>
                 (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f{a..c})(f{c..b}) = f {a..b}))"
 unfolding operative_1_lt[OF assms]
-proof safe fix a b c::"real^1" assume as:"\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a < c" "c < b"
-  show "opp (f {a..c}) (f {c..b}) = f {a..b}" apply(rule as(1)[rule_format]) using as(2-) unfolding vector_le_def vector_less_def by auto
-next fix a b c ::"real^1"
-  assume "\<forall>a b. b \<le> a \<longrightarrow> f {a..b} = neutral opp" "\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a \<le> c" "c \<le> b"
+proof safe fix a b c::"real" assume as:"\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a < c" "c < b"
+  show "opp (f {a..c}) (f {c..b}) = f {a..b}" apply(rule as(1)[rule_format]) using as(2-) by auto
+next fix a b c ::"real" assume "\<forall>a b. b \<le> a \<longrightarrow> f {a..b} = neutral opp"
+    "\<forall>a b c. a < c \<and> c < b \<longrightarrow> opp (f {a..c}) (f {c..b}) = f {a..b}" "a \<le> c" "c \<le> b"
   note as = this[rule_format]
   show "opp (f {a..c}) (f {c..b}) = f {a..b}"
   proof(cases "c = a \<or> c = b")
-    case False thus ?thesis apply-apply(subst as(2)) using as(3-) unfolding vector_le_def vector_less_def Cart_eq by(auto simp del:dest_vec1_eq)
+    case False thus ?thesis apply-apply(subst as(2)) using as(3-) by(auto)
     next case True thus ?thesis apply-
       proof(erule disjE) assume *:"c=a" hence "f {a..c} = neutral opp" apply-apply(rule as(1)[rule_format]) by auto
         thus ?thesis using assms unfolding * by auto
@@ -2898,16 +2942,16 @@
 subsection {* Special case of additivity we need for the FCT. *}
 
 lemma interval_bound_sing[simp]: "interval_upperbound {a} = a"  "interval_lowerbound {a} = a"
-  unfolding interval_upperbound_def interval_lowerbound_def unfolding Cart_eq by auto
-
-lemma additive_tagged_division_1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
-  assumes "dest_vec1 a \<le> dest_vec1 b" "p tagged_division_of {a..b}"
+  unfolding interval_upperbound_def interval_lowerbound_def  by auto
+
+lemma additive_tagged_division_1: fixes f::"real \<Rightarrow> 'a::real_normed_vector"
+  assumes "a \<le> b" "p tagged_division_of {a..b}"
   shows "setsum (\<lambda>(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a"
-proof- let ?f = "(\<lambda>k::(real^1) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
-  have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty_1
-    by(auto simp add:not_less vector_less_def)
+proof- let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
+  have ***:"\<forall>i<DIM(real). a $$ i \<le> b $$ i" using assms by auto 
+  have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty by auto
   have **:"{a..b} \<noteq> {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)]
-  note * = this[unfolded if_not_P[OF **] interval_bound_1[OF assms(1)],THEN sym ]
+  note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],THEN sym]
   show ?thesis unfolding * apply(subst setsum_iterate[THEN sym]) defer
     apply(rule setsum_cong2) unfolding split_paired_all split_conv using assms(2) by auto qed
 
@@ -2934,40 +2978,46 @@
 
 subsection {* Fundamental theorem of calculus. *}
 
-lemma fundamental_theorem_of_calculus: fixes f::"real^1 \<Rightarrow> 'a::banach"
-  assumes "a \<le> b"  "\<forall>x\<in>{a..b}. ((f o vec1) has_vector_derivative f'(vec1 x)) (at x within {a..b})"
-  shows "(f' has_integral (f(vec1 b) - f(vec1 a))) ({vec1 a..vec1 b})"
+lemma interval_bounds_real: assumes "a\<le>(b::real)"
+  shows "interval_upperbound {a..b} = b" "interval_lowerbound {a..b} = a"
+  apply(rule_tac[!] interval_bounds) using assms by auto
+
+lemma fundamental_theorem_of_calculus: fixes f::"real \<Rightarrow> 'a::banach"
+  assumes "a \<le> b"  "\<forall>x\<in>{a..b}. (f has_vector_derivative f' x) (at x within {a..b})"
+  shows "(f' has_integral (f b - f a)) ({a..b})"
 unfolding has_integral_factor_content
-proof safe fix e::real assume e:"e>0" have ab:"dest_vec1 (vec1 a) \<le> dest_vec1 (vec1 b)" using assms(1) by auto
+proof safe fix e::real assume e:"e>0"
   note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt]
   have *:"\<And>P Q. \<forall>x\<in>{a..b}. P x \<and> (\<forall>e>0. \<exists>d>0. Q x e d) \<Longrightarrow> \<forall>x. \<exists>(d::real)>0. x\<in>{a..b} \<longrightarrow> Q x e d" using e by blast
   note this[OF assm,unfolded gauge_existence_lemma] from choice[OF this,unfolded Ball_def[symmetric]]
   guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
-  show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
-                 norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \<le> e * content {vec1 a..vec1 b})"
-    apply(rule_tac x="\<lambda>x. ball x (d (dest_vec1 x))" in exI,safe)
+  show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow>
+                 norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b})"
+    apply(rule_tac x="\<lambda>x. ball x (d x)" in exI,safe)
     apply(rule gauge_ball_dependent,rule,rule d(1))
-  proof- fix p assume as:"p tagged_division_of {vec1 a..vec1 b}" "(\<lambda>x. ball x (d (dest_vec1 x))) fine p"
-    show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f (vec1 b) - f (vec1 a))) \<le> e * content {vec1 a..vec1 b}" 
-      unfolding content_1[OF ab] additive_tagged_division_1[OF ab as(1),of f,THEN sym]
-      unfolding vector_minus_component[THEN sym] additive_tagged_division_1[OF ab as(1),of "\<lambda>x. x",THEN sym]
-      apply(subst dest_vec1_setsum) unfolding setsum_right_distrib defer unfolding setsum_subtractf[THEN sym] 
+  proof- fix p assume as:"p tagged_division_of {a..b}" "(\<lambda>x. ball x (d x)) fine p"
+    show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b}" 
+      unfolding content_real[OF assms(1)] additive_tagged_division_1[OF assms(1) as(1),of f,THEN sym]
+      unfolding additive_tagged_division_1[OF assms(1) as(1),of "\<lambda>x. x",THEN sym]
+      unfolding setsum_right_distrib defer unfolding setsum_subtractf[THEN sym] 
     proof(rule setsum_norm_le,safe) fix x k assume "(x,k)\<in>p"
       note xk = tagged_division_ofD(2-4)[OF as(1) this] from this(3) guess u v apply-by(erule exE)+ note k=this
-      have *:"dest_vec1 u \<le> dest_vec1 v" using xk unfolding k by auto
-      have ball:"\<forall>xa\<in>k. xa \<in> ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
-      have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \<le> norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)"
+      have *:"u \<le> v" using xk unfolding k by auto
+      have ball:"\<forall>xa\<in>k. xa \<in> ball x (d x)" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,
+        unfolded split_conv subset_eq] .
+      have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le>
+        norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)"
         apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
         unfolding scaleR.diff_left by(auto simp add:algebra_simps)
-      also have "... \<le> e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)"
-        apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4
-        apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1])
+      also have "... \<le> e * norm (u - x) + e * norm (v - x)"
+        apply(rule add_mono) apply(rule d(2)[of "x" "u",unfolded o_def]) prefer 4
+        apply(rule d(2)[of "x" "v",unfolded o_def])
         using ball[rule_format,of u] ball[rule_format,of v] 
-        using xk(1-2) unfolding k subset_eq by(auto simp add:dist_norm norm_real)
-      also have "... \<le> e * dest_vec1 (interval_upperbound k - interval_lowerbound k)"
-        unfolding k interval_bound_1[OF *] using xk(1) unfolding k by(auto simp add:dist_norm norm_real field_simps)
+        using xk(1-2) unfolding k subset_eq by(auto simp add:dist_real_def) 
+      also have "... \<le> e * (interval_upperbound k - interval_lowerbound k)"
+        unfolding k interval_bounds_real[OF *] using xk(1) unfolding k by(auto simp add:dist_real_def field_simps)
       finally show "norm (content k *\<^sub>R f' x - (f (interval_upperbound k) - f (interval_lowerbound k))) \<le>
-        e * dest_vec1 (interval_upperbound k - interval_lowerbound k)" unfolding k interval_bound_1[OF *] content_1[OF *] .
+        e * (interval_upperbound k - interval_lowerbound k)" unfolding k interval_bounds_real[OF *] content_real[OF *] .
     qed(insert as, auto) qed qed
 
 subsection {* Attempt a systematic general set of "offset" results for components. *}
@@ -2980,11 +3030,11 @@
 
 subsection {* Only need trivial subintervals if the interval itself is trivial. *}
 
-lemma division_of_nontrivial: fixes s::"(real^'n) set set"
+lemma division_of_nontrivial: fixes s::"('a::ordered_euclidean_space) set set"
   assumes "s division_of {a..b}" "content({a..b}) \<noteq> 0"
   shows "{k. k \<in> s \<and> content k \<noteq> 0} division_of {a..b}" using assms(1) apply-
 proof(induct "card s" arbitrary:s rule:nat_less_induct)
-  fix s::"(real^'n) set set" assume assm:"s division_of {a..b}"
+  fix s::"'a set set" assume assm:"s division_of {a..b}"
     "\<forall>m<card s. \<forall>x. m = card x \<longrightarrow> x division_of {a..b} \<longrightarrow> {k \<in> x. content k \<noteq> 0} division_of {a..b}" 
   note s = division_ofD[OF assm(1)] let ?thesis = "{k \<in> s. content k \<noteq> 0} division_of {a..b}"
   { presume *:"{k \<in> s. content k \<noteq> 0} \<noteq> s \<Longrightarrow> ?thesis"
@@ -2999,30 +3049,32 @@
   have "k \<subseteq> \<Union>(s - {k})" apply safe apply(rule *[unfolded closed_limpt,rule_format]) unfolding islimpt_approachable
   proof safe fix x and e::real assume as:"x\<in>k" "e>0"
     from k(2)[unfolded k content_eq_0] guess i .. 
-    hence i:"c$i = d$i" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by smt
-    hence xi:"x$i = d$i" using as unfolding k mem_interval by smt
-    def y \<equiv> "(\<chi> j. if j = i then if c$i \<le> (a$i + b$i) / 2 then c$i + min e (b$i - c$i) / 2 else c$i - min e (c$i - a$i) / 2 else x$j)"
+    hence i:"c$$i = d$$i" "i<DIM('a)" using s(3)[OF k(1),unfolded k] unfolding interval_ne_empty by auto
+    hence xi:"x$$i = d$$i" using as unfolding k mem_interval by smt
+    def y \<equiv> "(\<chi>\<chi> j. if j = i then if c$$i \<le> (a$$i + b$$i) / 2 then c$$i +
+      min e (b$$i - c$$i) / 2 else c$$i - min e (c$$i - a$$i) / 2 else x$$j)::'a"
     show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> dist x' x < e" apply(rule_tac x=y in bexI) 
     proof have "d \<in> {c..d}" using s(3)[OF k(1)] unfolding k interval_eq_empty mem_interval by(fastsimp simp add: not_less)
       hence "d \<in> {a..b}" using s(2)[OF k(1)] unfolding k by auto note di = this[unfolded mem_interval,THEN spec[where x=i]]
-      hence xyi:"y$i \<noteq> x$i" unfolding y_def unfolding i xi Cart_lambda_beta if_P[OF refl]
-        apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2) using assms(2)[unfolded content_eq_0] by smt+ 
-      thus "y \<noteq> x" unfolding Cart_eq by auto
-      have *:"UNIV = insert i (UNIV - {i})" by auto
-      have "norm (y - x) < e + setsum (\<lambda>i. 0) (UNIV::'n set)" apply(rule le_less_trans[OF norm_le_l1])
+      hence xyi:"y$$i \<noteq> x$$i" unfolding y_def unfolding i xi euclidean_lambda_beta'[OF i(2)] if_P[OF refl]
+        apply(cases) apply(subst if_P,assumption) unfolding if_not_P not_le using as(2)
+        using assms(2)[unfolded content_eq_0] using i(2) by smt+ 
+      thus "y \<noteq> x" unfolding euclidean_eq[where 'a='a] using i by auto
+      have *:"{..<DIM('a)} = insert i ({..<DIM('a)} - {i})" using i by auto
+      have "norm (y - x) < e + setsum (\<lambda>i. 0) {..<DIM('a)}" apply(rule le_less_trans[OF norm_le_l1])
         apply(subst *,subst setsum_insert) prefer 3 apply(rule add_less_le_mono)
-      proof- show "\<bar>(y - x) $ i\<bar> < e" unfolding y_def Cart_lambda_beta vector_minus_component if_P[OF refl]
+      proof- show "\<bar>(y - x) $$ i\<bar> < e" unfolding y_def euclidean_simps euclidean_lambda_beta'[OF i(2)] if_P[OF refl]
           apply(cases) apply(subst if_P,assumption) unfolding if_not_P unfolding i xi using di as(2) by auto
-        show "(\<Sum>i\<in>UNIV - {i}. \<bar>(y - x) $ i\<bar>) \<le> (\<Sum>i\<in>UNIV. 0)" unfolding y_def by auto 
+        show "(\<Sum>i\<in>{..<DIM('a)} - {i}. \<bar>(y - x) $$ i\<bar>) \<le> (\<Sum>i\<in>{..<DIM('a)}. 0)" unfolding y_def by auto 
       qed auto thus "dist y x < e" unfolding dist_norm by auto
       have "y\<notin>k" unfolding k mem_interval apply rule apply(erule_tac x=i in allE) using xyi unfolding k i xi by auto
       moreover have "y \<in> \<Union>s" unfolding s mem_interval
-      proof note simps = y_def Cart_lambda_beta if_not_P
-        fix j::'n show "a $ j \<le> y $ j \<and> y $ j \<le> b $ j" 
+      proof(rule,rule) note simps = y_def euclidean_lambda_beta' if_not_P
+        fix j assume j:"j<DIM('a)" show "a $$ j \<le> y $$ j \<and> y $$ j \<le> b $$ j" 
         proof(cases "j = i") case False have "x \<in> {a..b}" using s(2)[OF k(1)] as(1) by auto
-          thus ?thesis unfolding simps if_not_P[OF False] unfolding mem_interval by auto
+          thus ?thesis using j apply- unfolding simps if_not_P[OF False] unfolding mem_interval by auto
         next case True note T = this show ?thesis
-          proof(cases "c $ i \<le> (a $ i + b $ i) / 2")
+          proof(cases "c $$ i \<le> (a $$ i + b $$ i) / 2")
             case True show ?thesis unfolding simps if_P[OF T] if_P[OF True] unfolding i
               using True as(2) di apply-apply rule unfolding T by (auto simp add:field_simps) 
           next case False thus ?thesis unfolding simps if_P[OF T] if_not_P[OF False] unfolding i
@@ -3036,20 +3088,20 @@
 
 subsection {* Integrabibility on subintervals. *}
 
-lemma operative_integrable: fixes f::"real^'n \<Rightarrow> 'a::banach" shows
+lemma operative_integrable: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach" shows
   "operative op \<and> (\<lambda>i. f integrable_on i)"
   unfolding operative_def neutral_and apply safe apply(subst integrable_on_def)
-  unfolding has_integral_null_eq apply(rule,rule refl) apply(rule,assumption)+
-  unfolding integrable_on_def by(auto intro: has_integral_split)
-
-lemma integrable_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach" 
+  unfolding has_integral_null_eq apply(rule,rule refl) apply(rule,assumption,assumption)+
+  unfolding integrable_on_def by(auto intro!: has_integral_split)
+
+lemma integrable_subinterval: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach" 
   assumes "f integrable_on {a..b}" "{c..d} \<subseteq> {a..b}" shows "f integrable_on {c..d}" 
   apply(cases "{c..d} = {}") defer apply(rule partial_division_extend_1[OF assms(2)],assumption)
   using operative_division_and[OF operative_integrable,THEN sym,of _ _ _ f] assms(1) by auto
 
 subsection {* Combining adjacent intervals in 1 dimension. *}
 
-lemma has_integral_combine: assumes "(a::real^1) \<le> c" "c \<le> b"
+lemma has_integral_combine: assumes "(a::real) \<le> c" "c \<le> b"
   "(f has_integral i) {a..c}" "(f has_integral (j::'a::banach)) {c..b}"
   shows "(f has_integral (i + j)) {a..b}"
 proof- note operative_integral[of f, unfolded operative_1_le[OF monoidal_lifted[OF monoidal_monoid]]]
@@ -3059,19 +3111,19 @@
   with * show ?thesis apply-apply(subst(asm) if_P) defer apply(subst(asm) if_P) defer apply(subst(asm) if_P)
     unfolding lifted.simps using assms(3-) by(auto simp add: integrable_on_def integral_unique) qed
 
-lemma integral_combine: fixes f::"real^1 \<Rightarrow> 'a::banach"
+lemma integral_combine: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "a \<le> c" "c \<le> b" "f integrable_on ({a..b})"
   shows "integral {a..c} f + integral {c..b} f = integral({a..b}) f"
   apply(rule integral_unique[THEN sym]) apply(rule has_integral_combine[OF assms(1-2)])
   apply(rule_tac[!] integrable_integral integrable_subinterval[OF assms(3)])+ using assms(1-2) by auto
 
-lemma integrable_combine: fixes f::"real^1 \<Rightarrow> 'a::banach"
+lemma integrable_combine: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "a \<le> c" "c \<le> b" "f integrable_on {a..c}" "f integrable_on {c..b}"
   shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by(fastsimp intro!:has_integral_combine)
 
 subsection {* Reduce integrability to "local" integrability. *}
 
-lemma integrable_on_little_subintervals: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_on_little_subintervals: fixes f::"'b::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "\<forall>x\<in>{a..b}. \<exists>d>0. \<forall>u v. x \<in> {u..v} \<and> {u..v} \<subseteq> ball x d \<and> {u..v} \<subseteq> {a..b} \<longrightarrow> f integrable_on {u..v}"
   shows "f integrable_on {a..b}"
 proof- have "\<forall>x. \<exists>d. x\<in>{a..b} \<longrightarrow> d>0 \<and> (\<forall>u v. x \<in> {u..v} \<and> {u..v} \<subseteq> ball x d \<and> {u..v} \<subseteq> {a..b} \<longrightarrow> f integrable_on {u..v})"
@@ -3089,53 +3141,46 @@
 
 lemma integral_has_vector_derivative: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "continuous_on {a..b} f" "x \<in> {a..b}"
-  shows "((\<lambda>u. integral {vec a..vec u} (f o dest_vec1)) has_vector_derivative f(x)) (at x within {a..b})"
+  shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})"
   unfolding has_vector_derivative_def has_derivative_within_alt
 apply safe apply(rule scaleR.bounded_linear_left)
 proof- fix e::real assume e:"e>0"
-  note compact_uniformly_continuous[OF assms(1) compact_real_interval,unfolded uniformly_continuous_on_def]
+  note compact_uniformly_continuous[OF assms(1) compact_interval,unfolded uniformly_continuous_on_def]
   from this[rule_format,OF e] guess d apply-by(erule conjE exE)+ note d=this[rule_format]
-  let ?I = "\<lambda>a b. integral {vec1 a..vec1 b} (f \<circ> dest_vec1)"
+  let ?I = "\<lambda>a b. integral {a..b} f"
   show "\<exists>d>0. \<forall>y\<in>{a..b}. norm (y - x) < d \<longrightarrow> norm (?I a y - ?I a x - (y - x) *\<^sub>R f x) \<le> e * norm (y - x)"
   proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x")
-      case False have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous)
-        apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
+      case False have "f integrable_on {a..y}" apply(rule integrable_subinterval,rule integrable_continuous)
+        apply(rule assms)  unfolding not_less using assms(2) goal1 by auto
       hence *:"?I a y - ?I a x = ?I x y" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
         using False unfolding not_less using assms(2) goal1 by auto
-      have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto
-      show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
+      have **:"norm (y - x) = content {x..y}" apply(subst content_real) using False unfolding not_less by auto
+      show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x)"]) unfolding * unfolding o_def
         defer apply(rule has_integral_sub) apply(rule integrable_integral)
-        apply(rule integrable_subinterval,rule integrable_continuous) apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+
-      proof- show "{vec1 x..vec1 y} \<subseteq> {vec1 a..vec1 b}" using goal1 assms(2) by auto
+        apply(rule integrable_subinterval,rule integrable_continuous) apply(rule assms)+
+      proof- show "{x..y} \<subseteq> {a..b}" using goal1 assms(2) by auto
         have *:"y - x = norm(y - x)" using False by auto
-        show "((\<lambda>xa. f x) has_integral (y - x) *\<^sub>R f x) {vec1 x..vec1 y}" apply(subst *) unfolding ** by auto
-        show "\<forall>xa\<in>{vec1 x..vec1 y}. norm (f (dest_vec1 xa) - f x) \<le> e" apply safe apply(rule less_imp_le)
+        show "((\<lambda>xa. f x) has_integral (y - x) *\<^sub>R f x) {x.. y}" apply(subst *) unfolding ** by auto
+        show "\<forall>xa\<in>{x..y}. norm (f xa - f x) \<le> e" apply safe apply(rule less_imp_le)
           apply(rule d(2)[unfolded dist_norm]) using assms(2) using goal1 by auto
       qed(insert e,auto)
-    next case True have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous)
-        apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
+    next case True have "f integrable_on {a..x}" apply(rule integrable_subinterval,rule integrable_continuous)
+        apply(rule assms)+  unfolding not_less using assms(2) goal1 by auto
       hence *:"?I a x - ?I a y = ?I y x" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
         using True using assms(2) goal1 by auto
-      have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto
+      have **:"norm (y - x) = content {y..x}" apply(subst content_real) using True unfolding not_less by auto
       have ***:"\<And>fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto 
       show ?thesis apply(subst ***) unfolding norm_minus_cancel **
-        apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
+        apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x)"]) unfolding * unfolding o_def
         defer apply(rule has_integral_sub) apply(subst minus_minus[THEN sym]) unfolding minus_minus
-        apply(rule integrable_integral) apply(rule integrable_subinterval,rule integrable_continuous)
-        apply(rule continuous_on_o_dest_vec1[unfolded o_def] assms)+
-      proof- show "{vec1 y..vec1 x} \<subseteq> {vec1 a..vec1 b}" using goal1 assms(2) by auto
+        apply(rule integrable_integral) apply(rule integrable_subinterval,rule integrable_continuous) apply(rule assms)+
+      proof- show "{y..x} \<subseteq> {a..b}" using goal1 assms(2) by auto
         have *:"x - y = norm(y - x)" using True by auto
-        show "((\<lambda>xa. f x) has_integral (x - y) *\<^sub>R f x) {vec1 y..vec1 x}" apply(subst *) unfolding ** by auto
-        show "\<forall>xa\<in>{vec1 y..vec1 x}. norm (f (dest_vec1 xa) - f x) \<le> e" apply safe apply(rule less_imp_le)
+        show "((\<lambda>xa. f x) has_integral (x - y) *\<^sub>R f x) {y..x}" apply(subst *) unfolding ** by auto
+        show "\<forall>xa\<in>{y..x}. norm (f xa - f x) \<le> e" apply safe apply(rule less_imp_le)
           apply(rule d(2)[unfolded dist_norm]) using assms(2) using goal1 by auto
       qed(insert e,auto) qed qed qed
 
-lemma integral_has_vector_derivative': fixes f::"real^1 \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a..b} f" "x \<in> {a..b}"
-  shows "((\<lambda>u. (integral {a..vec u} f)) has_vector_derivative f x) (at (x$1) within {a$1..b$1})"
-  using integral_has_vector_derivative[OF continuous_on_o_vec1[OF assms(1)], of "x$1"]
-  unfolding o_def vec1_dest_vec1 using assms(2) by auto
-
 lemma antiderivative_continuous: assumes "continuous_on {a..b::real} f"
   obtains g where "\<forall>x\<in> {a..b}. (g has_vector_derivative (f(x)::_::banach)) (at x within {a..b})"
   apply(rule that,rule) using integral_has_vector_derivative[OF assms] by auto
@@ -3143,13 +3188,12 @@
 subsection {* Combined fundamental theorem of calculus. *}
 
 lemma antiderivative_integral_continuous: fixes f::"real \<Rightarrow> 'a::banach" assumes "continuous_on {a..b} f"
-  obtains g where "\<forall>u\<in>{a..b}. \<forall>v \<in> {a..b}. u \<le> v \<longrightarrow> ((f o dest_vec1) has_integral (g v - g u)) {vec u..vec v}"
+  obtains g where "\<forall>u\<in>{a..b}. \<forall>v \<in> {a..b}. u \<le> v \<longrightarrow> (f has_integral (g v - g u)) {u..v}"
 proof- from antiderivative_continuous[OF assms] guess g . note g=this
   show ?thesis apply(rule that[of g])
   proof safe case goal1 have "\<forall>x\<in>{u..v}. (g has_vector_derivative f x) (at x within {u..v})"
       apply(rule,rule has_vector_derivative_within_subset) apply(rule g[rule_format]) using goal1(1-2) by auto
-    thus ?case using fundamental_theorem_of_calculus[OF goal1(3),of "g o dest_vec1" "f o dest_vec1"]
-      unfolding o_def vec1_dest_vec1 by auto qed qed
+    thus ?case using fundamental_theorem_of_calculus[OF goal1(3),of "g" "f"] by auto qed qed
 
 subsection {* General "twiddling" for interval-to-interval function image. *}
 
@@ -3206,34 +3250,33 @@
 
 subsection {* Special case of a basic affine transformation. *}
 
-lemma interval_image_affinity_interval: shows "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::real^'n) + c) ` {a..b} = {u..v}"
+lemma interval_image_affinity_interval: shows "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::'a::ordered_euclidean_space) + c) ` {a..b} = {u..v}"
   unfolding image_affinity_interval by auto
 
-lemmas Cart_simps = Cart_nth.add Cart_nth.minus Cart_nth.zero Cart_nth.diff Cart_nth.scaleR real_scaleR_def Cart_lambda_beta
-   Cart_eq vector_le_def vector_less_def
-
 lemma setprod_cong2: assumes "\<And>x. x \<in> A \<Longrightarrow> f x = g x" shows "setprod f A = setprod g A"
   apply(rule setprod_cong) using assms by auto
 
 lemma content_image_affinity_interval: 
- "content((\<lambda>x::real^'n. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ CARD('n) * content {a..b}" (is "?l = ?r")
+ "content((\<lambda>x::'a::ordered_euclidean_space. m *\<^sub>R x + c) ` {a..b}) = (abs m) ^ DIM('a) * content {a..b}" (is "?l = ?r")
 proof- { presume *:"{a..b}\<noteq>{} \<Longrightarrow> ?thesis" show ?thesis apply(cases,rule *,assumption)
       unfolding not_not using content_empty by auto }
+  have *:"DIM('a) = card {..<DIM('a)}" by auto
   assume as:"{a..b}\<noteq>{}" show ?thesis proof(cases "m \<ge> 0")
     case True show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_P[OF True]
-      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') 
-      defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym]
-      apply(rule setprod_cong2) using True as unfolding interval_ne_empty Cart_simps not_le  
+      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') defer apply(subst(2) *)
+      apply(subst setprod_constant[THEN sym]) apply(rule finite_lessThan) unfolding setprod_timesf[THEN sym]
+      apply(rule setprod_cong2) using True as unfolding interval_ne_empty euclidean_simps not_le  
       by(auto simp add:field_simps intro:mult_left_mono)
   next case False show ?thesis unfolding image_affinity_interval if_not_P[OF as] if_not_P[OF False]
-      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') 
-      defer apply(subst setprod_constant[THEN sym]) apply(rule finite_UNIV) unfolding setprod_timesf[THEN sym]
-      apply(rule setprod_cong2) using False as unfolding interval_ne_empty Cart_simps not_le 
+      unfolding content_closed_interval'[OF as] apply(subst content_closed_interval') defer apply(subst(2) *)
+      apply(subst setprod_constant[THEN sym]) apply(rule finite_lessThan) unfolding setprod_timesf[THEN sym]
+      apply(rule setprod_cong2) using False as unfolding interval_ne_empty euclidean_simps not_le 
       by(auto simp add:field_simps mult_le_cancel_left_neg) qed qed
 
-lemma has_integral_affinity: assumes "(f has_integral i) {a..b::real^'n}" "m \<noteq> 0"
-  shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ CARD('n::finite))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})"
-  apply(rule has_integral_twiddle,safe) unfolding Cart_eq Cart_simps apply(rule zero_less_power)
+lemma has_integral_affinity: fixes a::"'a::ordered_euclidean_space" assumes "(f has_integral i) {a..b}" "m \<noteq> 0"
+  shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (abs(m) ^ DIM('a))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` {a..b})"
+  apply(rule has_integral_twiddle,safe) apply(rule zero_less_power) unfolding euclidean_eq[where 'a='a]
+  unfolding scaleR_right_distrib euclidean_simps scaleR.scaleR_left[THEN sym]
   defer apply(insert assms(2), simp add:field_simps) apply(insert assms(2), simp add:field_simps)
   apply(rule continuous_intros)+ apply(rule interval_image_affinity_interval)+ apply(rule content_image_affinity_interval) using assms by auto
 
@@ -3244,64 +3287,68 @@
 subsection {* Special case of stretching coordinate axes separately. *}
 
 lemma image_stretch_interval:
-  "(\<lambda>x. \<chi> k. m k * x$k) ` {a..b::real^'n} =
-  (if {a..b} = {} then {} else {(\<chi> k. min (m(k) * a$k) (m(k) * b$k)) ..  (\<chi> k. max (m(k) * a$k) (m(k) * b$k))})" (is "?l = ?r")
+  "(\<lambda>x. \<chi>\<chi> k. m k * x$$k) ` {a..b::'a::ordered_euclidean_space} =
+  (if {a..b} = {} then {} else {(\<chi>\<chi> k. min (m(k) * a$$k) (m(k) * b$$k))::'a ..  (\<chi>\<chi> k. max (m(k) * a$$k) (m(k) * b$$k))})"
+  (is "?l = ?r")
 proof(cases "{a..b}={}") case True thus ?thesis unfolding True by auto
-next have *:"\<And>P Q. (\<forall>i. P i) \<and> (\<forall>i. Q i) \<longleftrightarrow> (\<forall>i. P i \<and> Q i)" by auto
+next have *:"\<And>P Q. (\<forall>i<DIM('a). P i) \<and> (\<forall>i<DIM('a). Q i) \<longleftrightarrow> (\<forall>i<DIM('a). P i \<and> Q i)" by auto
   case False note ab = this[unfolded interval_ne_empty]
   show ?thesis apply-apply(rule set_ext)
-  proof- fix x::"real^'n" have **:"\<And>P Q. (\<forall>i. P i = Q i) \<Longrightarrow> (\<forall>i. P i) = (\<forall>i. Q i)" by auto
+  proof- fix x::"'a" have **:"\<And>P Q. (\<forall>i<DIM('a). P i = Q i) \<Longrightarrow> (\<forall>i<DIM('a). P i) = (\<forall>i<DIM('a). Q i)" by auto
     show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" unfolding if_not_P[OF False] 
-      unfolding image_iff mem_interval Bex_def Cart_simps Cart_eq *
-      unfolding lambda_skolem[THEN sym,of "\<lambda> i xa. (a $ i \<le> xa \<and> xa \<le> b $ i) \<and> x $ i = m i * xa"]
-    proof(rule **,rule) fix i::'n show "(\<exists>xa. (a $ i \<le> xa \<and> xa \<le> b $ i) \<and> x $ i = m i * xa) =
-        (min (m i * a $ i) (m i * b $ i) \<le> x $ i \<and> x $ i \<le> max (m i * a $ i) (m i * b $ i))"
-      proof(cases "m i = 0") case True thus ?thesis using ab by auto
+      unfolding image_iff mem_interval Bex_def euclidean_simps euclidean_eq[where 'a='a] *
+      unfolding imp_conjR[THEN sym] apply(subst euclidean_lambda_beta'') apply(subst lambda_skolem'[THEN sym])
+      apply(rule **,rule,rule) unfolding euclidean_lambda_beta'
+    proof- fix i assume i:"i<DIM('a)" show "(\<exists>xa. (a $$ i \<le> xa \<and> xa \<le> b $$ i) \<and> x $$ i = m i * xa) =
+        (min (m i * a $$ i) (m i * b $$ i) \<le> x $$ i \<and> x $$ i \<le> max (m i * a $$ i) (m i * b $$ i))"
+      proof(cases "m i = 0") case True thus ?thesis using ab i by auto
       next case False hence "0 < m i \<or> 0 > m i" by auto thus ?thesis apply-
-        proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $ i) (m i * b $ i) = m i * a $ i"
-            "max (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab unfolding min_def max_def by auto
-          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI)
+        proof(erule disjE) assume as:"0 < m i" hence *:"min (m i * a $$ i) (m i * b $$ i) = m i * a $$ i"
+            "max (m i * a $$ i) (m i * b $$ i) = m i * b $$ i" using ab i unfolding min_def max_def by auto
+          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$$i" in exI)
             using as by(auto simp add:field_simps)
-        next assume as:"0 > m i" hence *:"max (m i * a $ i) (m i * b $ i) = m i * a $ i"
-            "min (m i * a $ i) (m i * b $ i) = m i * b $ i" using ab as unfolding min_def max_def 
+        next assume as:"0 > m i" hence *:"max (m i * a $$ i) (m i * b $$ i) = m i * a $$ i"
+            "min (m i * a $$ i) (m i * b $$ i) = m i * b $$ i" using ab as i unfolding min_def max_def 
             by(auto simp add:field_simps mult_le_cancel_left_neg intro: order_antisym)
-          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$i" in exI)
+          show ?thesis unfolding * apply rule defer apply(rule_tac x="1 / m i * x$$i" in exI)
             using as by(auto simp add:field_simps) qed qed qed qed qed 
 
-lemma interval_image_stretch_interval: "\<exists>u v. (\<lambda>x. \<chi> k. m k * x$k) ` {a..b::real^'n} = {u..v}"
+lemma interval_image_stretch_interval: "\<exists>u v. (\<lambda>x. \<chi>\<chi> k. m k * x$$k) ` {a..b::'a::ordered_euclidean_space} = {u..v::'a}"
   unfolding image_stretch_interval by auto 
 
 lemma content_image_stretch_interval:
-  "content((\<lambda>x::real^'n. \<chi> k. m k * x$k) ` {a..b}) = abs(setprod m UNIV) * content({a..b})"
+  "content((\<lambda>x::'a::ordered_euclidean_space. (\<chi>\<chi> k. m k * x$$k)::'a) ` {a..b}) = abs(setprod m {..<DIM('a)}) * content({a..b})"
 proof(cases "{a..b} = {}") case True thus ?thesis
     unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto
-next case False hence "(\<lambda>x. \<chi> k. m k * x $ k) ` {a..b} \<noteq> {}" by auto
+next case False hence "(\<lambda>x. (\<chi>\<chi> k. m k * x $$ k)::'a) ` {a..b} \<noteq> {}" by auto
   thus ?thesis using False unfolding content_def image_stretch_interval apply- unfolding interval_bounds' if_not_P
-    unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding Cart_lambda_beta
-  proof- fix i::'n have "(m i < 0 \<or> m i > 0) \<or> m i = 0" by auto
-    thus "max (m i * a $ i) (m i * b $ i) - min (m i * a $ i) (m i * b $ i) = \<bar>m i\<bar> * (b $ i - a $ i)"
-      apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] 
+    unfolding abs_setprod setprod_timesf[THEN sym] apply(rule setprod_cong2) unfolding lessThan_iff euclidean_lambda_beta'
+  proof- fix i assume i:"i<DIM('a)" have "(m i < 0 \<or> m i > 0) \<or> m i = 0" by auto
+    thus "max (m i * a $$ i) (m i * b $$ i) - min (m i * a $$ i) (m i * b $$ i) = \<bar>m i\<bar> * (b $$ i - a $$ i)"
+      apply-apply(erule disjE)+ unfolding min_def max_def using False[unfolded interval_ne_empty,rule_format,of i] i 
       by(auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos) qed qed
 
-lemma has_integral_stretch: assumes "(f has_integral i) {a..b}" "\<forall>k. ~(m k = 0)"
-  shows "((\<lambda>x. f(\<chi> k. m k * x$k)) has_integral
-             ((1/(abs(setprod m UNIV))) *\<^sub>R i)) ((\<lambda>x. \<chi> k. 1/(m k) * x$k) ` {a..b})"
-  apply(rule has_integral_twiddle) unfolding zero_less_abs_iff content_image_stretch_interval
-  unfolding image_stretch_interval empty_as_interval Cart_eq using assms
-proof- show "\<forall>x. continuous (at x) (\<lambda>x. \<chi> k. m k * x $ k)"
+lemma has_integral_stretch: fixes f::"'a::ordered_euclidean_space => 'b::real_normed_vector"
+  assumes "(f has_integral i) {a..b}" "\<forall>k<DIM('a). ~(m k = 0)"
+  shows "((\<lambda>x. f(\<chi>\<chi> k. m k * x$$k)) has_integral
+             ((1/(abs(setprod m {..<DIM('a)}))) *\<^sub>R i)) ((\<lambda>x. (\<chi>\<chi> k. 1/(m k) * x$$k)::'a) ` {a..b})"
+  apply(rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval
+  unfolding image_stretch_interval empty_as_interval euclidean_eq[where 'a='a] using assms
+proof- show "\<forall>y::'a. continuous (at y) (\<lambda>x. (\<chi>\<chi> k. m k * x $$ k)::'a)"
    apply(rule,rule linear_continuous_at) unfolding linear_linear
-   unfolding linear_def Cart_simps Cart_eq by(auto simp add:field_simps) qed auto
-
-lemma integrable_stretch: 
-  assumes "f integrable_on {a..b}" "\<forall>k. ~(m k = 0)"
-  shows "(\<lambda>x. f(\<chi> k. m k * x$k)) integrable_on ((\<lambda>x. \<chi> k. 1/(m k) * x$k) ` {a..b})"
-  using assms unfolding integrable_on_def apply-apply(erule exE) apply(drule has_integral_stretch) by auto
+   unfolding linear_def euclidean_simps euclidean_eq[where 'a='a] by(auto simp add:field_simps) qed auto
+
+lemma integrable_stretch:  fixes f::"'a::ordered_euclidean_space => 'b::real_normed_vector"
+  assumes "f integrable_on {a..b}" "\<forall>k<DIM('a). ~(m k = 0)"
+  shows "(\<lambda>x::'a. f(\<chi>\<chi> k. m k * x$$k)) integrable_on ((\<lambda>x. \<chi>\<chi> k. 1/(m k) * x$$k) ` {a..b})"
+  using assms unfolding integrable_on_def apply-apply(erule exE) 
+  apply(drule has_integral_stretch,assumption) by auto
 
 subsection {* even more special cases. *}
 
-lemma uminus_interval_vector[simp]:"uminus ` {a..b} = {-b .. -a::real^'n}"
+lemma uminus_interval_vector[simp]:"uminus ` {a..b} = {-b .. -a::'a::ordered_euclidean_space}"
   apply(rule set_ext,rule) defer unfolding image_iff
-  apply(rule_tac x="-x" in bexI) by(auto simp add:vector_le_def minus_le_iff le_minus_iff)
+  apply(rule_tac x="-x" in bexI) by(auto simp add:minus_le_iff le_minus_iff eucl_le[where 'a='a])
 
 lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) {a..b}"
   shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
@@ -3318,16 +3365,12 @@
 
 subsection {* Stronger form of FCT; quite a tedious proof. *}
 
-(** move this **)
-declare norm_triangle_ineq4[intro] 
-
 lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)" by(meson zero_less_one)
 
 lemma additive_tagged_division_1': fixes f::"real \<Rightarrow> 'a::real_normed_vector"
-  assumes "a \<le> b" "p tagged_division_of {vec1 a..vec1 b}"
-  shows "setsum (\<lambda>(x,k). f (dest_vec1 (interval_upperbound k)) - f(dest_vec1 (interval_lowerbound k))) p = f b - f a"
-  using additive_tagged_division_1[OF _ assms(2), of "f o dest_vec1"]
-  unfolding o_def vec1_dest_vec1 using assms(1) by auto
+  assumes "a \<le> b" "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). f (interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a"
+  using additive_tagged_division_1[OF _ assms(2), of f] using assms(1) by auto
 
 lemma split_minus[simp]:"(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x"
   unfolding split_def by(rule refl)
@@ -3336,17 +3379,17 @@
   apply(subst(asm)(2) norm_minus_cancel[THEN sym])
   apply(drule norm_triangle_le) by(auto simp add:algebra_simps)
 
-lemma fundamental_theorem_of_calculus_interior:
+lemma fundamental_theorem_of_calculus_interior: fixes f::"real => 'a::real_normed_vector"
   assumes"a \<le> b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_vector_derivative f'(x)) (at x)"
-  shows "((f' o dest_vec1) has_integral (f b - f a)) {vec a..vec b}"
+  shows "(f' has_integral (f b - f a)) {a..b}"
 proof- { presume *:"a < b \<Longrightarrow> ?thesis" 
     show ?thesis proof(cases,rule *,assumption)
       assume "\<not> a < b" hence "a = b" using assms(1) by auto
-      hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" by(auto simp add: Cart_eq vector_le_def order_antisym)
-      show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0_1 using * `a=b` by auto
+      hence *:"{a .. b} = {b}" "f b - f a = 0" by(auto simp add:  order_antisym)
+      show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0 using * `a=b` by auto
     qed } assume ab:"a < b"
-  let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
-                   norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f' \<circ> dest_vec1) x) - (f b - f a)) \<le> e * content {vec1 a..vec1 b})"
+  let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow>
+                   norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b})"
   { presume "\<And>e. e>0 \<Longrightarrow> ?P e" thus ?thesis unfolding has_integral_factor_content by auto }
   fix e::real assume e:"e>0"
   note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib]
@@ -3354,11 +3397,11 @@
   from this(2) have "\<forall>x\<in>{a<..<b}. \<exists>d>0. \<forall>y. norm (y - x) < d \<longrightarrow> norm (f y - f x - (y - x) *\<^sub>R f' x) \<le> e/2 * norm (y - x)"
     apply-apply safe apply(erule_tac x=x in ballE,erule_tac x="e/2" in allE) using e by auto note this[unfolded bgauge_existence_lemma]
   from choice[OF this] guess d .. note conjunctD2[OF this[rule_format]] note d = this[rule_format]
-  have "bounded (f ` {a..b})" apply(rule compact_imp_bounded compact_continuous_image)+ using compact_real_interval assms by auto
+  have "bounded (f ` {a..b})" apply(rule compact_imp_bounded compact_continuous_image)+ using compact_interval assms by auto
   from this[unfolded bounded_pos] guess B .. note B = this[rule_format]
 
   have "\<exists>da. 0 < da \<and> (\<forall>c. a \<le> c \<and> {a..c} \<subseteq> {a..b} \<and> {a..c} \<subseteq> ball a da
-    \<longrightarrow> norm(content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b - a)) / 4)"
+    \<longrightarrow> norm(content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b - a)) / 4)"
   proof- have "a\<in>{a..b}" using ab by auto
     note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
     note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps)
@@ -3366,9 +3409,8 @@
     have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' a) \<le> (e * (b - a)) / 8"
     proof(cases "f' a = 0") case True
       thus ?thesis apply(rule_tac x=1 in exI) using ab e by(auto intro!:mult_nonneg_nonneg) 
-    next case False thus ?thesis 
-        apply(rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI)
-        using ab e by(auto simp add:field_simps)
+    next case False thus ?thesis
+        apply(rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI) using ab e by(auto simp add:field_simps) 
     qed then guess l .. note l = conjunctD2[OF this]
     show ?thesis apply(rule_tac x="min k l" in exI) apply safe unfolding min_less_iff_conj apply(rule,(rule l k)+)
     proof- fix c assume as:"a \<le> c" "{a..c} \<subseteq> {a..b}" "{a..c} \<subseteq> ball a (min k l)" 
@@ -3379,13 +3421,16 @@
         thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto
       next case goal2 show ?case apply(rule less_imp_le) apply(cases "a = c") defer
           apply(rule k(2)[unfolded dist_norm]) using as' e ab by(auto simp add:field_simps)
-      qed finally show "norm (content {vec1 a..vec1 c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto
+      qed finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b - a) / 4"
+        unfolding content_real[OF as(1)] by auto
     qed qed then guess da .. note da=conjunctD2[OF this,rule_format]
 
-  have "\<exists>db>0. \<forall>c\<le>b. {c..b} \<subseteq> {a..b} \<and> {c..b} \<subseteq> ball b db \<longrightarrow> norm(content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b - a)) / 4"
+  have "\<exists>db>0. \<forall>c\<le>b. {c..b} \<subseteq> {a..b} \<and> {c..b} \<subseteq> ball b db \<longrightarrow>
+    norm(content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b - a)) / 4"
   proof- have "b\<in>{a..b}" using ab by auto
     note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
-    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0" using e ab by(auto simp add:field_simps)
+    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0"
+      using e ab by(auto simp add:field_simps)
     from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
     have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' b) \<le> (e * (b - a)) / 8"
     proof(cases "f' b = 0") case True
@@ -3403,129 +3448,137 @@
         thus ?case apply-apply(rule order_trans[OF _ l(2)]) unfolding norm_scaleR apply(rule mult_right_mono) by auto
       next case goal2 show ?case apply(rule less_imp_le) apply(cases "b = c") defer apply(subst norm_minus_commute)
           apply(rule k(2)[unfolded dist_norm]) using as' e ab by(auto simp add:field_simps)
-      qed finally show "norm (content {vec1 c..vec1 b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b - a) / 4" unfolding content_1'[OF as(1)] by auto
+      qed finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b - a) / 4"
+        unfolding content_real[OF as(1)] by auto
     qed qed then guess db .. note db=conjunctD2[OF this,rule_format]
 
-  let ?d = "(\<lambda>x. ball x (if x=vec1 a then da else if x=vec b then db else d (dest_vec1 x)))"
+  let ?d = "(\<lambda>x. ball x (if x=a then da else if x=b then db else d x))"
   show "?P e" apply(rule_tac x="?d" in exI)
   proof safe case goal1 show ?case apply(rule gauge_ball_dependent) using ab db(1) da(1) d(1) by auto
-  next case goal2 note as=this let ?A = "{t. fst t \<in> {vec1 a, vec1 b}}" note p = tagged_division_ofD[OF goal2(1)]
+  next case goal2 note as=this let ?A = "{t. fst t \<in> {a, b}}" note p = tagged_division_ofD[OF goal2(1)]
     have pA:"p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"  using goal2 by auto
     note * = additive_tagged_division_1'[OF assms(1) goal2(1), THEN sym]
     have **:"\<And>n1 s1 n2 s2::real. n2 \<le> s2 / 2 \<Longrightarrow> n1 - s1 \<le> s2 / 2 \<Longrightarrow> n1 + n2 \<le> s1 + s2" by arith
-    show ?case unfolding content_1'[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[THEN sym] split_minus
+    show ?case unfolding content_real[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[THEN sym] split_minus
       unfolding setsum_right_distrib apply(subst(2) pA,subst pA) unfolding setsum_Un_disjoint[OF pA(2-)]
     proof(rule norm_triangle_le,rule **) 
       case goal1 show ?case apply(rule order_trans,rule setsum_norm_le) apply(rule pA) defer apply(subst divide.setsum)
       proof(rule order_refl,safe,unfold not_le o_def split_conv fst_conv,rule ccontr) fix x k assume as:"(x,k) \<in> p"
-          "e * (dest_vec1 (interval_upperbound k) - dest_vec1 (interval_lowerbound k)) / 2
-          < norm (content k *\<^sub>R f' (dest_vec1 x) - (f (dest_vec1 (interval_upperbound k)) - f (dest_vec1 (interval_lowerbound k))))"
+          "e * (interval_upperbound k -  interval_lowerbound k) / 2
+          < norm (content k *\<^sub>R f' x - (f (interval_upperbound k) - f (interval_lowerbound k)))"
         from p(4)[OF this(1)] guess u v apply-by(erule exE)+ note k=this
-        hence "\<forall>i. u$i \<le> v$i" and uv:"{u,v}\<subseteq>{u..v}" using p(2)[OF as(1)] by auto note this(1) this(1)[unfolded forall_1]
-        note result = as(2)[unfolded k interval_bounds[OF this(1)] content_1[OF this(2)]]
-
-        assume as':"x \<noteq> vec1 a" "x \<noteq> vec1 b" hence "x$1 \<in> {a<..<b}" using p(2-3)[OF as(1)] by(auto simp add: Cart_eq) note  * = d(2)[OF this]
-        have "norm ((v$1 - u$1) *\<^sub>R f' (x$1) - (f (v$1) - f (u$1))) =
-          norm ((f (u$1) - f (x$1) - (u$1 - x$1) *\<^sub>R f' (x$1)) - (f (v$1) - f (x$1) - (v$1 - x$1) *\<^sub>R f' (x$1)))" 
+        hence "u \<le> v" and uv:"{u,v}\<subseteq>{u..v}" using p(2)[OF as(1)] by auto
+        note result = as(2)[unfolded k interval_bounds_real[OF this(1)] content_real[OF this(1)]]
+
+        assume as':"x \<noteq> a" "x \<noteq> b" hence "x \<in> {a<..<b}" using p(2-3)[OF as(1)] by auto
+        note  * = d(2)[OF this]
+        have "norm ((v - u) *\<^sub>R f' (x) - (f (v) - f (u))) =
+          norm ((f (u) - f (x) - (u - x) *\<^sub>R f' (x)) - (f (v) - f (x) - (v - x) *\<^sub>R f' (x)))" 
           apply(rule arg_cong[of _ _ norm]) unfolding scaleR_left.diff by auto 
-        also have "... \<le> e / 2 * norm (u$1 - x$1) + e / 2 * norm (v$1 - x$1)" apply(rule norm_triangle_le_sub)
+        also have "... \<le> e / 2 * norm (u - x) + e / 2 * norm (v - x)" apply(rule norm_triangle_le_sub)
           apply(rule add_mono) apply(rule_tac[!] *) using fineD[OF goal2(2) as(1)] as' unfolding k subset_eq
-          apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) using uv by(auto simp add:dist_real)
-        also have "... \<le> e / 2 * norm (v$1 - u$1)" using p(2)[OF as(1)] unfolding k by(auto simp add:field_simps)
-        finally have "e * (dest_vec1 v - dest_vec1 u) / 2 < e * (dest_vec1 v - dest_vec1 u) / 2"
+          apply- apply(erule_tac x=u in ballE,erule_tac[3] x=v in ballE) using uv by(auto simp:dist_real_def)
+        also have "... \<le> e / 2 * norm (v - u)" using p(2)[OF as(1)] unfolding k by(auto simp add:field_simps)
+        finally have "e * (v - u) / 2 < e * (v - u) / 2"
           apply- apply(rule less_le_trans[OF result]) using uv by auto thus False by auto qed
 
     next have *:"\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2) / 2 \<Longrightarrow> x - s1 \<le> s2 / 2" by auto
       case goal2 show ?case apply(rule *) apply(rule setsum_nonneg) apply(rule,unfold split_paired_all split_conv)
         defer unfolding setsum_Un_disjoint[OF pA(2-),THEN sym] pA(1)[THEN sym] unfolding setsum_right_distrib[THEN sym] 
-        apply(subst additive_tagged_division_1[OF _ as(1)]) unfolding vec1_dest_vec1 apply(rule assms)
-      proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {vec1 a, vec1 b}}" note xk=IntD1[OF this]
+        apply(subst additive_tagged_division_1[OF _ as(1)]) apply(rule assms)
+      proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {a, b}}" note xk=IntD1[OF this]
         from p(4)[OF this] guess u v apply-by(erule exE)+ note uv=this
         with p(2)[OF xk] have "{u..v} \<noteq> {}" by auto
-        thus "0 \<le> e * ((interval_upperbound k)$1 - (interval_lowerbound k)$1)"
+        thus "0 \<le> e * ((interval_upperbound k) - (interval_lowerbound k))"
           unfolding uv using e by(auto simp add:field_simps)
       next have *:"\<And>s f t e. setsum f s = setsum f t \<Longrightarrow> norm(setsum f t) \<le> e \<Longrightarrow> norm(setsum f s) \<le> e" by auto
-        show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R (f' \<circ> dest_vec1) x -
-          (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) \<le> e * (b - a) / 2" 
-          apply(rule *[where t="p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content(snd t) \<noteq> 0}"])
+        show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x -
+          (f ((interval_upperbound k)) - f ((interval_lowerbound k)))) \<le> e * (b - a) / 2" 
+          apply(rule *[where t="p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0}"])
           apply(rule setsum_mono_zero_right[OF pA(2)]) defer apply(rule) unfolding split_paired_all split_conv o_def
-        proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {vec1 a, vec1 b}} - p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content (snd t) \<noteq> 0}"
+        proof- fix x k assume "(x,k) \<in> p \<inter> {t. fst t \<in> {a, b}} - p \<inter> {t. fst t \<in> {a, b} \<and> content (snd t) \<noteq> 0}"
           hence xk:"(x,k)\<in>p" "content k = 0" by auto from p(4)[OF xk(1)] guess u v apply-by(erule exE)+ note uv=this
-          have "k\<noteq>{}" using p(2)[OF xk(1)] by auto hence *:"u = v" using xk unfolding uv content_eq_0_1 interval_eq_empty by auto
-          thus "content k *\<^sub>R (f' (x$1)) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1)) = 0" using xk unfolding uv by auto
-        next have *:"p \<inter> {t. fst t \<in> {vec1 a, vec1 b} \<and> content(snd t) \<noteq> 0} = 
-            {t. t\<in>p \<and> fst t = vec1 a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = vec1 b \<and> content(snd t) \<noteq> 0}" by blast
-          have **:"\<And>s f. \<And>e::real. (\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y) \<Longrightarrow> (\<forall>x. x \<in> s \<longrightarrow> norm(f x) \<le> e) \<Longrightarrow> e>0 \<Longrightarrow> norm(setsum f s) \<le> e"
+          have "k\<noteq>{}" using p(2)[OF xk(1)] by auto hence *:"u = v" using xk
+            unfolding uv content_eq_0 interval_eq_empty by auto
+          thus "content k *\<^sub>R (f' (x)) - (f ((interval_upperbound k)) - f ((interval_lowerbound k))) = 0" using xk unfolding uv by auto
+        next have *:"p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0} = 
+            {t. t\<in>p \<and> fst t = a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = b \<and> content(snd t) \<noteq> 0}" by blast
+          have **:"\<And>s f. \<And>e::real. (\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y) \<Longrightarrow> (\<forall>x. x \<in> s \<longrightarrow> norm(f x) \<le> e)
+            \<Longrightarrow> e>0 \<Longrightarrow> norm(setsum f s) \<le> e"
           proof(case_tac "s={}") case goal2 then obtain x where "x\<in>s" by auto hence *:"s = {x}" using goal2(1) by auto
             thus ?case using `x\<in>s` goal2(2) by auto
           qed auto
-          case goal2 show ?case apply(subst *, subst setsum_Un_disjoint) prefer 4 apply(rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) 
+          case goal2 show ?case apply(subst *, subst setsum_Un_disjoint) prefer 4
+            apply(rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"]) 
             apply(rule norm_triangle_le,rule add_mono) apply(rule_tac[1-2] **)
-          proof- let ?B = "\<lambda>x. {t \<in> p. fst t = vec1 x \<and> content (snd t) \<noteq> 0}"
-            have pa:"\<And>k. (vec1 a, k) \<in> p \<Longrightarrow> \<exists>v. k = {vec1 a .. v} \<and> vec1 a \<le> v" 
+          proof- let ?B = "\<lambda>x. {t \<in> p. fst t = x \<and> content (snd t) \<noteq> 0}"
+            have pa:"\<And>k. (a, k) \<in> p \<Longrightarrow> \<exists>v. k = {a .. v} \<and> a \<le> v" 
             proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this
               have *:"u \<le> v" using p(2)[OF goal1] unfolding uv by auto
-              have u:"u = vec1 a" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
-                have "u \<ge> vec1 a" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "u\<noteq>vec1 a" ultimately
-                have "u > vec1 a" unfolding Cart_simps by auto
-                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps)
+              have u:"u = a" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
+                have "u \<ge> a" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "u\<noteq>a" ultimately
+                have "u > a" by auto
+                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:)
               qed thus ?case apply(rule_tac x=v in exI) unfolding uv using * by auto
             qed
-            have pb:"\<And>k. (vec1 b, k) \<in> p \<Longrightarrow> \<exists>v. k = {v .. vec1 b} \<and> vec1 b \<ge> v" 
+            have pb:"\<And>k. (b, k) \<in> p \<Longrightarrow> \<exists>v. k = {v .. b} \<and> b \<ge> v" 
             proof- case goal1 guess u v using p(4)[OF goal1] apply-by(erule exE)+ note uv=this
               have *:"u \<le> v" using p(2)[OF goal1] unfolding uv by auto
-              have u:"v = vec1 b" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
-                have "v \<le> vec1 b" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "v\<noteq>vec1 b" ultimately
-                have "v < vec1 b" unfolding Cart_simps by auto
-                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:Cart_simps)
+              have u:"v =  b" proof(rule ccontr)  have "u \<in> {u..v}" using p(2-3)[OF goal1(1)] unfolding uv by auto 
+                have "v \<le>  b" using p(2-3)[OF goal1(1)] unfolding uv subset_eq by auto moreover assume "v\<noteq> b" ultimately
+                have "v <  b" by auto
+                thus False using p(2)[OF goal1(1)] unfolding uv by(auto simp add:)
               qed thus ?case apply(rule_tac x=u in exI) unfolding uv using * by auto
             qed
 
             show "\<forall>x y. x \<in> ?B a \<and> y \<in> ?B a \<longrightarrow> x = y" apply(rule,rule,rule,unfold split_paired_all)
               unfolding mem_Collect_eq fst_conv snd_conv apply safe
-            proof- fix x k k' assume k:"(vec1 a, k) \<in> p" "(vec1 a, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
+            proof- fix x k k' assume k:"( a, k) \<in> p" "( a, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
               guess v using pa[OF k(1)] .. note v = conjunctD2[OF this]
-              guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (min (v$1) (v'$1))"
-              have "{vec1 a <..< ?v} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter]
-              moreover have "vec1 ((a + ?v$1)/2) \<in> {vec1 a <..< ?v}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps)
-              ultimately have "vec1 ((a + ?v$1)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
+              guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = " (min (v) (v'))"
+              have "{ a <..< ?v} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:) note subset_interior[OF this,unfolded interior_inter]
+              moreover have " ((a + ?v)/2) \<in> { a <..< ?v}" using k(3-)
+                unfolding v v' content_eq_0 not_le by(auto simp add:not_le)
+              ultimately have " ((a + ?v)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
               hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto
               { assume "x\<in>k" thus "x\<in>k'" unfolding * . } { assume "x\<in>k'" thus "x\<in>k" unfolding * . }
             qed 
             show "\<forall>x y. x \<in> ?B b \<and> y \<in> ?B b \<longrightarrow> x = y" apply(rule,rule,rule,unfold split_paired_all)
               unfolding mem_Collect_eq fst_conv snd_conv apply safe
-            proof- fix x k k' assume k:"(vec1 b, k) \<in> p" "(vec1 b, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
+            proof- fix x k k' assume k:"( b, k) \<in> p" "( b, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
               guess v using pb[OF k(1)] .. note v = conjunctD2[OF this]
-              guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "vec1 (max (v$1) (v'$1))"
-              have "{?v <..< vec1 b} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:Cart_simps) note subset_interior[OF this,unfolded interior_inter]
-              moreover have "vec1 ((b + ?v$1)/2) \<in> {?v <..< vec1 b}" using k(3-) unfolding v v' content_eq_0_1 not_le by(auto simp add:Cart_simps)
-              ultimately have "vec1 ((b + ?v$1)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
+              guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = " (max (v) (v'))"
+              have "{?v <..<  b} \<subseteq> k \<inter> k'" unfolding v v' by(auto simp add:) note subset_interior[OF this,unfolded interior_inter]
+              moreover have " ((b + ?v)/2) \<in> {?v <..<  b}" using k(3-) unfolding v v' content_eq_0 not_le by auto
+              ultimately have " ((b + ?v)/2) \<in> interior k \<inter> interior k'" unfolding interior_open[OF open_interval] by auto
               hence *:"k = k'" apply- apply(rule ccontr) using p(5)[OF k(1-2)] by auto
               { assume "x\<in>k" thus "x\<in>k'" unfolding * . } { assume "x\<in>k'" thus "x\<in>k" unfolding * . }
             qed
 
             let ?a = a and ?b = b (* a is something else while proofing the next theorem. *)
-            show "\<forall>x. x \<in> ?B a \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x)
-              \<le> e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1
-            proof- case goal1 guess v using pa[OF goal1(1)] .. note v = conjunctD2[OF this]
-              have "vec1 ?a\<in>{vec1 ?a..v}" using v(2) by auto hence "dest_vec1 v \<le> ?b" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
-              moreover have "{?a..dest_vec1 v} \<subseteq> ball ?a da" using fineD[OF as(2) goal1(1)]
-                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE)
-                by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately
-              show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply-
-                apply(rule da(2)[of "v$1",unfolded vec1_dest_vec1])
-                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto
+            show "\<forall>x. x \<in> ?B a \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x) - (f ((interval_upperbound k)) -
+              f ((interval_lowerbound k)))) x) \<le> e * (b - a) / 4" apply(rule,rule) unfolding mem_Collect_eq
+              unfolding split_paired_all fst_conv snd_conv 
+            proof safe case goal1 guess v using pa[OF goal1(1)] .. note v = conjunctD2[OF this]
+              have " ?a\<in>{ ?a..v}" using v(2) by auto hence "v \<le> ?b" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
+              moreover have "{?a..v} \<subseteq> ball ?a da" using fineD[OF as(2) goal1(1)]
+                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x=" x" in ballE)
+                by(auto simp add:subset_eq dist_real_def v) ultimately
+              show ?case unfolding v interval_bounds_real[OF v(2)] apply- apply(rule da(2)[of "v"])
+                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0 by auto
             qed
-            show "\<forall>x. x \<in> ?B b \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x$1) - (f ((interval_upperbound k)$1) - f ((interval_lowerbound k)$1))) x)
-              \<le> e * (b - a) / 4" apply safe unfolding fst_conv snd_conv apply safe unfolding vec1_dest_vec1
-            proof- case goal1 guess v using pb[OF goal1(1)] .. note v = conjunctD2[OF this]
-              have "vec1 ?b\<in>{v..vec1 ?b}" using v(2) by auto hence "dest_vec1 v \<ge> ?a" using p(3)[OF goal1(1)] unfolding subset_eq v by auto
-              moreover have "{dest_vec1 v..?b} \<subseteq> ball ?b db" using fineD[OF as(2) goal1(1)]
-                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe apply(erule_tac x="vec1 x" in ballE) using ab
-                by(auto simp add:Cart_simps subset_eq dist_real v dist_real_def) ultimately
-              show ?case unfolding v unfolding interval_bounds[OF v(2)[unfolded v vector_le_def]] vec1_dest_vec1 apply-
-                apply(rule db(2)[of "v$1",unfolded vec1_dest_vec1])
-                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0_1 by auto
+            show "\<forall>x. x \<in> ?B b \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' (x) -
+              (f ((interval_upperbound k)) - f ((interval_lowerbound k)))) x) \<le> e * (b - a) / 4"
+              apply(rule,rule) unfolding mem_Collect_eq unfolding split_paired_all fst_conv snd_conv 
+            proof safe case goal1 guess v using pb[OF goal1(1)] .. note v = conjunctD2[OF this]
+              have " ?b\<in>{v.. ?b}" using v(2) by auto hence "v \<ge> ?a" using p(3)[OF goal1(1)]
+                unfolding subset_eq v by auto
+              moreover have "{v..?b} \<subseteq> ball ?b db" using fineD[OF as(2) goal1(1)]
+                apply-apply(subst(asm) if_P,rule refl) unfolding subset_eq apply safe
+                apply(erule_tac x=" x" in ballE) using ab
+                by(auto simp add:subset_eq v dist_real_def) ultimately
+              show ?case unfolding v unfolding interval_bounds_real[OF v(2)] apply- apply(rule db(2)[of "v"])
+                using goal1 fineD[OF as(2) goal1(1)] unfolding v content_eq_0 by auto
             qed
           qed(insert p(1) ab e, auto simp add:field_simps) qed auto qed qed qed qed
 
@@ -3534,7 +3587,7 @@
 lemma fundamental_theorem_of_calculus_interior_strong: fixes f::"real \<Rightarrow> 'a::banach"
   assumes"finite s" "a \<le> b" "continuous_on {a..b} f"
   "\<forall>x\<in>{a<..<b} - s. (f has_vector_derivative f'(x)) (at x)"
-  shows "((f' o dest_vec1) has_integral (f b - f a)) {vec a..vec b}" using assms apply- 
+  shows "(f' has_integral (f b - f a)) {a..b}" using assms apply- 
 proof(induct "card s" arbitrary:s a b)
   case 0 show ?case apply(rule fundamental_theorem_of_calculus_interior) using 0 by auto
 next case (Suc n) from this(2) guess c s' apply-apply(subst(asm) eq_commute) unfolding card_Suc_eq
@@ -3543,7 +3596,7 @@
     case False thus ?thesis apply- apply(rule Suc(1)[OF cs(3) _ Suc(4,5)]) apply safe defer
       apply(rule Suc(6)[rule_format]) using Suc(3) unfolding cs by auto
   next have *:"f b - f a = (f c - f a) + (f b - f c)" by auto
-    case True hence "vec1 a \<le> vec1 c" "vec1 c \<le> vec1 b" by auto
+    case True hence "a \<le> c" "c \<le> b" by auto
     thus ?thesis apply(subst *) apply(rule has_integral_combine) apply assumption+
       apply(rule_tac[!] Suc(1)[OF cs(3)]) using Suc(3) unfolding cs
     proof- show "continuous_on {a..c} f" "continuous_on {c..b} f"
@@ -3555,20 +3608,20 @@
 lemma fundamental_theorem_of_calculus_strong: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "finite s" "a \<le> b" "continuous_on {a..b} f"
   "\<forall>x\<in>{a..b} - s. (f has_vector_derivative f'(x)) (at x)"
-  shows "((f' o dest_vec1) has_integral (f(b) - f(a))) {vec1 a..vec1 b}"
+  shows "(f' has_integral (f(b) - f(a))) {a..b}"
   apply(rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f'])
   using assms(4) by auto
 
-lemma indefinite_integral_continuous_left: fixes f::"real^1 \<Rightarrow> 'a::banach"
+lemma indefinite_integral_continuous_left: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" "a < c" "c \<le> b" "0 < e"
-  obtains d where "0 < d" "\<forall>t. c$1 - d < t$1 \<and> t \<le> c \<longrightarrow> norm(integral {a..c} f - integral {a..t} f) < e"
-proof- have "\<exists>w>0. \<forall>t. c$1 - w < t$1 \<and> t < c \<longrightarrow> norm(f c) * norm(c - t) < e / 3"
+  obtains d where "0 < d" "\<forall>t. c - d < t \<and> t \<le> c \<longrightarrow> norm(integral {a..c} f - integral {a..t} f) < e"
+proof- have "\<exists>w>0. \<forall>t. c - w < t \<and> t < c \<longrightarrow> norm(f c) * norm(c - t) < e / 3"
   proof(cases "f c = 0") case False hence "0 < e / 3 / norm (f c)"
       apply-apply(rule divide_pos_pos) using `e>0` by auto
     thus ?thesis apply-apply(rule,rule,assumption,safe)
-    proof- fix t assume as:"t < c" and "c$1 - e / 3 / norm (f c) < t$(1::1)"
-      hence "c$1 - t$1 < e / 3 / norm (f c)" by auto
-      hence "norm (c - t) < e / 3 / norm (f c)" using as unfolding norm_vector_1 vector_less_def by auto
+    proof- fix t assume as:"t < c" and "c - e / 3 / norm (f c) < t"
+      hence "c - t < e / 3 / norm (f c)" by auto
+      hence "norm (c - t) < e / 3 / norm (f c)" using as by auto
       thus "norm (f c) * norm (c - t) < e / 3" using False apply-
         apply(subst mult_commute) apply(subst pos_less_divide_eq[THEN sym]) by auto
     qed next case True show ?thesis apply(rule_tac x=1 in exI) unfolding True using `e>0` by auto
@@ -3582,12 +3635,12 @@
   note this[unfolded gauge_def,rule_format,of c] note conjunctD2[OF this]
   from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k .. note k=conjunctD2[OF this]
 
-  let ?d = "min k (c$1 - a$1)/2" show ?thesis apply(rule that[of ?d])
-  proof safe show "?d > 0" using k(1) using assms(2) unfolding vector_less_def by auto
-    fix t assume as:"c$1 - ?d < t$1" "t \<le> c" let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e"
+  let ?d = "min k (c - a)/2" show ?thesis apply(rule that[of ?d])
+  proof safe show "?d > 0" using k(1) using assms(2) by auto
+    fix t assume as:"c - ?d < t" "t \<le> c" let ?thesis = "norm (integral {a..c} f - integral {a..t} f) < e"
     { presume *:"t < c \<Longrightarrow> ?thesis"
       show ?thesis apply(cases "t = c") defer apply(rule *)
-        unfolding vector_less_def apply(subst less_le) using `e>0` as(2) by auto } assume "t < c"
+        apply(subst less_le) using `e>0` as(2) by auto } assume "t < c"
 
     have "f integrable_on {a..t}" apply(rule integrable_subinterval[OF assms(1)]) using assms(2-3) as(2) by auto
     from integrable_integral[OF this,unfolded has_integral,rule_format,OF *] guess d2 ..
@@ -3596,82 +3649,79 @@
     have "gauge d3" using d2(1) d1(1) unfolding d3_def gauge_def by auto
     from fine_division_exists[OF this, of a t] guess p . note p=this
     note p'=tagged_division_ofD[OF this(1)]
-    have pt:"\<forall>(x,k)\<in>p. x$1 \<le> t$1" proof safe case goal1 from p'(2,3)[OF this] show ?case by auto qed
+    have pt:"\<forall>(x,k)\<in>p. x \<le> t" proof safe case goal1 from p'(2,3)[OF this] show ?case by auto qed
     with p(2) have "d2 fine p" unfolding fine_def d3_def apply safe apply(erule_tac x="(a,b)" in ballE)+ by auto
     note d2_fin = d2(2)[OF conjI[OF p(1) this]]
     
-    have *:"{a..c} \<inter> {x. x$1 \<le> t$1} = {a..t}" "{a..c} \<inter> {x. x$1 \<ge> t$1} = {t..c}"
+    have *:"{a..c} \<inter> {x. x $$0 \<le> t} = {a..t}" "{a..c} \<inter> {x. x$$0 \<ge> t} = {t..c}"
       using assms(2-3) as by(auto simp add:field_simps)
     have "p \<union> {(c, {t..c})} tagged_division_of {a..c} \<and> d1 fine p \<union> {(c, {t..c})}" apply rule
-      apply(rule tagged_division_union_interval[of _ _ _ 1 "t$1"]) unfolding * apply(rule p)
+      apply(rule tagged_division_union_interval[of _ _ _ 0 "t"]) unfolding * apply(rule p)
       apply(rule tagged_division_of_self) unfolding fine_def
     proof safe fix x k y assume "(x,k)\<in>p" "y\<in>k" thus "y\<in>d1 x"
         using p(2) pt unfolding fine_def d3_def apply- apply(erule_tac x="(x,k)" in ballE)+ by auto
-    next fix x assume "x\<in>{t..c}" hence "dist c x < k" unfolding dist_real
+    next fix x assume "x\<in>{t..c}" hence "dist c x < k" unfolding dist_real_def
         using as(1) by(auto simp add:field_simps) 
       thus "x \<in> d1 c" using k(2) unfolding d_def by auto
     qed(insert as(2), auto) note d1_fin = d1(2)[OF this]
 
-    have *:"integral{a..c} f - integral {a..t} f = -(((c$1 - t$1) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)) -
-        integral {a..c} f) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - integral {a..t} f) + (c$1 - t$1) *\<^sub>R f c" 
+    have *:"integral{a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)) -
+        integral {a..c} f) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - integral {a..t} f) + (c - t) *\<^sub>R f c" 
       "e = (e/3 + e/3) + e/3" by auto
-    have **:"(\<Sum>(x, k)\<in>p \<union> {(c, {t..c})}. content k *\<^sub>R f x) = (c$1 - t$1) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
+    have **:"(\<Sum>(x, k)\<in>p \<union> {(c, {t..c})}. content k *\<^sub>R f x) = (c - t) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
     proof- have **:"\<And>x F. F \<union> {x} = insert x F" by auto
       have "(c, {t..c}) \<notin> p" proof safe case goal1 from p'(2-3)[OF this]
-        have "c \<in> {a..t}" by auto thus False using `t<c` unfolding vector_less_def by auto
+        have "c \<in> {a..t}" by auto thus False using `t<c` by auto
       qed thus ?thesis unfolding ** apply- apply(subst setsum_insert) apply(rule p')
-        unfolding split_conv defer apply(subst content_1) using as(2) by auto qed 
-
-    have ***:"c$1 - w < t$1 \<and> t < c"
-    proof- have "c$1 - k < t$1" using `k>0` as(1) by(auto simp add:field_simps)
+        unfolding split_conv defer apply(subst content_real) using as(2) by auto qed 
+
+    have ***:"c - w < t \<and> t < c"
+    proof- have "c - k < t" using `k>0` as(1) by(auto simp add:field_simps)
       moreover have "k \<le> w" apply(rule ccontr) using k(2) 
-        unfolding subset_eq apply(erule_tac x="c + vec ((k + w)/2)" in ballE)
-        unfolding d_def using `k>0` `w>0` by(auto simp add:field_simps not_le not_less dist_real)
+        unfolding subset_eq apply(erule_tac x="c + ((k + w)/2)" in ballE)
+        unfolding d_def using `k>0` `w>0` by(auto simp add:field_simps not_le not_less dist_real_def)
       ultimately show  ?thesis using `t<c` by(auto simp add:field_simps) qed
 
     show ?thesis unfolding *(1) apply(subst *(2)) apply(rule norm_triangle_lt add_strict_mono)+
       unfolding norm_minus_cancel apply(rule d1_fin[unfolded **]) apply(rule d2_fin)
-      using w(2)[OF ***] unfolding norm_scaleR norm_real by(auto simp add:field_simps) qed qed 
-
-lemma indefinite_integral_continuous_right: fixes f::"real^1 \<Rightarrow> 'a::banach"
+      using w(2)[OF ***] unfolding norm_scaleR by(auto simp add:field_simps) qed qed 
+
+lemma indefinite_integral_continuous_right: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" "a \<le> c" "c < b" "0 < e"
-  obtains d where "0 < d" "\<forall>t. c \<le> t \<and> t$1 < c$1 + d \<longrightarrow> norm(integral{a..c} f - integral{a..t} f) < e"
-proof- have *:"(\<lambda>x. f (- x)) integrable_on {- b..- a}" "- b < - c" "- c \<le> - a"
-    using assms unfolding Cart_simps by auto
-  from indefinite_integral_continuous_left[OF * `e>0`] guess d . note d = this let ?d = "min d (b$1 - c$1)"
+  obtains d where "0 < d" "\<forall>t. c \<le> t \<and> t < c + d \<longrightarrow> norm(integral{a..c} f - integral{a..t} f) < e"
+proof- have *:"(\<lambda>x. f (- x)) integrable_on {- b..- a}" "- b < - c" "- c \<le> - a" using assms by auto
+  from indefinite_integral_continuous_left[OF * `e>0`] guess d . note d = this let ?d = "min d (b - c)"
   show ?thesis apply(rule that[of "?d"])
-  proof safe show "0 < ?d" using d(1) assms(3) unfolding Cart_simps by auto
-    fix t::"_^1" assume as:"c \<le> t" "t$1 < c$1 + ?d"
+  proof safe show "0 < ?d" using d(1) assms(3) by auto
+    fix t::"real" assume as:"c \<le> t" "t < c + ?d"
     have *:"integral{a..c} f = integral{a..b} f - integral{c..b} f"
       "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding algebra_simps
-      apply(rule_tac[!] integral_combine) using assms as unfolding Cart_simps by auto
-    have "(- c)$1 - d < (- t)$1 \<and> - t \<le> - c" using as by auto note d(2)[rule_format,OF this]
+      apply(rule_tac[!] integral_combine) using assms as by auto
+    have "(- c) - d < (- t) \<and> - t \<le> - c" using as by auto note d(2)[rule_format,OF this]
     thus "norm (integral {a..c} f - integral {a..t} f) < e" unfolding * 
       unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:algebra_simps) qed qed
-
-declare dest_vec1_eq[simp del] not_less[simp] not_le[simp]
-
-lemma indefinite_integral_continuous: fixes f::"real^1 \<Rightarrow> 'a::banach"
+   
+lemma indefinite_integral_continuous: fixes f::"real \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" shows  "continuous_on {a..b} (\<lambda>x. integral {a..x} f)"
 proof(unfold continuous_on_iff, safe)  fix x e assume as:"x\<in>{a..b}" "0<(e::real)"
   let ?thesis = "\<exists>d>0. \<forall>x'\<in>{a..b}. dist x' x < d \<longrightarrow> dist (integral {a..x'} f) (integral {a..x} f) < e"
   { presume *:"a<b \<Longrightarrow> ?thesis"
     show ?thesis apply(cases,rule *,assumption)
-    proof- case goal1 hence "{a..b} = {x}" using as(1) unfolding Cart_simps  
-        by(auto simp only:field_simps not_less Cart_eq forall_1 mem_interval) 
+    proof- case goal1 hence "{a..b} = {x}" using as(1) apply-apply(rule set_ext)
+        unfolding atLeastAtMost_iff by(auto simp only:field_simps not_less DIM_real)
       thus ?case using `e>0` by auto
     qed } assume "a<b"
-  have "(x=a \<or> x=b) \<or> (a<x \<and> x<b)" using as(1) by (auto simp add: Cart_simps)
+  have "(x=a \<or> x=b) \<or> (a<x \<and> x<b)" using as(1) by (auto simp add:)
   thus ?thesis apply-apply(erule disjE)+
   proof- assume "x=a" have "a \<le> a" by auto
     from indefinite_integral_continuous_right[OF assms(1) this `a<b` `e>0`] guess d . note d=this
     show ?thesis apply(rule,rule,rule d,safe) apply(subst dist_commute)
-      unfolding `x=a` dist_norm apply(rule d(2)[rule_format]) unfolding norm_real by auto
+      unfolding `x=a` dist_norm apply(rule d(2)[rule_format]) by auto
   next   assume "x=b" have "b \<le> b" by auto
     from indefinite_integral_continuous_left[OF assms(1) `a<b` this `e>0`] guess d . note d=this
     show ?thesis apply(rule,rule,rule d,safe) apply(subst dist_commute)
-      unfolding `x=b` dist_norm apply(rule d(2)[rule_format]) unfolding norm_real by auto
-  next assume "a<x \<and> x<b" hence xl:"a<x" "x\<le>b" and xr:"a\<le>x" "x<b" by(auto simp add: vector_less_def)
+      unfolding `x=b` dist_norm apply(rule d(2)[rule_format])  by auto
+  next assume "a<x \<and> x<b" hence xl:"a<x" "x\<le>b" and xr:"a\<le>x" "x<b" by(auto simp add: )
     from indefinite_integral_continuous_left [OF assms(1) xl `e>0`] guess d1 . note d1=this
     from indefinite_integral_continuous_right[OF assms(1) xr `e>0`] guess d2 . note d2=this
     show ?thesis apply(rule_tac x="min d1 d2" in exI)
@@ -3679,7 +3729,7 @@
       fix y assume "y\<in>{a..b}" "dist y x < min d1 d2"
       thus "dist (integral {a..y} f) (integral {a..x} f) < e" apply-apply(subst dist_commute)
         apply(cases "y < x") unfolding dist_norm apply(rule d1(2)[rule_format]) defer
-        apply(rule d2(2)[rule_format]) unfolding Cart_simps not_less norm_real by(auto simp add:field_simps)
+        apply(rule d2(2)[rule_format]) unfolding not_less by(auto simp add:field_simps)
     qed qed qed 
 
 subsection {* This doesn't directly involve integration, but that gives an easy proof. *}
@@ -3689,23 +3739,19 @@
   "\<forall>x\<in>({a..b} - k). (f has_derivative (\<lambda>h. 0)) (at x within {a..b})" "x \<in> {a..b}"
   shows "f x = y"
 proof- have ab:"a\<le>b" using assms by auto
-  have *:"(\<lambda>x. 0\<Colon>'a) \<circ> dest_vec1 = (\<lambda>x. 0)" unfolding o_def by auto have **:"a \<le> x" using assms by auto
-  have "((\<lambda>x. 0\<Colon>'a) \<circ> dest_vec1 has_integral f x - f a) {vec1 a..vec1 x}"
-    apply(rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) ** ])
+  have *:"a\<le>x" using assms(5) by auto
+  have "((\<lambda>x. 0\<Colon>'a) has_integral f x - f a) {a..x}"
+    apply(rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *])
     apply(rule continuous_on_subset[OF assms(2)]) defer
     apply safe unfolding has_vector_derivative_def apply(subst has_derivative_within_open[THEN sym])
-    apply assumption apply(rule open_interval_real) apply(rule has_derivative_within_subset[where s="{a..b}"])
+    apply assumption apply(rule open_interval) apply(rule has_derivative_within_subset[where s="{a..b}"])
     using assms(4) assms(5) by auto note this[unfolded *]
   note has_integral_unique[OF has_integral_0 this]
   thus ?thesis unfolding assms by auto qed
 
 subsection {* Generalize a bit to any convex set. *}
 
-lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
-  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
-  scaleR_cancel_left scaleR_cancel_right scaleR.add_right scaleR.add_left real_vector_class.scaleR_one
-
-lemma has_derivative_zero_unique_strong_convex: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_derivative_zero_unique_strong_convex: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
   assumes "convex s" "finite k" "continuous_on s f" "c \<in> s" "f c = y"
   "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)" "x \<in> s"
   shows "f x = y"
@@ -3741,7 +3787,7 @@
 subsection {* Also to any open connected set with finite set of exceptions. Could 
  generalize to locally convex set with limpt-free set of exceptions. *}
 
-lemma has_derivative_zero_unique_strong_connected: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_derivative_zero_unique_strong_connected: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
   assumes "connected s" "open s" "finite k" "continuous_on s f" "c \<in> s" "f c = y"
   "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)" "x\<in>s"
   shows "f x = y"
@@ -3762,7 +3808,7 @@
 
 subsection {* Integrating characteristic function of an interval. *}
 
-lemma has_integral_restrict_open_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_restrict_open_subinterval: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
   assumes "(f has_integral i) {c..d}" "{c..d} \<subseteq> {a..b}"
   shows "((\<lambda>x. if x \<in> {c<..<d} then f x else 0) has_integral i) {a..b}"
 proof- def g \<equiv> "\<lambda>x. if x \<in>{c<..<d} then f x else 0"
@@ -3799,14 +3845,14 @@
   ultimately show ?P unfolding operat apply- apply(subst *) apply(subst iterate_insert) apply rule+
     unfolding iterate defer apply(subst if_not_P) defer using p by auto qed
 
-lemma has_integral_restrict_closed_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_restrict_closed_subinterval: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach"
   assumes "(f has_integral i) ({c..d})" "{c..d} \<subseteq> {a..b}" 
   shows "((\<lambda>x. if x \<in> {c..d} then f x else 0) has_integral i) {a..b}"
 proof- note has_integral_restrict_open_subinterval[OF assms]
   note * = has_integral_spike[OF negligible_frontier_interval _ this]
   show ?thesis apply(rule *[of c d]) using interval_open_subset_closed[of c d] by auto qed
 
-lemma has_integral_restrict_closed_subintervals_eq: fixes f::"real^'n \<Rightarrow> 'a::banach" assumes "{c..d} \<subseteq> {a..b}" 
+lemma has_integral_restrict_closed_subintervals_eq: fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::banach" assumes "{c..d} \<subseteq> {a..b}" 
   shows "((\<lambda>x. if x \<in> {c..d} then f x else 0) has_integral i) {a..b} \<longleftrightarrow> (f has_integral i) {c..d}" (is "?l = ?r")
 proof(cases "{c..d} = {}") case False let ?g = "\<lambda>x. if x \<in> {c..d} then f x else 0"
   show ?thesis apply rule defer apply(rule has_integral_restrict_closed_subinterval[OF _ assms])
@@ -3819,7 +3865,7 @@
 
 subsection {* Hence we can apply the limit process uniformly to all integrals. *}
 
-lemma has_integral': fixes f::"real^'n \<Rightarrow> 'a::banach" shows
+lemma has_integral': fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" shows
  "(f has_integral i) s \<longleftrightarrow> (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b}
   \<longrightarrow> (\<exists>z. ((\<lambda>x. if x \<in> s then f(x) else 0) has_integral z) {a..b} \<and> norm(z - i) < e))" (is "?l \<longleftrightarrow> (\<forall>e>0. ?r e)")
 proof- { presume *:"\<exists>a b. s = {a..b} \<Longrightarrow> ?thesis"
@@ -3830,7 +3876,7 @@
   note B = conjunctD2[OF this,rule_format] show ?thesis apply safe
   proof- fix e assume ?l "e>(0::real)"
     show "?r e" apply(rule_tac x="B+1" in exI) apply safe defer apply(rule_tac x=i in exI)
-    proof fix c d assume as:"ball 0 (B+1) \<subseteq> {c..d::real^'n}"
+    proof fix c d assume as:"ball 0 (B+1) \<subseteq> {c..d::'n::ordered_euclidean_space}"
       thus "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) {c..d}" unfolding s
         apply-apply(rule has_integral_restrict_closed_subinterval) apply(rule `?l`[unfolded s])
         apply safe apply(drule B(2)[rule_format]) unfolding subset_eq apply(erule_tac x=x in ballE)
@@ -3838,7 +3884,7 @@
     qed(insert B `e>0`, auto)
   next assume as:"\<forall>e>0. ?r e" 
     from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format]
-    def c \<equiv> "(\<chi> i. - max B C)::real^'n" and d \<equiv> "(\<chi> i. max B C)::real^'n"
+    def c \<equiv> "(\<chi>\<chi> i. - max B C)::'n::ordered_euclidean_space" and d \<equiv> "(\<chi>\<chi> i. max B C)::'n::ordered_euclidean_space"
     have c_d:"{a..b} \<subseteq> {c..d}" apply safe apply(drule B(2)) unfolding mem_interval
     proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def
         by(auto simp add:field_simps) qed
@@ -3850,7 +3896,7 @@
 
     have "y = i" proof(rule ccontr) assume "y\<noteq>i" hence "0 < norm (y - i)" by auto
       from as[rule_format,OF this] guess C ..  note C=conjunctD2[OF this,rule_format]
-      def c \<equiv> "(\<chi> i. - max B C)::real^'n" and d \<equiv> "(\<chi> i. max B C)::real^'n"
+      def c \<equiv> "(\<chi>\<chi> i. - max B C)::'n::ordered_euclidean_space" and d \<equiv> "(\<chi>\<chi> i. max B C)::'n::ordered_euclidean_space"
       have c_d:"{a..b} \<subseteq> {c..d}" apply safe apply(drule B(2)) unfolding mem_interval
       proof case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def
           by(auto simp add:field_simps) qed
@@ -3862,7 +3908,7 @@
       thus False by auto qed
     thus ?l using y unfolding s by auto qed qed 
 
-lemma has_integral_trans[simp]: fixes f::"real^'n \<Rightarrow> real" shows
+(*lemma has_integral_trans[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real" shows
   "((\<lambda>x. vec1 (f x)) has_integral vec1 i) s \<longleftrightarrow> (f has_integral i) s"
   unfolding has_integral'[unfolded has_integral] 
 proof case goal1 thus ?case apply safe
@@ -3884,32 +3930,31 @@
     Cart_nth.scaleR vec1_dest_vec1 zero_index apply assumption
   apply(subst norm_vector_1) by auto qed
 
-lemma integral_trans[simp]: assumes "(f::real^'n \<Rightarrow> real) integrable_on s"
+lemma integral_trans[simp]: assumes "(f::'n::ordered_euclidean_space \<Rightarrow> real) integrable_on s"
   shows "integral s (\<lambda>x. vec1 (f x)) = vec1 (integral s f)"
   apply(rule integral_unique) using assms by auto
 
-lemma integrable_on_trans[simp]: fixes f::"real^'n \<Rightarrow> real" shows
+lemma integrable_on_trans[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real" shows
   "(\<lambda>x. vec1 (f x)) integrable_on s \<longleftrightarrow> (f integrable_on s)"
   unfolding integrable_on_def
   apply(subst(2) vec1_dest_vec1(1)[THEN sym]) unfolding has_integral_trans
-  apply safe defer apply(rule_tac x="vec1 y" in exI) by auto
-
-lemma has_integral_le: fixes f::"real^'n \<Rightarrow> real"
+  apply safe defer apply(rule_tac x="vec1 y" in exI) by auto *)
+
+lemma has_integral_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "(f has_integral i) s" "(g has_integral j) s"  "\<forall>x\<in>s. (f x) \<le> (g x)"
-  shows "i \<le> j" using has_integral_component_le[of "vec1 o f" "vec1 i" s "vec1 o g" "vec1 j" 1]
-  unfolding o_def using assms by auto 
-
-lemma integral_le: fixes f::"real^'n \<Rightarrow> real"
+  shows "i \<le> j" using has_integral_component_le[OF assms(1-2), of 0] using assms(3) by auto
+
+lemma integral_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. f x \<le> g x"
   shows "integral s f \<le> integral s g"
   using has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)] .
 
-lemma has_integral_nonneg: fixes f::"real^'n \<Rightarrow> real"
+lemma has_integral_nonneg: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "(f has_integral i) s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> i" 
-  using has_integral_component_nonneg[of "vec1 o f" "vec1 i" s 1]
+  using has_integral_component_nonneg[of "f" "i" s 0]
   unfolding o_def using assms by auto
 
-lemma integral_nonneg: fixes f::"real^'n \<Rightarrow> real"
+lemma integral_nonneg: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "f integrable_on s" "\<forall>x\<in>s. 0 \<le> f x" shows "0 \<le> integral s f" 
   using has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)] .
 
@@ -3920,10 +3965,10 @@
 proof- have *:"\<And>x. (if x \<in> t then if x \<in> s then f x else 0 else 0) =  (if x\<in>s then f x else 0)" using assms by auto
   show ?thesis apply(subst(2) has_integral') apply(subst has_integral') unfolding * by rule qed
 
-lemma has_integral_restrict_univ: fixes f::"real^'n \<Rightarrow> 'a::banach" shows
+lemma has_integral_restrict_univ: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" shows
   "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) UNIV \<longleftrightarrow> (f has_integral i) s" by auto
 
-lemma has_integral_on_superset: fixes f::"real^'n \<Rightarrow> 'a::banach" 
+lemma has_integral_on_superset: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" 
   assumes "\<forall>x. ~(x \<in> s) \<longrightarrow> f x = 0" "s \<subseteq> t" "(f has_integral i) s"
   shows "(f has_integral i) t"
 proof- have "(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. if x \<in> t then f x else 0)"
@@ -3931,16 +3976,16 @@
   thus ?thesis apply- using assms(3) apply(subst has_integral_restrict_univ[THEN sym])
   apply- apply(subst(asm) has_integral_restrict_univ[THEN sym]) by auto qed
 
-lemma integrable_on_superset: fixes f::"real^'n \<Rightarrow> 'a::banach" 
+lemma integrable_on_superset: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" 
   assumes "\<forall>x. ~(x \<in> s) \<longrightarrow> f x = 0" "s \<subseteq> t" "f integrable_on s"
   shows "f integrable_on t"
   using assms unfolding integrable_on_def by(auto intro:has_integral_on_superset)
 
-lemma integral_restrict_univ[intro]: fixes f::"real^'n \<Rightarrow> 'a::banach" 
+lemma integral_restrict_univ[intro]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" 
   shows "f integrable_on s \<Longrightarrow> integral UNIV (\<lambda>x. if x \<in> s then f x else 0) = integral s f"
   apply(rule integral_unique) unfolding has_integral_restrict_univ by auto
 
-lemma integrable_restrict_univ: fixes f::"real^'n \<Rightarrow> 'a::banach" shows
+lemma integrable_restrict_univ: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" shows
  "(\<lambda>x. if x \<in> s then f x else 0) integrable_on UNIV \<longleftrightarrow> f integrable_on s"
   unfolding integrable_on_def by auto
 
@@ -3949,21 +3994,21 @@
   proof safe case goal1 show ?case apply(rule has_integral_negligible[OF `?r`[rule_format,of a b]])
       unfolding indicator_def by auto qed qed auto
 
-lemma has_integral_spike_set_eq: fixes f::"real^'n \<Rightarrow> 'a::banach" 
+lemma has_integral_spike_set_eq: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" 
   assumes "negligible((s - t) \<union> (t - s))" shows "((f has_integral y) s \<longleftrightarrow> (f has_integral y) t)"
   unfolding has_integral_restrict_univ[THEN sym,of f] apply(rule has_integral_spike_eq[OF assms]) by (safe, auto split: split_if_asm)
 
-lemma has_integral_spike_set[dest]: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_spike_set[dest]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "negligible((s - t) \<union> (t - s))" "(f has_integral y) s"
   shows "(f has_integral y) t"
   using assms has_integral_spike_set_eq by auto
 
-lemma integrable_spike_set[dest]: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_spike_set[dest]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "negligible((s - t) \<union> (t - s))" "f integrable_on s"
   shows "f integrable_on t" using assms(2) unfolding integrable_on_def 
   unfolding has_integral_spike_set_eq[OF assms(1)] .
 
-lemma integrable_spike_set_eq: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_spike_set_eq: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "negligible((s - t) \<union> (t - s))"
   shows "(f integrable_on s \<longleftrightarrow> f integrable_on t)"
   apply(rule,rule_tac[!] integrable_spike_set) using assms by auto
@@ -4002,29 +4047,28 @@
 
 subsection {* More lemmas that are useful later. *}
 
-lemma has_integral_subset_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)$k"
-  shows "i$k \<le> j$k"
+lemma has_integral_subset_component_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+  assumes "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)$$k"
+  shows "i$$k \<le> j$$k"
 proof- note has_integral_restrict_univ[THEN sym, of f]
   note assms(2-3)[unfolded this] note * = has_integral_component_le[OF this]
   show ?thesis apply(rule *) using assms(1,4) by auto qed
 
-lemma has_integral_subset_le: fixes f::"real^'n \<Rightarrow> real"
+lemma has_integral_subset_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)"
-  shows "i \<le> j" using has_integral_subset_component_le[OF assms(1), of "vec1 o f" "vec1 i" "vec1 j" 1]
-  unfolding o_def using assms by auto
-
-lemma integral_subset_component_le: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "s \<subseteq> t" "f integrable_on s" "f integrable_on t" "\<forall>x \<in> t. 0 \<le> f(x)$k"
-  shows "(integral s f)$k \<le> (integral t f)$k"
+  shows "i \<le> j" using has_integral_subset_component_le[OF assms(1), of "f" "i" "j" 0] using assms by auto
+
+lemma integral_subset_component_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+  assumes "s \<subseteq> t" "f integrable_on s" "f integrable_on t" "\<forall>x \<in> t. 0 \<le> f(x)$$k"
+  shows "(integral s f)$$k \<le> (integral t f)$$k"
   apply(rule has_integral_subset_component_le) using assms by auto
 
-lemma integral_subset_le: fixes f::"real^'n \<Rightarrow> real"
+lemma integral_subset_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "s \<subseteq> t" "f integrable_on s" "f integrable_on t" "\<forall>x \<in> t. 0 \<le> f(x)"
   shows "(integral s f) \<le> (integral t f)"
   apply(rule has_integral_subset_le) using assms by auto
 
-lemma has_integral_alt': fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_alt': fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   shows "(f has_integral i) s \<longleftrightarrow> (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on {a..b}) \<and>
   (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> {a..b} \<longrightarrow> norm(integral {a..b} (\<lambda>x. if x \<in> s then f x else 0) - i) < e)" (is "?l = ?r")
 proof assume ?r
@@ -4036,9 +4080,9 @@
   qed next
   assume ?l note as = this[unfolded has_integral'[of f],rule_format]
   let ?f = "\<lambda>x. if x \<in> s then f x else 0"
-  show ?r proof safe fix a b::"real^'n"
+  show ?r proof safe fix a b::"'n::ordered_euclidean_space"
     from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format]
-    let ?a = "(\<chi> i. min (a$i) (-B))::real^'n" and ?b = "(\<chi> i. max (b$i) B)::real^'n"
+    let ?a = "(\<chi>\<chi> i. min (a$$i) (-B))::'n::ordered_euclidean_space" and ?b = "(\<chi>\<chi> i. max (b$$i) B)::'n::ordered_euclidean_space"
     show "?f integrable_on {a..b}" apply(rule integrable_subinterval[of _ ?a ?b])
     proof- have "ball 0 B \<subseteq> {?a..?b}" apply safe unfolding mem_ball mem_interval dist_norm
       proof case goal1 thus ?case using component_le_norm[of x i] by(auto simp add:field_simps) qed
@@ -4055,7 +4099,7 @@
 
 subsection {* Continuity of the integral (for a 1-dimensional interval). *}
 
-lemma integrable_alt: fixes f::"real^'n \<Rightarrow> 'a::banach" shows 
+lemma integrable_alt: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" shows 
   "f integrable_on s \<longleftrightarrow>
           (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on {a..b}) \<and>
           (\<forall>e>0. \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> {a..b} \<and> ball 0 B \<subseteq> {c..d}
@@ -4069,11 +4113,11 @@
         using B(2)[OF goal1(1)] B(2)[OF goal1(2)] by auto qed qed
         
 next assume ?r note as = conjunctD2[OF this,rule_format]
-  have "Cauchy (\<lambda>n. integral ({(\<chi> i. - real n) .. (\<chi> i. real n)}) (\<lambda>x. if x \<in> s then f x else 0))"
+  have "Cauchy (\<lambda>n. integral ({(\<chi>\<chi> i. - real n)::'n .. (\<chi>\<chi> i. real n)}) (\<lambda>x. if x \<in> s then f x else 0))"
   proof(unfold Cauchy_def,safe) case goal1
     from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format]
     from real_arch_simple[of B] guess N .. note N = this
-    { fix n assume n:"n \<ge> N" have "ball 0 B \<subseteq> {\<chi> i. - real n..\<chi> i. real n}" apply safe
+    { fix n assume n:"n \<ge> N" have "ball 0 B \<subseteq> {(\<chi>\<chi> i. - real n)::'n..\<chi>\<chi> i. real n}" apply safe
         unfolding mem_ball mem_interval dist_norm
       proof case goal1 thus ?case using component_le_norm[of x i]
           using n N by(auto simp add:field_simps) qed }
@@ -4088,33 +4132,33 @@
     from as(2)[OF *] guess B .. note B=conjunctD2[OF this,rule_format] let ?B = "max (real N) B"
     show ?case apply(rule_tac x="?B" in exI)
     proof safe show "0 < ?B" using B(1) by auto
-      fix a b assume ab:"ball 0 ?B \<subseteq> {a..b::real^'n}"
+      fix a b assume ab:"ball 0 ?B \<subseteq> {a..b::'n::ordered_euclidean_space}"
       from real_arch_simple[of ?B] guess n .. note n=this
       show "norm (integral {a..b} (\<lambda>x. if x \<in> s then f x else 0) - i) < e"
         apply(rule norm_triangle_half_l) apply(rule B(2)) defer apply(subst norm_minus_commute)
         apply(rule N[unfolded dist_norm, of n])
       proof safe show "N \<le> n" using n by auto
-        fix x::"real^'n" assume x:"x \<in> ball 0 B" hence "x\<in> ball 0 ?B" by auto
+        fix x::"'n::ordered_euclidean_space" assume x:"x \<in> ball 0 B" hence "x\<in> ball 0 ?B" by auto
         thus "x\<in>{a..b}" using ab by blast 
-        show "x\<in>{\<chi> i. - real n..\<chi> i. real n}" using x unfolding mem_interval mem_ball dist_norm apply-
+        show "x\<in>{\<chi>\<chi> i. - real n..\<chi>\<chi> i. real n}" using x unfolding mem_interval mem_ball dist_norm apply-
         proof case goal1 thus ?case using component_le_norm[of x i]
             using n by(auto simp add:field_simps) qed qed qed qed qed
 
-lemma integrable_altD: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_altD: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on s"
   shows "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on {a..b}"
   "\<And>e. e>0 \<Longrightarrow> \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> {a..b} \<and> ball 0 B \<subseteq> {c..d}
   \<longrightarrow> norm(integral {a..b} (\<lambda>x. if x \<in> s then f x else 0) - integral {c..d}  (\<lambda>x. if x \<in> s then f x else 0)) < e"
   using assms[unfolded integrable_alt[of f]] by auto
 
-lemma integrable_on_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_on_subinterval: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on s" "{a..b} \<subseteq> s" shows "f integrable_on {a..b}"
   apply(rule integrable_eq) defer apply(rule integrable_altD(1)[OF assms(1)])
   using assms(2) by auto
 
 subsection {* A straddling criterion for integrability. *}
 
-lemma integrable_straddle_interval: fixes f::"real^'n \<Rightarrow> real"
+lemma integrable_straddle_interval: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "\<forall>e>0. \<exists>g  h i j. (g has_integral i) ({a..b}) \<and> (h has_integral j) ({a..b}) \<and>
   norm(i - j) < e \<and> (\<forall>x\<in>{a..b}. (g x) \<le> (f x) \<and> (f x) \<le>(h x))"
   shows "f integrable_on {a..b}"
@@ -4150,7 +4194,7 @@
       apply(rule d2(2)[OF conjI[OF goal1(4,6)]])
       apply(rule d2(2)[OF conjI[OF goal1(1,3)]]) by auto qed qed 
      
-lemma integrable_straddle: fixes f::"real^'n \<Rightarrow> real"
+lemma integrable_straddle: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "\<forall>e>0. \<exists>g h i j. (g has_integral i) s \<and> (h has_integral j) s \<and>
   norm(i - j) < e \<and> (\<forall>x\<in>s. (g x) \<le>(f x) \<and>(f x) \<le>(h x))"
   shows "f integrable_on s"
@@ -4161,7 +4205,7 @@
     note g = this(1) and this(2)[OF *] from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format]
     note obt(2)[unfolded has_integral_alt'[of h]] note conjunctD2[OF this, rule_format]
     note h = this(1) and this(2)[OF *] from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
-    def c \<equiv> "\<chi> i. min (a$i) (- (max B1 B2))" and d \<equiv> "\<chi> i. max (b$i) (max B1 B2)"
+    def c \<equiv> "(\<chi>\<chi> i. min (a$$i) (- (max B1 B2)))::'n" and d \<equiv> "(\<chi>\<chi> i. max (b$$i) (max B1 B2))::'n"
     have *:"ball 0 B1 \<subseteq> {c..d}" "ball 0 B2 \<subseteq> {c..d}" apply safe unfolding mem_ball mem_interval dist_norm
     proof(rule_tac[!] allI)
       case goal1 thus ?case using component_le_norm[of x i] unfolding c_def d_def by auto next
@@ -4196,8 +4240,8 @@
     note obt(2)[unfolded has_integral_alt'[of h]] note conjunctD2[OF this, rule_format]
     note h = this(1) and this(2)[OF *] from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
     show ?case apply(rule_tac x="max B1 B2" in exI) apply safe apply(rule min_max.less_supI1,rule B1)
-    proof- fix a b c d::"real^'n" assume as:"ball 0 (max B1 B2) \<subseteq> {a..b}" "ball 0 (max B1 B2) \<subseteq> {c..d}"
-      have **:"ball 0 B1 \<subseteq> ball (0::real^'n) (max B1 B2)" "ball 0 B2 \<subseteq> ball (0::real^'n) (max B1 B2)" by auto
+    proof- fix a b c d::"'n::ordered_euclidean_space" assume as:"ball 0 (max B1 B2) \<subseteq> {a..b}" "ball 0 (max B1 B2) \<subseteq> {c..d}"
+      have **:"ball 0 B1 \<subseteq> ball (0::'n::ordered_euclidean_space) (max B1 B2)" "ball 0 B2 \<subseteq> ball (0::'n::ordered_euclidean_space) (max B1 B2)" by auto
       have *:"\<And>ga gc ha hc fa fc::real. abs(ga - i) < e / 3 \<and> abs(gc - i) < e / 3 \<and> abs(ha - j) < e / 3 \<and>
         abs(hc - j) < e / 3 \<and> abs(i - j) < e / 3 \<and> ga \<le> fa \<and> fa \<le> ha \<and> gc \<le> fc \<and> fc \<le> hc\<Longrightarrow> abs(fa - fc) < e" by smt
       show "norm (integral {a..b} (\<lambda>x. if x \<in> s then f x else 0) - integral {c..d} (\<lambda>x. if x \<in> s then f x else 0)) < e"
@@ -4211,14 +4255,14 @@
 
 subsection {* Adding integrals over several sets. *}
 
-lemma has_integral_union: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_union: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "(f has_integral i) s" "(f has_integral j) t" "negligible(s \<inter> t)"
   shows "(f has_integral (i + j)) (s \<union> t)"
 proof- note * = has_integral_restrict_univ[THEN sym, of f]
   show ?thesis unfolding * apply(rule has_integral_spike[OF assms(3)])
     defer apply(rule has_integral_add[OF assms(1-2)[unfolded *]]) by auto qed
 
-lemma has_integral_unions: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_unions: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "finite t" "\<forall>s\<in>t. (f has_integral (i s)) s"  "\<forall>s\<in>t. \<forall>s'\<in>t. ~(s = s') \<longrightarrow> negligible(s \<inter> s')"
   shows "(f has_integral (setsum i t)) (\<Union>t)"
 proof- note * = has_integral_restrict_univ[THEN sym, of f]
@@ -4236,7 +4280,7 @@
 
 subsection {* In particular adding integrals over a division, maybe not of an interval. *}
 
-lemma has_integral_combine_division: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_combine_division: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "d division_of s" "\<forall>k\<in>d. (f has_integral (i k)) k"
   shows "(f has_integral (setsum i d)) s"
 proof- note d = division_ofD[OF assms(1)]
@@ -4248,13 +4292,13 @@
       apply-apply(rule negligible_subset[of "({a..b}-{a<..<b}) \<union> ({c..d}-{c<..<d})"])
       apply(rule negligible_union negligible_frontier_interval)+ by auto qed qed
 
-lemma integral_combine_division_bottomup: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integral_combine_division_bottomup: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "d division_of s" "\<forall>k\<in>d. f integrable_on k"
   shows "integral s f = setsum (\<lambda>i. integral i f) d"
   apply(rule integral_unique) apply(rule has_integral_combine_division[OF assms(1)])
   using assms(2) unfolding has_integral_integral .
 
-lemma has_integral_combine_division_topdown: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_combine_division_topdown: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on s" "d division_of k" "k \<subseteq> s"
   shows "(f has_integral (setsum (\<lambda>i. integral i f) d)) k"
   apply(rule has_integral_combine_division[OF assms(2)])
@@ -4263,18 +4307,18 @@
   show ?case apply safe apply(rule integrable_on_subinterval)
     apply(rule assms) using assms(3) by auto qed
 
-lemma integral_combine_division_topdown: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integral_combine_division_topdown: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on s" "d division_of s"
   shows "integral s f = setsum (\<lambda>i. integral i f) d"
   apply(rule integral_unique,rule has_integral_combine_division_topdown) using assms by auto
 
-lemma integrable_combine_division: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_combine_division: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "d division_of s" "\<forall>i\<in>d. f integrable_on i"
   shows "f integrable_on s"
   using assms(2) unfolding integrable_on_def
   by(metis has_integral_combine_division[OF assms(1)])
 
-lemma integrable_on_subdivision: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integrable_on_subdivision: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "d division_of i" "f integrable_on s" "i \<subseteq> s"
   shows "f integrable_on i"
   apply(rule integrable_combine_division assms)+
@@ -4284,7 +4328,7 @@
 
 subsection {* Also tagged divisions. *}
 
-lemma has_integral_combine_tagged_division: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_combine_tagged_division: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "p tagged_division_of s" "\<forall>(x,k) \<in> p. (f has_integral (i k)) k"
   shows "(f has_integral (setsum (\<lambda>(x,k). i k) p)) s"
 proof- have *:"(f has_integral (setsum (\<lambda>k. integral k f) (snd ` p))) s"
@@ -4295,27 +4339,27 @@
     apply(rule setsum_over_tagged_division_lemma[OF assms(1)]) apply(rule integral_null,assumption)
     apply(rule setsum_cong2) using assms(2) by auto qed
 
-lemma integral_combine_tagged_division_bottomup: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integral_combine_tagged_division_bottomup: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "p tagged_division_of {a..b}" "\<forall>(x,k)\<in>p. f integrable_on k"
   shows "integral {a..b} f = setsum (\<lambda>(x,k). integral k f) p"
   apply(rule integral_unique) apply(rule has_integral_combine_tagged_division[OF assms(1)])
   using assms(2) by auto
 
-lemma has_integral_combine_tagged_division_topdown: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma has_integral_combine_tagged_division_topdown: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" "p tagged_division_of {a..b}"
   shows "(f has_integral (setsum (\<lambda>(x,k). integral k f) p)) {a..b}"
   apply(rule has_integral_combine_tagged_division[OF assms(2)])
 proof safe case goal1 note tagged_division_ofD(3-4)[OF assms(2) this]
   thus ?case using integrable_subinterval[OF assms(1)] by auto qed
 
-lemma integral_combine_tagged_division_topdown: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma integral_combine_tagged_division_topdown: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" "p tagged_division_of {a..b}"
   shows "integral {a..b} f = setsum (\<lambda>(x,k). integral k f) p"
   apply(rule integral_unique,rule has_integral_combine_tagged_division_topdown) using assms by auto
 
 subsection {* Henstock's lemma. *}
 
-lemma henstock_lemma_part1: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma henstock_lemma_part1: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" "0 < e" "gauge d"
   "(\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral({a..b}) f) < e)"
   and p:"p tagged_partial_division_of {a..b}" "d fine p"
@@ -4415,23 +4459,23 @@
       unfolding divide_inverse[THEN sym] using * by(auto simp add:field_simps real_eq_of_nat)
   qed finally show "?x \<le> e + k" . qed
 
-lemma henstock_lemma_part2: fixes f::"real^'m \<Rightarrow> real^'n"
+lemma henstock_lemma_part2: fixes f::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
   assumes "f integrable_on {a..b}" "0 < e" "gauge d"
   "\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow> norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p -
           integral({a..b}) f) < e"    "p tagged_partial_division_of {a..b}" "d fine p"
-  shows "setsum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (CARD('n)) * e"
-  unfolding split_def apply(rule vsum_norm_allsubsets_bound) defer 
+  shows "setsum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
+  unfolding split_def apply(rule setsum_norm_allsubsets_bound) defer 
   apply(rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)])
   apply safe apply(rule assms[rule_format,unfolded split_def]) defer
   apply(rule tagged_partial_division_subset,rule assms,assumption)
   apply(rule fine_subset,assumption,rule assms) using assms(5) by auto
   
-lemma henstock_lemma: fixes f::"real^'m \<Rightarrow> real^'n"
+lemma henstock_lemma: fixes f::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
   assumes "f integrable_on {a..b}" "e>0"
   obtains d where "gauge d"
   "\<forall>p. p tagged_partial_division_of {a..b} \<and> d fine p
   \<longrightarrow> setsum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e"
-proof- have *:"e / (2 * (real CARD('n) + 1)) > 0" apply(rule divide_pos_pos) using assms(2) by auto
+proof- have *:"e / (2 * (real DIM('n) + 1)) > 0" apply(rule divide_pos_pos) using assms(2) by auto
   from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this]
   guess d .. note d = conjunctD2[OF this] show thesis apply(rule that,rule d)
   proof safe case goal1 note * = henstock_lemma_part2[OF assms(1) * d this]
@@ -4439,16 +4483,16 @@
 
 subsection {* monotone convergence (bounded interval first). *}
 
-lemma monotone_convergence_interval: fixes f::"nat \<Rightarrow> real^'n \<Rightarrow> real^1"
+lemma monotone_convergence_interval: fixes f::"nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "\<forall>k. (f k) integrable_on {a..b}"
-  "\<forall>k. \<forall>x\<in>{a..b}. dest_vec1(f k x) \<le> dest_vec1(f (Suc k) x)"
+  "\<forall>k. \<forall>x\<in>{a..b}.(f k x) \<le> (f (Suc k) x)"
   "\<forall>x\<in>{a..b}. ((\<lambda>k. f k x) ---> g x) sequentially"
   "bounded {integral {a..b} (f k) | k . k \<in> UNIV}"
   shows "g integrable_on {a..b} \<and> ((\<lambda>k. integral ({a..b}) (f k)) ---> integral ({a..b}) g) sequentially"
 proof(case_tac[!] "content {a..b} = 0") assume as:"content {a..b} = 0"
   show ?thesis using integrable_on_null[OF as] unfolding integral_null[OF as] using Lim_const by auto
 next assume ab:"content {a..b} \<noteq> 0"
-  have fg:"\<forall>x\<in>{a..b}. \<forall> k. (f k x)$1 \<le> (g x)$1"
+  have fg:"\<forall>x\<in>{a..b}. \<forall> k. (f k x) $$ 0 \<le> (g x) $$ 0"
   proof safe case goal1 note assms(3)[rule_format,OF this]
     note * = Lim_component_ge[OF this trivial_limit_sequentially]
     show ?case apply(rule *) unfolding eventually_sequentially
@@ -4456,13 +4500,13 @@
       using assms(2)[rule_format,OF goal1] by auto qed
   have "\<exists>i. ((\<lambda>k. integral ({a..b}) (f k)) ---> i) sequentially"
     apply(rule bounded_increasing_convergent) defer
-    apply rule apply(rule integral_component_le) apply safe
+    apply rule apply(rule integral_le) apply safe
     apply(rule assms(1-2)[rule_format])+ using assms(4) by auto
   then guess i .. note i=this
-  have i':"\<And>k. dest_vec1(integral({a..b}) (f k)) \<le> dest_vec1 i"
+  have i':"\<And>k. (integral({a..b}) (f k)) \<le> i$$0"
     apply(rule Lim_component_ge,rule i) apply(rule trivial_limit_sequentially)
     unfolding eventually_sequentially apply(rule_tac x=k in exI)
-    apply(rule transitive_stepwise_le) prefer 3 apply(rule integral_dest_vec1_le)
+    apply(rule transitive_stepwise_le) prefer 3 unfolding Eucl_real_simps apply(rule integral_le)
     apply(rule assms(1-2)[rule_format])+ using assms(2) by auto
 
   have "(g has_integral i) {a..b}" unfolding has_integral
@@ -4473,23 +4517,22 @@
       apply(rule divide_pos_pos) by auto
     from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format]
 
-    have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i$1 - dest_vec1(integral {a..b} (f k)) \<and>
-                   i$1 - dest_vec1(integral {a..b} (f k)) < e / 4"
+    have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i$$0 - (integral {a..b} (f k)) \<and> i$$0 - (integral {a..b} (f k)) < e / 4"
     proof- case goal1 have "e/4 > 0" using e by auto
       from i[unfolded Lim_sequentially,rule_format,OF this] guess r ..
       thus ?case apply(rule_tac x=r in exI) apply rule
         apply(erule_tac x=k in allE)
-      proof- case goal1 thus ?case using i'[of k] unfolding dist_real by auto qed qed
+      proof- case goal1 thus ?case using i'[of k] unfolding dist_real_def by auto qed qed
     then guess r .. note r=conjunctD2[OF this[rule_format]]
 
-    have "\<forall>x\<in>{a..b}. \<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x)$1 - (f k x)$1 \<and>
-           (g x)$1 - (f k x)$1 < e / (4 * content({a..b}))"
+    have "\<forall>x\<in>{a..b}. \<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x)$$0 - (f k x)$$0 \<and>
+           (g x)$$0 - (f k x)$$0 < e / (4 * content({a..b}))"
     proof case goal1 have "e / (4 * content {a..b}) > 0" apply(rule divide_pos_pos,fact)
         using ab content_pos_le[of a b] by auto
       from assms(3)[rule_format,OF goal1,unfolded Lim_sequentially,rule_format,OF this]
       guess n .. note n=this
       thus ?case apply(rule_tac x="n + r" in exI) apply safe apply(erule_tac[2-3] x=k in allE)
-        unfolding dist_real using fg[rule_format,OF goal1] by(auto simp add:field_simps) qed
+        unfolding dist_real_def using fg[rule_format,OF goal1] by(auto simp add:field_simps) qed
     from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format]
     def d \<equiv> "\<lambda>x. c (m x) x" 
 
@@ -4516,7 +4559,7 @@
            from p'(4)[OF xk] guess u v apply-by(erule exE)+ note uv=this
            show " norm (content k *\<^sub>R (g x - f (m x) x)) \<le> content k * (e / (4 * content {a..b}))"
              unfolding norm_scaleR uv unfolding abs_of_nonneg[OF content_pos_le] 
-             apply(rule mult_left_mono) unfolding norm_real using m(2)[OF x,of "m x"] by auto
+             apply(rule mult_left_mono) using m(2)[OF x,of "m x"] by auto
          qed(insert ab,auto)
          
        next case goal2 show ?case apply(rule le_less_trans[of _ "norm (\<Sum>j = 0..s.
@@ -4547,47 +4590,47 @@
 
        next case goal3
          note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)]
-         have *:"\<And>sr sx ss ks kr::real^1. kr = sr \<longrightarrow> ks = ss \<longrightarrow> ks \<le> i \<and> sr \<le> sx \<and> sx \<le> ss \<and> 0 \<le> i$1 - kr$1
-           \<and> i$1 - kr$1 < e/4 \<longrightarrow> abs(sx$1 - i$1) < e/4" unfolding Cart_eq forall_1 vector_le_def by arith
-         show ?case unfolding norm_real Cart_nth.diff apply(rule *[rule_format],safe)
-           apply(rule comb[of r],rule comb[of s]) unfolding vector_le_def forall_1 setsum_component
-           apply(rule i') apply(rule_tac[1-2] setsum_mono) unfolding split_paired_all split_conv
-           apply(rule_tac[1-2] integral_component_le[OF ])
-         proof safe show "0 \<le> i$1 - (integral {a..b} (f r))$1" using r(1) by auto
-           show "i$1 - (integral {a..b} (f r))$1 < e / 4" using r(2) by auto
+         have *:"\<And>sr sx ss ks kr::real. kr = sr \<longrightarrow> ks = ss \<longrightarrow> ks \<le> i \<and> sr \<le> sx \<and> sx \<le> ss \<and> 0 \<le> i$$0 - kr$$0
+           \<and> i$$0 - kr$$0 < e/4 \<longrightarrow> abs(sx - i) < e/4" by auto 
+         show ?case unfolding real_norm_def apply(rule *[rule_format],safe)
+           apply(rule comb[of r],rule comb[of s]) apply(rule i'[unfolded Eucl_real_simps]) 
+           apply(rule_tac[1-2] setsum_mono) unfolding split_paired_all split_conv
+           apply(rule_tac[1-2] integral_le[OF ])
+         proof safe show "0 \<le> i$$0 - (integral {a..b} (f r))$$0" using r(1) by auto
+           show "i$$0 - (integral {a..b} (f r))$$0 < e / 4" using r(2) by auto
            fix x k assume xk:"(x,k)\<in>p" from p'(4)[OF this] guess u v apply-by(erule exE)+ note uv=this
            show "f r integrable_on k" "f s integrable_on k" "f (m x) integrable_on k" "f (m x) integrable_on k" 
              unfolding uv apply(rule_tac[!] integrable_on_subinterval[OF assms(1)[rule_format]])
              using p'(3)[OF xk] unfolding uv by auto 
            fix y assume "y\<in>k" hence "y\<in>{a..b}" using p'(3)[OF xk] by auto
-           hence *:"\<And>m. \<forall>n\<ge>m. (f m y)$1 \<le> (f n y)$1" apply-apply(rule transitive_stepwise_le) using assms(2) by auto
-           show "(f r y)$1 \<le> (f (m x) y)$1" "(f (m x) y)$1 \<le> (f s y)$1"
+           hence *:"\<And>m. \<forall>n\<ge>m. (f m y) \<le> (f n y)" apply-apply(rule transitive_stepwise_le) using assms(2) by auto
+           show "(f r y) \<le> (f (m x) y)" "(f (m x) y) \<le> (f s y)"
              apply(rule_tac[!] *[rule_format]) using s[rule_format,OF xk] m(1)[of x] p'(2-3)[OF xk] by auto
          qed qed qed qed note * = this 
 
    have "integral {a..b} g = i" apply(rule integral_unique) using * .
    thus ?thesis using i * by auto qed
 
-lemma monotone_convergence_increasing: fixes f::"nat \<Rightarrow> real^'n \<Rightarrow> real^1"
-  assumes "\<forall>k. (f k) integrable_on s"  "\<forall>k. \<forall>x\<in>s. (f k x)$1 \<le> (f (Suc k) x)$1"
+lemma monotone_convergence_increasing: fixes f::"nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
+  assumes "\<forall>k. (f k) integrable_on s"  "\<forall>k. \<forall>x\<in>s. (f k x) \<le> (f (Suc k) x)"
   "\<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially" "bounded {integral s (f k)| k. True}"
   shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
-proof- have lem:"\<And>f::nat \<Rightarrow> real^'n \<Rightarrow> real^1. \<And> g s. \<forall>k.\<forall>x\<in>s. 0\<le>dest_vec1 (f k x) \<Longrightarrow> \<forall>k. (f k) integrable_on s \<Longrightarrow>
-    \<forall>k. \<forall>x\<in>s. (f k x)$1 \<le> (f (Suc k) x)$1 \<Longrightarrow> \<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially  \<Longrightarrow>
+proof- have lem:"\<And>f::nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real. \<And> g s. \<forall>k.\<forall>x\<in>s. 0 \<le> (f k x) \<Longrightarrow> \<forall>k. (f k) integrable_on s \<Longrightarrow>
+    \<forall>k. \<forall>x\<in>s. (f k x) \<le> (f (Suc k) x) \<Longrightarrow> \<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially  \<Longrightarrow>
     bounded {integral s (f k)| k. True} \<Longrightarrow> g integrable_on s \<and> ((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
   proof- case goal1 note assms=this[rule_format]
-    have "\<forall>x\<in>s. \<forall>k. dest_vec1(f k x) \<le> dest_vec1(g x)" apply safe apply(rule Lim_component_ge)
+    have "\<forall>x\<in>s. \<forall>k. (f k x)$$0 \<le> (g x)$$0" apply safe apply(rule Lim_component_ge)
       apply(rule goal1(4)[rule_format],assumption) apply(rule trivial_limit_sequentially)
       unfolding eventually_sequentially apply(rule_tac x=k in exI)
       apply(rule transitive_stepwise_le) using goal1(3) by auto note fg=this[rule_format]
 
     have "\<exists>i. ((\<lambda>k. integral s (f k)) ---> i) sequentially" apply(rule bounded_increasing_convergent)
-      apply(rule goal1(5)) apply(rule,rule integral_component_le) apply(rule goal1(2)[rule_format])+
+      apply(rule goal1(5)) apply(rule,rule integral_le) apply(rule goal1(2)[rule_format])+
       using goal1(3) by auto then guess i .. note i=this
     have "\<And>k. \<forall>x\<in>s. \<forall>n\<ge>k. f k x \<le> f n x" apply(rule,rule transitive_stepwise_le) using goal1(3) by auto
-    hence i':"\<forall>k. (integral s (f k))$1 \<le> i$1" apply-apply(rule,rule Lim_component_ge)
+    hence i':"\<forall>k. (integral s (f k))$$0 \<le> i$$0" apply-apply(rule,rule Lim_component_ge)
       apply(rule i,rule trivial_limit_sequentially) unfolding eventually_sequentially
-      apply(rule_tac x=k in exI,safe) apply(rule integral_dest_vec1_le)
+      apply(rule_tac x=k in exI,safe) apply(rule integral_component_le)
       apply(rule goal1(2)[rule_format])+ by auto 
 
     note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format]
@@ -4602,14 +4645,14 @@
       case goal1 show ?case using int .
     next case goal2 thus ?case apply-apply(cases "x\<in>s") using assms(3) by auto
     next case goal3 thus ?case apply-apply(cases "x\<in>s") using assms(4) by auto
-    next case goal4 note * = integral_dest_vec1_nonneg[unfolded vector_le_def forall_1 zero_index]
+    next case goal4 note * = integral_nonneg
       have "\<And>k. norm (integral {a..b} (\<lambda>x. if x \<in> s then f k x else 0)) \<le> norm (integral s (f k))"
-        unfolding norm_real apply(subst abs_of_nonneg) apply(rule *[OF int])
+        unfolding real_norm_def apply(subst abs_of_nonneg) apply(rule *[OF int])
         apply(safe,case_tac "x\<in>s") apply(drule assms(1)) prefer 3
         apply(subst abs_of_nonneg) apply(rule *[OF assms(2) goal1(1)[THEN spec]])
         apply(subst integral_restrict_univ[THEN sym,OF int]) 
         unfolding ifif unfolding integral_restrict_univ[OF int']
-        apply(rule integral_subset_component_le[OF _ int' assms(2)]) using assms(1) by auto
+        apply(rule integral_subset_le[OF _ int' assms(2)]) using assms(1) by auto
       thus ?case using assms(5) unfolding bounded_iff apply safe
         apply(rule_tac x=aa in exI,safe) apply(erule_tac x="integral s (f k)" in ballE)
         apply(rule order_trans) apply assumption by auto qed note g = conjunctD2[OF this]
@@ -4620,31 +4663,31 @@
       note assms(2)[of N,unfolded has_integral_integral has_integral_alt'[of "f N"]]
       from this[THEN conjunct2,rule_format,OF `e/4>0`] guess B .. note B=conjunctD2[OF this]
       show ?case apply(rule,rule,rule B,safe)
-      proof- fix a b::"real^'n" assume ab:"ball 0 B \<subseteq> {a..b}"
+      proof- fix a b::"'n::ordered_euclidean_space" assume ab:"ball 0 B \<subseteq> {a..b}"
         from `e>0` have "e/2>0" by auto
         from g(2)[unfolded Lim_sequentially,of a b,rule_format,OF this] guess M .. note M=this
         have **:"norm (integral {a..b} (\<lambda>x. if x \<in> s then f N x else 0) - i) < e/2"
           apply(rule norm_triangle_half_l) using B(2)[rule_format,OF ab] N[rule_format,of N]
           unfolding dist_norm apply-defer apply(subst norm_minus_commute) by auto
-        have *:"\<And>f1 f2 g. abs(f1 - i$1) < e / 2 \<longrightarrow> abs(f2 - g) < e / 2 \<longrightarrow> f1 \<le> f2 \<longrightarrow> f2 \<le> i$1
-          \<longrightarrow> abs(g - i$1) < e" by arith
+        have *:"\<And>f1 f2 g. abs(f1 - i) < e / 2 \<longrightarrow> abs(f2 - g) < e / 2 \<longrightarrow> f1 \<le> f2 \<longrightarrow> f2 \<le> i
+          \<longrightarrow> abs(g - i) < e" unfolding Eucl_real_simps by arith
         show "norm (integral {a..b} (\<lambda>x. if x \<in> s then g x else 0) - i) < e" 
-          unfolding norm_real Cart_simps apply(rule *[rule_format])
-          apply(rule **[unfolded norm_real Cart_simps])
-          apply(rule M[rule_format,of "M + N",unfolded dist_real]) apply(rule le_add1)
-          apply(rule integral_component_le[OF int int]) defer
-          apply(rule order_trans[OF _ i'[rule_format,of "M + N"]])
-        proof safe case goal2 have "\<And>m. x\<in>s \<Longrightarrow> \<forall>n\<ge>m. (f m x)$1 \<le> (f n x)$1"
+          unfolding real_norm_def apply(rule *[rule_format])
+          apply(rule **[unfolded real_norm_def])
+          apply(rule M[rule_format,of "M + N",unfolded dist_real_def]) apply(rule le_add1)
+          apply(rule integral_le[OF int int]) defer
+          apply(rule order_trans[OF _ i'[rule_format,of "M + N",unfolded Eucl_real_simps]])
+        proof safe case goal2 have "\<And>m. x\<in>s \<Longrightarrow> \<forall>n\<ge>m. (f m x)$$0 \<le> (f n x)$$0"
             apply(rule transitive_stepwise_le) using assms(3) by auto thus ?case by auto
         next case goal1 show ?case apply(subst integral_restrict_univ[THEN sym,OF int]) 
             unfolding ifif integral_restrict_univ[OF int']
-            apply(rule integral_subset_component_le[OF _ int']) using assms by auto
+            apply(rule integral_subset_le[OF _ int']) using assms by auto
         qed qed qed 
     thus ?case apply safe defer apply(drule integral_unique) using i by auto qed
 
   have sub:"\<And>k. integral s (\<lambda>x. f k x - f 0 x) = integral s (f k) - integral s (f 0)"
     apply(subst integral_sub) apply(rule assms(1)[rule_format])+ by rule
-  have "\<And>x m. x\<in>s \<Longrightarrow> \<forall>n\<ge>m. dest_vec1 (f m x) \<le> dest_vec1 (f n x)" apply(rule transitive_stepwise_le)
+  have "\<And>x m. x\<in>s \<Longrightarrow> \<forall>n\<ge>m. (f m x) \<le> (f n x)" apply(rule transitive_stepwise_le)
     using assms(2) by auto note * = this[rule_format]
   have "(\<lambda>x. g x - f 0 x) integrable_on s \<and>((\<lambda>k. integral s (\<lambda>x. f (Suc k) x - f 0 x)) --->
       integral s (\<lambda>x. g x - f 0 x)) sequentially" apply(rule lem,safe)
@@ -4662,8 +4705,8 @@
   thus ?thesis unfolding sub apply-apply rule defer apply(subst(asm) integral_sub)
     using assms(1) apply auto apply(rule seq_offset_rev[where k=1]) by auto qed
 
-lemma monotone_convergence_decreasing: fixes f::"nat \<Rightarrow> real^'n \<Rightarrow> real^1"
-  assumes "\<forall>k. (f k) integrable_on s"  "\<forall>k. \<forall>x\<in>s. (f (Suc k) x)$1 \<le> (f k x)$1"
+lemma monotone_convergence_decreasing: fixes f::"nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
+  assumes "\<forall>k. (f k) integrable_on s"  "\<forall>k. \<forall>x\<in>s. (f (Suc k) x) \<le> (f k x)"
   "\<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially" "bounded {integral s (f k)| k. True}"
   shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
 proof- note assm = assms[rule_format]
@@ -4679,28 +4722,6 @@
     using Lim_neg[OF *(2)] apply- unfolding integral_neg[OF assm(1)]
     unfolding integral_neg[OF *(1),THEN sym] by auto qed
 
-lemma monotone_convergence_increasing_real: fixes f::"nat \<Rightarrow> real^'n \<Rightarrow> real"
-  assumes "\<forall>k. (f k) integrable_on s"  "\<forall>k. \<forall>x\<in>s. (f (Suc k) x) \<ge> (f k x)"
-  "\<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially" "bounded {integral s (f k)| k. True}"
-  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
-proof- have *:"{integral s (\<lambda>x. vec1 (f k x)) |k. True} =  vec1 ` {integral s (f k) |k. True}"
-    unfolding integral_trans[OF assms(1)[rule_format]] by auto
-  have "vec1 \<circ> g integrable_on s \<and> ((\<lambda>k. integral s (vec1 \<circ> f k)) ---> integral s (vec1 \<circ> g)) sequentially"
-    apply(rule monotone_convergence_increasing) unfolding o_def integrable_on_trans
-    unfolding vec1_dest_vec1 apply(rule assms)+ unfolding Lim_trans unfolding * using assms(3,4) by auto
-  thus ?thesis unfolding o_def unfolding integral_trans[OF assms(1)[rule_format]] by auto qed
-
-lemma monotone_convergence_decreasing_real: fixes f::"nat \<Rightarrow> real^'n \<Rightarrow> real"
-  assumes "\<forall>k. (f k) integrable_on s"  "\<forall>k. \<forall>x\<in>s. (f (Suc k) x) \<le> (f k x)"
-  "\<forall>x\<in>s. ((\<lambda>k. f k x) ---> g x) sequentially" "bounded {integral s (f k)| k. True}"
-  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) ---> integral s g) sequentially"
-proof- have *:"{integral s (\<lambda>x. vec1 (f k x)) |k. True} =  vec1 ` {integral s (f k) |k. True}"
-    unfolding integral_trans[OF assms(1)[rule_format]] by auto
-  have "vec1 \<circ> g integrable_on s \<and> ((\<lambda>k. integral s (vec1 \<circ> f k)) ---> integral s (vec1 \<circ> g)) sequentially"
-    apply(rule monotone_convergence_decreasing) unfolding o_def integrable_on_trans
-    unfolding vec1_dest_vec1 apply(rule assms)+ unfolding Lim_trans unfolding * using assms(3,4) by auto
-  thus ?thesis unfolding o_def unfolding integral_trans[OF assms(1)[rule_format]] by auto qed
-
 subsection {* absolute integrability (this is the same as Lebesgue integrability). *}
 
 definition absolutely_integrable_on (infixr "absolutely'_integrable'_on" 46) where
@@ -4714,11 +4735,11 @@
   shows "f integrable_on s" "(\<lambda>x. (norm(f x))) integrable_on s"
   using assms unfolding absolutely_integrable_on_def by auto
 
-lemma absolutely_integrable_on_trans[simp]: fixes f::"real^'n \<Rightarrow> real" shows
+(*lemma absolutely_integrable_on_trans[simp]: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real" shows
   "(vec1 o f) absolutely_integrable_on s \<longleftrightarrow> f absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def o_def by auto
-
-lemma integral_norm_bound_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
+  unfolding absolutely_integrable_on_def o_def by auto*)
+
+lemma integral_norm_bound_integral: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f integrable_on s" "g integrable_on s" "\<forall>x\<in>s. norm(f x) \<le> g x"
   shows "norm(integral s f) \<le> (integral s g)"
 proof- have *:"\<And>x y. (\<forall>e::real. 0 < e \<longrightarrow> x < y + e) \<longrightarrow> x \<le> y" apply(safe,rule ccontr)
@@ -4730,7 +4751,7 @@
       apply(rule add_le_less_mono) defer apply(subst norm_minus_commute,rule goal1)
       apply(rule order_trans[OF goal1(1)]) using goal1(2) by arith
   qed note norm=this[rule_format]
-  have lem:"\<And>f::real^'n \<Rightarrow> 'a. \<And> g a b. f integrable_on {a..b} \<Longrightarrow> g integrable_on {a..b} \<Longrightarrow>
+  have lem:"\<And>f::'n::ordered_euclidean_space \<Rightarrow> 'a. \<And> g a b. f integrable_on {a..b} \<Longrightarrow> g integrable_on {a..b} \<Longrightarrow>
     \<forall>x\<in>{a..b}. norm(f x) \<le> (g x) \<Longrightarrow> norm(integral({a..b}) f) \<le> (integral({a..b}) g)"
   proof(rule *[rule_format]) case goal1 hence *:"e/2>0" by auto
     from integrable_integral[OF goal1(1),unfolded has_integral[of f],rule_format,OF *]
@@ -4758,7 +4779,7 @@
   guess B1 .. note B1=conjunctD2[OF this[rule_format],rule_format]
   from integrable_integral[OF assms(2),unfolded has_integral'[of g],rule_format,OF e]
   guess B2 .. note B2=conjunctD2[OF this[rule_format],rule_format]
-  from bounded_subset_closed_interval[OF bounded_ball, of "0::real^'n" "max B1 B2"]
+  from bounded_subset_closed_interval[OF bounded_ball, of "0::'n::ordered_euclidean_space" "max B1 B2"]
   guess a b apply-by(erule exE)+ note ab=this[unfolded ball_max_Un]
 
   have "ball 0 B1 \<subseteq> {a..b}" using ab by auto
@@ -4771,22 +4792,24 @@
     defer apply(rule w(2)[unfolded real_norm_def],rule z(2))
     apply safe apply(case_tac "x\<in>s") unfolding if_P apply(rule assms(3)[rule_format]) by auto qed
 
-lemma integral_norm_bound_integral_component: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  assumes "f integrable_on s" "g integrable_on s"  "\<forall>x\<in>s. norm(f x) \<le> (g x)$k"
-  shows "norm(integral s f) \<le> (integral s g)$k"
-proof- have "norm (integral s f) \<le> integral s ((\<lambda>x. x $ k) o g)"
+lemma integral_norm_bound_integral_component: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
+  fixes g::"'n => 'b::ordered_euclidean_space"
+  assumes "f integrable_on s" "g integrable_on s"  "\<forall>x\<in>s. norm(f x) \<le> (g x)$$k"
+  shows "norm(integral s f) \<le> (integral s g)$$k"
+proof- have "norm (integral s f) \<le> integral s ((\<lambda>x. x $$ k) o g)"
     apply(rule integral_norm_bound_integral[OF assms(1)])
     apply(rule integrable_linear[OF assms(2)],rule)
     unfolding o_def by(rule assms)
   thus ?thesis unfolding o_def integral_component_eq[OF assms(2)] . qed
 
-lemma has_integral_norm_bound_integral_component: fixes f::"real^'n \<Rightarrow> 'a::banach"
-  assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. norm(f x) \<le> (g x)$k"
-  shows "norm(i) \<le> j$k" using integral_norm_bound_integral_component[of f s g k]
+lemma has_integral_norm_bound_integral_component: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
+  fixes g::"'n => 'b::ordered_euclidean_space"
+  assumes "(f has_integral i) s" "(g has_integral j) s" "\<forall>x\<in>s. norm(f x) \<le> (g x)$$k"
+  shows "norm(i) \<le> j$$k" using integral_norm_bound_integral_component[of f s g k]
   unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)]
   using assms by auto
 
-lemma absolutely_integrable_le: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma absolutely_integrable_le: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f absolutely_integrable_on s"
   shows "norm(integral s f) \<le> integral s (\<lambda>x. norm(f x))"
   apply(rule integral_norm_bound_integral) using assms by auto
@@ -4811,11 +4834,11 @@
  "f absolutely_integrable_on s \<Longrightarrow> (\<lambda>x. abs(f x::real)) absolutely_integrable_on s"
   apply(drule absolutely_integrable_norm) unfolding real_norm_def .
 
-lemma absolutely_integrable_on_subinterval: fixes f::"real^'n \<Rightarrow> 'a::banach" shows
+lemma absolutely_integrable_on_subinterval: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach" shows
   "f absolutely_integrable_on s \<Longrightarrow> {a..b} \<subseteq> s \<Longrightarrow> f absolutely_integrable_on {a..b}" 
   unfolding absolutely_integrable_on_def by(meson integrable_on_subinterval)
 
-lemma absolutely_integrable_bounded_variation: fixes f::"real^'n \<Rightarrow> 'a::banach"
+lemma absolutely_integrable_bounded_variation: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
   assumes "f absolutely_integrable_on UNIV"
   obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
   apply(rule that[of "integral UNIV (\<lambda>x. norm (f x))"])
@@ -4839,7 +4862,7 @@
   using norm_triangle_ineq3 .
 
 lemma bounded_variation_absolutely_integrable_interval:
-  fixes f::"real^'n \<Rightarrow> real^'m" assumes "f integrable_on {a..b}"
+  fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space" assumes "f integrable_on {a..b}"
   "\<forall>d. d division_of {a..b} \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
   shows "f absolutely_integrable_on {a..b}"
 proof- let ?S = "(\<lambda>d. setsum (\<lambda>k. norm(integral k f)) d) ` {d. d division_of {a..b} }" def i \<equiv> "Sup ?S"
@@ -4851,7 +4874,7 @@
         {d. d division_of {a..b}}))" using isLub_ubs[OF i,rule_format]
       unfolding setge_def ubs_def by auto 
     hence " \<exists>y. y division_of {a..b} \<and> i - e / 2 < (\<Sum>k\<in>y. norm (integral k f))"
-      unfolding mem_Collect_eq isUb_def setle_def by simp then guess d .. note d=conjunctD2[OF this]
+      unfolding mem_Collect_eq isUb_def setle_def by(simp add:not_le) then guess d .. note d=conjunctD2[OF this]
     note d' = division_ofD[OF this(1)]
 
     have "\<forall>x. \<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
@@ -5051,7 +5074,7 @@
         qed finally show ?case . 
       qed qed qed qed 
 
-lemma bounded_variation_absolutely_integrable:  fixes f::"real^'n \<Rightarrow> real^'m"
+lemma bounded_variation_absolutely_integrable:  fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assumes "f integrable_on UNIV" "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
   shows "f absolutely_integrable_on UNIV"
 proof(rule absolutely_integrable_onI,fact,rule)
@@ -5075,7 +5098,7 @@
     have "bounded (\<Union>d)" by(rule elementary_bounded,fact)
     from this[unfolded bounded_pos] guess K .. note K=conjunctD2[OF this]
     show ?case apply(rule_tac x="K + 1" in exI,safe)
-    proof- fix a b assume ab:"ball 0 (K + 1) \<subseteq> {a..b::real^'n}"
+    proof- fix a b assume ab:"ball 0 (K + 1) \<subseteq> {a..b::'n::ordered_euclidean_space}"
       have *:"\<forall>s s1. i - e < s1 \<and> s1 \<le> s \<and> s < i + e \<longrightarrow> abs(s - i) < (e::real)" by arith
       show "norm (integral {a..b} (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - i) < e"
         unfolding real_norm_def apply(rule *[rule_format],safe) apply(rule d(2))
@@ -5123,16 +5146,16 @@
  "(\<lambda>x. if x \<in> s then f x else (0::'a::banach)) absolutely_integrable_on UNIV \<longleftrightarrow> f absolutely_integrable_on s"
   unfolding absolutely_integrable_on_def if_distrib norm_zero integrable_restrict_univ ..
 
-lemma absolutely_integrable_add[intro]: fixes f g::"real^'n \<Rightarrow> real^'m"
+lemma absolutely_integrable_add[intro]: fixes f g::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
   shows "(\<lambda>x. f(x) + g(x)) absolutely_integrable_on s"
-proof- let ?P = "\<And>f g::real^'n \<Rightarrow> real^'m. f absolutely_integrable_on UNIV \<Longrightarrow>
+proof- let ?P = "\<And>f g::'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space. f absolutely_integrable_on UNIV \<Longrightarrow>
     g absolutely_integrable_on UNIV \<Longrightarrow> (\<lambda>x. f(x) + g(x)) absolutely_integrable_on UNIV"
   { presume as:"PROP ?P" note a = absolutely_integrable_restrict_univ[THEN sym]
     have *:"\<And>x. (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)
       = (if x \<in> s then f x + g x else 0)" by auto
     show ?thesis apply(subst a) using as[OF assms[unfolded a[of f] a[of g]]] unfolding * . }
-  fix f g::"real^'n \<Rightarrow> real^'m" assume assms:"f absolutely_integrable_on UNIV"
+  fix f g::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space" assume assms:"f absolutely_integrable_on UNIV"
     "g absolutely_integrable_on UNIV" 
   note absolutely_integrable_bounded_variation
   from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
@@ -5150,21 +5173,21 @@
     finally show ?case .
   qed(insert assms,auto) qed
 
-lemma absolutely_integrable_sub[intro]: fixes f g::"real^'n \<Rightarrow> real^'m"
+lemma absolutely_integrable_sub[intro]: fixes f g::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
   shows "(\<lambda>x. f(x) - g(x)) absolutely_integrable_on s"
   using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
   unfolding algebra_simps .
 
-lemma absolutely_integrable_linear: fixes f::"real^'m \<Rightarrow> real^'n" and h::"real^'n \<Rightarrow> real^'p"
+lemma absolutely_integrable_linear: fixes f::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space" and h::"'n::ordered_euclidean_space \<Rightarrow> 'p::ordered_euclidean_space"
   assumes "f absolutely_integrable_on s" "bounded_linear h"
   shows "(h o f) absolutely_integrable_on s"
-proof- { presume as:"\<And>f::real^'m \<Rightarrow> real^'n. \<And>h::real^'n \<Rightarrow> real^'p. 
+proof- { presume as:"\<And>f::'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space. \<And>h::'n::ordered_euclidean_space \<Rightarrow> 'p::ordered_euclidean_space. 
     f absolutely_integrable_on UNIV \<Longrightarrow> bounded_linear h \<Longrightarrow>
     (h o f) absolutely_integrable_on UNIV" note a = absolutely_integrable_restrict_univ[THEN sym]
     show ?thesis apply(subst a) using as[OF assms[unfolded a[of f] a[of g]]]
       unfolding o_def if_distrib linear_simps[OF assms(2)] . }
-  fix f::"real^'m \<Rightarrow> real^'n" and h::"real^'n \<Rightarrow> real^'p"
+  fix f::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space" and h::"'n::ordered_euclidean_space \<Rightarrow> 'p::ordered_euclidean_space"
   assume assms:"f absolutely_integrable_on UNIV" "bounded_linear h" 
   from absolutely_integrable_bounded_variation[OF assms(1)] guess B . note B=this
   from bounded_linear.pos_bounded[OF assms(2)] guess b .. note b=conjunctD2[OF this]
@@ -5185,125 +5208,112 @@
     finally show ?case .
   qed(insert assms,auto) qed
 
-lemma absolutely_integrable_setsum: fixes f::"'a \<Rightarrow> real^'n \<Rightarrow> real^'m"
+lemma absolutely_integrable_setsum: fixes f::"'a \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assumes "finite t" "\<And>a. a \<in> t \<Longrightarrow> (f a) absolutely_integrable_on s"
   shows "(\<lambda>x. setsum (\<lambda>a. f a x) t) absolutely_integrable_on s"
   using assms(1,2) apply induct defer apply(subst setsum.insert) apply assumption+ by(rule,auto)
 
-lemma absolutely_integrable_vector_abs:
+lemma absolutely_integrable_vector_abs: fixes f::"'a::ordered_euclidean_space => 'b::ordered_euclidean_space"
   assumes "f absolutely_integrable_on s"
-  shows "(\<lambda>x. (\<chi> i. abs(f x$i))::real^'n) absolutely_integrable_on s"
-proof- have *:"\<And>x. ((\<chi> i. abs(f x$i))::real^'n) = (setsum (\<lambda>i.
-    (((\<lambda>y. (\<chi> j. if j = i then y$1 else 0)) o
-    (vec1 o ((\<lambda>x. (norm((\<chi> j. if j = i then x$i else 0)::real^'n))) o f))) x)) UNIV)"
-    unfolding Cart_eq setsum_component Cart_lambda_beta
-  proof case goal1 have *:"\<And>i xa. ((if i = xa then f x $ xa else 0) \<bullet> (if i = xa then f x $ xa else 0)) =
-      (if i = xa then (f x $ xa) * (f x $ xa) else 0)" by auto
-    have "\<bar>f x $ i\<bar> = (setsum (\<lambda>k. if k = i then abs ((f x)$i) else 0) UNIV)"
-      unfolding setsum_delta[OF finite_UNIV] by auto
-    also have "... = (\<Sum>xa\<in>UNIV. ((\<lambda>y. \<chi> j. if j = xa then dest_vec1 y else 0) \<circ>
-                      (\<lambda>x. vec1 (norm (\<chi> j. if j = xa then x $ xa else 0))) \<circ> f) x $ i)"
-      unfolding norm_eq_sqrt_inner inner_vector_def Cart_lambda_beta o_def *
-      apply(rule setsum_cong2) unfolding setsum_delta[OF finite_UNIV] by auto 
+  shows "(\<lambda>x. (\<chi>\<chi> i. abs(f x$$i))::'c::ordered_euclidean_space) absolutely_integrable_on s"
+proof- have *:"\<And>x. ((\<chi>\<chi> i. abs(f x$$i))::'c::ordered_euclidean_space) = (setsum (\<lambda>i.
+    (((\<lambda>y. (\<chi>\<chi> j. if j = i then y else 0)) o
+    (((\<lambda>x. (norm((\<chi>\<chi> j. if j = i then x$$i else 0)::'c::ordered_euclidean_space))) o f))) x)) {..<DIM('c)})"
+    unfolding euclidean_eq[where 'a='c] euclidean_component.setsum apply safe
+    unfolding euclidean_lambda_beta'
+  proof- case goal1 have *:"\<And>i xa. ((if i = xa then f x $$ xa else 0) * (if i = xa then f x $$ xa else 0)) =
+      (if i = xa then (f x $$ xa) * (f x $$ xa) else 0)" by auto
+    have *:"\<And>xa. norm ((\<chi>\<chi> j. if j = xa then f x $$ xa else 0)::'c) = (if xa<DIM('c) then abs (f x $$ xa) else 0)"
+      unfolding norm_eq_sqrt_inner euclidean_inner[where 'a='c]
+      by(auto simp add:setsum_delta[OF finite_lessThan] *)
+    have "\<bar>f x $$ i\<bar> = (setsum (\<lambda>k. if k = i then abs ((f x)$$i) else 0) {..<DIM('c)})"
+      unfolding setsum_delta[OF finite_lessThan] using goal1 by auto
+    also have "... = (\<Sum>xa<DIM('c). ((\<lambda>y. (\<chi>\<chi> j. if j = xa then y else 0)::'c) \<circ>
+                      (\<lambda>x. (norm ((\<chi>\<chi> j. if j = xa then x $$ xa else 0)::'c))) \<circ> f) x $$ i)"
+      unfolding o_def * apply(rule setsum_cong2)
+      unfolding euclidean_lambda_beta'[OF goal1 ] by auto
     finally show ?case unfolding o_def . qed
-  show ?thesis unfolding * apply(rule absolutely_integrable_setsum) apply(rule finite_UNIV)
-    apply(rule absolutely_integrable_linear) 
-    unfolding absolutely_integrable_on_trans unfolding o_def apply(rule absolutely_integrable_norm)
+  show ?thesis unfolding * apply(rule absolutely_integrable_setsum) apply(rule finite_lessThan)
+    apply(rule absolutely_integrable_linear) unfolding o_def apply(rule absolutely_integrable_norm)
     apply(rule absolutely_integrable_linear[OF assms,unfolded o_def]) unfolding linear_linear
-    apply(rule_tac[!] linearI) unfolding Cart_eq by auto
+    apply(rule_tac[!] linearI) unfolding euclidean_eq[where 'a='c]
+    by(auto simp:euclidean_scaleR[where 'a=real,unfolded real_scaleR_def])
 qed
 
-lemma absolutely_integrable_max: fixes f::"real^'m \<Rightarrow> real^'n"
+lemma absolutely_integrable_max: fixes f g::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
   assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
-  shows "(\<lambda>x. (\<chi> i. max (f(x)$i) (g(x)$i))::real^'n) absolutely_integrable_on s"
-proof- have *:"\<And>x. (1 / 2) *\<^sub>R ((\<chi> i. \<bar>(f x - g x) $ i\<bar>) + (f x + g x)) = (\<chi> i. max (f(x)$i) (g(x)$i))"
-    unfolding Cart_eq by auto
+  shows "(\<lambda>x. (\<chi>\<chi> i. max (f(x)$$i) (g(x)$$i))::'n::ordered_euclidean_space) absolutely_integrable_on s"
+proof- have *:"\<And>x. (1 / 2) *\<^sub>R (((\<chi>\<chi> i. \<bar>(f x - g x) $$ i\<bar>)::'n) + (f x + g x)) = (\<chi>\<chi> i. max (f(x)$$i) (g(x)$$i))"
+    unfolding euclidean_eq[where 'a='n] by auto
   note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms]
   note absolutely_integrable_vector_abs[OF this(1)] this(2)
   note absolutely_integrable_add[OF this] note absolutely_integrable_cmul[OF this,of "1/2"]
   thus ?thesis unfolding * . qed
 
-lemma absolutely_integrable_max_real: fixes f::"real^'m \<Rightarrow> real"
+lemma absolutely_integrable_min: fixes f g::"'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
   assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
-  shows "(\<lambda>x. max (f x) (g x)) absolutely_integrable_on s"
-proof- have *:"(\<lambda>x. \<chi> i. max ((vec1 \<circ> f) x $ i) ((vec1 \<circ> g) x $ i)) = vec1 o (\<lambda>x. max (f x) (g x))"
-    apply rule unfolding Cart_eq by auto note absolutely_integrable_max[of "vec1 o f" s "vec1 o g"]
-  note this[unfolded absolutely_integrable_on_trans,OF assms]
-  thus ?thesis unfolding * by auto qed
-
-lemma absolutely_integrable_min: fixes f::"real^'m \<Rightarrow> real^'n"
-  assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
-  shows "(\<lambda>x. (\<chi> i. min (f(x)$i) (g(x)$i))::real^'n) absolutely_integrable_on s"
-proof- have *:"\<And>x. (1 / 2) *\<^sub>R ((f x + g x) - (\<chi> i. \<bar>(f x - g x) $ i\<bar>)) = (\<chi> i. min (f(x)$i) (g(x)$i))"
-    unfolding Cart_eq by auto
+  shows "(\<lambda>x. (\<chi>\<chi> i. min (f(x)$$i) (g(x)$$i))::'n::ordered_euclidean_space) absolutely_integrable_on s"
+proof- have *:"\<And>x. (1 / 2) *\<^sub>R ((f x + g x) - ((\<chi>\<chi> i. \<bar>(f x - g x) $$ i\<bar>)::'n)) = (\<chi>\<chi> i. min (f(x)$$i) (g(x)$$i))"
+    unfolding euclidean_eq[where 'a='n] by auto
   note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms]
   note this(1) absolutely_integrable_vector_abs[OF this(2)]
   note absolutely_integrable_sub[OF this] note absolutely_integrable_cmul[OF this,of "1/2"]
   thus ?thesis unfolding * . qed
 
-lemma absolutely_integrable_min_real: fixes f::"real^'m \<Rightarrow> real"
-  assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
-  shows "(\<lambda>x. min (f x) (g x)) absolutely_integrable_on s"
-proof- have *:"(\<lambda>x. \<chi> i. min ((vec1 \<circ> f) x $ i) ((vec1 \<circ> g) x $ i)) = vec1 o (\<lambda>x. min (f x) (g x))"
-    apply rule unfolding Cart_eq by auto note absolutely_integrable_min[of "vec1 o f" s "vec1 o g"]
-  note this[unfolded absolutely_integrable_on_trans,OF assms]
-  thus ?thesis unfolding * by auto qed
-
-lemma absolutely_integrable_abs_eq: fixes f::"real^'n \<Rightarrow> real^'m"
+lemma absolutely_integrable_abs_eq: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and>
-          (\<lambda>x. (\<chi> i. abs(f x$i))::real^'m) integrable_on s" (is "?l = ?r")
+          (\<lambda>x. (\<chi>\<chi> i. abs(f x$$i))::'m) integrable_on s" (is "?l = ?r")
 proof assume ?l thus ?r apply-apply rule defer
     apply(drule absolutely_integrable_vector_abs) by auto
-next assume ?r { presume lem:"\<And>f::real^'n \<Rightarrow> real^'m. f integrable_on UNIV \<Longrightarrow>
-    (\<lambda>x. (\<chi> i. abs(f(x)$i))::real^'m) integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
-    have *:"\<And>x. (\<chi> i. \<bar>(if x \<in> s then f x else 0) $ i\<bar>) = (if x\<in>s then (\<chi> i. \<bar>f x $ i\<bar>) else 0)"
-      unfolding Cart_eq by auto
+next assume ?r { presume lem:"\<And>f::'n \<Rightarrow> 'm. f integrable_on UNIV \<Longrightarrow>
+    (\<lambda>x. (\<chi>\<chi> i. abs(f(x)$$i))::'m) integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
+    have *:"\<And>x. (\<chi>\<chi> i. \<bar>(if x \<in> s then f x else 0) $$ i\<bar>) = (if x\<in>s then (\<chi>\<chi> i. \<bar>f x $$ i\<bar>) else (0::'m))"
+      unfolding euclidean_eq[where 'a='m] by auto
     show ?l apply(subst absolutely_integrable_restrict_univ[THEN sym]) apply(rule lem)
       unfolding integrable_restrict_univ * using `?r` by auto }
-  fix f::"real^'n \<Rightarrow> real^'m" assume assms:"f integrable_on UNIV"
-    "(\<lambda>x. (\<chi> i. abs(f(x)$i))::real^'m) integrable_on UNIV"
-  let ?B = "setsum (\<lambda>i. integral UNIV (\<lambda>x. (\<chi> j. abs(f x$j)) ::real^'m) $ i) UNIV"
+  fix f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space" assume assms:"f integrable_on UNIV"
+    "(\<lambda>x. (\<chi>\<chi> i. abs(f(x)$$i))::'m::ordered_euclidean_space) integrable_on UNIV"
+  let ?B = "setsum (\<lambda>i. integral UNIV (\<lambda>x. (\<chi>\<chi> j. abs(f x$$j)) ::'m::ordered_euclidean_space) $$ i) {..<DIM('m)}"
   show "f absolutely_integrable_on UNIV"
     apply(rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"],safe)
   proof- case goal1 note d=this and d'=division_ofD[OF this]
     have "(\<Sum>k\<in>d. norm (integral k f)) \<le>
-      (\<Sum>k\<in>d. setsum (op $ (integral k (\<lambda>x. \<chi> j. \<bar>f x $ j\<bar>))) UNIV)" apply(rule setsum_mono)
-      apply(rule order_trans[OF norm_le_l1],rule setsum_mono)
-    proof- fix k and i::'m assume "k\<in>d"
-      from d'(4)[OF this] guess a b apply-by(erule exE)+ note ab=this
-      show "\<bar>integral k f $ i\<bar> \<le> integral k (\<lambda>x. \<chi> j. \<bar>f x $ j\<bar>) $ i" apply(rule abs_leI)
-        unfolding vector_uminus_component[THEN sym] defer apply(subst integral_neg[THEN sym])
+      (\<Sum>k\<in>d. setsum (op $$ (integral k (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m))) {..<DIM('m)})" apply(rule setsum_mono)
+      apply(rule order_trans[OF norm_le_l1]) apply(rule setsum_mono) unfolding lessThan_iff
+    proof- fix k and i assume "k\<in>d" and i:"i<DIM('m)"
+      from d'(4)[OF this(1)] guess a b apply-by(erule exE)+ note ab=this
+      show "\<bar>integral k f $$ i\<bar> \<le> integral k (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) $$ i" apply(rule abs_leI)
+        unfolding euclidean_component.minus[THEN sym] defer apply(subst integral_neg[THEN sym])
         defer apply(rule_tac[1-2] integral_component_le) apply(rule integrable_neg)
         using integrable_on_subinterval[OF assms(1),of a b]
           integrable_on_subinterval[OF assms(2),of a b] unfolding ab by auto
-    qed also have "... \<le> setsum (op $ (integral UNIV (\<lambda>x. \<chi> j. \<bar>f x $ j\<bar>))) UNIV"
+    qed also have "... \<le> setsum (op $$ (integral UNIV (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>))::'m)) {..<DIM('m)}"
       apply(subst setsum_commute,rule setsum_mono)
-    proof- case goal1 have *:"(\<lambda>x. \<chi> j. \<bar>f x $ j\<bar>) integrable_on \<Union>d"
+    proof- case goal1 have *:"(\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) integrable_on \<Union>d"
         using integrable_on_subdivision[OF d assms(2)] by auto
-      have "(\<Sum>i\<in>d. integral i (\<lambda>x. \<chi> j. \<bar>f x $ j\<bar>) $ j)
-        = integral (\<Union>d) (\<lambda>x. (\<chi> j. abs(f x$j)) ::real^'m) $ j"
-        unfolding setsum_component[THEN sym]
-        apply(subst integral_combine_division_topdown[THEN sym,OF * d]) by auto
-      also have "... \<le> integral UNIV (\<lambda>x. \<chi> j. \<bar>f x $ j\<bar>) $ j"
+      have "(\<Sum>i\<in>d. integral i (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) $$ j)
+        = integral (\<Union>d) (\<lambda>x. (\<chi>\<chi> j. abs(f x$$j)) ::'m::ordered_euclidean_space) $$ j"
+        unfolding euclidean_component.setsum[THEN sym] integral_combine_division_topdown[OF * d] ..
+      also have "... \<le> integral UNIV (\<lambda>x. (\<chi>\<chi> j. \<bar>f x $$ j\<bar>)::'m) $$ j"
         apply(rule integral_subset_component_le) using assms * by auto
       finally show ?case .
     qed finally show ?case . qed qed
 
-lemma nonnegative_absolutely_integrable: fixes f::"real^'n \<Rightarrow> real^'m"
-  assumes "\<forall>x\<in>s. \<forall>i. 0 \<le> f(x)$i" "f integrable_on s"
+lemma nonnegative_absolutely_integrable: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+  assumes "\<forall>x\<in>s. \<forall>i. 0 \<le> f(x)$$i" "f integrable_on s"
   shows "f absolutely_integrable_on s"
   unfolding absolutely_integrable_abs_eq apply rule defer
-  apply(rule integrable_eq[of _ f]) unfolding Cart_eq using assms by auto
-
-lemma absolutely_integrable_integrable_bound: fixes f::"real^'n \<Rightarrow> real^'m"
+  apply(rule integrable_eq[of _ f]) using assms by auto
+
+lemma absolutely_integrable_integrable_bound: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assumes "\<forall>x\<in>s. norm(f x) \<le> g x" "f integrable_on s" "g integrable_on s"
   shows "f absolutely_integrable_on s"
-proof- { presume *:"\<And>f::real^'n \<Rightarrow> real^'m. \<And> g. \<forall>x. norm(f x) \<le> g x \<Longrightarrow> f integrable_on UNIV
+proof- { presume *:"\<And>f::'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space. \<And> g. \<forall>x. norm(f x) \<le> g x \<Longrightarrow> f integrable_on UNIV
     \<Longrightarrow> g integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
     show ?thesis apply(subst absolutely_integrable_restrict_univ[THEN sym])
       apply(rule *[of _ "\<lambda>x. if x\<in>s then g x else 0"])
       using assms unfolding integrable_restrict_univ by auto }
-  fix g and f :: "real^'n \<Rightarrow> real^'m"
+  fix g and f :: "'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
   assume assms:"\<forall>x. norm(f x) \<le> g x" "f integrable_on UNIV" "g integrable_on UNIV"
   show "f absolutely_integrable_on UNIV"
     apply(rule bounded_variation_absolutely_integrable[OF assms(2),where B="integral UNIV g"])
@@ -5319,15 +5329,14 @@
       apply(rule,rule_tac y="norm (f x)" in order_trans) using assms by auto
     finally show ?case . qed qed
 
-lemma absolutely_integrable_integrable_bound_real: fixes f::"real^'n \<Rightarrow> real"
+lemma absolutely_integrable_integrable_bound_real: fixes f::"'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "\<forall>x\<in>s. norm(f x) \<le> g x" "f integrable_on s" "g integrable_on s"
   shows "f absolutely_integrable_on s"
-  apply(subst absolutely_integrable_on_trans[THEN sym])
   apply(rule absolutely_integrable_integrable_bound[where g=g])
   using assms unfolding o_def by auto
 
 lemma absolutely_integrable_absolutely_integrable_bound:
-  fixes f::"real^'n \<Rightarrow> real^'m" and g::"real^'n \<Rightarrow> real^'p"
+  fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space" and g::"'n::ordered_euclidean_space \<Rightarrow> 'p::ordered_euclidean_space"
   assumes "\<forall>x\<in>s. norm(f x) \<le> norm(g x)" "f integrable_on s" "g absolutely_integrable_on s"
   shows "f absolutely_integrable_on s"
   apply(rule absolutely_integrable_integrable_bound[of s f "\<lambda>x. norm (g x)"])
@@ -5343,8 +5352,8 @@
     apply(subst Inf_insert_finite) apply(rule finite_imageI[OF insert(1)])
   proof(cases "k={}") case True
     thus ?P apply(subst if_P) defer apply(rule insert(5)[rule_format]) by auto
-  next case False thus ?P apply(subst if_not_P) defer
-      apply(rule absolutely_integrable_min_real) 
+  next case False thus ?P apply(subst if_not_P) defer      
+      apply(rule absolutely_integrable_min[where 'n=real,unfolded Eucl_real_simps])
       defer apply(rule insert(3)[OF False]) using insert(5) by auto
   qed qed auto
 
@@ -5359,13 +5368,13 @@
   proof(cases "k={}") case True
     thus ?P apply(subst if_P) defer apply(rule insert(5)[rule_format]) by auto
   next case False thus ?P apply(subst if_not_P) defer
-      apply(rule absolutely_integrable_max_real) 
+      apply(rule absolutely_integrable_max[where 'n=real,unfolded Eucl_real_simps]) 
       defer apply(rule insert(3)[OF False]) using insert(5) by auto
   qed qed auto
 
 subsection {* Dominated convergence. *}
 
-lemma dominated_convergence: fixes f::"nat \<Rightarrow> real^'n \<Rightarrow> real"
+lemma dominated_convergence: fixes f::"nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real"
   assumes "\<And>k. (f k) integrable_on s" "h integrable_on s"
   "\<And>k. \<forall>x \<in> s. norm(f k x) \<le> (h x)"
   "\<forall>x \<in> s. ((\<lambda>k. f k x) ---> g x) sequentially"
@@ -5373,7 +5382,7 @@
 proof- have "\<And>m. (\<lambda>x. Inf {f j x |j. m \<le> j}) integrable_on s \<and>
     ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) --->
     integral s (\<lambda>x. Inf {f j x |j. m \<le> j}))sequentially"
-  proof(rule monotone_convergence_decreasing_real,safe) fix m::nat
+  proof(rule monotone_convergence_decreasing,safe) fix m::nat
     show "bounded {integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}}) |k. True}"
       unfolding bounded_iff apply(rule_tac x="integral s h" in exI)
     proof safe fix k::nat
@@ -5418,7 +5427,7 @@
   have "\<And>m. (\<lambda>x. Sup {f j x |j. m \<le> j}) integrable_on s \<and>
     ((\<lambda>k. integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}})) --->
     integral s (\<lambda>x. Sup {f j x |j. m \<le> j})) sequentially"
-  proof(rule monotone_convergence_increasing_real,safe) fix m::nat
+  proof(rule monotone_convergence_increasing,safe) fix m::nat
     show "bounded {integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}}) |k. True}"
       unfolding bounded_iff apply(rule_tac x="integral s h" in exI)
     proof safe fix k::nat
@@ -5460,7 +5469,7 @@
       qed qed qed note inc1 = conjunctD2[OF this]
 
   have "g integrable_on s \<and> ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. k \<le> j})) ---> integral s g) sequentially"
-  apply(rule monotone_convergence_increasing_real,safe) apply fact 
+  apply(rule monotone_convergence_increasing,safe) apply fact 
   proof- show "bounded {integral s (\<lambda>x. Inf {f j x |j. k \<le> j}) |k. True}"
       unfolding bounded_iff apply(rule_tac x="integral s h" in exI)
     proof safe fix k::nat
@@ -5483,7 +5492,7 @@
     qed qed note inc2 = conjunctD2[OF this]
 
   have "g integrable_on s \<and> ((\<lambda>k. integral s (\<lambda>x. Sup {f j x |j. k \<le> j})) ---> integral s g) sequentially"
-  apply(rule monotone_convergence_decreasing_real,safe) apply fact 
+  apply(rule monotone_convergence_decreasing,safe) apply fact 
   proof- show "bounded {integral s (\<lambda>x. Sup {f j x |j. k \<le> j}) |k. True}"
       unfolding bounded_iff apply(rule_tac x="integral s h" in exI)
     proof safe fix k::nat
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -1,5 +1,5 @@
 theory Multivariate_Analysis
-imports Determinants Integration Real_Integration Fashoda
+imports Integration Fashoda
 begin
 
 end
--- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -11,21 +11,20 @@
 definition "onorm f = Sup {norm (f x)| x. norm x = 1}"
 
 lemma norm_bound_generalize:
-  fixes f:: "real ^'n \<Rightarrow> real^'m"
+  fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   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"
+    {fix x :: "'a" 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"
+    have bp: "b \<ge> 0" apply-apply(rule order_trans [OF norm_ge_zero])
+      apply(rule H[rule_format, of "basis 0::'a"]) by auto 
+    {fix x :: "'a"
       {assume "x = 0"
         then have "norm (f x) \<le> b * norm x" by (simp add: linear_0[OF lf] bp)}
       moreover
@@ -42,16 +41,18 @@
     then have ?rhs by blast}
   ultimately show ?thesis by blast
 qed
-
+ 
 lemma onorm:
-  fixes f:: "real ^'n \<Rightarrow> real ^'m"
+  fixes f:: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   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
+    have "norm (f (basis 0)) \<in> ?S" unfolding mem_Collect_eq
+      apply(rule_tac x="basis 0" in exI) by auto
+    hence Se: "?S \<noteq> {}" 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 Sup[OF Se b, unfolded onorm_def[symmetric]]
@@ -68,10 +69,11 @@
   }
 qed
 
-lemma onorm_pos_le: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" 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_pos_le: assumes lf: "linear (f::'n::euclidean_space \<Rightarrow> 'm::euclidean_space)" shows "0 <= onorm f"
+  using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis 0"]] 
+  using DIM_positive[where 'a='n] by auto
 
-lemma onorm_eq_0: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)"
+lemma onorm_eq_0: assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> 'b::euclidean_space)"
   shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"
   using onorm[OF lf]
   apply (auto simp add: onorm_pos_le)
@@ -81,24 +83,24 @@
   apply arith
   done
 
-lemma onorm_const: "onorm(\<lambda>x::real^'n. (y::real ^'m)) = norm y"
+lemma onorm_const: "onorm(\<lambda>x::'a::euclidean_space. (y::'b::euclidean_space)) = norm y"
 proof-
-  let ?f = "\<lambda>x::real^'n. (y::real ^ 'm)"
+  let ?f = "\<lambda>x::'a. (y::'b)"
   have th: "{norm (?f x)| x. norm x = 1} = {norm y}"
-    by(auto intro: vector_choose_size set_ext)
+    apply safe apply(rule_tac x="basis 0" in exI) by auto
   show ?thesis
     unfolding onorm_def th
     apply (rule Sup_unique) by (simp_all  add: setle_def)
 qed
 
-lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n \<Rightarrow> real ^'m)"
+lemma onorm_pos_lt: assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> 'b::euclidean_space)"
   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 \<Rightarrow> real ^'m)"
-  and lg: "linear (g::real^'k \<Rightarrow> real^'n)"
+  assumes lf: "linear (f::'n::euclidean_space \<Rightarrow> 'm::euclidean_space)"
+  and lg: "linear (g::'k::euclidean_space \<Rightarrow> 'n::euclidean_space)"
   shows "onorm (f o g) <= onorm f * onorm g"
   apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format])
   unfolding o_def
@@ -110,18 +112,18 @@
   apply (rule onorm_pos_le[OF lf])
   done
 
-lemma onorm_neg_lemma: assumes lf: "linear (f::real ^'n \<Rightarrow> real^'m)"
+lemma onorm_neg_lemma: assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> 'b::euclidean_space)"
   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 \<Rightarrow> real^'m)"
+lemma onorm_neg: assumes lf: "linear (f::'a::euclidean_space \<Rightarrow> 'b::euclidean_space)"
   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 \<Rightarrow> real ^'m)" and lg: "linear g"
+  assumes lf: "linear (f::'n::euclidean_space \<Rightarrow> 'm::euclidean_space)" 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)
@@ -132,14 +134,14 @@
   apply (rule onorm(1)[OF lg])
   done
 
-lemma onorm_triangle_le: "linear (f::real ^'n \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) <= e
+lemma onorm_triangle_le: "linear (f::'n::euclidean_space \<Rightarrow> 'm::euclidean_space) \<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 \<Rightarrow> real ^'m) \<Longrightarrow> linear g \<Longrightarrow> onorm(f) + onorm(g) < e
+lemma onorm_triangle_lt: "linear (f::'n::euclidean_space \<Rightarrow> 'm::euclidean_space) \<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)
--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -5,7 +5,7 @@
 header {* Continuous paths and path-connected sets *}
 
 theory Path_Connected
-imports Convex_Euclidean_Space
+imports Cartesian_Euclidean_Space
 begin
 
 subsection {* Paths. *}
@@ -66,7 +66,7 @@
 
 lemma compact_path_image[intro]: "path g \<Longrightarrow> compact(path_image g)"
   unfolding path_def path_image_def
-  by (erule compact_continuous_image, rule compact_real_interval)
+  by (erule compact_continuous_image, rule compact_interval)
 
 lemma reversepath_reversepath[simp]: "reversepath(reversepath g) = g"
   unfolding reversepath_def by auto
@@ -118,7 +118,7 @@
     fix x assume "x \<in> (\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {0::real..1}"
     hence "x \<ge> 1 / 2" unfolding image_iff by auto
     thus "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" proof(cases "x = 1 / 2")
-      case True hence "x = (1/2) *\<^sub>R 1" unfolding Cart_eq by auto
+      case True hence "x = (1/2) *\<^sub>R 1" by auto 
       thus ?thesis unfolding joinpaths_def using assms[unfolded pathstart_def pathfinish_def] by (auto simp add: mult_ac)
     qed (auto simp add:le_less joinpaths_def) qed
 next assume as:"continuous_on {0..1} g1" "continuous_on {0..1} g2"
@@ -504,12 +504,14 @@
 
 subsection {* sphere is path-connected. *}
 
+(** TODO covert this to ordered_euclidean_space **)
+
 lemma path_connected_punctured_universe:
  assumes "2 \<le> CARD('n::finite)" shows "path_connected((UNIV::(real^'n) 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}"
+  let ?basis = "\<lambda>k. cart_basis (\<psi> k)"
+  let ?A = "\<lambda>k. {x::real^'n. \<exists>i\<in>{1..k}. inner (cart_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)
@@ -537,15 +539,15 @@
           using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add: 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)"
+  have ***:"\<And>x::real^'n. (\<exists>i\<in>{1..CARD('n)}. inner (cart_basis (\<psi> i)) x \<noteq> 0) \<longleftrightarrow> (\<exists>i. inner (cart_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"
+    fix x::"real^'n" and i assume as:"inner (cart_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
+    thus "\<exists>i\<in>{1..CARD('n)}. inner (cart_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)
+  have **:"\<And>x::real^'n. x\<noteq>0 \<longleftrightarrow> (\<exists>i. inner (cart_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
 
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Jun 21 14:07:00 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Jun 21 19:33:51 2010 +0200
@@ -9,6 +9,16 @@
 imports SEQ Euclidean_Space Glbs
 begin
 
+(* to be moved elsewhere *)
+
+lemma euclidean_dist_l2:"dist x (y::'a::euclidean_space) = setL2 (\<lambda>i. dist(x$$i) (y$$i)) {..<DIM('a)}"
+  unfolding dist_norm norm_eq_sqrt_inner setL2_def apply(subst euclidean_inner)
+  apply(auto simp add:power2_eq_square) unfolding euclidean_component.diff ..
+
+lemma dist_nth_le: "dist (x $$ i) (y $$ i) \<le> dist x (y::'a::euclidean_space)"
+  apply(subst(2) euclidean_dist_l2) apply(cases "i<DIM('a)")
+  apply(rule member_le_setL2) by auto
+
 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)"
@@ -544,24 +554,22 @@
 apply (auto simp add: dist_norm)
 done
 
-instance cart :: (perfect_space, finite) perfect_space
-proof
-  fix x :: "'a ^ 'b"
-  {
-    fix e :: real assume "0 < e"
-    def a \<equiv> "x $ undefined"
+instance euclidean_space \<subseteq> perfect_space
+proof fix x::'a
+  { fix e :: real assume "0 < e"
+    def a \<equiv> "x $$ 0"
     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)
+    def y \<equiv> "\<chi>\<chi> i. if i = 0 then b else x$$i :: 'a"
+    from `b \<noteq> a` have "y \<noteq> x" unfolding a_def y_def apply(subst euclidean_eq) apply safe
+      apply(erule_tac x=0 in allE) using DIM_positive[where 'a='a] by auto
+
+    have *:"(\<Sum>i<DIM('a). (dist (y $$ i) (x $$ i))\<twosuperior>) = (\<Sum>i\<in>{0}. (dist (y $$ i) (x $$ i))\<twosuperior>)"
+      apply(rule setsum_mono_zero_right) unfolding y_def by auto
     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
+      apply(subst euclidean_dist_l2)
+      unfolding setL2_def * unfolding y_def a_def using `0 < e` by auto
     from `y \<noteq> x` and `dist y x < e`
     have "\<exists>y\<in>UNIV. y \<noteq> x \<and> dist y x < e" by auto
   }
@@ -577,26 +585,6 @@
 lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
   unfolding islimpt_def by auto
 
-lemma closed_positive_orthant: "closed {x::real^'n. \<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"
@@ -1324,8 +1312,8 @@
 qed
 
 lemma Lim_component:
-  fixes f :: "'a \<Rightarrow> 'b::metric_space ^ 'n"
-  shows "(f ---> l) net \<Longrightarrow> ((\<lambda>a. f a $i) ---> l$i) net"
+  fixes f :: "'a \<Rightarrow> ('a::euclidean_space)"
+  shows "(f ---> l) net \<Longrightarrow> ((\<lambda>a. f a $$i) ---> l$$i) net"
   unfolding tendsto_iff
   apply (clarify)
   apply (drule spec, drule (1) mp)
@@ -2287,6 +2275,23 @@
   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
 
+(* TODO: merge this lemma with the ones above *)
+lemma bounded_increasing_convergent: fixes s::"nat \<Rightarrow> real"
+  assumes "bounded {s n| n::nat. True}"  "\<forall>n. (s n) \<le>(s(Suc n))"
+  shows "\<exists>l. (s ---> l) sequentially"
+proof-
+  obtain a where a:"\<forall>n. \<bar> (s n)\<bar> \<le>  a" using assms(1)[unfolded bounded_iff] by auto
+  { fix m::nat
+    have "\<And> n. n\<ge>m \<longrightarrow>  (s m) \<le> (s n)"
+      apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE)
+      apply(case_tac "m \<le> na") unfolding not_less_eq_eq by(auto simp add: not_less_eq_eq)  }
+  hence "\<forall>m n. m \<le> n \<longrightarrow> (s m) \<le> (s n)" by auto
+  then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar> (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="l" in exI)
+    unfolding dist_norm  by auto
+qed
+
 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"
@@ -2310,69 +2315,84 @@
     by auto
 qed
 
-lemma bounded_component: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
+lemma bounded_component: "bounded s \<Longrightarrow>
+  bounded ((\<lambda>x. x $$ i) ` (s::'a::euclidean_space set))"
 unfolding bounded_def
 apply clarify
-apply (rule_tac x="x $ i" in exI)
+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)
+apply (rule order_trans[OF dist_nth_le],simp)
 done
 
 lemma compact_lemma:
-  fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n"
+  fixes f :: "nat \<Rightarrow> 'a::euclidean_space"
   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)"
+  shows "\<forall>d. \<exists>l::'a. \<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
-  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)"
+  fix d'::"nat set" def d \<equiv> "d' \<inter> {..<DIM('a)}"
+  have "finite d" "d\<subseteq>{..<DIM('a)}" unfolding d_def by auto
+  hence "\<exists>l::'a. \<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"
+  next case (insert k d) have k[intro]:"k<DIM('a)" using insert by auto
+    have s': "bounded ((\<lambda>x. x $$ k) ` s)" using `bounded s` by (rule bounded_component)
+    obtain l1::"'a" 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) using insert(4) 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"
+    def l \<equiv> "(\<chi>\<chi> i. if i = k then l2 else l1$$i)::'a"
     { 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"
+      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)
+      have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) $$ i) (l $$ i) < e) sequentially"
+        using N1' N2 apply(rule eventually_elim2) unfolding l_def r_def o_def
+        using insert.prems by auto
     }
     ultimately show ?case by auto
   qed
-qed
-
-instance cart :: (heine_borel, finite) heine_borel
+  thus "\<exists>l::'a. \<exists>r. subseq r \<and>
+      (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d'. dist (f (r n) $$ i) (l $$ i) < e) sequentially)"
+    apply safe apply(rule_tac x=l in exI,rule_tac x=r in exI) apply safe
+    apply(erule_tac x=e in allE) unfolding d_def eventually_sequentially apply safe 
+    apply(rule_tac x=N in exI) apply safe apply(erule_tac x=n in allE,safe)
+    apply(erule_tac x=i in ballE) 
+  proof- fix i and r::"nat=>nat" and n::nat and e::real and l::'a
+    assume "i\<in>d'" "i \<notin> d' \<inter> {..<DIM('a)}" and e:"e>0"
+    hence *:"i\<ge>DIM('a)" by auto
+    thus "dist (f (r n) $$ i) (l $$ i) < e" using e by auto
+  qed
+qed
+
+instance euclidean_space \<subseteq> heine_borel
 proof
-  fix s :: "('a ^ 'b) set" and f :: "nat \<Rightarrow> 'a ^ 'b"
+  fix s :: "'a set" and f :: "nat \<Rightarrow> 'a"
   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"
+  then obtain l::'a and 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"
+  let ?d = "{..<DIM('a)}"
   { 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"
+      using DIM_positive 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)
+    { 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))"
+        apply(subst euclidean_dist_l2) 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
+        apply(rule setsum_strict_mono) using n by auto
+      finally have "dist (f (r n)) l < e" unfolding setsum_constant
+        using DIM_positive[where 'a='a] by auto
     }
     ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
       by (rule eventually_elim1)
@@ -4161,12 +4181,6 @@
 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_def 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
@@ -4444,7 +4458,7 @@
     { 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)
+          using `c\<noteq>0` by auto
       }
       moreover
       { fix e::real assume "e>0"
@@ -4598,161 +4612,165 @@
 qed
 
 subsection {* Intervals *}
-
-lemma interval: fixes a :: "'a::ord^'n" 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_le_def)
-
-lemma mem_interval: fixes a :: "'a::ord^'n" 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_le_def)
-
-lemma interval_eq_empty: fixes a :: "real^'n" 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)
+  
+lemma interval: fixes a :: "'a::ordered_euclidean_space" shows
+  "{a <..< b} = {x::'a. \<forall>i<DIM('a). a$$i < x$$i \<and> x$$i < b$$i}" and
+  "{a .. b} = {x::'a. \<forall>i<DIM('a). a$$i \<le> x$$i \<and> x$$i \<le> b$$i}"
+  by(auto simp add:expand_set_eq eucl_le[where 'a='a] eucl_less[where 'a='a])
+
+lemma mem_interval: fixes a :: "'a::ordered_euclidean_space" shows
+  "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i < x$$i \<and> x$$i < b$$i)"
+  "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i \<le> x$$i \<and> x$$i \<le> b$$i)"
+  using interval[of a b] by(auto simp add: expand_set_eq eucl_le[where 'a='a] eucl_less[where 'a='a])
+
+lemma interval_eq_empty: fixes a :: "'a::ordered_euclidean_space" shows
+ "({a <..< b} = {} \<longleftrightarrow> (\<exists>i<DIM('a). b$$i \<le> a$$i))" (is ?th1) and
+ "({a  ..  b} = {} \<longleftrightarrow> (\<exists>i<DIM('a). 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
+  { fix i x assume i:"i<DIM('a)" and 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)"
+  { assume as:"\<forall>i<DIM('a). \<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  }
+    { fix i assume i:"i<DIM('a)" 
+      have "a$$i < b$$i" using as[THEN spec[where x=i]] using i by auto
+      hence "a$$i < ((1/2) *\<^sub>R (a+b)) $$ i" "((1/2) *\<^sub>R (a+b)) $$ i < b$$i"
+        unfolding euclidean_simps by auto }
     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
+  { fix i x assume i:"i<DIM('a)" and 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)"
+  { assume as:"\<forall>i<DIM('a). \<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  }
+    { fix i assume i:"i<DIM('a)"
+      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 euclidean_simps by auto }
     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" 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" 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
+lemma interval_ne_empty: fixes a :: "'a::ordered_euclidean_space" shows
+  "{a  ..  b} \<noteq> {} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i \<le> b$$i)" and
+  "{a <..< b} \<noteq> {} \<longleftrightarrow> (\<forall>i<DIM('a). a$$i < b$$i)"
+  unfolding interval_eq_empty[of a b] by fastsimp+
+
+lemma interval_sing: fixes a :: "'a::ordered_euclidean_space" shows
+ "{a .. a} = {a}" "{a<..<a} = {}"
+  apply(auto simp add: expand_set_eq euclidean_eq[where 'a='a] eucl_less[where 'a='a] eucl_le[where 'a='a])
+  apply (simp add: order_eq_iff) apply(rule_tac x=0 in exI) by (auto simp add: not_less less_imp_le)
+
+lemma subset_interval_imp: fixes a :: "'a::ordered_euclidean_space" shows
+ "(\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
+ "(\<forall>i<DIM('a). a$$i < c$$i \<and> d$$i < b$$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
+ "(\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
+ "(\<forall>i<DIM('a). 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" shows
- "{a .. a} = {a} \<and> {a<..<a} = {}"
-apply(auto simp add: expand_set_eq vector_less_def vector_le_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" shows
+lemma interval_open_subset_closed:  fixes a :: "'a::ordered_euclidean_space" 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_le_def Cart_eq)
+  { fix i assume "i<DIM('a)"
+    hence "a $$ i \<le> x $$ i"
+      using x order_less_imp_le[of "a$$i" "x$$i"] 
+      by(simp add: expand_set_eq eucl_less[where 'a='a] eucl_le[where 'a='a] euclidean_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_le_def Cart_eq)
+  { fix i assume "i<DIM('a)"
+    hence "x $$ i \<le> b $$ i"
+      using x order_less_imp_le[of "x$$i" "b$$i"]
+      by(simp add: expand_set_eq eucl_less[where 'a='a] eucl_le[where 'a='a] euclidean_eq)
   }
   ultimately
   show "a \<le> x \<and> x \<le> b"
-    by(simp add: expand_set_eq vector_less_def vector_le_def Cart_eq)
-qed
-
-lemma subset_interval: fixes a :: "real^'n" 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)
+    by(simp add: expand_set_eq eucl_less[where 'a='a] eucl_le[where 'a='a] euclidean_eq)
+qed
+
+lemma subset_interval: fixes a :: "'a::ordered_euclidean_space" shows
+ "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i \<le> d$$i) --> (\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i)" (is ?th1) and
+ "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i \<le> d$$i) --> (\<forall>i<DIM('a). a$$i < c$$i \<and> d$$i < b$$i)" (is ?th2) and
+ "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i < d$$i) --> (\<forall>i<DIM('a). a$$i \<le> c$$i \<and> d$$i \<le> b$$i)" (is ?th3) and
+ "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i<DIM('a). c$$i < d$$i) --> (\<forall>i<DIM('a). 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
+  { assume as: "{c<..<d} \<subseteq> {a .. b}" "\<forall>i<DIM('a). c$$i < d$$i"
+    hence "{c<..<d} \<noteq> {}" unfolding interval_eq_empty by auto
+    fix i assume i:"i<DIM('a)"
     (** 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
+    { let ?x = "(\<chi>\<chi> j. (if j=i then ((min (a$$j) (d$$j))+c$$j)/2 else (c$$j+d$$j)/2))::'a"
+      assume as2: "a$$i > c$$i"
+      { fix j assume j:"j<DIM('a)"
+        hence "c $$ j < ?x $$ j \<and> ?x $$ j < d $$ j"
+          apply(cases "j=i") using as(2)[THEN spec[where x=j]] i
+          by (auto simp add: as2)  }
+      hence "?x\<in>{c<..<d}" using i 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 i
+        by auto
+      ultimately have False using as by auto  }
+    hence "a$$i \<le> c$$i" by(rule ccontr)auto
+    moreover
+    { let ?x = "(\<chi>\<chi> j. (if j=i then ((max (b$$j) (c$$j))+d$$j)/2 else (c$$j+d$$j)/2))::'a"
+      assume as2: "b$$i < d$$i"
+      { fix j assume "j<DIM('a)"
+        hence "d $$ j > ?x $$ j \<and> ?x $$ j > c $$ j" 
           apply(cases "j=i") using as(2)[THEN spec[where x=j]]
           by (auto simp add: 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
-      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: 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
+        using as(2)[THEN spec[where x=i]] and as2 using i
         by auto
       ultimately have False using as by auto  }
-    hence "b$i \<ge> d$i" by(rule ccontr)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
+    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
+  show ?th3 unfolding subset_eq and Ball_def and mem_interval 
+    apply(rule,rule,rule,rule) apply(rule part1) unfolding subset_eq and Ball_def and mem_interval
+    prefer 4 apply auto by(erule_tac x=i in allE,erule_tac x=i in allE,fastsimp)+ 
+  { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i<DIM('a). c$$i < d$$i"
+    fix i assume i:"i<DIM('a)"
     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" 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)
+    hence "a$$i \<le> c$$i \<and> d$$i \<le> b$$i" using part1 and as(2) using i by auto  } note * = this
+  show ?th4 unfolding subset_eq and Ball_def and mem_interval 
+    apply(rule,rule,rule,rule) apply(rule *) unfolding subset_eq and Ball_def and mem_interval prefer 4
+    apply auto by(erule_tac x=i in allE, simp)+ 
+qed
+
+lemma disjoint_interval: fixes a::"'a::ordered_euclidean_space" shows
+  "{a .. b} \<inter> {c .. d} = {} \<longleftrightarrow> (\<exists>i<DIM('a). (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<DIM('a). (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<DIM('a). (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<DIM('a). (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" shows
- "{a .. b} \<inter> {c .. d} =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
+  let ?z = "(\<chi>\<chi> i. ((max (a$$i) (c$$i)) + (min (b$$i) (d$$i))) / 2)::'a"
+  note * = expand_set_eq Int_iff empty_iff mem_interval all_conj_distrib[THEN sym] eq_False
+  show ?th1 unfolding * apply safe apply(erule_tac x="?z" in allE)
+    unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
+  show ?th2 unfolding * apply safe apply(erule_tac x="?z" in allE)
+    unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
+  show ?th3 unfolding * apply safe apply(erule_tac x="?z" in allE)
+    unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
+  show ?th4 unfolding * apply safe apply(erule_tac x="?z" in allE)
+    unfolding not_all apply(erule exE,rule_tac x=x in exI) apply(erule_tac[2-] x=i in allE) by auto
+qed
+
+lemma inter_interval: fixes a :: "'a::ordered_euclidean_space" shows
+ "{a .. b} \<inter> {c .. d} =  {(\<chi>\<chi> i. max (a$$i) (c$$i)) .. (\<chi>\<chi> i. min (b$$i) (d$$i))}"
   unfolding expand_set_eq and Int_iff and mem_interval
   by auto
 
@@ -4762,57 +4780,55 @@
  "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[intro]: fixes a :: "real^'n" shows "open {a<..<b}"
+lemma open_interval[intro]: fixes a :: "'a::ordered_euclidean_space" 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"
+    { fix i assume "i<DIM('a)"
+      hence "\<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
+        using open_interval_lemma[of "a$$i" "x$$i" "b$$i"] by auto  }
+    hence "\<forall>i\<in>{..<DIM('a)}. \<exists>d>0. \<forall>x'. abs (x' - (x$$i)) < d \<longrightarrow> a$$i < x' \<and> x' < b$$i" by auto
+    from bchoice[OF this] guess d .. note d=this
+    let ?d = "Min (d ` {..<DIM('a)})"
+    have **:"finite (d ` {..<DIM('a)})" "d ` {..<DIM('a)} \<noteq> {}" by auto
+    have "?d>0" using 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"
+      { fix i assume i:"i<DIM('a)"
+        hence "\<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)
+          unfolding euclidean_simps Min_gr_iff[OF **] by auto
+        hence "a $$ i < x' $$ i" "x' $$ i < b $$ i" using i and d[THEN bspec[where x=i]] by auto  }
+      hence "a < x' \<and> x' < b" apply(subst(2) eucl_less,subst(1) eucl_less) by auto  }
+    ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by auto
   }
   thus ?thesis unfolding open_dist using open_interval_lemma by auto
 qed
 
-lemma open_interval_real[intro]: fixes a :: "real" shows "open {a<..<b}"
-  by (rule open_real_greaterThanLessThan)
-
-lemma closed_interval[intro]: fixes a :: "real^'n" shows "closed {a .. b}"
+lemma closed_interval[intro]: fixes a :: "'a::ordered_euclidean_space" 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
+  { fix x i assume i:"i<DIM('a)"
+    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
+        using component_le_norm[of "y-x" i, unfolded euclidean_simps] and xa using i
+        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
+    { 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
+        using component_le_norm[of "y-x" i, unfolded euclidean_simps] and xb using i
+        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 }
+    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[intro]: fixes a :: "real^'n" shows
+lemma interior_closed_interval[intro]: fixes a :: "'a::ordered_euclidean_space" 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
@@ -4820,91 +4836,87 @@
   { 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
+    { fix i assume i:"i<DIM('a)"
       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"
+        unfolding norm_minus_cancel using norm_basis 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 "a $$ i < x $$ i" and "x $$ i < b $$ i" unfolding euclidean_simps
+        unfolding basis_component using `e>0` i 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" shows
- "bounded {a .. b}"
+lemma bounded_closed_interval: fixes a :: "'a::ordered_euclidean_space" 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)
+  let ?b = "\<Sum>i<DIM('a). \<bar>a$$i\<bar> + \<bar>b$$i\<bar>"
+  { fix x::"'a" assume x:"\<forall>i<DIM('a). a $$ i \<le> x $$ i \<and> x $$ i \<le> b $$ i"
+    { fix i assume "i<DIM('a)"
+      hence "\<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<DIM('a). \<bar>x $$ i\<bar>) \<le> ?b" apply-apply(rule setsum_mono) by auto
     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" shows
+lemma bounded_interval: fixes a :: "'a::ordered_euclidean_space" 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" shows
+lemma not_interval_univ: fixes a :: "'a::ordered_euclidean_space" shows
  "({a .. b} \<noteq> UNIV) \<and> ({a<..<b} \<noteq> UNIV)"
-  using bounded_interval[of a b]
+  using bounded_interval[of a b] by auto
+
+lemma compact_interval: fixes a :: "'a::ordered_euclidean_space" shows "compact {a .. b}"
+  using bounded_closed_imp_compact[of "{a..b}"] using bounded_interval[of a b]
   by auto
 
-lemma compact_interval: fixes a :: "real^'n" 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"
+lemma open_interval_midpoint: fixes a :: "'a::ordered_euclidean_space"
   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"
+  { fix i assume "i<DIM('a)"
+    hence "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  }
+      unfolding euclidean_simps by auto  }
   thus ?thesis unfolding mem_interval by auto
 qed
 
-lemma open_closed_interval_convex: fixes x :: "real^'n"
+lemma open_closed_interval_convex: fixes x :: "'a::ordered_euclidean_space"
   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)
+  { fix i assume i:"i<DIM('a)"
+    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
+      using x unfolding mem_interval using i apply simp
+      using y unfolding mem_interval using i apply simp
       done
-    finally have "a $ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $ i" by auto
+    finally have "a $$ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i" unfolding euclidean_simps 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)
+    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
+      using x unfolding mem_interval using i apply simp
+      using y unfolding mem_interval using i 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 }
+    finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) $$ i < b $$ i" unfolding euclidean_simps 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"
+lemma closure_open_interval: fixes a :: "'a::ordered_euclidean_space"
   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
+  have ab:"a < b" using assms[unfolded interval_ne_empty] apply(subst eucl_less) 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)"
@@ -4914,7 +4926,7 @@
         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_ssub_ldistrib)  }
+      hence False using fn unfolding f_def using xc by auto  }
     moreover
     { assume "\<not> (f ---> x) sequentially"
       { fix e::real assume "e>0"
@@ -4932,26 +4944,25 @@
   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) set"
+lemma bounded_subset_open_interval_symmetric: fixes s::"('a::ordered_euclidean_space) 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"
+  def a \<equiv> "(\<chi>\<chi> i. b+1)::'a"
   { 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)
+    fix i assume i:"i<DIM('a)"
+    hence "(-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 euclidean_simps and a_def by auto  }
+  thus ?thesis by(auto intro: exI[where x=a] simp add: eucl_less[where 'a='a])
 qed
 
 lemma bounded_subset_open_interval:
-  fixes s :: "(real ^ 'n) set"
+  fixes s :: "('a::ordered_euclidean_space) 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) set"
+  fixes s :: "('a::ordered_euclidean_space) 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
@@ -4959,17 +4970,17 @@
 qed
 
 lemma bounded_subset_closed_interval:
-  fixes s :: "(real ^ 'n) set"
+  fixes s :: "('a::ordered_euclidean_space) 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 ^ _"
+  fixes a b :: "'a::ordered_euclidean_space"
   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 ^ _"
+  fixes a b :: "'a::ordered_euclidean_space"
   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
@@ -4977,42 +4988,47 @@
   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"
+lemma inter_interval_mixed_eq_empty: fixes a :: "'a::ordered_euclidean_space"
   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 stuff for half-infinite intervals too; FIXME: notation?  *)
 
-lemma closed_interval_left: fixes b::"real^'n"
-  shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
+lemma closed_interval_left: fixes b::"'a::ordered_euclidean_space"
+  shows "closed {x::'a. \<forall>i<DIM('a). 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  }
+  { fix i assume i:"i<DIM('a)"
+    fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i<DIM('a). 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"]] using i by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_norm euclidean_simps using i 
+        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"
-  shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
+lemma closed_interval_right: fixes a::"'a::ordered_euclidean_space"
+  shows "closed {x::'a. \<forall>i<DIM('a). 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  }
+  { fix i assume i:"i<DIM('a)"
+    fix x::"'a" assume x:"\<forall>e>0. \<exists>x'\<in>{x. \<forall>i<DIM('a). 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"]] i by auto
+      hence False using component_le_norm[of "y - x" i] unfolding dist_norm and euclidean_simps by auto   }
+    hence "a$$i \<le> x$$i" by(rule ccontr)auto  }
   thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast
 qed
 
 text {* 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}" (is ?th1) "is_interval {a<..<b}" (is ?th2) proof - 
+definition "is_interval (s::('a::ordered_euclidean_space) set) \<longleftrightarrow>
+  (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i<DIM('a). ((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::'a::ordered_euclidean_space}" (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 order_trans le_less_trans less_le_trans *)+ qed
@@ -5061,12 +5077,12 @@
 qed
 
 lemma closed_halfspace_component_le:
-  shows "closed {x::real^'n. x$i \<le> a}"
-  using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+  shows "closed {x::'a::ordered_euclidean_space. x$$i \<le> a}"
+  using closed_halfspace_le[of "(basis i)::'a" a] unfolding euclidean_component_def .
 
 lemma closed_halfspace_component_ge:
-  shows "closed {x::real^'n. x$i \<ge> a}"
-  using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+  shows "closed {x::'a::ordered_euclidean_space. x$$i \<ge> a}"
+  using closed_halfspace_ge[of a "(basis i)::'a"] unfolding euclidean_component_def .
 
 text{* Openness of halfspaces.                                                   *}
 
@@ -5083,38 +5099,39 @@
 qed
 
 lemma open_halfspace_component_lt:
-  shows "open {x::real^'n. x$i < a}"
-  using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto
+  shows "open {x::'a::ordered_euclidean_space. x$$i < a}"
+  using open_halfspace_lt[of "(basis i)::'a" a] unfolding euclidean_component_def .
 
 lemma open_halfspace_component_gt:
-  shows "open {x::real^'n. x$i  > a}"
-  using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding inner_basis[OF assms] by auto
+  shows "open {x::'a::ordered_euclidean_space. x$$i  > a}"
+  using open_halfspace_gt[of a "(basis i)::'a"] unfolding euclidean_component_def .
 
 text{* This gives a simple derivation of limit component bounds.                 *}
 
-lemma Lim_component_le: fixes f :: "'a \<Rightarrow> real^'n"
-  assumes "(f ---> l) net" "\<not> (trivial_limit net)"  "eventually (\<lambda>x. f(x)$i \<le> b) net"
-  shows "l$i \<le> b"
+lemma Lim_component_le: fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space"
+  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
+  { fix x have "x \<in> {x::'b. inner (basis i) x \<le> b} \<longleftrightarrow> x$$i \<le> b"
+      unfolding euclidean_component_def 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"
-  assumes "(f ---> l) net"  "\<not> (trivial_limit net)"  "eventually (\<lambda>x. b \<le> (f x)$i) net"
-  shows "b \<le> l$i"
+    using closed_halfspace_le[of "(basis i)::'b" b] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_ge: fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space"
+  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
+  { fix x have "x \<in> {x::'b. inner (basis i) x \<ge> b} \<longleftrightarrow> x$$i \<ge> b"
+      unfolding euclidean_component_def 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"
-  assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$i = b) net"
-  shows "l$i = b"
+    using closed_halfspace_ge[of b "(basis i)"] and assms(1,2,3) by auto
+qed
+
+lemma Lim_component_eq: fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space"
+  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
-
 text{* Limits relative to a union.                                               *}
 
 lemma eventually_within_Un:
@@ -5171,9 +5188,10 @@
   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" 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)
+lemma connected_ivt_component: fixes x::"'a::ordered_euclidean_space" 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)::'a" a]
+  unfolding euclidean_component_def by auto
 
 subsection {* Homeomorphisms *}
 
@@ -5266,7 +5284,7 @@
     (* 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 (auto intro: homeomorphism_compact)
+  unfolding homeomorphic_def by(auto intro: homeomorphism_compact)
 
 text{* Preservation of topological properties.                                   *}
 
@@ -5335,7 +5353,7 @@
 text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
 
 lemma cauchy_isometric:
-  fixes x :: "nat \<Rightarrow> real ^ 'n"
+  fixes x :: "nat \<Rightarrow> 'a::euclidean_space"
   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-
@@ -5355,7 +5373,7 @@
 qed
 
 lemma complete_isometric_image:
-  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
   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-
@@ -5378,10 +5396,10 @@
   shows "dist 0 x = norm x"
 unfolding dist_norm by simp
 
-lemma injective_imp_isometric: fixes f::"real^'m \<Rightarrow> real^'n"
+lemma injective_imp_isometric: fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes s:"closed s"  "subspace s"  and f:"bounded_linear f" "\<forall>x\<in>s. (f x = 0) \<longrightarrow> (x = 0)"
   shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm(x)"
-proof(cases "s \<subseteq> {0::real^'m}")
+proof(cases "s \<subseteq> {0::'a}")
   case True
   { fix x assume "x \<in> s"
     hence "x = 0" using True by auto
@@ -5393,8 +5411,8 @@
   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}"
+  let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}"
+  let ?S'' = "{x::'a. norm x = norm a}"
 
   have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by auto
   hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto
@@ -5419,7 +5437,7 @@
     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
+      have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" using s[unfolded subspace_def] 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`
@@ -5430,7 +5448,7 @@
 qed
 
 lemma closed_injective_image_subspace:
-  fixes f :: "real ^ _ \<Rightarrow> real ^ _"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 --> x = 0" "closed s"
   shows "closed(f ` s)"
 proof-
@@ -5441,82 +5459,79 @@
 
 subsection{* Some properties of a canonical subspace.                                  *}
 
+(** move **)
+declare euclidean_component.zero[simp]  
+
 lemma subspace_substandard:
- "subspace {x::real^_. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
-  unfolding subspace_def by auto
+  "subspace {x::'a::euclidean_space. (\<forall>i<DIM('a). P i \<longrightarrow> x$$i = 0)}"
+  unfolding subspace_def by(auto simp add: euclidean_simps)
 
 lemma closed_substandard:
- "closed {x::real^'n. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
+ "closed {x::'a::euclidean_space. \<forall>i<DIM('a). 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}"
+  let ?D = "{i. P i} \<inter> {..<DIM('a)}"
+  let ?Bs = "{{x::'a. 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) }
+      hence x:"\<forall>i\<in>?D. x $$ i = 0" by auto
+      hence "x\<in> \<Inter> ?Bs" by(auto simp add: x euclidean_component_def) }
     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  }
+        then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::'a. inner (basis i) x = 0}" by auto
+        hence "x $$ i = 0" unfolding B using x unfolding euclidean_component_def by auto  }
       hence "x\<in>?A" by auto }
     ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" .. }
   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. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d" (is "dim ?A = _")
+lemma dim_substandard: assumes "d\<subseteq>{..<DIM('a::euclidean_space)}"
+  shows "dim {x::'a::euclidean_space. \<forall>i<DIM('a). 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
-
+  let ?D = "{..<DIM('a)}"
+  let ?B = "(basis::nat => 'a) ` d"
+  let ?bas = "basis::nat \<Rightarrow> 'a"
+  have "?B \<subseteq> ?A" by(auto simp add:basis_component)
   moreover
-  { fix x::"real^'n" assume "x\<in>?A"
-    with finite[of d]
-    have "x\<in> span ?B"
+  { fix x::"'a" assume "x\<in>?A"
+    hence "finite d" "x\<in>?A" using assms by(auto intro:finite_subset)
+    hence "x\<in> span ?B"
     proof(induct d arbitrary: x)
-      case empty hence "x=0" unfolding Cart_eq by auto
+      case empty hence "x=0" apply(subst euclidean_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
+      hence *:"\<forall>i<DIM('a). 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
+      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)
+        hence "y $$ i = 0" unfolding y_def 
+          using *[THEN spec[where x=i]] by(auto simp add: euclidean_simps basis_component) }
+      hence "y \<in> span (basis ` F)" using insert(3) by auto
+      hence "y \<in> span (basis ` (insert k F))"
         using span_mono[of "?bas ` F" "?bas ` (insert k F)"]
-        using image_mono[OF **, of basis] by auto
+        using image_mono[OF **, of basis] using assms 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))"
+      hence "x$$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
         using span_mul by auto
       ultimately
-      have "y + x$k *\<^sub>R basis k \<in> span (?bas ` (insert k F))"
+      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
-
+    hence "x\<in>{(basis i)::'a |i. i \<in> ?D}" using assms by auto  }
+  hence "independent ?B" using independent_mono[OF independent_basis, 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
+  hence *:"inj_on (basis::nat\<Rightarrow>'a) d" using subset_inj_on[OF basis_inj, of "d"] by auto
   have "card ?B = card d" unfolding card_image[OF *] by auto
-
   ultimately show ?thesis using dim_unique[of "basis ` d" ?A] by auto
 qed
 
@@ -5532,32 +5547,33 @@
 apply (subgoal_tac "A \<noteq> UNIV", auto)
 done
 
-lemma closed_subspace: fixes s::"(real^'n) set"
+lemma closed_subspace: fixes s::"('a::euclidean_space) 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 *:"dim s \<le> DIM('a)" using dim_subset_UNIV by auto
+  def d \<equiv> "{..<dim s}" have t:"card d = dim s" unfolding d_def by auto
+  let ?t = "{x::'a. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x$$i = 0}"
+  have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x $$ i = 0} = s \<and>
+      inj_on f {x::'a. \<forall>i<DIM('a). i \<notin> d \<longrightarrow> x $$ i = 0}"
+    apply(rule subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]])
+    using dim_substandard[of d,where 'a='a] and t unfolding d_def using * assms by auto
+  then guess f apply-by(erule exE conjE)+ note f = this
+  interpret f: bounded_linear f using f unfolding linear_conv_bounded_linear by auto
   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
+    unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto
 qed
 
 lemma complete_subspace:
-  fixes s :: "(real ^ _) set" shows "subspace s ==> complete s"
+  fixes s :: "('a::euclidean_space) set" shows "subspace s ==> complete s"
   using complete_eq_closed closed_subspace
   by auto
 
 lemma dim_closure:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::euclidean_space) set"
   shows "dim(closure s) = dim s" (is "?dc = ?d")
 proof-
   have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
@@ -5568,14 +5584,6 @@
 
 subsection {* 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::linordered_field) ==> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
   by (simp add: field_simps inverse_eq_divide)
@@ -5600,66 +5608,47 @@
  "(m::'a::linordered_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: field_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"
+  fixes a b c :: "'a::ordered_euclidean_space"
   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")
+proof(cases "m=0")  
   { fix x assume "x \<le> c" "c \<le> x"
-    hence "x=c" unfolding vector_le_def and Cart_eq by (auto intro: order_antisym) }
+    hence "x=c" unfolding eucl_le[where 'a='a] apply-
+      apply(subst euclidean_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_le_def)
+  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: eucl_le[where 'a='a])
   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_le_def by auto
+      unfolding eucl_le[where 'a='a] by(auto simp add: euclidean_simps)
   } 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_le_def by(auto simp add: mult_left_mono_neg)
+      unfolding eucl_le[where 'a='a] by(auto simp add: mult_left_mono_neg euclidean_simps)
   } 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_le_def
-      apply(auto simp add: 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 mult_commute diff_le_iff)
+      unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
+      apply(auto simp add: 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 mult_commute diff_le_iff euclidean_simps)
   } 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_le_def
-      apply(auto simp add: vector_smult_assoc pth_3[symmetric]
+      unfolding image_iff Bex_def mem_interval eucl_le[where 'a='a]
+      apply(auto simp add: 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 mult_commute diff_le_iff)
+      by(auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff euclidean_simps)
   }
   ultimately show ?thesis using False by auto
 qed
 
-lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n)) ` {a..b} =
+lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::_::ordered_euclidean_space)) ` {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
 
@@ -5912,4 +5901,8 @@
   ultimately show "\<exists>!x\<in>s. g x = x" using `a\<in>s` by blast
 qed
 
+
+(** TODO move this someplace else within this theory **)
+instance euclidean_space \<subseteq> banach ..
+
 end
--- a/src/HOL/Multivariate_Analysis/Vec1.thy	Mon Jun 21 14:07:00 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,405 +0,0 @@
-(*  Title:      Multivariate_Analysis/Vec1.thy
-    Author:     Amine Chaieb, University of Cambridge
-    Author:     Robert Himmelmann, TU Muenchen
-*)
-
-header {* Vectors of size 1, 2, or 3 *}
-
-theory Vec1
-imports Topology_Euclidean_Space
-begin
-
-text{* Some common special cases.*}
-
-lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
-  by (metis num1_eq_iff)
-
-lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
-  by auto (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 [simp]: "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)
-
-instantiation num1 :: cart_one begin
-instance proof
-  show "CARD(1) = Suc 0" by auto
-qed end
-
-(* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *)
-
-abbreviation vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x \<equiv> vec x"
-
-abbreviation dest_vec1:: "'a ^1 \<Rightarrow> 'a"
-  where "dest_vec1 x \<equiv> (x$1)"
-
-lemma vec1_component[simp]: "(vec1 x)$1 = x"
-  by simp
-
-lemma vec1_dest_vec1: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
-  by (simp_all add:  Cart_eq)
-
-declare vec1_dest_vec1(1) [simp]
-
-lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))"
-  by (metis vec1_dest_vec1(1))
-
-lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))"
-  by (metis vec1_dest_vec1(1))
-
-lemma vec1_eq[simp]:  "vec1 x = vec1 y \<longleftrightarrow> x = y"
-  by (metis vec1_dest_vec1(2))
-
-lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y"
-  by (metis vec1_dest_vec1(1))
-
-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)
-
-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)
-
-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{* 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
-  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
-
-lemma range_vec1[simp]:"range vec1 = UNIV" apply(rule set_ext,rule) unfolding image_iff defer
-  apply(rule_tac x="dest_vec1 x" in bexI) by auto
-
-lemma dest_vec1_lambda: "dest_vec1(\<chi> i. x i) = x 1"
-  by (simp)
-
-lemma dest_vec1_vec: "dest_vec1(vec x) = x"
-  by (simp)
-
-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
-  apply auto
-  done
-
-lemma norm_vec1 [simp]: "norm(vec1 x) = abs(x)"
-  by (simp add: vec_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(1) norm_vec1)
-
-lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component
-   vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def
-
-lemma bounded_linear_vec1:"bounded_linear (vec1::real\<Rightarrow>real^1)"
-  unfolding bounded_linear_def additive_def bounded_linear_axioms_def 
-  unfolding smult_conv_scaleR[THEN sym] unfolding vec1_dest_vec1_simps
-  apply(rule conjI) defer apply(rule conjI) defer apply(rule_tac x=1 in exI) by auto
-
-lemma linear_vmul_dest_vec1:
-  fixes f:: "real^_ \<Rightarrow> real^1"
-  shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
-  unfolding smult_conv_scaleR
-  by (rule linear_vmul_component)
-
-lemma linear_from_scalars:
-  assumes lf: "linear (f::real^1 \<Rightarrow> real^_)"
-  shows "f = (\<lambda>x. dest_vec1 x *s column 1 (matrix f))"
-  unfolding smult_conv_scaleR
-  apply (rule ext)
-  apply (subst matrix_works[OF lf, symmetric])
-  apply (auto simp add: Cart_eq matrix_vector_mult_def column_def mult_commute)
-  done
-
-lemma linear_to_scalars: assumes lf: "linear (f::real ^'n \<Rightarrow> real^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 row_def inner_vector_def mult_commute)
-  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 vec_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{* Lifting and dropping *}
-
-lemma continuous_on_o_dest_vec1: fixes f::"real \<Rightarrow> 'a::real_normed_vector"
-  assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)"
-  using assms unfolding continuous_on_iff apply safe
-  apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe
-  apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real 
-  apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:vector_le_def)
-
-lemma continuous_on_o_vec1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
-  assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)"
-  using assms unfolding continuous_on_iff apply safe
-  apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe
-  apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real 
-  apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:vector_le_def)
-
-lemma continuous_on_vec1:"continuous_on A (vec1::real\<Rightarrow>real^1)"
-  by(rule linear_continuous_on[OF bounded_linear_vec1])
-
-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_le_def)
-
-lemma vec1_interval:fixes a::"real" shows
-  "vec1 ` {a .. b} = {vec1 a .. vec1 b}"
-  "vec1 ` {a<..<b} = {vec1 a<..<vec1 b}"
-  apply(rule_tac[!] set_ext) unfolding image_iff vector_less_def unfolding mem_interval
-  unfolding forall_1 unfolding vec1_dest_vec1_simps
-  apply rule defer apply(rule_tac x="dest_vec1 x" in bexI) prefer 3 apply rule defer
-  apply(rule_tac x="dest_vec1 x" in bexI) by auto
-
-(* Some special cases for intervals in R^1.                                  *)
-
-lemma interval_cases_1: fixes x :: "real^1" shows
- "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
-  unfolding Cart_eq vector_less_def vector_le_def mem_interval by(auto simp del:dest_vec1_eq)
-
-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)"
-  unfolding Cart_eq vector_less_def vector_le_def mem_interval by(auto simp del:dest_vec1_eq)
-
-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 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 forall_1 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"
-unfolding set_eq_subset[of "{a .. b}" "{c .. d}"]
-unfolding subset_interval_1(1)[of a b c d]
-unfolding subset_interval_1(1)[of c d a b]
-by auto
-
-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 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_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq)
-
-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_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq)
-
-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] 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] by auto
-
-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  by auto
-qed
-
-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_le_def Cart_eq)
-
-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_le_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 open_dest_vec1_vimage: "open S \<Longrightarrow> open (dest_vec1 -` S)"
-unfolding open_vector_def forall_1 by auto
-
-lemma tendsto_dest_vec1 [tendsto_intros]:
-  "(f ---> l) net \<Longrightarrow> ((\<lambda>x. dest_vec1 (f x)) ---> dest_vec1 l) net"
-by(rule tendsto_Cart_nth)
-
-lemma continuous_dest_vec1: "continuous net f \<Longrightarrow> continuous net (\<lambda>x. dest_vec1 (f x))"
-  unfolding continuous_def by (rule tendsto_dest_vec1)
-
-lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))" 
-  apply safe defer apply(erule_tac x="vec1 x" in allE) 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 dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding dist_norm by auto
-
-lemma bounded_linear_vec1_dest_vec1: fixes f::"real \<Rightarrow> real"
-  shows "linear (vec1 \<circ> f \<circ> dest_vec1) = bounded_linear f" (is "?l = ?r") proof-
-  { assume ?l guess K using linear_bounded[OF `?l`] ..
-    hence "\<exists>K. \<forall>x. \<bar>f x\<bar> \<le> \<bar>x\<bar> * K" apply(rule_tac x=K in exI)
-      unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) }
-  thus ?thesis unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def
-    unfolding vec1_dest_vec1_simps by auto qed
-
-lemma vec1_le[simp]:fixes a::real shows "vec1 a \<le> vec1 b \<longleftrightarrow> a \<le> b"
-  unfolding vector_le_def by auto
-lemma vec1_less[simp]:fixes a::real shows "vec1 a < vec1 b \<longleftrightarrow> a < b"
-  unfolding vector_less_def by auto
-
-end