merged
authorhuffman
Thu, 19 Nov 2009 06:01:02 -0800
changeset 33780 3e7ab843d817
parent 33779 b8efeea2cebd (current diff)
parent 33763 b1fbd5f3cfb4 (diff)
child 33781 c7d32e726bb9
merged
--- a/Admin/isatest/isatest-makedist	Wed Nov 18 16:57:58 2009 -0800
+++ b/Admin/isatest/isatest-makedist	Thu Nov 19 06:01:02 2009 -0800
@@ -104,15 +104,15 @@
 sleep 15
 $SSH macbroy23 "$MAKEALL -l HOL images $HOME/settings/at-sml-dev-e"
 sleep 15
-$SSH atbroy99 "$MAKEALL $HOME/settings/at64-poly"
+$SSH macbroy24 "$MAKEALL $HOME/settings/at64-poly"
 sleep 15
 $SSH macbroy2 "$MAKEALL $HOME/settings/mac-poly64-M4; $MAKEALL $HOME/settings/mac-poly64-M8; $MAKEALL $HOME/settings/mac-poly-M4; $MAKEALL $HOME/settings/mac-poly-M8"
 sleep 15
 $SSH macbroy5 "$MAKEALL $HOME/settings/mac-poly"
 sleep 15
 $SSH macbroy6 "sleep 10800; $MAKEALL $HOME/settings/at-mac-poly-5.1-para"
-sleep 15
-$SSH atbroy51 "$HOME/admin/isatest/isatest-annomaly"
+#sleep 15
+#$SSH atbroy51 "$HOME/admin/isatest/isatest-annomaly"
 
 echo ------------------- spawned tests successfully --- `date` --- $HOSTNAME >> $DISTLOG 2>&1
 
--- a/CONTRIBUTORS	Wed Nov 18 16:57:58 2009 -0800
+++ b/CONTRIBUTORS	Thu Nov 19 06:01:02 2009 -0800
@@ -7,6 +7,9 @@
 Contributions to this Isabelle version
 --------------------------------------
 
+* November 2009: Robert Himmelmann, TUM
+  Derivation and Brouwer's fixpoint theorem in Multivariate Analysis
+
 * November 2009: Stefan Berghofer, Lukas Bulwahn, TUM
   A tabled implementation of the reflexive transitive closure
 
--- a/NEWS	Wed Nov 18 16:57:58 2009 -0800
+++ b/NEWS	Thu Nov 19 06:01:02 2009 -0800
@@ -96,6 +96,14 @@
 zdiv_zmult_zmult2.  div_mult_mult1 is now [simp] by default.
 INCOMPATIBILITY.
 
+* Extended Multivariate Analysis to include derivation and Brouwer's fixpoint
+theorem.
+
+* Removed predicate "M hassize n" (<--> card M = n & finite M). Was only used
+in Multivariate Analysis. Renamed vector_less_eq_def to vector_le_def, the
+more usual name.
+INCOMPATIBILITY.
+
 * Added theorem List.map_map as [simp]. Removed List.map_compose.
 INCOMPATIBILITY.
 
--- a/doc-src/Nitpick/nitpick.tex	Wed Nov 18 16:57:58 2009 -0800
+++ b/doc-src/Nitpick/nitpick.tex	Thu Nov 19 06:01:02 2009 -0800
@@ -2019,9 +2019,11 @@
 you will find on Kodkod's web site \cite{kodkod-2009}. Unlike the standard
 version of MiniSat, the JNI version can be used incrementally.
 
+%%% No longer true:
+%%% "It is bundled with Kodkodi and requires no further installation or
+%%% configuration steps. Alternatively,"
 \item[$\bullet$] \textbf{\textit{PicoSAT}}: PicoSAT is an efficient solver
-written in C. It is bundled with Kodkodi and requires no further installation or
-configuration steps. Alternatively, you can install a standard version of
+written in C. You can install a standard version of
 PicoSAT and set the environment variable \texttt{PICOSAT\_HOME} to the directory
 that contains the \texttt{picosat} executable. The C sources for PicoSAT are
 available at \url{http://fmv.jku.at/picosat/} and are also bundled with Kodkodi.
@@ -2078,11 +2080,11 @@
 \url{http://cs.technion.ac.il/~gershman/HaifaSat.htm}.
 
 \item[$\bullet$] \textbf{\textit{smart}}: If \textit{sat\_solver} is set to
-\textit{smart}, Nitpick selects the first solver among MiniSatJNI, MiniSat,
-PicoSAT, zChaffJNI, zChaff, RSat, BerkMin, BerkMinAlloy, and Jerusat that is
-recognized by Isabelle. If none is found, it falls back on SAT4J, which should
-always be available. If \textit{verbose} is enabled, Nitpick displays which SAT
-solver was chosen.
+\textit{smart}, Nitpick selects the first solver among MiniSat,
+PicoSAT, zChaff, RSat, BerkMin, BerkMinAlloy, Jerusat, MiniSatJNI, and zChaffJNI
+that is recognized by Isabelle. If none is found, it falls back on SAT4J, which
+should always be available. If \textit{verbose} (\S\ref{output-format}) is
+enabled, Nitpick displays which SAT solver was chosen.
 \end{enum}
 \fussy
 
@@ -2469,8 +2471,8 @@
 
 \item[$\bullet$] Local definitions are not supported and result in an error.
 
-\item[$\bullet$] All constants and types whose names start with
-\textit{Nitpick}{.} are reserved for internal use.
+%\item[$\bullet$] All constants and types whose names start with
+%\textit{Nitpick}{.} are reserved for internal use.
 \end{enum}
 
 \let\em=\sl
--- a/doc-src/TutorialI/Rules/Primes.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/doc-src/TutorialI/Rules/Primes.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -112,13 +112,13 @@
 (*uniqueness of GCDs*)
 lemma is_gcd_unique: "\<lbrakk> is_gcd m a b; is_gcd n a b \<rbrakk> \<Longrightarrow> m=n"
 apply (simp add: is_gcd_def);
-apply (blast intro: dvd_anti_sym)
+apply (blast intro: dvd_antisym)
 done
 
 
 text {*
-@{thm[display] dvd_anti_sym}
-\rulename{dvd_anti_sym}
+@{thm[display] dvd_antisym}
+\rulename{dvd_antisym}
 
 \begin{isabelle}
 proof\ (prove):\ step\ 1\isanewline
@@ -154,7 +154,7 @@
   apply (auto intro: dvd_trans [of _ m])
   done
 
-(*This is half of the proof (by dvd_anti_sym) of*)
+(*This is half of the proof (by dvd_antisym) of*)
 lemma gcd_mult_cancel: "gcd k n = 1 \<Longrightarrow> gcd (k*m) n = gcd m n"
   oops
 
--- a/doc-src/TutorialI/Types/Numbers.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/doc-src/TutorialI/Types/Numbers.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -118,8 +118,8 @@
 @{thm[display] mod_by_0 [no_vars]}
 \rulename{mod_by_0}
 
-@{thm[display] dvd_anti_sym[no_vars]}
-\rulename{dvd_anti_sym}
+@{thm[display] dvd_antisym[no_vars]}
+\rulename{dvd_antisym}
 
 @{thm[display] dvd_add[no_vars]}
 \rulename{dvd_add}
--- a/doc-src/TutorialI/Types/document/Numbers.tex	Wed Nov 18 16:57:58 2009 -0800
+++ b/doc-src/TutorialI/Types/document/Numbers.tex	Thu Nov 19 06:01:02 2009 -0800
@@ -274,7 +274,7 @@
 \begin{isabelle}%
 {\isasymlbrakk}m\ dvd\ n{\isacharsemicolon}\ n\ dvd\ m{\isasymrbrakk}\ {\isasymLongrightarrow}\ m\ {\isacharequal}\ n%
 \end{isabelle}
-\rulename{dvd_anti_sym}
+\rulename{dvd_antisym}
 
 \begin{isabelle}%
 {\isasymlbrakk}a\ dvd\ b{\isacharsemicolon}\ a\ dvd\ c{\isasymrbrakk}\ {\isasymLongrightarrow}\ a\ dvd\ b\ {\isacharplus}\ c%
--- a/doc-src/TutorialI/Types/numerics.tex	Wed Nov 18 16:57:58 2009 -0800
+++ b/doc-src/TutorialI/Types/numerics.tex	Thu Nov 19 06:01:02 2009 -0800
@@ -1,4 +1,3 @@
-
 \section{Numbers}
 \label{sec:numbers}
 
@@ -192,7 +191,7 @@
 relation.  Here are some of the facts proved about it:
 \begin{isabelle}
 \isasymlbrakk m\ dvd\ n;\ n\ dvd\ m\isasymrbrakk \ \isasymLongrightarrow \ m\ =\ n%
-\rulenamedx{dvd_anti_sym}\isanewline
+\rulenamedx{dvd_antisym}\isanewline
 \isasymlbrakk k\ dvd\ m;\ k\ dvd\ n\isasymrbrakk \ \isasymLongrightarrow \ k\ dvd\ (m\ +\ n)
 \rulenamedx{dvd_add}
 \end{isabelle}
@@ -348,8 +347,7 @@
 \ 1.\ P\ (2\ /\ 5)
 \end{isabelle}
 Exponentiation can express floating-point values such as
-\isa{2 * 10\isacharcircum6}, but at present no special simplification
-is performed.
+\isa{2 * 10\isacharcircum6}, which will be simplified to integers.
 
 \begin{warn}
 Types \isa{rat}, \isa{real} and \isa{complex} are provided by theory HOL-Complex, which is
--- a/src/HOL/Divides.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Divides.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -2030,9 +2030,11 @@
                       split_neg_lemma [of concl: "%x y. P y"])
 done
 
-(* Enable arith to deal with div 2 and mod 2: *)
-declare split_zdiv [of _ _ "number_of k", simplified, standard, arith_split]
-declare split_zmod [of _ _ "number_of k", simplified, standard, arith_split]
+text {* Enable (lin)arith to deal with @{const div} and @{const mod}
+  when these are applied to some constant that is of the form
+  @{term "number_of k"}: *}
+declare split_zdiv [of _ _ "number_of k", standard, arith_split]
+declare split_zmod [of _ _ "number_of k", standard, arith_split]
 
 
 subsubsection{*Speeding up the Division Algorithm with Shifting*}
--- a/src/HOL/Extraction.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Extraction.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -13,20 +13,6 @@
 subsection {* Setup *}
 
 setup {*
-let
-fun realizes_set_proc (Const ("realizes", Type ("fun", [Type ("Null", []), _])) $ r $
-      (Const ("op :", _) $ x $ S)) = (case strip_comb S of
-        (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, U), ts @ [x]))
-      | (Free (s, U), ts) => SOME (list_comb (Free (s, U), ts @ [x]))
-      | _ => NONE)
-  | realizes_set_proc (Const ("realizes", Type ("fun", [T, _])) $ r $
-      (Const ("op :", _) $ x $ S)) = (case strip_comb S of
-        (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts @ [x]))
-      | (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts @ [x]))
-      | _ => NONE)
-  | realizes_set_proc _ = NONE;
-
-in
   Extraction.add_types
       [("bool", ([], NONE))] #>
   Extraction.set_preprocessor (fn thy =>
@@ -35,7 +21,6 @@
       Proofterm.rewrite_proof thy
         (RewriteHOLProof.rews, ProofRewriteRules.rprocs true) o
       ProofRewriteRules.elim_vars (curry Const @{const_name default}))
-end
 *}
 
 lemmas [extraction_expand] =
--- a/src/HOL/HOL.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/HOL.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -2060,7 +2060,6 @@
 setup {*
   Predicate_Compile_Alternative_Defs.setup
   #> Predicate_Compile_Inline_Defs.setup
-  #> Predicate_Compile_Preproc_Const_Defs.setup
 *}
 
 
--- a/src/HOL/IsaMakefile	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/IsaMakefile	Thu Nov 19 06:01:02 2009 -0800
@@ -1058,7 +1058,9 @@
   Multivariate_Analysis/Finite_Cartesian_Product.thy	\
   Multivariate_Analysis/Euclidean_Space.thy		\
   Multivariate_Analysis/Topology_Euclidean_Space.thy	\
-  Multivariate_Analysis/Convex_Euclidean_Space.thy
+  Multivariate_Analysis/Convex_Euclidean_Space.thy      \
+  Multivariate_Analysis/Brouwer_Fixpoint.thy            \
+  Multivariate_Analysis/Derivative.thy
 	@cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis
 
 
@@ -1340,7 +1342,9 @@
   SMT/Examples/cert/z3_linarith_19					\
   SMT/Examples/cert/z3_linarith_19.proof				\
   SMT/Examples/cert/z3_linarith_20					\
-  SMT/Examples/cert/z3_linarith_20.proof SMT/Examples/cert/z3_mono_01	\
+  SMT/Examples/cert/z3_linarith_20.proof 				\
+  SMT/Examples/cert/z3_linarith_21					\
+  SMT/Examples/cert/z3_linarith_21.proof SMT/Examples/cert/z3_mono_01	\
   SMT/Examples/cert/z3_mono_01.proof SMT/Examples/cert/z3_mono_02	\
   SMT/Examples/cert/z3_mono_02.proof SMT/Examples/cert/z3_nat_arith_01	\
   SMT/Examples/cert/z3_nat_arith_01.proof				\
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -0,0 +1,1983 @@
+
+(* ========================================================================= *)
+(* Results connected with topological dimension.                             *)
+(*                                                                           *)
+(* At the moment this is just Brouwer's fixpoint theorem. The proof is from  *)
+(* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518   *)
+(* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf".          *)
+(*                                                                           *)
+(* The script below is quite messy, but at least we avoid formalizing any    *)
+(* topological machinery; we don't even use barycentric subdivision; this is *)
+(* the big advantage of Kuhn's proof over the usual Sperner's lemma one.     *)
+(*                                                                           *)
+(*              (c) Copyright, John Harrison 1998-2008                       *)
+(* ========================================================================= *)
+
+(* Author:                     John Harrison
+   Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+
+header {* Results connected with topological dimension. *}
+
+theory Brouwer_Fixpoint
+  imports Convex_Euclidean_Space
+begin
+
+declare norm_scaleR[simp]
+ 
+lemma brouwer_compactness_lemma:
+  assumes "compact s" "continuous_on s f" "\<not> (\<exists>x\<in>s. (f x = (0::real^'n::finite)))"
+  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"
+    using continuous_attains_inf[OF assms(1), of "norm \<circ> f"] and False unfolding o_def by auto
+  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^'n. P x \<longrightarrow> P (f x))"  "\<forall>x. P x \<longrightarrow> (\<forall>i::'n. 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 
+ 
+subsection {* The key "counting" observation, somewhat abstracted. *}
+
+lemma setsum_Un_disjoint':assumes "finite A" "finite B" "A \<inter> B = {}" "A \<union> B = C"
+  shows "setsum g C = setsum g A + setsum g B"
+  using setsum_Un_disjoint[OF assms(1-3)] and assms(4) by auto
+
+lemma kuhn_counting_lemma: assumes "finite faces" "finite simplices"
+  "\<forall>f\<in>faces. bnd f  \<longrightarrow> (card {s \<in> simplices. face f s} = 1)"
+  "\<forall>f\<in>faces. \<not> bnd f \<longrightarrow> (card {s \<in> simplices. face f s} = 2)"
+  "\<forall>s\<in>simplices. compo s  \<longrightarrow> (card {f \<in> faces. face f s \<and> compo' f} = 1)"
+  "\<forall>s\<in>simplices. \<not> compo s \<longrightarrow> (card {f \<in> faces. face f s \<and> compo' f} = 0) \<or>
+                             (card {f \<in> faces. face f s \<and> compo' f} = 2)"
+    "odd(card {f \<in> faces. compo' f \<and> bnd f})"
+  shows "odd(card {s \<in> simplices. compo s})" proof-
+  have "\<And>x. {f\<in>faces. compo' f \<and> bnd f \<and> face f x} \<union> {f\<in>faces. compo' f \<and> \<not>bnd f \<and> face f x} = {f\<in>faces. compo' f \<and> face f x}"
+    "\<And>x. {f \<in> faces. compo' f \<and> bnd f \<and> face f x} \<inter> {f \<in> faces. compo' f \<and> \<not> bnd f \<and> face f x} = {}" by auto
+  hence lem1:"setsum (\<lambda>s. (card {f \<in> faces. face f s \<and> compo' f})) simplices =
+    setsum (\<lambda>s. card {f \<in> {f \<in> faces. compo' f \<and> bnd f}. face f s}) simplices +
+    setsum (\<lambda>s. card {f \<in> {f \<in> faces. compo' f \<and> \<not> (bnd f)}. face f s}) simplices"
+    unfolding setsum_addf[THEN sym] apply- apply(rule setsum_cong2)
+    using assms(1) by(auto simp add: card_Un_Int, auto simp add:conj_commute)
+  have lem2:"setsum (\<lambda>j. card {f \<in> {f \<in> faces. compo' f \<and> bnd f}. face f j}) simplices = 
+              1 * card {f \<in> faces. compo' f \<and> bnd f}"
+       "setsum (\<lambda>j. card {f \<in> {f \<in> faces. compo' f \<and> \<not> bnd f}. face f j}) simplices = 
+              2 * card {f \<in> faces. compo' f \<and> \<not> bnd f}"
+    apply(rule_tac[!] setsum_multicount) using assms by auto
+  have lem3:"setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) simplices =
+    setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices.   compo s}+
+    setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. \<not> compo s}"
+    apply(rule setsum_Un_disjoint') using assms(2) by auto
+  have lem4:"setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. compo s}
+    = setsum (\<lambda>s. 1) {s \<in> simplices. compo s}"
+    apply(rule setsum_cong2) using assms(5) by auto
+  have lem5: "setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f}) {s \<in> simplices. \<not> compo s} =
+    setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f})
+           {s \<in> simplices. (\<not> compo s) \<and> (card {f \<in> faces. face f s \<and> compo' f} = 0)} +
+    setsum (\<lambda>s. card {f \<in> faces. face f s \<and> compo' f})
+           {s \<in> simplices. (\<not> compo s) \<and> (card {f \<in> faces. face f s \<and> compo' f} = 2)}"
+    apply(rule setsum_Un_disjoint') using assms(2,6) by auto
+  have *:"int (\<Sum>s\<in>{s \<in> simplices. compo s}. card {f \<in> faces. face f s \<and> compo' f}) =
+    int (card {f \<in> faces. compo' f \<and> bnd f} + 2 * card {f \<in> faces. compo' f \<and> \<not> bnd f}) - 
+    int (card {s \<in> simplices. \<not> compo s \<and> card {f \<in> faces. face f s \<and> compo' f} = 2} * 2)"
+    using lem1[unfolded lem3 lem2 lem5] by auto
+  have even_minus_odd:"\<And>x y. even x \<Longrightarrow> odd (y::int) \<Longrightarrow> odd (x - y)" using assms by auto
+  have odd_minus_even:"\<And>x y. odd x \<Longrightarrow> even (y::int) \<Longrightarrow> odd (x - y)" using assms by auto
+  show ?thesis unfolding even_nat_def unfolding card_def and lem4[THEN sym] and *[unfolded card_def]
+    unfolding card_def[THEN sym] apply(rule odd_minus_even) unfolding zadd_int[THEN sym] apply(rule odd_plus_even)
+    apply(rule assms(7)[unfolded even_nat_def]) unfolding int_mult by auto qed
+
+subsection {* The odd/even result for faces of complete vertices, generalized. *}
+
+lemma card_1_exists: "card s = 1 \<longleftrightarrow> (\<exists>!x. x \<in> s)" unfolding One_nat_def
+  apply rule apply(drule card_eq_SucD) defer apply(erule ex1E) proof-
+  fix x assume as:"x \<in> s" "\<forall>y. y \<in> s \<longrightarrow> y = x"
+  have *:"s = insert x {}" apply- apply(rule set_ext,rule) unfolding singleton_iff
+    apply(rule as(2)[rule_format]) using as(1) by auto
+  show "card s = Suc 0" unfolding * using card_insert by auto qed auto
+
+lemma card_2_exists: "card s = 2 \<longleftrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. (z = x) \<or> (z = y)))" proof
+  assume "card s = 2" then obtain x y where obt:"s = {x, y}" "x\<noteq>y" unfolding numeral_2_eq_2 apply - apply(erule exE conjE|drule card_eq_SucD)+ by auto
+  show "\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y)" using obt by auto next
+  assume "\<exists>x\<in>s. \<exists>y\<in>s. x \<noteq> y \<and> (\<forall>z\<in>s. z = x \<or> z = y)" then guess x .. from this(2) guess y  ..
+  with `x\<in>s` have *:"s = {x, y}" "x\<noteq>y" by auto
+  from this(2) show "card s = 2" unfolding * by auto qed
+
+lemma image_lemma_0: assumes "card {a\<in>s. f ` (s - {a}) = t - {b}} = n"
+  shows "card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> (f ` s' = t - {b})} = n" proof-
+  have *:"{s'. \<exists>a\<in>s. (s' = s - {a}) \<and> (f ` s' = t - {b})} = (\<lambda>a. s - {a}) ` {a\<in>s. f ` (s - {a}) = t - {b}}" by auto
+  show ?thesis unfolding * unfolding assms[THEN sym] apply(rule card_image) unfolding inj_on_def 
+    apply(rule,rule,rule) unfolding mem_Collect_eq by auto qed
+
+lemma image_lemma_1: assumes "finite s" "finite t" "card s = card t" "f ` s = t" "b \<in> t"
+  shows "card {s'. \<exists>a\<in>s. s' = s - {a} \<and>  f ` s' = t - {b}} = 1" proof-
+  obtain a where a:"b = f a" "a\<in>s" using assms(4-5) by auto
+  have inj:"inj_on f s" apply(rule eq_card_imp_inj_on) using assms(1-4) by auto
+  have *:"{a \<in> s. f ` (s - {a}) = t - {b}} = {a}" apply(rule set_ext) unfolding singleton_iff
+    apply(rule,rule inj[unfolded inj_on_def,rule_format]) unfolding a using a(2) and assms and inj[unfolded inj_on_def] by auto
+  show ?thesis apply(rule image_lemma_0) unfolding *  by auto qed
+
+lemma image_lemma_2: assumes "finite s" "finite t" "card s = card t" "f ` s \<subseteq> t" "f ` s \<noteq> t" "b \<in> t"
+  shows "(card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 0) \<or>
+         (card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 2)" proof(cases "{a\<in>s. f ` (s - {a}) = t - {b}} = {}")
+  case True thus ?thesis apply-apply(rule disjI1, rule image_lemma_0) using assms(1) by(auto simp add:card_0_eq)
+next let ?M = "{a\<in>s. f ` (s - {a}) = t - {b}}"
+  case False then obtain a where "a\<in>?M" by auto hence a:"a\<in>s" "f ` (s - {a}) = t - {b}" by auto
+  have "f a \<in> t - {b}" using a and assms by auto
+  hence "\<exists>c \<in> s - {a}. f a = f c" unfolding image_iff[symmetric] and a by auto
+  then obtain c where c:"c \<in> s" "a \<noteq> c" "f a = f c" by auto
+  hence *:"f ` (s - {c}) = f ` (s - {a})" apply-apply(rule set_ext,rule) proof-
+    fix x assume "x \<in> f ` (s - {a})" then obtain y where y:"f y = x" "y\<in>s- {a}" by auto
+    thus "x \<in> f ` (s - {c})" unfolding image_iff apply(rule_tac x="if y = c then a else y" in bexI) using c a by auto qed auto
+  have "c\<in>?M" unfolding mem_Collect_eq and * using a and c(1) by auto
+  show ?thesis apply(rule disjI2, rule image_lemma_0) unfolding card_2_exists
+    apply(rule bexI[OF _ `a\<in>?M`], rule bexI[OF _ `c\<in>?M`],rule,rule `a\<noteq>c`) proof(rule,unfold mem_Collect_eq,erule conjE)
+    fix z assume as:"z \<in> s" "f ` (s - {z}) = t - {b}"
+    have inj:"inj_on f (s - {z})" apply(rule eq_card_imp_inj_on) unfolding as using as(1) and assms by auto
+    show "z = a \<or> z = c" proof(rule ccontr)
+      assume "\<not> (z = a \<or> z = c)" thus False using inj[unfolded inj_on_def,THEN bspec[where x=a],THEN bspec[where x=c]]
+	using `a\<in>s` `c\<in>s` `f a = f c` `a\<noteq>c` by auto qed qed qed
+
+subsection {* Combine this with the basic counting lemma. *}
+
+lemma kuhn_complete_lemma:
+  assumes "finite simplices"
+  "\<forall>f s. face f s \<longleftrightarrow> (\<exists>a\<in>s. f = s - {a})" "\<forall>s\<in>simplices. card s = n + 2" "\<forall>s\<in>simplices. (rl ` s) \<subseteq> {0..n+1}"
+  "\<forall>f\<in> {f. \<exists>s\<in>simplices. face f s}. bnd f  \<longrightarrow> (card {s\<in>simplices. face f s} = 1)"
+  "\<forall>f\<in> {f. \<exists>s\<in>simplices. face f s}. \<not>bnd f \<longrightarrow> (card {s\<in>simplices. face f s} = 2)"
+  "odd(card {f\<in>{f. \<exists>s\<in>simplices. face f s}. rl ` f = {0..n} \<and> bnd f})"
+  shows "odd (card {s\<in>simplices. (rl ` s = {0..n+1})})" 
+  apply(rule kuhn_counting_lemma) defer apply(rule assms)+ prefer 3 apply(rule assms) proof(rule_tac[1-2] ballI impI)+ 
+  have *:"{f. \<exists>s\<in>simplices. \<exists>a\<in>s. f = s - {a}} = (\<Union>s\<in>simplices. {f. \<exists>a\<in>s. (f = s - {a})})" by auto
+  have **: "\<forall>s\<in>simplices. card s = n + 2 \<and> finite s" using assms(3) by (auto intro: card_ge_0_finite)
+  show "finite {f. \<exists>s\<in>simplices. face f s}" unfolding assms(2)[rule_format] and *
+    apply(rule finite_UN_I[OF assms(1)]) using ** by auto
+  have *:"\<And> P f s. s\<in>simplices \<Longrightarrow> (f \<in> {f. \<exists>s\<in>simplices. \<exists>a\<in>s. f = s - {a}}) \<and>
+    (\<exists>a\<in>s. (f = s - {a})) \<and> P f \<longleftrightarrow> (\<exists>a\<in>s. (f = s - {a}) \<and> P f)" by auto
+  fix s assume s:"s\<in>simplices" let ?S = "{f \<in> {f. \<exists>s\<in>simplices. face f s}. face f s \<and> rl ` f = {0..n}}"
+    have "{0..n + 1} - {n + 1} = {0..n}" by auto
+    hence S:"?S = {s'. \<exists>a\<in>s. s' = s - {a} \<and> rl ` s' = {0..n + 1} - {n + 1}}" apply- apply(rule set_ext)
+      unfolding assms(2)[rule_format] mem_Collect_eq and *[OF s, unfolded mem_Collect_eq, where P="\<lambda>x. rl ` x = {0..n}"] by auto
+    show "rl ` s = {0..n+1} \<Longrightarrow> card ?S = 1" "rl ` s \<noteq> {0..n+1} \<Longrightarrow> card ?S = 0 \<or> card ?S = 2" unfolding S
+      apply(rule_tac[!] image_lemma_1 image_lemma_2) using ** assms(4) and s by auto qed
+
+subsection {*We use the following notion of ordering rather than pointwise indexing. *}
+
+definition "kle n x y \<longleftrightarrow> (\<exists>k\<subseteq>{1..n::nat}. (\<forall>j. y(j) = x(j) + (if j \<in> k then (1::nat) else 0)))"
+
+lemma kle_refl[intro]: "kle n x x" unfolding kle_def by auto
+
+lemma kle_antisym: "kle n x y \<and> kle n y x \<longleftrightarrow> (x = y)"
+  unfolding kle_def apply rule apply(rule ext) by auto
+
+lemma pointwise_minimal_pointwise_maximal: fixes s::"(nat\<Rightarrow>nat) set"
+  assumes  "finite s" "s \<noteq> {}" "\<forall>x\<in>s. \<forall>y\<in>s. (\<forall>j. x j \<le> y j) \<or> (\<forall>j. y j \<le> x j)"
+  shows "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<le> x j" "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. x j \<le> a j"
+  using assms unfolding atomize_conj apply- proof(induct s rule:finite_induct)
+  fix x and F::"(nat\<Rightarrow>nat) set" assume as:"finite F" "x \<notin> F" 
+    "\<lbrakk>F \<noteq> {}; \<forall>x\<in>F. \<forall>y\<in>F. (\<forall>j. x j \<le> y j) \<or> (\<forall>j. y j \<le> x j)\<rbrakk>
+        \<Longrightarrow> (\<exists>a\<in>F. \<forall>x\<in>F. \<forall>j. a j \<le> x j) \<and> (\<exists>a\<in>F. \<forall>x\<in>F. \<forall>j. x j \<le> a j)" "insert x F \<noteq> {}"
+    "\<forall>xa\<in>insert x F. \<forall>y\<in>insert x F. (\<forall>j. xa j \<le> y j) \<or> (\<forall>j. y j \<le> xa j)"
+  show "(\<exists>a\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. a j \<le> x j) \<and> (\<exists>a\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. x j \<le> a j)" proof(cases "F={}")
+    case True thus ?thesis apply-apply(rule,rule_tac[!] x=x in bexI) by auto next
+    case False obtain a b where a:"a\<in>insert x F" "\<forall>x\<in>F. \<forall>j. a j \<le> x j" and
+      b:"b\<in>insert x F" "\<forall>x\<in>F. \<forall>j. x j \<le> b j" using as(3)[OF False] using as(5) by auto
+    have "\<exists>a\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. a j \<le> x j"
+      using as(5)[rule_format,OF a(1) insertI1] apply- proof(erule disjE)
+      assume "\<forall>j. a j \<le> x j" thus ?thesis apply(rule_tac x=a in bexI) using a by auto next
+      assume "\<forall>j. x j \<le> a j" thus ?thesis apply(rule_tac x=x in bexI) apply(rule,rule) using a apply -
+	apply(erule_tac x=xa in ballE) apply(erule_tac x=j in allE)+ by auto qed moreover
+    have "\<exists>b\<in>insert x F. \<forall>x\<in>insert x F. \<forall>j. x j \<le> b j"
+      using as(5)[rule_format,OF b(1) insertI1] apply- proof(erule disjE)
+      assume "\<forall>j. x j \<le> b j" thus ?thesis apply(rule_tac x=b in bexI) using b by auto next
+      assume "\<forall>j. b j \<le> x j" thus ?thesis apply(rule_tac x=x in bexI) apply(rule,rule) using b apply -
+	apply(erule_tac x=xa in ballE) apply(erule_tac x=j in allE)+ by auto qed
+    ultimately show  ?thesis by auto qed qed(auto)
+
+lemma kle_imp_pointwise: "kle n x y \<Longrightarrow> (\<forall>j. x j \<le> y j)" unfolding kle_def by auto
+
+lemma pointwise_antisym: fixes x::"nat \<Rightarrow> nat"
+  shows "(\<forall>j. x j \<le> y j) \<and> (\<forall>j. y j \<le> x j) \<longleftrightarrow> (x = y)"
+  apply(rule, rule ext,erule conjE) apply(erule_tac x=xa in allE)+ by auto
+
+lemma kle_trans: assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x" shows "kle n x z"
+  using assms apply- apply(erule disjE) apply assumption proof- case goal1
+  hence "x=z" apply- apply(rule ext) apply(drule kle_imp_pointwise)+ 
+    apply(erule_tac x=xa in allE)+ by auto thus ?case by auto qed
+
+lemma kle_strict: assumes "kle n x y" "x \<noteq> y"
+  shows "\<forall>j. x j \<le> y j"  "\<exists>k. 1 \<le> k \<and> k \<le> n \<and> x(k) < y(k)"
+  apply(rule kle_imp_pointwise[OF assms(1)]) proof-
+  guess k using assms(1)[unfolded kle_def] .. note k = this
+  show "\<exists>k. 1 \<le> k \<and> k \<le> n \<and> x(k) < y(k)" proof(cases "k={}")
+    case True hence "x=y" apply-apply(rule ext) using k by auto
+    thus ?thesis using assms(2) by auto next
+    case False hence "(SOME k'. k' \<in> k) \<in> k" apply-apply(rule someI_ex) by auto
+    thus ?thesis apply(rule_tac x="SOME k'. k' \<in> k" in exI) using k by auto qed qed
+
+lemma kle_minimal: assumes "finite s" "s \<noteq> {}" "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x"
+  shows "\<exists>a\<in>s. \<forall>x\<in>s. kle n a x" proof-
+  have "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<le> x j" apply(rule pointwise_minimal_pointwise_maximal(1)[OF assms(1-2)])
+    apply(rule,rule) apply(drule_tac assms(3)[rule_format],assumption) using kle_imp_pointwise by auto
+  then guess a .. note a=this show ?thesis apply(rule_tac x=a in bexI) proof fix x assume "x\<in>s"
+    show "kle n a x" using assms(3)[rule_format,OF a(1) `x\<in>s`] apply- proof(erule disjE)
+      assume "kle n x a" hence "x = a" apply- unfolding pointwise_antisym[THEN sym]
+	apply(drule kle_imp_pointwise) using a(2)[rule_format,OF `x\<in>s`] by auto
+      thus ?thesis using kle_refl by auto  qed qed(insert a, auto) qed
+
+lemma kle_maximal: assumes "finite s" "s \<noteq> {}" "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x"
+  shows "\<exists>a\<in>s. \<forall>x\<in>s. kle n x a" proof-
+  have "\<exists>a\<in>s. \<forall>x\<in>s. \<forall>j. a j \<ge> x j" apply(rule pointwise_minimal_pointwise_maximal(2)[OF assms(1-2)])
+    apply(rule,rule) apply(drule_tac assms(3)[rule_format],assumption) using kle_imp_pointwise by auto
+  then guess a .. note a=this show ?thesis apply(rule_tac x=a in bexI) proof fix x assume "x\<in>s"
+    show "kle n x a" using assms(3)[rule_format,OF a(1) `x\<in>s`] apply- proof(erule disjE)
+      assume "kle n a x" hence "x = a" apply- unfolding pointwise_antisym[THEN sym]
+	apply(drule kle_imp_pointwise) using a(2)[rule_format,OF `x\<in>s`] by auto
+      thus ?thesis using kle_refl by auto  qed qed(insert a, auto) qed
+
+lemma kle_strict_set: assumes "kle n x y" "x \<noteq> y"
+  shows "1 \<le> card {k\<in>{1..n}. x k < y k}" proof-
+  guess i using kle_strict(2)[OF assms] ..
+  hence "card {i} \<le> card {k\<in>{1..n}. x k < y k}" apply- apply(rule card_mono) by auto
+  thus ?thesis by auto qed
+
+lemma kle_range_combine:
+  assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x"
+  "m1 \<le> card {k\<in>{1..n}. x k < y k}"
+  "m2 \<le> card {k\<in>{1..n}. y k < z k}"
+  shows "kle n x z \<and> m1 + m2 \<le> card {k\<in>{1..n}. x k < z k}"
+  apply(rule,rule kle_trans[OF assms(1-3)]) proof-
+  have "\<And>j. x j < y j \<Longrightarrow> x j < z j" apply(rule less_le_trans) using kle_imp_pointwise[OF assms(2)] by auto moreover
+  have "\<And>j. y j < z j \<Longrightarrow> x j < z j" apply(rule le_less_trans) using kle_imp_pointwise[OF assms(1)] by auto ultimately
+  have *:"{k\<in>{1..n}. x k < y k} \<union> {k\<in>{1..n}. y k < z k} = {k\<in>{1..n}. x k < z k}" by auto
+  have **:"{k \<in> {1..n}. x k < y k} \<inter> {k \<in> {1..n}. y k < z k} = {}" unfolding disjoint_iff_not_equal
+    apply(rule,rule,unfold mem_Collect_eq,rule ccontr) apply(erule conjE)+ proof-
+    fix i j assume as:"i \<in> {1..n}" "x i < y i" "j \<in> {1..n}" "y j < z j" "\<not> i \<noteq> j"
+    guess kx using assms(1)[unfolded kle_def] .. note kx=this
+    have "x i < y i" using as by auto hence "i \<in> kx" using as(1) kx apply(rule_tac ccontr) by auto 
+    hence "x i + 1 = y i" using kx by auto moreover
+    guess ky using assms(2)[unfolded kle_def] .. note ky=this
+    have "y i < z i" using as by auto hence "i \<in> ky" using as(1) ky apply(rule_tac ccontr) by auto 
+    hence "y i + 1 = z i" using ky by auto ultimately
+    have "z i = x i + 2" by auto
+    thus False using assms(3) unfolding kle_def by(auto simp add: split_if_eq1) qed
+  have fin:"\<And>P. finite {x\<in>{1..n::nat}. P x}" by auto
+  have "m1 + m2 \<le> card {k\<in>{1..n}. x k < y k} + card {k\<in>{1..n}. y k < z k}" using assms(4-5) by auto
+  also have "\<dots> \<le>  card {k\<in>{1..n}. x k < z k}" unfolding card_Un_Int[OF fin fin] unfolding * ** by auto
+  finally show " m1 + m2 \<le> card {k \<in> {1..n}. x k < z k}" by auto qed
+
+lemma kle_range_combine_l:
+  assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x" "m \<le> card {k\<in>{1..n}. y(k) < z(k)}"
+  shows "kle n x z \<and> m \<le> card {k\<in>{1..n}. x(k) < z(k)}"
+  using kle_range_combine[OF assms(1-3) _ assms(4), of 0] by auto
+
+lemma kle_range_combine_r:
+  assumes "kle n x y" "kle n y z" "kle n x z \<or> kle n z x" "m \<le> card {k\<in>{1..n}. x k < y k}"
+  shows "kle n x z \<and> m \<le> card {k\<in>{1..n}. x(k) < z(k)}"
+  using kle_range_combine[OF assms(1-3) assms(4), of 0] by auto
+
+lemma kle_range_induct:
+  assumes "card s = Suc m" "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x"
+  shows "\<exists>x\<in>s. \<exists>y\<in>s. kle n x y \<and> m \<le> card {k\<in>{1..n}. x k < y k}" proof-
+have "finite s" "s\<noteq>{}" using assms(1) by (auto intro: card_ge_0_finite)
+thus ?thesis using assms apply- proof(induct m arbitrary: s)
+  case 0 thus ?case using kle_refl by auto next
+  case (Suc m) then obtain a where a:"a\<in>s" "\<forall>x\<in>s. kle n a x" using kle_minimal[of s n] by auto
+  show ?case proof(cases "s \<subseteq> {a}") case False
+    hence "card (s - {a}) = Suc m" "s - {a} \<noteq> {}" using card_Diff_singleton[OF _ a(1)] Suc(4) `finite s` by auto
+    then obtain x b where xb:"x\<in>s - {a}" "b\<in>s - {a}" "kle n x b" "m \<le> card {k \<in> {1..n}. x k < b k}" 
+      using Suc(1)[of "s - {a}"] using Suc(5) `finite s` by auto
+    have "1 \<le> card {k \<in> {1..n}. a k < x k}" "m \<le> card {k \<in> {1..n}. x k < b k}"
+      apply(rule kle_strict_set) apply(rule a(2)[rule_format]) using a and xb by auto
+    thus ?thesis apply(rule_tac x=a in bexI, rule_tac x=b in bexI) 
+      using kle_range_combine[OF a(2)[rule_format] xb(3) Suc(5)[rule_format], of 1 "m"] using a(1) xb(1-2) by auto next
+    case True hence "s = {a}" using Suc(3) by auto hence "card s = 1" by auto
+    hence False using Suc(4) `finite s` by auto thus ?thesis by auto qed qed qed
+
+lemma kle_Suc: "kle n x y \<Longrightarrow> kle (n + 1) x y"
+  unfolding kle_def apply(erule exE) apply(rule_tac x=k in exI) by auto
+
+lemma kle_trans_1: assumes "kle n x y" shows "x j \<le> y j" "y j \<le> x j + 1"
+  using assms[unfolded kle_def] by auto 
+
+lemma kle_trans_2: assumes "kle n a b" "kle n b c" "\<forall>j. c j \<le> a j + 1" shows "kle n a c" proof-
+  guess kk1 using assms(1)[unfolded kle_def] .. note kk1=this
+  guess kk2 using assms(2)[unfolded kle_def] .. note kk2=this
+  show ?thesis unfolding kle_def apply(rule_tac x="kk1 \<union> kk2" in exI) apply(rule) defer proof
+    fix i show "c i = a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" proof(cases "i\<in>kk1 \<union> kk2")
+      case True hence "c i \<ge> a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)"
+	unfolding kk1[THEN conjunct2,rule_format,of i] kk2[THEN conjunct2,rule_format,of i] by auto
+      moreover have "c i \<le> a i + (if i \<in> kk1 \<union> kk2 then 1 else 0)" using True assms(3) by auto  
+      ultimately show ?thesis by auto next
+      case False thus ?thesis using kk1 kk2 by auto qed qed(insert kk1 kk2, auto) qed
+
+lemma kle_between_r: assumes "kle n a b" "kle n b c" "kle n a x" "kle n c x" shows "kle n b x"
+  apply(rule kle_trans_2[OF assms(2,4)]) proof have *:"\<And>c b x::nat. x \<le> c + 1 \<Longrightarrow> c \<le> b \<Longrightarrow> x \<le> b + 1" by auto
+  fix j show "x j \<le> b j + 1" apply(rule *)using kle_trans_1[OF assms(1),of j] kle_trans_1[OF assms(3), of j] by auto qed
+
+lemma kle_between_l: assumes "kle n a b" "kle n b c" "kle n x a" "kle n x c" shows "kle n x b"
+  apply(rule kle_trans_2[OF assms(3,1)]) proof have *:"\<And>c b x::nat. c \<le> x + 1 \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> x + 1" by auto
+  fix j show "b j \<le> x j + 1" apply(rule *) using kle_trans_1[OF assms(2),of j] kle_trans_1[OF assms(4), of j] by auto qed
+
+lemma kle_adjacent:
+  assumes "\<forall>j. b j = (if j = k then a(j) + 1 else a j)" "kle n a x" "kle n x b"
+  shows "(x = a) \<or> (x = b)" proof(cases "x k = a k")
+  case True show ?thesis apply(rule disjI1,rule ext) proof- fix j
+    have "x j \<le> a j" using kle_imp_pointwise[OF assms(3),THEN spec[where x=j]] 
+      unfolding assms(1)[rule_format] apply-apply(cases "j=k") using True by auto
+    thus "x j = a j" using kle_imp_pointwise[OF assms(2),THEN spec[where x=j]] by auto qed next
+  case False show ?thesis apply(rule disjI2,rule ext) proof- fix j
+    have "x j \<ge> b j" using kle_imp_pointwise[OF assms(2),THEN spec[where x=j]]
+      unfolding assms(1)[rule_format] apply-apply(cases "j=k") using False by auto
+    thus "x j = b j" using kle_imp_pointwise[OF assms(3),THEN spec[where x=j]] by auto qed qed
+
+subsection {* kuhn's notion of a simplex (a reformulation to avoid so much indexing). *}
+
+definition "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow>
+        (card s = n + 1 \<and>
+        (\<forall>x\<in>s. \<forall>j. x(j) \<le> p) \<and>
+        (\<forall>x\<in>s. \<forall>j. j\<notin>{1..n} \<longrightarrow> (x j = p)) \<and>
+        (\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x))"
+
+lemma ksimplexI:"card s = n + 1 \<Longrightarrow>  \<forall>x\<in>s. \<forall>j. x j \<le> p \<Longrightarrow> \<forall>x\<in>s. \<forall>j. j \<notin> {1..?n} \<longrightarrow> x j = ?p \<Longrightarrow> \<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x \<Longrightarrow> ksimplex p n s"
+  unfolding ksimplex_def by auto
+
+lemma ksimplex_eq: "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow>
+        (card s = n + 1 \<and> finite s \<and>
+        (\<forall>x\<in>s. \<forall>j. x(j) \<le> p) \<and>
+        (\<forall>x\<in>s. \<forall>j. j\<notin>{1..n} \<longrightarrow> (x j = p)) \<and>
+        (\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x))"
+  unfolding ksimplex_def by (auto intro: card_ge_0_finite)
+
+lemma ksimplex_extrema: assumes "ksimplex p n s" obtains a b where "a \<in> s" "b \<in> s"
+  "\<forall>x\<in>s. kle n a x \<and> kle n x b" "\<forall>i. b(i) = (if i \<in> {1..n} then a(i) + 1 else a(i))" proof(cases "n=0")
+  case True obtain x where *:"s = {x}" using assms[unfolded ksimplex_eq True,THEN conjunct1]
+    unfolding add_0_left card_1_exists by auto
+  show ?thesis apply(rule that[of x x]) unfolding * True by auto next
+  note assm = assms[unfolded ksimplex_eq]
+  case False have "s\<noteq>{}" using assm by auto
+  obtain a where a:"a\<in>s" "\<forall>x\<in>s. kle n a x" using `s\<noteq>{}` assm using kle_minimal[of s n] by auto
+  obtain b where b:"b\<in>s" "\<forall>x\<in>s. kle n x b" using `s\<noteq>{}` assm using kle_maximal[of s n] by auto
+  obtain c d where c_d:"c\<in>s" "d\<in>s" "kle n c d" "n \<le> card {k \<in> {1..n}. c k < d k}"
+    using kle_range_induct[of s n n] using assm by auto
+  have "kle n c b \<and> n \<le> card {k \<in> {1..n}. c k < b k}" apply(rule kle_range_combine_r[where y=d]) using c_d a b by auto
+  hence "kle n a b \<and> n \<le> card {k\<in>{1..n}. a(k) < b(k)}" apply-apply(rule kle_range_combine_l[where y=c]) using a `c\<in>s` `b\<in>s` by auto
+  moreover have "card {1..n} \<ge> card {k\<in>{1..n}. a(k) < b(k)}" apply(rule card_mono) by auto
+  ultimately have *:"{k\<in>{1 .. n}. a k < b k} = {1..n}" apply- apply(rule card_subset_eq) by auto
+  show ?thesis apply(rule that[OF a(1) b(1)]) defer apply(subst *[THEN sym]) unfolding mem_Collect_eq proof
+    guess k using a(2)[rule_format,OF b(1),unfolded kle_def] .. note k=this
+    fix i show "b i = (if i \<in> {1..n} \<and> a i < b i then a i + 1 else a i)" proof(cases "i \<in> {1..n}")
+      case True thus ?thesis unfolding k[THEN conjunct2,rule_format] by auto next
+      case False have "a i = p" using assm and False `a\<in>s` by auto
+      moreover   have "b i = p" using assm and False `b\<in>s` by auto
+      ultimately show ?thesis by auto qed qed(insert a(2) b(2) assm,auto) qed
+
+lemma ksimplex_extrema_strong:
+  assumes "ksimplex p n s" "n \<noteq> 0" obtains a b where "a \<in> s" "b \<in> s" "a \<noteq> b"
+  "\<forall>x\<in>s. kle n a x \<and> kle n x b" "\<forall>i. b(i) = (if i \<in> {1..n} then a(i) + 1 else a(i))" proof-
+  obtain a b where ab:"a \<in> s" "b \<in> s" "\<forall>x\<in>s. kle n a x \<and> kle n x b" "\<forall>i. b(i) = (if i \<in> {1..n} then a(i) + 1 else a(i))" 
+    apply(rule ksimplex_extrema[OF assms(1)]) by auto 
+  have "a \<noteq> b" apply(rule ccontr) unfolding not_not apply(drule cong[of _ _ 1 1]) using ab(4) assms(2) by auto
+  thus ?thesis apply(rule_tac that[of a b]) using ab by auto qed
+
+lemma ksimplexD:
+  assumes "ksimplex p n s"
+  shows "card s = n + 1" "finite s" "card s = n + 1" "\<forall>x\<in>s. \<forall>j. x j \<le> p" "\<forall>x\<in>s. \<forall>j. j \<notin> {1..?n} \<longrightarrow> x j = p"
+  "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" using assms unfolding ksimplex_eq by auto
+
+lemma ksimplex_successor:
+  assumes "ksimplex p n s" "a \<in> s"
+  shows "(\<forall>x\<in>s. kle n x a) \<or> (\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y(j) = (if j = k then a(j) + 1 else a(j)))"
+proof(cases "\<forall>x\<in>s. kle n x a") case True thus ?thesis by auto next note assm = ksimplexD[OF assms(1)]
+  case False then obtain b where b:"b\<in>s" "\<not> kle n b a" "\<forall>x\<in>{x \<in> s. \<not> kle n x a}. kle n b x"
+    using kle_minimal[of "{x\<in>s. \<not> kle n x a}" n] and assm by auto
+  hence  **:"1 \<le> card {k\<in>{1..n}. a k < b k}" apply- apply(rule kle_strict_set) using assm(6) and `a\<in>s` by(auto simp add:kle_refl)
+
+  let ?kle1 = "{x \<in> s. \<not> kle n x a}" have "card ?kle1 > 0" apply(rule ccontr) using assm(2) and False by auto
+  hence sizekle1: "card ?kle1 = Suc (card ?kle1 - 1)" using assm(2) by auto
+  obtain c d where c_d: "c \<in> s" "\<not> kle n c a" "d \<in> s" "\<not> kle n d a" "kle n c d" "card ?kle1 - 1 \<le> card {k \<in> {1..n}. c k < d k}"
+    using kle_range_induct[OF sizekle1, of n] using assm by auto
+
+  let ?kle2 = "{x \<in> s. kle n x a}"
+  have "card ?kle2 > 0" apply(rule ccontr) using assm(6)[rule_format,of a a] and `a\<in>s` and assm(2) by auto
+  hence sizekle2:"card ?kle2 = Suc (card ?kle2 - 1)" using assm(2) by auto
+  obtain e f where e_f: "e \<in> s" "kle n e a" "f \<in> s" "kle n f a" "kle n e f" "card ?kle2 - 1 \<le> card {k \<in> {1..n}. e k < f k}"
+    using kle_range_induct[OF sizekle2, of n] using assm by auto
+
+  have "card {k\<in>{1..n}. a k < b k} = 1" proof(rule ccontr) case goal1
+    hence as:"card {k\<in>{1..n}. a k < b k} \<ge> 2" using ** by auto
+    have *:"finite ?kle2" "finite ?kle1" "?kle2 \<union> ?kle1 = s" "?kle2 \<inter> ?kle1 = {}" using assm(2) by auto
+    have "(card ?kle2 - 1) + 2 + (card ?kle1 - 1) = card ?kle2 + card ?kle1" using sizekle1 sizekle2 by auto
+    also have "\<dots> = n + 1" unfolding card_Un_Int[OF *(1-2)] *(3-) using assm(3) by auto
+    finally have n:"(card ?kle2 - 1) + (2 + (card ?kle1 - 1)) = n + 1" by auto
+    have "kle n e a \<and> card {x \<in> s. kle n x a} - 1 \<le> card {k \<in> {1..n}. e k < a k}"
+      apply(rule kle_range_combine_r[where y=f]) using e_f using `a\<in>s` assm(6) by auto
+    moreover have "kle n b d \<and> card {x \<in> s. \<not> kle n x a} - 1 \<le> card {k \<in> {1..n}. b k < d k}"
+      apply(rule kle_range_combine_l[where y=c]) using c_d using assm(6) and b by auto
+    hence "kle n a d \<and> 2 + (card {x \<in> s. \<not> kle n x a} - 1) \<le> card {k \<in> {1..n}. a k < d k}" apply-
+      apply(rule kle_range_combine[where y=b]) using as and b assm(6) `a\<in>s` `d\<in>s` apply- by blast+
+    ultimately have "kle n e d \<and> (card ?kle2 - 1) + (2 + (card ?kle1 - 1)) \<le> card {k\<in>{1..n}. e k < d k}" apply-
+      apply(rule kle_range_combine[where y=a]) using assm(6)[rule_format,OF `e\<in>s` `d\<in>s`] apply - by blast+ 
+    moreover have "card {k \<in> {1..n}. e k < d k} \<le> card {1..n}" apply(rule card_mono) by auto
+    ultimately show False unfolding n by auto qed
+  then guess k unfolding card_1_exists .. note k=this[unfolded mem_Collect_eq]
+
+  show ?thesis apply(rule disjI2) apply(rule_tac x=b in bexI,rule_tac x=k in bexI) proof
+    fix j::nat have "kle n a b" using b and assm(6)[rule_format, OF `a\<in>s` `b\<in>s`] by auto
+    then guess kk unfolding kle_def .. note kk_raw=this note kk=this[THEN conjunct2,rule_format]
+    have kkk:"k\<in>kk" apply(rule ccontr) using k(1) unfolding kk by auto 
+    show "b j = (if j = k then a j + 1 else a j)" proof(cases "j\<in>kk")
+      case True hence "j=k" apply-apply(rule k(2)[rule_format]) using kk_raw kkk by auto
+      thus ?thesis unfolding kk using kkk by auto next
+      case False hence "j\<noteq>k" using k(2)[rule_format, of j k] using kk_raw kkk by auto
+      thus ?thesis unfolding kk using kkk using False by auto qed qed(insert k(1) `b\<in>s`, auto) qed
+
+lemma ksimplex_predecessor:
+  assumes "ksimplex p n s" "a \<in> s"
+  shows "(\<forall>x\<in>s. kle n a x) \<or> (\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a(j) = (if j = k then y(j) + 1 else y(j)))"
+proof(cases "\<forall>x\<in>s. kle n a x") case True thus ?thesis by auto next note assm = ksimplexD[OF assms(1)]
+  case False then obtain b where b:"b\<in>s" "\<not> kle n a b" "\<forall>x\<in>{x \<in> s. \<not> kle n a x}. kle n x b" 
+    using kle_maximal[of "{x\<in>s. \<not> kle n a x}" n] and assm by auto
+  hence  **:"1 \<le> card {k\<in>{1..n}. a k > b k}" apply- apply(rule kle_strict_set) using assm(6) and `a\<in>s` by(auto simp add:kle_refl)
+
+  let ?kle1 = "{x \<in> s. \<not> kle n a x}" have "card ?kle1 > 0" apply(rule ccontr) using assm(2) and False by auto
+  hence sizekle1:"card ?kle1 = Suc (card ?kle1 - 1)" using assm(2) by auto
+  obtain c d where c_d: "c \<in> s" "\<not> kle n a c" "d \<in> s" "\<not> kle n a d" "kle n d c" "card ?kle1 - 1 \<le> card {k \<in> {1..n}. c k > d k}"
+    using kle_range_induct[OF sizekle1, of n] using assm by auto
+
+  let ?kle2 = "{x \<in> s. kle n a x}"
+  have "card ?kle2 > 0" apply(rule ccontr) using assm(6)[rule_format,of a a] and `a\<in>s` and assm(2) by auto
+  hence sizekle2:"card ?kle2 = Suc (card ?kle2 - 1)" using assm(2) by auto
+  obtain e f where e_f: "e \<in> s" "kle n a e" "f \<in> s" "kle n a f" "kle n f e" "card ?kle2 - 1 \<le> card {k \<in> {1..n}. e k > f k}"
+    using kle_range_induct[OF sizekle2, of n] using assm by auto
+
+  have "card {k\<in>{1..n}. a k > b k} = 1" proof(rule ccontr) case goal1
+    hence as:"card {k\<in>{1..n}. a k > b k} \<ge> 2" using ** by auto
+    have *:"finite ?kle2" "finite ?kle1" "?kle2 \<union> ?kle1 = s" "?kle2 \<inter> ?kle1 = {}" using assm(2) by auto
+    have "(card ?kle2 - 1) + 2 + (card ?kle1 - 1) = card ?kle2 + card ?kle1" using sizekle1 sizekle2 by auto
+    also have "\<dots> = n + 1" unfolding card_Un_Int[OF *(1-2)] *(3-) using assm(3) by auto
+    finally have n:"(card ?kle1 - 1) + 2 + (card ?kle2 - 1) = n + 1" by auto
+    have "kle n a e \<and> card {x \<in> s. kle n a x} - 1 \<le> card {k \<in> {1..n}. e k > a k}"
+      apply(rule kle_range_combine_l[where y=f]) using e_f using `a\<in>s` assm(6) by auto
+    moreover have "kle n d b \<and> card {x \<in> s. \<not> kle n a x} - 1 \<le> card {k \<in> {1..n}. b k > d k}"
+      apply(rule kle_range_combine_r[where y=c]) using c_d using assm(6) and b by auto
+    hence "kle n d a \<and> (card {x \<in> s. \<not> kle n a x} - 1) + 2 \<le> card {k \<in> {1..n}. a k > d k}" apply-
+      apply(rule kle_range_combine[where y=b]) using as and b assm(6) `a\<in>s` `d\<in>s` by blast+
+    ultimately have "kle n d e \<and> (card ?kle1 - 1 + 2) + (card ?kle2 - 1) \<le> card {k\<in>{1..n}. e k > d k}" apply-
+      apply(rule kle_range_combine[where y=a]) using assm(6)[rule_format,OF `e\<in>s` `d\<in>s`] apply - by blast+
+    moreover have "card {k \<in> {1..n}. e k > d k} \<le> card {1..n}" apply(rule card_mono) by auto
+    ultimately show False unfolding n by auto qed
+  then guess k unfolding card_1_exists .. note k=this[unfolded mem_Collect_eq]
+
+  show ?thesis apply(rule disjI2) apply(rule_tac x=b in bexI,rule_tac x=k in bexI) proof
+    fix j::nat have "kle n b a" using b and assm(6)[rule_format, OF `a\<in>s` `b\<in>s`] by auto
+    then guess kk unfolding kle_def .. note kk_raw=this note kk=this[THEN conjunct2,rule_format]
+    have kkk:"k\<in>kk" apply(rule ccontr) using k(1) unfolding kk by auto 
+    show "a j = (if j = k then b j + 1 else b j)" proof(cases "j\<in>kk")
+      case True hence "j=k" apply-apply(rule k(2)[rule_format]) using kk_raw kkk by auto
+      thus ?thesis unfolding kk using kkk by auto next
+      case False hence "j\<noteq>k" using k(2)[rule_format, of j k] using kk_raw kkk by auto
+      thus ?thesis unfolding kk using kkk using False by auto qed qed(insert k(1) `b\<in>s`, auto) qed
+
+subsection {* The lemmas about simplices that we need. *}
+
+lemma card_funspace': assumes "finite s" "finite t" "card s = m" "card t = n"
+  shows "card {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)} = n ^ m" (is "card (?M s) = _")
+  using assms apply - proof(induct m arbitrary: s)
+  have *:"{f. \<forall>x. f x = d} = {\<lambda>x. d}" apply(rule set_ext,rule)unfolding mem_Collect_eq apply(rule,rule ext) by auto
+  case 0 thus ?case by(auto simp add: *) next
+  case (Suc m) guess a using card_eq_SucD[OF Suc(4)] .. then guess s0
+    apply(erule_tac exE) apply(erule conjE)+ . note as0 = this
+  have **:"card s0 = m" using as0 using Suc(2) Suc(4) by auto
+  let ?l = "(\<lambda>(b,g) x. if x = a then b else g x)" have *:"?M (insert a s0) = ?l ` {(b,g). b\<in>t \<and> g\<in>?M s0}"
+    apply(rule set_ext,rule) unfolding mem_Collect_eq image_iff apply(erule conjE)
+    apply(rule_tac x="(x a, \<lambda>y. if y\<in>s0 then x y else d)" in bexI) apply(rule ext) prefer 3 apply rule defer
+    apply(erule bexE,rule) unfolding mem_Collect_eq apply(erule splitE)+ apply(erule conjE)+ proof-
+    fix x xa xb xc y assume as:"x = (\<lambda>(b, g) x. if x = a then b else g x) xa" "xb \<in> UNIV - insert a s0" "xa = (xc, y)" "xc \<in> t"
+      "\<forall>x\<in>s0. y x \<in> t" "\<forall>x\<in>UNIV - s0. y x = d" thus "x xb = d" unfolding as by auto qed auto
+  have inj:"inj_on ?l {(b,g). b\<in>t \<and> g\<in>?M s0}" unfolding inj_on_def apply(rule,rule,rule) unfolding mem_Collect_eq apply(erule splitE conjE)+ proof-
+    case goal1 note as = this(1,4-)[unfolded goal1 split_conv]
+    have "xa = xb" using as(1)[THEN cong[of _ _ a]] by auto
+    moreover have "ya = yb" proof(rule ext) fix x show "ya x = yb x" proof(cases "x = a") 
+	case False thus ?thesis using as(1)[THEN cong[of _ _ x x]] by auto next
+	case True thus ?thesis using as(5,7) using as0(2) by auto qed qed 
+    ultimately show ?case unfolding goal1 by auto qed
+  have "finite s0" using `finite s` unfolding as0 by simp
+  show ?case unfolding as0 * card_image[OF inj] using assms
+    unfolding SetCompr_Sigma_eq apply-
+    unfolding card_cartesian_product
+    using Suc(1)[OF `finite s0` `finite t` ** `card t = n`] by auto
+qed
+
+lemma card_funspace: assumes  "finite s" "finite t"
+  shows "card {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)} = (card t) ^ (card s)"
+  using assms by (auto intro: card_funspace')
+
+lemma finite_funspace: assumes "finite s" "finite t"
+  shows "finite {f. (\<forall>x\<in>s. f x \<in> t) \<and> (\<forall>x\<in>UNIV - s. f x = d)}" (is "finite ?S")
+proof (cases "card t > 0")
+  case True
+  have "card ?S = (card t) ^ (card s)"
+    using assms by (auto intro!: card_funspace)
+  thus ?thesis using True by (auto intro: card_ge_0_finite)
+next
+  case False hence "t = {}" using `finite t` by auto
+  show ?thesis
+  proof (cases "s = {}")
+    have *:"{f. \<forall>x. f x = d} = {\<lambda>x. d}" by (auto intro: ext)
+    case True thus ?thesis using `t = {}` by (auto simp: *)
+  next
+    case False thus ?thesis using `t = {}` by simp
+  qed
+qed
+
+lemma finite_simplices: "finite {s. ksimplex p n s}"
+  apply(rule finite_subset[of _ "{s. s\<subseteq>{f. (\<forall>i\<in>{1..n}. f i \<in> {0..p}) \<and> (\<forall>i\<in>UNIV-{1..n}. f i = p)}}"])
+  unfolding ksimplex_def defer apply(rule finite_Collect_subsets) apply(rule finite_funspace) by auto
+
+lemma simplex_top_face: assumes "0<p" "\<forall>x\<in>f. x (n + 1) = p"
+  shows "(\<exists>s a. ksimplex p (n + 1) s \<and> a \<in> s \<and> (f = s - {a})) \<longleftrightarrow> ksimplex p n f" (is "?ls = ?rs") proof
+  assume ?ls then guess s .. then guess a apply-apply(erule exE,(erule conjE)+) . note sa=this
+  show ?rs unfolding ksimplex_def sa(3) apply(rule) defer apply rule defer apply(rule,rule,rule,rule) defer apply(rule,rule) proof-
+    fix x y assume as:"x \<in>s - {a}" "y \<in>s - {a}" have xyp:"x (n + 1) = y (n + 1)"
+	using as(1)[unfolded sa(3)[THEN sym], THEN assms(2)[rule_format]]
+	using as(2)[unfolded sa(3)[THEN sym], THEN assms(2)[rule_format]] by auto
+    show "kle n x y \<or> kle n y x" proof(cases "kle (n + 1) x y")
+      case True then guess k unfolding kle_def .. note k=this hence *:"n+1 \<notin> k" using xyp by auto
+      have "\<not> (\<exists>x\<in>k. x\<notin>{1..n})" apply(rule ccontr) unfolding not_not apply(erule bexE) proof-
+	fix x assume as:"x \<in> k" "x \<notin> {1..n}" have "x \<noteq> n+1" using as and * by auto
+	thus False using as and k[THEN conjunct1,unfolded subset_eq] by auto qed
+      thus ?thesis apply-apply(rule disjI1) unfolding kle_def using k apply(rule_tac x=k in exI) by auto next
+      case False hence "kle (n + 1) y x" using ksimplexD(6)[OF sa(1),rule_format, of x y] using as by auto
+      then guess k unfolding kle_def .. note k=this hence *:"n+1 \<notin> k" using xyp by auto
+      hence "\<not> (\<exists>x\<in>k. x\<notin>{1..n})" apply-apply(rule ccontr) unfolding not_not apply(erule bexE) proof-
+	fix x assume as:"x \<in> k" "x \<notin> {1..n}" have "x \<noteq> n+1" using as and * by auto
+	thus False using as and k[THEN conjunct1,unfolded subset_eq] by auto qed
+      thus ?thesis apply-apply(rule disjI2) unfolding kle_def using k apply(rule_tac x=k in exI) by auto qed next
+    fix x j assume as:"x\<in>s - {a}" "j\<notin>{1..n}"
+    thus "x j = p" using as(1)[unfolded sa(3)[THEN sym], THEN assms(2)[rule_format]]
+      apply(cases "j = n+1") using sa(1)[unfolded ksimplex_def] by auto qed(insert sa ksimplexD[OF sa(1)], auto) next
+  assume ?rs note rs=ksimplexD[OF this] guess a b apply(rule ksimplex_extrema[OF `?rs`]) . note ab = this
+  def c \<equiv> "\<lambda>i. if i = (n + 1) then p - 1 else a i"
+  have "c\<notin>f" apply(rule ccontr) unfolding not_not apply(drule assms(2)[rule_format]) unfolding c_def using assms(1) by auto
+  thus ?ls apply(rule_tac x="insert c f" in exI,rule_tac x=c in exI) unfolding ksimplex_def conj_assoc
+    apply(rule conjI) defer apply(rule conjI) defer apply(rule conjI) defer apply(rule conjI) defer  
+  proof(rule_tac[3-5] ballI allI)+
+    fix x j assume x:"x \<in> insert c f" thus "x j \<le> p" proof (cases "x=c")
+      case True show ?thesis unfolding True c_def apply(cases "j=n+1") using ab(1) and rs(4) by auto 
+    qed(insert x rs(4), auto simp add:c_def)
+    show "j \<notin> {1..n + 1} \<longrightarrow> x j = p" apply(cases "x=c") using x ab(1) rs(5) unfolding c_def by auto
+    { fix z assume z:"z \<in> insert c f" hence "kle (n + 1) c z" apply(cases "z = c") (*defer apply(rule kle_Suc)*) proof-
+	case False hence "z\<in>f" using z by auto
+	then guess k apply(drule_tac ab(3)[THEN bspec[where x=z], THEN conjunct1]) unfolding kle_def apply(erule exE) .
+	thus "kle (n + 1) c z" unfolding kle_def apply(rule_tac x="insert (n + 1) k" in exI) unfolding c_def
+	  using ab using rs(5)[rule_format,OF ab(1),of "n + 1"] assms(1) by auto qed auto } note * = this
+    fix y assume y:"y \<in> insert c f" show "kle (n + 1) x y \<or> kle (n + 1) y x" proof(cases "x = c \<or> y = c")
+      case False hence **:"x\<in>f" "y\<in>f" using x y by auto
+      show ?thesis using rs(6)[rule_format,OF **] by(auto dest: kle_Suc) qed(insert * x y, auto)
+  qed(insert rs, auto) qed
+
+lemma ksimplex_fix_plane:
+  assumes "a \<in> s" "j\<in>{1..n::nat}" "\<forall>x\<in>s - {a}. x j = q" "a0 \<in> s" "a1 \<in> s"
+  "\<forall>i. a1 i = ((if i\<in>{1..n} then a0 i + 1 else a0 i)::nat)"
+  shows "(a = a0) \<or> (a = a1)" proof- have *:"\<And>P A x y. \<forall>x\<in>A. P x \<Longrightarrow> x\<in>A \<Longrightarrow> y\<in>A \<Longrightarrow> P x \<and> P y" by auto
+  show ?thesis apply(rule ccontr) using *[OF assms(3), of a0 a1] unfolding assms(6)[THEN spec[where x=j]]
+    using assms(1-2,4-5) by auto qed
+
+lemma ksimplex_fix_plane_0:
+  assumes "a \<in> s" "j\<in>{1..n::nat}" "\<forall>x\<in>s - {a}. x j = 0" "a0 \<in> s" "a1 \<in> s"
+  "\<forall>i. a1 i = ((if i\<in>{1..n} then a0 i + 1 else a0 i)::nat)"
+  shows "a = a1" apply(rule ccontr) using ksimplex_fix_plane[OF assms]
+  using assms(3)[THEN bspec[where x=a1]] using assms(2,5)  
+  unfolding assms(6)[THEN spec[where x=j]] by simp
+
+lemma ksimplex_fix_plane_p:
+  assumes "ksimplex p n s" "a \<in> s" "j\<in>{1..n}" "\<forall>x\<in>s - {a}. x j = p" "a0 \<in> s" "a1 \<in> s"
+  "\<forall>i. a1 i = (if i\<in>{1..n} then a0 i + 1 else a0 i)"
+  shows "a = a0" proof(rule ccontr) note s = ksimplexD[OF assms(1),rule_format]
+  assume as:"a \<noteq> a0" hence *:"a0 \<in> s - {a}" using assms(5) by auto
+  hence "a1 = a" using ksimplex_fix_plane[OF assms(2-)] by auto
+  thus False using as using assms(3,5) and assms(7)[rule_format,of j]
+    unfolding assms(4)[rule_format,OF *] using s(4)[OF assms(6), of j] by auto qed
+
+lemma ksimplex_replace_0:
+  assumes "ksimplex p n s" "a \<in> s" "n \<noteq> 0" "j\<in>{1..n}" "\<forall>x\<in>s - {a}. x j = 0"
+  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1" proof-
+  have *:"\<And>s' a a'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> (s' = s)" by auto
+  have **:"\<And>s' a'. ksimplex p n s' \<Longrightarrow> a' \<in> s' \<Longrightarrow> s' - {a'} = s - {a} \<Longrightarrow> s' = s" proof- case goal1
+    guess a0 a1 apply(rule ksimplex_extrema_strong[OF assms(1,3)]) . note exta = this[rule_format]
+    have a:"a = a1" apply(rule ksimplex_fix_plane_0[OF assms(2,4-5)]) using exta(1-2,5) by auto moreover
+    guess b0 b1 apply(rule ksimplex_extrema_strong[OF goal1(1) assms(3)]) . note extb = this[rule_format]
+    have a':"a' = b1" apply(rule ksimplex_fix_plane_0[OF goal1(2) assms(4), of b0]) unfolding goal1(3) using assms extb goal1 by auto moreover
+    have "b0 = a0" unfolding kle_antisym[THEN sym, of b0 a0 n] using exta extb using goal1(3) unfolding a a' by blast
+    hence "b1 = a1" apply-apply(rule ext) unfolding exta(5) extb(5) by auto ultimately
+    show "s' = s" apply-apply(rule *[of _ a1 b1]) using exta(1-2) extb(1-2) goal1 by auto qed
+  show ?thesis unfolding card_1_exists apply-apply(rule ex1I[of _ s])
+    unfolding mem_Collect_eq defer apply(erule conjE bexE)+ apply(rule_tac a'=b in **) using assms(1,2) by auto qed
+
+lemma ksimplex_replace_1:
+  assumes "ksimplex p n s" "a \<in> s" "n \<noteq> 0" "j\<in>{1..n}" "\<forall>x\<in>s - {a}. x j = p"
+  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 1" proof-
+  have lem:"\<And>a a' s'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> s' = s" by auto
+  have lem:"\<And>s' a'. ksimplex p n s' \<Longrightarrow> a'\<in>s' \<Longrightarrow> s' - {a'} = s - {a} \<Longrightarrow> s' = s" proof- case goal1
+    guess a0 a1 apply(rule ksimplex_extrema_strong[OF assms(1,3)]) . note exta = this[rule_format]
+    have a:"a = a0" apply(rule ksimplex_fix_plane_p[OF assms(1-2,4-5) exta(1,2)]) unfolding exta by auto moreover
+    guess b0 b1 apply(rule ksimplex_extrema_strong[OF goal1(1) assms(3)]) . note extb = this[rule_format]
+    have a':"a' = b0" apply(rule ksimplex_fix_plane_p[OF goal1(1-2) assms(4), of _ b1]) unfolding goal1 extb using extb(1,2) assms(5) by auto
+    moreover have *:"b1 = a1" unfolding kle_antisym[THEN sym, of b1 a1 n] using exta extb using goal1(3) unfolding a a' by blast moreover
+    have "a0 = b0" apply(rule ext) proof- case goal1 show "a0 x = b0 x"
+	using *[THEN cong, of x x] unfolding exta extb apply-apply(cases "x\<in>{1..n}") by auto qed
+    ultimately show "s' = s" apply-apply(rule lem[OF goal1(3) _ goal1(2) assms(2)]) by auto qed 
+  show ?thesis unfolding card_1_exists apply(rule ex1I[of _ s]) unfolding mem_Collect_eq apply(rule,rule assms(1))
+    apply(rule_tac x=a in bexI) prefer 3 apply(erule conjE bexE)+ apply(rule_tac a'=b in lem) using assms(1-2) by auto qed
+
+lemma ksimplex_replace_2:
+  assumes "ksimplex p n s" "a \<in> s" "n \<noteq> 0" "~(\<exists>j\<in>{1..n}. \<forall>x\<in>s - {a}. x j = 0)" "~(\<exists>j\<in>{1..n}. \<forall>x\<in>s - {a}. x j = p)"
+  shows "card {s'. ksimplex p n s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})} = 2" (is "card ?A = 2")  proof-
+  have lem1:"\<And>a a' s s'. s' - {a'} = s - {a} \<Longrightarrow> a' = a \<Longrightarrow> a' \<in> s' \<Longrightarrow> a \<in> s \<Longrightarrow> s' = s" by auto
+  have lem2:"\<And>a b. a\<in>s \<Longrightarrow> b\<noteq>a \<Longrightarrow> s \<noteq> insert b (s - {a})" proof case goal1
+    hence "a\<in>insert b (s - {a})" by auto hence "a\<in> s - {a}" unfolding insert_iff using goal1 by auto
+    thus False by auto qed
+  guess a0 a1 apply(rule ksimplex_extrema_strong[OF assms(1,3)]) . note a0a1=this
+  { assume "a=a0"
+    have *:"\<And>P Q. (P \<or> Q) \<Longrightarrow> \<not> P \<Longrightarrow> Q" by auto
+    have "\<exists>x\<in>s. \<not> kle n x a0" apply(rule_tac x=a1 in bexI) proof assume as:"kle n a1 a0"
+      show False using kle_imp_pointwise[OF as,THEN spec[where x=1]] unfolding a0a1(5)[THEN spec[where x=1]]
+        using assms(3) by auto qed(insert a0a1,auto)
+    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a0 j + 1 else a0 j)"
+      apply(rule_tac *[OF ksimplex_successor[OF assms(1-2),unfolded `a=a0`]]) by auto
+    then guess a2 .. from this(2) guess k .. note k=this note a2=`a2\<in>s`
+    def a3 \<equiv> "\<lambda>j. if j = k then a1 j + 1 else a1 j"
+    have "a3 \<notin> s" proof assume "a3\<in>s" hence "kle n a3 a1" using a0a1(4) by auto
+      thus False apply(drule_tac kle_imp_pointwise) unfolding a3_def
+        apply(erule_tac x=k in allE) by auto qed
+    hence "a3 \<noteq> a0" "a3 \<noteq> a1" using a0a1 by auto
+    have "a2 \<noteq> a0" using k(2)[THEN spec[where x=k]] by auto
+    have lem3:"\<And>x. x\<in>(s - {a0}) \<Longrightarrow> kle n a2 x" proof(rule ccontr) case goal1 hence as:"x\<in>s" "x\<noteq>a0" by auto
+      have "kle n a2 x \<or> kle n x a2" using ksimplexD(6)[OF assms(1)] and as `a2\<in>s` by auto moreover
+      have "kle n a0 x" using a0a1(4) as by auto
+      ultimately have "x = a0 \<or> x = a2" apply-apply(rule kle_adjacent[OF k(2)]) using goal1(2) by auto
+      hence "x = a2" using as by auto thus False using goal1(2) using kle_refl by auto qed
+    let ?s = "insert a3 (s - {a0})" have "ksimplex p n ?s" apply(rule ksimplexI) proof(rule_tac[2-] ballI,rule_tac[4] ballI)
+      show "card ?s = n + 1" using ksimplexD(2-3)[OF assms(1)]
+        using `a3\<noteq>a0` `a3\<notin>s` `a0\<in>s` by(auto simp add:card_insert_if)
+      fix x assume x:"x \<in> insert a3 (s - {a0})"
+      show "\<forall>j. x j \<le> p" proof(rule,cases "x = a3")
+	fix j case False thus "x j\<le>p" using x ksimplexD(4)[OF assms(1)] by auto next
+	fix j case True show "x j\<le>p" unfolding True proof(cases "j=k") 
+	  case False thus "a3 j \<le>p" unfolding True a3_def using `a1\<in>s` ksimplexD(4)[OF assms(1)] by auto next
+	  guess a4 using assms(5)[unfolded bex_simps ball_simps,rule_format,OF k(1)] .. note a4=this
+	  have "a2 k \<le> a4 k" using lem3[OF a4(1)[unfolded `a=a0`],THEN kle_imp_pointwise] by auto
+	  also have "\<dots> < p" using ksimplexD(4)[OF assms(1),rule_format,of a4 k] using a4 by auto
+	  finally have *:"a0 k + 1 < p" unfolding k(2)[rule_format] by auto
+	  case True thus "a3 j \<le>p" unfolding a3_def unfolding a0a1(5)[rule_format]
+	    using k(1) k(2)assms(5) using * by simp qed qed
+      show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p" proof(rule,rule,cases "x=a3") fix j::nat assume j:"j\<notin>{1..n}"
+	{ case False thus "x j = p" using j x ksimplexD(5)[OF assms(1)] by auto }
+	case True show "x j = p" unfolding True a3_def using j k(1) 
+	  using ksimplexD(5)[OF assms(1),rule_format,OF `a1\<in>s` j] by auto qed
+      fix y assume y:"y\<in>insert a3 (s - {a0})"
+      have lem4:"\<And>x. x\<in>s \<Longrightarrow> x\<noteq>a0 \<Longrightarrow> kle n x a3" proof- case goal1
+	guess kk using a0a1(4)[rule_format,OF `x\<in>s`,THEN conjunct2,unfolded kle_def] 
+          apply-apply(erule exE,erule conjE) . note kk=this
+	have "k\<notin>kk" proof assume "k\<in>kk"
+	  hence "a1 k = x k + 1" using kk by auto
+	  hence "a0 k = x k" unfolding a0a1(5)[rule_format] using k(1) by auto
+	  hence "a2 k = x k + 1" unfolding k(2)[rule_format] by auto moreover
+	  have "a2 k \<le> x k" using lem3[of x,THEN kle_imp_pointwise] goal1 by auto 
+	  ultimately show False by auto qed
+	thus ?case unfolding kle_def apply(rule_tac x="insert k kk" in exI) using kk(1)
+	  unfolding a3_def kle_def kk(2)[rule_format] using k(1) by auto qed
+      show "kle n x y \<or> kle n y x" proof(cases "y=a3")
+	case True show ?thesis unfolding True apply(cases "x=a3") defer apply(rule disjI1,rule lem4)
+	  using x by auto next
+	case False show ?thesis proof(cases "x=a3") case True show ?thesis unfolding True
+	    apply(rule disjI2,rule lem4) using y False by auto next
+	  case False thus ?thesis apply(rule_tac ksimplexD(6)[OF assms(1),rule_format]) 
+	    using x y `y\<noteq>a3` by auto qed qed qed
+    hence "insert a3 (s - {a0}) \<in> ?A" unfolding mem_Collect_eq apply-apply(rule,assumption)
+      apply(rule_tac x="a3" in bexI) unfolding `a=a0` using `a3\<notin>s` by auto moreover
+    have "s \<in> ?A" using assms(1,2) by auto ultimately have  "?A \<supseteq> {s, insert a3 (s - {a0})}" by auto
+    moreover have "?A \<subseteq> {s, insert a3 (s - {a0})}" apply(rule) unfolding mem_Collect_eq proof(erule conjE)
+      fix s' assume as:"ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
+      from this(2) guess a' .. note a'=this
+      guess a_min a_max apply(rule ksimplex_extrema_strong[OF as assms(3)]) . note min_max=this
+      have *:"\<forall>x\<in>s' - {a'}. x k = a2 k" unfolding a' proof fix x assume x:"x\<in>s-{a}"
+	hence "kle n a2 x" apply-apply(rule lem3) using `a=a0` by auto
+	hence "a2 k \<le> x k" apply(drule_tac kle_imp_pointwise) by auto moreover
+	have "x k \<le> a2 k" unfolding k(2)[rule_format] using a0a1(4)[rule_format,of x,THEN conjunct1] 
+	  unfolding kle_def using x by auto ultimately show "x k = a2 k" by auto qed
+      have **:"a'=a_min \<or> a'=a_max" apply(rule ksimplex_fix_plane[OF a'(1) k(1) *]) using min_max by auto
+      show "s' \<in> {s, insert a3 (s - {a0})}" proof(cases "a'=a_min")
+	case True have "a_max = a1" unfolding kle_antisym[THEN sym,of a_max a1 n] apply(rule)
+	  apply(rule a0a1(4)[rule_format,THEN conjunct2]) defer  proof(rule min_max(4)[rule_format,THEN conjunct2])
+	  show "a1\<in>s'" using a' unfolding `a=a0` using a0a1 by auto
+	  show "a_max \<in> s" proof(rule ccontr) assume "a_max\<notin>s"
+	    hence "a_max = a'" using a' min_max by auto
+	    thus False unfolding True using min_max by auto qed qed
+	hence "\<forall>i. a_max i = a1 i" by auto
+	hence "a' = a" unfolding True `a=a0` apply-apply(subst expand_fun_eq,rule)
+	  apply(erule_tac x=x in allE) unfolding a0a1(5)[rule_format] min_max(5)[rule_format]
+	proof- case goal1 thus ?case apply(cases "x\<in>{1..n}") by auto qed
+	hence "s' = s" apply-apply(rule lem1[OF a'(2)]) using `a\<in>s` `a'\<in>s'` by auto
+	thus ?thesis by auto next
+	case False hence as:"a' = a_max" using ** by auto
+	have "a_min = a2" unfolding kle_antisym[THEN sym, of _ _ n] apply rule
+	  apply(rule min_max(4)[rule_format,THEN conjunct1]) defer proof(rule lem3)
+	  show "a_min \<in> s - {a0}" unfolding a'(2)[THEN sym,unfolded `a=a0`] 
+	    unfolding as using min_max(1-3) by auto
+	  have "a2 \<noteq> a" unfolding `a=a0` using k(2)[rule_format,of k] by auto
+	  hence "a2 \<in> s - {a}" using a2 by auto thus "a2 \<in> s'" unfolding a'(2)[THEN sym] by auto qed
+	hence "\<forall>i. a_min i = a2 i" by auto
+	hence "a' = a3" unfolding as `a=a0` apply-apply(subst expand_fun_eq,rule)
+	  apply(erule_tac x=x in allE) unfolding a0a1(5)[rule_format] min_max(5)[rule_format]
+	  unfolding a3_def k(2)[rule_format] unfolding a0a1(5)[rule_format] proof- case goal1
+	  show ?case unfolding goal1 apply(cases "x\<in>{1..n}") defer apply(cases "x=k")
+	    using `k\<in>{1..n}` by auto qed
+	hence "s' = insert a3 (s - {a0})" apply-apply(rule lem1) defer apply assumption
+	  apply(rule a'(1)) unfolding a' `a=a0` using `a3\<notin>s` by auto
+	thus ?thesis by auto qed qed
+    ultimately have *:"?A = {s, insert a3 (s - {a0})}" by blast
+    have "s \<noteq> insert a3 (s - {a0})" using `a3\<notin>s` by auto
+    hence ?thesis unfolding * by auto } moreover
+  { assume "a=a1"
+    have *:"\<And>P Q. (P \<or> Q) \<Longrightarrow> \<not> P \<Longrightarrow> Q" by auto
+    have "\<exists>x\<in>s. \<not> kle n a1 x" apply(rule_tac x=a0 in bexI) proof assume as:"kle n a1 a0"
+      show False using kle_imp_pointwise[OF as,THEN spec[where x=1]] unfolding a0a1(5)[THEN spec[where x=1]]
+        using assms(3) by auto qed(insert a0a1,auto)
+    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a1 j = (if j = k then y j + 1 else y j)"
+      apply(rule_tac *[OF ksimplex_predecessor[OF assms(1-2),unfolded `a=a1`]]) by auto
+    then guess a2 .. from this(2) guess k .. note k=this note a2=`a2\<in>s`
+    def a3 \<equiv> "\<lambda>j. if j = k then a0 j - 1 else a0 j"
+    have "a2 \<noteq> a1" using k(2)[THEN spec[where x=k]] by auto
+    have lem3:"\<And>x. x\<in>(s - {a1}) \<Longrightarrow> kle n x a2" proof(rule ccontr) case goal1 hence as:"x\<in>s" "x\<noteq>a1" by auto
+      have "kle n a2 x \<or> kle n x a2" using ksimplexD(6)[OF assms(1)] and as `a2\<in>s` by auto moreover
+      have "kle n x a1" using a0a1(4) as by auto
+      ultimately have "x = a2 \<or> x = a1" apply-apply(rule kle_adjacent[OF k(2)]) using goal1(2) by auto
+      hence "x = a2" using as by auto thus False using goal1(2) using kle_refl by auto qed
+    have "a0 k \<noteq> 0" proof-
+      guess a4 using assms(4)[unfolded bex_simps ball_simps,rule_format,OF `k\<in>{1..n}`] .. note a4=this
+      have "a4 k \<le> a2 k" using lem3[OF a4(1)[unfolded `a=a1`],THEN kle_imp_pointwise] by auto
+      moreover have "a4 k > 0" using a4 by auto ultimately have "a2 k > 0" by auto
+      hence "a1 k > 1" unfolding k(2)[rule_format] by simp
+      thus ?thesis unfolding a0a1(5)[rule_format] using k(1) by simp qed
+    hence lem4:"\<forall>j. a0 j = (if j=k then a3 j + 1 else a3 j)" unfolding a3_def by simp
+    have "\<not> kle n a0 a3" apply(rule ccontr) unfolding not_not apply(drule kle_imp_pointwise)
+      unfolding lem4[rule_format] apply(erule_tac x=k in allE) by auto
+    hence "a3 \<notin> s" using a0a1(4) by auto
+    hence "a3 \<noteq> a1" "a3 \<noteq> a0" using a0a1 by auto
+    let ?s = "insert a3 (s - {a1})" have "ksimplex p n ?s" apply(rule ksimplexI) proof(rule_tac[2-] ballI,rule_tac[4] ballI)
+      show "card ?s = n+1" using ksimplexD(2-3)[OF assms(1)]
+        using `a3\<noteq>a0` `a3\<notin>s` `a1\<in>s` by(auto simp add:card_insert_if)
+      fix x assume x:"x \<in> insert a3 (s - {a1})"
+      show "\<forall>j. x j \<le> p" proof(rule,cases "x = a3")
+	fix j case False thus "x j\<le>p" using x ksimplexD(4)[OF assms(1)] by auto next
+	fix j case True show "x j\<le>p" unfolding True proof(cases "j=k") 
+	  case False thus "a3 j \<le>p" unfolding True a3_def using `a0\<in>s` ksimplexD(4)[OF assms(1)] by auto next
+	  guess a4 using assms(5)[unfolded bex_simps ball_simps,rule_format,OF k(1)] .. note a4=this
+          case True have "a3 k \<le> a0 k" unfolding lem4[rule_format] by auto
+          also have "\<dots> \<le> p" using ksimplexD(4)[OF assms(1),rule_format,of a0 k] a0a1 by auto
+          finally show "a3 j \<le> p" unfolding True by auto qed qed
+      show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p" proof(rule,rule,cases "x=a3") fix j::nat assume j:"j\<notin>{1..n}"
+	{ case False thus "x j = p" using j x ksimplexD(5)[OF assms(1)] by auto }
+	case True show "x j = p" unfolding True a3_def using j k(1) 
+	  using ksimplexD(5)[OF assms(1),rule_format,OF `a0\<in>s` j] by auto qed
+      fix y assume y:"y\<in>insert a3 (s - {a1})"
+      have lem4:"\<And>x. x\<in>s \<Longrightarrow> x\<noteq>a1 \<Longrightarrow> kle n a3 x" proof- case goal1 hence *:"x\<in>s - {a1}" by auto
+        have "kle n a3 a2" proof- have "kle n a0 a1" using a0a1 by auto then guess kk unfolding kle_def ..
+          thus ?thesis unfolding kle_def apply(rule_tac x=kk in exI) unfolding lem4[rule_format] k(2)[rule_format]
+            apply(rule)defer proof(rule) case goal1 thus ?case apply-apply(erule conjE)
+              apply(erule_tac[!] x=j in allE) apply(cases "j\<in>kk") apply(case_tac[!] "j=k") by auto qed auto qed moreover
+        have "kle n a3 a0" unfolding kle_def lem4[rule_format] apply(rule_tac x="{k}" in exI) using k(1) by auto
+        ultimately show ?case apply-apply(rule kle_between_l[of _ a0 _ a2]) using lem3[OF *]
+          using a0a1(4)[rule_format,OF goal1(1)] by auto qed
+      show "kle n x y \<or> kle n y x" proof(cases "y=a3")
+	case True show ?thesis unfolding True apply(cases "x=a3") defer apply(rule disjI2,rule lem4)
+	  using x by auto next
+	case False show ?thesis proof(cases "x=a3") case True show ?thesis unfolding True
+	    apply(rule disjI1,rule lem4) using y False by auto next
+	  case False thus ?thesis apply(rule_tac ksimplexD(6)[OF assms(1),rule_format]) 
+	    using x y `y\<noteq>a3` by auto qed qed qed
+    hence "insert a3 (s - {a1}) \<in> ?A" unfolding mem_Collect_eq apply-apply(rule,assumption)
+      apply(rule_tac x="a3" in bexI) unfolding `a=a1` using `a3\<notin>s` by auto moreover
+    have "s \<in> ?A" using assms(1,2) by auto ultimately have  "?A \<supseteq> {s, insert a3 (s - {a1})}" by auto
+    moreover have "?A \<subseteq> {s, insert a3 (s - {a1})}" apply(rule) unfolding mem_Collect_eq proof(erule conjE)
+      fix s' assume as:"ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
+      from this(2) guess a' .. note a'=this
+      guess a_min a_max apply(rule ksimplex_extrema_strong[OF as assms(3)]) . note min_max=this
+      have *:"\<forall>x\<in>s' - {a'}. x k = a2 k" unfolding a' proof fix x assume x:"x\<in>s-{a}"
+	hence "kle n x a2" apply-apply(rule lem3) using `a=a1` by auto
+	hence "x k \<le> a2 k" apply(drule_tac kle_imp_pointwise) by auto moreover
+	{ have "a2 k \<le> a0 k" using k(2)[rule_format,of k] unfolding a0a1(5)[rule_format] using k(1) by simp
+	  also have "\<dots> \<le> x k" using a0a1(4)[rule_format,of x,THEN conjunct1,THEN kle_imp_pointwise] x by auto
+	  finally have "a2 k \<le> x k" . } ultimately show "x k = a2 k" by auto qed
+      have **:"a'=a_min \<or> a'=a_max" apply(rule ksimplex_fix_plane[OF a'(1) k(1) *]) using min_max by auto
+      have "a2 \<noteq> a1" proof assume as:"a2 = a1"
+	show False using k(2) unfolding as apply(erule_tac x=k in allE) by auto qed
+      hence a2':"a2 \<in> s' - {a'}" unfolding a' using a2 unfolding `a=a1` by auto
+      show "s' \<in> {s, insert a3 (s - {a1})}" proof(cases "a'=a_min")
+	case True have "a_max \<in> s - {a1}" using min_max unfolding a'(2)[unfolded `a=a1`,THEN sym] True by auto
+	hence "a_max = a2" unfolding kle_antisym[THEN sym,of a_max a2 n] apply-apply(rule)
+	  apply(rule lem3,assumption) apply(rule min_max(4)[rule_format,THEN conjunct2]) using a2' by auto
+	hence a_max:"\<forall>i. a_max i = a2 i" by auto
+	have *:"\<forall>j. a2 j = (if j\<in>{1..n} then a3 j + 1 else a3 j)" 
+	  using k(2) unfolding lem4[rule_format] a0a1(5)[rule_format] apply-apply(rule,erule_tac x=j in allE)
+	proof- case goal1 thus ?case apply(cases "j\<in>{1..n}",case_tac[!] "j=k") by auto qed
+	have "\<forall>i. a_min i = a3 i" using a_max apply-apply(rule,erule_tac x=i in allE)
+	  unfolding min_max(5)[rule_format] *[rule_format] proof- case goal1
+	  thus ?case apply(cases "i\<in>{1..n}") by auto qed hence "a_min = a3" unfolding expand_fun_eq .
+	hence "s' = insert a3 (s - {a1})" using a' unfolding `a=a1` True by auto thus ?thesis by auto next
+	case False hence as:"a'=a_max" using ** by auto
+	have "a_min = a0" unfolding kle_antisym[THEN sym,of _ _ n] apply(rule)
+	  apply(rule min_max(4)[rule_format,THEN conjunct1]) defer apply(rule a0a1(4)[rule_format,THEN conjunct1]) proof-
+	  have "a_min \<in> s - {a1}" using min_max(1,3) unfolding a'(2)[THEN sym,unfolded `a=a1`] as by auto
+	  thus "a_min \<in> s" by auto have "a0 \<in> s - {a1}" using a0a1(1-3) by auto thus "a0 \<in> s'"
+	    unfolding a'(2)[THEN sym,unfolded `a=a1`] by auto qed
+	hence "\<forall>i. a_max i = a1 i" unfolding a0a1(5)[rule_format] min_max(5)[rule_format] by auto
+	hence "s' = s" apply-apply(rule lem1[OF a'(2)]) using `a\<in>s` `a'\<in>s'` unfolding as `a=a1` unfolding expand_fun_eq by auto
+	thus ?thesis by auto qed qed 
+    ultimately have *:"?A = {s, insert a3 (s - {a1})}" by blast
+    have "s \<noteq> insert a3 (s - {a1})" using `a3\<notin>s` by auto
+    hence ?thesis unfolding * by auto } moreover
+  { assume as:"a\<noteq>a0" "a\<noteq>a1" have "\<not> (\<forall>x\<in>s. kle n a x)" proof case goal1
+      have "a=a0" unfolding kle_antisym[THEN sym,of _ _ n] apply(rule)
+	using goal1 a0a1 assms(2) by auto thus False using as by auto qed
+    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. a j = (if j = k then y j + 1 else y j)" using  ksimplex_predecessor[OF assms(1-2)] by blast
+    then guess u .. from this(2) guess k .. note k = this[rule_format] note u = `u\<in>s`
+    have "\<not> (\<forall>x\<in>s. kle n x a)" proof case goal1
+      have "a=a1" unfolding kle_antisym[THEN sym,of _ _ n] apply(rule)
+	using goal1 a0a1 assms(2) by auto thus False using as by auto qed
+    hence "\<exists>y\<in>s. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then a j + 1 else a j)" using  ksimplex_successor[OF assms(1-2)] by blast
+    then guess v .. from this(2) guess l .. note l = this[rule_format] note v = `v\<in>s`
+    def a' \<equiv> "\<lambda>j. if j = l then u j + 1 else u j"
+    have kl:"k \<noteq> l" proof assume "k=l" have *:"\<And>P. (if P then (1::nat) else 0) \<noteq> 2" by auto
+      thus False using ksimplexD(6)[OF assms(1),rule_format,OF u v] unfolding kle_def
+	unfolding l(2) k(2) `k=l` apply-apply(erule disjE)apply(erule_tac[!] exE conjE)+
+	apply(erule_tac[!] x=l in allE)+ by(auto simp add: *) qed
+    hence aa':"a'\<noteq>a" apply-apply rule unfolding expand_fun_eq unfolding a'_def k(2)
+      apply(erule_tac x=l in allE) by auto
+    have "a' \<notin> s" apply(rule) apply(drule ksimplexD(6)[OF assms(1),rule_format,OF `a\<in>s`]) proof(cases "kle n a a'")
+      case goal2 hence "kle n a' a" by auto thus False apply(drule_tac kle_imp_pointwise)
+	apply(erule_tac x=l in allE) unfolding a'_def k(2) using kl by auto next
+      case True thus False apply(drule_tac kle_imp_pointwise)
+	apply(erule_tac x=k in allE) unfolding a'_def k(2) using kl by auto qed
+    have kle_uv:"kle n u a" "kle n u a'" "kle n a v" "kle n a' v" unfolding kle_def apply-
+      apply(rule_tac[1] x="{k}" in exI,rule_tac[2] x="{l}" in exI)
+      apply(rule_tac[3] x="{l}" in exI,rule_tac[4] x="{k}" in exI)
+      unfolding l(2) k(2) a'_def using l(1) k(1) by auto
+    have uxv:"\<And>x. kle n u x \<Longrightarrow> kle n x v \<Longrightarrow> (x = u) \<or> (x = a) \<or> (x = a') \<or> (x = v)"
+    proof- case goal1 thus ?case proof(cases "x k = u k", case_tac[!] "x l = u l")
+      assume as:"x l = u l" "x k = u k"
+      have "x = u" unfolding expand_fun_eq
+	using goal1(2)[THEN kle_imp_pointwise,unfolded l(2)] unfolding k(2) apply-
+	using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1
+	thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as by auto qed thus ?case by auto next
+      assume as:"x l \<noteq> u l" "x k = u k"
+      have "x = a'" unfolding expand_fun_eq unfolding a'_def
+	using goal1(2)[THEN kle_imp_pointwise] unfolding l(2) k(2) apply-
+	using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1
+	thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as by auto qed thus ?case by auto next
+      assume as:"x l = u l" "x k \<noteq> u k"
+      have "x = a" unfolding expand_fun_eq
+	using goal1(2)[THEN kle_imp_pointwise] unfolding l(2) k(2) apply-
+	using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1
+	thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as by auto qed thus ?case by auto next
+      assume as:"x l \<noteq> u l" "x k \<noteq> u k"
+      have "x = v" unfolding expand_fun_eq
+	using goal1(2)[THEN kle_imp_pointwise] unfolding l(2) k(2) apply-
+	using goal1(1)[THEN kle_imp_pointwise] apply-apply rule apply(erule_tac x=xa in allE)+ proof- case goal1
+	thus ?case apply(cases "x=l") apply(case_tac[!] "x=k") using as `k\<noteq>l` by auto qed thus ?case by auto qed qed
+    have uv:"kle n u v" apply(rule kle_trans[OF kle_uv(1,3)]) using ksimplexD(6)[OF assms(1)] using u v by auto
+    have lem3:"\<And>x. x\<in>s \<Longrightarrow> kle n v x \<Longrightarrow> kle n a' x" apply(rule kle_between_r[of _ u _ v])
+      prefer 3 apply(rule kle_trans[OF uv]) defer apply(rule ksimplexD(6)[OF assms(1),rule_format])
+      using kle_uv `u\<in>s` by auto
+    have lem4:"\<And>x. x\<in>s \<Longrightarrow> kle n x u \<Longrightarrow> kle n x a'" apply(rule kle_between_l[of _ u _ v])
+      prefer 4 apply(rule kle_trans[OF _ uv]) defer apply(rule ksimplexD(6)[OF assms(1),rule_format])
+      using kle_uv `v\<in>s` by auto
+    have lem5:"\<And>x. x\<in>s \<Longrightarrow> x\<noteq>a \<Longrightarrow> kle n x a' \<or> kle n a' x" proof- case goal1 thus ?case
+      proof(cases "kle n v x \<or> kle n x u") case True thus ?thesis using goal1 by(auto intro:lem3 lem4) next
+	case False hence *:"kle n u x" "kle n x v" using ksimplexD(6)[OF assms(1)] using goal1 `u\<in>s` `v\<in>s` by auto
+	show ?thesis using uxv[OF *] using kle_uv using goal1 by auto qed qed
+    have "ksimplex p n (insert a' (s - {a}))" apply(rule ksimplexI) proof(rule_tac[2-] ballI,rule_tac[4] ballI)
+      show "card (insert a' (s - {a})) = n + 1" using ksimplexD(2-3)[OF assms(1)]
+        using `a'\<noteq>a` `a'\<notin>s` `a\<in>s` by(auto simp add:card_insert_if)
+      fix x assume x:"x \<in> insert a' (s - {a})"
+      show "\<forall>j. x j \<le> p" proof(rule,cases "x = a'")
+	fix j case False thus "x j\<le>p" using x ksimplexD(4)[OF assms(1)] by auto next
+	fix j case True show "x j\<le>p" unfolding True proof(cases "j=l") 
+	  case False thus "a' j \<le>p" unfolding True a'_def using `u\<in>s` ksimplexD(4)[OF assms(1)] by auto next
+	  case True have *:"a l = u l" "v l = a l + 1" using k(2)[of l] l(2)[of l] `k\<noteq>l` by auto
+	  have "u l + 1 \<le> p" unfolding *[THEN sym] using ksimplexD(4)[OF assms(1)] using `v\<in>s` by auto
+	  thus "a' j \<le>p" unfolding a'_def True by auto qed qed
+      show "\<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p" proof(rule,rule,cases "x=a'") fix j::nat assume j:"j\<notin>{1..n}"
+	{ case False thus "x j = p" using j x ksimplexD(5)[OF assms(1)] by auto }
+	case True show "x j = p" unfolding True a'_def using j l(1) 
+	  using ksimplexD(5)[OF assms(1),rule_format,OF `u\<in>s` j] by auto qed
+      fix y assume y:"y\<in>insert a' (s - {a})"
+      show "kle n x y \<or> kle n y x" proof(cases "y=a'")
+	case True show ?thesis unfolding True apply(cases "x=a'") defer apply(rule lem5) using x by auto next
+	case False show ?thesis proof(cases "x=a'") case True show ?thesis unfolding True
+	    using lem5[of y] using y by auto next
+	  case False thus ?thesis apply(rule_tac ksimplexD(6)[OF assms(1),rule_format]) 
+	    using x y `y\<noteq>a'` by auto qed qed qed
+    hence "insert a' (s - {a}) \<in> ?A" unfolding mem_Collect_eq apply-apply(rule,assumption)
+      apply(rule_tac x="a'" in bexI) using aa' `a'\<notin>s` by auto moreover
+    have "s \<in> ?A" using assms(1,2) by auto ultimately have  "?A \<supseteq> {s, insert a' (s - {a})}" by auto
+    moreover have "?A \<subseteq> {s, insert a' (s - {a})}" apply(rule) unfolding mem_Collect_eq proof(erule conjE)
+      fix s' assume as:"ksimplex p n s'" and "\<exists>b\<in>s'. s' - {b} = s - {a}"
+      from this(2) guess a'' .. note a''=this
+      have "u\<noteq>v" unfolding expand_fun_eq unfolding l(2) k(2) by auto
+      hence uv':"\<not> kle n v u" using uv using kle_antisym by auto
+      have "u\<noteq>a" "v\<noteq>a" unfolding expand_fun_eq k(2) l(2) by auto 
+      hence uvs':"u\<in>s'" "v\<in>s'" using `u\<in>s` `v\<in>s` using a'' by auto
+      have lem6:"a \<in> s' \<or> a' \<in> s'" proof(cases "\<forall>x\<in>s'. kle n x u \<or> kle n v x")
+	case False then guess w unfolding ball_simps .. note w=this
+	hence "kle n u w" "kle n w v" using ksimplexD(6)[OF as] uvs' by auto
+	hence "w = a' \<or> w = a" using uxv[of w] uvs' w by auto thus ?thesis using w by auto next
+	case True have "\<not> (\<forall>x\<in>s'. kle n x u)" unfolding ball_simps apply(rule_tac x=v in bexI)
+	  using uv `u\<noteq>v` unfolding kle_antisym[of n u v,THEN sym] using `v\<in>s'` by auto
+	hence "\<exists>y\<in>s'. \<exists>k\<in>{1..n}. \<forall>j. y j = (if j = k then u j + 1 else u j)" using ksimplex_successor[OF as `u\<in>s'`] by blast
+	then guess w .. note w=this from this(2) guess kk .. note kk=this[rule_format]
+	have "\<not> kle n w u" apply-apply(rule,drule kle_imp_pointwise) 
+	  apply(erule_tac x=kk in allE) unfolding kk by auto 
+	hence *:"kle n v w" using True[rule_format,OF w(1)] by auto
+	hence False proof(cases "kk\<noteq>l") case True thus False using *[THEN kle_imp_pointwise, unfolded l(2) kk k(2)]
+	    apply(erule_tac x=l in allE) using `k\<noteq>l` by auto  next
+	  case False hence "kk\<noteq>k" using `k\<noteq>l` by auto thus False using *[THEN kle_imp_pointwise, unfolded l(2) kk k(2)]
+	    apply(erule_tac x=k in allE) using `k\<noteq>l` by auto qed
+	thus ?thesis by auto qed
+      thus "s' \<in> {s, insert a' (s - {a})}" proof(cases "a\<in>s'")
+	case True hence "s' = s" apply-apply(rule lem1[OF a''(2)]) using a'' `a\<in>s` by auto
+	thus ?thesis by auto next case False hence "a'\<in>s'" using lem6 by auto
+	hence "s' = insert a' (s - {a})" apply-apply(rule lem1[of _ a'' _ a'])
+	  unfolding a''(2)[THEN sym] using a'' using `a'\<notin>s` by auto
+	thus ?thesis by auto qed qed 
+    ultimately have *:"?A = {s, insert a' (s - {a})}" by blast
+    have "s \<noteq> insert a' (s - {a})" using `a'\<notin>s` by auto
+    hence ?thesis unfolding * by auto } 
+  ultimately show ?thesis by auto qed
+
+subsection {* Hence another step towards concreteness. *}
+
+lemma kuhn_simplex_lemma:
+  assumes "\<forall>s. ksimplex p (n + 1) s \<longrightarrow> (rl ` s \<subseteq>{0..n+1})"
+  "odd (card{f. \<exists>s a. ksimplex p (n + 1) s \<and> a \<in> s \<and> (f = s - {a}) \<and>
+  (rl ` f = {0 .. n}) \<and> ((\<exists>j\<in>{1..n+1}.\<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}.\<forall>x\<in>f. x j = p))})"
+  shows "odd(card {s\<in>{s. ksimplex p (n + 1) s}. rl ` s = {0..n+1} })" proof-
+  have *:"\<And>x y. x = y \<Longrightarrow> odd (card x) \<Longrightarrow> odd (card y)" by auto
+  have *:"odd(card {f\<in>{f. \<exists>s\<in>{s. ksimplex p (n + 1) s}. (\<exists>a\<in>s. f = s - {a})}. 
+                (rl ` f = {0..n}) \<and>
+               ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or>
+                (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p))})" apply(rule *[OF _ assms(2)]) by auto
+  show ?thesis apply(rule kuhn_complete_lemma[OF finite_simplices]) prefer 6 apply(rule *) apply(rule,rule,rule)
+    apply(subst ksimplex_def) defer apply(rule,rule assms(1)[rule_format]) unfolding mem_Collect_eq apply assumption
+    apply default+ unfolding mem_Collect_eq apply(erule disjE bexE)+ defer apply(erule disjE bexE)+ defer 
+    apply default+ unfolding mem_Collect_eq apply(erule disjE bexE)+ unfolding mem_Collect_eq proof-
+    fix f s a assume as:"ksimplex p (n + 1) s" "a\<in>s" "f = s - {a}"
+    let ?S = "{s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})}"
+    have S:"?S = {s'. ksimplex p (n + 1) s' \<and> (\<exists>b\<in>s'. s' - {b} = s - {a})}" unfolding as by blast
+    { fix j assume j:"j \<in> {1..n + 1}" "\<forall>x\<in>f. x j = 0" thus "card {s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})} = 1" unfolding S
+	apply-apply(rule ksimplex_replace_0) apply(rule as)+ unfolding as by auto }
+    { fix j assume j:"j \<in> {1..n + 1}" "\<forall>x\<in>f. x j = p" thus "card {s. ksimplex p (n + 1) s \<and> (\<exists>a\<in>s. f = s - {a})} = 1" unfolding S
+	apply-apply(rule ksimplex_replace_1) apply(rule as)+ unfolding as by auto }
+    show "\<not> ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p)) \<Longrightarrow> card ?S = 2"
+      unfolding S apply(rule ksimplex_replace_2) apply(rule as)+ unfolding as by auto qed auto qed
+
+subsection {* Reduced labelling. *}
+
+definition "reduced label (n::nat) (x::nat\<Rightarrow>nat) =
+  (SOME k. k \<le> n \<and> (\<forall>i. 1\<le>i \<and> i<k+1 \<longrightarrow> label x i = 0) \<and> (k = n \<or> label x (k + 1) \<noteq> (0::nat)))"
+
+lemma reduced_labelling: shows "reduced label n x \<le> n" (is ?t1) and
+  "\<forall>i. 1\<le>i \<and> i < reduced label n x + 1 \<longrightarrow> (label x i = 0)" (is ?t2)
+  "(reduced label n x = n) \<or> (label x (reduced label n x + 1) \<noteq> 0)"  (is ?t3) proof-
+  have num_WOP:"\<And>P k. P (k::nat) \<Longrightarrow> \<exists>n. P n \<and> (\<forall>m<n. \<not> P m)"
+    apply(drule ex_has_least_nat[where m="\<lambda>x. x"]) apply(erule exE,rule_tac x=x in exI) by auto
+  have *:"n \<le> n \<and> (label x (n + 1) \<noteq> 0 \<or> n = n)" by auto
+  then guess N apply(drule_tac num_WOP[of "\<lambda>j. j\<le>n \<and> (label x (j+1) \<noteq> 0 \<or> n = j)"]) apply(erule exE) . note N=this
+  have N':"N \<le> n" "\<forall>i. 1 \<le> i \<and> i < N + 1 \<longrightarrow> label x i = 0" "N = n \<or> label x (N + 1) \<noteq> 0" defer proof(rule,rule)
+    fix i assume i:"1\<le>i \<and> i<N+1" thus "label x i = 0" using N[THEN conjunct2,THEN spec[where x="i - 1"]] using N by auto qed(insert N, auto)
+  show ?t1 ?t2 ?t3 unfolding reduced_def apply(rule_tac[!] someI2_ex) using N' by(auto intro!: exI[where x=N]) qed
+
+lemma reduced_labelling_unique: fixes x::"nat \<Rightarrow> nat"
+  assumes "r \<le> n"  "\<forall>i. 1 \<le> i \<and> i < r + 1 \<longrightarrow> (label x i = 0)" "(r = n) \<or> (label x (r + 1) \<noteq> 0)"
+  shows "reduced label n x = r" apply(rule le_antisym) apply(rule_tac[!] ccontr) unfolding not_le
+  using reduced_labelling[of label n x] using assms by auto
+
+lemma reduced_labelling_0: assumes "j\<in>{1..n}" "label x j = 0" shows "reduced label n x \<noteq> j - 1"
+  using reduced_labelling[of label n x] using assms by fastsimp 
+
+lemma reduced_labelling_1: assumes "j\<in>{1..n}" "label x j \<noteq> 0" shows "reduced label n x < j"
+  using assms and reduced_labelling[of label n x] apply(erule_tac x=j in allE) by auto
+
+lemma reduced_labelling_Suc:
+  assumes "reduced lab (n + 1) x \<noteq> n + 1" shows "reduced lab (n + 1) x = reduced lab n x"
+  apply(subst eq_commute) apply(rule reduced_labelling_unique)
+  using reduced_labelling[of lab "n+1" x] and assms by auto 
+
+lemma complete_face_top:
+  assumes "\<forall>x\<in>f. \<forall>j\<in>{1..n+1}. x j = 0 \<longrightarrow> lab x j = 0"
+          "\<forall>x\<in>f. \<forall>j\<in>{1..n+1}. x j = p \<longrightarrow> lab x j = 1"
+  shows "((reduced lab (n + 1)) ` f = {0..n}) \<and> ((\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n+1}. \<forall>x\<in>f. x j = p)) \<longleftrightarrow>
+  ((reduced lab (n + 1)) ` f = {0..n}) \<and> (\<forall>x\<in>f. x (n + 1) = p)" (is "?l = ?r") proof
+  assume ?l (is "?as \<and> (?a \<or> ?b)") thus ?r apply-apply(rule,erule conjE,assumption) proof(cases ?a)
+    case True then guess j .. note j=this {fix x assume x:"x\<in>f"
+      have "reduced lab (n+1) x \<noteq> j - 1" using j apply-apply(rule reduced_labelling_0) defer apply(rule assms(1)[rule_format]) using x by auto }
+    moreover have "j - 1 \<in> {0..n}" using j by auto
+    then guess y unfolding `?l`[THEN conjunct1,THEN sym] and image_iff .. note y = this
+    ultimately have False by auto thus "\<forall>x\<in>f. x (n + 1) = p" by auto next
+
+    case False hence ?b using `?l` by blast then guess j .. note j=this {fix x assume x:"x\<in>f"
+      have "reduced lab (n+1) x < j" using j apply-apply(rule reduced_labelling_1) using assms(2)[rule_format,of x j] and x by auto } note * = this
+    have "j = n + 1" proof(rule ccontr) case goal1 hence "j < n + 1" using j by auto moreover
+      have "n \<in> {0..n}" by auto then guess y unfolding `?l`[THEN conjunct1,THEN sym] image_iff ..
+      ultimately show False using *[of y] by auto qed
+    thus "\<forall>x\<in>f. x (n + 1) = p" using j by auto qed qed(auto)
+
+subsection {* Hence we get just about the nice induction. *}
+
+lemma kuhn_induction:
+  assumes "0 < p" "\<forall>x. \<forall>j\<in>{1..n+1}. (\<forall>j. x j \<le> p) \<and> (x j = 0) \<longrightarrow> (lab x j = 0)"
+                  "\<forall>x. \<forall>j\<in>{1..n+1}. (\<forall>j. x j \<le> p) \<and> (x j = p) \<longrightarrow> (lab x j = 1)"
+        "odd (card {f. ksimplex p n f \<and> ((reduced lab n) ` f = {0..n})})"
+  shows "odd (card {s. ksimplex p (n+1) s \<and>((reduced lab (n+1)) `  s = {0..n+1})})" proof-
+  have *:"\<And>s t. odd (card s) \<Longrightarrow> s = t \<Longrightarrow> odd (card t)" "\<And>s f. (\<And>x. f x \<le> n +1 ) \<Longrightarrow> f ` s \<subseteq> {0..n+1}" by auto
+  show ?thesis apply(rule kuhn_simplex_lemma[unfolded mem_Collect_eq]) apply(rule,rule,rule *,rule reduced_labelling)
+    apply(rule *(1)[OF assms(4)]) apply(rule set_ext) unfolding mem_Collect_eq apply(rule,erule conjE) defer apply(rule) proof-(*(rule,rule)*)
+    fix f assume as:"ksimplex p n f" "reduced lab n ` f = {0..n}"
+    have *:"\<forall>x\<in>f. \<forall>j\<in>{1..n + 1}. x j = 0 \<longrightarrow> lab x j = 0" "\<forall>x\<in>f. \<forall>j\<in>{1..n + 1}. x j = p \<longrightarrow> lab x j = 1"
+      using assms(2-3) using as(1)[unfolded ksimplex_def] by auto
+    have allp:"\<forall>x\<in>f. x (n + 1) = p" using assms(2) using as(1)[unfolded ksimplex_def] by auto
+    { fix x assume "x\<in>f" hence "reduced lab (n + 1) x < n + 1" apply-apply(rule reduced_labelling_1)
+	defer using assms(3) using as(1)[unfolded ksimplex_def] by auto
+      hence "reduced lab (n + 1) x = reduced lab n x" apply-apply(rule reduced_labelling_Suc) using reduced_labelling(1) by auto }
+    hence "reduced lab (n + 1) ` f = {0..n}" unfolding as(2)[THEN sym] apply- apply(rule set_ext) unfolding image_iff by auto
+    moreover guess s using as(1)[unfolded simplex_top_face[OF assms(1) allp,THEN sym]] .. then guess a ..
+    ultimately show "\<exists>s a. ksimplex p (n + 1) s \<and>
+      a \<in> s \<and> f = s - {a} \<and> reduced lab (n + 1) ` f = {0..n} \<and> ((\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p))" (is ?ex)
+      apply(rule_tac x=s in exI,rule_tac x=a in exI) unfolding complete_face_top[OF *] using allp as(1) by auto
+  next fix f assume as:"\<exists>s a. ksimplex p (n + 1) s \<and>
+      a \<in> s \<and> f = s - {a} \<and> reduced lab (n + 1) ` f = {0..n} \<and> ((\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0) \<or> (\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p))" (is ?ex)
+    then guess s .. then guess a apply-apply(erule exE,(erule conjE)+) . note sa=this
+    { fix x assume "x\<in>f" hence "reduced lab (n + 1) x \<in> reduced lab (n + 1) ` f" by auto
+      hence "reduced lab (n + 1) x < n + 1" using sa(4) by auto 
+      hence "reduced lab (n + 1) x = reduced lab n x" apply-apply(rule reduced_labelling_Suc)
+	using reduced_labelling(1) by auto }
+    thus part1:"reduced lab n ` f = {0..n}" unfolding sa(4)[THEN sym] apply-apply(rule set_ext) unfolding image_iff by auto
+    have *:"\<forall>x\<in>f. x (n + 1) = p" proof(cases "\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = 0")
+      case True then guess j .. hence "\<And>x. x\<in>f \<Longrightarrow> reduced lab (n + 1) x \<noteq> j - 1" apply-apply(rule reduced_labelling_0) apply assumption
+	apply(rule assms(2)[rule_format]) using sa(1)[unfolded ksimplex_def] unfolding sa by auto moreover
+      have "j - 1 \<in> {0..n}" using `j\<in>{1..n+1}` by auto
+      ultimately have False unfolding sa(4)[THEN sym] unfolding image_iff by fastsimp thus ?thesis by auto next
+      case False hence "\<exists>j\<in>{1..n + 1}. \<forall>x\<in>f. x j = p" using sa(5) by fastsimp then guess j .. note j=this
+      thus ?thesis proof(cases "j = n+1")
+	case False hence *:"j\<in>{1..n}" using j by auto
+	hence "\<And>x. x\<in>f \<Longrightarrow> reduced lab n x < j" apply(rule reduced_labelling_1) proof- fix x assume "x\<in>f"
+	  hence "lab x j = 1" apply-apply(rule assms(3)[rule_format,OF j(1)]) 
+	    using sa(1)[unfolded ksimplex_def] using j unfolding sa by auto thus "lab x j \<noteq> 0" by auto qed
+	moreover have "j\<in>{0..n}" using * by auto
+	ultimately have False unfolding part1[THEN sym] using * unfolding image_iff by auto thus ?thesis by auto qed auto qed 
+    thus "ksimplex p n f" using as unfolding simplex_top_face[OF assms(1) *,THEN sym] by auto qed qed
+
+lemma kuhn_induction_Suc:
+  assumes "0 < p" "\<forall>x. \<forall>j\<in>{1..Suc n}. (\<forall>j. x j \<le> p) \<and> (x j = 0) \<longrightarrow> (lab x j = 0)"
+                  "\<forall>x. \<forall>j\<in>{1..Suc n}. (\<forall>j. x j \<le> p) \<and> (x j = p) \<longrightarrow> (lab x j = 1)"
+        "odd (card {f. ksimplex p n f \<and> ((reduced lab n) ` f = {0..n})})"
+  shows "odd (card {s. ksimplex p (Suc n) s \<and>((reduced lab (Suc n)) `  s = {0..Suc n})})"
+  using assms unfolding Suc_eq_plus1 by(rule kuhn_induction)
+
+subsection {* And so we get the final combinatorial result. *}
+
+lemma ksimplex_0: "ksimplex p 0 s \<longleftrightarrow> s = {(\<lambda>x. p)}" (is "?l = ?r") proof
+  assume l:?l guess a using ksimplexD(3)[OF l, unfolded add_0] unfolding card_1_exists .. note a=this
+  have "a = (\<lambda>x. p)" using ksimplexD(5)[OF l, rule_format, OF a(1)] by(rule,auto) thus ?r using a by auto next
+  assume r:?r show ?l unfolding r ksimplex_eq by auto qed
+
+lemma reduce_labelling_0[simp]: "reduced lab 0 x = 0" apply(rule reduced_labelling_unique) by auto
+
+lemma kuhn_combinatorial:
+  assumes "0 < p" "\<forall>x j. (\<forall>j. x(j) \<le> p) \<and> 1 \<le> j \<and> j \<le> n \<and> (x j = 0) \<longrightarrow> (lab x j = 0)"
+  "\<forall>x j. (\<forall>j. x(j) \<le> p) \<and> 1 \<le> j \<and> j \<le> n  \<and> (x j = p) \<longrightarrow> (lab x j = 1)"
+  shows " odd (card {s. ksimplex p n s \<and> ((reduced lab n) ` s = {0..n})})" using assms proof(induct n)
+  let ?M = "\<lambda>n. {s. ksimplex p n s \<and> ((reduced lab n) ` s = {0..n})}"
+  { case 0 have *:"?M 0 = {{(\<lambda>x. p)}}" unfolding ksimplex_0 by auto show ?case unfolding * by auto }
+  case (Suc n) have "odd (card (?M n))" apply(rule Suc(1)[OF Suc(2)]) using Suc(3-) by auto
+  thus ?case apply-apply(rule kuhn_induction_Suc) using Suc(2-) by auto qed
+
+lemma kuhn_lemma: assumes "0 < (p::nat)" "0 < (n::nat)"
+  "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. (label x i = (0::nat)) \<or> (label x i = 1))"
+  "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. (x i = 0) \<longrightarrow> (label x i = 0))"
+  "\<forall>x. (\<forall>i\<in>{1..n}. x i \<le> p) \<longrightarrow> (\<forall>i\<in>{1..n}. (x i = p) \<longrightarrow> (label x i = 1))"
+  obtains q where "\<forall>i\<in>{1..n}. q i < p"
+  "\<forall>i\<in>{1..n}. \<exists>r s. (\<forall>j\<in>{1..n}. q(j) \<le> r(j) \<and> r(j) \<le> q(j) + 1) \<and>
+                               (\<forall>j\<in>{1..n}. q(j) \<le> s(j) \<and> s(j) \<le> q(j) + 1) \<and>
+                               ~(label r i = label s i)" proof-
+  let ?A = "{s. ksimplex p n s \<and> reduced label n ` s = {0..n}}" have "n\<noteq>0" using assms by auto
+  have conjD:"\<And>P Q. P \<and> Q \<Longrightarrow> P" "\<And>P Q. P \<and> Q \<Longrightarrow> Q" by auto
+  have "odd (card ?A)" apply(rule kuhn_combinatorial[of p n label]) using assms by auto
+  hence "card ?A \<noteq> 0" apply-apply(rule ccontr) by auto hence "?A \<noteq> {}" unfolding card_eq_0_iff by auto
+  then obtain s where "s\<in>?A" by auto note s=conjD[OF this[unfolded mem_Collect_eq]]
+  guess a b apply(rule ksimplex_extrema_strong[OF s(1) `n\<noteq>0`]) . note ab=this
+  show ?thesis apply(rule that[of a]) proof(rule_tac[!] ballI) fix i assume "i\<in>{1..n}"
+    hence "a i + 1 \<le> p" apply-apply(rule order_trans[of _ "b i"]) apply(subst ab(5)[THEN spec[where x=i]])
+      using s(1)[unfolded ksimplex_def] defer apply- apply(erule conjE)+ apply(drule_tac bspec[OF _ ab(2)])+ by auto
+    thus "a i < p" by auto
+    case goal2 hence "i \<in> reduced label n ` s" using s by auto then guess u unfolding image_iff .. note u=this
+    from goal2 have "i - 1 \<in> reduced label n ` s" using s by auto then guess v unfolding image_iff .. note v=this
+    show ?case apply(rule_tac x=u in exI, rule_tac x=v in exI) apply(rule conjI) defer apply(rule conjI) defer 2 proof(rule_tac[1-2] ballI)
+      show "label u i \<noteq> label v i" using reduced_labelling[of label n u] reduced_labelling[of label n v]
+        unfolding u(2)[THEN sym] v(2)[THEN sym] using goal2 by auto
+      fix j assume j:"j\<in>{1..n}" show "a j \<le> u j \<and> u j \<le> a j + 1" "a j \<le> v j \<and> v j \<le> a j + 1"
+        using conjD[OF ab(4)[rule_format, OF u(1)]] and conjD[OF ab(4)[rule_format, OF v(1)]] apply- 
+        apply(drule_tac[!] kle_imp_pointwise)+ apply(erule_tac[!] x=j in allE)+ unfolding ab(5)[rule_format] using j
+        by auto qed qed qed
+
+subsection {* The main result for the unit cube. *}
+
+lemma kuhn_labelling_lemma':
+  assumes "(\<forall>x::nat\<Rightarrow>real. P x \<longrightarrow> P (f x))"  "\<forall>x. P x \<longrightarrow> (\<forall>i::nat. 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 brouwer_cube: fixes f::"real^'n::finite \<Rightarrow> real^'n::finite"
+  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
+  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)"
+    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 *:"\<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
+      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 "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 
+    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-
+    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])
+    guess e using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] apply-apply(erule exE,(erule conjE)+) .
+    note e=this[rule_format,unfolded vector_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"
+      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)+
+        have tria:"norm (y - x) \<le> norm (y - z) + norm (x - z)" using dist_triangle[of y x z,unfolded vector_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
+        finally show "norm (f y - f x) < d / real n / 8" apply- apply(rule e(2)) using as by auto
+        have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" apply(rule add_strict_mono) using as by auto
+        thus "norm (y - x) < 2 * (d / real n / 8)" using tria by auto
+        show "norm (f x - f z) < d / real n / 8" apply(rule e(2)) using as e(1) by auto qed(insert as, auto) qed qed
+  then guess e apply-apply(erule exE,(erule conjE)+) . note e=this[rule_format] 
+  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> "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  
+    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 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)"
+    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
+          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
+          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
+  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
+    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
+    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
+  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])
+    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])
+    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)
+  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)" 
+      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)" 
+      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
+  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
+  thus False using i by auto qed
+
+subsection {* Retractions. *}
+
+definition "retraction s t (r::real^'n::finite\<Rightarrow>real^'n) \<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
+  "(t retract_of s) \<longleftrightarrow> (\<exists>r. retraction s t r)"
+
+lemma retraction_idempotent: "retraction s t r \<Longrightarrow> x \<in> s \<Longrightarrow>  r(r x) = r x"
+  unfolding retraction_def by auto
+
+subsection {*preservation of fixpoints under (more general notion of) retraction. *}
+
+lemma invertible_fixpoint_property: fixes s::"(real^'n::finite) set" and t::"(real^'m::finite) 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-
+  have "\<exists>x\<in>s. (i \<circ> g \<circ> r) x = x" apply(rule assms(6)[rule_format],rule)
+    apply(rule continuous_on_compose assms)+ apply((rule continuous_on_subset)?,rule assms)+
+    using assms(2,4,8) unfolding image_compose by(auto,blast)
+    then guess x .. note x = this hence *:"g (r x) \<in> t" using assms(4,8) by auto
+    have "r ((i \<circ> g \<circ> r) x) = r x" using x by auto
+    thus ?thesis apply(rule_tac that[of "r x"]) using x unfolding o_def
+      unfolding assms(5)[rule_format,OF *] using assms(4) by auto qed
+
+lemma homeomorphic_fixpoint_property:
+  fixes s::"(real^'n::finite) set" and t::"(real^'m::finite) 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 ..
+  thus ?thesis apply- apply rule apply(rule_tac[!] allI impI)+ 
+    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:
+  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-
+    apply(rule invertible_fixpoint_property[OF continuous_on_id _ _ _ _ assms(2), of t h g]) prefer 7
+    apply(rule_tac y=y in that) using assms by auto qed
+
+subsection {*So the Brouwer theorem for any set with nonempty interior. *}
+
+lemma brouwer_weak: fixes f::"real^'n::finite \<Rightarrow> real^'n::finite"
+  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
+  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::finite \<Rightarrow> real^'n::finite"
+  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
+  using assms by auto
+
+text {*Still more general form; could derive this directly without using the 
+  rather involved HOMEOMORPHIC_CONVEX_COMPACT theorem, just using
+  a scaling and translation to put the set inside the unit cube. *}
+
+lemma brouwer: fixes f::"real^'n::finite \<Rightarrow> real^'n::finite"
+  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
+    apply(erule_tac exE,rule_tac x=b in exI) by(auto simp add: vector_dist_norm) 
+  then guess e apply-apply(erule exE,(erule conjE)+) . note e=this
+  have "\<exists>x\<in> cball 0 e. (f \<circ> closest_point s) x = x"
+    apply(rule_tac brouwer_ball[OF e(1), of 0 "f \<circ> closest_point s"]) apply(rule continuous_on_compose )
+    apply(rule continuous_on_closest_point[OF assms(2) compact_imp_closed[OF assms(1)] assms(3)])
+    apply(rule continuous_on_subset[OF assms(4)])
+    using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3)] apply - defer
+    using assms(5)[unfolded subset_eq] using e(2)[unfolded subset_eq mem_cball] by(auto simp add:vector_dist_norm)
+  then guess x .. note x=this
+  have *:"closest_point s x = x" apply(rule closest_point_self) 
+    apply(rule assms(5)[unfolded subset_eq,THEN bspec[where x="x"],unfolded image_iff])
+    apply(rule_tac x="closest_point s x" in bexI) using x unfolding o_def
+    using closest_point_in_set[OF compact_imp_closed[OF assms(1)] assms(3), of x] by auto
+  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. *}                                      
+
+lemma no_retraction_cball: assumes "0 < e" 
+  shows "\<not> (frontier(cball a e) retract_of (cball 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+
+    apply(rule_tac x=x in bexI) apply assumption+ apply(rule continuous_on_intros)+
+    unfolding frontier_cball subset_eq Ball_def image_iff apply(rule,rule,erule bexE)
+    unfolding vector_dist_norm apply(simp add: * norm_minus_commute) . note x = this
+  hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:group_simps)
+  hence "a = x" unfolding scaleR_left_distrib[THEN sym] by auto 
+  thus False using x using assms by auto qed
+
+subsection {*Bijections between intervals. *}
+
+definition "interval_bij = (\<lambda> (a,b) (u,v) (x::real^'n::finite).
+    (\<chi> i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)"
+
+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:group_simps field_simps add_divide_distrib[THEN sym]) 
+
+lemma continuous_interval_bij:
+  "continuous (at x) (interval_bij (a,b::real^'n::finite) (u,v))" 
+  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
+    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))"
+  apply(rule continuous_at_imp_continuous_on) by(rule, rule continuous_interval_bij)
+
+(** move this **)
+lemma divide_nonneg_nonneg:assumes "a \<ge> 0" "b \<ge> 0" shows "0 \<le> a / (b::real)"
+  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::finite}" 
+  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-
+    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)"
+    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)"
+    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
+
+lemma interval_bij_bij: 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
+  apply(rule,insert assms,erule_tac x=i in allE) by auto
+
+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
+
+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))"
+  unfolding infnorm_2 by auto
+
+lemma infnorm_eq_1_imp: assumes "infnorm (x::real^2) = 1" shows "abs(x$1) \<le> 1" "abs(x$2) \<le> 1"
+  using assms unfolding infnorm_eq_1_2 by auto
+
+lemma fashoda_unit: fixes f g::"real^1 \<Rightarrow> real^2"
+  assumes "f ` {- 1..1} \<subseteq> {- 1..1}" "g ` {- 1..1} \<subseteq> {- 1..1}"
+  "continuous_on {- 1..1} f"  "continuous_on {- 1..1} g"
+  "f (- 1)$1 = - 1" "f 1$1 = 1" "g (- 1) $2 = -1" "g 1 $2 = 1"
+  shows "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. f s = g t" proof(rule ccontr)
+  case goal1 note as = this[unfolded bex_simps,rule_format]
+  def sqprojection \<equiv> "\<lambda>z::real^2. (inverse (infnorm z)) *\<^sub>R z" 
+  def negatex \<equiv> "\<lambda>x::real^2. (vector [-(x$1), x$2])::real^2" 
+  have lem1:"\<forall>z::real^2. infnorm(negatex z) = infnorm z"
+    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
+  let ?F = "(\<lambda>w::real^2. (f \<circ> vec1 \<circ> (\<lambda>x. x$1)) w - (g \<circ> vec1 \<circ> (\<lambda>x. x$2)) w)"
+  have *:"\<And>i. vec1 ` (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real^1}"
+    apply(rule set_ext) unfolding image_iff Bex_def mem_interval apply rule defer 
+    apply(rule_tac x="dest_vec1 x" in exI) apply rule apply(rule_tac x="vec (dest_vec1 x)" in exI)
+    by(auto simp add: dest_vec1_def[THEN sym]) 
+  { fix x assume "x \<in> (\<lambda>w. (f \<circ> vec1 \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> vec1 \<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 "vec1 (w$1)" "vec1 (w$2)"] unfolding mem_interval 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 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)
+    apply(rule continuous_on_mul ) apply(rule continuous_at_imp_continuous_on,rule) apply(rule continuous_at_inv[unfolded o_def])
+    apply(rule continuous_at_infnorm) unfolding infnorm_eq_0 defer apply(rule continuous_on_id) apply(rule linear_continuous_on) proof-
+    show "bounded_linear negatex" apply(rule bounded_linearI') unfolding Cart_eq proof(rule_tac[!] allI) fix i::2 and x y::"real^2" and c::real
+      show "negatex (x + y) $ i = (negatex x + negatex y) $ i" "negatex (c *s x) $ i = (c *s negatex x) $ i"
+	apply-apply(case_tac[!] "i\<noteq>1") prefer 3 apply(drule_tac[1-2] 21) 
+	unfolding negatex_def by(auto simp add:vector_2 ) qed qed(insert x0, auto)
+  have 3:"(negatex \<circ> sqprojection \<circ> ?F) ` {- 1..1} \<subseteq> {- 1..1}" unfolding subset_eq apply rule proof-
+    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
+    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
+    apply(rule 1 2 3)+ . note x=this
+  have "?F x \<noteq> 0" apply(rule x0) using x(1) by auto
+  hence *:"infnorm (sqprojection (?F x)) = 1" unfolding o_def by(rule lem2[rule_format])
+  have nx:"infnorm x = 1" apply(subst x(2)[THEN sym]) unfolding *[THEN sym] o_def by(rule lem1[rule_format])
+  have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)"    "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)"
+    apply- apply(rule_tac[!] allI impI)+ proof- fix x::"real^2" and i::2 assume x:"x\<noteq>0"
+    have "inverse (infnorm x) > 0" using x[unfolded infnorm_pos_lt[THEN sym]] by auto
+    thus "(0 < sqprojection x $ i) = (0 < x $ i)"   "(sqprojection x $ i < 0) = (x $ i < 0)"
+      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:"vec1 (x $ 1) \<in> {- 1..1::real^1}" "vec1 (x $ 2) \<in> {- 1..1::real^1}" using x(1) unfolding mem_interval by auto
+  hence nz:"f (vec1 (x $ 1)) - g (vec1 (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 
+    presume "P \<or> Q \<or> R \<or> S" "P\<Longrightarrow>False" "Q\<Longrightarrow>False" "R\<Longrightarrow>False" "S\<Longrightarrow>False" thus False by auto
+  next assume as:"x$1 = 1" hence "vec1 (x$1) = 1" unfolding Cart_eq by auto
+    hence *:"f (vec1 (x $ 1)) $ 1 = 1" using assms(6) by auto
+    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 1 < 0"
+      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 (vec1 (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 
+      apply(erule_tac x=1 in allE) by auto 
+  next assume as:"x$1 = -1" hence "vec1 (x$1) = - 1" unfolding Cart_eq by auto
+    hence *:"f (vec1 (x $ 1)) $ 1 = - 1" using assms(5) by auto
+    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 1 > 0"
+      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 (vec1 (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 
+      apply(erule_tac x=1 in allE) by auto
+  next assume as:"x$2 = 1" hence "vec1 (x$2) = 1" unfolding Cart_eq by auto
+    hence *:"g (vec1 (x $ 2)) $ 2 = 1" using assms(8) by auto
+    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 2 > 0"
+      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 (vec1 (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 
+     apply(erule_tac x=2 in allE) by auto
+ next assume as:"x$2 = -1" hence "vec1 (x$2) = - 1" unfolding Cart_eq by auto
+    hence *:"g (vec1 (x $ 2)) $ 2 = - 1" using assms(7) by auto
+    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 2 < 0"
+      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 (vec1 (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 
+      apply(erule_tac x=2 in allE) by auto qed(auto) qed
+
+lemma fashoda_unit_path: fixes f ::"real^1 \<Rightarrow> real^2" and g ::"real^1 \<Rightarrow> real^2"
+  assumes "path f" "path g" "path_image f \<subseteq> {- 1..1}" "path_image g \<subseteq> {- 1..1}"
+  "(pathstart f)$1 = -1" "(pathfinish f)$1 = 1"  "(pathstart g)$2 = -1" "(pathfinish g)$2 = 1"
+  obtains z where "z \<in> path_image f" "z \<in> path_image g" proof-
+  note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]
+  def iscale \<equiv> "\<lambda>z::real^1. inverse 2 *\<^sub>R (z + 1)"
+  have isc:"iscale ` {- 1..1} \<subseteq> {0..1}" unfolding iscale_def by(auto simp add:dest_vec1_add dest_vec1_neg)
+  have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t" proof(rule fashoda_unit) 
+    show "(f \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}" "(g \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}"
+      using isc and assms(3-4) unfolding image_compose by auto
+    have *:"continuous_on {- 1..1} iscale" unfolding iscale_def by(rule continuous_on_intros)+
+    show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)"
+      apply-apply(rule_tac[!] continuous_on_compose[OF *]) apply(rule_tac[!] continuous_on_subset[OF _ isc])
+      by(rule assms)+ have *:"(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1" unfolding Cart_eq by auto
+    show "(f \<circ> iscale) (- 1) $ 1 = - 1" "(f \<circ> iscale) 1 $ 1 = 1" "(g \<circ> iscale) (- 1) $ 2 = -1" "(g \<circ> iscale) 1 $ 2 = 1"
+      unfolding o_def iscale_def using assms by(auto simp add:*) qed
+  then guess s .. from this(2) guess t .. note st=this
+  show thesis apply(rule_tac z="f (iscale s)" in that)
+    using st `s\<in>{- 1..1}` unfolding o_def path_image_def image_iff apply-
+    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
+
+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"
+  "(pathstart g)$2 = a$2" "(pathfinish g)$2 = b$2"
+  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)
+  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)
+    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
+  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)
+    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
+  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
+  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)+
+    unfolding subset_eq apply(rule_tac[1-2] ballI)
+  proof- fix x assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
+    then guess y unfolding image_iff .. note y=this
+    show "x \<in> {- 1..1}" unfolding y o_def apply(rule in_interval_interval_bij)
+      using y(1) using assms(3)[unfolded path_image_def subset_eq] int_nem by auto
+  next fix x assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
+    then guess y unfolding image_iff .. note y=this
+    show "x \<in> {- 1..1}" unfolding y o_def apply(rule in_interval_interval_bij)
+      using y(1) using assms(4)[unfolded path_image_def subset_eq] int_nem by auto
+  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
+      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
+    using zf(1) zg(1) by auto qed
+
+subsection {*Some slightly ad hoc lemmas I use below*}
+
+lemma segment_vertical: fixes a::"real^2" assumes "a$1 = b$1"
+  shows "x \<in> closed_segment a b \<longleftrightarrow> (x$1 = a$1 \<and> x$1 = b$1 \<and>
+           (a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2))" (is "_ = ?R")
+proof- 
+  let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
+  { presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq
+      unfolding Cart_eq forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast }
+  { assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this
+    { fix b a assume "b + u * a > a + u * b"
+      hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps)
+      hence "b \<ge> a" apply(drule_tac mult_less_imp_less_left) using u by auto
+      hence "u * a \<le> u * b" apply-apply(rule mult_left_mono[OF _ u(3)]) 
+        using u(3-4) by(auto simp add:field_simps) } note * = this
+    { fix a b assume "u * b > u * a" hence "(1 - u) * a \<le> (1 - u) * b" apply-apply(rule mult_left_mono)
+        apply(drule mult_less_imp_less_left) using u by auto
+      hence "a + u * b \<le> b + u * a" by(auto simp add:field_simps) } note ** = this
+    thus ?R unfolding u assms using u by(auto simp add:field_simps not_le intro:* **) }
+  { assume ?R thus ?L proof(cases "x$2 = b$2")
+      case True thus ?L apply(rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI) unfolding assms True
+        using `?R` by(auto simp add:field_simps)
+    next case False thus ?L apply(rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI) unfolding assms using `?R`
+        by(auto simp add:field_simps)
+    qed } qed
+
+lemma segment_horizontal: fixes a::"real^2" assumes "a$2 = b$2"
+  shows "x \<in> closed_segment a b \<longleftrightarrow> (x$2 = a$2 \<and> x$2 = b$2 \<and>
+           (a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1))" (is "_ = ?R")
+proof- 
+  let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
+  { presume "?L \<Longrightarrow> ?R" "?R \<Longrightarrow> ?L" thus ?thesis unfolding closed_segment_def mem_Collect_eq
+      unfolding Cart_eq forall_2 smult_conv_scaleR[THEN sym] vector_component_simps by blast }
+  { assume ?L then guess u apply-apply(erule exE)apply(erule conjE)+ . note u=this
+    { fix b a assume "b + u * a > a + u * b"
+      hence "(1 - u) * b > (1 - u) * a" by(auto simp add:field_simps)
+      hence "b \<ge> a" apply(drule_tac mult_less_imp_less_left) using u by auto
+      hence "u * a \<le> u * b" apply-apply(rule mult_left_mono[OF _ u(3)]) 
+        using u(3-4) by(auto simp add:field_simps) } note * = this
+    { fix a b assume "u * b > u * a" hence "(1 - u) * a \<le> (1 - u) * b" apply-apply(rule mult_left_mono)
+        apply(drule mult_less_imp_less_left) using u by auto
+      hence "a + u * b \<le> b + u * a" by(auto simp add:field_simps) } note ** = this
+    thus ?R unfolding u assms using u by(auto simp add:field_simps not_le intro:* **) }
+  { assume ?R thus ?L proof(cases "x$1 = b$1")
+      case True thus ?L apply(rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI) unfolding assms True
+        using `?R` by(auto simp add:field_simps)
+    next case False thus ?L apply(rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI) unfolding assms using `?R`
+        by(auto simp add:field_simps)
+    qed } qed
+
+subsection {*useful Fashoda corollary pointed out to me by Tom Hales. *}
+
+lemma fashoda_interlace: fixes a::"real^2"
+  assumes "path f" "path g"
+  "path_image f \<subseteq> {a..b}" "path_image g \<subseteq> {a..b}"
+  "(pathstart f)$2 = a$2" "(pathfinish f)$2 = a$2"
+  "(pathstart g)$2 = a$2" "(pathfinish g)$2 = a$2"
+  "(pathstart f)$1 < (pathstart g)$1" "(pathstart g)$1 < (pathfinish f)$1"
+  "(pathfinish f)$1 < (pathfinish g)$1"
+  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]
+  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]
+  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]) +++
+     linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" 
+  let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++
+     linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++
+     linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++
+     linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])"
+  let ?a = "vector[a$1 - 2, a$2 - 3]"
+  let ?b = "vector[b$1 + 2, b$2 + 3]"
+  have P1P2:"path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union>
+      path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union>
+      path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union>
+      path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))"
+    "path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union>
+      path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union>
+      path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union>
+      path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2)
+      by(auto simp add: pathstart_join pathfinish_join path_image_join path_image_linepath path_join path_linepath) 
+  have abab: "{a..b} \<subseteq> {?a..?b}" by(auto simp add:vector_le_def forall_2 vector_2)
+  guess z apply(rule fashoda[of ?P1 ?P2 ?a ?b])
+    unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 proof-
+    show "path ?P1" "path ?P2" using assms by(auto simp add: pathstart_join pathfinish_join path_join)
+    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)
+      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)
+      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"
+      by(auto simp add: assms)
+  qed note z=this[unfolded P1P2 path_image_linepath]
+  show thesis apply(rule that[of z]) proof-
+    have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or>
+     z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or>
+   z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or>
+  z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow>
+  (((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or>
+    z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or>
+   z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or>
+  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 "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 "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]
+      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)
+      thus False unfolding * using ab by auto
+    qed hence "z \<in> path_image f \<or> z \<in> path_image g" using z unfolding Un_iff by blast
+    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-
+      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-
+      apply(simp only: segment_vertical segment_horizontal vector_2) unfolding assms by auto
+  qed qed
+
+(** The Following still needs to be translated. Maybe I will do that later.
+
+(* ------------------------------------------------------------------------- *)
+(* Complement in dimension N >= 2 of set homeomorphic to any interval in     *)
+(* any dimension is (path-)connected. This naively generalizes the argument  *)
+(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer        *)
+(* fixed point theorem", American Mathematical Monthly 1984.                 *)
+(* ------------------------------------------------------------------------- *)
+
+let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove
+ (`!p:real^M->real^N a b.
+        ~(interval[a,b] = {}) /\
+        p continuous_on interval[a,b] /\
+        (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y)
+        ==> ?f. f continuous_on (:real^N) /\
+                IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\
+                (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`,
+  REPEAT STRIP_TAC THEN
+  FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN
+  DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN
+  SUBGOAL_THEN `(q:real^N->real^M) continuous_on
+                (IMAGE p (interval[a:real^M,b]))`
+  ASSUME_TAC THENL
+   [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL];
+    ALL_TAC] THEN
+  MP_TAC(ISPECL [`q:real^N->real^M`;
+                 `IMAGE (p:real^M->real^N)
+                 (interval[a,b])`;
+                 `a:real^M`; `b:real^M`]
+        TIETZE_CLOSED_INTERVAL) THEN
+  ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE;
+               COMPACT_IMP_CLOSED] THEN
+  ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN
+  EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN
+  REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN
+  CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN
+  MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN
+  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ]
+        CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);;
+
+let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
+ (`!s:real^N->bool a b:real^M.
+        s homeomorphic (interval[a,b])
+        ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`,
+  REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN
+  REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+  MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN
+  DISCH_TAC THEN
+  SUBGOAL_THEN
+   `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\
+          (p:real^M->real^N) x = p y ==> x = y`
+  ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN
+  FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN
+  DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN
+  ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN
+  ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV;
+                  NOT_BOUNDED_UNIV] THEN
+  ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN
+  X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN
+  SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+  SUBGOAL_THEN `bounded((path_component s c) UNION
+                        (IMAGE (p:real^M->real^N) (interval[a,b])))`
+  MP_TAC THENL
+   [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED;
+                 COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
+    ALL_TAC] THEN
+  DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN
+  REWRITE_TAC[UNION_SUBSET] THEN
+  DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN
+  MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`]
+    RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN
+  ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN
+  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN
+  DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC
+   (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN
+  REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN
+  ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN
+  SUBGOAL_THEN
+    `(q:real^N->real^N) continuous_on
+     (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))`
+  MP_TAC THENL
+   [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN
+    REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN
+    REPEAT CONJ_TAC THENL
+     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
+      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
+                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
+      ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV];
+      ALL_TAC] THEN
+    X_GEN_TAC `z:real^N` THEN
+    REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN
+    STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
+    MP_TAC(ISPECL
+     [`path_component s (z:real^N)`; `path_component s (c:real^N)`]
+     OPEN_INTER_CLOSURE_EQ_EMPTY) THEN
+    ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL
+     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
+      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
+                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
+      REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN
+      DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN
+      GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN
+      REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]];
+    ALL_TAC] THEN
+  SUBGOAL_THEN
+   `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) =
+    (:real^N)`
+  SUBST1_TAC THENL
+   [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN
+    REWRITE_TAC[CLOSURE_SUBSET];
+    DISCH_TAC] THEN
+  MP_TAC(ISPECL
+   [`(\x. &2 % c - x) o
+     (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`;
+    `cball(c:real^N,B)`]
+    BROUWER) THEN
+  REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN
+  ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN
+  SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL
+   [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN
+    REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN
+    ASM SET_TAC[PATH_COMPONENT_REFL_EQ];
+    ALL_TAC] THEN
+  REPEAT CONJ_TAC THENL
+   [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
+    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL
+     [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_MUL THEN
+    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
+    REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN
+    MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN
+    ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN
+    SUBGOAL_THEN
+     `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)`
+    SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
+    ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST;
+                 CONTINUOUS_ON_LIFT_NORM];
+    REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN
+    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+    REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN
+    REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
+    ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
+    ASM_REAL_ARITH_TAC;
+    REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN
+    REWRITE_TAC[IN_CBALL; o_THM; dist] THEN
+    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+    REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN
+    ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL
+     [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN
+      REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
+      ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
+      ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN
+      UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN
+      REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB];
+      EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN
+      REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN
+      ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN
+      SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL
+       [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN
+      ASM_REWRITE_TAC[] THEN
+      MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN
+      ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);;
+
+let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
+ (`!s:real^N->bool a b:real^M.
+        2 <= dimindex(:N) /\ s homeomorphic interval[a,b]
+        ==> path_connected((:real^N) DIFF s)`,
+  REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
+  FIRST_ASSUM(MP_TAC o MATCH_MP
+    UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
+  ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN
+  ABBREV_TAC `t = (:real^N) DIFF s` THEN
+  DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN
+  STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN
+  REWRITE_TAC[COMPACT_INTERVAL] THEN
+  DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN
+  REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN
+  X_GEN_TAC `B:real` THEN STRIP_TAC THEN
+  SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\
+                (?v:real^N. v IN path_component t y /\ B < norm(v))`
+  STRIP_ASSUME_TAC THENL
+   [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN
+  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN
+  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
+  MATCH_MP_TAC PATH_COMPONENT_SYM THEN
+  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN
+  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
+  MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN
+  EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL
+   [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE
+     `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN
+    ASM_REWRITE_TAC[SUBSET; IN_CBALL_0];
+    MP_TAC(ISPEC `cball(vec 0:real^N,B)`
+       PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN
+    ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN
+    REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
+    DISCH_THEN MATCH_MP_TAC THEN
+    ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);;
+
+(* ------------------------------------------------------------------------- *)
+(* In particular, apply all these to the special case of an arc.             *)
+(* ------------------------------------------------------------------------- *)
+
+let RETRACTION_ARC = prove
+ (`!p. arc p
+       ==> ?f. f continuous_on (:real^N) /\
+               IMAGE f (:real^N) SUBSET path_image p /\
+               (!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]);;
+
+let PATH_CONNECTED_ARC_COMPLEMENT = prove
+ (`!p. 2 <= dimindex(:N) /\ arc p
+       ==> path_connected((:real^N) DIFF path_image p)`,
+  REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN
+  MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`]
+    PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
+  ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN
+  ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN
+  MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN
+  EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);;
+
+let CONNECTED_ARC_COMPLEMENT = prove
+ (`!p. 2 <= dimindex(:N) /\ arc p
+       ==> connected((:real^N) DIFF path_image p)`,
+  SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *)
+
+end 
+   
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -23,7 +23,7 @@
 
 lemma dim1in[intro]:"Suc 0 \<in> {1::nat .. CARD(1)}" by auto
 
-lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_less_eq_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component
+lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_le_def Cart_lambda_beta dest_vec1_def basis_component vector_uminus_component
 
 lemmas continuous_intros = continuous_add continuous_vmul continuous_cmul continuous_const continuous_sub continuous_at_id continuous_within_id
 
@@ -61,7 +61,7 @@
 lemma mem_interval_1: fixes x :: "real^1" shows
  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
  "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp_all add: Cart_eq vector_less_def vector_less_eq_def dest_vec1_def all_1)
+by(simp_all add: Cart_eq vector_less_def vector_le_def dest_vec1_def all_1)
 
 lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n::finite)) ` {a..b} =
   (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})"
@@ -77,7 +77,7 @@
   apply(rule_tac [!] allI)apply(rule_tac [!] impI)
   apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI)
   apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI)
-  by (auto simp add: vector_less_def vector_less_eq_def all_1 dest_vec1_def
+  by (auto simp add: vector_less_def vector_le_def all_1 dest_vec1_def
     vec1_dest_vec1[unfolded dest_vec1_def One_nat_def])
 
 lemma dest_vec1_setsum: assumes "finite S"
@@ -2354,7 +2354,7 @@
   assumes "dest_vec1 a \<le> dest_vec1 b" "continuous_on {a .. b} f" "(f a)$k \<le> y" "y \<le> (f b)$k"
   shows "\<exists>x\<in>{a..b}. (f x)$k = y"
 proof- have "f a \<in> f ` {a..b}" "f b \<in> f ` {a..b}" apply(rule_tac[!] imageI) 
-    using assms(1) by(auto simp add: vector_less_eq_def dest_vec1_def)
+    using assms(1) by(auto simp add: vector_le_def dest_vec1_def)
   thus ?thesis using connected_ivt_component[of "f ` {a..b}" "f a" "f b" k y]
     using connected_continuous_image[OF assms(2) convex_connected[OF convex_interval(1)]]
     using assms by(auto intro!: imageI) qed
@@ -2415,10 +2415,10 @@
       show ?thesis proof(cases "x$i=1")
         case True have "\<forall>j\<in>{i. x$i \<noteq> 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof-
           fix j assume "x $ j \<noteq> 0" "x $ j \<noteq> 1"
-          hence j:"x$j \<in> {0<..<1}" using Suc(2) by(auto simp add: vector_less_eq_def elim!:allE[where x=j])
+          hence 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_less_eq_def elim!:ballE[where x=j]) qed
+          thus False using True Suc(2) j by(auto simp add: vector_le_def elim!:ballE[where x=j]) qed
         thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
           by(auto simp add: Cart_lambda_beta)
       next let ?y = "\<lambda>j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)"
@@ -2439,7 +2439,7 @@
   show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule 
     apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
     unfolding mem_interval apply rule unfolding mem_Collect_eq apply(erule_tac x=i in allE)
-    by(auto simp add: vector_less_eq_def mem_def[of _ convex]) qed
+    by(auto simp add: vector_le_def mem_def[of _ convex]) qed
 
 subsection {* And this is a finite set of vertices. *}
 
@@ -2469,7 +2469,7 @@
     hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..1}" unfolding mem_interval using assms
       by(auto simp add: Cart_eq vector_component_simps field_simps)
     thus "\<exists>z\<in>{0..1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) 
-      using assms by(auto simp add: Cart_eq vector_less_eq_def Cart_lambda_beta)
+      using assms by(auto simp add: Cart_eq vector_le_def Cart_lambda_beta)
   next
     fix y z assume as:"z\<in>{0..1}" "y = x - ?d + (2*d) *\<^sub>R z" 
     have "\<And>i. 0 \<le> d * z $ i \<and> d * z $ i \<le> d" using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
@@ -2911,7 +2911,7 @@
 lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof-
   have *:"\<And>g. path_image(reversepath g) \<subseteq> path_image g"
     unfolding path_image_def subset_eq reversepath_def Ball_def image_iff apply(rule,rule,erule bexE)  
-    apply(rule_tac x="1 - xa" in bexI) by(auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
+    apply(rule_tac x="1 - xa" in bexI) by(auto simp add:vector_le_def vector_component_simps elim!:ballE)
   show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
 
 lemma path_reversepath[simp]: "path(reversepath g) \<longleftrightarrow> path g" proof-
@@ -2919,7 +2919,7 @@
     apply(rule continuous_on_compose[unfolded o_def, of _ "\<lambda>x. 1 - x"])
     apply(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id)
     apply(rule continuous_on_subset[of "{0..1}"], assumption)
-    by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
+    by (auto, auto simp add:vector_le_def vector_component_simps elim!:ballE)
   show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
 
 lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
@@ -2930,7 +2930,7 @@
   have *:"g1 = (\<lambda>x. g1 (2 *\<^sub>R x)) \<circ> (\<lambda>x. (1/2) *\<^sub>R x)" 
          "g2 = (\<lambda>x. g2 (2 *\<^sub>R x - 1)) \<circ> (\<lambda>x. (1/2) *\<^sub>R (x + 1))" unfolding o_def by auto
   have "op *\<^sub>R (1 / 2) ` {0::real^1..1} \<subseteq> {0..1}"  "(\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {(0::real^1)..1} \<subseteq> {0..1}"
-    unfolding image_smult_interval by (auto, auto simp add:vector_less_eq_def vector_component_simps elim!:ballE)
+    unfolding image_smult_interval by (auto, auto simp add:vector_le_def vector_component_simps elim!:ballE)
   thus "continuous_on {0..1} g1 \<and> continuous_on {0..1} g2" apply -apply rule
     apply(subst *) defer apply(subst *) apply (rule_tac[!] continuous_on_compose)
     apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -0,0 +1,1332 @@
+(*  Title:      HOL/Library/Convex_Euclidean_Space.thy
+    Author:                     John Harrison
+    Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+
+header {* Multivariate calculus in Euclidean space. *}
+
+theory Derivative
+  imports Brouwer_Fixpoint RealVector
+begin
+
+
+(* Because I do not want to type this all the time *)
+lemmas linear_linear = linear_conv_bounded_linear[THEN sym]
+
+subsection {* Derivatives *}
+
+text {* The definition is slightly tricky since we make it work over
+  nets of a particular form. This lets us prove theorems generally and use 
+  "at a" or "at a within s" for restriction to a set (1-sided on R etc.) *}
+
+definition has_derivative :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a net \<Rightarrow> bool)"
+(infixl "(has'_derivative)" 12) where
+ "(f has_derivative f') net \<equiv> bounded_linear f' \<and> ((\<lambda>y. (1 / (norm (y - netlimit net))) *\<^sub>R
+   (f y - (f (netlimit net) + f'(y - netlimit net)))) ---> 0) net"
+
+lemma derivative_linear[dest]:"(f has_derivative f') net \<Longrightarrow> bounded_linear f'"
+  unfolding has_derivative_def by auto
+
+lemma FDERIV_conv_has_derivative:"FDERIV f (x::'a::{real_normed_vector,perfect_space}) :> f' = (f has_derivative f') (at x)" (is "?l = ?r") proof 
+  assume ?l note as = this[unfolded fderiv_def]
+  show ?r unfolding has_derivative_def Lim_at apply-apply(rule,rule as[THEN conjunct1]) proof(rule,rule)
+    fix e::real assume "e>0"
+    guess d using as[THEN conjunct2,unfolded LIM_def,rule_format,OF`e>0`] ..
+    thus "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow>
+      dist ((1 / norm (xa - netlimit (at x))) *\<^sub>R (f xa - (f (netlimit (at x)) + f' (xa - netlimit (at x))))) (0) < e"
+      apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa - x" in allE)
+      unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed next
+  assume ?r note as = this[unfolded has_derivative_def]
+  show ?l unfolding fderiv_def LIM_def apply-apply(rule,rule as[THEN conjunct1]) proof(rule,rule)
+    fix e::real assume "e>0"
+    guess d using as[THEN conjunct2,unfolded Lim_at,rule_format,OF`e>0`] ..
+    thus "\<exists>s>0. \<forall>xa. xa \<noteq> 0 \<and> dist xa 0 < s \<longrightarrow> dist (norm (f (x + xa) - f x - f' xa) / norm xa) 0 < e" apply-
+      apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa + x" in allE)
+      unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed qed
+
+subsection {* These are the only cases we'll care about, probably. *}
+
+lemma has_derivative_within: "(f has_derivative f') (at x within s) \<longleftrightarrow>
+         bounded_linear f' \<and> ((\<lambda>y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) ---> 0) (at x within s)"
+  unfolding has_derivative_def and Lim by(auto simp add:netlimit_within)
+
+lemma has_derivative_at: "(f has_derivative f') (at x) \<longleftrightarrow>
+         bounded_linear f' \<and> ((\<lambda>y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) ---> 0) (at x)"
+  apply(subst within_UNIV[THEN sym]) unfolding has_derivative_within unfolding within_UNIV by auto
+
+subsection {* More explicit epsilon-delta forms. *}
+
+lemma has_derivative_within':
+  "(f has_derivative f')(at x within s) \<longleftrightarrow> bounded_linear f' \<and>
+        (\<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 vector_dist_norm
+  unfolding diff_0_right norm_mul by(simp add: group_simps)
+
+lemma has_derivative_at':
+ "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
+   (\<forall>e>0. \<exists>d>0. \<forall>x'. 0 < norm(x' - x) \<and> norm(x' - x) < d
+        \<longrightarrow> norm(f x' - f x - f'(x' - x)) / norm(x' - x) < e)"
+  apply(subst within_UNIV[THEN sym]) unfolding has_derivative_within' by auto
+
+lemma has_derivative_at_within: "(f has_derivative f') (at x) \<Longrightarrow> (f has_derivative f') (at x within s)"
+  unfolding has_derivative_within' has_derivative_at' by meson
+
+lemma has_derivative_within_open:
+  "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 real^1 *}
+
+lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding vector_dist_norm by(auto simp add:vec1_dest_vec1_simps)
+
+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 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::finite \<Rightarrow> real^'b::finite" 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"
+  unfolding has_derivative_def apply(rule,rule bounded_linear_axioms)
+  unfolding diff by(simp add: Lim_const)
+
+lemma has_derivative_id: "((\<lambda>x. x) has_derivative (\<lambda>h. h)) net"
+  apply(rule bounded_linear.has_derivative) using bounded_linear_ident[unfolded id_def] by simp
+
+lemma has_derivative_const: "((\<lambda>x. c) has_derivative (\<lambda>h. 0)) net"
+  unfolding has_derivative_def apply(rule,rule bounded_linear_zero) by(simp add: Lim_const)
+
+lemma (in bounded_linear) cmul: shows "bounded_linear (\<lambda>x. (c::real) *\<^sub>R f x)" proof
+  guess K using pos_bounded ..
+  thus "\<exists>K. \<forall>x. norm ((c::real) *\<^sub>R f x) \<le> norm x * K" apply(rule_tac x="abs c * K" in exI) proof
+    fix x case goal1
+    hence "abs c * norm (f x) \<le> abs c * (norm x * K)" apply-apply(erule conjE,erule_tac x=x in allE)  
+      apply(rule mult_left_mono) by auto
+    thus ?case by(auto simp add:field_simps)
+  qed qed(auto simp add: scaleR.add_right add scaleR)
+
+lemma has_derivative_cmul: assumes "(f has_derivative f') net" shows "((\<lambda>x. c *\<^sub>R f(x)) has_derivative (\<lambda>h. c *\<^sub>R f'(h))) net"
+  unfolding has_derivative_def apply(rule,rule bounded_linear.cmul)
+  using assms[unfolded has_derivative_def] using Lim_cmul[OF assms[unfolded has_derivative_def,THEN conjunct2]]
+  unfolding scaleR_right_diff_distrib scaleR_right_distrib by auto 
+
+lemma has_derivative_cmul_eq: assumes "c \<noteq> 0" 
+  shows "(((\<lambda>x. c *\<^sub>R f(x)) has_derivative (\<lambda>h. c *\<^sub>R f'(h))) net \<longleftrightarrow> (f has_derivative f') net)"
+  apply(rule) defer apply(rule has_derivative_cmul,assumption) 
+  apply(drule has_derivative_cmul[where c="1/c"]) using assms by auto
+
+lemma has_derivative_neg:
+ "(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. -(f x)) has_derivative (\<lambda>h. -(f' h))) net"
+  apply(drule has_derivative_cmul[where c="-1"]) by auto
+
+lemma has_derivative_neg_eq: "((\<lambda>x. -(f x)) has_derivative (\<lambda>h. -(f' h))) net \<longleftrightarrow> (f has_derivative f') net"
+  apply(rule, drule_tac[!] has_derivative_neg) by auto
+
+lemma has_derivative_add: assumes "(f has_derivative f') net" "(g has_derivative g') net"
+  shows "((\<lambda>x. f(x) + g(x)) has_derivative (\<lambda>h. f'(h) + g'(h))) net" proof-
+  note as = assms[unfolded has_derivative_def]
+  show ?thesis unfolding has_derivative_def apply(rule,rule bounded_linear_add)
+    using Lim_add[OF as(1)[THEN conjunct2] as(2)[THEN conjunct2]] and as
+    by(auto simp add:group_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
+
+lemma has_derivative_add_const:"(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net"
+  apply(drule has_derivative_add) apply(rule has_derivative_const) by auto
+
+lemma has_derivative_sub:
+ "(f has_derivative f') net \<Longrightarrow> (g has_derivative g') net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_derivative (\<lambda>h. f'(h) - g'(h))) net"
+  apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:group_simps)
+
+lemma has_derivative_setsum: assumes "finite s" "\<forall>a\<in>s. ((f a) has_derivative (f' a)) net"
+  shows "((\<lambda>x. setsum (\<lambda>a. f a x) s) has_derivative (\<lambda>h. setsum (\<lambda>a. f' a h) s)) net"
+  apply(induct_tac s rule:finite_subset_induct[where A=s]) apply(rule assms(1)) 
+proof- fix x F assume as:"finite F" "x \<notin> F" "x\<in>s" "((\<lambda>x. \<Sum>a\<in>F. f a x) has_derivative (\<lambda>h. \<Sum>a\<in>F. f' a h)) net" 
+  thus "((\<lambda>xa. \<Sum>a\<in>insert x F. f a xa) has_derivative (\<lambda>h. \<Sum>a\<in>insert x F. f' a h)) net"
+    unfolding setsum_insert[OF as(1,2)] apply-apply(rule has_derivative_add) apply(rule assms(2)[rule_format]) by auto
+qed(auto intro!: has_derivative_const)
+
+lemma has_derivative_setsum_numseg:
+  "\<forall>i. m \<le> i \<and> i \<le> n \<longrightarrow> ((f i) has_derivative (f' i)) net \<Longrightarrow>
+  ((\<lambda>x. setsum (\<lambda>i. f i x) {m..n::nat}) has_derivative (\<lambda>h. setsum (\<lambda>i. f' i h) {m..n})) net"
+  apply(rule has_derivative_setsum) by auto
+
+subsection {* somewhat different results for derivative of scalar multiplier. *}
+
+lemma has_derivative_vmul_component: fixes c::"real^'a::finite \<Rightarrow> real^'b::finite" and v::"real^'c::finite"
+  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" 
+    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)
+    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)
+    fix e assume "\<not> trivial_limit net" "0 < (e::real)"
+    then obtain A where A:"A\<in>Rep_net net" "\<forall>x\<in>A. dist ((1 / norm (x - netlimit net)) *\<^sub>R (c x - (c (netlimit net) + c' (x - netlimit net)))) 0 < e"
+      using assms[unfolded has_derivative_def Lim] unfolding eventually_def by auto
+    show "eventually (\<lambda>x. dist (1 / norm (x - netlimit net) * (c x $ k - (c (netlimit net) $ k + c' (x - netlimit net) $ k))) 0 < e) net"
+      unfolding eventually_def apply(rule_tac x=A in bexI) apply rule proof-
+      case goal1 thus ?case apply -apply(drule A(2)[rule_format]) unfolding vector_dist_norm vec1_vec apply(rule le_less_trans) prefer 2 apply assumption unfolding * ** and norm_vec1[unfolded vec1_vec]
+	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(insert A, auto) 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::finite"
+  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
+
+lemma has_derivative_vmul_at: fixes c::"real \<Rightarrow> real" and v::"real^'a::finite"
+  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)
+
+lemma has_derivative_lift_dot:
+  assumes "(f has_derivative f') net"
+  shows "((\<lambda>x. inner v (f x)) has_derivative (\<lambda>t. inner v (f' t))) net" proof-
+  show ?thesis using assms unfolding has_derivative_def apply- apply(erule conjE,rule)
+    apply(rule bounded_linear_compose[of _ f']) apply(rule inner.bounded_linear_right,assumption)
+    apply(drule Lim_inner[where a=v]) unfolding o_def
+    by(auto simp add:inner.scaleR_right inner.add_right inner.diff_right) qed
+
+lemmas has_derivative_intros = has_derivative_sub has_derivative_add has_derivative_cmul has_derivative_id has_derivative_const
+   has_derivative_neg has_derivative_vmul_component has_derivative_vmul_at has_derivative_vmul_within has_derivative_cmul 
+   bounded_linear.has_derivative has_derivative_lift_dot
+
+subsection {* limit transformation for derivatives. *}
+
+lemma has_derivative_transform_within:
+  assumes "0 < d" "x \<in> s" "\<forall>x'\<in>s. dist x' x < d \<longrightarrow> f x' = g x'" "(f has_derivative f') (at x within s)"
+  shows "(g has_derivative f') (at x within s)"
+  using assms(4) unfolding has_derivative_within apply- apply(erule conjE,rule,assumption)
+  apply(rule Lim_transform_within[OF assms(1)]) defer apply assumption
+  apply(rule,rule) apply(drule assms(3)[rule_format]) using assms(3)[rule_format, OF assms(2)] by auto
+
+lemma has_derivative_transform_at:
+  assumes "0 < d" "\<forall>x'. dist x' x < d \<longrightarrow> f x' = g x'" "(f has_derivative f') (at x)"
+  shows "(g has_derivative f') (at x)"
+  apply(subst within_UNIV[THEN sym]) apply(rule has_derivative_transform_within[OF assms(1)])
+  using assms(2-3) unfolding within_UNIV by auto
+
+lemma has_derivative_transform_within_open:
+  assumes "open s" "x \<in> s" "\<forall>y\<in>s. f y = g y" "(f has_derivative f') (at x)"
+  shows "(g has_derivative f') (at x)"
+  using assms(4) unfolding has_derivative_at apply- apply(erule conjE,rule,assumption)
+  apply(rule Lim_transform_within_open[OF assms(1,2)]) defer apply assumption
+  apply(rule,rule) apply(drule assms(3)[rule_format]) using assms(3)[rule_format, OF assms(2)] by auto
+
+subsection {* differentiability. *}
+
+definition differentiable :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a net \<Rightarrow> bool" (infixr "differentiable" 30) where
+  "f differentiable net \<equiv> (\<exists>f'. (f has_derivative f') net)"
+
+definition differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "differentiable'_on" 30) where
+  "f differentiable_on s \<equiv> (\<forall>x\<in>s. f differentiable (at x within s))"
+
+lemma differentiableI: "(f has_derivative f') net \<Longrightarrow> f differentiable net"
+  unfolding differentiable_def by auto
+
+lemma differentiable_at_withinI: "f differentiable (at x) \<Longrightarrow> f differentiable (at x within s)"
+  unfolding differentiable_def using has_derivative_at_within by blast
+
+lemma differentiable_within_open: assumes "a \<in> s" "open s" shows 
+  "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::finite) 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)
+
+lemma differentiable_transform_within:
+  assumes "0 < d" "x \<in> s" "\<forall>x'\<in>s. dist x' x < d \<longrightarrow> f x' = g x'" "f differentiable (at x within s)"
+  shows "g differentiable (at x within s)"
+  using assms(4) unfolding differentiable_def by(auto intro!: has_derivative_transform_within[OF assms(1-3)])
+
+lemma differentiable_transform_at:
+  assumes "0 < d" "\<forall>x'. dist x' x < d \<longrightarrow> f x' = g x'" "f differentiable at x"
+  shows "g differentiable at x"
+  using assms(3) unfolding differentiable_def using has_derivative_transform_at[OF assms(1-2)] by auto
+
+subsection {* Frechet derivative and Jacobian matrix. *}
+
+definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)"
+
+lemma frechet_derivative_works:
+ "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::finite \<Rightarrow> real^'b::finite"
+  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::finite) \<Rightarrow> (real^'b::finite)) 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 vector_dist_norm diff_0_right norm_mul
+  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)" 
+  shows "continuous (at x within s) f" proof-
+  from assms guess f' unfolding differentiable_def has_derivative_within .. note f'=this
+  then interpret bounded_linear f' by auto
+  have *:"\<And>xa. x\<noteq>xa \<Longrightarrow> (f' \<circ> (\<lambda>y. y - x)) xa + norm (xa - x) *\<^sub>R ((1 / norm (xa - x)) *\<^sub>R (f xa - (f x + f' (xa - x)))) - ((f' \<circ> (\<lambda>y. y - x)) x + 0) = f xa - f x"
+    using zero by auto
+  have **:"continuous (at x within s) (f' \<circ> (\<lambda>y. y - x))"
+    apply(rule continuous_within_compose) apply(rule continuous_intros)+
+    by(rule linear_continuous_within[OF f'[THEN conjunct1]])
+  show ?thesis unfolding continuous_within using f'[THEN conjunct2, THEN Lim_mul_norm_within]
+    apply-apply(drule Lim_add) apply(rule **[unfolded continuous_within]) unfolding Lim_within and vector_dist_norm
+    apply(rule,rule) apply(erule_tac x=e in allE) apply(erule impE|assumption)+ apply(erule exE,rule_tac x=d in exI)
+    by(auto simp add:zero * elim!:allE) qed
+
+lemma differentiable_imp_continuous_at: "f differentiable at x \<Longrightarrow> continuous (at x) f"
+ by(rule differentiable_imp_continuous_within[of _ x UNIV, unfolded within_UNIV])
+
+lemma differentiable_imp_continuous_on: "f differentiable_on s \<Longrightarrow> continuous_on s f"
+  unfolding differentiable_on_def continuous_on_eq_continuous_within
+  using differentiable_imp_continuous_within by blast
+
+lemma has_derivative_within_subset:
+ "(f has_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_derivative f') (at x within t)"
+  unfolding has_derivative_within using Lim_within_subset by blast
+
+lemma differentiable_within_subset:
+  "f differentiable (at x within t) \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f differentiable (at x within s)"
+  unfolding differentiable_def using has_derivative_within_subset by blast
+
+lemma differentiable_on_subset: "f differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f differentiable_on s"
+  unfolding differentiable_on_def using differentiable_within_subset by blast
+
+lemma differentiable_on_empty: "f differentiable_on {}"
+  unfolding differentiable_on_def by auto
+
+subsection {* Several results are easier using a "multiplied-out" variant.              *)
+(* (I got this idea from Dieudonne's proof of the chain rule). *}
+
+lemma has_derivative_within_alt:
+ "(f has_derivative f') (at x within s) \<longleftrightarrow> bounded_linear f' \<and>
+  (\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm(y - x) < d \<longrightarrow> norm(f(y) - f(x) - f'(y - x)) \<le> e * norm(y - x))" (is "?lhs \<longleftrightarrow> ?rhs")
+proof assume ?lhs thus ?rhs unfolding has_derivative_within apply-apply(erule conjE,rule,assumption)
+    unfolding Lim_within apply(rule,erule_tac x=e in allE,rule,erule impE,assumption)
+    apply(erule exE,rule_tac x=d in exI) apply(erule conjE,rule,assumption,rule,rule) proof-
+    fix x y e d assume as:"0 < e" "0 < d" "norm (y - x) < d" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow>
+      dist ((1 / norm (xa - x)) *\<^sub>R (f xa - (f x + f' (xa - x)))) 0 < e" "y \<in> s" "bounded_linear f'"
+    then interpret bounded_linear f' by auto
+    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 vector_dist_norm diff_0_right norm_mul using as(3)
+	using pos_divide_less_eq[OF False[unfolded dist_nz], unfolded vector_dist_norm]
+	by(auto simp add:linear_0 linear_sub group_simps)
+      thus ?thesis by(auto simp add:group_simps) qed qed next
+  assume ?rhs thus ?lhs unfolding has_derivative_within Lim_within apply-apply(erule conjE,rule,assumption)
+    apply(rule,erule_tac x="e/2" in allE,rule,erule impE) defer apply(erule exE,rule_tac x=d in exI)
+    apply(erule conjE,rule,assumption,rule,rule) unfolding vector_dist_norm diff_0_right norm_scaleR
+    apply(erule_tac x=xa in ballE,erule impE) proof-
+    fix e d y assume "bounded_linear f'" "0 < e" "0 < d" "y \<in> s" "0 < norm (y - x) \<and> norm (y - x) < d"
+        "norm (f y - f x - f' (y - x)) \<le> e / 2 * norm (y - x)"
+    thus "\<bar>1 / norm (y - x)\<bar> * norm (f y - (f x + f' (y - x))) < e"
+      apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:group_simps) qed auto qed
+
+lemma has_derivative_at_alt:
+  "(f has_derivative f') (at (x::real^'n::finite)) \<longleftrightarrow> bounded_linear f' \<and>
+  (\<forall>e>0. \<exists>d>0. \<forall>y. norm(y - x) < d \<longrightarrow> norm(f y - f x - f'(y - x)) \<le> e * norm(y - x))"
+  using has_derivative_within_alt[where s=UNIV] unfolding within_UNIV by auto
+
+subsection {* The chain rule. *}
+
+lemma diff_chain_within:
+  assumes "(f has_derivative f') (at x within s)" "(g has_derivative g') (at (f x) within (f ` s))"
+  shows "((g o f) has_derivative (g' o f'))(at x within s)"
+  unfolding has_derivative_within_alt apply(rule,rule bounded_linear_compose[unfolded o_def[THEN sym]])
+  apply(rule assms(2)[unfolded has_derivative_def,THEN conjE],assumption)
+  apply(rule assms(1)[unfolded has_derivative_def,THEN conjE],assumption) proof(rule,rule)
+  note assms = assms[unfolded has_derivative_within_alt]
+  fix e::real assume "0<e"
+  guess B1 using bounded_linear.pos_bounded[OF assms(1)[THEN conjunct1, rule_format]] .. note B1 = this
+  guess B2 using bounded_linear.pos_bounded[OF assms(2)[THEN conjunct1, rule_format]] .. note B2 = this
+  have *:"e / 2 / B2 > 0" using `e>0` B2 apply-apply(rule divide_pos_pos) by auto
+  guess d1 using assms(1)[THEN conjunct2, rule_format, OF *] .. note d1 = this
+  have *:"e / 2 / (1 + B1) > 0" using `e>0` B1 apply-apply(rule divide_pos_pos) by auto
+  guess de using assms(2)[THEN conjunct2, rule_format, OF *] .. note de = this
+  guess d2 using assms(1)[THEN conjunct2, rule_format, OF zero_less_one] .. note d2 = this
+
+  def d0 \<equiv> "(min d1 d2)/2" have d0:"d0>0" "d0 < d1" "d0 < d2" unfolding d0_def using d1 d2 by auto
+  def d \<equiv> "(min d0 (de / (B1 + 1))) / 2" have "de * 2 / (B1 + 1) > de / (B1 + 1)" apply(rule divide_strict_right_mono) using B1 de by auto
+  hence d:"d>0" "d < d1" "d < d2" "d < (de / (B1 + 1))" unfolding d_def using d0 d1 d2 de B1 by(auto intro!: divide_pos_pos simp add:min_less_iff_disj not_less)
+
+  show "\<exists>d>0. \<forall>y\<in>s. norm (y - x) < d \<longrightarrow> norm ((g \<circ> f) y - (g \<circ> f) x - (g' \<circ> f') (y - x)) \<le> e * norm (y - x)" apply(rule_tac x=d in exI)
+    proof(rule,rule `d>0`,rule,rule) 
+    fix y assume as:"y \<in> s" "norm (y - x) < d"
+    hence 1:"norm (f y - f x - f' (y - x)) \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x))" using d1 d2 d by auto
+
+    have "norm (f y - f x) \<le> norm (f y - f x - f' (y - x)) + norm (f' (y - x))"
+      using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:group_simps)
+    also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:group_simps)
+    also have "\<dots> \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x)) + B1 * norm (y - x)" apply(rule add_right_mono) using d1 d2 d as by auto
+    also have "\<dots> \<le> norm (y - x) + B1 * norm (y - x)" by auto
+    also have "\<dots> = norm (y - x) * (1 + B1)" by(auto simp add:field_simps)
+    finally have 3:"norm (f y - f x) \<le> norm (y - x) * (1 + B1)" by auto 
+
+    hence "norm (f y - f x) \<le> d * (1 + B1)" apply- apply(rule order_trans,assumption,rule mult_right_mono) using as B1 by auto 
+    also have "\<dots> < de" using d B1 by(auto simp add:field_simps) 
+    finally have "norm (g (f y) - g (f x) - g' (f y - f x)) \<le> e / 2 / (1 + B1) * norm (f y - f x)"
+      apply-apply(rule de[THEN conjunct2,rule_format]) using `y\<in>s` using d as by auto 
+    also have "\<dots> = (e / 2) * (1 / (1 + B1) * norm (f y - f x))" by auto 
+    also have "\<dots> \<le> e / 2 * norm (y - x)" apply(rule mult_left_mono) using `e>0` and 3 using B1 and `e>0` by(auto simp add:divide_le_eq)
+    finally have 4:"norm (g (f y) - g (f x) - g' (f y - f x)) \<le> e / 2 * norm (y - x)" by auto
+    
+    interpret g': bounded_linear g' using assms(2) by auto
+    interpret f': bounded_linear f' using assms(1) by auto
+    have "norm (- g' (f' (y - x)) + g' (f y - f x)) = norm (g' (f y - f x - f' (y - x)))"
+      by(auto simp add:group_simps f'.diff g'.diff g'.add)
+    also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:group_simps)
+    also have "\<dots> \<le> B2 * (e / 2 / B2 * norm (y - x))" apply(rule mult_left_mono) using as d1 d2 d B2 by auto 
+    also have "\<dots> \<le> e / 2 * norm (y - x)" using B2 by auto
+    finally have 5:"norm (- g' (f' (y - x)) + g' (f y - f x)) \<le> e / 2 * norm (y - x)" by auto
+    
+    have "norm (g (f y) - g (f x) - g' (f y - f x)) + norm (g (f y) - g (f x) - g' (f' (y - x)) - (g (f y) - g (f x) - g' (f y - f x))) \<le> e * norm (y - x)" using 5 4 by auto
+    thus "norm ((g \<circ> f) y - (g \<circ> f) x - (g' \<circ> f') (y - x)) \<le> e * norm (y - x)" unfolding o_def apply- apply(rule order_trans, rule norm_triangle_sub) by assumption qed qed
+
+lemma diff_chain_at:
+  "(f has_derivative f') (at x) \<Longrightarrow> (g has_derivative g') (at (f x)) \<Longrightarrow> ((g o f) has_derivative (g' o f')) (at x)"
+  using diff_chain_within[of f f' x UNIV g g'] using has_derivative_within_subset[of g g' "f x" UNIV "range f"] unfolding within_UNIV by auto
+
+subsection {* Composition rules stated just for differentiability. *}
+
+lemma differentiable_const[intro]: "(\<lambda>z. c) differentiable (net::'a::real_normed_vector net)"
+  unfolding differentiable_def using has_derivative_const by auto
+
+lemma differentiable_id[intro]: "(\<lambda>z. z) differentiable (net::'a::real_normed_vector net)"
+    unfolding differentiable_def using has_derivative_id by auto
+
+lemma differentiable_cmul[intro]: "f differentiable net \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) differentiable (net::'a::real_normed_vector net)"
+  unfolding differentiable_def apply(erule exE, drule has_derivative_cmul) by auto
+
+lemma differentiable_neg[intro]: "f differentiable net \<Longrightarrow> (\<lambda>z. -(f z)) differentiable (net::'a::real_normed_vector net)"
+  unfolding differentiable_def apply(erule exE, drule has_derivative_neg) by auto
+
+lemma differentiable_add: "f differentiable net \<Longrightarrow> g differentiable net
+   \<Longrightarrow> (\<lambda>z. f z + g z) differentiable (net::'a::real_normed_vector net)"
+    unfolding differentiable_def apply(erule exE)+ apply(rule_tac x="\<lambda>z. f' z + f'a z" in exI)
+    apply(rule has_derivative_add) by auto
+
+lemma differentiable_sub: "f differentiable net \<Longrightarrow> g differentiable net
+  \<Longrightarrow> (\<lambda>z. f z - g z) differentiable (net::'a::real_normed_vector net)"
+  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::finite \<Rightarrow>real^'n)"
+  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::finite \<Rightarrow>real^'n)"
+  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
+
+lemma differentiable_chain_at:
+  "f differentiable (at x) \<Longrightarrow> g differentiable (at(f x)) \<Longrightarrow> (g o f) differentiable (at x)"
+  unfolding differentiable_def by(meson diff_chain_at)
+
+lemma differentiable_chain_within:
+  "f differentiable (at x within s) \<Longrightarrow> g differentiable (at(f x) within (f ` s))
+   \<Longrightarrow> (g o f) differentiable (at x within s)"
+  unfolding differentiable_def by(meson diff_chain_within)
+
+subsection {* Uniqueness of derivative.                                                 *)
+(*                                                                           *)
+(* 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::finite \<Rightarrow> real^'b::finite"
+  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-
+  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 vector_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))"
+    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
+    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:group_simps)
+    finally show False using c using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] using norm_basis[of i] unfolding vector_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
+
+lemma frechet_derivative_unique_at: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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
+ 
+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::finite \<Rightarrow> real^'b::finite"
+  assumes "\<forall>i. 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)
+      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)
+      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
+
+lemma frechet_derivative_unique_within_open_interval: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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
+  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
+
+lemma frechet_derivative_at: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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::finite \<Rightarrow> real^'b::finite"
+  assumes "\<forall>i. 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. *}
+
+lemma differential_zero_maxmin_component: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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" 
+    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 dot_rmult dot_basis unfolding smult_conv_scaleR by simp  } 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 vector_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>"]) 
+    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  group_simps by auto qed
+
+subsection {* In particular if we have a mapping into R^1. *}
+
+lemma differential_zero_maxmin: fixes f::"real^'a::finite \<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]
+    unfolding dest_vec1_def[THEN sym] vec1_dest_vec1 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 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 
+
+lemma rolle: fixes f::"real\<Rightarrow>real"
+  assumes "a < b" "f a = f b" "continuous_on {a..b} f"
+  "\<forall>x\<in>{a<..<b}. (f has_derivative f'(x)) (at x)"
+  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
+    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"
+      case False hence "f d = f c" using d c assms(2) by auto
+      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 dest_vec1_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
+
+subsection {* One-dimensional mean value theorem. *}
+
+lemma mvt: fixes f::"real \<Rightarrow> real"
+  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}. (f b - f a = (f' x) (b - a))" proof-
+  have "\<exists>x\<in>{a<..<b}. (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa) = (\<lambda>v. 0)"
+    apply(rule rolle[OF assms(1), of "\<lambda>x. f x - (f b - f a) / (b - a) * x"]) defer
+    apply(rule continuous_on_intros assms(2) continuous_on_cmul[where 'b=real, unfolded real_scaleR_def])+ proof
+    fix x assume x:"x \<in> {a<..<b}"
+    show "((\<lambda>x. f x - (f b - f a) / (b - a) * x) has_derivative (\<lambda>xa. f' x xa - (f b - f a) / (b - a) * xa)) (at x)"
+      by(rule has_derivative_intros assms(3)[rule_format,OF x]
+        has_derivative_cmul[where 'b=real, unfolded real_scaleR_def])+ 
+  qed(insert assms(1), auto simp add:field_simps)
+  then guess x .. thus ?thesis apply(rule_tac x=x in bexI) apply(drule fun_cong[of _ _ "b - a"]) by auto qed
+
+lemma mvt_simple: fixes f::"real \<Rightarrow> real"
+  assumes "a<b"  "\<forall>x\<in>{a..b}. (f has_derivative f' x) (at x within {a..b})"
+  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] 
+    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"
+  assumes "a \<le> b" "\<forall>x\<in>{a..b}. (f has_derivative f'(x)) (at x within {a..b})"
+  shows "\<exists>x\<in>{a..b}. f b - f a = f' x (b - a)" proof(cases "a = b")
+  interpret bounded_linear "f' b" using assms(2) assms(1) by auto
+  case True thus ?thesis apply(rule_tac x=a in bexI)
+    using assms(2)[THEN bspec[where x=a]] unfolding has_derivative_def
+    unfolding True using zero by auto next
+  case False thus ?thesis using mvt_simple[OF _ assms(2)] using assms(1) by auto qed
+
+subsection {* A nice generalization (see Havin's proof of 5.19 from Rudin's book). *}
+
+lemma inner_eq_dot: fixes a::"real^'n::finite"
+  shows "a \<bullet> b = inner a b" unfolding inner_vector_def dot_def by auto
+
+lemma mvt_general: fixes f::"real\<Rightarrow>real^'n::finite"
+  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)"
+    apply(rule mvt) apply(rule assms(1))unfolding inner_eq_dot apply(rule continuous_on_inner continuous_on_intros assms(2))+ 
+    unfolding o_def apply(rule,rule has_derivative_lift_dot) using assms(3) by auto
+  then guess x .. note x=this
+  show ?thesis proof(cases "f a = f b")
+    case False have "norm (f b - f a) * norm (f b - f a) = norm (f b - f a)^2" by(simp add:class_semiring.semiring_rules)
+    also have "\<dots> = (f b - f a) \<bullet> (f b - f a)" unfolding norm_pow_2 ..
+    also have "\<dots> = (f b - f a) \<bullet> f' x (b - a)" using x by auto
+    also have "\<dots> \<le> norm (f b - f a) * norm (f' x (b - a))" by(rule norm_cauchy_schwarz)
+    finally show ?thesis using False x(1) by(auto simp add: real_mult_left_cancel) next
+    case True thus ?thesis using assms(1) apply(rule_tac x="(a + b) /2" in bexI) by auto qed qed
+
+subsection {* Still more general bound theorem. *}
+
+lemma differentiable_bound: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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)"
+  have *:"\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s"
+    using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:group_simps)
+  hence 1:"continuous_on {0..1} (f \<circ> ?p)" apply- apply(rule continuous_on_intros continuous_on_vmul)+
+    unfolding continuous_on_eq_continuous_within apply(rule,rule differentiable_imp_continuous_within)
+    unfolding differentiable_def apply(rule_tac x="f' xa" in exI)
+    apply(rule has_derivative_within_subset) apply(rule assms(2)[rule_format]) by auto
+  have 2:"\<forall>u\<in>{0<..<1}. ((f \<circ> ?p) has_derivative f' (x + u *\<^sub>R (y - x)) \<circ> (\<lambda>u. 0 + u *\<^sub>R (y - x))) (at u)" proof rule case goal1
+    let ?u = "x + u *\<^sub>R (y - x)"
+    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
+  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"
+      using onorm(1)[OF derivative_is_linear[OF assms(2)[rule_format,OF goal1]]] by assumption
+    also have "\<dots> \<le> B * norm y" apply(rule mult_right_mono)
+      using assms(3)[rule_format,OF goal1] by(auto simp add:field_simps)
+    finally show ?case by simp qed
+  have "norm (f x - f y) = norm ((f \<circ> (\<lambda>u. x + u *\<^sub>R (y - x))) 1 - (f \<circ> (\<lambda>u. x + u *\<^sub>R (y - x))) 0)"
+    by(auto simp add:norm_minus_commute) 
+  also have "\<dots> \<le> norm (f' (x + u *\<^sub>R (y - x)) (y - x))" using u by auto
+  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 unfolding norm_vec1 by auto
+  hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by(auto simp add:norm_vec1)
+  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 norm_vec1) qed
+
+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
+ 
+subsection {* In particular. *}
+
+lemma has_derivative_zero_constant: fixes f::"real\<Rightarrow>real"
+  assumes "convex s" "\<forall>x\<in>s. (f has_derivative (\<lambda>h. 0)) (at x within s)"
+  shows "\<exists>c. \<forall>x\<in>s. f x = c" proof(cases "s={}")
+  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
+  thus ?thesis apply(rule_tac x="f x" in exI) by auto qed auto
+
+lemma has_derivative_zero_unique: fixes f::"real\<Rightarrow>real"
+  assumes "convex s" "a \<in> s" "f a = c" "\<forall>x\<in>s. (f has_derivative (\<lambda>h. 0)) (at x within s)" "x\<in>s"
+  shows "f x = c" using has_derivative_zero_constant[OF assms(1,4)] using assms(2-3,5) by auto
+
+subsection {* Differentiability of inverse function (most basic form). *}
+
+lemma has_derivative_inverse_basic: fixes f::"real^'b::finite \<Rightarrow> real^'c::finite"
+  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-
+  interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto
+  interpret g': bounded_linear g' using assms by auto
+  guess C using bounded_linear.pos_bounded[OF assms(2)] .. note C = this
+(*  have fgid:"\<And>x. g' (f' x) = x" using assms(3) unfolding o_def id_def apply()*)
+  have lem1:"\<forall>e>0. \<exists>d>0. \<forall>z. norm(z - y) < d \<longrightarrow> norm(g z - g y - g'(z - y)) \<le> e * norm(g z - g y)" proof(rule,rule) case goal1
+    have *:"e / C > 0" apply(rule divide_pos_pos) using `e>0` C by auto
+    guess d0 using assms(1)[unfolded has_derivative_at_alt,THEN conjunct2,rule_format,OF *] .. note d0=this
+    guess d1 using assms(4)[unfolded continuous_at Lim_at,rule_format,OF d0[THEN conjunct1]] .. note d1=this
+    guess d2 using assms(5)[unfolded open_dist,rule_format,OF assms(6)] .. note d2=this
+    guess d using real_lbound_gt_zero[OF d1[THEN conjunct1] d2[THEN conjunct1]] .. note d=this
+    thus ?case apply(rule_tac x=d in exI) apply rule defer proof(rule,rule)
+      fix z assume as:"norm (z - y) < d" hence "z\<in>t" using d2 d unfolding vector_dist_norm by auto
+      have "norm (g z - g y - g' (z - y)) \<le> norm (g' (f (g z) - y - f' (g z - g y)))"
+        unfolding g'.diff f'.diff unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] 
+	unfolding assms(7)[rule_format,OF `z\<in>t`] apply(subst norm_minus_cancel[THEN sym]) by auto
+      also have "\<dots> \<le> norm(f (g z) - y - f' (g z - g y)) * C" by(rule C[THEN conjunct2,rule_format]) 
+      also have "\<dots> \<le> (e / C) * norm (g z - g y) * C" apply(rule mult_right_mono)
+	apply(rule d0[THEN conjunct2,rule_format,unfolded assms(7)[rule_format,OF `y\<in>t`]]) apply(cases "z=y") defer
+	apply(rule d1[THEN conjunct2, unfolded vector_dist_norm,rule_format]) using as d C d0 by auto
+      also have "\<dots> \<le> e * norm (g z - g y)" using C by(auto simp add:field_simps)
+      finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (g z - g y)" by simp qed auto qed
+  have *:"(0::real) < 1 / 2" by auto guess d using lem1[rule_format,OF *] .. note d=this def B\<equiv>"C*2"
+  have "B>0" unfolding B_def using C by auto
+  have lem2:"\<forall>z. norm(z - y) < d \<longrightarrow> norm(g z - g y) \<le> B * norm(z - y)" proof(rule,rule) case goal1
+    have "norm (g z - g y) \<le> norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" by(rule norm_triangle_sub)
+    also have "\<dots> \<le> norm(g' (z - y)) + 1 / 2 * norm (g z - g y)" apply(rule add_left_mono) using d and goal1 by auto
+    also have "\<dots> \<le> norm (z - y) * C + 1 / 2 * norm (g z - g y)" apply(rule add_right_mono) using C by auto
+    finally show ?case unfolding B_def by(auto simp add:field_simps) qed
+  show ?thesis unfolding has_derivative_at_alt proof(rule,rule assms,rule,rule) case goal1
+    hence *:"e/B >0" apply-apply(rule divide_pos_pos) using `B>0` by auto
+    guess d' using lem1[rule_format,OF *] .. note d'=this
+    guess k using real_lbound_gt_zero[OF d[THEN conjunct1] d'[THEN conjunct1]] .. note k=this
+    show ?case apply(rule_tac x=k in exI,rule) defer proof(rule,rule) fix z assume as:"norm(z - y) < k"
+      hence "norm (g z - g y - g' (z - y)) \<le> e / B * norm(g z - g y)" using d' k by auto
+      also have "\<dots> \<le> e * norm(z - y)" unfolding mult_frac_num pos_divide_le_eq[OF `B>0`]
+	using lem2[THEN spec[where x=z]] using k as using `e>0` by(auto simp add:field_simps)
+      finally show "norm (g z - g y - g' (z - y)) \<le> e * norm (z - y)" by simp qed(insert k, auto) qed qed
+
+subsection {* Simply rewrite that based on the domain point x. *}
+
+lemma has_derivative_inverse_basic_x: fixes f::"real^'b::finite \<Rightarrow> real^'c::finite"
+  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)))"
+  apply(rule has_derivative_inverse_basic) using assms by auto
+
+subsection {* This is the version in Dieudonne', assuming continuity of f and g. *}
+
+lemma has_derivative_inverse_dieudonne: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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))"
+  apply(rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)])
+  using assms(3-6) unfolding continuous_on_eq_continuous_at[OF assms(1)]  continuous_on_eq_continuous_at[OF assms(2)] by auto
+
+subsection {* Here's the simplest way of not assuming much about g. *}
+
+lemma has_derivative_inverse: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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-
+  { fix y assume "y\<in>interior (f ` s)" 
+    then obtain x where "x\<in>s" and *:"y = f x" unfolding image_iff using interior_subset by auto
+    have "f (g y) = y" unfolding * and assms(5)[rule_format,OF `x\<in>s`] .. } note * = this
+  show ?thesis apply(rule has_derivative_inverse_basic_x[OF assms(6-8)])
+    apply(rule continuous_on_interior[OF _ assms(3)]) apply(rule continuous_on_inverse[OF assms(4,1)])
+    apply(rule assms(2,5) assms(5)[rule_format] open_interior assms(3))+ by(rule, rule *, assumption)  qed
+
+subsection {* Proving surjectivity via Brouwer fixpoint theorem. *}
+
+lemma brouwer_surjective: fixes f::"real^'n::finite \<Rightarrow> real^'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-
+  have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:group_simps)
+  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::finite \<Rightarrow> real^'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)+
+  unfolding cball_eq_empty using assms by auto 
+
+text {* See Sussmann: "Multidifferential calculus", Theorem 2.1.1 *}
+
+lemma sussmann_open_mapping: fixes f::"real^'a::finite \<Rightarrow> real^'b::finite"
+  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"
+  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
+  guess B using bounded_linear.pos_bounded[OF assms(5)] .. note B=this hence *:"1/(2*B)>0" by(auto intro!: divide_pos_pos)
+  guess e0 using assms(4)[unfolded has_derivative_at_alt,THEN conjunct2,rule_format,OF *] .. note e0=this
+  guess e1 using assms(8)[unfolded mem_interior_cball] .. note e1=this
+  have *:"0<e0/B" "0<e1/B" apply(rule_tac[!] divide_pos_pos) using e0 e1 B by auto
+  guess e using real_lbound_gt_zero[OF *] .. note e=this
+  have "\<forall>z\<in>cball (f x) (e/2). \<exists>y\<in>cball (f x) e. f (x + g' (y - f x)) = z"
+    apply(rule,rule brouwer_surjective_cball[where s="cball (f x) (e/2)"])
+    prefer 3 apply(rule,rule) proof- 
+    show "continuous_on (cball (f x) e) (\<lambda>y. f (x + g' (y - f x)))" unfolding g'.diff
+      apply(rule continuous_on_compose[of _ _ f, unfolded o_def])
+      apply(rule continuous_on_intros linear_continuous_on[OF assms(5)])+
+      apply(rule continuous_on_subset[OF assms(2)]) apply(rule,unfold image_iff,erule bexE) proof-
+      fix y z assume as:"y \<in>cball (f x) e"  "z = x + (g' y - g' (f x))"
+      have "dist x z = norm (g' (f x) - g' y)" unfolding as(2) and vector_dist_norm by auto
+      also have "\<dots> \<le> norm (f x - y) * B" unfolding g'.diff[THEN sym] using B by auto
+      also have "\<dots> \<le> e * B" using as(1)[unfolded mem_cball vector_dist_norm] using B by auto
+      also have "\<dots> \<le> e1" using e unfolding less_divide_eq using B by auto
+      finally have "z\<in>cball x e1" unfolding mem_cball by force
+      thus "z \<in> s" using e1 assms(7) by auto qed next
+    fix y z assume as:"y \<in> cball (f x) (e / 2)" "z \<in> cball (f x) e"
+    have "norm (g' (z - f x)) \<le> norm (z - f x) * B" using B by auto
+    also have "\<dots> \<le> e * B" apply(rule mult_right_mono) using as(2)[unfolded mem_cball vector_dist_norm] and B unfolding norm_minus_commute by auto
+    also have "\<dots> < e0" using e and B unfolding less_divide_eq by auto
+    finally have *:"norm (x + g' (z - f x) - x) < e0" by auto
+    have **:"f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto
+    have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> norm (f (x + g' (z - f x)) - z) + norm (f x - y)"
+      using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:group_simps)
+    also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding group_simps ** by auto 
+    also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" using as(1)[unfolded mem_cball vector_dist_norm] by auto
+    also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" using * and B by(auto simp add:field_simps)
+    also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" by auto
+    also have "\<dots> \<le> e/2 + e/2" apply(rule add_right_mono) using as(2)[unfolded mem_cball vector_dist_norm] unfolding norm_minus_commute by auto
+    finally show "y + (z - f (x + g' (z - f x))) \<in> cball (f x) e" unfolding mem_cball vector_dist_norm by auto
+  qed(insert e, auto) note lem = this
+  show ?thesis unfolding mem_interior apply(rule_tac x="e/2" in exI)
+    apply(rule,rule divide_pos_pos) prefer 3 proof 
+    fix y assume "y \<in> ball (f x) (e/2)" hence *:"y\<in>cball (f x) (e/2)" by auto
+    guess z using lem[rule_format,OF *] .. note z=this
+    hence "norm (g' (z - f x)) \<le> norm (z - f x) * B" using B by(auto simp add:field_simps)
+    also have "\<dots> \<le> e * B" apply(rule mult_right_mono) using z(1) unfolding mem_cball vector_dist_norm norm_minus_commute using B by auto
+    also have "\<dots> \<le> e1"  using e B unfolding less_divide_eq by auto
+    finally have "x + g'(z - f x) \<in> t" apply- apply(rule e1[THEN conjunct2,unfolded subset_eq,rule_format]) 
+      unfolding mem_cball vector_dist_norm by auto
+    thus "y \<in> f ` t" using z by auto qed(insert e, auto) qed
+
+text {* Hence the following eccentric variant of the inverse function theorem.    *)
+(* This has no continuity assumptions, but we do need the inverse function.  *)
+(* 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::finite \<Rightarrow> real^'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-
+  have linf:"bounded_linear f'" using assms(5) unfolding has_derivative_def by auto
+  hence ling:"bounded_linear g'" unfolding linear_conv_bounded_linear[THEN sym]
+    apply- apply(rule right_inverse_linear) using assms(6) by auto 
+  moreover have "g' \<circ> f' = id" using assms(6) linf ling unfolding linear_conv_bounded_linear[THEN sym]
+    using linear_inverse_left by auto
+  moreover have *:"\<forall>t\<subseteq>s. x\<in>interior t \<longrightarrow> f x \<in> interior (f ` t)" apply(rule,rule,rule,rule sussmann_open_mapping )
+    apply(rule assms ling)+ by auto
+  have "continuous (at (f x)) g" unfolding continuous_at Lim_at proof(rule,rule)
+    fix e::real assume "e>0"
+    hence "f x \<in> interior (f ` (ball x e \<inter> s))" using *[rule_format,of "ball x e \<inter> s"] `x\<in>s`
+      by(auto simp add: interior_open[OF open_ball] interior_open[OF assms(1)])
+    then guess d unfolding mem_interior .. note d=this
+    show "\<exists>d>0. \<forall>y. 0 < dist y (f x) \<and> dist y (f x) < d \<longrightarrow> dist (g y) (g (f x)) < e"
+      apply(rule_tac x=d in exI) apply(rule,rule d[THEN conjunct1]) proof(rule,rule) case goal1
+      hence "g y \<in> g ` f ` (ball x e \<inter> s)" using d[THEN conjunct2,unfolded subset_eq,THEN bspec[where x=y]]
+	by(auto simp add:dist_commute)
+      hence "g y \<in> ball x e \<inter> s" using assms(4) by auto
+      thus "dist (g y) (g (f x)) < e" using assms(4)[rule_format,OF `x\<in>s`] by(auto simp add:dist_commute) qed qed
+  moreover have "f x \<in> interior (f ` s)" apply(rule sussmann_open_mapping)
+    apply(rule assms ling)+ using interior_open[OF assms(1)] and `x\<in>s` by auto
+  moreover have "\<And>y. y \<in> interior (f ` s) \<Longrightarrow> f (g y) = y" proof- case goal1
+    hence "y\<in>f ` s" using interior_subset by auto then guess z unfolding image_iff ..
+    thus ?case using assms(4) by auto qed
+  ultimately show ?thesis apply- apply(rule has_derivative_inverse_basic_x[OF assms(5)]) using assms by auto qed 
+
+subsection {* A rewrite based on the other domain. *}
+
+lemma has_derivative_inverse_strong_x: fixes f::"real^'n::finite \<Rightarrow> real^'n"
+  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)"
+  using has_derivative_inverse_strong[OF assms(1-6)] unfolding assms(7) by simp
+
+subsection {* On a region. *}
+
+lemma has_derivative_inverse_on: fixes f::"real^'n::finite \<Rightarrow> real^'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)+
+  unfolding continuous_on_eq_continuous_at[OF assms(1)]
+  apply(rule,rule differentiable_imp_continuous_at) unfolding differentiable_def using assms by auto
+
+subsection {* Invertible derivative continous at a point implies local injectivity.     *)
+(* It's only for this we need continuity of the derivative, except of course *)
+(* if we want the fact that the inverse derivative is also continuous. So if *)
+(* we know for some other reason that the inverse function exists, it's OK. *}
+
+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:group_simps)
+
+lemma has_derivative_locally_injective: fixes f::"real^'n::finite \<Rightarrow> real^'m::finite"
+  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
+  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
+  from `open s` obtain d2 where "d2>0" "ball a d2 \<subseteq> s" using `a\<in>s` ..
+  obtain d2 where "d2>0" "ball a d2 \<subseteq> s" using assms(2,1) ..
+  guess d2 using assms(2)[unfolded open_contains_ball,rule_format,OF `a\<in>s`] .. note d2=this
+  guess d using real_lbound_gt_zero[OF d1[THEN conjunct1] d2[THEN conjunct1]] .. note d = this
+  show ?thesis proof show "a\<in>ball a d" using d by auto
+    show "\<forall>x\<in>ball a d. \<forall>x'\<in>ball a d. f x' = f x \<longrightarrow> x' = x" proof(intro strip)
+      fix x y assume as:"x\<in>ball a d" "y\<in>ball a d" "f x = f y"
+      def ph \<equiv> "\<lambda>w. w - g'(f w - f x)" have ph':"ph = g' \<circ> (\<lambda>w. f' a w - (f w - f x))"
+	unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:group_simps)
+      have "norm (ph x - ph y) \<le> (1/2) * norm (x - y)"
+	apply(rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"])
+	apply(rule_tac[!] ballI) proof- fix u assume u:"u \<in> ball a d" hence "u\<in>s" using d d2 by auto
+	have *:"(\<lambda>v. v - g' (f' u v)) = g' \<circ> (\<lambda>w. f' a w - f' u w)" unfolding o_def and diff using f'g' by auto
+	show "(ph has_derivative (\<lambda>v. v - g' (f' u v))) (at u within ball a d)"
+	  unfolding ph' * apply(rule diff_chain_within) defer apply(rule bounded_linear.has_derivative[OF assms(3)])
+	  apply(rule has_derivative_intros) defer apply(rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right])
+	  apply(rule has_derivative_at_within) using assms(5) and `u\<in>s` `a\<in>s`
+	  by(auto intro!: has_derivative_intros derivative_linear)
+	have **:"bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)" apply(rule_tac[!] bounded_linear_sub)
+	  apply(rule_tac[!] derivative_linear) using assms(5) `u\<in>s` `a\<in>s` by auto
+	have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)" unfolding * apply(rule onorm_compose)
+	  unfolding linear_conv_bounded_linear by(rule assms(3) **)+ 
+	also have "\<dots> \<le> onorm g' * k" apply(rule mult_left_mono) 
+	  using d1[THEN conjunct2,rule_format,of u] using onorm_neg[OF **(1)[unfolded linear_linear]]
+	  using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:group_simps) 
+	also have "\<dots> \<le> 1/2" unfolding k_def by auto
+	finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" by assumption qed
+      moreover have "norm (ph y - ph x) = norm (y - x)" apply(rule arg_cong[where f=norm])
+	unfolding ph_def using diff unfolding as by auto
+      ultimately show "x = y" unfolding norm_minus_commute by auto qed qed auto qed
+
+subsection {* Uniformly convergent sequence of derivatives. *}
+
+lemma has_derivative_sequence_lipschitz_lemma: fixes f::"nat \<Rightarrow> real^'m::finite \<Rightarrow> real^'n::finite"
+  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)+ 
+  fix m n x y assume as:"N\<le>m" "N\<le>n" "x\<in>s" "y\<in>s"
+  show "norm((f m x - f n x) - (f m y - f n y)) \<le> 2 * e * norm(x - y)"
+    apply(rule differentiable_bound[where f'="\<lambda>x h. f' m x h - f' n x h", OF assms(1) _ _ as(3-4)]) apply(rule_tac[!] ballI) proof-
+    fix x assume "x\<in>s" show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)"
+      by(rule has_derivative_intros assms(2)[rule_format] `x\<in>s`)+
+    { fix h have "norm (f' m x h - f' n x h) \<le> norm (f' m x h - g' x h) + norm (f' n x h - g' x h)"
+	using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:group_simps) 
+      also have "\<dots> \<le> e * norm h+ e * norm h"  using assms(3)[rule_format,OF `N\<le>m` `x\<in>s`, of h] assms(3)[rule_format,OF `N\<le>n` `x\<in>s`, of h]
+	by(auto simp add:field_simps)
+      finally have "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" by auto }
+    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::finite \<Rightarrow> real^'n::finite"
+  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)
+  case goal1 have *:"2 * (1/2* e) = e" "1/2 * e >0" using `e>0` by auto
+  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::finite\<Rightarrow>real^'n::finite"
+  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"
+  shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) ---> g x) sequentially \<and> (g has_derivative g'(x)) (at x within s)" proof-
+  have lem1:"\<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)"
+    apply(rule has_derivative_sequence_lipschitz[where e="42::nat"]) apply(rule assms)+ by auto
+  have "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) ---> g x) sequentially" apply(rule bchoice) unfolding convergent_eq_cauchy proof
+    fix x assume "x\<in>s" show "Cauchy (\<lambda>n. f n x)" proof(cases "x=x0")
+      case True thus ?thesis using convergent_imp_cauchy[OF assms(5)] by auto next
+      case False show ?thesis unfolding Cauchy_def proof(rule,rule)
+	fix e::real assume "e>0" hence *:"e/2>0" "e/2/norm(x-x0)>0" using False by(auto intro!:divide_pos_pos)
+	guess M using convergent_imp_cauchy[OF assms(5), unfolded Cauchy_def, rule_format,OF *(1)] .. note M=this
+	guess N using lem1[rule_format,OF *(2)] .. note N = this
+	show " \<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e" apply(rule_tac x="max M N" in exI) proof(default+)
+	  fix m n assume as:"max M N \<le>m" "max M N\<le>n"
+	  have "dist (f m x) (f n x) \<le> norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))"
+	    unfolding vector_dist_norm by(rule norm_triangle_sub)
+	  also have "\<dots> \<le> norm (f m x0 - f n x0) + e / 2" using N[rule_format,OF _ _ `x\<in>s` `x0\<in>s`, of m n] and as and False by auto
+	  also have "\<dots> < e / 2 + e / 2" apply(rule add_strict_right_mono) using as and M[rule_format] unfolding vector_dist_norm by auto 
+	  finally show "dist (f m x) (f n x) < e" by auto qed qed qed qed
+  then guess g .. note g = this
+  have lem2:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm((f n x - f n y) - (g x - g y)) \<le> e * norm(x - y)" proof(rule,rule)
+    fix e::real assume *:"e>0" guess N using lem1[rule_format,OF *] .. note N=this
+    show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>y\<in>s. norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" apply(rule_tac x=N in exI) proof(default+)
+      fix n x y assume as:"N \<le> n" "x \<in> s" "y \<in> s"
+      have "eventually (\<lambda>xa. norm (f n x - f n y - (f xa x - f xa y)) \<le> e * norm (x - y)) sequentially" 
+	unfolding eventually_sequentially apply(rule_tac x=N in exI) proof(rule,rule)
+	fix m assume "N\<le>m" thus "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)"
+	  using N[rule_format, of n m x y] and as by(auto simp add:group_simps) qed
+      thus "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" apply-
+	apply(rule Lim_norm_ubound[OF trivial_limit_sequentially, where f="\<lambda>m. (f n x - f n y) - (f m x - f m y)"])
+	apply(rule Lim_sub Lim_const g[rule_format] as)+ by assumption qed qed
+  show ?thesis unfolding has_derivative_within_alt apply(rule_tac x=g in exI)
+    apply(rule,rule,rule g[rule_format],assumption) proof fix x assume "x\<in>s"
+    have lem3:"\<forall>u. ((\<lambda>n. f' n x u) ---> g' x u) sequentially" unfolding Lim_sequentially proof(rule,rule,rule)
+      fix u and e::real assume "e>0" show "\<exists>N. \<forall>n\<ge>N. dist (f' n x u) (g' x u) < e" proof(cases "u=0")
+	case True guess N using assms(3)[rule_format,OF `e>0`] .. note N=this
+	show ?thesis apply(rule_tac x=N in exI) unfolding True 
+	  using N[rule_format,OF _ `x\<in>s`,of _ 0] and `e>0` by auto next
+	case False hence *:"e / 2 / norm u > 0" using `e>0` by(auto intro!: divide_pos_pos)
+	guess N using assms(3)[rule_format,OF *] .. note N=this
+	show ?thesis apply(rule_tac x=N in exI) proof(rule,rule) case goal1
+	  show ?case unfolding vector_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
+      note lin = assms(2)[rule_format,OF `x\<in>s`,THEN derivative_linear]
+      show "g' x (c *s x') = c *s g' x x'" apply(rule Lim_unique[OF trivial_limit_sequentially])
+	apply(rule lem3[rule_format]) unfolding smult_conv_scaleR 
+        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])
+	apply(rule lem3[rule_format]) unfolding lin[unfolded bounded_linear_def additive_def,THEN conjunct1,rule_format]
+        apply(rule Lim_add) by(rule lem3[rule_format])+ qed 
+    show "\<forall>e>0. \<exists>d>0. \<forall>y\<in>s. norm (y - x) < d \<longrightarrow> norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)" proof(rule,rule) case goal1
+      have *:"e/3>0" using goal1 by auto guess N1 using assms(3)[rule_format,OF *] .. note N1=this
+      guess N2 using lem2[rule_format,OF *] .. note N2=this
+      guess d1 using assms(2)[unfolded has_derivative_within_alt, rule_format,OF `x\<in>s`, of "max N1 N2",THEN conjunct2,rule_format,OF *] .. note d1=this
+      show ?case apply(rule_tac x=d1 in exI) apply(rule,rule d1[THEN conjunct1]) proof(rule,rule)
+	fix y assume as:"y \<in> s" "norm (y - x) < d1" let ?N ="max N1 N2"
+	have "norm (g y - g x - (f ?N y - f ?N x)) \<le> e /3 * norm (y - x)" apply(subst norm_minus_cancel[THEN sym])
+	  using N2[rule_format, OF _ `y\<in>s` `x\<in>s`, of ?N] by auto moreover
+	have "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" using d1 and as by auto ultimately
+	have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)" 
+	  using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] 
+	  by (auto simp add:group_simps) moreover
+	have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)" using N1 `x\<in>s` by auto
+	ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
+	  using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:group_simps)
+	qed qed qed qed
+
+subsection {* Can choose to line up antiderivatives if we want. *}
+
+lemma has_antiderivative_sequence: fixes f::"nat\<Rightarrow> real^'m::finite \<Rightarrow> real^'n::finite"
+  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={}")
+  case False then obtain a where "a\<in>s" by auto have *:"\<And>P Q. \<exists>g. \<forall>x\<in>s. P g x \<and> Q g x \<Longrightarrow> \<exists>g. \<forall>x\<in>s. Q g x" by auto
+  show ?thesis  apply(rule *) apply(rule has_derivative_sequence[OF assms(1) _ assms(3), of "\<lambda>n x. f n x + (f 0 a - f n a)"])
+    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::finite \<Rightarrow> real^'m::finite \<Rightarrow> real^'n::finite"
+  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))"
+    apply(rule) using assms(2) apply(erule_tac x="inverse (real (Suc n))" in allE) by auto
+  guess f using *[THEN choice] .. note * = this guess f' using *[THEN choice] .. note f=this
+  show ?thesis apply(rule has_antiderivative_sequence[OF assms(1), of f f']) defer proof(rule,rule)
+    fix e::real assume "0<e" guess  N using reals_Archimedean[OF `e>0`] .. note N=this 
+    show "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (f' n x h - g' x h) \<le> e * norm h"  apply(rule_tac x=N in exI) proof(default+) case goal1
+      have *:"inverse (real (Suc n)) \<le> e" apply(rule order_trans[OF _ N[THEN less_imp_le]])
+	using goal1(1) by(auto simp add:field_simps) 
+      show ?case using f[rule_format,THEN conjunct2,OF goal1(2), of n, THEN spec[where x=h]] 
+	apply(rule order_trans) using N * apply(cases "h=0") by auto qed qed(insert f,auto) qed
+
+subsection {* Differentiation of a series. *}
+
+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::finite \<Rightarrow> real^'n::finite"
+  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"
+  shows "\<exists>g. \<forall>x\<in>s. ((\<lambda>n. f n x) sums_seq (g x)) k \<and> (g has_derivative g'(x)) (at x within s)"
+  unfolding sums_seq_def apply(rule has_derivative_sequence[OF assms(1) _ assms(3)]) apply(rule,rule)
+  apply(rule has_derivative_setsum) defer apply(rule,rule assms(2)[rule_format],assumption)
+  using assms(4-5) unfolding sums_seq_def by auto
+
+subsection {* Derivative with composed bilinear function. *}
+
+lemma has_derivative_bilinear_within: fixes h::"real^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real^'p::finite" and f::"real^'q::finite \<Rightarrow> real^'m"
+  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])
+    using assms(2) unfolding differentiable_def by auto moreover
+  interpret f':bounded_linear f' using assms unfolding has_derivative_def by auto
+  interpret g':bounded_linear g' using assms unfolding has_derivative_def by auto
+  interpret h:bounded_bilinear h using assms by auto
+  have "((\<lambda>y. f' (y - x)) ---> 0) (at x within s)" unfolding f'.zero[THEN sym]
+    apply(rule Lim_linear[of "\<lambda>y. y - x" 0 "at x within s" f']) using Lim_sub[OF Lim_within_id Lim_const, of x x s]
+    unfolding id_def using assms(1) unfolding has_derivative_def by auto
+  hence "((\<lambda>y. f x + f' (y - x)) ---> f x) (at x within s)"
+    using Lim_add[OF Lim_const, of "\<lambda>y. f' (y - x)" 0 "at x within s" "f x"] by auto ultimately
+  have *:"((\<lambda>x'. h (f x + f' (x' - x)) ((1/(norm (x' - x))) *\<^sub>R (g x' - (g x + g' (x' - x))))
+             + h ((1/ (norm (x' - x))) *\<^sub>R (f x' - (f x + f' (x' - x)))) (g x')) ---> h (f x) 0 + h 0 (g x)) (at x within s)"
+    apply-apply(rule Lim_add) apply(rule_tac[!] Lim_bilinear[OF _ _ assms(3)]) using assms(1-2)  unfolding has_derivative_within by auto
+  guess B using bounded_bilinear.pos_bounded[OF assms(3)] .. note B=this
+  guess C using f'.pos_bounded .. note C=this
+  guess D using g'.pos_bounded .. note D=this
+  have bcd:"B * C * D > 0" using B C D by (auto intro!: mult_pos_pos)
+  have **:"((\<lambda>y. (1/(norm(y - x))) *\<^sub>R (h (f'(y - x)) (g'(y - x)))) ---> 0) (at x within s)" unfolding Lim_within proof(rule,rule) case goal1
+    hence "e/(B*C*D)>0" using B C D by(auto intro!:divide_pos_pos mult_pos_pos)
+    thus ?case apply(rule_tac x="e/(B*C*D)" in exI,rule) proof(rule,rule,erule conjE)
+      fix y assume as:"y \<in> s" "0 < dist y x" "dist y x < e / (B * C * D)"
+      have "norm (h (f' (y - x)) (g' (y - x))) \<le> norm (f' (y - x)) * norm (g' (y - x)) * B" using B by auto
+      also have "\<dots> \<le> (norm (y - x) * C) * (D * norm (y - x)) * B" apply(rule mult_right_mono)
+	apply(rule pordered_semiring_class.mult_mono) using B C D by (auto simp add: field_simps intro!:mult_nonneg_nonneg)
+      also have "\<dots> = (B * C * D * norm (y - x)) * norm (y - x)" by(auto simp add:field_simps)
+      also have "\<dots> < e * norm (y - x)" apply(rule mult_strict_right_mono)
+	using as(3)[unfolded vector_dist_norm] and as(2) unfolding pos_less_divide_eq[OF bcd] by (auto simp add:field_simps)
+      finally show "dist ((1 / norm (y - x)) *\<^sub>R h (f' (y - x)) (g' (y - x))) 0 < e"
+	unfolding vector_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 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
+     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::finite \<Rightarrow> real^'n::finite \<Rightarrow> real^'p::finite" and f::"real^'p::finite \<Rightarrow> real^'m"
+  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 R(^1)->R^n as a vector. *}
+
+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"
+
+definition "vector_derivative f net \<equiv> (SOME f'. (f has_vector_derivative f') net)"
+
+lemma vector_derivative_works: fixes f::"real \<Rightarrow> 'a::real_normed_vector"
+  shows "f differentiable net \<longleftrightarrow> (f has_vector_derivative (vector_derivative f net)) net" (is "?l = ?r")
+proof assume ?l guess f' using `?l`[unfolded differentiable_def] .. note f' = this
+  then interpret bounded_linear f' by auto
+  thus ?r unfolding vector_derivative_def has_vector_derivative_def
+    apply-apply(rule someI_ex,rule_tac x="f' 1" in exI)
+    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::finite"
+  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
+  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
+
+lemma vector_derivative_unique_within_closed_interval: fixes f::"real \<Rightarrow> real^'n::finite"
+  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
+  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
+
+lemma vector_derivative_at: fixes f::"real \<Rightarrow> real^'a::finite" 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::finite"
+  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)
+  using vector_derivative_works[unfolded differentiable_def]
+  using assms by(auto simp add:has_vector_derivative_def)
+
+lemma has_vector_derivative_within_subset: fixes f::"real \<Rightarrow> real^'a::finite" shows
+ "(f has_vector_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_vector_derivative f') (at x within t)"
+  unfolding has_vector_derivative_def apply(rule has_derivative_within_subset) by auto
+
+lemma has_vector_derivative_const: fixes c::"real^'n::finite" shows
+ "((\<lambda>x. c) has_vector_derivative 0) net"
+  unfolding has_vector_derivative_def using has_derivative_const by auto
+
+lemma has_vector_derivative_id: "((\<lambda>x::real. x) has_vector_derivative 1) net"
+  unfolding has_vector_derivative_def using has_derivative_id by auto
+
+lemma has_vector_derivative_cmul: fixes f::"real \<Rightarrow> real^'a::finite"
+  shows "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net"
+  unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:group_simps)
+
+lemma has_vector_derivative_cmul_eq: fixes f::"real \<Rightarrow> real^'a::finite" assumes "c \<noteq> 0"
+  shows "(((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net \<longleftrightarrow> (f has_vector_derivative f') net)"
+  apply rule apply(drule has_vector_derivative_cmul[where c="1/c"]) defer
+  apply(rule has_vector_derivative_cmul) using assms by auto
+
+lemma has_vector_derivative_neg:
+ "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. -(f x)) has_vector_derivative (- f')) net"
+  unfolding has_vector_derivative_def apply(drule has_derivative_neg) by auto
+
+lemma has_vector_derivative_add:
+  assumes "(f has_vector_derivative f') net" "(g has_vector_derivative g') net"
+  shows "((\<lambda>x. f(x) + g(x)) has_vector_derivative (f' + g')) net"
+  using has_derivative_add[OF assms[unfolded has_vector_derivative_def]]
+  unfolding has_vector_derivative_def unfolding scaleR_right_distrib by auto
+
+lemma has_vector_derivative_sub:
+  assumes "(f has_vector_derivative f') net" "(g has_vector_derivative g') net"
+  shows "((\<lambda>x. f(x) - g(x)) has_vector_derivative (f' - g')) net"
+  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::finite \<Rightarrow> real^'n::finite \<Rightarrow> real^'p::finite"
+  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)"]
+    using assms(3) unfolding scaleR_right scaleR_left scaleR_right_distrib by auto qed
+
+lemma has_vector_derivative_bilinear_at: fixes h::"real^'m::finite \<Rightarrow> real^'n::finite \<Rightarrow> real^'p::finite"
+  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
+
+lemma has_vector_derivative_at_within: "(f has_vector_derivative f') (at x) \<Longrightarrow> (f has_vector_derivative f') (at x within s)"
+  unfolding has_vector_derivative_def apply(rule has_derivative_at_within) by auto
+
+lemma has_vector_derivative_transform_within:
+  assumes "0 < d" "x \<in> s" "\<forall>x'\<in>s. dist x' x < d \<longrightarrow> f x' = g x'" "(f has_vector_derivative f') (at x within s)"
+  shows "(g has_vector_derivative f') (at x within s)"
+  using assms unfolding has_vector_derivative_def by(rule has_derivative_transform_within)
+
+lemma has_vector_derivative_transform_at:
+  assumes "0 < d" "\<forall>x'. dist x' x < d \<longrightarrow> f x' = g x'" "(f has_vector_derivative f') (at x)"
+  shows "(g has_vector_derivative f') (at x)"
+  using assms unfolding has_vector_derivative_def by(rule has_derivative_transform_at)
+
+lemma has_vector_derivative_transform_within_open:
+  assumes "open s" "x \<in> s" "\<forall>y\<in>s. f y = g y" "(f has_vector_derivative f') (at x)"
+  shows "(g has_vector_derivative f') (at x)"
+  using assms unfolding has_vector_derivative_def by(rule has_derivative_transform_within_open)
+
+lemma vector_diff_chain_at:
+  assumes "(f has_vector_derivative f') (at x)" "(g has_vector_derivative g') (at (f x))"
+  shows "((g \<circ> f) has_vector_derivative (f' *\<^sub>R g')) (at x)"
+  using assms(2) unfolding has_vector_derivative_def apply- apply(drule diff_chain_at[OF assms(1)[unfolded has_vector_derivative_def]])
+  unfolding o_def scaleR.scaleR_left by auto
+
+lemma vector_diff_chain_within:
+  assumes "(f has_vector_derivative f') (at x within s)" "(g has_vector_derivative g') (at (f x) within f ` s)"
+  shows "((g o f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)"
+  using assms(2) unfolding has_vector_derivative_def apply- apply(drule diff_chain_within[OF assms(1)[unfolded has_vector_derivative_def]])
+  unfolding o_def scaleR.scaleR_left by auto
+
+end
+
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -94,7 +94,7 @@
 
 instantiation "^" :: (ord,type) ord
  begin
-definition vector_less_eq_def:
+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)"
 
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -1,5 +1,5 @@
 theory Multivariate_Analysis
-imports Convex_Euclidean_Space Determinants
+imports Determinants Derivative
 begin
 
 end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -4619,17 +4619,17 @@
 lemma interval: fixes a :: "'a::ord^'n::finite" shows
   "{a <..< b} = {x::'a^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}" and
   "{a .. b} = {x::'a^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
-  by (auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
+  by (auto simp add: expand_set_eq vector_less_def vector_le_def)
 
 lemma mem_interval: fixes a :: "'a::ord^'n::finite" shows
   "x \<in> {a<..<b} \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
   "x \<in> {a .. b} \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
-  using interval[of a b] by(auto simp add: expand_set_eq vector_less_def vector_less_eq_def)
+  using interval[of a b] by(auto simp add: expand_set_eq vector_less_def vector_le_def)
 
 lemma mem_interval_1: fixes x :: "real^1" shows
  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
  "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp_all add: Cart_eq vector_less_def vector_less_eq_def dest_vec1_def forall_1)
+by(simp_all add: Cart_eq vector_less_def vector_le_def dest_vec1_def forall_1)
 
 lemma vec1_interval:fixes a::"real" shows
   "vec1 ` {a .. b} = {vec1 a .. vec1 b}"
@@ -4690,7 +4690,7 @@
 
 lemma interval_sing: fixes a :: "'a::linorder^'n::finite" shows
  "{a .. a} = {a} \<and> {a<..<a} = {}"
-apply(auto simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+apply(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
@@ -4703,17 +4703,17 @@
   { fix i
     have "a $ i \<le> x $ i"
       using x order_less_imp_le[of "a$i" "x$i"]
-      by(simp add: expand_set_eq vector_less_def vector_less_eq_def Cart_eq)
+      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_less_eq_def Cart_eq)
+      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_less_eq_def Cart_eq)
+    by(simp add: expand_set_eq vector_less_def vector_le_def Cart_eq)
 qed
 
 lemma subset_interval: fixes a :: "real^'n::finite" shows
@@ -5026,12 +5026,12 @@
 
 lemma interval_cases_1: fixes x :: "real^1" shows
  "x \<in> {a .. b} ==> x \<in> {a<..<b} \<or> (x = a) \<or> (x = b)"
-  by(simp add:  Cart_eq vector_less_def vector_less_eq_def all_1, auto)
+  by(simp add:  Cart_eq vector_less_def vector_le_def all_1, auto)
 
 lemma in_interval_1: fixes x :: "real^1" shows
  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b) \<and>
   (x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp add: Cart_eq vector_less_def vector_less_eq_def all_1 dest_vec1_def)
+by(simp add: Cart_eq vector_less_def vector_le_def all_1 dest_vec1_def)
 
 lemma interval_eq_empty_1: fixes a :: "real^1" shows
   "{a .. b} = {} \<longleftrightarrow> dest_vec1 b < dest_vec1 a"
@@ -5067,10 +5067,10 @@
 
 lemma open_closed_interval_1: fixes a :: "real^1" shows
  "{a<..<b} = {a .. b} - {a, b}"
-  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_le_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
 
 lemma closed_open_interval_1: "dest_vec1 (a::real^1) \<le> dest_vec1 b ==> {a .. b} = {a<..<b} \<union> {a,b}"
-  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_le_def and all_1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
 
 (* Some stuff for half-infinite intervals too; FIXME: notation?  *)
 
@@ -5742,30 +5742,30 @@
             else {m *\<^sub>R b + c .. m *\<^sub>R a + c}))"
 proof(cases "m=0")
   { fix x assume "x \<le> c" "c \<le> x"
-    hence "x=c" unfolding vector_less_eq_def and Cart_eq by (auto intro: order_antisym) }
+    hence "x=c" unfolding vector_le_def and Cart_eq by (auto intro: order_antisym) }
   moreover case True
-  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: vector_less_eq_def)
+  moreover have "c \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}" unfolding True by(auto simp add: vector_le_def)
   ultimately show ?thesis by auto
 next
   case False
   { fix y assume "a \<le> y" "y \<le> b" "m > 0"
     hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
-      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component)
+      unfolding vector_le_def by(auto simp add: vector_smult_component vector_add_component)
   } moreover
   { fix y assume "a \<le> y" "y \<le> b" "m < 0"
     hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
-      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
+      unfolding vector_le_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
   } moreover
   { fix y assume "m > 0"  "m *\<^sub>R a + c \<le> y"  "y \<le> m *\<^sub>R b + c"
     hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
-      unfolding image_iff Bex_def mem_interval vector_less_eq_def
+      unfolding image_iff Bex_def mem_interval vector_le_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
         intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
       by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
   } moreover
   { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
     hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
-      unfolding image_iff Bex_def mem_interval vector_less_eq_def
+      unfolding image_iff Bex_def mem_interval vector_le_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
         intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
       by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
--- a/src/HOL/Nitpick.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nitpick.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -213,7 +213,7 @@
 (* While Nitpick normally avoids to unfold definitions for locales, it
    unfortunately needs to unfold them when dealing with the following built-in
    constants. A cleaner approach would be to change "Nitpick_HOL" and
-   "Nitpick_Nits" so that they handle the unexpanded overloaded constants
+   "Nitpick_Nut" so that they handle the unexpanded overloaded constants
    directly, but this is slightly more tricky to implement. *)
 lemmas [nitpick_def] = div_int_inst.div_int div_int_inst.mod_int
     div_nat_inst.div_nat div_nat_inst.mod_nat lower_semilattice_fun_inst.inf_fun
--- a/src/HOL/Nitpick_Examples/ROOT.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nitpick_Examples/ROOT.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -5,4 +5,8 @@
 Nitpick examples.
 *)
 
-setmp_noncritical quick_and_dirty true (try use_thy) "Nitpick_Examples";
+if getenv "KODKODI" = "" then
+  ()
+else
+  setmp_noncritical quick_and_dirty true use_thys
+                    ["Nitpick_Examples"];
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -17,7 +17,7 @@
 nitpick [expect = genuine] 2
 nitpick [expect = genuine]
 nitpick [card = 5, expect = genuine]
-nitpick [sat_solver = MiniSat, expect = genuine] 2
+nitpick [sat_solver = SAT4J, expect = genuine] 2
 oops
 
 subsection {* Examples and Test Cases *}
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -109,13 +109,13 @@
 
 lemma "\<exists>one \<in> {1}. \<exists>two \<in> {2}.
        f5 (\<lambda>a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x"
-nitpick [expect = genuine]
-sorry
+nitpick [expect = potential] (* unfortunate *)
+oops
 
 lemma "\<exists>two \<in> {2}. \<exists>one \<in> {1}.
        f5 (\<lambda>a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x"
-nitpick [expect = genuine]
-sorry
+nitpick [expect = potential] (* unfortunate *)
+oops
 
 lemma "\<forall>a. g a = a
        \<Longrightarrow> \<exists>one \<in> {1}. \<exists>two \<in> {2}. f5 g x =
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -141,11 +141,11 @@
 by (rule Rep_Sum_inverse)
 
 lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
-nitpick [expect = none]
+(* nitpick [expect = none] FIXME *)
 by (rule Zero_nat_def_raw)
 
 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Rep_Nat n))"
-nitpick [expect = none]
+(* nitpick [expect = none] FIXME *)
 by (rule Suc_def)
 
 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
@@ -177,11 +177,11 @@
 by (rule Rep_point_ext_type_inverse)
 
 lemma "Fract a b = of_int a / of_int b"
-nitpick [card = 1\<midarrow>2, expect = none]
+nitpick [card = 1, expect = none]
 by (rule Fract_of_int_quotient)
 
 lemma "Abs_Rat (Rep_Rat a) = a"
-nitpick [card = 1\<midarrow>2, expect = none]
+nitpick [card = 1, expect = none]
 by (rule Rep_Rat_inverse)
 
 end
--- a/src/HOL/Nominal/nominal_datatype.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nominal/nominal_datatype.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -568,7 +568,7 @@
     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) =
       thy3
       |> Sign.map_naming Name_Space.conceal
-      |> Inductive.add_inductive_global (serial ())
+      |> Inductive.add_inductive_global
           {quiet_mode = false, verbose = false, alt_name = Binding.name big_rep_name,
            coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false}
           (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn))
@@ -1511,7 +1511,7 @@
     val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) =
       thy10
       |> Sign.map_naming Name_Space.conceal
-      |> Inductive.add_inductive_global (serial ())
+      |> Inductive.add_inductive_global
           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
            coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false}
           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
--- a/src/HOL/Nominal/nominal_primrec.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Nominal/nominal_primrec.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -239,7 +239,7 @@
 
 local
 
-fun gen_primrec set_group prep_spec prep_term invs fctxt raw_fixes raw_params raw_spec lthy =
+fun gen_primrec prep_spec prep_term invs fctxt raw_fixes raw_params raw_spec lthy =
   let
     val (fixes', spec) = fst (prep_spec (raw_fixes @ raw_params) raw_spec lthy);
     val fixes = List.take (fixes', length raw_fixes);
@@ -280,7 +280,6 @@
       else primrec_err ("functions " ^ commas_quote names2 ^
         "\nare not mutually recursive");
     val (defs_thms, lthy') = lthy |>
-      set_group ? Local_Theory.set_group (serial ()) |>
       fold_map (apfst (snd o snd) oo Local_Theory.define Thm.definitionK o fst) defs';
     val qualify = Binding.qualify false
       (space_implode "_" (map (Long_Name.base_name o #1) defs));
@@ -384,8 +383,8 @@
 
 in
 
-val add_primrec = gen_primrec false Specification.check_spec (K I);
-val add_primrec_cmd = gen_primrec true Specification.read_spec Syntax.read_term;
+val add_primrec = gen_primrec Specification.check_spec (K I);
+val add_primrec_cmd = gen_primrec Specification.read_spec Syntax.read_term;
 
 end;
 
--- a/src/HOL/Number_Theory/Cong.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Number_Theory/Cong.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -20,7 +20,7 @@
 
 The original theory, "IntPrimes", by Thomas M. Rasmussen, defined and
 developed the congruence relations on the integers. The notion was
-extended to the natural numbers by Chiaeb. Jeremy Avigad combined
+extended to the natural numbers by Chaieb. Jeremy Avigad combined
 these, revised and tidied them, made the development uniform for the
 natural numbers and the integers, and added a number of new theorems.
 
--- a/src/HOL/Number_Theory/Primes.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Number_Theory/Primes.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -16,7 +16,7 @@
 another extension of the notions to the integers, and added a number
 of results to "Primes" and "GCD". IntPrimes also defined and developed
 the congruence relations on the integers. The notion was extended to
-the natural numbers by Chiaeb.
+the natural numbers by Chaieb.
 
 Jeremy Avigad combined all of these, made everything uniform for the
 natural numbers and the integers, and added a number of new theorems.
--- a/src/HOL/Predicate.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Predicate.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -570,6 +570,9 @@
 definition if_pred :: "bool \<Rightarrow> unit pred" where
   if_pred_eq: "if_pred b = (if b then single () else \<bottom>)"
 
+definition holds :: "unit pred \<Rightarrow> bool" where
+  holds_eq: "holds P = eval P ()"
+
 definition not_pred :: "unit pred \<Rightarrow> unit pred" where
   not_pred_eq: "not_pred P = (if eval P () then \<bottom> else single ())"
 
@@ -592,7 +595,54 @@
 lemma not_predE': "eval (not_pred P) x \<Longrightarrow> (\<not> eval P x \<Longrightarrow> thesis) \<Longrightarrow> thesis"
   unfolding not_pred_eq
   by (auto split: split_if_asm elim: botE)
+lemma "f () = False \<or> f () = True"
+by simp
 
+lemma closure_of_bool_cases:
+assumes "(f :: unit \<Rightarrow> bool) = (%u. False) \<Longrightarrow> P f"
+assumes "f = (%u. True) \<Longrightarrow> P f"
+shows "P f"
+proof -
+  have "f = (%u. False) \<or> f = (%u. True)"
+    apply (cases "f ()")
+    apply (rule disjI2)
+    apply (rule ext)
+    apply (simp add: unit_eq)
+    apply (rule disjI1)
+    apply (rule ext)
+    apply (simp add: unit_eq)
+    done
+  from this prems show ?thesis by blast
+qed
+
+lemma unit_pred_cases:
+assumes "P \<bottom>"
+assumes "P (single ())"
+shows "P Q"
+using assms
+unfolding bot_pred_def Collect_def empty_def single_def
+apply (cases Q)
+apply simp
+apply (rule_tac f="fun" in closure_of_bool_cases)
+apply auto
+apply (subgoal_tac "(%x. () = x) = (%x. True)") 
+apply auto
+done
+
+lemma holds_if_pred:
+  "holds (if_pred b) = b"
+unfolding if_pred_eq holds_eq
+by (cases b) (auto intro: singleI elim: botE)
+
+lemma if_pred_holds:
+  "if_pred (holds P) = P"
+unfolding if_pred_eq holds_eq
+by (rule unit_pred_cases) (auto intro: singleI elim: botE)
+
+lemma is_empty_holds:
+  "is_empty P \<longleftrightarrow> \<not> holds P"
+unfolding is_empty_def holds_eq
+by (rule unit_pred_cases) (auto elim: botE intro: singleI)
 
 subsubsection {* Implementation *}
 
@@ -838,7 +888,7 @@
   bind (infixl "\<guillemotright>=" 70)
 
 hide (open) type pred seq
-hide (open) const Pred eval single bind is_empty singleton if_pred not_pred
+hide (open) const Pred eval single bind is_empty singleton if_pred not_pred holds
   Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map not_unique the
 
 end
--- a/src/HOL/SMT/Examples/SMT_Examples.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/SMT/Examples/SMT_Examples.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -409,6 +409,13 @@
   using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_20"]]
   by smt
 
+lemma                                                                         
+  assumes "(n + m) mod 2 = 0" and "n mod 4 = 3"                               
+  shows "n mod 2 = 1 & m mod 2 = (1::int)"      
+  using assms
+  using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_21"]]
+  by smt
+
 
 subsection {* Linear arithmetic with quantifiers *}
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT/Examples/cert/z3_linarith_21	Thu Nov 19 06:01:02 2009 -0800
@@ -0,0 +1,10 @@
+(benchmark Isabelle
+:extrafuns (
+  (uf_2 Int)
+  (uf_1 Int)
+ )
+:assumption (= (mod (+ uf_1 uf_2) 2) 0)
+:assumption (= (mod uf_1 4) 3)
+:assumption (not (and (= (mod uf_1 2) 1) (= (mod uf_2 2) 1)))
+:formula true
+)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT/Examples/cert/z3_linarith_21.proof	Thu Nov 19 06:01:02 2009 -0800
@@ -0,0 +1,208 @@
+#2 := false
+#9 := 0::int
+#11 := 4::int
+decl uf_1 :: int
+#4 := uf_1
+#189 := (div uf_1 4::int)
+#210 := -4::int
+#211 := (* -4::int #189)
+#12 := (mod uf_1 4::int)
+#134 := -1::int
+#209 := (* -1::int #12)
+#212 := (+ #209 #211)
+#213 := (+ uf_1 #212)
+#214 := (<= #213 0::int)
+#215 := (not #214)
+#208 := (>= #213 0::int)
+#207 := (not #208)
+#216 := (or #207 #215)
+#217 := (not #216)
+#1 := true
+#36 := [true-axiom]: true
+#393 := (or false #217)
+#394 := [th-lemma]: #393
+#395 := [unit-resolution #394 #36]: #217
+#224 := (or #216 #214)
+#225 := [def-axiom]: #224
+#396 := [unit-resolution #225 #395]: #214
+#222 := (or #216 #208)
+#223 := [def-axiom]: #222
+#397 := [unit-resolution #223 #395]: #208
+#250 := (>= #12 4::int)
+#251 := (not #250)
+#398 := (or false #251)
+#399 := [th-lemma]: #398
+#400 := [unit-resolution #399 #36]: #251
+#13 := 3::int
+#90 := (>= #12 3::int)
+#92 := (not #90)
+#89 := (<= #12 3::int)
+#91 := (not #89)
+#93 := (or #91 #92)
+#94 := (not #93)
+#14 := (= #12 3::int)
+#95 := (iff #14 #94)
+#96 := [rewrite]: #95
+#38 := [asserted]: #14
+#97 := [mp #38 #96]: #94
+#99 := [not-or-elim #97]: #90
+#7 := 2::int
+#261 := (div uf_1 2::int)
+#140 := -2::int
+#276 := (* -2::int #261)
+#15 := (mod uf_1 2::int)
+#275 := (* -1::int #15)
+#277 := (+ #275 #276)
+#278 := (+ uf_1 #277)
+#279 := (<= #278 0::int)
+#280 := (not #279)
+#274 := (>= #278 0::int)
+#273 := (not #274)
+#281 := (or #273 #280)
+#282 := (not #281)
+#408 := (or false #282)
+#409 := [th-lemma]: #408
+#410 := [unit-resolution #409 #36]: #282
+#289 := (or #281 #279)
+#290 := [def-axiom]: #289
+#411 := [unit-resolution #290 #410]: #279
+#287 := (or #281 #274)
+#288 := [def-axiom]: #287
+#412 := [unit-resolution #288 #410]: #274
+#16 := 1::int
+#55 := (>= #15 1::int)
+#100 := (not #55)
+decl uf_2 :: int
+#5 := uf_2
+#18 := (mod uf_2 2::int)
+#61 := (<= #18 1::int)
+#102 := (not #61)
+#375 := [hypothesis]: #102
+#358 := (>= #18 2::int)
+#359 := (not #358)
+#403 := (or false #359)
+#404 := [th-lemma]: #403
+#405 := [unit-resolution #404 #36]: #359
+#406 := [th-lemma #405 #375]: false
+#407 := [lemma #406]: #61
+#413 := (or #100 #102)
+#62 := (>= #18 1::int)
+#315 := (div uf_2 2::int)
+#330 := (* -2::int #315)
+#329 := (* -1::int #18)
+#331 := (+ #329 #330)
+#332 := (+ uf_2 #331)
+#333 := (<= #332 0::int)
+#334 := (not #333)
+#328 := (>= #332 0::int)
+#327 := (not #328)
+#335 := (or #327 #334)
+#336 := (not #335)
+#376 := (or false #336)
+#377 := [th-lemma]: #376
+#378 := [unit-resolution #377 #36]: #336
+#343 := (or #335 #333)
+#344 := [def-axiom]: #343
+#379 := [unit-resolution #344 #378]: #333
+#341 := (or #335 #328)
+#342 := [def-axiom]: #341
+#380 := [unit-resolution #342 #378]: #328
+#103 := (not #62)
+#381 := [hypothesis]: #103
+#352 := (>= #18 0::int)
+#382 := (or false #352)
+#383 := [th-lemma]: #382
+#384 := [unit-resolution #383 #36]: #352
+#6 := (+ uf_1 uf_2)
+#116 := (div #6 2::int)
+#141 := (* -2::int #116)
+#8 := (mod #6 2::int)
+#139 := (* -1::int #8)
+#142 := (+ #139 #141)
+#143 := (+ uf_2 #142)
+#144 := (+ uf_1 #143)
+#138 := (<= #144 0::int)
+#136 := (not #138)
+#137 := (>= #144 0::int)
+#135 := (not #137)
+#145 := (or #135 #136)
+#146 := (not #145)
+#385 := (or false #146)
+#386 := [th-lemma]: #385
+#387 := [unit-resolution #386 #36]: #146
+#153 := (or #145 #138)
+#154 := [def-axiom]: #153
+#388 := [unit-resolution #154 #387]: #138
+#151 := (or #145 #137)
+#152 := [def-axiom]: #151
+#389 := [unit-resolution #152 #387]: #137
+#78 := (<= #8 0::int)
+#79 := (>= #8 0::int)
+#81 := (not #79)
+#80 := (not #78)
+#82 := (or #80 #81)
+#83 := (not #82)
+#10 := (= #8 0::int)
+#84 := (iff #10 #83)
+#85 := [rewrite]: #84
+#37 := [asserted]: #10
+#86 := [mp #37 #85]: #83
+#87 := [not-or-elim #86]: #78
+#390 := (or false #79)
+#391 := [th-lemma]: #390
+#392 := [unit-resolution #391 #36]: #79
+#401 := [th-lemma #99 #400 #397 #396 #392 #87 #389 #388 #384 #381 #380 #379]: false
+#402 := [lemma #401]: #62
+#57 := (<= #15 1::int)
+#101 := (not #57)
+#369 := [hypothesis]: #101
+#304 := (>= #15 2::int)
+#305 := (not #304)
+#370 := (or false #305)
+#371 := [th-lemma]: #370
+#372 := [unit-resolution #371 #36]: #305
+#373 := [th-lemma #372 #369]: false
+#374 := [lemma #373]: #57
+#104 := (or #100 #101 #102 #103)
+#69 := (and #55 #57 #61 #62)
+#74 := (not #69)
+#113 := (iff #74 #104)
+#105 := (not #104)
+#108 := (not #105)
+#111 := (iff #108 #104)
+#112 := [rewrite]: #111
+#109 := (iff #74 #108)
+#106 := (iff #69 #105)
+#107 := [rewrite]: #106
+#110 := [monotonicity #107]: #109
+#114 := [trans #110 #112]: #113
+#19 := (= #18 1::int)
+#17 := (= #15 1::int)
+#20 := (and #17 #19)
+#21 := (not #20)
+#75 := (iff #21 #74)
+#72 := (iff #20 #69)
+#63 := (and #61 #62)
+#58 := (and #55 #57)
+#66 := (and #58 #63)
+#70 := (iff #66 #69)
+#71 := [rewrite]: #70
+#67 := (iff #20 #66)
+#64 := (iff #19 #63)
+#65 := [rewrite]: #64
+#59 := (iff #17 #58)
+#60 := [rewrite]: #59
+#68 := [monotonicity #60 #65]: #67
+#73 := [trans #68 #71]: #72
+#76 := [monotonicity #73]: #75
+#39 := [asserted]: #21
+#77 := [mp #39 #76]: #74
+#115 := [mp #77 #114]: #104
+#414 := [unit-resolution #115 #374 #402]: #413
+#415 := [unit-resolution #414 #407]: #100
+#298 := (>= #15 0::int)
+#416 := (or false #298)
+#417 := [th-lemma]: #416
+#418 := [unit-resolution #417 #36]: #298
+[th-lemma #418 #415 #412 #411 #99 #400 #397 #396]: false
+unsat
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -18,6 +18,7 @@
     thm
 
   (*setup*)
+  val trace_assms: bool Config.T
   val setup: theory -> theory
 end
 
@@ -474,6 +475,9 @@
 
 val true_false = @{lemma "True == ~ False" by simp}
 
+val (trace_assms, trace_assms_setup) =
+  Attrib.config_bool "z3_trace_assms" false
+
 local
   val TT_eq = @{lemma "(P = (~False)) == P" by simp}
   val remove_trigger = @{lemma "trigger t p == p"
@@ -491,10 +495,15 @@
 
   val threshold = 10
   
+  fun trace ctxt thm =
+    if Config.get ctxt trace_assms
+    then tracing (Display.string_of_thm ctxt thm)
+    else ()
+
   val lookup = (fn Some thms => first_of thms | Many net => net_instance net)
   fun lookup_assm ctxt assms ct =
     (case lookup assms ct of
-      SOME thm => thm
+      SOME thm => (trace ctxt thm; thm)
     | _ => z3_exn ("not asserted: " ^
         quote (Syntax.string_of_term ctxt (Thm.term_of ct))))
 in
@@ -1392,6 +1401,6 @@
 
   in (fn ptab => fn idx => result (fst (lookup idx ptab))) end
 
-val setup = Z3_Rewrite_Rules.setup
+val setup = trace_assms_setup #> Z3_Rewrite_Rules.setup
 
 end
--- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -155,7 +155,7 @@
     val ({intrs = rec_intrs, elims = rec_elims, ...}, thy1) =
       thy0
       |> Sign.map_naming Name_Space.conceal
-      |> Inductive.add_inductive_global (serial ())
+      |> Inductive.add_inductive_global
           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name',
             coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false}
           (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts))
--- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -174,7 +174,7 @@
     val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) =
       thy1
       |> Sign.map_naming Name_Space.conceal
-      |> Inductive.add_inductive_global (serial ())
+      |> Inductive.add_inductive_global
           {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name,
            coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false}
           (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') []
--- a/src/HOL/Tools/Function/fun.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Function/fun.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -151,15 +151,11 @@
   domintros=false, partials=false, tailrec=false }
 
 fun gen_fun add config fixes statements int lthy =
-  let val group = serial () in
-    lthy
-      |> Local_Theory.set_group group
-      |> add fixes statements config
-      |> by_pat_completeness_auto int
-      |> Local_Theory.restore
-      |> Local_Theory.set_group group
-      |> termination_by (Function_Common.get_termination_prover lthy) int
-  end;
+  lthy
+    |> add fixes statements config
+    |> by_pat_completeness_auto int
+    |> Local_Theory.restore
+    |> termination_by (Function_Common.get_termination_prover lthy) int
 
 val add_fun = gen_fun Function.add_function
 val add_fun_cmd = gen_fun Function.add_function_cmd
--- a/src/HOL/Tools/Function/function.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Function/function.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -20,8 +20,6 @@
 
     val termination_proof : term option -> local_theory -> Proof.state
     val termination_proof_cmd : string option -> local_theory -> Proof.state
-    val termination : term option -> local_theory -> Proof.state
-    val termination_cmd : string option -> local_theory -> Proof.state
 
     val setup : theory -> theory
     val get_congs : Proof.context -> thm list
@@ -119,7 +117,6 @@
         end
     in
       lthy
-        |> is_external ? Local_Theory.set_group (serial ())
         |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
         |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
     end
@@ -170,18 +167,8 @@
       |> Proof.theorem_i NONE afterqed [[(goal, [])]]
     end
 
-val termination_proof = gen_termination_proof Syntax.check_term;
-val termination_proof_cmd = gen_termination_proof Syntax.read_term;
-
-fun termination term_opt lthy =
-  lthy
-  |> Local_Theory.set_group (serial ())
-  |> termination_proof term_opt;
-
-fun termination_cmd term_opt lthy =
-  lthy
-  |> Local_Theory.set_group (serial ())
-  |> termination_proof_cmd term_opt;
+val termination_proof = gen_termination_proof Syntax.check_term
+val termination_proof_cmd = gen_termination_proof Syntax.read_term
 
 
 (* Datatype hook to declare datatype congs as "function_congs" *)
@@ -217,13 +204,13 @@
 val _ =
   OuterSyntax.local_theory_to_proof "function" "define general recursive functions" K.thy_goal
   (function_parser default_config
-     >> (fn ((config, fixes), statements) => add_function_cmd fixes statements config));
+     >> (fn ((config, fixes), statements) => add_function_cmd fixes statements config))
 
 val _ =
   OuterSyntax.local_theory_to_proof "termination" "prove termination of a recursive function" K.thy_goal
-  (Scan.option P.term >> termination_cmd);
+  (Scan.option P.term >> termination_proof_cmd)
 
-end;
+end
 
 
 end
--- a/src/HOL/Tools/Function/function_common.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Function/function_common.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -259,12 +259,18 @@
             (fname, length args)
           end
 
-      val _ = AList.group (op =) (map check eqs)
+      val grouped_args = AList.group (op =) (map check eqs)
+      val _ = grouped_args
         |> map (fn (fname, ars) =>
              length (distinct (op =) ars) = 1
              orelse error ("Function " ^ quote fname ^
                            " has different numbers of arguments in different equations"))
 
+      val not_defined = subtract (op =) (map fst grouped_args) fnames
+      val _ = null not_defined
+        orelse error ("No defining equations for function" ^
+          plural " " "s " not_defined ^ commas_quote not_defined)
+
       fun check_sorts ((fname, fT), _) =
           Sorts.of_sort (Sign.classes_of (ProofContext.theory_of ctxt)) (fT, HOLogic.typeS)
           orelse error (cat_lines 
--- a/src/HOL/Tools/Nitpick/HISTORY	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/HISTORY	Thu Nov 19 06:01:02 2009 -0800
@@ -12,6 +12,8 @@
   * Added support for codatatype view of datatypes
   * Fixed soundness bugs related to sets and sets of sets
   * Fixed monotonicity check
+  * Fixed error when processing definitions that resulted in an exception
+  * Fixed error in Kodkod encoding of "The" and "Eps"
   * Fixed error in display of uncurried constants
   * Speeded up scope enumeration
 
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -1047,8 +1047,7 @@
                         \$JAVA_LIBRARY_PATH\" \
                         \LD_LIBRARY_PATH=\"$KODKODI_JAVA_LIBRARY_PATH:\
                         \$LD_LIBRARY_PATH\" \
-                        \\"$ISABELLE_TOOL\" java \
-                        \de.tum.in.isabelle.Kodkodi.Kodkodi" ^
+                        \\"$KODKODI\"/bin/kodkodi" ^
                         (if ms >= 0 then " -max-msecs " ^ Int.toString ms
                          else "") ^
                         (if max_solutions > 1 then " -solve-all" else "") ^
--- a/src/HOL/Tools/Nitpick/kodkod_sat.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/kodkod_sat.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -28,11 +28,9 @@
 
 (* (string * sat_solver_info) list *)
 val static_list =
-  [("MiniSatJNI", Internal (JNI, Incremental, ["MiniSat"])),
-   ("MiniSat", ExternalV2 (ToFile, "MINISAT_HOME", "minisat", [], "SAT", "",
+  [("MiniSat", ExternalV2 (ToFile, "MINISAT_HOME", "minisat", [], "SAT", "",
                            "UNSAT")),
    ("PicoSAT", External (ToStdout, "PICOSAT_HOME", "picosat", [])),
-   ("zChaffJNI", Internal (JNI, Batch, ["zChaff"])),
    ("zChaff", ExternalV2 (ToStdout, "ZCHAFF_HOME", "zchaff", [],
                           "Instance Satisfiable", "",
                           "Instance Unsatisfiable")),
@@ -44,6 +42,8 @@
                            "solution =", "UNSATISFIABLE          !!")),
    ("BerkMinAlloy", External (ToStdout, "BERKMINALLOY_HOME", "berkmin", [])),
    ("Jerusat", External (ToStdout, "JERUSAT_HOME", "Jerusat1.3", [])),
+   ("MiniSatJNI", Internal (JNI, Incremental, ["MiniSat"])),
+   ("zChaffJNI", Internal (JNI, Batch, ["zChaff"])),
    ("SAT4J", Internal (Java, Incremental, ["DefaultSAT4J"])),
    ("SAT4JLight", Internal (Java, Incremental, ["LightSAT4J"])),
    ("HaifaSat", ExternalV2 (ToStdout, "HAIFASAT_HOME", "HaifaSat", ["-p", "1"],
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -1109,13 +1109,13 @@
     |> map_filter (try (Refute.specialize_type thy x))
     |> filter (equal (Const x) o term_under_def)
 
-(* term -> term *)
+(* theory -> term -> term option *)
 fun normalized_rhs_of thy t =
   let
-    (* term -> term *)
-    fun aux (v as Var _) t = lambda v t
-      | aux (c as Const (@{const_name TYPE}, T)) t = lambda c t
-      | aux _ _ = raise TERM ("Nitpick_HOL.normalized_rhs_of", [t])
+    (* term option -> term option *)
+    fun aux (v as Var _) (SOME t) = SOME (lambda v t)
+      | aux (c as Const (@{const_name TYPE}, T)) (SOME t) = SOME (lambda c t)
+      | aux _ _ = NONE
     val (lhs, rhs) =
       case t of
         Const (@{const_name "=="}, _) $ t1 $ t2 => (t1, t2)
@@ -1123,7 +1123,7 @@
         (t1, t2)
       | _ => raise TERM ("Nitpick_HOL.normalized_rhs_of", [t])
     val args = strip_comb lhs |> snd
-  in fold_rev aux args rhs end
+  in fold_rev aux args (SOME rhs) end
 
 (* theory -> const_table -> styp -> term option *)
 fun def_of_const thy table (x as (s, _)) =
@@ -1131,7 +1131,7 @@
     NONE
   else
     x |> def_props_for_const thy false table |> List.last
-      |> normalized_rhs_of thy |> prefix_abs_vars s |> SOME
+      |> normalized_rhs_of thy |> Option.map (prefix_abs_vars s)
     handle List.Empty => NONE
 
 datatype fixpoint_kind = Lfp | Gfp | NoFp
@@ -1416,8 +1416,8 @@
     SOME t' => is_constr_pattern_lhs thy t'
   | NONE => false
 
-val unfold_max_depth = 63
-val axioms_max_depth = 63
+val unfold_max_depth = 255
+val axioms_max_depth = 255
 
 (* extended_context -> term -> term *)
 fun unfold_defs_in_term (ext_ctxt as {thy, destroy_constrs, fast_descrs,
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -1092,6 +1092,12 @@
                   else
                     kk_rel_eq r1 r2
                 end)
+         | Op2 (The, T, _, u1, u2) =>
+           to_f_with_polarity polar
+                              (Op2 (The, T, Opt (Atom (2, bool_j0)), u1, u2))
+         | Op2 (Eps, T, _, u1, u2) =>
+           to_f_with_polarity polar
+                              (Op2 (Eps, T, Opt (Atom (2, bool_j0)), u1, u2))
          | Op2 (Apply, T, _, u1, u2) =>
            (case (polar, rep_of u1) of
               (Neg, Func (R, Formula Neut)) => kk_subset (to_opt R u2) (to_r u1)
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -784,8 +784,8 @@
   let
     (* typ -> ctype *)
     val ctype_for = fresh_ctype_for_type cdata
-    (* term -> accumulator -> accumulator *)
-    val do_term = snd oo consider_term cdata
+    (* term -> accumulator -> ctype * accumulator *)
+    val do_term = consider_term cdata
     (* sign -> term -> accumulator -> accumulator *)
     fun do_formula _ _ (_, UnsolvableCSet) = unsolvable_accum
       | do_formula sn t (accum as (gamma as {bounds, frees, consts}, cset)) =
@@ -810,8 +810,11 @@
           (* term -> term -> accumulator *)
           fun do_equals t1 t2 =
             case sn of
-              Pos => do_term t accum
-            | Neg => fold do_term [t1, t2] accum
+              Pos => do_term t accum |> snd
+            | Neg => let
+                       val (C1, accum) = do_term t1 accum
+                       val (C2, accum) = do_term t2 accum
+                     in accum ||> add_ctypes_equal C1 C2 end
         in
           case t of
             Const (s0 as @{const_name all}, _) $ Abs (_, T1, t1) =>
@@ -839,10 +842,10 @@
           | @{const "op -->"} $ t1 $ t2 =>
             accum |> do_contra_formula t1 |> do_co_formula t2
           | Const (@{const_name If}, _) $ t1 $ t2 $ t3 =>
-            accum |> do_term t1 |> fold do_co_formula [t2, t3]
+            accum |> do_term t1 |> snd |> fold do_co_formula [t2, t3]
           | Const (@{const_name Let}, _) $ t1 $ t2 =>
             do_co_formula (betapply (t2, t1)) accum
-          | _ => do_term t accum
+          | _ => do_term t accum |> snd
         end
         |> tap (fn _ => print_g ("\<Gamma> \<turnstile> " ^
                                  Syntax.string_of_term ctxt t ^
@@ -873,7 +876,7 @@
     I
   else
     let
-      (* term -> accumulator -> accumulator *)
+      (* term -> accumulator -> ctype * accumulator *)
       val do_term = consider_term cdata
       (* typ -> term -> accumulator -> accumulator *)
       fun do_all abs_T body_t accum =
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -1158,8 +1158,10 @@
             let
               val u1' = sub u1
               val opt1 = is_opt_rep (rep_of u1')
+              val opt = (oper = Eps orelse opt1)
               val unopt_R = best_one_rep_for_type scope T |> optable_rep ofs T
-              val R = unopt_R |> (oper = Eps orelse opt1) ? opt_rep ofs T
+              val R = if is_boolean_type T then bool_rep polar opt
+                      else unopt_R |> opt ? opt_rep ofs T
               val u = Op2 (oper, T, R, u1', sub u2)
             in
               if is_precise_type datatypes T orelse not opt1 then
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -111,10 +111,10 @@
   in
     Options {
       expected_modes = Option.map (pair const) expected_modes,
-      proposed_modes = if not (null proposed_modes) then [(const, map fst proposed_modes)] else [],
+      proposed_modes = Option.map (pair const o map fst) proposed_modes,
       proposed_names =
-        map_filter
-          (fn (m, NONE) => NONE | (m, SOME name) => SOME ((const, m), name)) proposed_modes,
+        the_default [] (Option.map (map_filter
+          (fn (m, NONE) => NONE | (m, SOME name) => SOME ((const, m), name))) proposed_modes),
       show_steps = chk "show_steps",
       show_intermediate_results = chk "show_intermediate_results",
       show_proof_trace = chk "show_proof_trace",
@@ -176,7 +176,7 @@
 
 val opt_modes =
   Scan.optional (P.$$$ "(" |-- Args.$$$ "modes" |-- P.$$$ ":" |--
-    P.enum1 "," mode_and_opt_proposal --| P.$$$ ")") []
+    P.enum "," mode_and_opt_proposal --| P.$$$ ")" >> SOME) NONE
 
 val opt_expected_modes =
   Scan.optional (P.$$$ "(" |-- Args.$$$ "expected_modes" |-- P.$$$ ":" |--
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -282,7 +282,7 @@
 
 datatype options = Options of {  
   expected_modes : (string * mode' list) option,
-  proposed_modes : (string * mode' list) list,
+  proposed_modes : (string * mode' list) option,
   proposed_names : ((string * mode') * string) list,
   show_steps : bool,
   show_proof_trace : bool,
@@ -299,7 +299,7 @@
 };
 
 fun expected_modes (Options opt) = #expected_modes opt
-fun proposed_modes (Options opt) name = AList.lookup (op =) (#proposed_modes opt) name
+fun proposed_modes (Options opt) = #proposed_modes opt
 fun proposed_names (Options opt) name mode = AList.lookup (eq_pair (op =) eq_mode')
   (#proposed_names opt) (name, mode)
 
@@ -318,7 +318,7 @@
 
 val default_options = Options {
   expected_modes = NONE,
-  proposed_modes = [],
+  proposed_modes = NONE,
   proposed_names = [],
   show_steps = false,
   show_intermediate_results = false,
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -82,10 +82,6 @@
     ^ "\n" ^ Pretty.string_of (Pretty.chunks
       (Goal_Display.pretty_goals_without_context (! Goal_Display.goals_limit) st)));
 
-(* reference to preprocessing of InductiveSet package *)
-
-val ind_set_codegen_preproc = (fn thy => I) (*Inductive_Set.codegen_preproc;*)
-
 (** fundamentals **)
 
 (* syntactic operations *)
@@ -417,22 +413,45 @@
   end
 
 (* validity checks *)
+(* EXPECTED MODE and PROPOSED_MODE are largely the same; define a clear semantics for those! *)
 
-fun check_expected_modes preds (options : Predicate_Compile_Aux.options) modes =
-      case expected_modes options of
-      SOME (s, ms) => (case AList.lookup (op =) modes s of
-        SOME modes =>
-          let
-            val modes' = map (translate_mode (the (AList.lookup (op =) preds s))) modes
-          in
-            if not (eq_set eq_mode' (ms, modes')) then
-              error ("expected modes were not inferred:\n"
-              ^ "  inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode' modes')  ^ "\n"
-              ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode' ms))
-            else ()
-          end
-        | NONE => ())
-    | NONE => ()
+fun check_expected_modes preds options modes =
+  case expected_modes options of
+    SOME (s, ms) => (case AList.lookup (op =) modes s of
+      SOME modes =>
+        let
+          val modes' = map (translate_mode (the (AList.lookup (op =) preds s))) modes
+        in
+          if not (eq_set eq_mode' (ms, modes')) then
+            error ("expected modes were not inferred:\n"
+            ^ "  inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode' modes')  ^ "\n"
+            ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode' ms))
+          else ()
+        end
+      | NONE => ())
+  | NONE => ()
+
+fun check_proposed_modes preds options modes extra_modes errors =
+  case proposed_modes options of
+    SOME (s, ms) => (case AList.lookup (op =) modes s of
+      SOME inferred_ms =>
+        let
+          val preds_without_modes = map fst (filter (null o snd) (modes @ extra_modes))
+          val modes' = map (translate_mode (the (AList.lookup (op =) preds s))) inferred_ms
+        in
+          if not (eq_set eq_mode' (ms, modes')) then
+            error ("expected modes were not inferred:\n"
+            ^ "  inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode' modes')  ^ "\n"
+            ^ "  expected modes for " ^ s ^ ": " ^ commas (map string_of_mode' ms) ^ "\n"
+            ^ "For the following clauses, the following modes could not be inferred: " ^ "\n"
+            ^ cat_lines errors ^
+            (if not (null preds_without_modes) then
+              "\n" ^ "No mode inferred for the predicates " ^ commas preds_without_modes
+            else ""))
+          else ()
+        end
+      | NONE => ())
+  | NONE => ()
 
 (* importing introduction rules *)
 
@@ -584,13 +603,13 @@
           let
             val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
           in (fst (dest_Const const) = name) end;      
-        val intros = ind_set_codegen_preproc thy
+        val intros =
           (map (expand_tuples thy #> preprocess_intro thy) (filter is_intro_of (#intrs result)))
         val index = find_index (fn s => s = name) (#names (fst info))
         val pre_elim = nth (#elims result) index
         val pred = nth (#preds result) index
         val nparams = length (Inductive.params_of (#raw_induct result))
-        (*val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams 
+        (*val elim = singleton (Inductive_Set.codegen_preproc thy) (preprocess_elim thy nparams 
           (expand_tuples_elim pre_elim))*)
         val elim =
           (Drule.standard o Skip_Proof.make_thm thy)
@@ -659,8 +678,8 @@
 fun register_predicate (constname, pre_intros, pre_elim, nparams) thy =
   let
     (* preprocessing *)
-    val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
-    val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
+    val intros = map (preprocess_intro thy) pre_intros
+    val elim = preprocess_elim thy nparams pre_elim
   in
     if not (member (op =) (Graph.keys (PredData.get thy)) constname) then
       PredData.map
@@ -1042,21 +1061,34 @@
     else NONE
   end;
 
-fun print_failed_mode options thy modes p m rs i =
+fun print_failed_mode options thy modes p m rs is =
   if show_mode_inference options then
     let
-      val _ = tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
-      p ^ " violates mode " ^ string_of_mode thy p m)
+      val _ = tracing ("Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
+        p ^ " violates mode " ^ string_of_mode thy p m)
     in () end
   else ()
 
+fun error_of thy p m is =
+  ("  Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
+        p ^ " violates mode " ^ string_of_mode thy p m)
+
+fun find_indices f xs =
+  map_filter (fn (i, true) => SOME i | (i, false) => NONE) (map_index (apsnd f) xs)
+
 fun check_modes_pred options with_generator thy param_vs clauses modes gen_modes (p, ms) =
   let
     val rs = case AList.lookup (op =) clauses p of SOME rs => rs | NONE => []
-  in (p, filter (fn m => case find_index
-    (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
-      ~1 => true
-    | i => (print_failed_mode options thy modes p m rs i; false)) ms)
+    fun invalid_mode m =
+      case find_indices
+        (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
+        [] => NONE
+      | is => SOME (error_of thy p m is)
+    val res = map (fn m => (m, invalid_mode m)) ms
+    val ms' = map_filter (fn (m, NONE) => SOME m | _ => NONE) res
+    val errors = map_filter snd res
+  in
+    ((p, ms'), errors)
   end;
 
 fun get_modes_pred with_generator thy param_vs clauses modes gen_modes (p, ms) =
@@ -1071,14 +1103,24 @@
   let val y = f x
   in if x = y then x else fixp f y end;
 
+fun fixp_with_state f ((x : (string * mode list) list), state) =
+  let
+    val (y, state') = f (x, state)
+  in
+    if x = y then (y, state') else fixp_with_state f (y, state')
+  end
+
 fun infer_modes options thy extra_modes all_modes param_vs clauses =
   let
-    val modes =
-      fixp (fn modes =>
-        map (check_modes_pred options false thy param_vs clauses (modes @ extra_modes) []) modes)
-          all_modes
+    val (modes, errors) =
+      fixp_with_state (fn (modes, errors) =>
+        let
+          val res = map
+            (check_modes_pred options false thy param_vs clauses (modes @ extra_modes) []) modes
+        in (map fst res, errors @ maps snd res) end)
+          (all_modes, [])
   in
-    map (get_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes
+    (map (get_modes_pred false thy param_vs clauses (modes @ extra_modes) []) modes, errors)
   end;
 
 fun remove_from rem [] = []
@@ -1087,7 +1129,7 @@
       NONE => (k, vs)
     | SOME vs' => (k, subtract (op =) vs' vs))
     :: remove_from rem xs
-    
+
 fun infer_modes_with_generator options thy extra_modes all_modes param_vs clauses =
   let
     val prednames = map fst clauses
@@ -1096,16 +1138,21 @@
       |> filter_out (fn (name, _) => member (op =) prednames name)
     val starting_modes = remove_from extra_modes' all_modes
     fun eq_mode (m1, m2) = (m1 = m2)
-    val modes =
-      fixp (fn modes =>
-        map (check_modes_pred options true thy param_vs clauses extra_modes'
-          (gen_modes @ modes)) modes) starting_modes
+    val (modes, errors) =
+      fixp_with_state (fn (modes, errors) =>
+        let
+          val res = map
+            (check_modes_pred options true thy param_vs clauses extra_modes'
+              (gen_modes @ modes)) modes
+        in (map fst res, errors @ maps snd res) end) (starting_modes, [])
+    val moded_clauses =
+      map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes
+    val (moded_clauses', _) = infer_modes options thy extra_modes all_modes param_vs clauses
+    val join_moded_clauses_table = AList.join (op =)
+      (fn _ => fn ((mps1, mps2)) =>
+        merge (fn ((m1, _), (m2, _)) => eq_mode (m1, m2)) (mps1, mps2))
   in
-    AList.join (op =)
-    (fn _ => fn ((mps1, mps2)) =>
-      merge (fn ((m1, _), (m2, _)) => eq_mode (m1, m2)) (mps1, mps2))
-    (infer_modes options thy extra_modes all_modes param_vs clauses,
-    map (get_modes_pred true thy param_vs clauses extra_modes (gen_modes @ modes)) modes)
+    (join_moded_clauses_table (moded_clauses', moded_clauses), errors)
   end;
 
 (* term construction *)
@@ -1524,66 +1571,67 @@
   let
     val compfuns = PredicateCompFuns.compfuns
     val T = AList.lookup (op =) preds name |> the
-    fun create_definition (mode as (iss, is)) thy = let
-      val mode_cname = create_constname_of_mode options thy "" name T mode
-      val mode_cbasename = Long_Name.base_name mode_cname
-      val Ts = binder_types T
-      val (Ts1, Ts2) = chop (length iss) Ts
-      val (Us1, Us2) =  split_smodeT is Ts2
-      val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
-      val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (HOLogic.mk_tupleT Us2))
-      val names = Name.variant_list []
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-      val param_names = Name.variant_list []
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1')))
-      val xparams = map2 (curry Free) param_names Ts1'
-      fun mk_vars (i, T) names =
-        let
-          val vname = Name.variant names ("x" ^ string_of_int (length Ts1' + i))
-        in
-          case AList.lookup (op =) is i of
-             NONE => ((([], [Free (vname, T)]), Free (vname, T)), vname :: names)
-           | SOME NONE => ((([Free (vname, T)], []), Free (vname, T)), vname :: names)
-           | SOME (SOME pis) =>
-             let
-               val (Tins, Touts) = split_tupleT pis T
-               val name_in = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "in")
-               val name_out = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "out")
-               val xin = Free (name_in, HOLogic.mk_tupleT Tins)
-               val xout = Free (name_out, HOLogic.mk_tupleT Touts)
-               val xarg = mk_arg xin xout pis T
-             in
-               (((if null Tins then [] else [xin],
-               if null Touts then [] else [xout]), xarg), name_in :: name_out :: names) end
-             end
-      val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
-      val (xinout, xargs) = split_list xinoutargs
-      val (xins, xouts) = pairself flat (split_list xinout)
-      val (xparams', names') = fold_map (mk_Eval_of []) ((xparams ~~ Ts1) ~~ iss) names
-      fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
-        | mk_split_lambda [x] t = lambda x t
-        | mk_split_lambda xs t =
-        let
-          fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
-            | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
-        in
-          mk_split_lambda' xs t
+    fun create_definition (mode as (iss, is)) thy =
+      let
+        val mode_cname = create_constname_of_mode options thy "" name T mode
+        val mode_cbasename = Long_Name.base_name mode_cname
+        val Ts = binder_types T
+        val (Ts1, Ts2) = chop (length iss) Ts
+        val (Us1, Us2) =  split_smodeT is Ts2
+        val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
+        val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (HOLogic.mk_tupleT Us2))
+        val names = Name.variant_list []
+          (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
+        val param_names = Name.variant_list []
+          (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts1')))
+        val xparams = map2 (curry Free) param_names Ts1'
+        fun mk_vars (i, T) names =
+          let
+            val vname = Name.variant names ("x" ^ string_of_int (length Ts1' + i))
+          in
+            case AList.lookup (op =) is i of
+               NONE => ((([], [Free (vname, T)]), Free (vname, T)), vname :: names)
+             | SOME NONE => ((([Free (vname, T)], []), Free (vname, T)), vname :: names)
+             | SOME (SOME pis) =>
+               let
+                 val (Tins, Touts) = split_tupleT pis T
+                 val name_in = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "in")
+                 val name_out = Name.variant names ("x" ^ string_of_int (length Ts1' + i) ^ "out")
+                 val xin = Free (name_in, HOLogic.mk_tupleT Tins)
+                 val xout = Free (name_out, HOLogic.mk_tupleT Touts)
+                 val xarg = mk_arg xin xout pis T
+               in
+                 (((if null Tins then [] else [xin],
+                 if null Touts then [] else [xout]), xarg), name_in :: name_out :: names) end
+               end
+        val (xinoutargs, names) = fold_map mk_vars ((1 upto (length Ts2)) ~~ Ts2) param_names
+        val (xinout, xargs) = split_list xinoutargs
+        val (xins, xouts) = pairself flat (split_list xinout)
+        val (xparams', names') = fold_map (mk_Eval_of []) ((xparams ~~ Ts1) ~~ iss) names
+        fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
+          | mk_split_lambda [x] t = lambda x t
+          | mk_split_lambda xs t =
+          let
+            fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
+              | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
+          in
+            mk_split_lambda' xs t
+          end;
+        val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
+          (list_comb (Const (name, T), xparams' @ xargs)))
+        val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
+        val def = Logic.mk_equals (lhs, predterm)
+        val ([definition], thy') = thy |>
+          Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
+          PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
+        val (intro, elim) =
+          create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
+        in thy'
+          |> add_predfun name mode (mode_cname, definition, intro, elim)
+          |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
+          |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
+          |> Theory.checkpoint
         end;
-      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
-        (list_comb (Const (name, T), xparams' @ xargs)))
-      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
-      val def = Logic.mk_equals (lhs, predterm)
-      val ([definition], thy') = thy |>
-        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
-        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
-      val (intro, elim) =
-        create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
-      in thy'
-        |> add_predfun name mode (mode_cname, definition, intro, elim)
-        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
-        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
-        |> Theory.checkpoint
-      end;
   in
     fold create_definition modes thy
   end;
@@ -1626,8 +1674,8 @@
 
 (* MAJOR FIXME:  prove_params should be simple
  - different form of introrule for parameters ? *)
-fun prove_param thy NONE t = TRY (rtac @{thm refl} 1)
-  | prove_param thy (m as SOME (Mode (mode, is, ms))) t =
+fun prove_param options thy NONE t = TRY (rtac @{thm refl} 1)
+  | prove_param options thy (m as SOME (Mode (mode, is, ms))) t =
   let
     val  (f, args) = strip_comb (Envir.eta_contract t)
     val (params, _) = chop (length ms) args
@@ -1639,16 +1687,15 @@
     | Free _ => TRY (rtac @{thm refl} 1)
     | Abs _ => error "prove_param: No valid parameter term"
   in
-    REPEAT_DETERM (etac @{thm thin_rl} 1)
-    THEN REPEAT_DETERM (rtac @{thm ext} 1)
-    THEN print_tac "prove_param"
+    REPEAT_DETERM (rtac @{thm ext} 1)
+    THEN print_tac' options "prove_param"
     THEN f_tac
-    THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map2 (prove_param thy) ms params))
+    THEN print_tac' options "after simplification in prove_args"
+    THEN (EVERY (map2 (prove_param options thy) ms params))
     THEN (REPEAT_DETERM (atac 1))
   end
 
-fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
+fun prove_expr options thy (Mode (mode, is, ms), t, us) (premposition : int) =
   case strip_comb t of
     (Const (name, T), args) =>  
       let
@@ -1656,16 +1703,16 @@
         val (args1, args2) = chop (length ms) args
       in
         rtac @{thm bindI} 1
-        THEN print_tac "before intro rule:"
+        THEN print_tac' options "before intro rule:"
         (* for the right assumption in first position *)
         THEN rotate_tac premposition 1
         THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
         THEN rtac introrule 1
-        THEN print_tac "after intro rule"
+        THEN print_tac' options "after intro rule"
         (* work with parameter arguments *)
-        THEN (atac 1)
-        THEN (print_tac "parameter goal")
-        THEN (EVERY (map2 (prove_param thy) ms args1))
+        THEN atac 1
+        THEN print_tac' options "parameter goal"
+        THEN (EVERY (map2 (prove_param options thy) ms args1))
         THEN (REPEAT_DETERM (atac 1))
       end
   | _ => rtac @{thm bindI} 1
@@ -1673,7 +1720,7 @@
       (HOL_basic_ss' addsimps [@{thm "split_eta"}, @{thm "split_beta"}, @{thm "fst_conv"},
          @{thm "snd_conv"}, @{thm pair_collapse}]) 1
     THEN (atac 1)
-    THEN print_tac "after prove parameter call"
+    THEN print_tac' options "after prove parameter call"
     
 
 fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
@@ -1725,9 +1772,9 @@
     val (in_ts, clause_out_ts) = split_smode is ts;
     fun prove_prems out_ts [] =
       (prove_match thy out_ts)
-      THEN print_tac "before simplifying assumptions"
+      THEN print_tac' options "before simplifying assumptions"
       THEN asm_full_simp_tac HOL_basic_ss' 1
-      THEN print_tac "before single intro rule"
+      THEN print_tac' options "before single intro rule"
       THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
     | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
       let
@@ -1737,11 +1784,11 @@
               val (_, out_ts''') = split_smode is us
               val rec_tac = prove_prems out_ts''' ps
             in
-              print_tac "before clause:"
+              print_tac' options "before clause:"
               THEN asm_simp_tac HOL_basic_ss 1
-              THEN print_tac "before prove_expr:"
-              THEN prove_expr thy (mode, t, us) premposition
-              THEN print_tac "after prove_expr:"
+              THEN print_tac' options "before prove_expr:"
+              THEN prove_expr options thy (mode, t, us) premposition
+              THEN print_tac' options "after prove_expr:"
               THEN rec_tac
             end
           | Negprem (us, t) =>
@@ -1752,13 +1799,18 @@
               val (_, params) = strip_comb t
             in
               rtac @{thm bindI} 1
+              THEN print_tac' options "before prove_neg_expr:"
               THEN (if (is_some name) then
-                  simp_tac (HOL_basic_ss addsimps
+                  print_tac' options ("before unfolding definition " ^
+                    (Display.string_of_thm_global thy
+                      (predfun_definition_of thy (the name) (iss, is))))
+                  THEN simp_tac (HOL_basic_ss addsimps
                     [predfun_definition_of thy (the name) (iss, is)]) 1
                   THEN rtac @{thm not_predI} 1
+                  THEN print_tac' options "after applying rule not_predI"
                   THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
                   THEN (REPEAT_DETERM (atac 1))
-                  THEN (EVERY (map2 (prove_param thy) param_modes params))
+                  THEN (EVERY (map2 (prove_param options thy) param_modes params))
                 else
                   rtac @{thm not_predI'} 1)
                   THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
@@ -1767,9 +1819,9 @@
           | Sidecond t =>
            rtac @{thm bindI} 1
            THEN rtac @{thm if_predI} 1
-           THEN print_tac "before sidecond:"
+           THEN print_tac' options "before sidecond:"
            THEN prove_sidecond thy modes t
-           THEN print_tac "after sidecond:"
+           THEN print_tac' options "after sidecond:"
            THEN prove_prems [] ps)
       in (prove_match thy out_ts)
           THEN rest_tac
@@ -1800,7 +1852,7 @@
            (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
              (1 upto (length moded_clauses))))
     THEN (EVERY (map2 (prove_clause options thy nargs modes mode) clauses moded_clauses))
-    THEN print_tac "proved one direction"
+    THEN print_tac' options "proved one direction"
   end;
 
 (** Proof in the other direction **)
@@ -2106,9 +2158,10 @@
             map SOME (all_smodes_of_typs Rs) | _ => [NONE]) Ts), all_smodes_of_typs Us)
       end
     val all_modes = map (fn (s, T) =>
-      case proposed_modes options s of
+      case proposed_modes options of
         NONE => (s, modes_of_typ T)
-      | SOME modes' => (s, map (translate_mode' nparams) modes'))
+      | SOME (s', modes') =>
+          if s = s' then (s, map (translate_mode' nparams) modes') else (s, modes_of_typ T))
         preds
   in (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) end;
 
@@ -2173,12 +2226,12 @@
             val args = map2 (curry Free) arg_names Ts
             val predfun = Const (predfun_name_of thy predname full_mode,
               Ts ---> PredicateCompFuns.mk_predT @{typ unit})
-            val rhs = PredicateCompFuns.mk_Eval (list_comb (predfun, args), @{term "Unity"})
+            val rhs = @{term Predicate.holds} $ (list_comb (predfun, args))
             val eq_term = HOLogic.mk_Trueprop
               (HOLogic.mk_eq (list_comb (Const (predname, T), args), rhs))
             val def = predfun_definition_of thy predname full_mode
             val tac = fn _ => Simplifier.simp_tac
-              (HOL_basic_ss addsimps [def, @{thm eval_pred}]) 1
+              (HOL_basic_ss addsimps [def, @{thm holds_eq}, @{thm eval_pred}]) 1
             val eq = Goal.prove (ProofContext.init thy) arg_names [] eq_term tac
           in
             (pred, result_thms @ [eq])
@@ -2199,7 +2252,7 @@
   define_functions : options -> (string * typ) list -> string * mode list -> theory -> theory,
   infer_modes : options -> theory -> (string * mode list) list -> (string * mode list) list
     -> string list -> (string * (term list * indprem list) list) list
-    -> moded_clause list pred_mode_table,
+    -> moded_clause list pred_mode_table * string list,
   prove : options -> theory -> (string * (term list * indprem list) list) list
     -> (string * typ) list -> (string * mode list) list
     -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
@@ -2220,10 +2273,11 @@
     val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) =
       prepare_intrs options thy prednames (maps (intros_of thy) prednames)
     val _ = print_step options "Infering modes..."
-    val moded_clauses =
+    val (moded_clauses, errors) =
       #infer_modes (dest_steps steps) options thy extra_modes all_modes param_vs clauses
     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
     val _ = check_expected_modes preds options modes
+    val _ = check_proposed_modes preds options modes extra_modes errors
     val _ = print_modes options thy modes
       (*val _ = print_moded_clauses thy moded_clauses*)
     val _ = print_step options "Defining executable functions..."
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -137,7 +137,7 @@
     th'
   end
 
-fun store_thm_in_table options ignore_consts thy th=
+fun store_thm_in_table options ignore thy th=
   let
     val th = th
       |> inline_equations options thy
@@ -153,29 +153,29 @@
       else if (is_introlike th) then (defining_const_of_introrule th, th)
       else error "store_thm: unexpected definition format"
   in
-    if not (member (op =) ignore_consts const) then
-      Symtab.cons_list (const, th)
-    else I
+    if ignore const then I else Symtab.cons_list (const, th)
   end
 
 fun make_const_spec_table options thy =
   let
-    fun store ignore_const f =
-      fold (store_thm_in_table options ignore_const thy)
+    fun store ignore f =
+      fold (store_thm_in_table options ignore thy)
         (map (Thm.transfer thy) (f (ProofContext.init thy)))
     val table = Symtab.empty
-      |> store [] Predicate_Compile_Alternative_Defs.get
-    val ignore_consts = Symtab.keys table
+      |> store (K false) Predicate_Compile_Alternative_Defs.get
+    val ignore = Symtab.defined table
   in
     table
-    |> store ignore_consts Predicate_Compile_Preproc_Const_Defs.get
-    |> store ignore_consts Nitpick_Simps.get
-    |> store ignore_consts Nitpick_Intros.get
+    |> store ignore (fn ctxt => maps
+      (fn (roughly, (ts, ths)) => if roughly = Spec_Rules.Equational then ths else [])
+        (Spec_Rules.get ctxt))
+    |> store ignore Nitpick_Simps.get
+    |> store ignore Nitpick_Intros.get
   end
 
 fun get_specification table constname =
   case Symtab.lookup table constname of
-    SOME thms => thms
+    SOME thms => thms                  
   | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
 
 val logic_operator_names =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -89,9 +89,7 @@
 
 fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
   | mk_param thy lookup_pred t =
-  let
-  val _ = tracing ("called param with " ^ (Syntax.string_of_term_global thy t))
-  in if Predicate_Compile_Aux.is_predT (fastype_of t) then
+  if Predicate_Compile_Aux.is_predT (fastype_of t) then
     t
   else
     let
@@ -109,7 +107,7 @@
       val pred_body = HOLogic.mk_eq (body', resvar)
       val param = fold_rev lambda (vs' @ [resvar]) pred_body
     in param end
-  end
+    
 (* creates the list of premises for every intro rule *)
 (* theory -> term -> (string list, term list list) *)
 
@@ -354,7 +352,7 @@
             val (ind_result, thy') =
               thy
               |> Sign.map_naming Name_Space.conceal
-              |> Inductive.add_inductive_global (serial ())
+              |> Inductive.add_inductive_global
                 {quiet_mode = false, verbose = false, alt_name = Binding.empty, coind = false,
                   no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
                 (map (fn (s, T) =>
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -23,7 +23,7 @@
 
 val options = Options {
   expected_modes = NONE,
-  proposed_modes = [],
+  proposed_modes = NONE,
   proposed_names = [],
   show_steps = true,
   show_intermediate_results = true,
@@ -83,7 +83,7 @@
     val intro = Goal.prove (ProofContext.init thy') (map fst vs') [] t tac
     val _ = tracing (Display.string_of_thm ctxt' intro)
     val thy'' = thy'
-      |> Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm intro)
+      |> Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro)
       |> Predicate_Compile.preprocess options full_constname
       |> Predicate_Compile_Core.add_equations options [full_constname]
       (*  |> Predicate_Compile_Core.add_depth_limited_equations Predicate_Compile_Aux.default_options [full_constname]*)
--- a/src/HOL/Tools/inductive.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/inductive.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -51,7 +51,7 @@
     (Attrib.binding * string) list ->
     (Facts.ref * Attrib.src list) list ->
     bool -> local_theory -> inductive_result * local_theory
-  val add_inductive_global: serial -> inductive_flags ->
+  val add_inductive_global: inductive_flags ->
     ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
     thm list -> theory -> inductive_result * theory
   val arities_of: thm -> (string * int) list
@@ -886,19 +886,17 @@
       coind = coind, no_elim = false, no_ind = false, skip_mono = false, fork_mono = not int};
   in
     lthy
-    |> Local_Theory.set_group (serial ())
     |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos
   end;
 
 val add_inductive_i = gen_add_inductive_i add_ind_def;
 val add_inductive = gen_add_inductive add_ind_def;
 
-fun add_inductive_global group flags cnames_syn pnames pre_intros monos thy =
+fun add_inductive_global flags cnames_syn pnames pre_intros monos thy =
   let
     val name = Sign.full_name thy (fst (fst (hd cnames_syn)));
     val ctxt' = thy
       |> Theory_Target.init NONE
-      |> Local_Theory.set_group group
       |> add_inductive_i flags cnames_syn pnames pre_intros monos |> snd
       |> Local_Theory.exit;
     val info = #2 (the_inductive ctxt' name);
--- a/src/HOL/Tools/inductive_realizer.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/inductive_realizer.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -350,7 +350,7 @@
     (** realizability predicate **)
 
     val (ind_info, thy3') = thy2 |>
-      Inductive.add_inductive_global (serial ())
+      Inductive.add_inductive_global
         {quiet_mode = false, verbose = false, alt_name = Binding.empty, coind = false,
           no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
         rlzpreds rlzparams (map (fn (rintr, intr) =>
--- a/src/HOL/Tools/lin_arith.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/lin_arith.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -438,7 +438,7 @@
       in
         SOME [(Ts, subgoal1), (split_type :: Ts, subgoal2)]
       end
-    (* ?P (nat ?i) = ((ALL n. ?i = int n --> ?P n) & (?i < 0 --> ?P 0)) *)
+    (* ?P (nat ?i) = ((ALL n. ?i = of_nat n --> ?P n) & (?i < 0 --> ?P 0)) *)
     | (Const ("Int.nat", _), [t1]) =>
       let
         val rev_terms   = rev terms
@@ -449,17 +449,17 @@
                             (map (incr_boundvars 1) rev_terms)
         val terms2      = map (subst_term [(split_term, zero_nat)]) rev_terms
         val t1'         = incr_boundvars 1 t1
-        val t1_eq_int_n = Const ("op =", HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $
+        val t1_eq_nat_n = Const ("op =", HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $
                             (Const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) $ n)
         val t1_lt_zero  = Const (@{const_name HOL.less},
                             HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1 $ zero_int
         val not_false   = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
-        val subgoal1    = (HOLogic.mk_Trueprop t1_eq_int_n) :: terms1 @ [not_false]
+        val subgoal1    = (HOLogic.mk_Trueprop t1_eq_nat_n) :: terms1 @ [not_false]
         val subgoal2    = (HOLogic.mk_Trueprop t1_lt_zero) :: terms2 @ [not_false]
       in
         SOME [(HOLogic.natT :: Ts, subgoal1), (Ts, subgoal2)]
       end
-    (* "?P ((?n::nat) mod (number_of ?k)) =
+    (* ?P ((?n::nat) mod (number_of ?k)) =
          ((number_of ?k = 0 --> ?P ?n) & (~ (number_of ?k = 0) -->
            (ALL i j. j < number_of ?k --> ?n = number_of ?k * i + j --> ?P j))) *)
     | (Const ("Divides.div_class.mod", Type ("fun", [Type ("nat", []), _])), [t1, t2]) =>
@@ -491,7 +491,7 @@
       in
         SOME [(Ts, subgoal1), (split_type :: split_type :: Ts, subgoal2)]
       end
-    (* "?P ((?n::nat) div (number_of ?k)) =
+    (* ?P ((?n::nat) div (number_of ?k)) =
          ((number_of ?k = 0 --> ?P 0) & (~ (number_of ?k = 0) -->
            (ALL i j. j < number_of ?k --> ?n = number_of ?k * i + j --> ?P i))) *)
     | (Const ("Divides.div_class.div", Type ("fun", [Type ("nat", []), _])), [t1, t2]) =>
@@ -523,14 +523,16 @@
       in
         SOME [(Ts, subgoal1), (split_type :: split_type :: Ts, subgoal2)]
       end
-    (* "?P ((?n::int) mod (number_of ?k)) =
-         ((iszero (number_of ?k) --> ?P ?n) &
-          (neg (number_of (uminus ?k)) -->
-            (ALL i j. 0 <= j & j < number_of ?k & ?n = number_of ?k * i + j --> ?P j)) &
-          (neg (number_of ?k) -->
-            (ALL i j. number_of ?k < j & j <= 0 & ?n = number_of ?k * i + j --> ?P j))) *)
+    (* ?P ((?n::int) mod (number_of ?k)) =
+         ((number_of ?k = 0 --> ?P ?n) &
+          (0 < number_of ?k -->
+            (ALL i j.
+              0 <= j & j < number_of ?k & ?n = number_of ?k * i + j --> ?P j)) &
+          (number_of ?k < 0 -->
+            (ALL i j.
+              number_of ?k < j & j <= 0 & ?n = number_of ?k * i + j --> ?P j))) *)
     | (Const ("Divides.div_class.mod",
-        Type ("fun", [Type ("Int.int", []), _])), [t1, t2 as (number_of $ k)]) =>
+        Type ("fun", [Type ("Int.int", []), _])), [t1, t2]) =>
       let
         val rev_terms               = rev terms
         val zero                    = Const (@{const_name HOL.zero}, split_type)
@@ -540,33 +542,33 @@
         val terms2_3                = map (subst_term [(incr_boundvars 2 split_term, j)])
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
-        val (t2' as (_ $ k'))       = incr_boundvars 2 t2
-        val iszero_t2               = Const ("Int.iszero", split_type --> HOLogic.boolT) $ t2
-        val neg_minus_k             = Const ("Int.neg", split_type --> HOLogic.boolT) $
-                                        (number_of $
-                                          (Const (@{const_name HOL.uminus},
-                                            HOLogic.intT --> HOLogic.intT) $ k'))
+        val t2'                     = incr_boundvars 2 t2
+        val t2_eq_zero              = Const ("op =",
+                                        split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
+        val zero_lt_t2              = Const (@{const_name HOL.less},
+                                        split_type --> split_type --> HOLogic.boolT) $ zero $ t2'
+        val t2_lt_zero              = Const (@{const_name HOL.less},
+                                        split_type --> split_type --> HOLogic.boolT) $ t2' $ zero
         val zero_leq_j              = Const (@{const_name HOL.less_eq},
                                         split_type --> split_type --> HOLogic.boolT) $ zero $ j
+        val j_leq_zero              = Const (@{const_name HOL.less_eq},
+                                        split_type --> split_type --> HOLogic.boolT) $ j $ zero
         val j_lt_t2                 = Const (@{const_name HOL.less},
                                         split_type --> split_type--> HOLogic.boolT) $ j $ t2'
+        val t2_lt_j                 = Const (@{const_name HOL.less},
+                                        split_type --> split_type--> HOLogic.boolT) $ t2' $ j
         val t1_eq_t2_times_i_plus_j = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
                                        (Const (@{const_name HOL.plus}, split_type --> split_type --> split_type) $
                                          (Const (@{const_name HOL.times},
                                            split_type --> split_type --> split_type) $ t2' $ i) $ j)
-        val neg_t2                  = Const ("Int.neg", split_type --> HOLogic.boolT) $ t2'
-        val t2_lt_j                 = Const (@{const_name HOL.less},
-                                        split_type --> split_type--> HOLogic.boolT) $ t2' $ j
-        val j_leq_zero              = Const (@{const_name HOL.less_eq},
-                                        split_type --> split_type --> HOLogic.boolT) $ j $ zero
         val not_false               = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
-        val subgoal1                = (HOLogic.mk_Trueprop iszero_t2) :: terms1 @ [not_false]
-        val subgoal2                = (map HOLogic.mk_Trueprop [neg_minus_k, zero_leq_j])
+        val subgoal1                = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false]
+        val subgoal2                = (map HOLogic.mk_Trueprop [zero_lt_t2, zero_leq_j])
                                         @ hd terms2_3
                                         :: (if tl terms2_3 = [] then [not_false] else [])
                                         @ (map HOLogic.mk_Trueprop [j_lt_t2, t1_eq_t2_times_i_plus_j])
                                         @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false])
-        val subgoal3                = (map HOLogic.mk_Trueprop [neg_t2, t2_lt_j])
+        val subgoal3                = (map HOLogic.mk_Trueprop [t2_lt_zero, t2_lt_j])
                                         @ hd terms2_3
                                         :: (if tl terms2_3 = [] then [not_false] else [])
                                         @ (map HOLogic.mk_Trueprop [j_leq_zero, t1_eq_t2_times_i_plus_j])
@@ -575,14 +577,16 @@
       in
         SOME [(Ts, subgoal1), (Ts', subgoal2), (Ts', subgoal3)]
       end
-    (* "?P ((?n::int) div (number_of ?k)) =
-         ((iszero (number_of ?k) --> ?P 0) &
-          (neg (number_of (uminus ?k)) -->
-            (ALL i. (EX j. 0 <= j & j < number_of ?k & ?n = number_of ?k * i + j) --> ?P i)) &
-          (neg (number_of ?k) -->
-            (ALL i. (EX j. number_of ?k < j & j <= 0 & ?n = number_of ?k * i + j) --> ?P i))) *)
+    (* ?P ((?n::int) div (number_of ?k)) =
+         ((number_of ?k = 0 --> ?P 0) &
+          (0 < number_of ?k -->
+            (ALL i j.
+              0 <= j & j < number_of ?k & ?n = number_of ?k * i + j --> ?P i)) &
+          (number_of ?k < 0 -->
+            (ALL i j.
+              number_of ?k < j & j <= 0 & ?n = number_of ?k * i + j --> ?P i))) *)
     | (Const ("Divides.div_class.div",
-        Type ("fun", [Type ("Int.int", []), _])), [t1, t2 as (number_of $ k)]) =>
+        Type ("fun", [Type ("Int.int", []), _])), [t1, t2]) =>
       let
         val rev_terms               = rev terms
         val zero                    = Const (@{const_name HOL.zero}, split_type)
@@ -592,38 +596,37 @@
         val terms2_3                = map (subst_term [(incr_boundvars 2 split_term, i)])
                                         (map (incr_boundvars 2) rev_terms)
         val t1'                     = incr_boundvars 2 t1
-        val (t2' as (_ $ k'))       = incr_boundvars 2 t2
-        val iszero_t2               = Const ("Int.iszero", split_type --> HOLogic.boolT) $ t2
-        val neg_minus_k             = Const ("Int.neg", split_type --> HOLogic.boolT) $
-                                        (number_of $
-                                          (Const (@{const_name HOL.uminus},
-                                            HOLogic.intT --> HOLogic.intT) $ k'))
+        val t2'                     = incr_boundvars 2 t2
+        val t2_eq_zero              = Const ("op =",
+                                        split_type --> split_type --> HOLogic.boolT) $ t2 $ zero
+        val zero_lt_t2              = Const (@{const_name HOL.less},
+                                        split_type --> split_type --> HOLogic.boolT) $ zero $ t2'
+        val t2_lt_zero              = Const (@{const_name HOL.less},
+                                        split_type --> split_type --> HOLogic.boolT) $ t2' $ zero
         val zero_leq_j              = Const (@{const_name HOL.less_eq},
                                         split_type --> split_type --> HOLogic.boolT) $ zero $ j
+        val j_leq_zero              = Const (@{const_name HOL.less_eq},
+                                        split_type --> split_type --> HOLogic.boolT) $ j $ zero
         val j_lt_t2                 = Const (@{const_name HOL.less},
                                         split_type --> split_type--> HOLogic.boolT) $ j $ t2'
-        val t1_eq_t2_times_i_plus_j = Const ("op =",
-                                        split_type --> split_type --> HOLogic.boolT) $ t1' $
+        val t2_lt_j                 = Const (@{const_name HOL.less},
+                                        split_type --> split_type--> HOLogic.boolT) $ t2' $ j
+        val t1_eq_t2_times_i_plus_j = Const ("op =", split_type --> split_type --> HOLogic.boolT) $ t1' $
                                        (Const (@{const_name HOL.plus}, split_type --> split_type --> split_type) $
                                          (Const (@{const_name HOL.times},
                                            split_type --> split_type --> split_type) $ t2' $ i) $ j)
-        val neg_t2                  = Const ("Int.neg", split_type --> HOLogic.boolT) $ t2'
-        val t2_lt_j                 = Const (@{const_name HOL.less},
-                                        split_type --> split_type--> HOLogic.boolT) $ t2' $ j
-        val j_leq_zero              = Const (@{const_name HOL.less_eq},
-                                        split_type --> split_type --> HOLogic.boolT) $ j $ zero
         val not_false               = HOLogic.mk_Trueprop (HOLogic.Not $ HOLogic.false_const)
-        val subgoal1                = (HOLogic.mk_Trueprop iszero_t2) :: terms1 @ [not_false]
-        val subgoal2                = (HOLogic.mk_Trueprop neg_minus_k)
-                                        :: terms2_3
-                                        @ not_false
-                                        :: (map HOLogic.mk_Trueprop
-                                             [zero_leq_j, j_lt_t2, t1_eq_t2_times_i_plus_j])
-        val subgoal3                = (HOLogic.mk_Trueprop neg_t2)
-                                        :: terms2_3
-                                        @ not_false
-                                        :: (map HOLogic.mk_Trueprop
-                                             [t2_lt_j, j_leq_zero, t1_eq_t2_times_i_plus_j])
+        val subgoal1                = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false]
+        val subgoal2                = (map HOLogic.mk_Trueprop [zero_lt_t2, zero_leq_j])
+                                        @ hd terms2_3
+                                        :: (if tl terms2_3 = [] then [not_false] else [])
+                                        @ (map HOLogic.mk_Trueprop [j_lt_t2, t1_eq_t2_times_i_plus_j])
+                                        @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false])
+        val subgoal3                = (map HOLogic.mk_Trueprop [t2_lt_zero, t2_lt_j])
+                                        @ hd terms2_3
+                                        :: (if tl terms2_3 = [] then [not_false] else [])
+                                        @ (map HOLogic.mk_Trueprop [j_leq_zero, t1_eq_t2_times_i_plus_j])
+                                        @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false])
         val Ts'                     = split_type :: split_type :: Ts
       in
         SOME [(Ts, subgoal1), (Ts', subgoal2), (Ts', subgoal3)]
@@ -886,12 +889,12 @@
      (l <= min m n + k) = (l <= m+k & l <= n+k)
   *)
   refute_tac (K true)
-    (* Splitting is also done inside simple_tac, but not completely --   *)
-    (* split_tac may use split theorems that have not been implemented in    *)
-    (* simple_tac (cf. pre_decomp and split_once_items above), and       *)
-    (* split_limit may trigger.                                   *)
-    (* Therefore splitting outside of simple_tac may allow us to prove   *)
-    (* some goals that simple_tac alone would fail on.                   *)
+    (* Splitting is also done inside simple_tac, but not completely --    *)
+    (* split_tac may use split theorems that have not been implemented in *)
+    (* simple_tac (cf. pre_decomp and split_once_items above), and        *)
+    (* split_limit may trigger.                                           *)
+    (* Therefore splitting outside of simple_tac may allow us to prove    *)
+    (* some goals that simple_tac alone would fail on.                    *)
     (REPEAT_DETERM o split_tac (#splits (get_arith_data ctxt)))
     (lin_arith_tac ctxt ex);
 
--- a/src/HOL/Tools/primrec.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/primrec.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -198,7 +198,7 @@
       (list_comb (Const (rec_name, dummyT), fs @ map Bound (0 :: (length ls downto 1))))
     val rhs = singleton (Syntax.check_terms ctxt)
       (TypeInfer.constrain varT raw_rhs);
-  in (var, ((Binding.name def_name, []), rhs)) end;
+  in (var, ((Binding.conceal (Binding.name def_name), []), rhs)) end;
 
 
 (* find datatypes which contain all datatypes in tnames' *)
@@ -265,7 +265,7 @@
 
 local
 
-fun gen_primrec set_group prep_spec raw_fixes raw_spec lthy =
+fun gen_primrec prep_spec raw_fixes raw_spec lthy =
   let
     val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
     fun attr_bindings prefix = map (fn ((b, attrs), _) =>
@@ -275,7 +275,6 @@
         map (Attrib.internal o K) [Simplifier.simp_add, Nitpick_Simps.add]);
   in
     lthy
-    |> set_group ? Local_Theory.set_group (serial ())
     |> add_primrec_simple fixes (map snd spec)
     |-> (fn (prefix, simps) => fold_map Local_Theory.note (attr_bindings prefix ~~ simps)
       #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps')))
@@ -284,8 +283,8 @@
 
 in
 
-val add_primrec = gen_primrec false Specification.check_spec;
-val add_primrec_cmd = gen_primrec true Specification.read_spec;
+val add_primrec = gen_primrec Specification.check_spec;
+val add_primrec_cmd = gen_primrec Specification.read_spec;
 
 end;
 
--- a/src/HOL/Tools/record.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/record.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -1009,14 +1009,20 @@
 (** record simprocs **)
 
 fun future_forward_prf_standard thy prf prop () =
-   let val thm = if !quick_and_dirty then Skip_Proof.make_thm thy prop 
-                 else Goal.future_result (ProofContext.init thy) (Future.fork_pri ~1 prf) prop;
-   in Drule.standard thm end;
+  let val thm =
+    if ! quick_and_dirty then Skip_Proof.make_thm thy prop
+    else if Goal.future_enabled () then
+      Goal.future_result (ProofContext.init thy) (Future.fork_pri ~1 prf) prop
+    else prf ()
+  in Drule.standard thm end;
 
 fun prove_common immediate stndrd thy asms prop tac =
-  let val prv = if !quick_and_dirty then Skip_Proof.prove 
-                else if immediate then Goal.prove else Goal.prove_future;
-      val prf = prv (ProofContext.init thy) [] asms prop tac
+  let
+    val prv =
+      if ! quick_and_dirty then Skip_Proof.prove
+      else if immediate orelse not (Goal.future_enabled ()) then Goal.prove
+      else Goal.prove_future;
+    val prf = prv (ProofContext.init thy) [] asms prop tac;
   in if stndrd then Drule.standard prf else prf end;
 
 val prove_future_global = prove_common false;
--- a/src/HOL/Tools/rewrite_hol_proof.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/Tools/rewrite_hol_proof.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -7,7 +7,7 @@
 signature REWRITE_HOL_PROOF =
 sig
   val rews: (Proofterm.proof * Proofterm.proof) list
-  val elim_cong: typ list -> Proofterm.proof -> Proofterm.proof option
+  val elim_cong: typ list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option
 end;
 
 structure RewriteHOLProof : REWRITE_HOL_PROOF =
@@ -309,17 +309,19 @@
 
 fun mk_AbsP P t = AbsP ("H", Option.map HOLogic.mk_Trueprop P, t);
 
-fun elim_cong Ts (PThm (_, (("HOL.iffD1", _, _), _)) % _ % _ %% prf1 %% prf2) =
+fun elim_cong_aux Ts (PThm (_, (("HOL.iffD1", _, _), _)) % _ % _ %% prf1 %% prf2) =
       Option.map (make_subst Ts prf2 []) (strip_cong [] prf1)
-  | elim_cong Ts (PThm (_, (("HOL.iffD1", _, _), _)) % P % _ %% prf) =
+  | elim_cong_aux Ts (PThm (_, (("HOL.iffD1", _, _), _)) % P % _ %% prf) =
       Option.map (mk_AbsP P o make_subst Ts (PBound 0) [])
         (strip_cong [] (incr_pboundvars 1 0 prf))
-  | elim_cong Ts (PThm (_, (("HOL.iffD2", _, _), _)) % _ % _ %% prf1 %% prf2) =
+  | elim_cong_aux Ts (PThm (_, (("HOL.iffD2", _, _), _)) % _ % _ %% prf1 %% prf2) =
       Option.map (make_subst Ts prf2 [] o
         apsnd (map (make_sym Ts))) (strip_cong [] prf1)
-  | elim_cong Ts (PThm (_, (("HOL.iffD2", _, _), _)) % _ % P %% prf) =
+  | elim_cong_aux Ts (PThm (_, (("HOL.iffD2", _, _), _)) % _ % P %% prf) =
       Option.map (mk_AbsP P o make_subst Ts (PBound 0) [] o
         apsnd (map (make_sym Ts))) (strip_cong [] (incr_pboundvars 1 0 prf))
-  | elim_cong _ _ = NONE;
+  | elim_cong_aux _ _ = NONE;
+
+fun elim_cong Ts prf = Option.map (rpair no_skel) (elim_cong_aux Ts prf);
 
 end;
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Thu Nov 19 06:01:02 2009 -0800
@@ -87,6 +87,9 @@
 values "{(c, a, b). JamesBond a b c}"
 values "{(c, b, a). JamesBond a b c}"
 
+values "{(a, b). JamesBond 0 b a}"
+values "{(c, a). JamesBond a 0 c}"
+values "{(a, c). JamesBond a 0 c}"
 
 subsection {* Alternative Rules *}
 
@@ -476,6 +479,8 @@
 
 subsection {* transitive predicate *}
 
+text {* Also look at the tabled transitive closure in the Library *}
+
 code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,
   (o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool,
   (o => o => bool) => o => i => bool, (o => o => bool) => o => o => bool) tranclp
@@ -509,6 +514,28 @@
 values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
 values 20 "{(n, m). tranclp succ n m}"
 
+inductive example_graph :: "int => int => bool"
+where
+  "example_graph 0 1"
+| "example_graph 1 2"
+| "example_graph 1 3"
+| "example_graph 4 7"
+| "example_graph 4 5"
+| "example_graph 5 6"
+| "example_graph 7 6"
+| "example_graph 7 8"
+ 
+inductive not_reachable_in_example_graph :: "int => int => bool"
+where "\<not> (tranclp example_graph x y) ==> not_reachable_in_example_graph x y"
+
+code_pred (expected_modes: i => i => bool) not_reachable_in_example_graph .
+
+thm not_reachable_in_example_graph.equation
+
+value "not_reachable_in_example_graph 0 3"
+value "not_reachable_in_example_graph 4 8"
+value "not_reachable_in_example_graph 5 6"
+
 subsection {* IMP *}
 
 types
@@ -724,7 +751,7 @@
 
 subsection {* Inverting list functions *}
 
-code_pred [inductify, show_intermediate_results] length .
+code_pred [inductify] length .
 code_pred [inductify, random] length .
 thm size_listP.equation
 thm size_listP.random_equation
@@ -804,8 +831,6 @@
 thm spliceP.equation
 
 values "{xs. spliceP xs [1, 2, 3] [1, 1, 1, 2, 1, (3::nat)]}"
-(* TODO: correct error messages:*)
-(* values {(xs, ys, zs). spliceP xs ... } *)
 
 code_pred [inductify] List.rev .
 code_pred [inductify] map .
@@ -965,13 +990,8 @@
 | objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)"
 | plus: "\<lbrakk> eval_var l conf (IntVal vl); eval_var r conf (IntVal vr) \<rbrakk> \<Longrightarrow> eval_var (Add l r) conf (IntVal (vl+vr))"
 
-(* TODO: breaks if code_pred_intro is used? *)
-(*
-lemmas [code_pred_intro] = irconst objaddr plus
-*)
-thm eval_var.cases
 
-code_pred eval_var . (*by (rule eval_var.cases)*)
+code_pred eval_var .
 thm eval_var.equation
 
 values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
--- a/src/HOLCF/Tools/fixrec.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/HOLCF/Tools/fixrec.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -421,7 +421,6 @@
 (* code adapted from HOL/Tools/primrec.ML *)
 
 fun gen_fixrec
-  (set_group : bool)
   prep_spec
   (strict : bool)
   raw_fixes
@@ -473,8 +472,8 @@
 
 in
 
-val add_fixrec = gen_fixrec false Specification.check_spec;
-val add_fixrec_cmd = gen_fixrec true Specification.read_spec;
+val add_fixrec = gen_fixrec Specification.check_spec;
+val add_fixrec_cmd = gen_fixrec Specification.read_spec;
 
 end; (* local *)
 
--- a/src/Pure/General/name_space.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/General/name_space.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -34,8 +34,11 @@
   type naming
   val default_naming: naming
   val conceal: naming -> naming
-  val set_group: serial -> naming -> naming
+  val get_group: naming -> serial option
+  val set_group: serial option -> naming -> naming
   val set_theory_name: string -> naming -> naming
+  val new_group: naming -> naming
+  val reset_group: naming -> naming
   val add_path: string -> naming -> naming
   val root_path: naming -> naming
   val parent_path: naming -> naming
@@ -241,12 +244,18 @@
 val conceal = map_naming (fn (_, group, theory_name, path) =>
   (true, group, theory_name, path));
 
-fun set_group group = map_naming (fn (conceal, _, theory_name, path) =>
-  (conceal, SOME group, theory_name, path));
-
 fun set_theory_name theory_name = map_naming (fn (conceal, group, _, path) =>
   (conceal, group, theory_name, path));
 
+
+fun get_group (Naming {group, ...}) = group;
+
+fun set_group group = map_naming (fn (conceal, _, theory_name, path) =>
+  (conceal, group, theory_name, path));
+
+fun new_group naming = set_group (SOME (serial ())) naming;
+val reset_group = set_group NONE;
+
 fun add_path elems = map_path (fn path => path @ [(elems, false)]);
 val root_path = map_path (fn _ => []);
 val parent_path = map_path (perhaps (try (#1 o split_last)));
--- a/src/Pure/Isar/code.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Isar/code.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -775,49 +775,4 @@
 
 end;
 
-(** datastructure to log definitions for preprocessing of the predicate compiler **)
-(* obviously a clone of Named_Thms *)
-
-signature PREDICATE_COMPILE_PREPROC_CONST_DEFS =
-sig
-  val get: Proof.context -> thm list
-  val add_thm: thm -> Context.generic -> Context.generic
-  val del_thm: thm -> Context.generic -> Context.generic
-  
-  val add_attribute : attribute
-  val del_attribute : attribute
-  
-  val add_attrib : Attrib.src
-  
-  val setup: theory -> theory
-end;
-
-structure Predicate_Compile_Preproc_Const_Defs : PREDICATE_COMPILE_PREPROC_CONST_DEFS =
-struct
-
-structure Data = Generic_Data
-(
-  type T = thm list;
-  val empty = [];
-  val extend = I;
-  val merge = Thm.merge_thms;
-);
-
-val get = Data.get o Context.Proof;
-
-val add_thm = Data.map o Thm.add_thm;
-val del_thm = Data.map o Thm.del_thm;
-
-val add_attribute = Thm.declaration_attribute add_thm;
-val del_attribute = Thm.declaration_attribute del_thm;
-
-val add_attrib = Attrib.internal (K add_attribute)
-
-val setup =
-  Attrib.setup (Binding.name "pred_compile_preproc") (Attrib.add_del add_attribute del_attribute)
-    ("declaration of definition for preprocessing of the predicate compiler") #>
-  PureThy.add_thms_dynamic (Binding.name "pred_compile_preproc", Data.get);
-
-end;
-
 structure Code : CODE = struct open Code; end;
--- a/src/Pure/Isar/constdefs.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Isar/constdefs.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -55,8 +55,7 @@
       |> PureThy.add_defs false [((name, def), atts)]
       |-> (fn [thm] =>  (* FIXME thm vs. atts !? *)
         Spec_Rules.add_global Spec_Rules.Equational ([Logic.varify head], [thm]) #>
-        Code.add_default_eqn thm #>
-        Context.theory_map (Predicate_Compile_Preproc_Const_Defs.add_thm thm));
+        Code.add_default_eqn thm);
   in ((c, T), thy') end;
 
 fun gen_constdefs prep_vars prep_prop prep_att (raw_structs, specs) thy =
--- a/src/Pure/Isar/local_theory.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Isar/local_theory.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -14,7 +14,8 @@
   val full_name: local_theory -> binding -> string
   val map_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory
   val conceal: local_theory -> local_theory
-  val set_group: serial -> local_theory -> local_theory
+  val new_group: local_theory -> local_theory
+  val reset_group: local_theory -> local_theory
   val restore_naming: local_theory -> local_theory -> local_theory
   val target_of: local_theory -> Proof.context
   val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory
@@ -42,7 +43,7 @@
   val term_syntax: bool -> declaration -> local_theory -> local_theory
   val declaration: bool -> declaration -> local_theory -> local_theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
-  val init: string -> operations -> Proof.context -> local_theory
+  val init: serial option -> string -> operations -> Proof.context -> local_theory
   val restore: local_theory -> local_theory
   val reinit: local_theory -> local_theory
   val exit: local_theory -> Proof.context
@@ -114,7 +115,8 @@
   (f naming, theory_prefix, operations, target));
 
 val conceal = map_naming Name_Space.conceal;
-val set_group = map_naming o Name_Space.set_group;
+val new_group = map_naming Name_Space.new_group;
+val reset_group = map_naming Name_Space.reset_group;
 
 val restore_naming = map_naming o K o naming_of;
 
@@ -203,9 +205,10 @@
 
 (* init *)
 
-fun init theory_prefix operations target =
+fun init group theory_prefix operations target =
   let val naming =
     Sign.naming_of (ProofContext.theory_of target)
+    |> Name_Space.set_group group
     |> Name_Space.mandatory_path theory_prefix;
   in
     target |> Data.map
@@ -215,8 +218,8 @@
   end;
 
 fun restore lthy =
-  let val {theory_prefix, operations, target, ...} = get_lthy lthy
-  in init theory_prefix operations target end;
+  let val {naming, theory_prefix, operations, target} = get_lthy lthy
+  in init (Name_Space.get_group naming) theory_prefix operations target end;
 
 val reinit = checkpoint o operation #reinit;
 
--- a/src/Pure/Isar/specification.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Isar/specification.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -212,8 +212,7 @@
 
     val ([(def_name, [th'])], lthy4) = lthy3
       |> Local_Theory.notes_kind Thm.definitionK
-        [((name, Predicate_Compile_Preproc_Const_Defs.add_attrib ::
-            Code.add_default_eqn_attrib :: atts), [([th], [])])];
+        [((name, Code.add_default_eqn_attrib :: atts), [([th], [])])];
 
     val lhs' = Morphism.term (Local_Theory.target_morphism lthy4) lhs;
     val _ =
--- a/src/Pure/Isar/theory_target.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Isar/theory_target.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -324,7 +324,7 @@
 
 fun init_lthy (ta as Target {target, instantiation, overloading, ...}) =
   Data.put ta #>
-  Local_Theory.init (Long_Name.base_name target)
+  Local_Theory.init NONE (Long_Name.base_name target)
    {pretty = pretty ta,
     abbrev = abbrev ta,
     define = define ta,
--- a/src/Pure/Isar/toplevel.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Isar/toplevel.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -302,7 +302,8 @@
 local
 
 fun apply_tr int (Init (_, f, exit)) (State (NONE, _)) =
-      State (SOME (Theory (Context.Theory (Theory.checkpoint (f int)), NONE), exit), NONE)
+      State (SOME (Theory (Context.Theory
+          (Theory.checkpoint (Runtime.controlled_execution f int)), NONE), exit), NONE)
   | apply_tr _ Exit (State (prev as SOME (Theory (Context.Theory _, _), _), _)) =
       State (NONE, prev)
   | apply_tr _ CommitExit (State (NONE, SOME (Theory (Context.Theory thy, _), exit))) =
@@ -419,7 +420,13 @@
     | _ => raise UNDEF));
 
 fun theory' f = transaction (fn int =>
-  (fn Theory (Context.Theory thy, _) => Theory (Context.Theory (f int thy), NONE)
+  (fn Theory (Context.Theory thy, _) =>
+      let val thy' = thy
+        |> Sign.new_group
+        |> Theory.checkpoint
+        |> f int
+        |> Sign.reset_group;
+      in Theory (Context.Theory thy', NONE) end
     | _ => raise UNDEF));
 
 fun theory f = theory' (K f);
@@ -442,7 +449,10 @@
   (fn Theory (gthy, _) =>
         let
           val finish = loc_finish loc gthy;
-          val lthy' = f int (loc_begin loc gthy);
+          val lthy' = loc_begin loc gthy
+            |> Local_Theory.new_group
+            |> f int
+            |> Local_Theory.reset_group;
         in Theory (finish lthy', SOME lthy') end
     | _ => raise UNDEF));
 
@@ -491,14 +501,14 @@
 in
 
 fun local_theory_to_proof' loc f = begin_proof
-  (fn int => fn gthy => f int (loc_begin loc gthy))
-  (loc_finish loc);
+  (fn int => fn gthy => f int (Local_Theory.new_group (loc_begin loc gthy)))
+  (fn gthy => loc_finish loc gthy o Local_Theory.reset_group);
 
 fun local_theory_to_proof loc f = local_theory_to_proof' loc (K f);
 
 fun theory_to_proof f = begin_proof
-  (K (fn Context.Theory thy => f thy | _ => raise UNDEF))
-  (K (Context.Theory o ProofContext.theory_of));
+  (K (fn Context.Theory thy => f (Theory.checkpoint (Sign.new_group thy)) | _ => raise UNDEF))
+  (K (Context.Theory o Sign.reset_group o ProofContext.theory_of));
 
 end;
 
@@ -531,7 +541,7 @@
 
 fun skip_proof_to_theory pred = transaction (fn _ =>
   (fn SkipProof (d, (gthy, _)) => if pred d then Theory (gthy, NONE) else raise UNDEF
-  | _ => raise UNDEF));
+    | _ => raise UNDEF));
 
 
 
--- a/src/Pure/Proof/proof_rewrite_rules.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/Proof/proof_rewrite_rules.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -6,8 +6,9 @@
 
 signature PROOF_REWRITE_RULES =
 sig
-  val rew : bool -> typ list -> Proofterm.proof -> Proofterm.proof option
-  val rprocs : bool -> (typ list -> Proofterm.proof -> Proofterm.proof option) list
+  val rew : bool -> typ list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option
+  val rprocs : bool ->
+    (typ list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option) list
   val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof
   val elim_defs : theory -> bool -> thm list -> Proofterm.proof -> Proofterm.proof
   val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof
@@ -183,7 +184,7 @@
         end
 
       | rew' _ = NONE;
-  in rew' end;
+  in rew' #> Option.map (rpair no_skel) end;
 
 fun rprocs b = [rew b];
 val _ = Context.>> (Context.map_theory (fold Proofterm.add_prf_rproc (rprocs false)));
--- a/src/Pure/defs.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/defs.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -10,10 +10,10 @@
   val pretty_const: Pretty.pp -> string * typ list -> Pretty.T
   val plain_args: typ list -> bool
   type T
-  val all_specifications_of: T -> (string * {def: string option, description: string,
-    lhs: typ list, rhs: (string * typ list) list} list) list
-  val specifications_of: T -> string -> {def: string option, description: string,
-    lhs: typ list, rhs: (string * typ list) list} list
+  type spec =
+    {def: string option, description: string, lhs: typ list, rhs: (string * typ list) list}
+  val all_specifications_of: T -> (string * spec list) list
+  val specifications_of: T -> string -> spec list
   val dest: T ->
    {restricts: ((string * typ list) * string) list,
     reducts: ((string * typ list) * (string * typ list) list) list}
--- a/src/Pure/item_net.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/item_net.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -40,12 +40,14 @@
 
 (* standard operations *)
 
-fun member (Items {eq, index, net, ...}) x =
-  exists (fn t => exists (fn (_, y) => eq (x, y)) (Net.unify_term net t)) (index x);
+fun member (Items {eq, index, content, net, ...}) x =
+  (case index x of
+    [] => Library.member eq content x
+  | t :: _ => exists (fn (_, y) => eq (x, y)) (Net.unify_term net t));
 
 fun cons x (Items {eq, index, content, next, net}) =
   mk_items eq index (x :: content) (next - 1)
-    (fold (fn t => Net.insert_term_safe (eq o pairself #2) (t, (next, x))) (index x) net);
+    (fold (fn t => Net.insert_term (K false) (t, (next, x))) (index x) net);
 
 
 fun merge (items1, Items {content = content2, ...}) =
--- a/src/Pure/more_thm.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/more_thm.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -75,9 +75,6 @@
   val untag_rule: string -> thm -> thm
   val tag: Properties.property -> attribute
   val untag: string -> attribute
-  val position_of: thm -> Position.T
-  val default_position: Position.T -> thm -> thm
-  val default_position_of: thm -> thm -> thm
   val def_name: string -> string
   val def_name_optional: string -> string -> string
   val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding
@@ -380,14 +377,6 @@
 fun untag s x = rule_attribute (K (untag_rule s)) x;
 
 
-(* position *)
-
-val position_of = Position.of_properties o Thm.get_tags;
-
-fun default_position pos = Thm.map_tags (Position.default_properties pos);
-val default_position_of = default_position o position_of;
-
-
 (* def_name *)
 
 fun def_name c = c ^ "_def";
--- a/src/Pure/proofterm.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/proofterm.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -109,18 +109,20 @@
   val axm_proof: string -> term -> proof
   val oracle_proof: string -> term -> oracle * proof
   val promise_proof: theory -> serial -> term -> proof
-  val fulfill_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
+  val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body
   val thm_proof: theory -> string -> term list -> term ->
     (serial * proof_body future) list -> proof_body -> pthm * proof
   val get_name: term list -> term -> proof -> string
 
   (** rewriting on proof terms **)
   val add_prf_rrule: proof * proof -> theory -> theory
-  val add_prf_rproc: (typ list -> proof -> proof option) -> theory -> theory
+  val add_prf_rproc: (typ list -> proof -> (proof * proof) option) -> theory -> theory
+  val no_skel: proof
+  val normal_skel: proof
   val rewrite_proof: theory -> (proof * proof) list *
-    (typ list -> proof -> proof option) list -> proof -> proof
+    (typ list -> proof -> (proof * proof) option) list -> proof -> proof
   val rewrite_proof_notypes: (proof * proof) list *
-    (typ list -> proof -> proof option) list -> proof -> proof
+    (typ list -> proof -> (proof * proof) option) list -> proof -> proof
   val rew_proof: theory -> proof -> proof
 end
 
@@ -1169,28 +1171,30 @@
 
 (**** rewriting on proof terms ****)
 
-val skel0 = PBound 0;
+val no_skel = PBound 0;
+val normal_skel = Hyp (Var ((Name.uu, 0), propT));
 
 fun rewrite_prf tymatch (rules, procs) prf =
   let
-    fun rew _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, skel0)
-      | rew _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, skel0)
-      | rew Ts prf = (case get_first (fn r => r Ts prf) procs of
-          SOME prf' => SOME (prf', skel0)
-        | NONE => get_first (fn (prf1, prf2) => SOME (prf_subst
-            (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2)
-               handle PMatch => NONE) (filter (could_unify prf o fst) rules));
+    fun rew _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, no_skel)
+      | rew _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, no_skel)
+      | rew Ts prf =
+          (case get_first (fn r => r Ts prf) procs of
+            NONE => get_first (fn (prf1, prf2) => SOME (prf_subst
+              (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2)
+                 handle PMatch => NONE) (filter (could_unify prf o fst) rules)
+          | some => some);
 
     fun rew0 Ts (prf as AbsP (_, _, prf' %% PBound 0)) =
           if prf_loose_Pbvar1 prf' 0 then rew Ts prf
           else
             let val prf'' = incr_pboundvars (~1) 0 prf'
-            in SOME (the_default (prf'', skel0) (rew Ts prf'')) end
+            in SOME (the_default (prf'', no_skel) (rew Ts prf'')) end
       | rew0 Ts (prf as Abst (_, _, prf' % SOME (Bound 0))) =
           if prf_loose_bvar1 prf' 0 then rew Ts prf
           else
             let val prf'' = incr_pboundvars 0 (~1) prf'
-            in SOME (the_default (prf'', skel0) (rew Ts prf'')) end
+            in SOME (the_default (prf'', no_skel) (rew Ts prf'')) end
       | rew0 Ts prf = rew Ts prf;
 
     fun rew1 _ (Hyp (Var _)) _ = NONE
@@ -1205,20 +1209,20 @@
     and rew2 Ts skel (prf % SOME t) = (case prf of
             Abst (_, _, body) =>
               let val prf' = prf_subst_bounds [t] body
-              in SOME (the_default prf' (rew2 Ts skel0 prf')) end
-          | _ => (case rew1 Ts (case skel of skel' % _ => skel' | _ => skel0) prf of
+              in SOME (the_default prf' (rew2 Ts no_skel prf')) end
+          | _ => (case rew1 Ts (case skel of skel' % _ => skel' | _ => no_skel) prf of
               SOME prf' => SOME (prf' % SOME t)
             | NONE => NONE))
       | rew2 Ts skel (prf % NONE) = Option.map (fn prf' => prf' % NONE)
-          (rew1 Ts (case skel of skel' % _ => skel' | _ => skel0) prf)
+          (rew1 Ts (case skel of skel' % _ => skel' | _ => no_skel) prf)
       | rew2 Ts skel (prf1 %% prf2) = (case prf1 of
             AbsP (_, _, body) =>
               let val prf' = prf_subst_pbounds [prf2] body
-              in SOME (the_default prf' (rew2 Ts skel0 prf')) end
+              in SOME (the_default prf' (rew2 Ts no_skel prf')) end
           | _ =>
             let val (skel1, skel2) = (case skel of
                 skel1 %% skel2 => (skel1, skel2)
-              | _ => (skel0, skel0))
+              | _ => (no_skel, no_skel))
             in case rew1 Ts skel1 prf1 of
                 SOME prf1' => (case rew1 Ts skel2 prf2 of
                     SOME prf2' => SOME (prf1' %% prf2')
@@ -1228,16 +1232,16 @@
                   | NONE => NONE)
             end)
       | rew2 Ts skel (Abst (s, T, prf)) = (case rew1 (the_default dummyT T :: Ts)
-              (case skel of Abst (_, _, skel') => skel' | _ => skel0) prf of
+              (case skel of Abst (_, _, skel') => skel' | _ => no_skel) prf of
             SOME prf' => SOME (Abst (s, T, prf'))
           | NONE => NONE)
       | rew2 Ts skel (AbsP (s, t, prf)) = (case rew1 Ts
-              (case skel of AbsP (_, _, skel') => skel' | _ => skel0) prf of
+              (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of
             SOME prf' => SOME (AbsP (s, t, prf'))
           | NONE => NONE)
       | rew2 _ _ _ = NONE;
 
-  in the_default prf (rew1 [] skel0 prf) end;
+  in the_default prf (rew1 [] no_skel prf) end;
 
 fun rewrite_proof thy = rewrite_prf (fn (tyenv, f) =>
   Sign.typ_match thy (f ()) tyenv handle Type.TYPE_MATCH => raise PMatch);
@@ -1249,7 +1253,9 @@
 
 structure ProofData = Theory_Data
 (
-  type T = (stamp * (proof * proof)) list * (stamp * (typ list -> proof -> proof option)) list;
+  type T =
+    (stamp * (proof * proof)) list *
+    (stamp * (typ list -> proof -> (proof * proof) option)) list;
 
   val empty = ([], []);
   val extend = I;
@@ -1277,27 +1283,26 @@
         | _ => false));
   in Promise (i, prop, map TVar (Term.add_tvars prop [])) end;
 
-fun fulfill_proof _ [] body0 = body0
-  | fulfill_proof thy ps body0 =
-      let
-        val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0;
-        val oracles = fold (fn (_, PBody {oracles, ...}) => merge_oracles oracles) ps oracles0;
-        val thms = fold (fn (_, PBody {thms, ...}) => merge_thms thms) ps thms0;
-        val proofs = fold (fn (i, PBody {proof, ...}) => Inttab.update (i, proof)) ps Inttab.empty;
+fun fulfill_norm_proof thy ps body0 =
+  let
+    val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0;
+    val oracles = fold (fn (_, PBody {oracles, ...}) => merge_oracles oracles) ps oracles0;
+    val thms = fold (fn (_, PBody {thms, ...}) => merge_thms thms) ps thms0;
+    val proofs = fold (fn (i, PBody {proof, ...}) => Inttab.update (i, proof)) ps Inttab.empty;
 
-        fun fill (Promise (i, prop, Ts)) =
-            (case Inttab.lookup proofs i of
-              NONE => NONE
-            | SOME prf => SOME (instantiate (Term.add_tvars prop [] ~~ Ts, []) prf))
-          | fill _ = NONE;
-        val (rules, procs) = get_data thy;
-        val proof = rewrite_prf fst (rules, K fill :: procs) proof0;
-      in PBody {oracles = oracles, thms = thms, proof = proof} end;
+    fun fill (Promise (i, prop, Ts)) =
+          (case Inttab.lookup proofs i of
+            NONE => NONE
+          | SOME prf => SOME (instantiate (Term.add_tvars prop [] ~~ Ts, []) prf, normal_skel))
+      | fill _ = NONE;
+    val (rules, procs) = get_data thy;
+    val proof = rewrite_prf fst (rules, K fill :: procs) proof0;
+  in PBody {oracles = oracles, thms = thms, proof = proof} end;
 
 fun fulfill_proof_future _ [] body = Future.value body
   | fulfill_proof_future thy promises body =
       Future.fork_deps (map snd promises) (fn () =>
-        fulfill_proof thy (map (apsnd Future.join) promises) body);
+        fulfill_norm_proof thy (map (apsnd Future.join) promises) body);
 
 
 (***** theorems *****)
--- a/src/Pure/sign.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/sign.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -121,6 +121,8 @@
   val del_trrules: (xstring * string) Syntax.trrule list -> theory -> theory
   val add_trrules_i: ast Syntax.trrule list -> theory -> theory
   val del_trrules_i: ast Syntax.trrule list -> theory -> theory
+  val new_group: theory -> theory
+  val reset_group: theory -> theory
   val add_path: string -> theory -> theory
   val root_path: theory -> theory
   val parent_path: theory -> theory
@@ -610,6 +612,9 @@
 
 (* naming *)
 
+val new_group = map_naming Name_Space.new_group;
+val reset_group = map_naming Name_Space.reset_group;
+
 val add_path = map_naming o Name_Space.add_path;
 val root_path = map_naming Name_Space.root_path;
 val parent_path = map_naming Name_Space.parent_path;
--- a/src/Pure/thm.ML	Wed Nov 18 16:57:58 2009 -0800
+++ b/src/Pure/thm.ML	Thu Nov 19 06:01:02 2009 -0800
@@ -540,7 +540,7 @@
 fun raw_body (Thm (Deriv {body, ...}, _)) = body;
 
 fun fulfill_body (Thm (Deriv {promises, body}, {thy_ref, ...})) =
-  Pt.fulfill_proof (Theory.deref thy_ref)
+  Pt.fulfill_norm_proof (Theory.deref thy_ref)
     (map #1 promises ~~ fulfill_bodies (map #2 promises)) body
 and fulfill_bodies futures = map fulfill_body (Exn.release_all (Future.join_results futures));