merged
authorhaftmann
Thu, 12 Mar 2009 23:01:25 +0100
changeset 30499 1a1a9ca977d6
parent 30490 d09b7f0c2c14 (diff)
parent 30498 55f2933bef6e (current diff)
child 30500 072daf3914c0
merged
src/HOL/Divides.thy
src/HOL/IsaMakefile
--- a/NEWS	Thu Mar 12 18:01:27 2009 +0100
+++ b/NEWS	Thu Mar 12 23:01:25 2009 +0100
@@ -186,7 +186,7 @@
 stated. Any theorems that could solve the lemma directly are listed
 underneath the goal.
 
-* New command find_consts searches for constants based on type and
+* New command 'find_consts' searches for constants based on type and
 name patterns, e.g.
 
     find_consts "_ => bool"
@@ -198,8 +198,8 @@
 
     find_consts strict: "_ => bool" name: "Int" -"int => int"
 
-* Linear arithmetic now ignores all inequalities when fast_arith_neq_limit
-is exceeded, instead of giving up entirely.
+* New command 'local_setup' is similar to 'setup', but operates on a
+local theory context.
 
 
 *** Document preparation ***
@@ -537,6 +537,9 @@
 * ATP selection (E/Vampire/Spass) is now via Proof General's settings
 menu.
 
+* Linear arithmetic now ignores all inequalities when
+fast_arith_neq_limit is exceeded, instead of giving up entirely.
+
 
 *** HOL-Algebra ***
 
--- a/doc-src/IsarRef/Thy/Generic.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc-src/IsarRef/Thy/Generic.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -142,9 +142,8 @@
   compose rules by resolution.  @{attribute THEN} resolves with the
   first premise of @{text a} (an alternative position may be also
   specified); the @{attribute COMP} version skips the automatic
-  lifting process that is normally intended (cf.\ @{ML [source=false]
-  "op RS"} and @{ML [source=false] "op COMP"} in
-  \cite{isabelle-implementation}).
+  lifting process that is normally intended (cf.\ @{ML "op RS"} and
+  @{ML "op COMP"} in \cite{isabelle-implementation}).
   
   \item @{attribute unfolded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} and @{attribute
   folded}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} expand and fold back again the given
--- a/doc-src/IsarRef/Thy/Proof.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc-src/IsarRef/Thy/Proof.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -769,11 +769,11 @@
   \item @{attribute rule}~@{text del} undeclares introduction,
   elimination, or destruct rules.
   
-  \item @{attribute OF}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} applies some theorem to all
-  of the given rules @{text "a\<^sub>1, \<dots>, a\<^sub>n"} (in parallel).  This
-  corresponds to the @{ML [source=false] "op MRS"} operation in ML,
-  but note the reversed order.  Positions may be effectively skipped
-  by including ``@{text _}'' (underscore) as argument.
+  \item @{attribute OF}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} applies some
+  theorem to all of the given rules @{text "a\<^sub>1, \<dots>, a\<^sub>n"}
+  (in parallel).  This corresponds to the @{ML "op MRS"} operation in
+  ML, but note the reversed order.  Positions may be effectively
+  skipped by including ``@{text _}'' (underscore) as argument.
   
   \item @{attribute of}~@{text "t\<^sub>1 \<dots> t\<^sub>n"} performs positional
   instantiation of term variables.  The terms @{text "t\<^sub>1, \<dots>, t\<^sub>n"} are
@@ -886,11 +886,11 @@
 
   \item @{command "method_setup"}~@{text "name = text description"}
   defines a proof method in the current theory.  The given @{text
-  "text"} has to be an ML expression of type @{ML_type [source=false]
-  "Args.src -> Proof.context -> Proof.method"}.  Parsing concrete
-  method syntax from @{ML_type Args.src} input can be quite tedious in
-  general.  The following simple examples are for methods without any
-  explicit arguments, or a list of theorems, respectively.
+  "text"} has to be an ML expression of type @{ML_type "Args.src ->
+  Proof.context -> Proof.method"}.  Parsing concrete method syntax
+  from @{ML_type Args.src} input can be quite tedious in general.  The
+  following simple examples are for methods without any explicit
+  arguments, or a list of theorems, respectively.
 
 %FIXME proper antiquotations
 {\footnotesize
--- a/doc-src/IsarRef/Thy/Spec.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc-src/IsarRef/Thy/Spec.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -799,6 +799,7 @@
     @{command_def "ML_val"} & : & @{text "any \<rightarrow>"} \\
     @{command_def "ML_command"} & : & @{text "any \<rightarrow>"} \\
     @{command_def "setup"} & : & @{text "theory \<rightarrow> theory"} \\
+    @{command_def "local_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   \end{matharray}
 
   \begin{mldecls}
@@ -809,7 +810,7 @@
   \begin{rail}
     'use' name
     ;
-    ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup') text
+    ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup' | 'local\_setup') text
     ;
   \end{rail}
 
@@ -817,7 +818,7 @@
 
   \item @{command "use"}~@{text "file"} reads and executes ML
   commands from @{text "file"}.  The current theory context is passed
-  down to the ML toplevel and may be modified, using @{ML [source=false]
+  down to the ML toplevel and may be modified, using @{ML
   "Context.>>"} or derived ML commands.  The file name is checked with
   the @{keyword_ref "uses"} dependency declaration given in the theory
   header (see also \secref{sec:begin-thy}).
@@ -845,9 +846,15 @@
   
   \item @{command "setup"}~@{text "text"} changes the current theory
   context by applying @{text "text"}, which refers to an ML expression
-  of type @{ML_type [source=false] "theory -> theory"}.  This enables
-  to initialize any object-logic specific tools and packages written
-  in ML, for example.
+  of type @{ML_type "theory -> theory"}.  This enables to initialize
+  any object-logic specific tools and packages written in ML, for
+  example.
+
+  \item @{command "local_setup"} is similar to @{command "setup"} for
+  a local theory context, and an ML expression of type @{ML_type
+  "local_theory -> local_theory"}.  This allows to
+  invoke local theory specification packages without going through
+  concrete outer syntax, for example.
 
   \item @{ML bind_thms}~@{text "(name, thms)"} stores a list of
   theorems produced in ML both in the theory context and the ML
--- a/doc-src/IsarRef/Thy/document/Generic.tex	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc-src/IsarRef/Thy/document/Generic.tex	Thu Mar 12 23:01:25 2009 +0100
@@ -160,8 +160,8 @@
   compose rules by resolution.  \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}} resolves with the
   first premise of \isa{a} (an alternative position may be also
   specified); the \hyperlink{attribute.COMP}{\mbox{\isa{COMP}}} version skips the automatic
-  lifting process that is normally intended (cf.\ \verb|op RS| and \verb|op COMP| in
-  \cite{isabelle-implementation}).
+  lifting process that is normally intended (cf.\ \verb|op RS| and
+  \verb|op COMP| in \cite{isabelle-implementation}).
   
   \item \hyperlink{attribute.unfolded}{\mbox{\isa{unfolded}}}~\isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}} and \hyperlink{attribute.folded}{\mbox{\isa{folded}}}~\isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}} expand and fold back again the given
   definitions throughout a rule.
--- a/doc-src/IsarRef/Thy/document/Proof.tex	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc-src/IsarRef/Thy/document/Proof.tex	Thu Mar 12 23:01:25 2009 +0100
@@ -774,11 +774,11 @@
   \item \hyperlink{attribute.rule}{\mbox{\isa{rule}}}~\isa{del} undeclares introduction,
   elimination, or destruct rules.
   
-  \item \hyperlink{attribute.OF}{\mbox{\isa{OF}}}~\isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}} applies some theorem to all
-  of the given rules \isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n{\isachardoublequote}} (in parallel).  This
-  corresponds to the \verb|op MRS| operation in ML,
-  but note the reversed order.  Positions may be effectively skipped
-  by including ``\isa{{\isacharunderscore}}'' (underscore) as argument.
+  \item \hyperlink{attribute.OF}{\mbox{\isa{OF}}}~\isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}\ {\isasymdots}\ a\isactrlsub n{\isachardoublequote}} applies some
+  theorem to all of the given rules \isa{{\isachardoublequote}a\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ a\isactrlsub n{\isachardoublequote}}
+  (in parallel).  This corresponds to the \verb|op MRS| operation in
+  ML, but note the reversed order.  Positions may be effectively
+  skipped by including ``\isa{{\isacharunderscore}}'' (underscore) as argument.
   
   \item \hyperlink{attribute.of}{\mbox{\isa{of}}}~\isa{{\isachardoublequote}t\isactrlsub {\isadigit{1}}\ {\isasymdots}\ t\isactrlsub n{\isachardoublequote}} performs positional
   instantiation of term variables.  The terms \isa{{\isachardoublequote}t\isactrlsub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ t\isactrlsub n{\isachardoublequote}} are
@@ -890,10 +890,11 @@
   \begin{description}
 
   \item \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}}~\isa{{\isachardoublequote}name\ {\isacharequal}\ text\ description{\isachardoublequote}}
-  defines a proof method in the current theory.  The given \isa{{\isachardoublequote}text{\isachardoublequote}} has to be an ML expression of type \verb|Args.src -> Proof.context -> Proof.method|.  Parsing concrete
-  method syntax from \verb|Args.src| input can be quite tedious in
-  general.  The following simple examples are for methods without any
-  explicit arguments, or a list of theorems, respectively.
+  defines a proof method in the current theory.  The given \isa{{\isachardoublequote}text{\isachardoublequote}} has to be an ML expression of type \verb|Args.src ->|\isasep\isanewline%
+\verb|  Proof.context -> Proof.method|.  Parsing concrete method syntax
+  from \verb|Args.src| input can be quite tedious in general.  The
+  following simple examples are for methods without any explicit
+  arguments, or a list of theorems, respectively.
 
 %FIXME proper antiquotations
 {\footnotesize
--- a/doc-src/IsarRef/Thy/document/Spec.tex	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc-src/IsarRef/Thy/document/Spec.tex	Thu Mar 12 23:01:25 2009 +0100
@@ -807,6 +807,7 @@
     \indexdef{}{command}{ML\_val}\hypertarget{command.ML-val}{\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}}} & : & \isa{{\isachardoublequote}any\ {\isasymrightarrow}{\isachardoublequote}} \\
     \indexdef{}{command}{ML\_command}\hypertarget{command.ML-command}{\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}}} & : & \isa{{\isachardoublequote}any\ {\isasymrightarrow}{\isachardoublequote}} \\
     \indexdef{}{command}{setup}\hypertarget{command.setup}{\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+    \indexdef{}{command}{local\_setup}\hypertarget{command.local-setup}{\hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isacharunderscore}setup}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
   \end{matharray}
 
   \begin{mldecls}
@@ -817,7 +818,7 @@
   \begin{rail}
     'use' name
     ;
-    ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup') text
+    ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup' | 'local\_setup') text
     ;
   \end{rail}
 
@@ -852,9 +853,14 @@
   
   \item \hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}~\isa{{\isachardoublequote}text{\isachardoublequote}} changes the current theory
   context by applying \isa{{\isachardoublequote}text{\isachardoublequote}}, which refers to an ML expression
-  of type \verb|theory -> theory|.  This enables
-  to initialize any object-logic specific tools and packages written
-  in ML, for example.
+  of type \verb|theory -> theory|.  This enables to initialize
+  any object-logic specific tools and packages written in ML, for
+  example.
+
+  \item \hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isacharunderscore}setup}}}} is similar to \hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}} for
+  a local theory context, and an ML expression of type \verb|local_theory -> local_theory|.  This allows to
+  invoke local theory specification packages without going through
+  concrete outer syntax, for example.
 
   \item \verb|bind_thms|~\isa{{\isachardoublequote}{\isacharparenleft}name{\isacharcomma}\ thms{\isacharparenright}{\isachardoublequote}} stores a list of
   theorems produced in ML both in the theory context and the ML
--- a/doc/Contents	Thu Mar 12 18:01:27 2009 +0100
+++ b/doc/Contents	Thu Mar 12 23:01:25 2009 +0100
@@ -6,7 +6,7 @@
   classes         Tutorial on Type Classes
   functions       Tutorial on Function Definitions
   codegen         Tutorial on Code Generation
-  sugar           LaTeX sugar for proof documents
+  sugar           LaTeX Sugar for Isabelle documents
 
 Reference Manuals
   isar-ref        The Isabelle/Isar Reference Manual
--- a/etc/isar-keywords-ZF.el	Thu Mar 12 18:01:27 2009 +0100
+++ b/etc/isar-keywords-ZF.el	Thu Mar 12 23:01:25 2009 +0100
@@ -100,6 +100,7 @@
     "let"
     "linear_undo"
     "local"
+    "local_setup"
     "locale"
     "method_setup"
     "moreover"
@@ -379,6 +380,7 @@
     "judgment"
     "lemmas"
     "local"
+    "local_setup"
     "locale"
     "method_setup"
     "no_notation"
--- a/etc/isar-keywords.el	Thu Mar 12 18:01:27 2009 +0100
+++ b/etc/isar-keywords.el	Thu Mar 12 23:01:25 2009 +0100
@@ -127,6 +127,7 @@
     "let"
     "linear_undo"
     "local"
+    "local_setup"
     "locale"
     "method_setup"
     "moreover"
@@ -468,6 +469,7 @@
     "judgment"
     "lemmas"
     "local"
+    "local_setup"
     "locale"
     "method_setup"
     "no_notation"
--- a/lib/jedit/isabelle.xml	Thu Mar 12 18:01:27 2009 +0100
+++ b/lib/jedit/isabelle.xml	Thu Mar 12 23:01:25 2009 +0100
@@ -185,6 +185,7 @@
       <OPERATOR>let</OPERATOR>
       <INVALID>linear_undo</INVALID>
       <OPERATOR>local</OPERATOR>
+      <OPERATOR>local_setup</OPERATOR>
       <OPERATOR>locale</OPERATOR>
       <OPERATOR>method_setup</OPERATOR>
       <KEYWORD4>module_name</KEYWORD4>
--- a/src/HOL/Divides.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Divides.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -295,6 +295,27 @@
 
 end
 
+lemma div_mult_div_if_dvd: "(y::'a::{semiring_div,no_zero_divisors}) dvd x \<Longrightarrow> 
+  z dvd w \<Longrightarrow> (x div y) * (w div z) = (x * w) div (y * z)"
+unfolding dvd_def
+  apply clarify
+  apply (case_tac "y = 0")
+  apply simp
+  apply (case_tac "z = 0")
+  apply simp
+  apply (simp add: algebra_simps)
+  apply (subst mult_assoc [symmetric])
+  apply (simp add: no_zero_divisors)
+done
+
+
+lemma div_power: "(y::'a::{semiring_div,no_zero_divisors,recpower}) dvd x \<Longrightarrow>
+    (x div y)^n = x^n div y^n"
+apply (induct n)
+ apply simp
+apply(simp add: div_mult_div_if_dvd dvd_power_same)
+done
+
 class ring_div = semiring_div + comm_ring_1
 begin
 
--- a/src/HOL/FunDef.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/FunDef.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -314,13 +314,11 @@
 use "Tools/function_package/scnp_reconstruct.ML"
 
 setup {* ScnpReconstruct.setup *}
-(*
-setup {*
+
+ML_val -- "setup inactive"
+{*
   Context.theory_map (FundefCommon.set_termination_prover (ScnpReconstruct.decomp_scnp 
   [ScnpSolve.MAX, ScnpSolve.MIN, ScnpSolve.MS])) 
 *}
-*)
-
-
 
 end
--- a/src/HOL/Import/shuffler.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Import/shuffler.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -662,7 +662,7 @@
 fun search_meth ctxt =
     let
         val thy = ProofContext.theory_of ctxt
-        val prems = Assumption.prems_of ctxt
+        val prems = Assumption.all_prems_of ctxt
     in
         Method.SIMPLE_METHOD' (gen_shuffle_tac thy true (map (pair "premise") prems))
     end
--- a/src/HOL/IsaMakefile	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/IsaMakefile	Thu Mar 12 23:01:25 2009 +0100
@@ -345,7 +345,8 @@
   Library/Product_plus.thy \
   Library/Product_Vector.thy \
   Library/Enum.thy Library/Float.thy $(SRC)/Tools/float.ML $(SRC)/HOL/Tools/float_arith.ML \
-  Library/reify_data.ML Library/reflection.ML
+  Library/reify_data.ML Library/reflection.ML \
+  Library/LaTeXsugar.thy Library/OptionalSugar.thy
 	@cd Library; $(ISABELLE_TOOL) usedir $(OUT)/HOL Library
 
 
--- a/src/HOL/Library/Determinants.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Determinants.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -40,7 +40,7 @@
 lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp
 
 lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)"
-  "setprod f {m .. Suc n} = (if m \<le> Suc n then f (Suc n) * setprod f {m..n} 
+  "setprod f {m .. Suc n} = (if m \<le> Suc n then f (Suc n) * setprod f {m..n}
                              else setprod f {m..n})"
   by (auto simp add: atLeastAtMostSuc_conv)
 
@@ -98,20 +98,20 @@
 (* A few general lemmas we need below.                                       *)
 (* ------------------------------------------------------------------------- *)
 
-lemma Cart_lambda_beta_perm: assumes p: "p permutes {1..dimindex(UNIV::'n set)}" 
-  and i: "i \<in> {1..dimindex(UNIV::'n set)}" 
+lemma Cart_lambda_beta_perm: assumes p: "p permutes {1..dimindex(UNIV::'n set)}"
+  and i: "i \<in> {1..dimindex(UNIV::'n set)}"
   shows "Cart_nth (Cart_lambda g ::'a^'n) (p i) = g(p i)"
   using permutes_in_image[OF p] i
   by (simp add:  Cart_lambda_beta permutes_in_image[OF p])
 
 lemma setprod_permute:
-  assumes p: "p permutes S" 
+  assumes p: "p permutes S"
   shows "setprod f S = setprod (f o p) S"
 proof-
   {assume "\<not> finite S" hence ?thesis by simp}
   moreover
   {assume fS: "finite S"
-    then have ?thesis 
+    then have ?thesis
       apply (simp add: setprod_def)
       apply (rule ab_semigroup_mult.fold_image_permute)
       apply (auto simp add: p)
@@ -134,9 +134,9 @@
   have fU: "finite ?U" by blast
   {fix p assume p: "p \<in> {p. p permutes ?U}"
     from p have pU: "p permutes ?U" by blast
-    have sth: "sign (inv p) = sign p" 
+    have sth: "sign (inv p) = sign p"
       by (metis sign_inverse fU p mem_def Collect_def permutation_permutes)
-    from permutes_inj[OF pU] 
+    from permutes_inj[OF pU]
     have pi: "inj_on p ?U" by (blast intro: subset_inj_on)
     from permutes_image[OF pU]
     have "setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U = setprod (\<lambda>i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp
@@ -148,7 +148,7 @@
 	from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
 	have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
 	  unfolding transp_def by (simp add: Cart_lambda_beta expand_fun_eq)}
-      then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)  
+      then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
     qed
     finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
       by simp}
@@ -156,7 +156,7 @@
   apply (rule setsum_cong2) by blast
 qed
 
-lemma det_lowerdiagonal: 
+lemma det_lowerdiagonal:
   fixes A :: "'a::comm_ring_1^'n^'n"
   assumes ld: "\<And>i j. i \<in> {1 .. dimindex (UNIV:: 'n set)} \<Longrightarrow> j \<in> {1 .. dimindex(UNIV:: 'n set)} \<Longrightarrow> i < j \<Longrightarrow> A$i$j = 0"
   shows "det A = setprod (\<lambda>i. A$i$i) {1..dimindex(UNIV:: 'n set)}"
@@ -179,7 +179,7 @@
     unfolding det_def by (simp add: sign_id)
 qed
 
-lemma det_upperdiagonal: 
+lemma det_upperdiagonal:
   fixes A :: "'a::comm_ring_1^'n^'n"
   assumes ld: "\<And>i j. i \<in> {1 .. dimindex (UNIV:: 'n set)} \<Longrightarrow> j \<in> {1 .. dimindex(UNIV:: 'n set)} \<Longrightarrow> i > j \<Longrightarrow> A$i$j = 0"
   shows "det A = setprod (\<lambda>i. A$i$i) {1..dimindex(UNIV:: 'n set)}"
@@ -216,7 +216,7 @@
   then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_lowerdiagonal
     by blast
   also have "\<dots> = 1" unfolding th setprod_1 ..
-  finally show ?thesis . 
+  finally show ?thesis .
 qed
 
 lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0"
@@ -235,7 +235,7 @@
   then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_lowerdiagonal
     by blast
   also have "\<dots> = 0" unfolding th  ..
-  finally show ?thesis . 
+  finally show ?thesis .
 qed
 
 lemma det_permute_rows:
@@ -243,7 +243,7 @@
   assumes p: "p permutes {1 .. dimindex (UNIV :: 'n set)}"
   shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
   apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric] del: One_nat_def)
-  apply (subst sum_permutations_compose_right[OF p])  
+  apply (subst sum_permutations_compose_right[OF p])
 proof(rule setsum_cong2)
   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
   let ?PU = "{p. p permutes ?U}"
@@ -259,14 +259,14 @@
       have "?Ap$i$ (q o p) i = A $ p i $ (q o p) i " by simp}
     hence "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$p i$(q o p) i) ?U"
       by (auto intro: setprod_cong)
-    also have "\<dots> = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U" 
+    also have "\<dots> = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U"
       by (simp only: setprod_permute[OF ip, symmetric])
     also have "\<dots> = setprod (\<lambda>i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U"
       by (simp only: o_def)
     also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p])
-    finally   have thp: "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U" 
+    finally   have thp: "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U"
       by blast
-  show "of_int (sign (q o p)) * setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U" 
+  show "of_int (sign (q o p)) * setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
     by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
 qed
 
@@ -282,18 +282,18 @@
   moreover
   have "?Ap = transp (\<chi> i. transp A $ p i)"
     by (simp add: transp_def Cart_eq Cart_lambda_beta Cart_lambda_beta_perm[OF p])
-  ultimately show ?thesis by simp 
+  ultimately show ?thesis by simp
 qed
 
 lemma det_identical_rows:
   fixes A :: "'a::ordered_idom^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
+  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
   and j: "j\<in>{1 .. dimindex (UNIV :: 'n set)}"
   and ij: "i \<noteq> j"
   and r: "row i A = row j A"
   shows	"det A = 0"
 proof-
-  have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0" 
+  have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0"
     by simp
   have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min)
   let ?p = "Fun.swap i j id"
@@ -302,12 +302,12 @@
   hence "det A = det ?A" by simp
   moreover have "det A = - det ?A"
     by (simp add: det_permute_rows[OF permutes_swap_id[OF i j]] sign_swap_id ij th1)
-  ultimately show "det A = 0" by (metis tha) 
+  ultimately show "det A = 0" by (metis tha)
 qed
 
 lemma det_identical_columns:
   fixes A :: "'a::ordered_idom^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
+  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
   and j: "j\<in>{1 .. dimindex (UNIV :: 'n set)}"
   and ij: "i \<noteq> j"
   and r: "column i A = column j A"
@@ -316,9 +316,9 @@
 apply (rule det_identical_rows[OF i j ij])
 by (metis row_transp i j r)
 
-lemma det_zero_row: 
+lemma det_zero_row:
   fixes A :: "'a::{idom, ring_char_0}^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
+  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
   and r: "row i A = 0"
   shows "det A = 0"
 using i r
@@ -332,16 +332,16 @@
 apply (subgoal_tac "(0\<Colon>'a ^ 'n) $ a i = 0")
 apply simp
 apply (rule zero_index)
-apply (drule permutes_in_image[of _ _ i]) 
+apply (drule permutes_in_image[of _ _ i])
 apply simp
-apply (drule permutes_in_image[of _ _ i]) 
+apply (drule permutes_in_image[of _ _ i])
 apply simp
 apply simp
 done
 
 lemma det_zero_column:
   fixes A :: "'a::{idom,ring_char_0}^'n^'n"
-  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
+  assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}"
   and r: "column i A = 0"
   shows "det A = 0"
   apply (subst det_transp[symmetric])
@@ -361,7 +361,7 @@
 proof(rule setprod_cong[OF refl])
   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
   fix i assume i: "i \<in> ?U"
-  from Cart_lambda_beta'[OF i, of g] have 
+  from Cart_lambda_beta'[OF i, of g] have
     "((\<chi> i. g i) :: 'a^'n^'n) $ i = g i" .
   hence "((\<chi> i. g i) :: 'a^'n^'n) $ i $ f i = g i $ f i" by simp
   then
@@ -369,7 +369,7 @@
 qed
 
 lemma det_row_add:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" 
+  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
   shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
              det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
              det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)"
@@ -387,7 +387,7 @@
   note pin[simp] = permutes_in_image[OF pU]
   have kU: "?U = insert k ?Uk" using k by blast
   {fix j assume j: "j \<in> ?Uk"
-    from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" 
+    from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j"
       by simp_all}
   then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
     and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk"
@@ -411,7 +411,7 @@
 qed
 
 lemma det_row_mul:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" 
+  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
   shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
              c* det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
 
@@ -451,7 +451,7 @@
 qed
 
 lemma det_row_0:
-  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" 
+  assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
   shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0"
 using det_row_mul[OF k, of 0 "\<lambda>i. 1" b]
 apply (simp)
@@ -483,8 +483,8 @@
   let ?S = "{row j A |j. j\<in> ?U \<and> j\<noteq> i}"
   let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)"
   let ?P = "\<lambda>x. ?d (row i A + x) = det A"
-  {fix k 
-    
+  {fix k
+
     have "(if k = i then row i A + 0 else row k A) = row k A" by simp}
   then have P0: "?P 0"
     apply -
@@ -499,10 +499,10 @@
       apply (rule det_identical_rows[OF i j(2,3)])
       using i j by (vector row_def)
     have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 ..
-    then have "?P (c*s z + y)" unfolding thz Py det_row_mul[OF i] det_row_add[OF i] 
+    then have "?P (c*s z + y)" unfolding thz Py det_row_mul[OF i] det_row_add[OF i]
       by simp }
 
-  ultimately show ?thesis 
+  ultimately show ?thesis
     apply -
     apply (rule span_induct_alt[of ?P ?S, OF P0])
     apply blast
@@ -524,7 +524,7 @@
   from d obtain i where i: "i \<in> ?U" "row i A \<in> span (rows A - {row i A})"
     unfolding dependent_def rows_def by blast
   {fix j k assume j: "j \<in>?U" and k: "k \<in> ?U" and jk: "j \<noteq> k"
-    and c: "row j A = row k A" 
+    and c: "row j A = row k A"
     from det_identical_rows[OF j k jk c] have ?thesis .}
   moreover
   {assume H: "\<And> i j. i\<in> ?U \<Longrightarrow> j \<in> ?U \<Longrightarrow> i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
@@ -537,7 +537,7 @@
     from det_row_span[OF i(1) th0]
     have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)"
       unfolding right_minus vector_smult_lzero ..
-    with det_row_mul[OF i(1), of "0::'a" "\<lambda>i. 1"] 
+    with det_row_mul[OF i(1), of "0::'a" "\<lambda>i. 1"]
     have "det A = 0" by simp}
   ultimately show ?thesis by blast
 qed
@@ -552,7 +552,7 @@
 lemma Cart_lambda_cong: "(\<And>x. x \<in> {1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
   apply (rule iffD1[OF Cart_lambda_unique]) by vector
 
-lemma det_linear_row_setsum: 
+lemma det_linear_row_setsum:
   assumes fS: "finite S" and k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
   shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
   using k
@@ -567,7 +567,7 @@
   assumes fS: "finite S"
   shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
 proof(induct k)
-  case 0 
+  case 0
   have th: "{f. \<forall>i. f i = i} = {id}" by (auto intro: ext)
   show ?case by (auto simp add: th)
 next
@@ -581,7 +581,7 @@
     apply (auto intro: ext)
     done
   with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f]
-  show ?case by metis 
+  show ?case by metis
 qed
 
 
@@ -608,9 +608,9 @@
   from Suc.prems have k': "k \<le> dimindex (UNIV :: 'n set)" by arith
   have thif: "\<And>a b c d. (if b \<or> a then c else d) = (if a then c else if b then c else d)" by simp
   have thif2: "\<And>a b c d e. (if a then b else if c then d else e) =
-     (if c then (if a then b else d) else (if a then b else e))" by simp 
-  have "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) = 
-        det (\<chi> i. if i = Suc k then setsum (a i) S 
+     (if c then (if a then b else d) else (if a then b else e))" by simp
+  have "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) =
+        det (\<chi> i. if i = Suc k then setsum (a i) S
                  else if i \<le> k then setsum (a i) S else c i)"
     unfolding le_Suc_eq thif  ..
   also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<le> k then setsum (a i) S
@@ -618,14 +618,14 @@
     unfolding det_linear_row_setsum[OF fS Sk]
     apply (subst thif2)
     by (simp cong del: if_weak_cong cong add: if_cong)
-  finally have tha: 
-    "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) = 
+  finally have tha:
+    "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) =
      (\<Sum>(j, f)\<in>S \<times> ?F k. det (\<chi> i. if i \<le> k then a i (f i)
                                 else if i = Suc k then a i j
-                                else c i))" 
+                                else c i))"
     unfolding  Suc.hyps[OF k'] unfolding setsum_cartesian_product by blast
   show ?case unfolding tha
-    apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"], 
+    apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"],
       blast intro: finite_cartesian_product fS finite_bounded_functions[OF fS],
       blast intro: finite_cartesian_product fS finite_bounded_functions[OF fS], auto intro: ext)
     apply (rule cong[OF refl[of det]])
@@ -637,7 +637,7 @@
   shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. (\<forall>i \<in> {1 .. dimindex (UNIV :: 'n set)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. dimindex (UNIV :: 'n set)} \<longrightarrow> f i = i)}"
 proof-
   have th0: "\<And>x y. ((\<chi> i. if i <= dimindex(UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
-  
+
   from det_linear_rows_setsum_lemma[OF fS, of "dimindex (UNIV :: 'n set)" a, unfolded th0, OF order_refl] show ?thesis by blast
 qed
 
@@ -674,25 +674,25 @@
   have fU: "finite ?U" by simp
   have fF: "finite ?F"  using finite_bounded_functions[OF fU] .
   {fix p assume p: "p permutes ?U"
-    
+
     have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p]
       using p[unfolded permutes_def] by simp}
-  then have PUF: "?PU \<subseteq> ?F"  by blast 
+  then have PUF: "?PU \<subseteq> ?F"  by blast
   {fix f assume fPU: "f \<in> ?F - ?PU"
     have fUU: "f ` ?U \<subseteq> ?U" using fPU by auto
     from fPU have f: "\<forall>i \<in> ?U. f i \<in> ?U"
-      "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" unfolding permutes_def 
+      "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" unfolding permutes_def
       by auto
-    
+
     let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n"
     let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n"
     {assume fni: "\<not> inj_on f ?U"
       then obtain i j where ij: "i \<in> ?U" "j \<in> ?U" "f i = f j" "i \<noteq> j"
 	unfolding inj_on_def by blast
-      from ij 
+      from ij
       have rth: "row i ?B = row j ?B" by (vector row_def)
-      from det_identical_rows[OF ij(1,2,4) rth] 
-      have "det (\<chi> i. A$i$f i *s B$f i) = 0" 
+      from det_identical_rows[OF ij(1,2,4) rth]
+      have "det (\<chi> i. A$i$f i *s B$f i) = 0"
 	unfolding det_rows_mul by simp}
     moreover
     {assume fi: "inj_on f ?U"
@@ -701,7 +701,7 @@
 	apply (case_tac "i \<in> ?U")
 	apply (case_tac "j \<in> ?U") by metis+
       note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]]
-      
+
       {fix y
 	from fs f have "\<exists>x. f x = y" by (cases "y \<in> ?U") blast+
 	then obtain x where x: "f x = y" by blast
@@ -724,17 +724,17 @@
       fix q assume qU: "q \<in> ?PU"
       hence q: "q permutes ?U" by blast
       from p q have pp: "permutation p" and pq: "permutation q"
-	unfolding permutation_permutes by auto 
-      have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" 
-	"\<And>a. of_int (sign p) * (of_int (sign p) * a) = a" 
-	unfolding mult_assoc[symmetric]	unfolding of_int_mult[symmetric] 
+	unfolding permutation_permutes by auto
+      have th00: "of_int (sign p) * of_int (sign p) = (1::'a)"
+	"\<And>a. of_int (sign p) * (of_int (sign p) * a) = a"
+	unfolding mult_assoc[symmetric]	unfolding of_int_mult[symmetric]
 	by (simp_all add: sign_idempotent)
       have ths: "?s q = ?s p * ?s (q o inv p)"
 	using pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
 	by (simp add:  th00 mult_ac sign_idempotent sign_compose)
       have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) o p) ?U"
 	by (rule setprod_permute[OF p])
-      have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U" 
+      have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U"
 	unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p]
 	apply (rule setprod_cong[OF refl])
 	using permutes_in_image[OF q] by vector
@@ -743,16 +743,16 @@
 	by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
     qed
   }
-  then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B" 
+  then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
     unfolding det_def setsum_product
-    by (rule setsum_cong2) 
+    by (rule setsum_cong2)
   have "det (A**B) = setsum (\<lambda>f.  det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
-    unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] .. 
+    unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] ..
   also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
-    using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric] 
+    using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric]
     unfolding det_rows_mul by auto
   finally show ?thesis unfolding th2 .
-qed  
+qed
 
 (* ------------------------------------------------------------------------- *)
 (* Relation to invertibility.                                                *)
@@ -768,7 +768,7 @@
   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)"
   by (metis invertible_def matrix_left_right_inverse)
 
-lemma invertible_det_nz: 
+lemma invertible_det_nz:
   fixes A::"real ^'n^'n"
   shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
 proof-
@@ -782,7 +782,7 @@
   {assume H: "\<not> invertible A"
     let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
     have fU: "finite ?U" by simp
-    from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0" 
+    from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0"
       and iU: "i \<in> ?U" and ci: "c i \<noteq> 0"
       unfolding invertible_righ_inverse
       unfolding matrix_right_invertible_independent_rows by blast
@@ -791,14 +791,14 @@
       apply (simp only: ab_left_minus add_assoc[symmetric])
       apply simp
       done
-    from c ci 
+    from c ci
     have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s c j *s row j A) (?U - {i})"
-      unfolding setsum_diff1'[OF fU iU] setsum_cmul 
+      unfolding setsum_diff1'[OF fU iU] setsum_cmul
       apply (simp add: field_simps)
       apply (rule vector_mul_lcancel_imp[OF ci])
       apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps)
       unfolding stupid ..
-    have thr: "- row i A \<in> span {row j A| j. j\<in> ?U \<and> j \<noteq> i}" 
+    have thr: "- row i A \<in> span {row j A| j. j\<in> ?U \<and> j \<noteq> i}"
       unfolding thr0
       apply (rule span_setsum)
       apply simp
@@ -808,8 +808,8 @@
       apply auto
       done
     let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n"
-    have thrb: "row i ?B = 0" using iU by (vector row_def) 
-    have "det A = 0" 
+    have thrb: "row i ?B = 0" using iU by (vector row_def)
+    have "det A = 0"
       unfolding det_row_span[OF iU thr, symmetric] right_minus
       unfolding  det_zero_row[OF iU thrb]  ..}
   ultimately show ?thesis by blast
@@ -823,8 +823,8 @@
   fixes A:: "'a::ordered_idom^'n^'n" and x :: "'a ^'n"
   assumes k: "k \<in> {1 .. dimindex(UNIV ::'n set)}"
   shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) {1 .. dimindex(UNIV::'n set)}
-                           else row i A)::'a^'n^'n) = x$k * det A" 
-  (is "?lhs = ?rhs") 
+                           else row i A)::'a^'n^'n) = x$k * det A"
+  (is "?lhs = ?rhs")
 proof-
   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
   let ?Uk = "?U - {k}"
@@ -835,7 +835,7 @@
     by (vector ring_simps)
   have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" by (auto intro: ext)
   have "(\<chi> i. row i A) = A" by (vector row_def)
-  then have thd1: "det (\<chi> i. row i A) = det A"  by simp 
+  then have thd1: "det (\<chi> i. row i A) = det A"  by simp
   have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A"
     apply (rule det_row_span[OF k])
     apply (rule span_setsum[OF fUk])
@@ -846,7 +846,7 @@
     done
   show "?lhs = x$k * det A"
     apply (subst U)
-    unfolding setsum_insert[OF fUk kUk] 
+    unfolding setsum_insert[OF fUk kUk]
     apply (subst th00)
     unfolding add_assoc
     apply (subst det_row_add[OF k])
@@ -863,8 +863,8 @@
 proof-
   have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
     by (auto simp add: row_transp intro: setsum_cong2)
-  show ?thesis 
-  unfolding matrix_mult_vsum 
+  show ?thesis
+  unfolding matrix_mult_vsum
   unfolding cramer_lemma_transp[OF k, of x "transp A", unfolded det_transp, symmetric]
   unfolding stupid[of "\<lambda>i. x$i"]
   apply (subst det_transp[symmetric])
@@ -873,10 +873,10 @@
 
 lemma cramer:
   fixes A ::"real^'n^'n"
-  assumes d0: "det A \<noteq> 0" 
+  assumes d0: "det A \<noteq> 0"
   shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
 proof-
-  from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"  
+  from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
     unfolding invertible_det_nz[symmetric] invertible_def by blast
   have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid)
   hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc)
@@ -896,10 +896,10 @@
 
 lemma orthogonal_transformation: "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^'n). norm (f v) = norm v)"
   unfolding orthogonal_transformation_def
-  apply auto 
+  apply auto
   apply (erule_tac x=v in allE)+
   apply (simp add: real_vector_norm_def)
-  by (simp add: dot_norm  linear_add[symmetric]) 
+  by (simp add: dot_norm  linear_add[symmetric])
 
 definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
 
@@ -909,12 +909,12 @@
 lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1)"
   by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
 
-lemma orthogonal_matrix_mul: 
+lemma orthogonal_matrix_mul:
   fixes A :: "real ^'n^'n"
   assumes oA : "orthogonal_matrix A"
-  and oB: "orthogonal_matrix B" 
+  and oB: "orthogonal_matrix B"
   shows "orthogonal_matrix(A ** B)"
-  using oA oB 
+  using oA oB
   unfolding orthogonal_matrix matrix_transp_mul
   apply (subst matrix_mul_assoc)
   apply (subst matrix_mul_assoc[symmetric])
@@ -939,7 +939,7 @@
 	"\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
 	by simp_all
       from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] i j
-      have "?A$i$j = ?m1 $ i $ j" 
+      have "?A$i$j = ?m1 $ i $ j"
 	by (simp add: Cart_lambda_beta' dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def del: One_nat_def)}
     hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector
     with lf have ?rhs by blast}
@@ -953,17 +953,17 @@
   ultimately show ?thesis by blast
 qed
 
-lemma det_orthogonal_matrix: 
+lemma det_orthogonal_matrix:
   fixes Q:: "'a::ordered_idom^'n^'n"
   assumes oQ: "orthogonal_matrix Q"
   shows "det Q = 1 \<or> det Q = - 1"
 proof-
-  
-  have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x") 
-  proof- 
+
+  have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
+  proof-
     fix x:: 'a
     have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps)
-    have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0" 
+    have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0"
       apply (subst eq_iff_diff_eq_0) by simp
     have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp
     also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" unfolding th0 th1 by simp
@@ -972,27 +972,27 @@
   from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def)
   hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp
   hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp)
-  then show ?thesis unfolding th . 
+  then show ?thesis unfolding th .
 qed
 
 (* ------------------------------------------------------------------------- *)
 (* Linearity of scaling, and hence isometry, that preserves origin.          *)
 (* ------------------------------------------------------------------------- *)
-lemma scaling_linear: 
+lemma scaling_linear:
   fixes f :: "real ^'n \<Rightarrow> real ^'n"
   assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
   shows "linear f"
 proof-
-  {fix v w 
+  {fix v w
     {fix x note fd[rule_format, of x 0, unfolded dist_def f0 diff_0_right] }
     note th0 = this
-    have "f v \<bullet> f w = c^2 * (v \<bullet> w)" 
+    have "f v \<bullet> f w = c^2 * (v \<bullet> w)"
       unfolding dot_norm_neg dist_def[symmetric]
       unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
   note fc = this
   show ?thesis unfolding linear_def vector_eq
     by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps)
-qed    
+qed
 
 lemma isometry_linear:
   "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
@@ -1005,7 +1005,7 @@
 
 lemma orthogonal_transformation_isometry:
   "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
-  unfolding orthogonal_transformation 
+  unfolding orthogonal_transformation
   apply (rule iffI)
   apply clarify
   apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_def)
@@ -1028,12 +1028,12 @@
   and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y"
   shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
 proof-
-  {fix x y x' y' x0 y0 x0' y0' :: "real ^'n" 
+  {fix x y x' y' x0 y0 x0' y0' :: "real ^'n"
     assume H: "x = norm x *s x0" "y = norm y *s y0"
-    "x' = norm x *s x0'" "y' = norm y *s y0'" 
+    "x' = norm x *s x0'" "y' = norm y *s y0'"
     "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1"
     "norm(x0' - y0') = norm(x0 - y0)"
-    
+
     have "norm(x' - y') = norm(x - y)"
       apply (subst H(1))
       apply (subst H(2))
@@ -1055,13 +1055,13 @@
       then have "dist (?g x) (?g y) = dist x y" by simp }
     moreover
     {assume "x = 0" "y \<noteq> 0"
-      then have "dist (?g x) (?g y) = dist x y" 
+      then have "dist (?g x) (?g y) = dist x y"
 	apply (simp add: dist_def norm_mul)
 	apply (rule f1[rule_format])
 	by(simp add: norm_mul field_simps)}
     moreover
     {assume "x \<noteq> 0" "y = 0"
-      then have "dist (?g x) (?g y) = dist x y" 
+      then have "dist (?g x) (?g y) = dist x y"
 	apply (simp add: dist_def norm_mul)
 	apply (rule f1[rule_format])
 	by(simp add: norm_mul field_simps)}
@@ -1077,14 +1077,14 @@
 	norm (inverse (norm x) *s x - inverse (norm y) *s y)"
 	using z
 	by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_def])
-      from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" 
+      from z th0[OF th00] have "dist (?g x) (?g y) = dist x y"
 	by (simp add: dist_def)}
     ultimately have "dist (?g x) (?g y) = dist x y" by blast}
   note thd = this
-    show ?thesis 
+    show ?thesis
     apply (rule exI[where x= ?g])
     unfolding orthogonal_transformation_isometry
-      using  g0 thfg thd by metis 
+      using  g0 thfg thd by metis
 qed
 
 (* ------------------------------------------------------------------------- *)
@@ -1094,7 +1094,7 @@
 definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
 definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
 
-lemma orthogonal_rotation_or_rotoinversion: 
+lemma orthogonal_rotation_or_rotoinversion:
   fixes Q :: "'a::ordered_idom^'n^'n"
   shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
   by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
@@ -1104,9 +1104,9 @@
 
 lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp
 
-lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" 
+lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2"
   by (simp add: nat_number setprod_numseg mult_commute)
-lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" 
+lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3"
   by (simp add: nat_number setprod_numseg mult_commute)
 
 lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
@@ -1116,7 +1116,7 @@
 proof-
   have f12: "finite {2::nat}" "1 \<notin> {2::nat}" by auto
   have th12: "{1 .. 2} = insert (1::nat) {2}" by auto
-  show ?thesis 
+  show ?thesis
   apply (simp add: det_def dimindex_def th12 del: One_nat_def)
   unfolding setsum_over_permutations_insert[OF f12]
   unfolding permutes_sing
@@ -1124,7 +1124,7 @@
   by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
 qed
 
-lemma det_3: "det (A::'a::comm_ring_1^3^3) = 
+lemma det_3: "det (A::'a::comm_ring_1^3^3) =
   A$1$1 * A$2$2 * A$3$3 +
   A$1$2 * A$2$3 * A$3$1 +
   A$1$3 * A$2$1 * A$3$2 -
@@ -1136,7 +1136,7 @@
   have f23: "finite {(3::nat)}" "2 \<notin> {(3::nat)}" by auto
   have th12: "{1 .. 3} = insert (1::nat) (insert 2 {3})" by auto
 
-  show ?thesis 
+  show ?thesis
   apply (simp add: det_def dimindex_def th12 del: One_nat_def)
   unfolding setsum_over_permutations_insert[OF f123]
   unfolding setsum_over_permutations_insert[OF f23]
--- a/src/HOL/Library/Euclidean_Space.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Euclidean_Space.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -5,7 +5,7 @@
 header {* (Real) Vectors in Euclidean space, and elementary linear algebra.*}
 
 theory Euclidean_Space
-  imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Complex_Main 
+  imports "~~/src/HOL/Decision_Procs/Dense_Linear_Order" Complex_Main
   Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type
   Inner_Product
   uses ("normarith.ML")
@@ -31,26 +31,26 @@
 qed
 
 lemma setsum_singleton[simp]: "setsum f {x} = f x" by simp
-lemma setsum_1: "setsum f {(1::'a::{order,one})..1} = f 1" 
+lemma setsum_1: "setsum f {(1::'a::{order,one})..1} = f 1"
   by (simp add: atLeastAtMost_singleton)
 
-lemma setsum_2: "setsum f {1::nat..2} = f 1 + f 2" 
+lemma setsum_2: "setsum f {1::nat..2} = f 1 + f 2"
   by (simp add: nat_number  atLeastAtMostSuc_conv add_commute)
 
-lemma setsum_3: "setsum f {1::nat..3} = f 1 + f 2 + f 3" 
+lemma setsum_3: "setsum f {1::nat..3} = f 1 + f 2 + f 3"
   by (simp add: nat_number  atLeastAtMostSuc_conv add_commute)
 
 subsection{* Basic componentwise operations on vectors. *}
 
 instantiation "^" :: (plus,type) plus
 begin
-definition  vector_add_def : "op + \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) + (y$i)))" 
+definition  vector_add_def : "op + \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) + (y$i)))"
 instance ..
 end
 
 instantiation "^" :: (times,type) times
 begin
-  definition vector_mult_def : "op * \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))" 
+  definition vector_mult_def : "op * \<equiv> (\<lambda> x y.  (\<chi> i. (x$i) * (y$i)))"
   instance ..
 end
 
@@ -64,12 +64,12 @@
 instance ..
 end
 instantiation "^" :: (zero,type) zero begin
-  definition vector_zero_def : "0 \<equiv> (\<chi> i. 0)" 
+  definition vector_zero_def : "0 \<equiv> (\<chi> i. 0)"
 instance ..
 end
 
 instantiation "^" :: (one,type) one begin
-  definition vector_one_def : "1 \<equiv> (\<chi> i. 1)" 
+  definition vector_one_def : "1 \<equiv> (\<chi> i. 1)"
 instance ..
 end
 
@@ -80,13 +80,13 @@
   x$i <= y$i)"
 definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i : {1 ..
   dimindex (UNIV :: 'b set)}. x$i < y$i)"
- 
+
 instance by (intro_classes)
 end
 
 instantiation "^" :: (scaleR, type) scaleR
 begin
-definition vector_scaleR_def: "scaleR = (\<lambda> r x.  (\<chi> i. scaleR r (x$i)))" 
+definition vector_scaleR_def: "scaleR = (\<lambda> r x.  (\<chi> i. scaleR r (x$i)))"
 instance ..
 end
 
@@ -117,19 +117,19 @@
 lemmas Cart_lambda_beta' = Cart_lambda_beta[rule_format]
 method_setup vector = {*
 let
-  val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym, 
-  @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, 
+  val ss1 = HOL_basic_ss addsimps [@{thm dot_def}, @{thm setsum_addf} RS sym,
+  @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
   @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym]
-  val ss2 = @{simpset} addsimps 
-             [@{thm vector_add_def}, @{thm vector_mult_def},  
-              @{thm vector_minus_def}, @{thm vector_uminus_def}, 
-              @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def}, 
+  val ss2 = @{simpset} addsimps
+             [@{thm vector_add_def}, @{thm vector_mult_def},
+              @{thm vector_minus_def}, @{thm vector_uminus_def},
+              @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def},
               @{thm vector_scaleR_def},
               @{thm Cart_lambda_beta'}, @{thm vector_scalar_mult_def}]
- fun vector_arith_tac ths = 
+ fun vector_arith_tac ths =
    simp_tac ss1
    THEN' (fn i => rtac @{thm setsum_cong2} i
-         ORELSE rtac @{thm setsum_0'} i 
+         ORELSE rtac @{thm setsum_0'} i
          ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i)
    (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
    THEN' asm_full_simp_tac (ss2 addsimps ths)
@@ -145,30 +145,30 @@
 
 text{* Obvious "component-pushing". *}
 
-lemma vec_component: " i \<in> {1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (vec x :: 'a ^ 'n)$i = x" 
-  by (vector vec_def) 
-
-lemma vector_add_component: 
+lemma vec_component: " i \<in> {1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (vec x :: 'a ^ 'n)$i = x"
+  by (vector vec_def)
+
+lemma vector_add_component:
   fixes x y :: "'a::{plus} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
   shows "(x + y)$i = x$i + y$i"
   using i by vector
 
-lemma vector_minus_component: 
+lemma vector_minus_component:
   fixes x y :: "'a::{minus} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
   shows "(x - y)$i = x$i - y$i"
   using i  by vector
 
-lemma vector_mult_component: 
+lemma vector_mult_component:
   fixes x y :: "'a::{times} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
   shows "(x * y)$i = x$i * y$i"
   using i by vector
 
-lemma vector_smult_component: 
+lemma vector_smult_component:
   fixes y :: "'a::{times} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
   shows "(c *s y)$i = c * (y$i)"
   using i by vector
 
-lemma vector_uminus_component: 
+lemma vector_uminus_component:
   fixes x :: "'a::{uminus} ^ 'n"  assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
   shows "(- x)$i = - (x$i)"
   using i by vector
@@ -188,26 +188,26 @@
 
 subsection {* Some frequently useful arithmetic lemmas over vectors. *}
 
-instance "^" :: (semigroup_add,type) semigroup_add 
+instance "^" :: (semigroup_add,type) semigroup_add
   apply (intro_classes) by (vector add_assoc)
 
 
-instance "^" :: (monoid_add,type) monoid_add 
-  apply (intro_classes) by vector+ 
-
-instance "^" :: (group_add,type) group_add 
-  apply (intro_classes) by (vector algebra_simps)+ 
-
-instance "^" :: (ab_semigroup_add,type) ab_semigroup_add 
+instance "^" :: (monoid_add,type) monoid_add
+  apply (intro_classes) by vector+
+
+instance "^" :: (group_add,type) group_add
+  apply (intro_classes) by (vector algebra_simps)+
+
+instance "^" :: (ab_semigroup_add,type) ab_semigroup_add
   apply (intro_classes) by (vector add_commute)
 
 instance "^" :: (comm_monoid_add,type) comm_monoid_add
   apply (intro_classes) by vector
 
-instance "^" :: (ab_group_add,type) ab_group_add 
+instance "^" :: (ab_group_add,type) ab_group_add
   apply (intro_classes) by vector+
 
-instance "^" :: (cancel_semigroup_add,type) cancel_semigroup_add 
+instance "^" :: (cancel_semigroup_add,type) cancel_semigroup_add
   apply (intro_classes)
   by (vector Cart_eq)+
 
@@ -218,30 +218,30 @@
 instance "^" :: (real_vector, type) real_vector
   by default (vector scaleR_left_distrib scaleR_right_distrib)+
 
-instance "^" :: (semigroup_mult,type) semigroup_mult 
+instance "^" :: (semigroup_mult,type) semigroup_mult
   apply (intro_classes) by (vector mult_assoc)
 
-instance "^" :: (monoid_mult,type) monoid_mult 
+instance "^" :: (monoid_mult,type) monoid_mult
   apply (intro_classes) by vector+
 
-instance "^" :: (ab_semigroup_mult,type) ab_semigroup_mult 
+instance "^" :: (ab_semigroup_mult,type) ab_semigroup_mult
   apply (intro_classes) by (vector mult_commute)
 
-instance "^" :: (ab_semigroup_idem_mult,type) ab_semigroup_idem_mult 
+instance "^" :: (ab_semigroup_idem_mult,type) ab_semigroup_idem_mult
   apply (intro_classes) by (vector mult_idem)
 
-instance "^" :: (comm_monoid_mult,type) comm_monoid_mult 
+instance "^" :: (comm_monoid_mult,type) comm_monoid_mult
   apply (intro_classes) by vector
 
 fun vector_power :: "('a::{one,times} ^'n) \<Rightarrow> nat \<Rightarrow> 'a^'n" where
   "vector_power x 0 = 1"
   | "vector_power x (Suc n) = x * vector_power x n"
 
-instantiation "^" :: (recpower,type) recpower 
+instantiation "^" :: (recpower,type) recpower
 begin
   definition vec_power_def: "op ^ \<equiv> vector_power"
-  instance 
-  apply (intro_classes) by (simp_all add: vec_power_def) 
+  instance
+  apply (intro_classes) by (simp_all add: vec_power_def)
 end
 
 instance "^" :: (semiring,type) semiring
@@ -250,16 +250,16 @@
 instance "^" :: (semiring_0,type) semiring_0
   apply (intro_classes) by (vector ring_simps)+
 instance "^" :: (semiring_1,type) semiring_1
-  apply (intro_classes) apply vector using dimindex_ge_1 by auto 
+  apply (intro_classes) apply vector using dimindex_ge_1 by auto
 instance "^" :: (comm_semiring,type) comm_semiring
   apply (intro_classes) by (vector ring_simps)+
 
-instance "^" :: (comm_semiring_0,type) comm_semiring_0 by (intro_classes) 
+instance "^" :: (comm_semiring_0,type) comm_semiring_0 by (intro_classes)
 instance "^" :: (cancel_comm_monoid_add, type) cancel_comm_monoid_add ..
-instance "^" :: (semiring_0_cancel,type) semiring_0_cancel by (intro_classes) 
-instance "^" :: (comm_semiring_0_cancel,type) comm_semiring_0_cancel by (intro_classes) 
-instance "^" :: (ring,type) ring by (intro_classes) 
-instance "^" :: (semiring_1_cancel,type) semiring_1_cancel by (intro_classes) 
+instance "^" :: (semiring_0_cancel,type) semiring_0_cancel by (intro_classes)
+instance "^" :: (comm_semiring_0_cancel,type) comm_semiring_0_cancel by (intro_classes)
+instance "^" :: (ring,type) ring by (intro_classes)
+instance "^" :: (semiring_1_cancel,type) semiring_1_cancel by (intro_classes)
 instance "^" :: (comm_semiring_1,type) comm_semiring_1 by (intro_classes)
 
 instance "^" :: (ring_1,type) ring_1 ..
@@ -273,31 +273,31 @@
 
 instance "^" :: (real_algebra_1,type) real_algebra_1 ..
 
-lemma of_nat_index: 
+lemma of_nat_index:
   "i\<in>{1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
   apply (induct n)
   apply vector
   apply vector
   done
-lemma zero_index[simp]: 
+lemma zero_index[simp]:
   "i\<in>{1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (0 :: 'a::zero ^'n)$i = 0" by vector
 
-lemma one_index[simp]: 
+lemma one_index[simp]:
   "i\<in>{1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> (1 :: 'a::one ^'n)$i = 1" by vector
 
 lemma one_plus_of_nat_neq_0: "(1::'a::semiring_char_0) + of_nat n \<noteq> 0"
 proof-
   have "(1::'a) + of_nat n = 0 \<longleftrightarrow> of_nat 1 + of_nat n = (of_nat 0 :: 'a)" by simp
-  also have "\<dots> \<longleftrightarrow> 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff) 
-  finally show ?thesis by simp 
+  also have "\<dots> \<longleftrightarrow> 1 + n = 0" by (simp only: of_nat_add[symmetric] of_nat_eq_iff)
+  finally show ?thesis by simp
 qed
 
-instance "^" :: (semiring_char_0,type) semiring_char_0 
-proof (intro_classes) 
+instance "^" :: (semiring_char_0,type) semiring_char_0
+proof (intro_classes)
   fix m n ::nat
   show "(of_nat m :: 'a^'b) = of_nat n \<longleftrightarrow> m = n"
   proof(induct m arbitrary: n)
-    case 0 thus ?case apply vector 
+    case 0 thus ?case apply vector
       apply (induct n,auto simp add: ring_simps)
       using dimindex_ge_1 apply auto
       apply vector
@@ -323,24 +323,24 @@
 instance "^" :: (comm_ring_1,type) comm_ring_1 by intro_classes
 instance "^" :: (ring_char_0,type) ring_char_0 by intro_classes
 
-lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"  
+lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
   by (vector mult_assoc)
-lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" 
+lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
   by (vector ring_simps)
-lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" 
+lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
   by (vector ring_simps)
 lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
 lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
-lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" 
+lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
   by (vector ring_simps)
 lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
 lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
 lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector
 lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
-lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" 
+lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
   by (vector ring_simps)
 
-lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)" 
+lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
   apply (auto simp add: vec_def Cart_eq vec_component Cart_lambda_beta )
   using dimindex_ge_1 apply auto done
 
@@ -581,15 +581,15 @@
 
 subsection{* Properties of the dot product.  *}
 
-lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \<bullet> y = y \<bullet> x" 
+lemma dot_sym: "(x::'a:: {comm_monoid_add, ab_semigroup_mult} ^ 'n) \<bullet> y = y \<bullet> x"
   by (vector mult_commute)
 lemma dot_ladd: "((x::'a::ring ^ 'n) + y) \<bullet> z = (x \<bullet> z) + (y \<bullet> z)"
   by (vector ring_simps)
-lemma dot_radd: "x \<bullet> (y + (z::'a::ring ^ 'n)) = (x \<bullet> y) + (x \<bullet> z)" 
+lemma dot_radd: "x \<bullet> (y + (z::'a::ring ^ 'n)) = (x \<bullet> y) + (x \<bullet> z)"
   by (vector ring_simps)
-lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \<bullet> z = (x \<bullet> z) - (y \<bullet> z)" 
+lemma dot_lsub: "((x::'a::ring ^ 'n) - y) \<bullet> z = (x \<bullet> z) - (y \<bullet> z)"
   by (vector ring_simps)
-lemma dot_rsub: "(x::'a::ring ^ 'n) \<bullet> (y - z) = (x \<bullet> y) - (x \<bullet> z)" 
+lemma dot_rsub: "(x::'a::ring ^ 'n) \<bullet> (y - z) = (x \<bullet> y) - (x \<bullet> z)"
   by (vector ring_simps)
 lemma dot_lmult: "(c *s x) \<bullet> y = (c::'a::ring) * (x \<bullet> y)" by (vector ring_simps)
 lemma dot_rmult: "x \<bullet> (c *s y) = (c::'a::comm_ring) * (x \<bullet> y)" by (vector ring_simps)
@@ -625,8 +625,8 @@
   ultimately show ?thesis by metis
 qed
 
-lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x] 
-  by (auto simp add: le_less) 
+lemma dot_pos_lt[simp]: "(0 < x \<bullet> x) \<longleftrightarrow> (x::'a::{ordered_ring_strict,ring_no_zero_divisors} ^ 'n) \<noteq> 0" using dot_eq_0[of x] dot_pos_le[of x]
+  by (auto simp add: le_less)
 
 subsection{* The collapse of the general concepts to dimension one. *}
 
@@ -642,13 +642,13 @@
 lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
   by (simp add: vector_norm_def dimindex_def)
 
-lemma norm_real: "norm(x::real ^ 1) = abs(x$1)" 
+lemma norm_real: "norm(x::real ^ 1) = abs(x$1)"
   by (simp add: norm_vector_1)
 
 text{* Metric *}
 
 text {* FIXME: generalize to arbitrary @{text real_normed_vector} types *}
-definition dist:: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> real" where 
+definition dist:: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> real" where
   "dist x y = norm (x - y)"
 
 lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))"
@@ -667,14 +667,14 @@
   shows "\<exists>x \<ge> a. x <= b \<and> f x \<notin> e1 \<and> f x \<notin> e2" (is "\<exists> x. ?P x")
 proof-
   let ?S = "{c. \<forall>x \<ge> a. x <= c \<longrightarrow> f x \<in> e1}"
-  have Se: " \<exists>x. x \<in> ?S" apply (rule exI[where x=a]) by (auto simp add: fa) 
-  have Sub: "\<exists>y. isUb UNIV ?S y" 
+  have Se: " \<exists>x. x \<in> ?S" apply (rule exI[where x=a]) by (auto simp add: fa)
+  have Sub: "\<exists>y. isUb UNIV ?S y"
     apply (rule exI[where x= b])
-    using ab fb e12 by (auto simp add: isUb_def setle_def)  
-  from reals_complete[OF Se Sub] obtain l where 
+    using ab fb e12 by (auto simp add: isUb_def setle_def)
+  from reals_complete[OF Se Sub] obtain l where
     l: "isLub UNIV ?S l"by blast
   have alb: "a \<le> l" "l \<le> b" using l ab fa fb e12
-    apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)    
+    apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
     by (metis linorder_linear)
   have ale1: "\<forall>z \<ge> a. z < l \<longrightarrow> f z \<in> e1" using l
     apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
@@ -685,11 +685,11 @@
     {assume le2: "f l \<in> e2"
       from le2 fa fb e12 alb have la: "l \<noteq> a" by metis
       hence lap: "l - a > 0" using alb by arith
-      from e2[rule_format, OF le2] obtain e where 
+      from e2[rule_format, OF le2] obtain e where
 	e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e2" by metis
-      from dst[OF alb e(1)] obtain d where 
+      from dst[OF alb e(1)] obtain d where
 	d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
-      have "\<exists>d'. d' < d \<and> d' >0 \<and> l - d' > a" using lap d(1) 
+      have "\<exists>d'. d' < d \<and> d' >0 \<and> l - d' > a" using lap d(1)
 	apply ferrack by arith
       then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis
       from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e2" by metis
@@ -701,16 +701,16 @@
     {assume le1: "f l \<in> e1"
     from le1 fa fb e12 alb have lb: "l \<noteq> b" by metis
       hence blp: "b - l > 0" using alb by arith
-      from e1[rule_format, OF le1] obtain e where 
+      from e1[rule_format, OF le1] obtain e where
 	e: "e > 0" "\<forall>y. dist y (f l) < e \<longrightarrow> y \<in> e1" by metis
-      from dst[OF alb e(1)] obtain d where 
+      from dst[OF alb e(1)] obtain d where
 	d: "d > 0" "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> dist (f y) (f l) < e" by metis
-      have "\<exists>d'. d' < d \<and> d' >0" using d(1) by dlo 
+      have "\<exists>d'. d' < d \<and> d' >0" using d(1) by dlo
       then obtain d' where d': "d' > 0" "d' < d" by metis
       from d e have th0: "\<forall>y. \<bar>y - l\<bar> < d \<longrightarrow> f y \<in> e1" by auto
       hence "\<forall>y. l \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" using d' by auto
       with ale1 have "\<forall>y. a \<le> y \<and> y \<le> l + d' \<longrightarrow> f y \<in> e1" by auto
-      with l d' have False 
+      with l d' have False
 	by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) }
     ultimately show ?thesis using alb by metis
 qed
@@ -719,7 +719,7 @@
 
 lemma square_bound_lemma: "(x::real) < (1 + x) * (1 + x)"
 proof-
-  have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith 
+  have "(x + 1/2)^2 + 3/4 > 0" using zero_le_power2[of "x+1/2"] by arith
   thus ?thesis by (simp add: ring_simps power2_eq_square)
 qed
 
@@ -740,14 +740,14 @@
 lemma real_less_rsqrt: "x^2 < y \<Longrightarrow> x < sqrt y"
   using real_sqrt_less_mono[of "x^2" y] by simp
 
-lemma sqrt_even_pow2: assumes n: "even n" 
+lemma sqrt_even_pow2: assumes n: "even n"
   shows "sqrt(2 ^ n) = 2 ^ (n div 2)"
 proof-
-  from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2 
-    by (auto simp add: nat_number) 
+  from n obtain m where m: "n = 2*m" unfolding even_nat_equiv_def2
+    by (auto simp add: nat_number)
   from m  have "sqrt(2 ^ n) = sqrt ((2 ^ m) ^ 2)"
     by (simp only: power_mult[symmetric] mult_commute)
-  then show ?thesis  using m by simp 
+  then show ?thesis  using m by simp
 qed
 
 lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)"
@@ -786,7 +786,7 @@
   {assume "norm x = 0"
     hence ?thesis by (simp add: dot_lzero dot_rzero)}
   moreover
-  {assume "norm y = 0" 
+  {assume "norm y = 0"
     hence ?thesis by (simp add: dot_lzero dot_rzero)}
   moreover
   {assume h: "norm x \<noteq> 0" "norm y \<noteq> 0"
@@ -829,7 +829,7 @@
 lemma norm_le_l1: "norm (x:: real ^'n) <= setsum(\<lambda>i. \<bar>x$i\<bar>) {1..dimindex(UNIV::'n set)}"
   by (simp add: vector_norm_def setL2_le_setsum)
 
-lemma real_abs_norm[simp]: "\<bar> norm x\<bar> = norm (x :: real ^'n)" 
+lemma real_abs_norm[simp]: "\<bar> norm x\<bar> = norm (x :: real ^'n)"
   by (rule abs_norm_cancel)
 lemma real_abs_sub_norm: "\<bar>norm(x::real ^'n) - norm y\<bar> <= norm(x - y)"
   by (rule norm_triangle_ineq3)
@@ -863,7 +863,7 @@
   apply arith
   done
 
-lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2" 
+lemma norm_ge_square: "norm(x) >= a \<longleftrightarrow> a <= 0 \<or> x \<bullet> x >= a ^ 2"
   apply (simp add: dot_square_norm real_abs_le_square_iff[symmetric])
   using norm_ge_zero[of x]
   apply arith
@@ -891,7 +891,7 @@
 next
   assume ?rhs
   then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y\<bullet> y = 0" by simp
-  hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" 
+  hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0"
     by (simp add: dot_rsub dot_lsub dot_sym)
   then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps dot_lsub dot_rsub)
   then show "x = y" by (simp add: dot_eq_0)
@@ -919,13 +919,13 @@
 lemma pth_4: "0 *s (x::real^'n) == 0" "c *s 0 = (0::real ^ 'n)" by vector+
 lemma pth_5: "c *s (d *s x) == (c * d) *s (x::real ^ 'n)" by (atomize (full)) vector
 lemma pth_6: "(c::real) *s (x + y) == c *s x + c *s y" by (atomize (full)) (vector ring_simps)
-lemma pth_7: "0 + x == (x::real^'n)" "x + 0 == x" by simp_all 
-lemma pth_8: "(c::real) *s x + d *s x == (c + d) *s x" by (atomize (full)) (vector ring_simps) 
+lemma pth_7: "0 + x == (x::real^'n)" "x + 0 == x" by simp_all
+lemma pth_8: "(c::real) *s x + d *s x == (c + d) *s x" by (atomize (full)) (vector ring_simps)
 lemma pth_9: "((c::real) *s x + z) + d *s x == (c + d) *s x + z"
   "c *s x + (d *s x + z) == (c + d) *s x + z"
   "(c *s x + w) + (d *s x + z) == (c + d) *s x + (w + z)" by ((atomize (full)), vector ring_simps)+
 lemma pth_a: "(0::real) *s x + y == y" by (atomize (full)) vector
-lemma pth_b: "(c::real) *s x + d *s y == c *s x + d *s y" 
+lemma pth_b: "(c::real) *s x + d *s y == c *s x + d *s y"
   "(c *s x + z) + d *s y == c *s x + (z + d *s y)"
   "c *s x + (d *s y + z) == c *s x + (d *s y + z)"
   "(c *s x + w) + (d *s y + z) == c *s x + (w + (d *s y + z))"
@@ -941,7 +941,7 @@
 
 lemma real_eq_0_iff_le_ge_0: "(x::real) = 0 == x \<ge> 0 \<and> -x \<ge> 0" by arith
 
-lemma norm_pths: 
+lemma norm_pths:
   "(x::real ^'n) = y \<longleftrightarrow> norm (x - y) \<le> 0"
   "x \<noteq> y \<longleftrightarrow> \<not> (norm (x - y) \<le> 0)"
   using norm_ge_zero[of "x - y"] by auto
@@ -967,26 +967,26 @@
 
 lemma dist_eq_0[simp]: "dist x y = 0 \<longleftrightarrow> x = y" by norm
 
-lemma dist_pos_lt: "x \<noteq> y ==> 0 < dist x y" by norm 
-lemma dist_nz:  "x \<noteq> y \<longleftrightarrow> 0 < dist x y" by norm 
-
-lemma dist_triangle_le: "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e" by norm 
-
-lemma dist_triangle_lt: "dist x z + dist y z < e ==> dist x y < e" by norm 
-
-lemma dist_triangle_half_l: "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 ==> dist x1 x2 < e" by norm 
-
-lemma dist_triangle_half_r: "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 ==> dist x1 x2 < e" by norm 
+lemma dist_pos_lt: "x \<noteq> y ==> 0 < dist x y" by norm
+lemma dist_nz:  "x \<noteq> y \<longleftrightarrow> 0 < dist x y" by norm
+
+lemma dist_triangle_le: "dist x z + dist y z <= e \<Longrightarrow> dist x y <= e" by norm
+
+lemma dist_triangle_lt: "dist x z + dist y z < e ==> dist x y < e" by norm
+
+lemma dist_triangle_half_l: "dist x1 y < e / 2 \<Longrightarrow> dist x2 y < e / 2 ==> dist x1 x2 < e" by norm
+
+lemma dist_triangle_half_r: "dist y x1 < e / 2 \<Longrightarrow> dist y x2 < e / 2 ==> dist x1 x2 < e" by norm
 
 lemma dist_triangle_add: "dist (x + y) (x' + y') <= dist x x' + dist y y'"
-  by norm 
-
-lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y" 
-  unfolding dist_def vector_ssub_ldistrib[symmetric] norm_mul .. 
-
-lemma dist_triangle_add_half: " dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 ==> dist(x + y) (x' + y') < e" by norm 
-
-lemma dist_le_0[simp]: "dist x y <= 0 \<longleftrightarrow> x = y" by norm 
+  by norm
+
+lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"
+  unfolding dist_def vector_ssub_ldistrib[symmetric] norm_mul ..
+
+lemma dist_triangle_add_half: " dist x x' < e / 2 \<Longrightarrow> dist y y' < e / 2 ==> dist(x + y) (x' + y') < e" by norm
+
+lemma dist_le_0[simp]: "dist x y <= 0 \<longleftrightarrow> x = y" by norm
 
 lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"
   apply vector
@@ -996,24 +996,24 @@
   apply (auto simp add: vector_component zero_index)
   done
 
-lemma setsum_clauses: 
+lemma setsum_clauses:
   shows "setsum f {} = 0"
   and "finite S \<Longrightarrow> setsum f (insert x S) =
                  (if x \<in> S then setsum f S else f x + setsum f S)"
   by (auto simp add: insert_absorb)
 
-lemma setsum_cmul: 
+lemma setsum_cmul:
   fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
   shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
   by (simp add: setsum_eq Cart_eq Cart_lambda_beta vector_component setsum_right_distrib)
 
-lemma setsum_component: 
+lemma setsum_component:
   fixes f:: " 'a \<Rightarrow> ('b::semiring_1) ^'n"
   assumes i: "i \<in> {1 .. dimindex(UNIV:: 'n set)}"
   shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
   using i by (simp add: setsum_eq Cart_lambda_beta)
 
-lemma setsum_norm: 
+lemma setsum_norm:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
   assumes fS: "finite S"
   shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
@@ -1027,7 +1027,7 @@
   finally  show ?case  using "2.hyps" by simp
 qed
 
-lemma real_setsum_norm: 
+lemma real_setsum_norm:
   fixes f :: "'a \<Rightarrow> real ^'n"
   assumes fS: "finite S"
   shows "norm (setsum f S) <= setsum (\<lambda>x. norm(f x)) S"
@@ -1041,25 +1041,25 @@
   finally  show ?case  using "2.hyps" by simp
 qed
 
-lemma setsum_norm_le: 
+lemma setsum_norm_le:
   fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
   assumes fS: "finite S"
   and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
   shows "norm (setsum f S) \<le> setsum g S"
 proof-
-  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S" 
+  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"
     by - (rule setsum_mono, simp)
   then show ?thesis using setsum_norm[OF fS, of f] fg
     by arith
 qed
 
-lemma real_setsum_norm_le: 
+lemma real_setsum_norm_le:
   fixes f :: "'a \<Rightarrow> real ^ 'n"
   assumes fS: "finite S"
   and fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
   shows "norm (setsum f S) \<le> setsum g S"
 proof-
-  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S" 
+  from fg have "setsum (\<lambda>x. norm(f x)) S <= setsum g S"
     by - (rule setsum_mono, simp)
   then show ?thesis using real_setsum_norm[OF fS, of f] fg
     by arith
@@ -1089,9 +1089,9 @@
   case 1 then show ?case by (simp add: vector_smult_lzero)
 next
   case (2 x F)
-  from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v" 
+  from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v"
     by simp
-  also have "\<dots> = f x *s v + setsum f F *s v" 
+  also have "\<dots> = f x *s v + setsum f F *s v"
     by (simp add: vector_sadd_rdistrib)
   also have "\<dots> = setsum (\<lambda>x. f x *s v) (insert x F)" using "2.hyps" by simp
   finally show ?case .
@@ -1105,20 +1105,20 @@
 proof-
   let ?A = "{m .. n}"
   let ?B = "{n + 1 .. n + p}"
-  have eq: "{m .. n+p} = ?A \<union> ?B" using mn by auto 
+  have eq: "{m .. n+p} = ?A \<union> ?B" using mn by auto
   have d: "?A \<inter> ?B = {}" by auto
   from setsum_Un_disjoint[of "?A" "?B" f] eq d show ?thesis by auto
 qed
 
 lemma setsum_natinterval_left:
-  assumes mn: "(m::nat) <= n" 
+  assumes mn: "(m::nat) <= n"
   shows "setsum f {m..n} = f m + setsum f {m + 1..n}"
 proof-
   from mn have "{m .. n} = insert m {m+1 .. n}" by auto
   then show ?thesis by auto
 qed
 
-lemma setsum_natinterval_difff: 
+lemma setsum_natinterval_difff:
   fixes f:: "nat \<Rightarrow> ('a::ab_group_add)"
   shows  "setsum (\<lambda>k. f k - f(k + 1)) {(m::nat) .. n} =
           (if m <= n then f m - f(n + 1) else 0)"
@@ -1136,8 +1136,8 @@
 proof-
   {fix x assume "x \<in> S" then have "{y. y\<in> f`S \<and> f x = y} = {f x}" by auto}
   note th0 = this
-  have "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S" 
-    apply (rule setsum_cong2) 
+  have "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
+    apply (rule setsum_cong2)
     by (simp add: th0)
   also have "\<dots> = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
     apply (rule setsum_setsum_restrict[OF fS])
@@ -1149,14 +1149,14 @@
 lemma setsum_group:
   assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
   shows "setsum (\<lambda>y. setsum g {x. x\<in> S \<and> f x = y}) T = setsum g S"
-  
+
 apply (subst setsum_image_gen[OF fS, of g f])
 apply (rule setsum_mono_zero_right[OF fT fST])
 by (auto intro: setsum_0')
 
 lemma vsum_norm_allsubsets_bound:
   fixes f:: "'a \<Rightarrow> real ^'n"
-  assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e" 
+  assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
   shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real (dimindex(UNIV :: 'n set)) *  e"
 proof-
   let ?d = "real (dimindex (UNIV ::'n set))"
@@ -1183,9 +1183,9 @@
     have Pne: "setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn \<le> e"
       using i component_le_norm[OF i, of "setsum (\<lambda>x. - f x) ?Pn"]  fPs[OF PnP]
       by (auto simp add: setsum_negf setsum_component vector_component intro: abs_le_D1)
-    have "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn" 
+    have "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P = setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pp + setsum (\<lambda>x. \<bar>f x $ i\<bar>) ?Pn"
       apply (subst thp)
-      apply (rule setsum_Un_zero) 
+      apply (rule setsum_Un_zero)
       using fP thp0 by auto
     also have "\<dots> \<le> 2*e" using Pne Ppe by arith
     finally show "setsum (\<lambda>x. \<bar>f x $ i\<bar>) P \<le> 2*e" .
@@ -1204,13 +1204,13 @@
 
 definition "basis k = (\<chi> i. if i = k then 1 else 0)"
 
-lemma delta_mult_idempotent: 
+lemma delta_mult_idempotent:
   "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto)
 
 lemma norm_basis:
   assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
   shows "norm (basis k :: real ^'n) = 1"
-  using k 
+  using k
   apply (simp add: basis_def real_vector_norm_def dot_def)
   apply (vector delta_mult_idempotent)
   using setsum_delta[of "{1 .. dimindex (UNIV :: 'n set)}" "k" "\<lambda>k. 1::real"]
@@ -1228,7 +1228,7 @@
   apply (rule exI[where x="c *s basis 1"])
   by (simp only: norm_mul norm_basis_1)
 
-lemma vector_choose_dist: assumes e: "0 <= e" 
+lemma vector_choose_dist: assumes e: "0 <= e"
   shows "\<exists>(y::real^'n). dist x y = e"
 proof-
   from vector_choose_size[OF e] obtain c:: "real ^'n"  where "norm c = e"
@@ -1250,7 +1250,7 @@
   "setsum (\<lambda>i. (x$i) *s basis i) {1 .. dimindex (UNIV :: 'n set)} = (x::('a::ring_1) ^'n)" (is "?lhs = ?rhs" is "setsum ?f ?S = _")
   by (auto simp add: Cart_eq basis_component[where ?'n = "'n"] setsum_component vector_component cond_value_iff setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong)
 
-lemma basis_expansion_unique: 
+lemma basis_expansion_unique:
   "setsum (\<lambda>i. f i *s basis i) {1 .. dimindex (UNIV :: 'n set)} = (x::('a::comm_ring_1) ^'n) \<longleftrightarrow> (\<forall>i\<in>{1 .. dimindex(UNIV:: 'n set)}. f i = x$i)"
   by (simp add: Cart_eq setsum_component vector_component basis_component setsum_delta cond_value_iff cong del: if_weak_cong)
 
@@ -1266,7 +1266,7 @@
 lemma basis_eq_0: "basis i = (0::'a::semiring_1^'n) \<longleftrightarrow> i \<notin> {1..dimindex(UNIV ::'n set)}"
   by (auto simp add: Cart_eq basis_component zero_index)
 
-lemma basis_nonzero: 
+lemma basis_nonzero:
   assumes k: "k \<in> {1 .. dimindex(UNIV ::'n set)}"
   shows "basis k \<noteq> (0:: 'a::semiring_1 ^'n)"
   using k by (simp add: basis_eq_0)
@@ -1294,15 +1294,15 @@
 definition "orthogonal x y \<longleftrightarrow> (x \<bullet> y = 0)"
 
 lemma orthogonal_basis:
-  assumes i:"i \<in> {1 .. dimindex(UNIV ::'n set)}" 
+  assumes i:"i \<in> {1 .. dimindex(UNIV ::'n set)}"
   shows "orthogonal (basis i :: 'a^'n) x \<longleftrightarrow> x$i = (0::'a::ring_1)"
   using i
   by (auto simp add: orthogonal_def dot_def basis_def Cart_lambda_beta cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
 
 lemma orthogonal_basis_basis:
-  assumes i:"i \<in> {1 .. dimindex(UNIV ::'n set)}" 
-  and j: "j \<in> {1 .. dimindex(UNIV ::'n set)}" 
-  shows "orthogonal (basis i :: 'a::ring_1^'n) (basis j) \<longleftrightarrow> i \<noteq> j" 
+  assumes i:"i \<in> {1 .. dimindex(UNIV ::'n set)}"
+  and j: "j \<in> {1 .. dimindex(UNIV ::'n set)}"
+  shows "orthogonal (basis i :: 'a::ring_1^'n) (basis j) \<longleftrightarrow> i \<noteq> j"
   unfolding orthogonal_basis[OF i] basis_component[OF i] by simp
 
   (* FIXME : Maybe some of these require less than comm_ring, but not all*)
@@ -1443,14 +1443,14 @@
 
 lemma linear_neg: "linear (f :: 'a::ring_1 ^'n \<Rightarrow> _) ==> f (-x) = - f x"
   unfolding vector_sneg_minus1
-  using linear_cmul[of f] by auto 
-
-lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def) 
+  using linear_cmul[of f] by auto
+
+lemma linear_add: "linear f ==> f(x + y) = f x + f y" by (metis linear_def)
 
 lemma linear_sub: "linear (f::'a::ring_1 ^'n \<Rightarrow> _) ==> f(x - y) = f x - f y"
   by (simp add: diff_def linear_add linear_neg)
 
-lemma linear_setsum: 
+lemma linear_setsum:
   fixes f:: "'a::semiring_1^'n \<Rightarrow> _"
   assumes lf: "linear f" and fS: "finite S"
   shows "f (setsum g S) = setsum (f o g) S"
@@ -1470,7 +1470,7 @@
   assumes lf: "linear f" and fS: "finite S"
   shows "f (setsum (\<lambda>i. c i *s v i) S) = setsum (\<lambda>i. c i *s f (v i)) S"
   using linear_setsum[OF lf fS, of "\<lambda>i. c i *s v i" , unfolded o_def]
-  linear_cmul[OF lf] by simp 
+  linear_cmul[OF lf] by simp
 
 lemma linear_injective_0:
   assumes lf: "linear (f:: 'a::ring_1 ^ 'n \<Rightarrow> _)"
@@ -1478,7 +1478,7 @@
 proof-
   have "inj f \<longleftrightarrow> (\<forall> x y. f x = f y \<longrightarrow> x = y)" by (simp add: inj_on_def)
   also have "\<dots> \<longleftrightarrow> (\<forall> x y. f x - f y = 0 \<longrightarrow> x - y = 0)" by simp
-  also have "\<dots> \<longleftrightarrow> (\<forall> x y. f (x - y) = 0 \<longrightarrow> x - y = 0)" 
+  also have "\<dots> \<longleftrightarrow> (\<forall> x y. f (x - y) = 0 \<longrightarrow> x - y = 0)"
     by (simp add: linear_sub[OF lf])
   also have "\<dots> \<longleftrightarrow> (\<forall> x. f x = 0 \<longrightarrow> x = 0)" by auto
   finally show ?thesis .
@@ -1518,7 +1518,7 @@
   assumes lf: "linear f"
   shows "\<exists>B > 0. \<forall>x. norm (f x) \<le> B * norm x"
 proof-
-  from linear_bounded[OF lf] obtain B where 
+  from linear_bounded[OF lf] obtain B where
     B: "\<forall>x. norm (f x) \<le> B * norm x" by blast
   let ?K = "\<bar>B\<bar> + 1"
   have Kp: "?K > 0" by arith
@@ -1562,15 +1562,15 @@
 
 lemma  (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0"
   using add_imp_eq[of x y 0] by auto
-    
-lemma bilinear_lzero: 
+
+lemma bilinear_lzero:
   fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"
-  using bilinear_ladd[OF bh, of 0 0 x] 
+  using bilinear_ladd[OF bh, of 0 0 x]
     by (simp add: eq_add_iff ring_simps)
 
-lemma bilinear_rzero: 
+lemma bilinear_rzero:
   fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"
-  using bilinear_radd[OF bh, of x 0 0 ] 
+  using bilinear_radd[OF bh, of x 0 0 ]
     by (simp add: eq_add_iff ring_simps)
 
 lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ 'n)) z = h x z - h y z"
@@ -1583,7 +1583,7 @@
   fixes h:: "'a ^'n \<Rightarrow> 'a::semiring_1^'m \<Rightarrow> 'a ^ 'k"
   assumes bh: "bilinear h" and fS: "finite S" and fT: "finite T"
   shows "h (setsum f S) (setsum g T) = setsum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "
-proof- 
+proof-
   have "h (setsum f S) (setsum g T) = setsum (\<lambda>x. h (f x) (setsum g T)) S"
     apply (rule linear_setsum[unfolded o_def])
     using bh fS by (auto simp add: bilinear_def)
@@ -1598,7 +1598,7 @@
   fixes h:: "real ^'m \<Rightarrow> real^'n \<Rightarrow> real ^ 'k"
   assumes bh: "bilinear h"
   shows "\<exists>B. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
-proof- 
+proof-
   let ?M = "{1 .. dimindex (UNIV :: 'm set)}"
   let ?N = "{1 .. dimindex (UNIV :: 'n set)}"
   let ?B = "setsum (\<lambda>(i,j). norm (h (basis i) (basis j))) (?M \<times> ?N)"
@@ -1626,7 +1626,7 @@
   assumes bh: "bilinear h"
   shows "\<exists>B > 0. \<forall>x y. norm (h x y) \<le> B * norm x * norm y"
 proof-
-  from bilinear_bounded[OF bh] obtain B where 
+  from bilinear_bounded[OF bh] obtain B where
     B: "\<forall>x y. norm (h x y) \<le> B * norm x * norm y" by blast
   let ?K = "\<bar>B\<bar> + 1"
   have Kp: "?K > 0" by arith
@@ -1634,11 +1634,11 @@
   {fix x::"real ^'m" and y :: "real ^'n"
     from KB Kp
     have "B * norm x * norm y \<le> ?K * norm x * norm y"
-      apply - 
+      apply -
       apply (rule mult_right_mono, rule mult_right_mono)
       by (auto simp add: norm_ge_zero)
     then have "norm (h x y) \<le> ?K * norm x * norm y"
-      using B[rule_format, of x y] by simp} 
+      using B[rule_format, of x y] by simp}
   with Kp show ?thesis by blast
 qed
 
@@ -1663,14 +1663,14 @@
       have "f x \<bullet> y = f (setsum (\<lambda>i. (x$i) *s basis i) ?N) \<bullet> y"
 	by (simp only: basis_expansion)
       also have "\<dots> = (setsum (\<lambda>i. (x$i) *s f (basis i)) ?N) \<bullet> y"
-	unfolding linear_setsum[OF lf fN] 
+	unfolding linear_setsum[OF lf fN]
 	by (simp add: linear_cmul[OF lf])
       finally have "f x \<bullet> y = x \<bullet> ?w"
 	apply (simp only: )
 	apply (simp add: dot_def setsum_component Cart_lambda_beta setsum_left_distrib setsum_right_distrib vector_component setsum_commute[of _ ?M ?N] ring_simps del: One_nat_def)
 	done}
   }
-  then show ?thesis unfolding adjoint_def 
+  then show ?thesis unfolding adjoint_def
     some_eq_ex[of "\<lambda>f'. \<forall>x y. f x \<bullet> y = x \<bullet> f' y"]
     using choice_iff[of "\<lambda>a b. \<forall>x. f x \<bullet> a = x \<bullet> b "]
     by metis
@@ -1715,27 +1715,27 @@
 
 consts generic_mult :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" (infixr "\<star>" 75)
 
-defs (overloaded) 
+defs (overloaded)
 matrix_matrix_mult_def: "(m:: ('a::semiring_1) ^'n^'m) \<star> (m' :: 'a ^'p^'n) \<equiv> (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) {1 .. dimindex (UNIV :: 'n set)}) ::'a ^ 'p ^'m"
 
-abbreviation 
+abbreviation
   matrix_matrix_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"  (infixl "**" 70)
   where "m ** m' == m\<star> m'"
 
-defs (overloaded) 
+defs (overloaded)
   matrix_vector_mult_def: "(m::('a::semiring_1) ^'n^'m) \<star> (x::'a ^'n) \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) {1..dimindex(UNIV ::'n set)}) :: 'a^'m"
 
-abbreviation 
+abbreviation
   matrix_vector_mult' :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"  (infixl "*v" 70)
-  where 
+  where
   "m *v v == m \<star> v"
 
-defs (overloaded) 
+defs (overloaded)
   vector_matrix_mult_def: "(x::'a^'m) \<star> (m::('a::semiring_1) ^'n^'m) \<equiv> (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (x$i)) {1..dimindex(UNIV :: 'm set)}) :: 'a^'n"
 
-abbreviation 
+abbreviation
   vactor_matrix_mult' :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "  (infixl "v*" 70)
-  where 
+  where
   "v v* m == v \<star> m"
 
 definition "(mat::'a::zero => 'a ^'n^'m) k = (\<chi> i j. if i = j then k else 0)"
@@ -1749,11 +1749,11 @@
 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A \<star> B) + (A \<star> C)"
   by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
 
-lemma setsum_delta': 
-  assumes fS: "finite S" shows 
-  "setsum (\<lambda>k. if a = k then b k else 0) S = 
+lemma setsum_delta':
+  assumes fS: "finite S" shows
+  "setsum (\<lambda>k. if a = k then b k else 0) S =
      (if a\<in> S then b a else 0)"
-  using setsum_delta[OF fS, of a b, symmetric] 
+  using setsum_delta[OF fS, of a b, symmetric]
   by (auto intro: setsum_cong)
 
 lemma matrix_mul_lid: "mat 1 ** A = A"
@@ -1781,7 +1781,7 @@
 
 lemma matrix_vector_mul_lid: "mat 1 *v x = x"
   apply (vector matrix_vector_mult_def mat_def)
-  by (simp add: cond_value_iff cond_application_beta 
+  by (simp add: cond_value_iff cond_application_beta
     setsum_delta' cong del: if_weak_cong)
 
 lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^'m^'n)"
@@ -1796,7 +1796,7 @@
   apply (erule_tac x="i" in ballE)
   by (auto simp add: basis_def cond_value_iff cond_application_beta Cart_lambda_beta setsum_delta[OF finite_atLeastAtMost] cong del: if_weak_cong)
 
-lemma matrix_vector_mul_component: 
+lemma matrix_vector_mul_component:
   assumes k: "k \<in> {1.. dimindex (UNIV :: 'm set)}"
   shows "((A::'a::semiring_1^'n'^'m) *v x)$k = (A$k) \<bullet> x"
   using k
@@ -1813,18 +1813,18 @@
 lemma transp_transp: "transp(transp A) = A"
   by (vector transp_def)
 
-lemma row_transp: 
+lemma row_transp:
   fixes A:: "'a::semiring_1^'n^'m"
   assumes i: "i \<in> {1.. dimindex (UNIV :: 'n set)}"
   shows "row i (transp A) = column i A"
-  using i 
+  using i
   by (simp add: row_def column_def transp_def Cart_eq Cart_lambda_beta)
 
 lemma column_transp:
   fixes A:: "'a::semiring_1^'n^'m"
   assumes i: "i \<in> {1.. dimindex (UNIV :: 'm set)}"
   shows "column i (transp A) = row i A"
-  using i 
+  using i
   by (simp add: row_def column_def transp_def Cart_eq Cart_lambda_beta)
 
 lemma rows_transp: "rows(transp (A::'a::semiring_1^'n^'m)) = columns A"
@@ -1890,8 +1890,8 @@
 lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: 'a:: comm_ring_1 ^ 'n)) = A"
   by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
 
-lemma matrix_compose: 
-  assumes lf: "linear (f::'a::comm_ring_1^'n \<Rightarrow> _)" and lg: "linear g" 
+lemma matrix_compose:
+  assumes lf: "linear (f::'a::comm_ring_1^'n \<Rightarrow> _)" and lg: "linear g"
   shows "matrix (g o f) = matrix g ** matrix f"
   using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]
   by (simp  add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
@@ -1923,9 +1923,9 @@
   done
 
 
-lemma real_convex_bound_lt: 
+lemma real_convex_bound_lt:
   assumes xa: "(x::real) < a" and ya: "y < a" and u: "0 <= u" and v: "0 <= v"
-  and uv: "u + v = 1" 
+  and uv: "u + v = 1"
   shows "u * x + v * y < a"
 proof-
   have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith
@@ -1937,7 +1937,7 @@
     apply (cases "u = 0", simp_all add: uv')
     apply(rule mult_strict_left_mono)
     using uv' apply simp_all
-    
+
     apply (rule add_less_le_mono)
     apply(rule mult_strict_left_mono)
     apply simp_all
@@ -1947,9 +1947,9 @@
   thus ?thesis unfolding th .
 qed
 
-lemma real_convex_bound_le: 
+lemma real_convex_bound_le:
   assumes xa: "(x::real) \<le> a" and ya: "y \<le> a" and u: "0 <= u" and v: "0 <= v"
-  and uv: "u + v = 1" 
+  and uv: "u + v = 1"
   shows "u * x + v * y \<le> a"
 proof-
   from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)
@@ -1969,7 +1969,7 @@
 done
 
 
-lemma triangle_lemma: 
+lemma triangle_lemma:
   assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2"
   shows "x <= y + z"
 proof-
@@ -1992,12 +1992,12 @@
     let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"
     {fix i assume i: "i \<in> ?S"
       with f i have "P i (f i)" by metis
-      then have "P i (?x$i)" using Cart_lambda_beta[of f, rule_format, OF i] by auto 
+      then have "P i (?x$i)" using Cart_lambda_beta[of f, rule_format, OF i] by auto
     }
     hence "\<forall>i \<in> ?S. P i (?x$i)" by metis
     hence ?rhs by metis }
   ultimately show ?thesis by metis
-qed 
+qed
 
 (* Supremum and infimum of real sets *)
 
@@ -2019,7 +2019,7 @@
 lemma rsup_le: assumes Se: "S \<noteq> {}" and Sb: "S *<= b" shows "rsup S \<le> b"
 proof-
   from Sb have bu: "isUb UNIV S b" by (simp add: isUb_def setle_def)
-  from rsup[OF Se] Sb have "isLub UNIV S (rsup S)"  by blast 
+  from rsup[OF Se] Sb have "isLub UNIV S (rsup S)"  by blast
   then show ?thesis using bu by (auto simp add: isLub_def leastP_def setle_def setge_def)
 qed
 
@@ -2030,12 +2030,12 @@
   let ?m = "Max S"
   from Max_ge[OF fS] have Sm: "\<forall> x\<in> S. x \<le> ?m" by metis
   with rsup[OF Se] have lub: "isLub UNIV S (rsup S)" by (metis setle_def)
-  from Max_in[OF fS Se] lub have mrS: "?m \<le> rsup S" 
+  from Max_in[OF fS Se] lub have mrS: "?m \<le> rsup S"
     by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def)
-  moreover 
+  moreover
   have "rsup S \<le> ?m" using Sm lub
     by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-  ultimately  show ?thesis by arith 
+  ultimately  show ?thesis by arith
 qed
 
 lemma rsup_finite_in: assumes fS: "finite S" and Se: "S \<noteq> {}"
@@ -2065,7 +2065,7 @@
 
 lemma rsup_unique: assumes b: "S *<= b" and S: "\<forall>b' < b. \<exists>x \<in> S. b' < x"
   shows "rsup S = b"
-using b S  
+using b S
 unfolding setle_def rsup_alt
 apply -
 apply (rule some_equality)
@@ -2104,7 +2104,7 @@
 lemma rsup_asclose: assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>rsup S - l\<bar> \<le> e"
 proof-
   have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
-  show ?thesis using S b rsup_bounds[of S "l - e" "l+e"] unfolding th 
+  show ?thesis using S b rsup_bounds[of S "l - e" "l+e"] unfolding th
     by  (auto simp add: setge_def setle_def)
 qed
 
@@ -2142,7 +2142,7 @@
 lemma rinf_ge: assumes Se: "S \<noteq> {}" and Sb: "b <=* S" shows "rinf S \<ge> b"
 proof-
   from Sb have bu: "isLb UNIV S b" by (simp add: isLb_def setge_def)
-  from rinf[OF Se] Sb have "isGlb UNIV S (rinf S)"  by blast 
+  from rinf[OF Se] Sb have "isGlb UNIV S (rinf S)"  by blast
   then show ?thesis using bu by (auto simp add: isGlb_def greatestP_def setle_def setge_def)
 qed
 
@@ -2153,12 +2153,12 @@
   let ?m = "Min S"
   from Min_le[OF fS] have Sm: "\<forall> x\<in> S. x \<ge> ?m" by metis
   with rinf[OF Se] have glb: "isGlb UNIV S (rinf S)" by (metis setge_def)
-  from Min_in[OF fS Se] glb have mrS: "?m \<ge> rinf S" 
+  from Min_in[OF fS Se] glb have mrS: "?m \<ge> rinf S"
     by (auto simp add: isGlb_def greatestP_def setle_def setge_def isLb_def)
-  moreover 
+  moreover
   have "rinf S \<ge> ?m" using Sm glb
     by (auto simp add: isGlb_def greatestP_def isLb_def setle_def setge_def)
-  ultimately  show ?thesis by arith 
+  ultimately  show ?thesis by arith
 qed
 
 lemma rinf_finite_in: assumes fS: "finite S" and Se: "S \<noteq> {}"
@@ -2188,7 +2188,7 @@
 
 lemma rinf_unique: assumes b: "b <=* S" and S: "\<forall>b' > b. \<exists>x \<in> S. b' > x"
   shows "rinf S = b"
-using b S  
+using b S
 unfolding setge_def rinf_alt
 apply -
 apply (rule some_equality)
@@ -2226,7 +2226,7 @@
 lemma rinf_asclose: assumes S:"S \<noteq> {}" and b: "\<forall>x\<in>S. \<bar>x - l\<bar> \<le> e" shows "\<bar>rinf S - l\<bar> \<le> e"
 proof-
   have th: "\<And>(x::real) l e. \<bar>x - l\<bar> \<le> e \<longleftrightarrow> l - e \<le> x \<and> x \<le> l + e" by arith
-  show ?thesis using S b rinf_bounds[of S "l - e" "l+e"] unfolding th 
+  show ?thesis using S b rinf_bounds[of S "l - e" "l+e"] unfolding th
     by  (auto simp add: setge_def setle_def)
 qed
 
@@ -2248,7 +2248,7 @@
 
   moreover
   {assume H: ?lhs
-    from H[rule_format, of "basis 1"] 
+    from H[rule_format, of "basis 1"]
     have bp: "b \<ge> 0" using norm_ge_zero[of "f (basis 1)"] dimindex_ge_1[of "UNIV:: 'n set"]
       by (auto simp add: norm_basis elim: order_trans [OF norm_ge_zero])
     {fix x :: "real ^'n"
@@ -2260,9 +2260,9 @@
 	let ?c = "1/ norm x"
 	have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
 	with H have "norm (f(?c*s x)) \<le> b" by blast
-	hence "?c * norm (f x) \<le> b" 
+	hence "?c * norm (f x) \<le> b"
 	  by (simp add: linear_cmul[OF lf] norm_mul)
-	hence "norm (f x) \<le> b * norm x" 
+	hence "norm (f x) \<le> b * norm x"
 	  using n0 norm_ge_zero[of x] by (auto simp add: field_simps)}
       ultimately have "norm (f x) \<le> b * norm x" by blast}
     then have ?rhs by blast}
@@ -2278,16 +2278,16 @@
   {
     let ?S = "{norm (f x) |x. norm x = 1}"
     have Se: "?S \<noteq> {}" using  norm_basis_1 by auto
-    from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b" 
+    from linear_bounded[OF lf] have b: "\<exists> b. ?S *<= b"
       unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def)
     {from rsup[OF Se b, unfolded onorm_def[symmetric]]
-      show "norm (f x) <= onorm f * norm x" 
-	apply - 
+      show "norm (f x) <= onorm f * norm x"
+	apply -
 	apply (rule spec[where x = x])
 	unfolding norm_bound_generalize[OF lf, symmetric]
 	by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
     {
-      show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"  
+      show "\<forall>x. norm (f x) <= b * norm x \<Longrightarrow> onorm f <= b"
 	using rsup[OF Se b, unfolded onorm_def[symmetric]]
 	unfolding norm_bound_generalize[OF lf, symmetric]
 	by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)}
@@ -2297,7 +2297,7 @@
 lemma onorm_pos_le: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" shows "0 <= onorm f"
   using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "basis 1"], unfolded norm_basis_1] by simp
 
-lemma onorm_eq_0: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" 
+lemma onorm_eq_0: assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)"
   shows "onorm f = 0 \<longleftrightarrow> (\<forall>x. f x = 0)"
   using onorm[OF lf]
   apply (auto simp add: onorm_pos_le)
@@ -2317,7 +2317,7 @@
     apply (rule rsup_unique) by (simp_all  add: setle_def)
 qed
 
-lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n \<Rightarrow> real ^'m)" 
+lemma onorm_pos_lt: assumes lf: "linear (f::real ^ 'n \<Rightarrow> real ^'m)"
   shows "0 < onorm f \<longleftrightarrow> ~(\<forall>x. f x = 0)"
   unfolding onorm_eq_0[OF lf, symmetric]
   using onorm_pos_le[OF lf] by arith
@@ -2374,7 +2374,7 @@
 
 definition vec1:: "'a \<Rightarrow> 'a ^ 1" where "vec1 x = (\<chi> i. x)"
 
-definition dest_vec1:: "'a ^1 \<Rightarrow> 'a" 
+definition dest_vec1:: "'a ^1 \<Rightarrow> 'a"
   where "dest_vec1 x = (x$1)"
 
 lemma vec1_component[simp]: "(vec1 x)$1 = x"
@@ -2385,7 +2385,7 @@
 
 lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" by (metis vec1_dest_vec1)
 
-lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" by (metis vec1_dest_vec1) 
+lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" by (metis vec1_dest_vec1)
 
 lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))"  by (metis vec1_dest_vec1)
 
@@ -2446,7 +2446,7 @@
 lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>"
   by (metis vec1_dest_vec1 norm_vec1)
 
-lemma linear_vmul_dest_vec1: 
+lemma linear_vmul_dest_vec1:
   fixes f:: "'a::semiring_1^'n \<Rightarrow> 'a^1"
   shows "linear f \<Longrightarrow> linear (\<lambda>x. dest_vec1(f x) *s v)"
   unfolding dest_vec1_def
@@ -2563,10 +2563,10 @@
   have th_0: "1 \<le> ?n +1" by simp
   have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
     by (simp add: pastecart_fst_snd)
-  have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)" 
+  have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
     by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square del: One_nat_def)
   then show ?thesis
-    unfolding th0 
+    unfolding th0
     unfolding real_vector_norm_def real_sqrt_le_iff id_def
     by (simp add: dot_def dimindex_finite_sum Cart_lambda_beta)
 qed
@@ -2592,13 +2592,13 @@
     using dimindex_nonzero[of "UNIV :: 'n set"] dimindex_nonzero[of "UNIV :: 'm set"]
     apply (simp add: Ball_def atLeastAtMost_iff inj_on_def dimindex_finite_sum del: One_nat_def)
     by arith
-  have fS: "?f ` ?S = ?M" 
+  have fS: "?f ` ?S = ?M"
     apply (rule set_ext)
     apply (simp add: image_iff Bex_def) using dimindex_nonzero[of "UNIV :: 'n set"] dimindex_nonzero[of "UNIV :: 'm set"] by arith
-  have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)" 
-    by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square setsum_reindex[OF finj, unfolded fS] del: One_nat_def)    
+  have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
+    by (simp add: dot_def setsum_add_split[OF th_0, of _ ?m] pastecart_def dimindex_finite_sum Cart_lambda_beta setsum_nonneg zero_le_square setsum_reindex[OF finj, unfolded fS] del: One_nat_def)
   then show ?thesis
-    unfolding th0 
+    unfolding th0
     unfolding real_vector_norm_def real_sqrt_le_iff id_def
     by (simp add: dot_def dimindex_finite_sum Cart_lambda_beta)
 qed
@@ -2644,14 +2644,14 @@
     done
   let ?r = "\<lambda>n. n - ?n"
   have rinj: "inj_on ?r ?S" apply (simp add: inj_on_def Ball_def thnm) by arith
-  have rS: "?r ` ?S = ?M" apply (rule set_ext) 
+  have rS: "?r ` ?S = ?M" apply (rule set_ext)
     apply (simp add: thnm image_iff Bex_def) by arith
   have "pastecart x1 x2 \<bullet> (pastecart y1 y2) = setsum ?g ?NM" by (simp add: dot_def)
   also have "\<dots> = setsum ?g ?N + setsum ?g ?S"
     by (simp add: dot_def thnm setsum_add_split[OF th_0, of _ ?m] del: One_nat_def)
   also have "\<dots> = setsum (?f x1 y1) ?N + setsum (?f x2 y2) ?M"
     unfolding setsum_reindex[OF rinj, unfolded rS o_def] th2 th3 ..
-  finally 
+  finally
   show ?thesis by (simp add: dot_def)
 qed
 
@@ -2679,7 +2679,7 @@
 unfolding hull_def subset_iff by auto
 
 lemma hull_eq: "(\<And>T. T \<subseteq> S ==> Inter T \<in> S) ==> (S hull s) = s \<longleftrightarrow> s \<in> S"
-using hull_same[of s S] hull_in[of S s] by metis  
+using hull_same[of s S] hull_in[of S s] by metis
 
 
 lemma hull_hull: "S hull (S hull s) = S hull s"
@@ -2749,12 +2749,12 @@
 lemma real_pow_lbound: "0 <= x ==> 1 + real n * x <= (1 + x) ^ n"
 proof(induct n)
   case 0 thus ?case by simp
-next 
+next
   case (Suc n)
   hence h: "1 + real n * x \<le> (1 + x) ^ n" by simp
   from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp
   from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp
-  also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric]) 
+  also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])
     apply (simp add: ring_simps)
     using mult_left_mono[OF p Suc.prems] by simp
   finally show ?case  by (simp add: real_of_nat_Suc ring_simps)
@@ -2763,13 +2763,13 @@
 lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"
 proof-
   from x have x0: "x - 1 > 0" by arith
-  from real_arch[OF x0, rule_format, of y] 
+  from real_arch[OF x0, rule_format, of y]
   obtain n::nat where n:"y < real n * (x - 1)" by metis
   from x0 have x00: "x- 1 \<ge> 0" by arith
-  from real_pow_lbound[OF x00, of n] n 
+  from real_pow_lbound[OF x00, of n] n
   have "y < x^n" by auto
   then show ?thesis by metis
-qed 
+qed
 
 lemma real_arch_pow2: "\<exists>n. (x::real) < 2^ n"
   using real_arch_pow[of 2 x] by simp
@@ -2777,13 +2777,13 @@
 lemma real_arch_pow_inv: assumes y: "(y::real) > 0" and x1: "x < 1"
   shows "\<exists>n. x^n < y"
 proof-
-  {assume x0: "x > 0" 
+  {assume x0: "x > 0"
     from x0 x1 have ix: "1 < 1/x" by (simp add: field_simps)
     from real_arch_pow[OF ix, of "1/y"]
     obtain n where n: "1/y < (1/x)^n" by blast
-    then 
+    then
     have ?thesis using y x0 by (auto simp add: field_simps power_divide) }
-  moreover 
+  moreover
   {assume "\<not> x > 0" with y x1 have ?thesis apply auto by (rule exI[where x=1], auto)}
   ultimately show ?thesis by metis
 qed
@@ -2821,18 +2821,18 @@
   have "max x y \<le> rsup {x,y}" using rsup_finite_ge_iff[OF f, of "max x y"]
     by (simp add: linorder_linear)
   ultimately show ?thesis by arith
-qed 
+qed
 
 lemma real_min_rinf: "min x y = rinf {x,y}"
 proof-
   have f: "finite {x, y}" "{x,y} \<noteq> {}"  by simp_all
-  from rinf_finite_le_iff[OF f, of "min x y"] have "rinf {x,y} \<le> min x y" 
+  from rinf_finite_le_iff[OF f, of "min x y"] have "rinf {x,y} \<le> min x y"
     by (simp add: linorder_linear)
   moreover
   have "min x y \<le> rinf {x,y}" using rinf_finite_ge_iff[OF f, of "min x y"]
     by simp
   ultimately show ?thesis by arith
-qed 
+qed
 
 (* ------------------------------------------------------------------------- *)
 (* Geometric progression.                                                    *)
@@ -2863,9 +2863,9 @@
   from mn have mn': "n - m \<ge> 0" by arith
   let ?f = "op + m"
   have i: "inj_on ?f ?S" unfolding inj_on_def by auto
-  have f: "?f ` ?S = {m..n}" 
+  have f: "?f ` ?S = {m..n}"
     using mn apply (auto simp add: image_iff Bex_def) by arith
-  have th: "op ^ x o op + m = (\<lambda>i. x^m * x^i)" 
+  have th: "op ^ x o op + m = (\<lambda>i. x^m * x^i)"
     by (rule ext, simp add: power_add power_mult)
   from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
   have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
@@ -2873,8 +2873,8 @@
     by (simp add: ring_simps power_add[symmetric])
 qed
 
-lemma sum_gp: "setsum (op ^ (x::'a::{field, recpower})) {m .. n} = 
-   (if n < m then 0 else if x = 1 then of_nat ((n + 1) - m) 
+lemma sum_gp: "setsum (op ^ (x::'a::{field, recpower})) {m .. n} =
+   (if n < m then 0 else if x = 1 then of_nat ((n + 1) - m)
                     else (x^ m - x^ (Suc n)) / (1 - x))"
 proof-
   {assume nm: "n < m" hence ?thesis by simp}
@@ -2889,7 +2889,7 @@
   ultimately show ?thesis by metis
 qed
 
-lemma sum_gp_offset: "setsum (op ^ (x::'a::{field,recpower})) {m .. m+n} = 
+lemma sum_gp_offset: "setsum (op ^ (x::'a::{field,recpower})) {m .. m+n} =
   (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
   unfolding sum_gp[of x m "m + n"] power_Suc
   by (simp add: ring_simps power_add)
@@ -2908,7 +2908,7 @@
 
 lemma subspace_0: "subspace S ==> 0 \<in> S" by (metis subspace_def)
 
-lemma subspace_add: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S ==> x + y \<in> S" 
+lemma subspace_add: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S ==> x + y \<in> S"
   by (metis subspace_def)
 
 lemma subspace_mul: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> c *s x \<in> S"
@@ -2926,10 +2926,10 @@
   shows "setsum f B \<in> A"
   using  fB f sA
   apply(induct rule: finite_induct[OF fB])
-  by (simp add: subspace_def sA, auto simp add: sA subspace_add) 
-
-lemma subspace_linear_image: 
-  assumes lf: "linear (f::'a::semiring_1^'n \<Rightarrow> _)" and sS: "subspace S" 
+  by (simp add: subspace_def sA, auto simp add: sA subspace_add)
+
+lemma subspace_linear_image:
+  assumes lf: "linear (f::'a::semiring_1^'n \<Rightarrow> _)" and sS: "subspace S"
   shows "subspace(f ` S)"
   using lf sS linear_0[OF lf]
   unfolding linear_def subspace_def
@@ -2986,7 +2986,7 @@
   from SP have SP': "S \<subseteq> P" by (simp add: mem_def subset_eq)
   from P have P': "P \<in> subspace" by (simp add: mem_def)
   from x hull_minimal[OF SP' P', unfolded span_def[symmetric]]
-  show "P x" by (metis mem_def subset_eq) 
+  show "P x" by (metis mem_def subset_eq)
 qed
 
 lemma span_empty: "span {} = {(0::'a::semiring_0 ^ 'n)}"
@@ -3016,11 +3016,11 @@
   using span_induct SP P by blast
 
 inductive span_induct_alt_help for S:: "'a::semiring_1^'n \<Rightarrow> bool"
-  where 
+  where
   span_induct_alt_help_0: "span_induct_alt_help S 0"
   | span_induct_alt_help_S: "x \<in> S \<Longrightarrow> span_induct_alt_help S z \<Longrightarrow> span_induct_alt_help S (c *s x + z)"
 
-lemma span_induct_alt': 
+lemma span_induct_alt':
   assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" shows "\<forall>x \<in> span S. h x"
 proof-
   {fix x:: "'a^'n" assume x: "span_induct_alt_help S x"
@@ -3031,7 +3031,7 @@
       done}
   note th0 = this
   {fix x assume x: "x \<in> span S"
-    
+
     have "span_induct_alt_help S x"
       proof(rule span_induct[where x=x and S=S])
 	show "x \<in> span S" using x .
@@ -3043,7 +3043,7 @@
 	have "span_induct_alt_help S 0" by (rule span_induct_alt_help_0)
 	moreover
 	{fix x y assume h: "span_induct_alt_help S x" "span_induct_alt_help S y"
-	  from h 
+	  from h
 	  have "span_induct_alt_help S (x + y)"
 	    apply (induct rule: span_induct_alt_help.induct)
 	    apply simp
@@ -3054,7 +3054,7 @@
 	    done}
 	moreover
 	{fix c x assume xt: "span_induct_alt_help S x"
-	  then have "span_induct_alt_help S (c*s x)" 
+	  then have "span_induct_alt_help S (c*s x)"
 	    apply (induct rule: span_induct_alt_help.induct)
 	    apply (simp add: span_induct_alt_help_0)
 	    apply (simp add: vector_smult_assoc vector_add_ldistrib)
@@ -3063,13 +3063,13 @@
 	    apply simp
 	    done
 	}
-	ultimately show "subspace (span_induct_alt_help S)" 
+	ultimately show "subspace (span_induct_alt_help S)"
 	  unfolding subspace_def mem_def Ball_def by blast
       qed}
   with th0 show ?thesis by blast
-qed 
-
-lemma span_induct_alt: 
+qed
+
+lemma span_induct_alt:
   assumes h0: "h (0::'a::semiring_1^'n)" and hS: "\<And>c x y. x \<in> S \<Longrightarrow> h y \<Longrightarrow> h (c*s x + y)" and x: "x \<in> span S"
   shows "h x"
 using span_induct_alt'[of h S] h0 hS x by blast
@@ -3118,9 +3118,9 @@
       apply (rule subspace_span)
       apply (rule x)
       done}
-  moreover 
+  moreover
   {fix x assume x: "x \<in> span S"
-    have th0:"(\<lambda>a. f a \<in> span (f ` S)) = {x. f x \<in> span (f ` S)}" apply (rule set_ext) 
+    have th0:"(\<lambda>a. f a \<in> span (f ` S)) = {x. f x \<in> span (f ` S)}" apply (rule set_ext)
       unfolding mem_def Collect_def ..
     have "f x \<in> span (f ` S)"
       apply (rule span_induct[where S=S])
@@ -3146,15 +3146,15 @@
 	apply (rule exI[where x="1"], simp)
 	by (rule span_0)}
     moreover
-    {assume ab: "x \<noteq> b" 
+    {assume ab: "x \<noteq> b"
       then have "?P x"  using xS
 	apply -
 	apply (rule exI[where x=0])
 	apply (rule span_superset)
 	by simp}
     ultimately have "?P x" by blast}
-  moreover have "subspace ?P" 
-    unfolding subspace_def 
+  moreover have "subspace ?P"
+    unfolding subspace_def
     apply auto
     apply (simp add: mem_def)
     apply (rule exI[where x=0])
@@ -3174,7 +3174,7 @@
     apply (rule span_mul[unfolded mem_def])
     apply assumption
     by (vector ring_simps)
-  ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis 
+  ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis
 qed
 
 lemma span_breakdown_eq:
@@ -3186,7 +3186,7 @@
       apply (rule_tac x= "k" in exI)
       apply (rule set_rev_mp[of _ "span (S - {a})" _])
       apply assumption
-      apply (rule span_mono)      
+      apply (rule span_mono)
       apply blast
       done}
   moreover
@@ -3196,7 +3196,7 @@
       apply (rule span_add)
       apply (rule set_rev_mp[of _ "span S" _])
       apply (rule k)
-      apply (rule span_mono)      
+      apply (rule span_mono)
       apply blast
       apply (rule span_mul)
       apply (rule span_superset)
@@ -3224,7 +3224,7 @@
       done
     with na  have ?thesis by blast}
   moreover
-  {assume k0: "k \<noteq> 0" 
+  {assume k0: "k \<noteq> 0"
     have eq: "b = (1/k) *s a - ((1/k) *s a - b)" by vector
     from k0 have eq': "(1/k) *s (a - k*s b) = (1/k) *s a - b"
       by (vector field_simps)
@@ -3247,8 +3247,8 @@
   ultimately show ?thesis by blast
 qed
 
-lemma in_span_delete: 
-  assumes a: "(a::'a::field^'n) \<in> span S" 
+lemma in_span_delete:
+  assumes a: "(a::'a::field^'n) \<in> span S"
   and na: "a \<notin> span (S-{b})"
   shows "b \<in> span (insert a (S - {b}))"
   apply (rule in_span_insert)
@@ -3268,7 +3268,7 @@
   from span_breakdown[of x "insert x S" y, OF insertI1 y]
   obtain k where k: "y -k*s x \<in> span (S - {x})" by auto
   have eq: "y = (y - k *s x) + k *s x" by vector
-  show ?thesis 
+  show ?thesis
     apply (subst eq)
     apply (rule span_add)
     apply (rule set_rev_mp)
@@ -3304,18 +3304,18 @@
   next
     fix c x y
     assume x: "x \<in> P" and hy: "?h y"
-    from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P" 
+    from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P"
       and u: "setsum (\<lambda>v. u v *s v) S = y" by blast
     let ?S = "insert x S"
     let ?u = "\<lambda>y. if y = x then (if x \<in> S then u y + c else c)
                   else u y"
     from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P" by blast+
     {assume xS: "x \<in> S"
-      have S1: "S = (S - {x}) \<union> {x}" 
+      have S1: "S = (S - {x}) \<union> {x}"
 	and Sss:"finite (S - {x})" "finite {x}" "(S -{x}) \<inter> {x} = {}" using xS fS by auto
       have "setsum (\<lambda>v. ?u v *s v) ?S =(\<Sum>v\<in>S - {x}. u v *s v) + (u x + c) *s x"
-	using xS 
-	by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]] 
+	using xS
+	by (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]]
 	  setsum_clauses(2)[OF fS] cong del: if_weak_cong)
       also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
 	apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
@@ -3324,7 +3324,7 @@
 	by (simp add: add_commute u)
       finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
     then have "?Q ?S ?u (c*s x + y)" using th0 by blast}
-  moreover 
+  moreover
   {assume xS: "x \<notin> S"
     have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *s v) = y"
       unfolding u[symmetric]
@@ -3334,7 +3334,7 @@
       by (simp add: th00 setsum_clauses add_commute cong del: if_weak_cong)}
   ultimately have "?Q ?S ?u (c*s x + y)"
     by (cases "x \<in> S", simp, simp)
-    then show "?h (c*s x + y)" 
+    then show "?h (c*s x + y)"
       apply -
       apply (rule exI[where x="?S"])
       apply (rule exI[where x="?u"]) by metis
@@ -3346,11 +3346,11 @@
   "dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>(v::'a::{idom,field}^'n) \<in>S. u v \<noteq> 0 \<and> setsum (\<lambda>v. u v *s v) S = 0))" (is "?lhs = ?rhs")
 proof-
   {assume dP: "dependent P"
-    then obtain a S u where aP: "a \<in> P" and fS: "finite S" 
-      and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *s v) S = a" 
+    then obtain a S u where aP: "a \<in> P" and fS: "finite S"
+      and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *s v) S = a"
       unfolding dependent_def span_explicit by blast
-    let ?S = "insert a S" 
-    let ?u = "\<lambda>y. if y = a then - 1 else u y" 
+    let ?S = "insert a S"
+    let ?u = "\<lambda>y. if y = a then - 1 else u y"
     let ?v = a
     from aP SP have aS: "a \<notin> S" by blast
     from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto
@@ -3366,16 +3366,16 @@
       apply (rule exI[where x= "?u"])
       by clarsimp}
   moreover
-  {fix S u v assume fS: "finite S" 
-      and SP: "S \<subseteq> P" and vS: "v \<in> S" and uv: "u v \<noteq> 0" 
+  {fix S u v assume fS: "finite S"
+      and SP: "S \<subseteq> P" and vS: "v \<in> S" and uv: "u v \<noteq> 0"
     and u: "setsum (\<lambda>v. u v *s v) S = 0"
-    let ?a = v 
+    let ?a = v
     let ?S = "S - {v}"
     let ?u = "\<lambda>i. (- u i) / u v"
-    have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"       using fS SP vS by auto 
+    have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"       using fS SP vS by auto
     have "setsum (\<lambda>v. ?u v *s v) ?S = setsum (\<lambda>v. (- (inverse (u ?a))) *s (u v *s v)) S - ?u v *s v"
-      using fS vS uv 
-      by (simp add: setsum_diff1 vector_smult_lneg divide_inverse 
+      using fS vS uv
+      by (simp add: setsum_diff1 vector_smult_lneg divide_inverse
 	vector_smult_assoc field_simps)
     also have "\<dots> = ?a"
       unfolding setsum_cmul u
@@ -3398,7 +3398,7 @@
   (is "_ = ?rhs")
 proof-
   {fix y assume y: "y \<in> span S"
-    from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and 
+    from y obtain S' u where fS': "finite S'" and SS': "S' \<subseteq> S" and
       u: "setsum (\<lambda>v. u v *s v) S' = y" unfolding span_explicit by blast
     let ?u = "\<lambda>x. if x \<in> S' then u x else 0"
     from setsum_restrict_set[OF fS, of "\<lambda>v. u v *s v" S', symmetric] SS'
@@ -3410,7 +3410,7 @@
       done
     hence "setsum (\<lambda>v. ?u v *s v) S = y" by (metis u)
     hence "y \<in> ?rhs" by auto}
-  moreover 
+  moreover
   {fix y u assume u: "setsum (\<lambda>v. u v *s v) S = y"
     then have "y \<in> span S" using fS unfolding span_explicit by auto}
   ultimately show ?thesis by blast
@@ -3431,7 +3431,7 @@
 apply (auto simp add: Collect_def mem_def)
 done
 
-  
+
 lemma has_size_stdbasis: "{basis i ::real ^'n | i. i \<in> {1 .. dimindex (UNIV :: 'n set)}} hassize (dimindex(UNIV :: 'n set))" (is "?S hassize ?n")
 proof-
   have eq: "?S = basis ` {1 .. ?n}" by blast
@@ -3461,10 +3461,10 @@
  {fix x::"'a^'n" assume xS: "x\<in> ?B"
    from xS have "?P x" by (auto simp add: basis_component)}
  moreover
- have "subspace ?P" 
+ have "subspace ?P"
    by (auto simp add: subspace_def Collect_def mem_def zero_index vector_component)
  ultimately show ?thesis
-   using x span_induct[of ?B ?P x] i iS by blast 
+   using x span_induct[of ?B ?P x] i iS by blast
 qed
 
 lemma independent_stdbasis: "independent {basis i ::real^'n |i. i\<in> {1 .. dimindex(UNIV :: 'n set)}}"
@@ -3508,7 +3508,7 @@
 	apply assumption
 	apply blast
 	by (simp add: dependent_def)}
-    moreover 
+    moreover
     {assume i: ?rhs
       have ?lhs using i aS
 	apply simp
@@ -3541,7 +3541,7 @@
   by (metis subset_eq span_superset)
 
 lemma spanning_subset_independent:
-  assumes BA: "B \<subseteq> A" and iA: "independent (A::('a::field ^'n) set)" 
+  assumes BA: "B \<subseteq> A" and iA: "independent (A::('a::field ^'n) set)"
   and AsB: "A \<subseteq> span B"
   shows "A = B"
 proof
@@ -3569,7 +3569,7 @@
 
 lemma exchange_lemma:
   assumes f:"finite (t:: ('a::field^'n) set)" and i: "independent s"
-  and sp:"s \<subseteq> span t" 
+  and sp:"s \<subseteq> span t"
   shows "\<exists>t'. (t' hassize card t) \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
 using f i sp
 proof(induct c\<equiv>"card(t - s)" arbitrary: s t rule: nat_less_induct)
@@ -3584,15 +3584,15 @@
     and ft: "finite t" and s: "independent s" and sp: "s \<subseteq> span t"
     and n: "n = card (t - s)"
   let ?P = "\<lambda>t'. (t' hassize card t) \<and> s \<subseteq> t' \<and> t' \<subseteq> s \<union> t \<and> s \<subseteq> span t'"
-  let ?ths = "\<exists>t'. ?P t'" 
-  {assume st: "s \<subseteq> t" 
-    from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t]) 
+  let ?ths = "\<exists>t'. ?P t'"
+  {assume st: "s \<subseteq> t"
+    from st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])
       by (auto simp add: hassize_def intro: span_superset)}
   moreover
   {assume st: "t \<subseteq> s"
-    
-    from spanning_subset_independent[OF st s sp] 
-      st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t]) 
+
+    from spanning_subset_independent[OF st s sp]
+      st ft span_mono[OF st] have ?ths apply - apply (rule exI[where x=t])
       by (auto simp add: hassize_def intro: span_superset)}
   moreover
   {assume st: "\<not> s \<subseteq> t" "\<not> t \<subseteq> s"
@@ -3603,28 +3603,28 @@
       from b ft have ct0: "card t \<noteq> 0" by auto
     {assume stb: "s \<subseteq> span(t -{b})"
       from ft have ftb: "finite (t -{b})" by auto
-      from H[rule_format, OF cardlt ftb s stb] 
+      from H[rule_format, OF cardlt ftb s stb]
       obtain u where u: "u hassize card (t-{b})" "s \<subseteq> u" "u \<subseteq> s \<union> (t - {b})" "s \<subseteq> span u" by blast
       let ?w = "insert b u"
       have th0: "s \<subseteq> insert b u" using u by blast
-      from u(3) b have "u \<subseteq> s \<union> t" by blast 
+      from u(3) b have "u \<subseteq> s \<union> t" by blast
       then have th1: "insert b u \<subseteq> s \<union> t" using u b by blast
       have bu: "b \<notin> u" using b u by blast
       from u(1) have fu: "finite u" by (simp add: hassize_def)
       from u(1) ft b have "u hassize (card t - 1)" by auto
-      then 
-      have th2: "insert b u hassize card t" 
+      then
+      have th2: "insert b u hassize card t"
 	using  card_insert_disjoint[OF fu bu] ct0 by (auto simp add: hassize_def)
       from u(4) have "s \<subseteq> span u" .
       also have "\<dots> \<subseteq> span (insert b u)" apply (rule span_mono) by blast
       finally have th3: "s \<subseteq> span (insert b u)" .      from th0 th1 th2 th3 have th: "?P ?w"  by blast
       from th have ?ths by blast}
     moreover
-    {assume stb: "\<not> s \<subseteq> span(t -{b})" 
+    {assume stb: "\<not> s \<subseteq> span(t -{b})"
       from stb obtain a where a: "a \<in> s" "a \<notin> span (t - {b})" by blast
       have ab: "a \<noteq> b" using a b by blast
       have at: "a \<notin> t" using a ab span_superset[of a "t- {b}"] by auto
-      have mlt: "card ((insert a (t - {b})) - s) < n" 
+      have mlt: "card ((insert a (t - {b})) - s) < n"
 	using cardlt ft n  a b by auto
       have ft': "finite (insert a (t - {b}))" using ft by auto
       {fix x assume xs: "x \<in> s"
@@ -3637,15 +3637,15 @@
 	have x: "x \<in> span (insert b (insert a (t - {b})))" ..
 	from span_trans[OF bs x] have "x \<in> span (insert a (t - {b}))"  .}
       then have sp': "s \<subseteq> span (insert a (t - {b}))" by blast
-      
-      from H[rule_format, OF mlt ft' s sp' refl] obtain u where 
+
+      from H[rule_format, OF mlt ft' s sp' refl] obtain u where
 	u: "u hassize card (insert a (t -{b}))" "s \<subseteq> u" "u \<subseteq> s \<union> insert a (t -{b})"
 	"s \<subseteq> span u" by blast
       from u a b ft at ct0 have "?P u" by (auto simp add: hassize_def)
       then have ?ths by blast }
     ultimately have ?ths by blast
   }
-  ultimately 
+  ultimately
   show ?ths  by blast
 qed
 
@@ -3659,7 +3659,7 @@
 lemma finite_Atleast_Atmost[simp]: "finite {f x |x. x\<in> {(i::'a::finite_intvl_succ) .. j}}"
 proof-
   have eq: "{f x |x. x\<in> {i .. j}} = f ` {i .. j}" by auto
-  show ?thesis unfolding eq 
+  show ?thesis unfolding eq
     apply (rule finite_imageI)
     apply (rule finite_intvl)
     done
@@ -3668,7 +3668,7 @@
 lemma finite_Atleast_Atmost_nat[simp]: "finite {f x |x. x\<in> {(i::nat) .. j}}"
 proof-
   have eq: "{f x |x. x\<in> {i .. j}} = f ` {i .. j}" by auto
-  show ?thesis unfolding eq 
+  show ?thesis unfolding eq
     apply (rule finite_imageI)
     apply (rule finite_atLeastAtMost)
     done
@@ -3682,7 +3682,7 @@
   apply (rule independent_span_bound)
   apply (rule finite_Atleast_Atmost_nat)
   apply assumption
-  unfolding span_stdbasis 
+  unfolding span_stdbasis
   apply (rule subset_UNIV)
   done
 
@@ -3710,14 +3710,14 @@
     from VS obtain a where a: "a \<in> V" "a \<notin> span S" by blast
     from a have aS: "a \<notin> S" by (auto simp add: span_superset)
     have th0: "insert a S \<subseteq> V" using a sv by blast
-    from independent_insert[of a S]  i a 
+    from independent_insert[of a S]  i a
     have th1: "independent (insert a S)" by auto
-    have mlt: "?d - card (insert a S) < n" 
-      using aS a n independent_bound[OF th1] dimindex_ge_1[of "UNIV :: 'n set"] 
-      by auto 
-      
-    from H[rule_format, OF mlt th0 th1 refl] 
-    obtain B where B: "insert a S \<subseteq> B" "B \<subseteq> V" "independent B" " V \<subseteq> span B" 
+    have mlt: "?d - card (insert a S) < n"
+      using aS a n independent_bound[OF th1] dimindex_ge_1[of "UNIV :: 'n set"]
+      by auto
+
+    from H[rule_format, OF mlt th0 th1 refl]
+    obtain B where B: "insert a S \<subseteq> B" "B \<subseteq> V" "independent B" " V \<subseteq> span B"
       by blast
     from B have "?P B" by auto
     then have ?ths by blast}
@@ -3732,7 +3732,7 @@
 
 definition "dim V = (SOME n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n))"
 
-lemma basis_exists:  "\<exists>B. (B :: (real ^'n) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize dim V)" 
+lemma basis_exists:  "\<exists>B. (B :: (real ^'n) set) \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize dim V)"
 unfolding dim_def some_eq_ex[of "\<lambda>n. \<exists>B. B \<subseteq> V \<and> independent B \<and> V \<subseteq> span B \<and> (B hassize n)"]
 unfolding hassize_def
 using maximal_independent_subset[of V] independent_bound
@@ -3784,7 +3784,7 @@
 qed
 
 lemma card_le_dim_spanning:
-  assumes BV: "(B:: (real ^'n) set) \<subseteq> V" and VB: "V \<subseteq> span B" 
+  assumes BV: "(B:: (real ^'n) set) \<subseteq> V" and VB: "V \<subseteq> span B"
   and fB: "finite B" and dVB: "dim V \<ge> card B"
   shows "independent B"
 proof-
@@ -3794,10 +3794,10 @@
     from BV a have th0: "B -{a} \<subseteq> V" by blast
     {fix x assume x: "x \<in> V"
       from a have eq: "insert a (B -{a}) = B" by blast
-      from x VB have x': "x \<in> span B" by blast 
+      from x VB have x': "x \<in> span B" by blast
       from span_trans[OF a(2), unfolded eq, OF x']
       have "x \<in> span (B -{a})" . }
-    then have th1: "V \<subseteq> span (B -{a})" by blast 
+    then have th1: "V \<subseteq> span (B -{a})" by blast
     have th2: "finite (B -{a})" using fB by auto
     from span_card_ge_dim[OF th0 th1 th2]
     have c: "dim V \<le> card (B -{a})" .
@@ -3806,7 +3806,7 @@
 qed
 
 lemma card_eq_dim: "(B:: (real ^'n) set) \<subseteq> V \<Longrightarrow> B hassize dim V \<Longrightarrow> independent B \<longleftrightarrow> V \<subseteq> span B"
-  by (metis hassize_def order_eq_iff card_le_dim_spanning 
+  by (metis hassize_def order_eq_iff card_le_dim_spanning
     card_ge_dim_independent)
 
 (* ------------------------------------------------------------------------- *)
@@ -3818,18 +3818,18 @@
   by (metis independent_card_le_dim independent_bound subset_refl)
 
 lemma dependent_biggerset_general: "(finite (S:: (real^'n) set) \<Longrightarrow> card S > dim S) \<Longrightarrow> dependent S"
-  using independent_bound_general[of S] by (metis linorder_not_le) 
+  using independent_bound_general[of S] by (metis linorder_not_le)
 
 lemma dim_span: "dim (span (S:: (real ^'n) set)) = dim S"
 proof-
-  have th0: "dim S \<le> dim (span S)" 
+  have th0: "dim S \<le> dim (span S)"
     by (auto simp add: subset_eq intro: dim_subset span_superset)
-  from basis_exists[of S] 
+  from basis_exists[of S]
   obtain B where B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
   from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+
-  have bSS: "B \<subseteq> span S" using B(1) by (metis subset_eq span_inc) 
-  have sssB: "span S \<subseteq> span B" using span_mono[OF B(3)] by (simp add: span_span) 
-  from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis 
+  have bSS: "B \<subseteq> span S" using B(1) by (metis subset_eq span_inc)
+  have sssB: "span S \<subseteq> span B" using span_mono[OF B(3)] by (simp add: span_span)
+  from span_card_ge_dim[OF bSS sssB fB(1)] th0 show ?thesis
     using fB(2)  by arith
 qed
 
@@ -3847,7 +3847,7 @@
 
 lemma dim_image_le: assumes lf: "linear f" shows "dim (f ` S) \<le> dim (S:: (real ^'n) set)"
 proof-
-  from basis_exists[of S] obtain B where 
+  from basis_exists[of S] obtain B where
     B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
   from B have fB: "finite B" "card B = dim S" unfolding hassize_def by blast+
   have "dim (f ` S) \<le> card (f ` B)"
@@ -3860,7 +3860,7 @@
 (* Relation between bases and injectivity/surjectivity of map.               *)
 
 lemma spanning_surjective_image:
-  assumes us: "UNIV \<subseteq> span (S:: ('a::semiring_1 ^'n) set)" 
+  assumes us: "UNIV \<subseteq> span (S:: ('a::semiring_1 ^'n) set)"
   and lf: "linear f" and sf: "surj f"
   shows "UNIV \<subseteq> span (f ` S)"
 proof-
@@ -3881,7 +3881,7 @@
     hence "a \<in> span (S -{a})" using fi by (auto simp add: inj_on_def)
     with a(1) iS  have False by (simp add: dependent_def) }
   then show ?thesis unfolding dependent_def by blast
-qed 
+qed
 
 (* ------------------------------------------------------------------------- *)
 (* Picking an orthogonal replacement for a spanning set.                     *)
@@ -3904,15 +3904,15 @@
   case 1 thus ?case apply (rule exI[where x="{}"]) by (auto simp add: pairwise_def)
 next
   case (2 a B)
-  note fB = `finite B` and aB = `a \<notin> B` 
-  from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C` 
-  obtain C where C: "finite C" "card C \<le> card B" 
+  note fB = `finite B` and aB = `a \<notin> B`
+  from `\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C`
+  obtain C where C: "finite C" "card C \<le> card B"
     "span C = span B" "pairwise orthogonal C" by blast
   let ?a = "a - setsum (\<lambda>x. (x\<bullet>a / (x\<bullet>x)) *s x) C"
   let ?C = "insert ?a C"
   from C(1) have fC: "finite ?C" by simp
   from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)
-  {fix x k 
+  {fix x k
     have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)
     have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"
       apply (simp only: vector_ssub_ldistrib th0)
@@ -3924,18 +3924,18 @@
       by (rule span_superset)}
   then have SC: "span ?C = span (insert a B)"
     unfolding expand_set_eq span_breakdown_eq C(3)[symmetric] by auto
-  thm pairwise_def 
+  thm pairwise_def
   {fix x y assume xC: "x \<in> ?C" and yC: "y \<in> ?C" and xy: "x \<noteq> y"
-    {assume xa: "x = ?a" and ya: "y = ?a" 
+    {assume xa: "x = ?a" and ya: "y = ?a"
       have "orthogonal x y" using xa ya xy by blast}
     moreover
-    {assume xa: "x = ?a" and ya: "y \<noteq> ?a" "y \<in> C" 
+    {assume xa: "x = ?a" and ya: "y \<noteq> ?a" "y \<in> C"
       from ya have Cy: "C = insert y (C - {y})" by blast
       have fth: "finite (C - {y})" using C by simp
       have "orthogonal x y"
 	using xa ya
 	unfolding orthogonal_def xa dot_lsub dot_rsub diff_eq_0_iff_eq
-	apply simp 
+	apply simp
 	apply (subst Cy)
 	using C(1) fth
 	apply (simp only: setsum_clauses)
@@ -3946,13 +3946,13 @@
 	apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
 	by auto}
     moreover
-    {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a" 
+    {assume xa: "x \<noteq> ?a" "x \<in> C" and ya: "y = ?a"
       from xa have Cx: "C = insert x (C - {x})" by blast
       have fth: "finite (C - {x})" using C by simp
       have "orthogonal x y"
 	using xa ya
 	unfolding orthogonal_def ya dot_rsub dot_lsub diff_eq_0_iff_eq
-	apply simp 
+	apply simp
 	apply (subst Cx)
 	using C(1) fth
 	apply (simp only: setsum_clauses)
@@ -3963,12 +3963,12 @@
 	apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
 	by auto}
     moreover
-    {assume xa: "x \<in> C" and ya: "y \<in> C" 
+    {assume xa: "x \<in> C" and ya: "y \<in> C"
       have "orthogonal x y" using xa ya xy C(4) unfolding pairwise_def by blast}
     ultimately have "orthogonal x y" using xC yC by blast}
   then have CPO: "pairwise orthogonal ?C" unfolding pairwise_def by blast
   from fC cC SC CPO have "?P (insert a B) ?C" by blast
-  then show ?case by blast 
+  then show ?case by blast
 qed
 
 lemma orthogonal_basis_exists:
@@ -3977,18 +3977,18 @@
 proof-
   from basis_exists[of V] obtain B where B: "B \<subseteq> V" "independent B" "V \<subseteq> span B" "B hassize dim V" by blast
   from B have fB: "finite B" "card B = dim V" by (simp_all add: hassize_def)
-  from basis_orthogonal[OF fB(1)] obtain C where 
+  from basis_orthogonal[OF fB(1)] obtain C where
     C: "finite C" "card C \<le> card B" "span C = span B" "pairwise orthogonal C" by blast
-  from C B 
-  have CSV: "C \<subseteq> span V" by (metis span_inc span_mono subset_trans) 
+  from C B
+  have CSV: "C \<subseteq> span V" by (metis span_inc span_mono subset_trans)
   from span_mono[OF B(3)]  C have SVC: "span V \<subseteq> span C" by (simp add: span_span)
   from card_le_dim_spanning[OF CSV SVC C(1)] C(2,3) fB
-  have iC: "independent C" by (simp add: dim_span) 
+  have iC: "independent C" by (simp add: dim_span)
   from C fB have "card C \<le> dim V" by simp
   moreover have "dim V \<le> card C" using span_card_ge_dim[OF CSV SVC C(1)]
     by (simp add: dim_span)
   ultimately have CdV: "C hassize dim V" unfolding hassize_def using C(1) by simp
-  from C B CSV CdV iC show ?thesis by auto 
+  from C B CSV CdV iC show ?thesis by auto
 qed
 
 lemma span_eq: "span S = span T \<longleftrightarrow> S \<subseteq> span T \<and> T \<subseteq> span S"
@@ -4003,8 +4003,8 @@
   shows "\<exists>(a:: real ^'n). a \<noteq>0 \<and> (\<forall>x \<in> span S. a \<bullet> x = 0)"
 proof-
   from sU obtain a where a: "a \<notin> span S" by blast
-  from orthogonal_basis_exists obtain B where 
-    B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "B hassize dim S" "pairwise orthogonal B" 
+  from orthogonal_basis_exists obtain B where
+    B: "independent B" "B \<subseteq> span S" "S \<subseteq> span B" "B hassize dim S" "pairwise orthogonal B"
     by blast
   from B have fB: "finite B" "card B = dim S" by (simp_all add: hassize_def)
   from span_mono[OF B(2)] span_mono[OF B(3)]
@@ -4020,12 +4020,12 @@
   have "\<forall>x\<in>span B. ?a \<bullet> x = 0"
   proof(rule span_induct')
     show "subspace (\<lambda>x. ?a \<bullet> x = 0)"
-      by (auto simp add: subspace_def mem_def dot_radd dot_rmult) 
+      by (auto simp add: subspace_def mem_def dot_radd dot_rmult)
   next
     {fix x assume x: "x \<in> B"
       from x have B': "B = insert x (B - {x})" by blast
       have fth: "finite (B - {x})" using fB by simp
-      have "?a \<bullet> x = 0" 
+      have "?a \<bullet> x = 0"
 	apply (subst B') using fB fth
 	unfolding setsum_clauses(2)[OF fth]
 	apply simp
@@ -4038,7 +4038,7 @@
   with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])
 qed
 
-lemma span_not_univ_subset_hyperplane: 
+lemma span_not_univ_subset_hyperplane:
   assumes SU: "span S \<noteq> (UNIV ::(real^'n) set)"
   shows "\<exists> a. a \<noteq>0 \<and> span S \<subseteq> {x. a \<bullet> x = 0}"
   using span_not_univ_orthogonal[OF SU] by auto
@@ -4058,9 +4058,9 @@
 (* We can extend a linear basis-basis injection to the whole set.            *)
 
 lemma linear_indep_image_lemma:
-  assumes lf: "linear f" and fB: "finite B" 
+  assumes lf: "linear f" and fB: "finite B"
   and ifB: "independent (f ` B)"
-  and fi: "inj_on f B" and xsB: "x \<in> span B" 
+  and fi: "inj_on f B" and xsB: "x \<in> span B"
   and fx: "f (x::'a::field^'n) = 0"
   shows "x = 0"
   using fB ifB fi xsB fx
@@ -4070,11 +4070,11 @@
   case (2 a b x)
   have fb: "finite b" using "2.prems" by simp
   have th0: "f ` b \<subseteq> f ` (insert a b)"
-    apply (rule image_mono) by blast 
+    apply (rule image_mono) by blast
   from independent_mono[ OF "2.prems"(2) th0]
   have ifb: "independent (f ` b)"  .
-  have fib: "inj_on f b" 
-    apply (rule subset_inj_on [OF "2.prems"(3)]) 
+  have fib: "inj_on f b"
+    apply (rule subset_inj_on [OF "2.prems"(3)])
     by blast
   from span_breakdown[of a "insert a b", simplified, OF "2.prems"(4)]
   obtain k where k: "x - k*s a \<in> span (b -{a})" by blast
@@ -4084,16 +4084,16 @@
     using k span_mono[of "b-{a}" b] by blast
   hence "f x - k*s f a \<in> span (f ` b)"
     by (simp add: linear_sub[OF lf] linear_cmul[OF lf])
-  hence th: "-k *s f a \<in> span (f ` b)" 
+  hence th: "-k *s f a \<in> span (f ` b)"
     using "2.prems"(5) by (simp add: vector_smult_lneg)
-  {assume k0: "k = 0" 
+  {assume k0: "k = 0"
     from k0 k have "x \<in> span (b -{a})" by simp
     then have "x \<in> span b" using span_mono[of "b-{a}" b]
       by blast}
   moreover
   {assume k0: "k \<noteq> 0"
     from span_mul[OF th, of "- 1/ k"] k0
-    have th1: "f a \<in> span (f ` b)" 
+    have th1: "f a \<in> span (f ` b)"
       by (auto simp add: vector_smult_assoc)
     from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric]
     have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast
@@ -4112,17 +4112,17 @@
 
 lemma linear_independent_extend_lemma:
   assumes fi: "finite B" and ib: "independent B"
-  shows "\<exists>g. (\<forall>x\<in> span B. \<forall>y\<in> span B. g ((x::'a::field^'n) + y) = g x + g y) 
+  shows "\<exists>g. (\<forall>x\<in> span B. \<forall>y\<in> span B. g ((x::'a::field^'n) + y) = g x + g y)
            \<and> (\<forall>x\<in> span B. \<forall>c. g (c*s x) = c *s g x)
            \<and> (\<forall>x\<in> B. g x = f x)"
 using ib fi
 proof(induct rule: finite_induct[OF fi])
-  case 1 thus ?case by (auto simp add: span_empty) 
+  case 1 thus ?case by (auto simp add: span_empty)
 next
   case (2 a b)
   from "2.prems" "2.hyps" have ibf: "independent b" "finite b"
     by (simp_all add: independent_insert)
-  from "2.hyps"(3)[OF ibf] obtain g where 
+  from "2.hyps"(3)[OF ibf] obtain g where
     g: "\<forall>x\<in>span b. \<forall>y\<in>span b. g (x + y) = g x + g y"
     "\<forall>x\<in>span b. \<forall>c. g (c *s x) = c *s g x" "\<forall>x\<in>b. g x = f x" by blast
   let ?h = "\<lambda>z. SOME k. (z - k *s a) \<in> span b"
@@ -4132,12 +4132,12 @@
       unfolding span_breakdown_eq[symmetric]
       using z .
     {fix k assume k: "z - k *s a \<in> span b"
-      have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a" 
+      have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
 	by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
       from span_sub[OF th0 k]
       have khz: "(k - ?h z) *s a \<in> span b" by (simp add: eq)
       {assume "k \<noteq> ?h z" hence k0: "k - ?h z \<noteq> 0" by simp
-	from k0 span_mul[OF khz, of "1 /(k - ?h z)"] 
+	from k0 span_mul[OF khz, of "1 /(k - ?h z)"]
 	have "a \<in> span b" by (simp add: vector_smult_assoc)
 	with "2.prems"(1) "2.hyps"(2) have False
 	  by (auto simp add: dependent_def)}
@@ -4146,26 +4146,26 @@
   note h = this
   let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"
   {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
-    have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)" 
+    have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"
       by (vector ring_simps)
     have addh: "?h (x + y) = ?h x + ?h y"
       apply (rule conjunct2[OF h, rule_format, symmetric])
       apply (rule span_add[OF x y])
       unfolding tha
       by (metis span_add x y conjunct1[OF h, rule_format])
-    have "?g (x + y) = ?g x + ?g y" 
+    have "?g (x + y) = ?g x + ?g y"
       unfolding addh tha
       g(1)[rule_format,OF conjunct1[OF h, OF x] conjunct1[OF h, OF y]]
       by (simp add: vector_sadd_rdistrib)}
   moreover
   {fix x:: "'a^'n" and c:: 'a  assume x: "x \<in> span (insert a b)"
-    have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)" 
+    have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"
       by (vector ring_simps)
-    have hc: "?h (c *s x) = c * ?h x" 
+    have hc: "?h (c *s x) = c * ?h x"
       apply (rule conjunct2[OF h, rule_format, symmetric])
       apply (metis span_mul x)
       by (metis tha span_mul x conjunct1[OF h])
-    have "?g (c *s x) = c*s ?g x" 
+    have "?g (c *s x) = c*s ?g x"
       unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]
       by (vector ring_simps)}
   moreover
@@ -4177,7 +4177,7 @@
 	using conjunct1[OF h, OF span_superset, OF insertI1]
 	by (auto simp add: span_0)
 
-      from xa ha1[symmetric] have "?g x = f x" 
+      from xa ha1[symmetric] have "?g x = f x"
 	apply simp
 	using g(2)[rule_format, OF span_0, of 0]
 	by simp}
@@ -4201,12 +4201,12 @@
 proof-
   from maximal_independent_subset_extend[of B UNIV] iB
   obtain C where C: "B \<subseteq> C" "independent C" "\<And>x. x \<in> span C" by auto
-  
+
   from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f]
-  obtain g where g: "(\<forall>x\<in> span C. \<forall>y\<in> span C. g (x + y) = g x + g y) 
+  obtain g where g: "(\<forall>x\<in> span C. \<forall>y\<in> span C. g (x + y) = g x + g y)
            \<and> (\<forall>x\<in> span C. \<forall>c. g (c*s x) = c *s g x)
            \<and> (\<forall>x\<in> C. g x = f x)" by blast
-  from g show ?thesis unfolding linear_def using C 
+  from g show ?thesis unfolding linear_def using C
     apply clarsimp by blast
 qed
 
@@ -4218,7 +4218,7 @@
 proof(induct arbitrary: B rule: finite_induct[OF fA])
   case 1 thus ?case by simp
 next
-  case (2 x s t) 
+  case (2 x s t)
   thus ?case
   proof(induct rule: finite_induct[OF "2.prems"(1)])
     case 1    then show ?case by simp
@@ -4234,7 +4234,7 @@
   qed
 qed
 
-lemma card_subset_eq: assumes fB: "finite B" and AB: "A \<subseteq> B" and 
+lemma card_subset_eq: assumes fB: "finite B" and AB: "A \<subseteq> B" and
   c: "card A = card B"
   shows "A = B"
 proof-
@@ -4245,27 +4245,27 @@
   from card_Un_disjoint[OF fA fBA e, unfolded eq c]
   have "card (B - A) = 0" by arith
   hence "B - A = {}" unfolding card_eq_0_iff using fA fB by simp
-  with AB show "A = B" by blast  
+  with AB show "A = B" by blast
 qed
 
 lemma subspace_isomorphism:
-  assumes s: "subspace (S:: (real ^'n) set)" and t: "subspace T" 
+  assumes s: "subspace (S:: (real ^'n) set)" and t: "subspace T"
   and d: "dim S = dim T"
   shows "\<exists>f. linear f \<and> f ` S = T \<and> inj_on f S"
 proof-
-  from basis_exists[of S] obtain B where 
+  from basis_exists[of S] obtain B where
     B: "B \<subseteq> S" "independent B" "S \<subseteq> span B" "B hassize dim S" by blast
-  from basis_exists[of T] obtain C where 
+  from basis_exists[of T] obtain C where
     C: "C \<subseteq> T" "independent C" "T \<subseteq> span C" "C hassize dim T" by blast
   from B(4) C(4) card_le_inj[of B C] d obtain f where
-    f: "f ` B \<subseteq> C" "inj_on f B" unfolding hassize_def by auto 
+    f: "f ` B \<subseteq> C" "inj_on f B" unfolding hassize_def by auto
   from linear_independent_extend[OF B(2)] obtain g where
     g: "linear g" "\<forall>x\<in> B. g x = f x" by blast
   from B(4) have fB: "finite B" by (simp add: hassize_def)
   from C(4) have fC: "finite C" by (simp add: hassize_def)
-  from inj_on_iff_eq_card[OF fB, of f] f(2) 
+  from inj_on_iff_eq_card[OF fB, of f] f(2)
   have "card (f ` B) = card B" by simp
-  with B(4) C(4) have ceq: "card (f ` B) = card C" using d 
+  with B(4) C(4) have ceq: "card (f ` B) = card C" using d
     by (simp add: hassize_def)
   have "g ` B = f ` B" using g(2)
     by (auto simp add: image_iff)
@@ -4277,9 +4277,9 @@
   {fix x y assume x: "x \<in> S" and y: "y \<in> S" and gxy:"g x = g y"
     from B(3) x y have x': "x \<in> span B" and y': "y \<in> span B" by blast+
     from gxy have th0: "g (x - y) = 0" by (simp add: linear_sub[OF g(1)])
-    have th1: "x - y \<in> span B" using x' y' by (metis span_sub) 
+    have th1: "x - y \<in> span B" using x' y' by (metis span_sub)
     have "x=y" using g0[OF th1 th0] by simp }
-  then have giS: "inj_on g S" 
+  then have giS: "inj_on g S"
     unfolding inj_on_def by blast
   from span_subspace[OF B(1,3) s]
   have "g ` S = span (g ` B)" by (simp add: span_linear_image[OF g(1)])
@@ -4308,20 +4308,20 @@
 qed
 
 lemma linear_eq_0:
-  assumes lf: "linear f" and SB: "S \<subseteq> span B" and f0: "\<forall>x\<in>B. f x = 0" 
+  assumes lf: "linear f" and SB: "S \<subseteq> span B" and f0: "\<forall>x\<in>B. f x = 0"
   shows "\<forall>x \<in> S. f x = (0::'a::semiring_1^'n)"
   by (metis linear_eq_0_span[OF lf] subset_eq SB f0)
 
 lemma linear_eq:
   assumes lf: "linear (f::'a::ring_1^'n \<Rightarrow> _)" and lg: "linear g" and S: "S \<subseteq> span B"
-  and fg: "\<forall> x\<in> B. f x = g x" 
+  and fg: "\<forall> x\<in> B. f x = g x"
   shows "\<forall>x\<in> S. f x = g x"
 proof-
   let ?h = "\<lambda>x. f x - g x"
   from fg have fg': "\<forall>x\<in> B. ?h x = 0" by simp
   from linear_eq_0[OF linear_compose_sub[OF lf lg] S fg']
   show ?thesis by simp
-qed    
+qed
 
 lemma linear_eq_stdbasis:
   assumes lf: "linear (f::'a::ring_1^'m \<Rightarrow> 'a^'n)" and lg: "linear g"
@@ -4329,7 +4329,7 @@
   shows "f = g"
 proof-
   let ?U = "UNIV :: 'm set"
-  let ?I = "{basis i:: 'a^'m|i. i \<in> {1 .. dimindex ?U}}" 
+  let ?I = "{basis i:: 'a^'m|i. i \<in> {1 .. dimindex ?U}}"
   {fix x assume x: "x \<in> (UNIV :: ('a^'m) set)"
     from equalityD2[OF span_stdbasis]
     have IU: " (UNIV :: ('a^'m) set) \<subseteq> span ?I" by blast
@@ -4341,27 +4341,27 @@
 (* Similar results for bilinear functions.                                   *)
 
 lemma bilinear_eq:
-  assumes bf: "bilinear (f:: 'a::ring^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)" 
+  assumes bf: "bilinear (f:: 'a::ring^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)"
   and bg: "bilinear g"
   and SB: "S \<subseteq> span B" and TC: "T \<subseteq> span C"
   and fg: "\<forall>x\<in> B. \<forall>y\<in> C. f x y = g x y"
   shows "\<forall>x\<in>S. \<forall>y\<in>T. f x y = g x y "
 proof-
   let ?P = "\<lambda>x. \<forall>y\<in> span C. f x y = g x y"
-  from bf bg have sp: "subspace ?P" 
-    unfolding bilinear_def linear_def subspace_def bf bg  
+  from bf bg have sp: "subspace ?P"
+    unfolding bilinear_def linear_def subspace_def bf bg
     by(auto simp add: span_0 mem_def bilinear_lzero[OF bf] bilinear_lzero[OF bg] span_add Ball_def intro:  bilinear_ladd[OF bf])
 
-  have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y" 
+  have "\<forall>x \<in> span B. \<forall>y\<in> span C. f x y = g x y"
     apply -
     apply (rule ballI)
-    apply (rule span_induct[of B ?P]) 
+    apply (rule span_induct[of B ?P])
     defer
     apply (rule sp)
     apply assumption
     apply (clarsimp simp add: Ball_def)
     apply (rule_tac P="\<lambda>y. f xa y = g xa y" and S=C in span_induct)
-    using fg 
+    using fg
     apply (auto simp add: subspace_def)
     using bf bg unfolding bilinear_def linear_def
     by(auto simp add: span_0 mem_def bilinear_rzero[OF bf] bilinear_rzero[OF bg] span_add Ball_def intro:  bilinear_ladd[OF bf])
@@ -4369,7 +4369,7 @@
 qed
 
 lemma bilinear_eq_stdbasis:
-  assumes bf: "bilinear (f:: 'a::ring_1^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)" 
+  assumes bf: "bilinear (f:: 'a::ring_1^'m \<Rightarrow> 'a^'n \<Rightarrow> 'a^'p)"
   and bg: "bilinear g"
   and fg: "\<forall>i\<in> {1 .. dimindex (UNIV :: 'm set)}. \<forall>j\<in>  {1 .. dimindex (UNIV :: 'n set)}. f (basis i) (basis j) = g (basis i) (basis j)"
   shows "f = g"
@@ -4394,16 +4394,16 @@
 proof-
   from linear_independent_extend[OF independent_injective_image, OF independent_stdbasis, OF lf fi]
   obtain h:: "real ^'m \<Rightarrow> real ^'n" where h: "linear h" " \<forall>x \<in> f ` {basis i|i. i \<in> {1 .. dimindex (UNIV::'n set)}}. h x = inv f x" by blast
-  from h(2) 
+  from h(2)
   have th: "\<forall>i\<in>{1..dimindex (UNIV::'n set)}. (h \<circ> f) (basis i) = id (basis i)"
     using inv_o_cancel[OF fi, unfolded stupid_ext[symmetric] id_def o_def]
     apply auto
     apply (erule_tac x="basis i" in allE)
     by auto
-  
+
   from linear_eq_stdbasis[OF linear_compose[OF lf h(1)] linear_id th]
   have "h o f = id" .
-  then show ?thesis using h(1) by blast  
+  then show ?thesis using h(1) by blast
 qed
 
 lemma linear_surjective_right_inverse:
@@ -4411,18 +4411,18 @@
   shows "\<exists>g. linear g \<and> f o g = id"
 proof-
   from linear_independent_extend[OF independent_stdbasis]
-  obtain h:: "real ^'n \<Rightarrow> real ^'m" where 
+  obtain h:: "real ^'n \<Rightarrow> real ^'m" where
     h: "linear h" "\<forall> x\<in> {basis i| i. i\<in> {1 .. dimindex (UNIV :: 'n set)}}. h x = inv f x" by blast
-  from h(2) 
+  from h(2)
   have th: "\<forall>i\<in>{1..dimindex (UNIV::'n set)}. (f o h) (basis i) = id (basis i)"
     using sf
     apply (auto simp add: surj_iff o_def stupid_ext[symmetric])
     apply (erule_tac x="basis i" in allE)
     by auto
-  
+
   from linear_eq_stdbasis[OF linear_compose[OF h(1) lf] linear_id th]
   have "f o h = id" .
-  then show ?thesis using h(1) by blast  
+  then show ?thesis using h(1) by blast
 qed
 
 lemma matrix_left_invertible_injective:
@@ -4434,7 +4434,7 @@
       unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .}
   moreover
   {assume A: "\<forall>x y. A *v x = A *v y \<longrightarrow> x = y"
-    hence i: "inj (op *v A)" unfolding inj_on_def by auto 
+    hence i: "inj (op *v A)" unfolding inj_on_def by auto
     from linear_injective_left_inverse[OF matrix_vector_mul_linear i]
     obtain g where g: "linear g" "g o op *v A = id" by blast
     have "matrix g ** A = mat 1"
@@ -4454,25 +4454,25 @@
 "(\<exists>B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
 proof-
   {fix B :: "real ^'m^'n"  assume AB: "A ** B = mat 1"
-    {fix x :: "real ^ 'm" 
+    {fix x :: "real ^ 'm"
       have "A *v (B *v x) = x"
 	by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)}
     hence "surj (op *v A)" unfolding surj_def by metis }
   moreover
   {assume sf: "surj (op *v A)"
     from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf]
-    obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A o g = id" 
+    obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A o g = id"
       by blast
 
     have "A ** (matrix g) = mat 1"
-      unfolding matrix_eq  matrix_vector_mul_lid 
-	matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] 
+      unfolding matrix_eq  matrix_vector_mul_lid
+	matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
       using g(2) unfolding o_def stupid_ext[symmetric] id_def
       .
     hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast
   }
   ultimately show ?thesis unfolding surj_def by blast
-qed    
+qed
 
 lemma matrix_left_invertible_independent_columns:
   fixes A :: "real^'n^'m"
@@ -4481,7 +4481,7 @@
 proof-
   let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
   {assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
-    {fix c i assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0" 
+    {fix c i assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0"
       and i: "i \<in> ?U"
       let ?x = "\<chi> i. c i"
       have th0:"A *v ?x = 0"
@@ -4493,11 +4493,11 @@
     hence ?rhs by blast}
   moreover
   {assume H: ?rhs
-    {fix x assume x: "A *v x = 0" 
+    {fix x assume x: "A *v x = 0"
       let ?c = "\<lambda>i. ((x$i ):: real)"
       from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x]
       have "x = 0" by vector}}
-  ultimately show ?thesis unfolding matrix_left_invertible_ker by blast 
+  ultimately show ?thesis unfolding matrix_left_invertible_ker by blast
 qed
 
 lemma matrix_right_invertible_independent_rows:
@@ -4514,13 +4514,13 @@
   have fU: "finite ?U" by simp
   have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"
     unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def
-    apply (subst eq_commute) ..    
+    apply (subst eq_commute) ..
   have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast
   {assume h: ?lhs
-    {fix x:: "real ^'n" 
+    {fix x:: "real ^'n"
 	from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m"
 	  where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
-	have "x \<in> span (columns A)"  
+	have "x \<in> span (columns A)"
 	  unfolding y[symmetric]
 	  apply (rule span_setsum[OF fU])
 	  apply clarify
@@ -4532,21 +4532,21 @@
   moreover
   {assume h:?rhs
     let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
-    {fix y have "?P y" 
+    {fix y have "?P y"
       proof(rule span_induct_alt[of ?P "columns A"])
 	show "\<exists>x\<Colon>real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
 	  apply (rule exI[where x=0])
 	  by (simp add: zero_index vector_smult_lzero)
       next
 	fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"
-	from y1 obtain i where i: "i \<in> ?U" "y1 = column i A" 
+	from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
 	  unfolding columns_def by blast
-	from y2 obtain x:: "real ^'m" where 
+	from y2 obtain x:: "real ^'m" where
 	  x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
 	let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
 	show "?P (c*s y1 + y2)"
 	  proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric]Cart_lambda_beta setsum_component cond_value_iff right_distrib cond_application_beta vector_component cong del: if_weak_cong, simp only: One_nat_def[symmetric])
-	    fix j 
+	    fix j
 	    have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1)
 	      by (simp add: ring_simps)
@@ -4558,7 +4558,7 @@
 	      by (simp add: setsum_addf)
 	    also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
 	      unfolding setsum_delta[OF fU]
-	      using i(1) by simp 
+	      using i(1) by simp
 	    finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
 	  qed
@@ -4579,12 +4579,12 @@
 (* An injective map real^'n->real^'n is also surjective.                       *)
 
 lemma linear_injective_imp_surjective:
-  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and fi: "inj f" 
+  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and fi: "inj f"
   shows "surj f"
 proof-
   let ?U = "UNIV :: (real ^'n) set"
-  from basis_exists[of ?U] obtain B 
-    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U" 
+  from basis_exists[of ?U] obtain B
+    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U"
     by blast
   from B(4) have d: "dim ?U = card B" by (simp add: hassize_def)
   have th: "?U \<subseteq> span (f ` B)"
@@ -4604,7 +4604,7 @@
 
 (* And vice versa.                                                           *)
 
-lemma surjective_iff_injective_gen: 
+lemma surjective_iff_injective_gen:
   assumes fS: "finite S" and fT: "finite T" and c: "card S = card T"
   and ST: "f ` S \<subseteq> T"
   shows "(\<forall>y \<in> T. \<exists>x \<in> S. f x = y) \<longleftrightarrow> inj_on f S" (is "?lhs \<longleftrightarrow> ?rhs")
@@ -4641,17 +4641,17 @@
 qed
 
 lemma linear_surjective_imp_injective:
-  assumes lf: "linear (f::real ^'n => real ^'n)" and sf: "surj f" 
+  assumes lf: "linear (f::real ^'n => real ^'n)" and sf: "surj f"
   shows "inj f"
 proof-
   let ?U = "UNIV :: (real ^'n) set"
-  from basis_exists[of ?U] obtain B 
-    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U" 
+  from basis_exists[of ?U] obtain B
+    where B: "B \<subseteq> ?U" "independent B" "?U \<subseteq> span B" "B hassize dim ?U"
     by blast
   {fix x assume x: "x \<in> span B" and fx: "f x = 0"
     from B(4) have fB: "finite B" by (simp add: hassize_def)
     from B(4) have d: "dim ?U = card B" by (simp add: hassize_def)
-    have fBi: "independent (f ` B)" 
+    have fBi: "independent (f ` B)"
       apply (rule card_le_dim_spanning[of "f ` B" ?U])
       apply blast
       using sf B(3)
@@ -4676,12 +4676,12 @@
     moreover have "card (f ` B) \<le> card B"
       by (rule card_image_le, rule fB)
     ultimately have th1: "card B = card (f ` B)" unfolding d by arith
-    have fiB: "inj_on f B" 
+    have fiB: "inj_on f B"
       unfolding surjective_iff_injective_gen[OF fB finite_imageI[OF fB] th1 subset_refl, symmetric] by blast
     from linear_indep_image_lemma[OF lf fB fBi fiB x] fx
     have "x = 0" by blast}
   note th = this
-  from th show ?thesis unfolding linear_injective_0[OF lf] 
+  from th show ?thesis unfolding linear_injective_0[OF lf]
     using B(3) by blast
 qed
 
@@ -4689,7 +4689,7 @@
 
 lemma left_right_inverse_eq:
   assumes fg: "f o g = id" and gh: "g o h = id"
-  shows "f = h" 
+  shows "f = h"
 proof-
   have "f = f o (g o h)" unfolding gh by simp
   also have "\<dots> = (f o g) o h" by (simp add: o_assoc)
@@ -4723,7 +4723,7 @@
   {fix f f':: "real ^'n \<Rightarrow> real ^'n"
     assume lf: "linear f" "linear f'" and f: "f o f' = id"
     from f have sf: "surj f"
-      
+
       apply (auto simp add: o_def stupid_ext[symmetric] id_def surj_def)
       by metis
     from linear_surjective_isomorphism[OF lf(1) sf] lf f
@@ -4735,13 +4735,13 @@
 (* Moreover, a one-sided inverse is automatically linear.                    *)
 
 lemma left_inverse_linear:
-  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and gf: "g o f = id" 
+  assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'n)" and gf: "g o f = id"
   shows "linear g"
 proof-
   from gf have fi: "inj f" apply (auto simp add: inj_on_def o_def id_def stupid_ext[symmetric])
     by metis
-  from linear_injective_isomorphism[OF lf fi] 
-  obtain h:: "real ^'n \<Rightarrow> real ^'n" where 
+  from linear_injective_isomorphism[OF lf fi]
+  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
     h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast
   have "h = g" apply (rule ext) using gf h(2,3)
     apply (simp add: o_def id_def stupid_ext[symmetric])
@@ -4750,13 +4750,13 @@
 qed
 
 lemma right_inverse_linear:
-  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and gf: "f o g = id" 
+  assumes lf: "linear (f:: real ^'n \<Rightarrow> real ^'n)" and gf: "f o g = id"
   shows "linear g"
 proof-
   from gf have fi: "surj f" apply (auto simp add: surj_def o_def id_def stupid_ext[symmetric])
     by metis
-  from linear_surjective_isomorphism[OF lf fi] 
-  obtain h:: "real ^'n \<Rightarrow> real ^'n" where 
+  from linear_surjective_isomorphism[OF lf fi]
+  obtain h:: "real ^'n \<Rightarrow> real ^'n" where
     h: "linear h" "\<forall>x. h (f x) = x" "\<forall>x. f (h x) = x" by blast
   have "h = g" apply (rule ext) using gf h(2,3)
     apply (simp add: o_def id_def stupid_ext[symmetric])
@@ -4767,7 +4767,7 @@
 (* The same result in terms of square matrices.                              *)
 
 lemma matrix_left_right_inverse:
-  fixes A A' :: "real ^'n^'n" 
+  fixes A A' :: "real ^'n^'n"
   shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
 proof-
   {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1"
@@ -4779,7 +4779,7 @@
     from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA]
     obtain f' :: "real ^'n \<Rightarrow> real ^'n"
       where f': "linear f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast
-    have th: "matrix f' ** A = mat 1" 
+    have th: "matrix f' ** A = mat 1"
       by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format])
     hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp
     hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid)
@@ -4846,17 +4846,17 @@
   have th: "\<And>x y (z::real). x - y <= z \<longleftrightarrow> x - z <= y" by arith
   have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
   have th2: "\<And>x (y::real). abs(x + y) - abs(x) <= abs(y)" by arith
-  show ?thesis 
+  show ?thesis
   unfolding infnorm_def
   unfolding rsup_finite_le_iff[ OF infnorm_set_lemma]
   apply (subst diff_le_eq[symmetric])
   unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
-  unfolding infnorm_set_image bex_simps 
+  unfolding infnorm_set_image bex_simps
   apply (subst th)
-  unfolding th1 
+  unfolding th1
   unfolding rsup_finite_ge_iff[ OF infnorm_set_lemma]
-  
-  unfolding infnorm_set_image ball_simps bex_simps 
+
+  unfolding infnorm_set_image ball_simps bex_simps
   apply (simp add: vector_add_component)
   apply (metis numseg_dimindex_nonempty th2)
   done
@@ -4885,7 +4885,7 @@
   apply (simp add: vector_component)
   done
 
-lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)" 
+lemma infnorm_sub: "infnorm (x - y) = infnorm (y - x)"
 proof-
   have "y - x = - (x - y)" by simp
   then show ?thesis  by (metis infnorm_neg)
@@ -4896,7 +4896,7 @@
   have th: "\<And>(nx::real) n ny. nx <= n + ny \<Longrightarrow> ny <= n + nx ==> \<bar>nx - ny\<bar> <= n"
     by arith
   from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
-  have ths: "infnorm x \<le> infnorm (x - y) + infnorm y" 
+  have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
     "infnorm y \<le> infnorm (x - y) + infnorm x"
     by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])
   from th[OF ths]  show ?thesis .
@@ -4911,11 +4911,11 @@
   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
   let ?S = "{\<bar>x$i\<bar> |i. i\<in> ?U}"
   have fS: "finite ?S" unfolding image_Collect[symmetric]
-    apply (rule finite_imageI) unfolding Collect_def mem_def by simp  
+    apply (rule finite_imageI) unfolding Collect_def mem_def by simp
   have S0: "?S \<noteq> {}" using numseg_dimindex_nonempty by blast
   have th1: "\<And>S f. f ` S = { f i| i. i \<in> S}" by blast
   from rsup_finite_in[OF fS S0] rsup_finite_Ub[OF fS S0] i
-  show ?thesis unfolding infnorm_def isUb_def setle_def 
+  show ?thesis unfolding infnorm_def isUb_def setle_def
     unfolding infnorm_set_image ball_simps by auto
 qed
 
@@ -4942,7 +4942,7 @@
     have "infnorm x \<le> 1/\<bar>a\<bar> * infnorm (a*s x)"
       unfolding th by simp
     with ap have "\<bar>a\<bar> * infnorm x \<le> \<bar>a\<bar> * (1/\<bar>a\<bar> * infnorm (a *s x))" by (simp add: field_simps)
-    then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*s x)" 
+    then have "\<bar>a\<bar> * infnorm x \<le> infnorm (a*s x)"
       using ap by (simp add: field_simps)
     with infnorm_mul_lemma[of a x] have ?thesis by arith }
   ultimately show ?thesis by blast
@@ -4954,7 +4954,7 @@
 (* Prove that it differs only up to a bound from Euclidean norm.             *)
 
 lemma infnorm_le_norm: "infnorm x \<le> norm x"
-  unfolding infnorm_def rsup_finite_le_iff[OF infnorm_set_lemma] 
+  unfolding infnorm_def rsup_finite_le_iff[OF infnorm_set_lemma]
   unfolding infnorm_set_image  ball_simps
   by (metis component_le_norm)
 lemma card_enum: "card {1 .. n} = n" by auto
@@ -4968,20 +4968,20 @@
   have th: "sqrt (real ?d) * infnorm x \<ge> 0"
     by (simp add: dimindex_ge_1 zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)
   have th1: "x\<bullet>x \<le> (sqrt (real ?d) * infnorm x)^2"
-    unfolding power_mult_distrib d2 
+    unfolding power_mult_distrib d2
     apply (subst d)
     apply (subst power2_abs[symmetric])
     unfolding real_of_nat_def dot_def power2_eq_square[symmetric]
     apply (subst power2_abs[symmetric])
     apply (rule setsum_bounded)
     apply (rule power_mono)
-    unfolding abs_of_nonneg[OF infnorm_pos_le] 
+    unfolding abs_of_nonneg[OF infnorm_pos_le]
     unfolding infnorm_def  rsup_finite_ge_iff[OF infnorm_set_lemma]
     unfolding infnorm_set_image bex_simps
     apply blast
     by (rule abs_ge_zero)
   from real_le_lsqrt[OF dot_pos_le th th1]
-  show ?thesis unfolding real_vector_norm_def id_def . 
+  show ?thesis unfolding real_vector_norm_def id_def .
 qed
 
 (* Equality in Cauchy-Schwarz and triangle inequalities.                     *)
@@ -5037,7 +5037,7 @@
   {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
     hence "norm x \<noteq> 0" "norm y \<noteq> 0"
       by simp_all
-    hence n: "norm x > 0" "norm y > 0" 
+    hence n: "norm x > 0" "norm y > 0"
       using norm_ge_zero[of x] norm_ge_zero[of y]
       by arith+
     have th: "\<And>(a::real) b c. a + b + c \<noteq> 0 ==> (a = b + c \<longleftrightarrow> a^2 = (b + c)^2)" by algebra
@@ -5058,7 +5058,7 @@
 
 lemma collinear_empty:  "collinear {}" by (simp add: collinear_def)
 
-lemma collinear_sing: "collinear {(x::'a::ring_1^'n)}" 
+lemma collinear_sing: "collinear {(x::'a::ring_1^'n)}"
   apply (simp add: collinear_def)
   apply (rule exI[where x=0])
   by simp
@@ -5075,20 +5075,20 @@
 
 lemma collinear_lemma: "collinear {(0::real^'n),x,y} \<longleftrightarrow> x = 0 \<or> y = 0 \<or> (\<exists>c. y = c *s x)" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
-  {assume "x=0 \<or> y = 0" hence ?thesis 
+  {assume "x=0 \<or> y = 0" hence ?thesis
       by (cases "x = 0", simp_all add: collinear_2 insert_commute)}
   moreover
   {assume x: "x \<noteq> 0" and y: "y \<noteq> 0"
     {assume h: "?lhs"
       then obtain u where u: "\<forall> x\<in> {0,x,y}. \<forall>y\<in> {0,x,y}. \<exists>c. x - y = c *s u" unfolding collinear_def by blast
       from u[rule_format, of x 0] u[rule_format, of y 0]
-      obtain cx and cy where 
+      obtain cx and cy where
 	cx: "x = cx*s u" and cy: "y = cy*s u"
 	by auto
       from cx x have cx0: "cx \<noteq> 0" by auto
       from cy y have cy0: "cy \<noteq> 0" by auto
       let ?d = "cy / cx"
-      from cx cy cx0 have "y = ?d *s x" 
+      from cx cy cx0 have "y = ?d *s x"
 	by (simp add: vector_smult_assoc)
       hence ?rhs using x y by blast}
     moreover
--- a/src/HOL/Library/Finite_Cartesian_Product.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Finite_Cartesian_Product.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -18,11 +18,11 @@
 translations "DIM(t)" => "CONST dimindex (CONST UNIV :: t set)"
 
 lemma dimindex_nonzero: "dimindex S \<noteq>  0"
-unfolding dimindex_def 
+unfolding dimindex_def
 by (simp add: neq0_conv[symmetric] del: neq0_conv)
 
 lemma dimindex_ge_1: "dimindex S \<ge> 1"
-  using dimindex_nonzero[of S] by arith 
+  using dimindex_nonzero[of S] by arith
 lemma dimindex_univ: "dimindex (S :: 'a set) = DIM('a)" by (simp add: dimindex_def)
 
 definition hassize (infixr "hassize" 12) where
@@ -48,7 +48,7 @@
 
 text{* Dimension of such a type, and indexing over it. *}
 
-lemma inj_on_Abs_finite_image: 
+lemma inj_on_Abs_finite_image:
   "inj_on (Abs_finite_image:: _ \<Rightarrow> 'a finite_image) {1 .. DIM('a)}"
 by (auto simp add: inj_on_def finite_image_def Abs_finite_image_inject[where ?'a='a])
 
@@ -72,23 +72,23 @@
 unfolding card_finite_image[of T, symmetric]
 by (auto simp add: dimindex_def finite_finite_image)
 
-lemma Abs_finite_image_works: 
+lemma Abs_finite_image_works:
   fixes i:: "'a finite_image"
   shows " \<exists>!n \<in> {1 .. DIM('a)}. Abs_finite_image n = i"
   unfolding Bex1_def Ex1_def
   apply (rule_tac x="Rep_finite_image i" in exI)
-  using Rep_finite_image_inverse[where ?'a = 'a] 
-    Rep_finite_image[where ?'a = 'a] 
+  using Rep_finite_image_inverse[where ?'a = 'a]
+    Rep_finite_image[where ?'a = 'a]
   Abs_finite_image_inverse[where ?'a='a, symmetric]
   by (auto simp add: finite_image_def)
 
-lemma Abs_finite_image_inj: 
+lemma Abs_finite_image_inj:
  "i \<in> {1 .. DIM('a)} \<Longrightarrow> j \<in> {1 .. DIM('a)}
   \<Longrightarrow> (((Abs_finite_image i ::'a finite_image) = Abs_finite_image j) \<longleftrightarrow> (i = j))"
-  using Abs_finite_image_works[where ?'a = 'a] 
+  using Abs_finite_image_works[where ?'a = 'a]
   by (auto simp add: atLeastAtMost_iff Bex1_def)
 
-lemma forall_Abs_finite_image: 
+lemma forall_Abs_finite_image:
   "(\<forall>k:: 'a finite_image. P k) \<longleftrightarrow> (\<forall>i \<in> {1 .. DIM('a)}. P(Abs_finite_image i))"
 unfolding Ball_def atLeastAtMost_iff Ex1_def
 using Abs_finite_image_works[where ?'a = 'a, unfolded atLeastAtMost_iff Bex1_def]
@@ -115,7 +115,7 @@
   unfolding Cart_nth_def forall_Abs_finite_image[symmetric, where P = "\<lambda>i. Rep_Cart x i = Rep_Cart y i"] stupid_ext
   using Rep_Cart_inject[of x y] ..
 
-consts Cart_lambda :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a ^ 'b" 
+consts Cart_lambda :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a ^ 'b"
 notation (xsymbols) Cart_lambda (binder "\<chi>" 10)
 
 defs Cart_lambda_def: "Cart_lambda g == (SOME (f:: 'a ^ 'b). \<forall>i \<in> {1 .. DIM('b)}. f$i = g i)"
@@ -127,13 +127,13 @@
   let ?f = "Abs_Cart (\<lambda>k. g (THE i. ?p i k)):: 'a ^ 'b"
   let ?P = "\<lambda>f i. f$i = g i"
   let ?Q = "\<lambda>(f::'a ^ 'b). \<forall> i \<in> {1 .. DIM('b)}. ?P f i"
-  {fix i 
+  {fix i
     assume i: "i \<in> {1 .. DIM('b)}"
     let ?j = "THE j. ?p j (Abs_finite_image i)"
     from theI'[where P = "\<lambda>j. ?p (j::nat) (Abs_finite_image i :: 'b finite_image)", OF Abs_finite_image_works[of "Abs_finite_image i :: 'b finite_image", unfolded Bex1_def]]
     have j: "?j \<in> {1 .. DIM('b)}" "(Abs_finite_image ?j :: 'b finite_image) = Abs_finite_image i" by blast+
     from i j Abs_finite_image_inject[of i ?j, where ?'a = 'b]
-    have th: "?j = i" by (simp add: finite_image_def)  
+    have th: "?j = i" by (simp add: finite_image_def)
     have "?P ?f i"
       using th
       by (simp add: Cart_nth_def Abs_Cart_inverse Rep_Cart_inverse Cart_def) }
@@ -175,7 +175,7 @@
 apply (simp add: Rep_finite_sum)
 done
 
-lemma inj_on_Abs_finite_sum: "inj_on (Abs_finite_sum :: _ \<Rightarrow> ('a,'b) finite_sum) {1 .. DIM('a) + DIM('b)}" 
+lemma inj_on_Abs_finite_sum: "inj_on (Abs_finite_sum :: _ \<Rightarrow> ('a,'b) finite_sum) {1 .. DIM('a) + DIM('b)}"
   using Abs_finite_sum_inject[where ?'a = 'a and ?'b = 'b]
   by (auto simp add: inj_on_def finite_sum_def)
 
@@ -196,7 +196,7 @@
 lemma pastecart_fst_snd: "pastecart (fstcart z) (sndcart z) = z"
 proof -
  {fix i
-  assume H: "i \<le> DIM('b) + DIM('c)" 
+  assume H: "i \<le> DIM('b) + DIM('c)"
     "\<not> i \<le> DIM('b)"
     from H have ith: "i - DIM('b) \<in> {1 .. DIM('c)}"
       apply simp by arith
@@ -226,8 +226,8 @@
   {fix n
     assume n: "n \<le> DIM('n)"
     have "finite {v:: 'a ^ 'n . (\<forall>i\<in> {1 .. DIM('n)}. i \<le> n \<longrightarrow> P i (v$i))
-                              \<and> (\<forall>i\<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}" 
-      using n 
+                              \<and> (\<forall>i\<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}"
+      using n
       proof(induct n)
 	case 0
 	have th0: "{v . (\<forall>i \<in> {1 .. DIM('n)}. v$i = (SOME x. False))} =
@@ -241,7 +241,7 @@
             (\<forall>i\<Colon>nat\<in>{1\<Colon>nat..DIM('n)}.
                 Suc n < i \<longrightarrow> v$i = (SOME x\<Colon>'a. False))}"
 	let ?S = "{x::'a . P (Suc  n) x} \<times> {v:: 'a^'n. (\<forall>i \<in> {1 .. DIM('n)}. i <= n \<longrightarrow> P i (v$i)) \<and> (\<forall>i \<in> {1 .. DIM('n)}. n < i \<longrightarrow> v$i = (SOME x. False))}"
-	have th0: " ?T \<subseteq> (?h ` ?S)" 
+	have th0: " ?T \<subseteq> (?h ` ?S)"
 	  using Suc.prems
 	  apply (auto simp add: image_def)
 	  apply (rule_tac x = "x$(Suc n)" in exI)
@@ -252,11 +252,11 @@
 	  apply simp
 	  apply (rule_tac x= "\<chi> i. if i = Suc n then (SOME x:: 'a. False) else (x:: 'a ^ 'n)$i:: 'a ^ 'n" in exI)
 	  by (simp add: Cart_eq Cart_lambda_beta)
-	have th1: "finite ?S" 
-	  apply (rule finite_cartesian_product) 
-	  using f Suc.hyps Suc.prems by auto 
-	from finite_imageI[OF th1] have th2: "finite (?h ` ?S)" . 
-	from finite_subset[OF th0 th2] show ?case by blast 
+	have th1: "finite ?S"
+	  apply (rule finite_cartesian_product)
+	  using f Suc.hyps Suc.prems by auto
+	from finite_imageI[OF th1] have th2: "finite (?h ` ?S)" .
+	from finite_subset[OF th0 th2] show ?case by blast
       qed}
 
   note th = this
--- a/src/HOL/Library/Formal_Power_Series.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -1,5 +1,5 @@
 (*  Title:      Formal_Power_Series.thy
-    ID:         
+    ID:
     Author:     Amine Chaieb, University of Cambridge
 *)
 
@@ -49,7 +49,7 @@
 instance ..
 end
 
-lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)" 
+lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)"
   unfolding fps_one_def by simp
 
 instantiation fps :: (plus)  plus
@@ -119,7 +119,7 @@
 lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
   by auto
 
-subsection{* Formal power series form a commutative ring with unity, if the range of sequences 
+subsection{* Formal power series form a commutative ring with unity, if the range of sequences
   they represent is a commutative ring with unity*}
 
 instance fps :: (semigroup_add) semigroup_add
@@ -303,7 +303,7 @@
 lemma fps_eq_iff: "f = g \<longleftrightarrow> (\<forall>n. f $ n = g $n)"
   by (rule expand_fps_eq)
 
-lemma fps_setsum_nth: "(setsum f S) $ n = setsum (\<lambda>k. (f k) $ n) S" 
+lemma fps_setsum_nth: "(setsum f S) $ n = setsum (\<lambda>k. (f k) $ n) S"
 proof (cases "finite S")
   assume "\<not> finite S" then show ?thesis by simp
 next
@@ -397,17 +397,17 @@
 instantiation fps :: ("{comm_monoid_add,inverse, times, uminus}") inverse
 begin
 
-fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a" where 
+fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a" where
   "natfun_inverse f 0 = inverse (f$0)"
-| "natfun_inverse f n = - inverse (f$0) * setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}" 
+| "natfun_inverse f n = - inverse (f$0) * setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
 
-definition fps_inverse_def: 
+definition fps_inverse_def:
   "inverse f = (if f$0 = 0 then 0 else Abs_fps (natfun_inverse f))"
 definition fps_divide_def: "divide = (\<lambda>(f::'a fps) g. f * inverse g)"
 instance ..
 end
 
-lemma fps_inverse_zero[simp]: 
+lemma fps_inverse_zero[simp]:
   "inverse (0 :: 'a::{comm_monoid_add,inverse, times, uminus} fps) = 0"
   by (simp add: fps_ext fps_inverse_def)
 
@@ -422,7 +422,7 @@
   shows "inverse f * f = 1"
 proof-
   have c: "inverse f * f = f * inverse f" by (simp add: mult_commute)
-  from f0 have ifn: "\<And>n. inverse f $ n = natfun_inverse f n" 
+  from f0 have ifn: "\<And>n. inverse f $ n = natfun_inverse f n"
     by (simp add: fps_inverse_def)
   from f0 have th0: "(inverse f * f) $ 0 = 1"
     by (simp add: fps_mult_nth fps_inverse_def)
@@ -430,16 +430,16 @@
     from np have eq: "{0..n} = {0} \<union> {1 .. n}" by auto
     have d: "{0} \<inter> {1 .. n} = {}" by auto
     have f: "finite {0::nat}" "finite {1..n}" by auto
-    from f0 np have th0: "- (inverse f$n) = 
+    from f0 np have th0: "- (inverse f$n) =
       (setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
       by (cases n, simp, simp add: divide_inverse fps_inverse_def)
     from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
-    have th1: "setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} = 
-      - (f$0) * (inverse f)$n" 
+    have th1: "setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} =
+      - (f$0) * (inverse f)$n"
       by (simp add: ring_simps)
-    have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))" 
+    have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
       unfolding fps_mult_nth ifn ..
-    also have "\<dots> = f$0 * natfun_inverse f n 
+    also have "\<dots> = f$0 * natfun_inverse f n
       + (\<Sum>i = 1..n. f$i * natfun_inverse f (n-i))"
       unfolding setsum_Un_disjoint[OF f d, unfolded eq[symmetric]]
       by simp
@@ -464,12 +464,12 @@
   shows "inverse (inverse f) = f"
 proof-
   from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
-  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0] 
+  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
   have th0: "inverse f * f = inverse f * inverse (inverse f)"   by (simp add: mult_ac)
   then show ?thesis using f0 unfolding mult_cancel_left by simp
 qed
 
-lemma fps_inverse_unique: assumes f0: "f$0 \<noteq> (0::'a::field)" and fg: "f*g = 1" 
+lemma fps_inverse_unique: assumes f0: "f$0 \<noteq> (0::'a::field)" and fg: "f*g = 1"
   shows "inverse f = g"
 proof-
   from inverse_mult_eq_1[OF f0] fg
@@ -478,7 +478,7 @@
     by (auto simp add: expand_fps_eq)
 qed
 
-lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field))) 
+lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
   = Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
   apply (rule fps_inverse_unique)
   apply simp
@@ -488,15 +488,15 @@
   let ?f = "\<lambda>i. if n = i then (1\<Colon>'a) else if n - i = 1 then - 1 else 0"
   let ?g = "\<lambda>i. if i = n then 1 else if i=n - 1 then - 1 else 0"
   let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
-  have th1: "setsum ?f {0..n} = setsum ?g {0..n}"  
+  have th1: "setsum ?f {0..n} = setsum ?g {0..n}"
     by (rule setsum_cong2) auto
-  have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"  
+  have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"
     using n apply - by (rule setsum_cong2) auto
   have eq: "{0 .. n} = {0.. n - 1} \<union> {n}" by auto
-  from n have d: "{0.. n - 1} \<inter> {n} = {}" by auto 
+  from n have d: "{0.. n - 1} \<inter> {n} = {}" by auto
   have f: "finite {0.. n - 1}" "finite {n}" by auto
   show "setsum ?f {0..n} = 0"
-    unfolding th1 
+    unfolding th1
     apply (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
     unfolding th2
     by(simp add: setsum_delta)
@@ -511,7 +511,7 @@
 lemma fps_deriv_linear[simp]: "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) = fps_const a * fps_deriv f + fps_const b * fps_deriv g"
   unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: ring_simps)
 
-lemma fps_deriv_mult[simp]: 
+lemma fps_deriv_mult[simp]:
   fixes f :: "('a :: comm_ring_1) fps"
   shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
 proof-
@@ -562,7 +562,7 @@
       apply (rule setsum_cong2)
       by (auto simp add: of_nat_diff ring_simps)
     finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .}
-  then show ?thesis unfolding fps_eq_iff by auto 
+  then show ?thesis unfolding fps_eq_iff by auto
 qed
 
 lemma fps_deriv_neg[simp]: "fps_deriv (- (f:: ('a:: comm_ring_1) fps)) = - (fps_deriv f)"
@@ -571,7 +571,7 @@
   using fps_deriv_linear[of 1 f 1 g] by simp
 
 lemma fps_deriv_sub[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) - g) = fps_deriv f - fps_deriv g"
-  unfolding diff_minus by simp 
+  unfolding diff_minus by simp
 
 lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
   by (simp add: fps_ext fps_deriv_def fps_const_def)
@@ -612,7 +612,7 @@
   ultimately show ?thesis by blast
 qed
 
-lemma fps_deriv_eq_iff: 
+lemma fps_deriv_eq_iff:
   fixes f:: "('a::{idom,semiring_char_0}) fps"
   shows "fps_deriv f = fps_deriv g \<longleftrightarrow> (f = fps_const(f$0 - g$0) + g)"
 proof-
@@ -623,7 +623,7 @@
 
 lemma fps_deriv_eq_iff_ex: "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>(c::'a::{idom,semiring_char_0}). f = fps_const c + g)"
   apply auto unfolding fps_deriv_eq_iff by blast
-  
+
 
 fun fps_nth_deriv :: "nat \<Rightarrow> ('a::semiring_1) fps \<Rightarrow> 'a fps" where
   "fps_nth_deriv 0 f = f"
@@ -642,7 +642,7 @@
   using fps_nth_deriv_linear[of n 1 f 1 g] by simp
 
 lemma fps_nth_deriv_sub[simp]: "fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
-  unfolding diff_minus fps_nth_deriv_add by simp 
+  unfolding diff_minus fps_nth_deriv_add by simp
 
 lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
   by (induct n, simp_all )
@@ -700,7 +700,7 @@
 next
   case (Suc n)
   note h = Suc.hyps[OF `a$0 = 1`]
-  show ?case unfolding power_Suc fps_mult_nth 
+  show ?case unfolding power_Suc fps_mult_nth
     using h `a$0 = 1`  fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: ring_simps)
 qed
 
@@ -719,10 +719,10 @@
 apply (induct n, auto simp add: power_Suc fps_mult_nth)
 by (rule startsby_zero_power, simp_all)
 
-lemma startsby_zero_power_prefix: 
+lemma startsby_zero_power_prefix:
   assumes a0: "a $0 = (0::'a::idom)"
   shows "\<forall>n < k. a ^ k $ n = 0"
-  using a0 
+  using a0
 proof(induct k rule: nat_less_induct)
   fix k assume H: "\<forall>m<k. a $0 =  0 \<longrightarrow> (\<forall>n<m. a ^ m $ n = 0)" and a0: "a $0 = (0\<Colon>'a)"
   let ?ths = "\<forall>m<k. a ^ k $ m = 0"
@@ -730,7 +730,7 @@
   moreover
   {fix l assume k: "k = Suc l"
     {fix m assume mk: "m < k"
-      {assume "m=0" hence "a^k $ m = 0" using startsby_zero_power[of a k] k a0 
+      {assume "m=0" hence "a^k $ m = 0" using startsby_zero_power[of a k] k a0
 	  by simp}
       moreover
       {assume m0: "m \<noteq> 0"
@@ -742,14 +742,14 @@
 	  using a0
 	  apply simp
 	  apply (rule H[rule_format])
-	  using a0 k mk by auto 
+	  using a0 k mk by auto
 	finally have "a^k $ m = 0" .}
     ultimately have "a^k $ m = 0" by blast}
     hence ?ths by blast}
   ultimately show ?ths by (cases k, auto)
 qed
 
-lemma startsby_zero_setsum_depends: 
+lemma startsby_zero_setsum_depends:
   assumes a0: "a $0 = (0::'a::idom)" and kn: "n \<ge> k"
   shows "setsum (\<lambda>i. (a ^ i)$k) {0 .. n} = setsum (\<lambda>i. (a ^ i)$k) {0 .. k}"
   apply (rule setsum_mono_zero_right)
@@ -786,7 +786,7 @@
     {assume "n = 0" hence ?thesis by simp}
     moreover
     {assume n: "n > 0"
-      from startsby_zero_power[OF a0 n] eq a0 n have ?thesis 
+      from startsby_zero_power[OF a0 n] eq a0 n have ?thesis
 	by (simp add: fps_inverse_def)}
     ultimately have ?thesis by blast}
   moreover
@@ -806,7 +806,7 @@
   apply (induct n, auto simp add: power_Suc ring_simps fps_const_add[symmetric] simp del: fps_const_add)
   by (case_tac n, auto simp add: power_Suc ring_simps)
 
-lemma fps_inverse_deriv: 
+lemma fps_inverse_deriv:
   fixes a:: "('a :: field) fps"
   assumes a0: "a$0 \<noteq> 0"
   shows "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2"
@@ -825,7 +825,7 @@
   then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: ring_simps)
 qed
 
-lemma fps_inverse_mult: 
+lemma fps_inverse_mult:
   fixes a::"('a :: field) fps"
   shows "inverse (a * b) = inverse a * inverse b"
 proof-
@@ -839,7 +839,7 @@
   moreover
   {assume a0: "a$0 \<noteq> 0" and b0: "b$0 \<noteq> 0"
     from a0 b0 have ab0:"(a*b) $ 0 \<noteq> 0" by (simp  add: fps_mult_nth)
-    from inverse_mult_eq_1[OF ab0] 
+    from inverse_mult_eq_1[OF ab0]
     have "inverse (a*b) * (a*b) * inverse a * inverse b = 1 * inverse a * inverse b" by simp
     then have "inverse (a*b) * (inverse a * a) * (inverse b * b) = inverse a * inverse b"
       by (simp add: ring_simps)
@@ -847,7 +847,7 @@
 ultimately show ?thesis by blast
 qed
 
-lemma fps_inverse_deriv': 
+lemma fps_inverse_deriv':
   fixes a:: "('a :: field) fps"
   assumes a0: "a$0 \<noteq> 0"
   shows "fps_deriv (inverse a) = - fps_deriv a / a ^ 2"
@@ -864,7 +864,7 @@
   shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b ^ 2"
   using fps_inverse_deriv[OF a0]
   by (simp add: fps_divide_def ring_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
-  
+
 subsection{* The eXtractor series X*}
 
 lemma minus_one_power_iff: "(- (1::'a :: {recpower, comm_ring_1})) ^ n = (if even n then 1 else - 1)"
@@ -872,7 +872,7 @@
 
 definition "X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
 
-lemma fps_inverse_gp': "inverse (Abs_fps(\<lambda>n. (1::'a::field))) 
+lemma fps_inverse_gp': "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
   = 1 - X"
   by (simp add: fps_inverse_gp fps_eq_iff X_def)
 
@@ -897,7 +897,7 @@
   case 0 thus ?case by (simp add: X_def fps_power_def fps_eq_iff)
 next
   case (Suc k)
-  {fix m 
+  {fix m
     have "(X^Suc k) $ m = (if m = 0 then (0::'a) else (X^k) $ (m - 1))"
       by (simp add: power_Suc del: One_nat_def)
     then     have "(X^Suc k) $ m = (if m = Suc k then (1::'a) else 0)"
@@ -908,7 +908,7 @@
 lemma X_power_mult_nth: "(X^k * (f :: ('a::comm_ring_1) fps)) $n = (if n < k then 0 else f $ (n - k))"
   apply (induct k arbitrary: n)
   apply (simp)
-  unfolding power_Suc mult_assoc 
+  unfolding power_Suc mult_assoc
   by (case_tac n, auto)
 
 lemma X_power_mult_right_nth: "((f :: ('a::comm_ring_1) fps) * X^k) $n = (if n < k then 0 else f $ (n - k))"
@@ -933,7 +933,7 @@
   show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
 qed
 
-  
+
 subsection{* Integration *}
 definition "fps_integral a a0 = Abs_fps (\<lambda>n. if n = 0 then a0 else (a$(n - 1) / of_nat n))"
 
@@ -947,7 +947,7 @@
   ultimately show ?thesis
     unfolding fps_deriv_eq_iff by auto
 qed
-  
+
 subsection {* Composition of FPSs *}
 definition fps_compose :: "('a::semiring_1) fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl "oo" 55) where
   fps_compose_def: "a oo b = Abs_fps (\<lambda>n. setsum (\<lambda>i. a$i * (b^i$n)) {0..n})"
@@ -956,8 +956,8 @@
 
 lemma fps_compose_X[simp]: "a oo X = (a :: ('a :: comm_ring_1) fps)"
   by (simp add: fps_ext fps_compose_def mult_delta_right setsum_delta')
- 
-lemma fps_const_compose[simp]: 
+
+lemma fps_const_compose[simp]:
   "fps_const (a::'a::{comm_ring_1}) oo b = fps_const (a)"
   by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
 
@@ -971,11 +971,11 @@
 subsubsection {* Rule 1 *}
   (* {a_{n+k}}_0^infty Corresponds to (f - setsum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
 
-lemma fps_power_mult_eq_shift: 
+lemma fps_power_mult_eq_shift:
   "X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) = Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: field) * X^i) {0 .. k}" (is "?lhs = ?rhs")
 proof-
   {fix n:: nat
-    have "?lhs $ n = (if n < Suc k then 0 else a n)" 
+    have "?lhs $ n = (if n < Suc k then 0 else a n)"
       unfolding X_power_mult_nth by auto
     also have "\<dots> = ?rhs $ n"
     proof(induct k)
@@ -985,7 +985,7 @@
       note th = Suc.hyps[symmetric]
       have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: field) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: field) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps)
       also  have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
-	using th 
+	using th
 	unfolding fps_sub_nth by simp
       also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
 	unfolding X_power_mult_right_nth
@@ -1001,7 +1001,7 @@
 subsubsection{* Rule 2*}
 
   (* We can not reach the form of Wilf, but still near to it using rewrite rules*)
-  (* If f reprents {a_n} and P is a polynomial, then 
+  (* If f reprents {a_n} and P is a polynomial, then
         P(xD) f represents {P(n) a_n}*)
 
 definition "XD = op * X o fps_deriv"
@@ -1033,18 +1033,18 @@
   let ?sa = "Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
   have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)" by simp
   {fix n:: nat
-    {assume "n=0" hence "a$n = ((1 - ?X) * ?sa) $ n" 
+    {assume "n=0" hence "a$n = ((1 - ?X) * ?sa) $ n"
 	by (simp add: fps_mult_nth)}
     moreover
     {assume n0: "n \<noteq> 0"
       then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1}\<union>{2..n} = {1..n}"
 	"{0..n - 1}\<union>{n} = {0..n}"
 	apply (simp_all add: expand_set_eq) by presburger+
-      have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}" 
+      have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}"
 	"{0..n - 1}\<inter>{n} ={}" using n0
 	by (simp_all add: expand_set_eq, presburger+)
-      have f: "finite {0}" "finite {1}" "finite {2 .. n}" 
-	"finite {0 .. n - 1}" "finite {n}" by simp_all 
+      have f: "finite {0}" "finite {1}" "finite {2 .. n}"
+	"finite {0 .. n - 1}" "finite {n}" by simp_all
     have "((1 - ?X) * ?sa) $ n = setsum (\<lambda>i. (1 - ?X)$ i * ?sa $ (n - i)) {0 .. n}"
       by (simp add: fps_mult_nth)
     also have "\<dots> = a$n" unfolding th0
@@ -1055,7 +1055,7 @@
       by simp
     finally have "a$n = ((1 - ?X) * ?sa) $ n" by simp}
   ultimately have "a$n = ((1 - ?X) * ?sa) $ n" by blast}
-then show ?thesis 
+then show ?thesis
   unfolding fps_eq_iff by blast
 qed
 
@@ -1072,7 +1072,7 @@
   finally show ?thesis by (simp add: inverse_mult_eq_1[OF th0])
 qed
 
-subsubsection{* Rule 4 in its more general form: generalizes Rule 3 for an arbitrary 
+subsubsection{* Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
   finite product of FPS, also the relvant instance of powers of a FPS*}
 
 definition "natpermute n k = {l:: nat list. length l = k \<and> foldl op + 0 l = n}"
@@ -1082,7 +1082,7 @@
   apply (case_tac x, auto)
   done
 
-lemma foldl_add_start0: 
+lemma foldl_add_start0:
   "foldl op + x xs = x + foldl op + (0::nat) xs"
   apply (induct xs arbitrary: x)
   apply simp
@@ -1118,7 +1118,7 @@
     apply (subst foldl_add_start0)    by simp
   also have "\<dots> = x + a + setsum (op ! as) {0..<length as}" unfolding Cons.hyps by simp
   also have "\<dots> = x + setsum (op ! (a#as)) {0..<length (a#as)}"
-    unfolding eq[symmetric] 
+    unfolding eq[symmetric]
     unfolding setsum_Un_disjoint[OF f d, unfolded seq]
     by simp
   finally show ?case  .
@@ -1139,16 +1139,16 @@
   assumes mn: "h \<le> k"
   shows "natpermute n k = (\<Union>m \<in>{0..n}. {l1 @ l2 |l1 l2. l1 \<in> natpermute m h \<and> l2 \<in> natpermute (n - m) (k - h)})" (is "?L = ?R" is "?L = (\<Union>m \<in>{0..n}. ?S m)")
 proof-
-  {fix l assume l: "l \<in> ?R" 
+  {fix l assume l: "l \<in> ?R"
     from l obtain m xs ys where h: "m \<in> {0..n}" and xs: "xs \<in> natpermute m h" and ys: "ys \<in> natpermute (n - m) (k - h)"  and leq: "l = xs@ys" by blast
     from xs have xs': "foldl op + 0 xs = m" by (simp add: natpermute_def)
     from ys have ys': "foldl op + 0 ys = n - m" by (simp add: natpermute_def)
-    have "l \<in> ?L" using leq xs ys h 
+    have "l \<in> ?L" using leq xs ys h
       apply simp
       apply (clarsimp simp add: natpermute_def simp del: foldl_append)
       apply (simp add: foldl_add_append[unfolded foldl_append])
       unfolding xs' ys'
-      using mn xs ys 
+      using mn xs ys
       unfolding natpermute_def by simp}
   moreover
   {fix l assume l: "l \<in> natpermute n k"
@@ -1156,16 +1156,16 @@
     let ?ys = "drop h l"
     let ?m = "foldl op + 0 ?xs"
     from l have ls: "foldl op + 0 (?xs @ ?ys) = n" by (simp add: natpermute_def)
-    have xs: "?xs \<in> natpermute ?m h" using l mn by (simp add: natpermute_def)     
+    have xs: "?xs \<in> natpermute ?m h" using l mn by (simp add: natpermute_def)
     have ys: "?ys \<in> natpermute (n - ?m) (k - h)" using l mn ls[unfolded foldl_add_append]
       by (simp add: natpermute_def)
     from ls have m: "?m \<in> {0..n}"  unfolding foldl_add_append by simp
-    from xs ys ls have "l \<in> ?R" 
+    from xs ys ls have "l \<in> ?R"
       apply auto
       apply (rule bexI[where x = "?m"])
       apply (rule exI[where x = "?xs"])
       apply (rule exI[where x = "?ys"])
-      using ls l unfolding foldl_add_append 
+      using ls l unfolding foldl_add_append
       by (auto simp add: natpermute_def)}
   ultimately show ?thesis by blast
 qed
@@ -1179,7 +1179,7 @@
 
 lemma natpermute_finite: "finite (natpermute n k)"
 proof(induct k arbitrary: n)
-  case 0 thus ?case 
+  case 0 thus ?case
     apply (subst natpermute_split[of 0 0, simplified])
     by (simp add: natpermute_0)
 next
@@ -1203,7 +1203,7 @@
 proof-
   {fix xs assume H: "xs \<in> natpermute n (k+1)" and n: "n \<in> set xs"
     from n obtain i where i: "i \<in> {0.. k}" "xs!i = n" using H
-      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def) 
+      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
     have eqs: "({0..k} - {i}) \<union> {i} = {0..k}" using i by auto
     have f: "finite({0..k} - {i})" "finite {i}" by auto
     have d: "({0..k} - {i}) \<inter> {i} = {}" using i by auto
@@ -1235,14 +1235,14 @@
     also have "\<dots> = setsum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
       apply (rule setsum_cong2) by (simp del: replicate.simps)
     also have "\<dots> = n" using i by (simp add: setsum_delta)
-    finally 
+    finally
     have "?xs \<in> natpermute n (k+1)" using xsl unfolding natpermute_def Collect_def mem_def
       by blast
     then have "?xs \<in> ?A"  using nxs  by blast}
   ultimately show ?thesis by auto
 qed
 
-    (* The general form *)	
+    (* The general form *)
 lemma fps_setprod_nth:
   fixes m :: nat and a :: "nat \<Rightarrow> ('a::comm_ring_1) fps"
   shows "(setprod a {0 .. m})$n = setsum (\<lambda>v. setprod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
@@ -1266,7 +1266,7 @@
       apply (simp add: k)
       unfolding natpermute_split[of m "m + 1", simplified, of n, unfolded natlist_trivial_1[unfolded One_nat_def] k]
       apply (subst setsum_UN_disjoint)
-      apply simp 
+      apply simp
       apply simp
       unfolding image_Collect[symmetric]
       apply clarsimp
@@ -1305,7 +1305,7 @@
   shows "(a ^m)$n = (if m=0 then 1$n else setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
   by (cases m, simp_all add: fps_power_nth_Suc del: power_Suc)
 
-lemma fps_nth_power_0: 
+lemma fps_nth_power_0:
   fixes m :: nat and a :: "('a::{comm_ring_1, recpower}) fps"
   shows "(a ^m)$0 = (a$0) ^ m"
 proof-
@@ -1323,7 +1323,7 @@
  ultimately show ?thesis by (cases m, auto)
 qed
 
-lemma fps_compose_inj_right: 
+lemma fps_compose_inj_right:
   assumes a0: "a$0 = (0::'a::{recpower,idom})"
   and a1: "a$1 \<noteq> 0"
   shows "(b oo a = c oo a) \<longleftrightarrow> b = c" (is "?lhs \<longleftrightarrow>?rhs")
@@ -1331,7 +1331,7 @@
   {assume ?rhs then have "?lhs" by simp}
   moreover
   {assume h: ?lhs
-    {fix n have "b$n = c$n" 
+    {fix n have "b$n = c$n"
       proof(induct n rule: nat_less_induct)
 	fix n assume H: "\<forall>m<n. b$m = c$m"
 	{assume n0: "n=0"
@@ -1393,9 +1393,9 @@
 	unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
 	by simp
       finally have False using c' by simp}
-    then show "((r,Suc k,a,xs!i), r, Suc k, a, Suc n) \<in> ?R" 
+    then show "((r,Suc k,a,xs!i), r, Suc k, a, Suc n) \<in> ?R"
       apply auto by (metis not_less)}
-  {fix r k a n 
+  {fix r k a n
     show "((r,Suc k, a, 0),r, Suc k, a, Suc n) \<in> ?R" by simp}
 qed
 
@@ -1407,13 +1407,13 @@
 lemma fps_radical_nth_0[simp]: "fps_radical r n a $ 0 = (if n=0 then 1 else r n (a$0))"
   by (cases n, simp_all add: fps_radical_def)
 
-lemma fps_radical_power_nth[simp]: 
+lemma fps_radical_power_nth[simp]:
   assumes r: "(r k (a$0)) ^ k = a$0"
   shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)"
 proof-
   {assume "k=0" hence ?thesis by simp }
   moreover
-  {fix h assume h: "k = Suc h" 
+  {fix h assume h: "k = Suc h"
     have fh: "finite {0..h}" by simp
     have eq1: "fps_radical r k a ^ k $ 0 = (\<Prod>j\<in>{0..h}. fps_radical r k a $ (replicate k 0) ! j)"
       unfolding fps_power_nth h by simp
@@ -1427,9 +1427,9 @@
       unfolding setprod_constant[OF fh] using r by (simp add: h)
     finally have ?thesis using h by simp}
   ultimately show ?thesis by (cases k, auto)
-qed 
+qed
 
-lemma natpermute_max_card: assumes n0: "n\<noteq>0" 
+lemma natpermute_max_card: assumes n0: "n\<noteq>0"
   shows "card {xs \<in> natpermute n (k+1). n \<in> set xs} = k+1"
   unfolding natpermute_contain_maximal
 proof-
@@ -1448,14 +1448,14 @@
     then show "{replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
       by auto
   qed
-  from card_UN_disjoint[OF fK fAK d] 
+  from card_UN_disjoint[OF fK fAK d]
   show "card (\<Union>i\<in>{0..k}. {replicate (k + 1) 0[i := n]}) = k+1" by simp
 qed
-  
-lemma power_radical: 
+
+lemma power_radical:
   fixes a:: "'a ::{field, ring_char_0, recpower} fps"
   assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \<noteq> 0"
-  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a" 
+  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a"
 proof-
   let ?r = "fps_radical r (Suc k) a"
   from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
@@ -1473,11 +1473,11 @@
 	let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
 	have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
 	have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-	have f: "finite ?Pnkn" "finite ?Pnknn" 
+	have f: "finite ?Pnkn" "finite ?Pnknn"
 	  using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
 	  by (metis natpermute_finite)+
 	let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
-	have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn" 
+	have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
 	proof(rule setsum_cong2)
 	  fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
 	  let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
@@ -1490,12 +1490,12 @@
 	    unfolding setprod_gen_delta[OF fK] using i r0 by simp
 	  finally show ?ths .
 	qed
-	then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"  
-	  by (simp add: natpermute_max_card[OF nz, simplified]) 
+	then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+	  by (simp add: natpermute_max_card[OF nz, simplified])
 	also have "\<dots> = a$n - setsum ?f ?Pnknn"
 	  unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
 	finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
-	have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn" 
+	have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
 	  unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
 	also have "\<dots> = a$n" unfolding fn by simp
 	finally have "?r ^ Suc k $ n = a $n" .}
@@ -1505,15 +1505,15 @@
 qed
 
 lemma eq_divide_imp': assumes c0: "(c::'a::field) ~= 0" and eq: "a * c = b"
-  shows "a = b / c" 
+  shows "a = b / c"
 proof-
   from eq have "a * c * inverse c = b * inverse c" by simp
   hence "a * (inverse c * c) = b/c" by (simp only: field_simps divide_inverse)
   then show "a = b/c" unfolding  field_inverse[OF c0] by simp
 qed
 
-lemma radical_unique:  
-  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0" 
+lemma radical_unique:
+  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0"
   and a0: "r (Suc k) (b$0 ::'a::{field, ring_char_0, recpower}) = a$0" and b0: "b$0 \<noteq> 0"
   shows "a^(Suc k) = b \<longleftrightarrow> a = fps_radical r (Suc k) b"
 proof-
@@ -1540,12 +1540,12 @@
 	let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
 	have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
 	have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
-	have f: "finite ?Pnkn" "finite ?Pnknn" 
+	have f: "finite ?Pnkn" "finite ?Pnknn"
 	  using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
 	  by (metis natpermute_finite)+
 	let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
 	let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
-	have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn" 
+	have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
 	proof(rule setsum_cong2)
 	  fix v assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
 	  let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
@@ -1558,7 +1558,7 @@
 	    unfolding  setprod_gen_delta[OF fK] using i by simp
 	  finally show ?ths .
 	qed
-	then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"  
+	then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
 	  by (simp add: natpermute_max_card[OF nz, simplified])
 	have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
 	proof (rule setsum_cong2, rule setprod_cong, simp)
@@ -1578,20 +1578,20 @@
 	      by simp
 	    finally have False using c' by simp}
 	  then have thn: "xs!i < n" by arith
-	  from h[rule_format, OF thn]  
+	  from h[rule_format, OF thn]
 	  show "a$(xs !i) = ?r$(xs!i)" .
 	qed
 	have th00: "\<And>(x::'a). of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
 	  by (simp add: field_simps del: of_nat_Suc)
 	from H have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff)
 	also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
-	  unfolding fps_power_nth_Suc 
-	  using setsum_Un_disjoint[OF f d, unfolded Suc_plus1[symmetric], 
+	  unfolding fps_power_nth_Suc
+	  using setsum_Un_disjoint[OF f d, unfolded Suc_plus1[symmetric],
 	    unfolded eq, of ?g] by simp
 	also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn" unfolding th0 th1 ..
 	finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - setsum ?f ?Pnknn" by simp
 	then have "a$n = (b$n - setsum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
-	  apply - 
+	  apply -
 	  apply (rule eq_divide_imp')
 	  using r00
 	  apply (simp del: of_nat_Suc)
@@ -1607,8 +1607,8 @@
 qed
 
 
-lemma radical_power: 
-  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0" 
+lemma radical_power:
+  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0"
   and a0: "(a$0 ::'a::{field, ring_char_0, recpower}) \<noteq> 0"
   shows "(fps_radical r (Suc k) (a ^ Suc k)) = a"
 proof-
@@ -1620,7 +1620,7 @@
   from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis by metis
 qed
 
-lemma fps_deriv_radical: 
+lemma fps_deriv_radical:
   fixes a:: "'a ::{field, ring_char_0, recpower} fps"
   assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \<noteq> 0"
   shows "fps_deriv (fps_radical r (Suc k) a) = fps_deriv a / (fps_const (of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)"
@@ -1638,13 +1638,13 @@
   hence "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a" by simp
   hence "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w"
     by (simp add: fps_divide_def)
-  then show ?thesis unfolding th0 by simp 
+  then show ?thesis unfolding th0 by simp
 qed
 
-lemma radical_mult_distrib: 
+lemma radical_mult_distrib:
   fixes a:: "'a ::{field, ring_char_0, recpower} fps"
-  assumes 
-  ra0: "r (k) (a $ 0) ^ k = a $ 0" 
+  assumes
+  ra0: "r (k) (a $ 0) ^ k = a $ 0"
   and rb0: "r (k) (b $ 0) ^ k = b $ 0"
   and r0': "r (k) ((a * b) $ 0) = r (k) (a $ 0) * r (k) (b $ 0)"
   and a0: "a$0 \<noteq> 0"
@@ -1658,10 +1658,10 @@
   {fix h assume k: "k = Suc h"
   let ?ra = "fps_radical r (Suc h) a"
   let ?rb = "fps_radical r (Suc h) b"
-  have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0" 
+  have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
     using r0' k by (simp add: fps_mult_nth)
   have ab0: "(a*b) $ 0 \<noteq> 0" using a0 b0 by (simp add: fps_mult_nth)
-  from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric] 
+  from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric]
     power_radical[of r, OF ra0[unfolded k] a0] power_radical[of r, OF rb0[unfolded k] b0] k
   have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)}
 ultimately show ?thesis by (cases k, auto)
@@ -1669,10 +1669,10 @@
 
 lemma radical_inverse:
   fixes a:: "'a ::{field, ring_char_0, recpower} fps"
-  assumes 
-  ra0: "r (k) (a $ 0) ^ k = a $ 0" 
+  assumes
+  ra0: "r (k) (a $ 0) ^ k = a $ 0"
   and ria0: "r (k) (inverse (a $ 0)) = inverse (r (k) (a $ 0))"
-  and r1: "(r (k) 1) = 1" 
+  and r1: "(r (k) 1) = 1"
   and a0: "a$0 \<noteq> 0"
   shows "fps_radical r (k) (inverse a) = inverse (fps_radical r (k) a)"
 proof-
@@ -1686,7 +1686,7 @@
     using ria0 ra0 a0
     by (simp add: fps_inverse_def  nonzero_power_inverse[OF th00, symmetric]
              del: power_Suc)
-  from inverse_mult_eq_1[OF a0] have th0: "a * inverse a = 1" 
+  from inverse_mult_eq_1[OF a0] have th0: "a * inverse a = 1"
     by (simp add: mult_commute)
   from radical_unique[where a=1 and b=1 and r=r and k=h, simplified, OF r1[unfolded k]]
   have th01: "fps_radical r (Suc h) 1 = 1" .
@@ -1711,13 +1711,13 @@
 
 lemma radical_divide:
   fixes a:: "'a ::{field, ring_char_0, recpower} fps"
-  assumes 
-      ra0: "r k (a $ 0) ^ k = a $ 0" 
+  assumes
+      ra0: "r k (a $ 0) ^ k = a $ 0"
   and rb0: "r k (b $ 0) ^ k = b $ 0"
   and r1: "r k 1 = 1"
-  and rb0': "r k (inverse (b $ 0)) = inverse (r k (b $ 0))" 
+  and rb0': "r k (inverse (b $ 0)) = inverse (r k (b $ 0))"
   and raib': "r k (a$0 / (b$0)) = r k (a$0) / r k (b$0)"
-  and a0: "a$0 \<noteq> 0" 
+  and a0: "a$0 \<noteq> 0"
   and b0: "b$0 \<noteq> 0"
   shows "fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b"
 proof-
@@ -1730,12 +1730,12 @@
   {assume k0: "k\<noteq> 0"
     from b0 k0 rb0 have rbn0: "r k (b $0) \<noteq> 0"
       by (auto simp add: power_0_left)
-    
+
     from rb0 rb0' have rib0: "(r k (inverse (b $ 0)))^k = inverse (b$0)"
     by (simp add: nonzero_power_inverse[OF rbn0, symmetric])
   from rib0 have th0: "r k (inverse b $ 0) ^ k = inverse b $ 0"
     by (simp add:fps_inverse_def b0)
-  from raib 
+  from raib
   have th1: "r k ((a * inverse b) $ 0) = r k (a $ 0) * r k (inverse b $ 0)"
     by (simp add: divide_inverse fps_inverse_def  b0 fps_mult_nth)
   from nonzero_imp_inverse_nonzero[OF b0] b0 have th2: "inverse b $ 0 \<noteq> 0"
@@ -1750,7 +1750,7 @@
 
 subsection{* Derivative of composition *}
 
-lemma fps_compose_deriv: 
+lemma fps_compose_deriv:
   fixes a:: "('a::idom) fps"
   assumes b0: "b$0 = 0"
   shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * (fps_deriv b)"
@@ -1773,7 +1773,7 @@
     apply (rule setsum_reindex_cong[where f="Suc"])
     by (auto simp add: mult_assoc)
   finally have th0: "(fps_deriv (a oo b))$n = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
-  
+
   have "(((fps_deriv a) oo b) * (fps_deriv b))$n = setsum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
     unfolding fps_mult_nth by (simp add: mult_ac)
   also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
@@ -1816,11 +1816,11 @@
 
 subsection{* Finite FPS (i.e. polynomials) and X *}
 lemma fps_poly_sum_X:
-  assumes z: "\<forall>i > n. a$i = (0::'a::comm_ring_1)" 
+  assumes z: "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
   shows "a = setsum (\<lambda>i. fps_const (a$i) * X^i) {0..n}" (is "a = ?r")
 proof-
   {fix i
-    have "a$i = ?r$i" 
+    have "a$i = ?r$i"
       unfolding fps_setsum_nth fps_mult_left_const_nth X_power_nth
       by (simp add: mult_delta_right setsum_delta' z)
   }
@@ -1841,10 +1841,10 @@
 proof-
   let ?i = "fps_inv a oo a"
   {fix n
-    have "?i $n = X$n" 
+    have "?i $n = X$n"
     proof(induct n rule: nat_less_induct)
       fix n assume h: "\<forall>m<n. ?i$m = X$m"
-      {assume "n=0" hence "?i $n = X$n" using a0 
+      {assume "n=0" hence "?i $n = X$n" using a0
 	  by (simp add: fps_compose_nth fps_inv_def)}
       moreover
       {fix n1 assume n1: "n = Suc n1"
@@ -1853,7 +1853,7 @@
                    del: power_Suc)
 	also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + (X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
 	  using a0 a1 n1 by (simp add: fps_inv_def)
-	also have "\<dots> = X$n" using n1 by simp 
+	also have "\<dots> = X$n" using n1 by simp
 	finally have "?i $ n = X$n" .}
       ultimately show "?i $ n = X$n" by (cases n, auto)
     qed}
@@ -1872,10 +1872,10 @@
 proof-
   let ?i = "fps_ginv b a oo a"
   {fix n
-    have "?i $n = b$n" 
+    have "?i $n = b$n"
     proof(induct n rule: nat_less_induct)
       fix n assume h: "\<forall>m<n. ?i$m = b$m"
-      {assume "n=0" hence "?i $n = b$n" using a0 
+      {assume "n=0" hence "?i $n = b$n" using a0
 	  by (simp add: fps_compose_nth fps_ginv_def)}
       moreover
       {fix n1 assume n1: "n = Suc n1"
@@ -1884,7 +1884,7 @@
                    del: power_Suc)
 	also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + (b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
 	  using a0 a1 n1 by (simp add: fps_ginv_def)
-	also have "\<dots> = b$n" using n1 by simp 
+	also have "\<dots> = b$n" using n1 by simp
 	finally have "?i $ n = b$n" .}
       ultimately show "?i $ n = b$n" by (cases n, auto)
     qed}
@@ -1927,10 +1927,10 @@
       show "setsum f (insert x F) oo a  = setsum (\<lambda>i. f i oo a) (insert x F)"
 	using fF xF h by (simp add: fps_compose_add_distrib)
     qed}
-  ultimately show ?thesis by blast 
+  ultimately show ?thesis by blast
 qed
 
-lemma convolution_eq: 
+lemma convolution_eq:
   "setsum (%i. a (i :: nat) * b (n - i)) {0 .. n} = setsum (%(i,j). a i * b j) {(i,j). i <= n \<and> j \<le> n \<and> i + j = n}"
   apply (rule setsum_reindex_cong[where f=fst])
   apply (clarsimp simp add: inj_on_def)
@@ -1946,8 +1946,8 @@
   shows "((a oo c) * (b oo d))$n = setsum (%(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}" (is "?l = ?r")
 proof-
   let ?S = "{(k\<Colon>nat, m\<Colon>nat). k + m \<le> n}"
-  have s: "?S \<subseteq> {0..n} <*> {0..n}" by (auto simp add: subset_eq)  
-  have f: "finite {(k\<Colon>nat, m\<Colon>nat). k + m \<le> n}" 
+  have s: "?S \<subseteq> {0..n} <*> {0..n}" by (auto simp add: subset_eq)
+  have f: "finite {(k\<Colon>nat, m\<Colon>nat). k + m \<le> n}"
     apply (rule finite_subset[OF s])
     by auto
   have "?r =  setsum (%i. setsum (%(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
@@ -1955,7 +1955,7 @@
     apply (subst setsum_commute)
     apply (rule setsum_cong2)
     by (auto simp add: ring_simps)
-  also have "\<dots> = ?l" 
+  also have "\<dots> = ?l"
     apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
     apply (rule setsum_cong2)
     apply (simp add: setsum_cartesian_product mult_assoc)
@@ -1998,9 +1998,9 @@
   apply simp
   apply arith
   done
-  
+
 
-lemma setsum_pair_less_iff: 
+lemma setsum_pair_less_iff:
   "setsum (%((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} = setsum (%s. setsum (%i. a i * b (s - i) * c s) {0..s}) {0..n}" (is "?l = ?r")
 proof-
   let ?KM=  "{(k,m). k + m \<le> n}"
@@ -2025,12 +2025,12 @@
   unfolding setsum_pair_less_iff[where a = "%k. a$k" and b="%m. b$m" and c="%s. (c ^ s)$n" and n = n] ..
 
 
-lemma fps_compose_mult_distrib: 
+lemma fps_compose_mult_distrib:
   assumes c0: "c$0 = (0::'a::idom)"
   shows "(a * b) oo c = (a oo c) * (b oo c)" (is "?l = ?r")
   apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma[OF c0])
   by (simp add: fps_compose_nth fps_mult_nth setsum_left_distrib)
-lemma fps_compose_setprod_distrib: 
+lemma fps_compose_setprod_distrib:
   assumes c0: "c$0 = (0::'a::idom)"
   shows "(setprod a S) oo c = setprod (%k. a k oo c) S" (is "?l = ?r")
   apply (cases "finite S")
@@ -2061,7 +2061,7 @@
   "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
   by (auto simp add: fps_const_mult_apply_left mult_commute)
 
-lemma fps_compose_assoc: 
+lemma fps_compose_assoc:
   assumes c0: "c$0 = (0::'a::idom)" and b0: "b$0 = 0"
   shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r")
 proof-
@@ -2088,7 +2088,7 @@
   moreover
   {fix h assume h: "k = Suc h"
     {fix n
-      {assume kn: "k>n" hence "?l $ n = ?r $n" using a0 startsby_zero_power_prefix[OF a0] h 
+      {assume kn: "k>n" hence "?l $ n = ?r $n" using a0 startsby_zero_power_prefix[OF a0] h
 	  by (simp add: fps_compose_nth del: power_Suc)}
       moreover
       {assume kn: "k \<le> n"
@@ -2106,14 +2106,14 @@
   let ?ia = "fps_inv a"
   let ?iaa = "a oo fps_inv a"
   have th0: "?ia $ 0 = 0" by (simp add: fps_inv_def)
-  have th1: "?iaa $ 0 = 0" using a0 a1 
+  have th1: "?iaa $ 0 = 0" using a0 a1
     by (simp add: fps_inv_def fps_compose_nth)
   have th2: "X$0 = 0" by simp
   from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo X" by simp
   then have "(a oo fps_inv a) oo a = X oo a"
     by (simp add: fps_compose_assoc[OF a0 th0] X_fps_compose_startby0[OF a0])
   with fps_compose_inj_right[OF a0 a1]
-  show ?thesis by simp 
+  show ?thesis by simp
 qed
 
 lemma fps_inv_deriv:
@@ -2135,7 +2135,7 @@
 subsection{* Elementary series *}
 
 subsubsection{* Exponential series *}
-definition "E x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"   
+definition "E x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"
 
 lemma E_deriv[simp]: "fps_deriv (E a) = fps_const (a::'a::{field, recpower, ring_char_0}) * E a" (is "?l = ?r")
 proof-
@@ -2146,27 +2146,27 @@
 then show ?thesis by (simp add: fps_eq_iff)
 qed
 
-lemma E_unique_ODE: 
+lemma E_unique_ODE:
   "fps_deriv a = fps_const c * a \<longleftrightarrow> a = fps_const (a$0) * E (c :: 'a::{field, ring_char_0, recpower})"
   (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   {assume d: ?lhs
-  from d have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)" 
+  from d have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
     by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
   {fix n have "a$n = a$0 * c ^ n/ (of_nat (fact n))"
       apply (induct n)
       apply simp
-      unfolding th 
+      unfolding th
       using fact_gt_zero
       apply (simp add: field_simps del: of_nat_Suc fact.simps)
       apply (drule sym)
       by (simp add: ring_simps of_nat_mult power_Suc)}
   note th' = this
-  have ?rhs 
+  have ?rhs
     by (auto simp add: fps_eq_iff fps_const_mult_left E_def intro : th')}
 moreover
 {assume h: ?rhs
-  have ?lhs 
+  have ?lhs
     apply (subst h)
     apply simp
     apply (simp only: h[symmetric])
@@ -2197,13 +2197,13 @@
   from fps_inverse_unique[OF th1 th0] show ?thesis by simp
 qed
 
-lemma E_nth_deriv[simp]: "fps_nth_deriv n (E (a::'a::{field, recpower, ring_char_0})) = (fps_const a)^n * (E a)"  
+lemma E_nth_deriv[simp]: "fps_nth_deriv n (E (a::'a::{field, recpower, ring_char_0})) = (fps_const a)^n * (E a)"
   by (induct n, auto simp add: power_Suc)
 
 lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
   by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_negf[symmetric])
 
-lemma fps_compose_sub_distrib: 
+lemma fps_compose_sub_distrib:
   shows "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
   unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
 
@@ -2213,8 +2213,8 @@
 lemma X_compose_E[simp]: "X oo E (a::'a::{field, recpower}) = E a - 1"
   by (simp add: fps_eq_iff X_fps_compose)
 
-lemma LE_compose: 
-  assumes a: "a\<noteq>0" 
+lemma LE_compose:
+  assumes a: "a\<noteq>0"
   shows "fps_inv (E a - 1) oo (E a - 1) = X"
   and "(E a - 1) oo fps_inv (E a - 1) = X"
 proof-
@@ -2226,12 +2226,12 @@
 qed
 
 
-lemma fps_const_inverse: 
+lemma fps_const_inverse:
   "inverse (fps_const (a::'a::{field, division_by_zero})) = fps_const (inverse a)"
   apply (auto simp add: fps_eq_iff fps_inverse_def) by (case_tac "n", auto)
 
 
-lemma inverse_one_plus_X: 
+lemma inverse_one_plus_X:
   "inverse (1 + X) = Abs_fps (\<lambda>n. (- 1 ::'a::{field, recpower})^n)"
   (is "inverse ?l = ?r")
 proof-
@@ -2246,8 +2246,8 @@
 lemma E_power_mult: "(E (c::'a::{field,recpower,ring_char_0}))^n = E (of_nat n * c)"
   by (induct n, auto simp add: ring_simps E_add_mult power_Suc)
 
-subsubsection{* Logarithmic series *}  
-definition "(L::'a::{field, ring_char_0,recpower} fps) 
+subsubsection{* Logarithmic series *}
+definition "(L::'a::{field, ring_char_0,recpower} fps)
   = Abs_fps (\<lambda>n. (- 1) ^ Suc n / of_nat n)"
 
 lemma fps_deriv_L: "fps_deriv L = inverse (1 + X)"
@@ -2258,7 +2258,7 @@
   by (simp add: L_def)
 
 lemma L_E_inv:
-  assumes a: "a\<noteq> (0::'a::{field,division_by_zero,ring_char_0,recpower})" 
+  assumes a: "a\<noteq> (0::'a::{field,division_by_zero,ring_char_0,recpower})"
   shows "L = fps_const a * fps_inv (E a - 1)" (is "?l = ?r")
 proof-
   let ?b = "E a - 1"
@@ -2274,7 +2274,7 @@
     by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult)
   hence "fps_deriv (fps_const a * fps_inv ?b) = inverse (X + 1)"
     using a by (simp add: fps_divide_def field_simps)
-  hence "fps_deriv ?l = fps_deriv ?r" 
+  hence "fps_deriv ?l = fps_deriv ?r"
     by (simp add: fps_deriv_L add_commute)
   then show ?thesis unfolding fps_deriv_eq_iff
     by (simp add: L_nth fps_inv_def)
@@ -2282,33 +2282,33 @@
 
 subsubsection{* Formal trigonometric functions  *}
 
-definition "fps_sin (c::'a::{field, recpower, ring_char_0}) = 
+definition "fps_sin (c::'a::{field, recpower, ring_char_0}) =
   Abs_fps (\<lambda>n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))"
 
 definition "fps_cos (c::'a::{field, recpower, ring_char_0}) = Abs_fps (\<lambda>n. if even n then (- 1) ^ (n div 2) * c^n / (of_nat (fact n)) else 0)"
 
-lemma fps_sin_deriv: 
+lemma fps_sin_deriv:
   "fps_deriv (fps_sin c) = fps_const c * fps_cos c"
   (is "?lhs = ?rhs")
 proof-
   {fix n::nat
     {assume en: "even n"
       have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
-      also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))" 
+      also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))"
 	using en by (simp add: fps_sin_def)
       also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
 	unfolding fact_Suc of_nat_mult
 	by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
 	by (simp add: field_simps del: of_nat_add of_nat_Suc)
-      finally have "?lhs $n = ?rhs$n" using en 
+      finally have "?lhs $n = ?rhs$n" using en
 	by (simp add: fps_cos_def ring_simps power_Suc )}
-    then have "?lhs $ n = ?rhs $ n" 
+    then have "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def fps_cos_def) }
   then show ?thesis by (auto simp add: fps_eq_iff)
 qed
 
-lemma fps_cos_deriv: 
+lemma fps_cos_deriv:
   "fps_deriv (fps_cos c) = fps_const (- c)* (fps_sin c)"
   (is "?lhs = ?rhs")
 proof-
@@ -2318,7 +2318,7 @@
     {assume en: "odd n"
       from en have n0: "n \<noteq>0 " by presburger
       have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
-      also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))" 
+      also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))"
 	using en by (simp add: fps_cos_def)
       also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
 	unfolding fact_Suc of_nat_mult
@@ -2327,10 +2327,10 @@
 	by (simp add: field_simps del: of_nat_add of_nat_Suc)
       also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
 	unfolding th0 unfolding th1[OF en] by simp
-      finally have "?lhs $n = ?rhs$n" using en 
+      finally have "?lhs $n = ?rhs$n" using en
 	by (simp add: fps_sin_def ring_simps power_Suc)}
-    then have "?lhs $ n = ?rhs $ n" 
-      by (cases "even n", simp_all add: fps_deriv_def fps_sin_def 
+    then have "?lhs $ n = ?rhs $ n"
+      by (cases "even n", simp_all add: fps_deriv_def fps_sin_def
 	fps_cos_def) }
   then show ?thesis by (auto simp add: fps_eq_iff)
 qed
@@ -2353,7 +2353,7 @@
 lemma fps_tan_deriv: "fps_deriv(fps_tan c) = fps_const c/ (fps_cos c ^ 2)"
 proof-
   have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
-  show ?thesis 
+  show ?thesis
     using fps_sin_cos_sum_of_squares[of c]
     apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv add: fps_const_neg[symmetric] ring_simps power2_eq_square del: fps_const_neg)
     unfolding right_distrib[symmetric]
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -18,12 +18,12 @@
 proof-
   obtain x y where xy: "z = Complex x y" by (cases z)
   {assume y0: "y = 0"
-    {assume x0: "x \<ge> 0" 
+    {assume x0: "x \<ge> 0"
       then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
 	by (simp add: csqrt_def power2_eq_square)}
     moreover
     {assume "\<not> x \<ge> 0" hence x0: "- x \<ge> 0" by arith
-      then have ?thesis using y0 xy real_sqrt_pow2[OF x0] 
+      then have ?thesis using y0 xy real_sqrt_pow2[OF x0]
 	by (simp add: csqrt_def power2_eq_square) }
     ultimately have ?thesis by blast}
   moreover
@@ -31,20 +31,20 @@
     {fix x y
       let ?z = "Complex x y"
       from abs_Re_le_cmod[of ?z] have tha: "abs x \<le> cmod ?z" by auto
-      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+ 
+      hence "cmod ?z - x \<ge> 0" "cmod ?z + x \<ge> 0" by arith+
       hence "(sqrt (x * x + y * y) + x) / 2 \<ge> 0" "(sqrt (x * x + y * y) - x) / 2 \<ge> 0" by (simp_all add: power2_eq_square) }
     note th = this
-    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2" 
-      by (simp add: power2_eq_square) 
+    have sq4: "\<And>x::real. x^2 / 4 = (x / 2) ^ 2"
+      by (simp add: power2_eq_square)
     from th[of x y]
     have sq4': "sqrt (((sqrt (x * x + y * y) + x)^2 / 4)) = (sqrt (x * x + y * y) + x) / 2" "sqrt (((sqrt (x * x + y * y) - x)^2 / 4)) = (sqrt (x * x + y * y) - x) / 2" unfolding sq4 by simp_all
     then have th1: "sqrt ((sqrt (x * x + y * y) + x) * (sqrt (x * x + y * y) + x) / 4) - sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) - x) / 4) = x"
-      unfolding power2_eq_square by simp 
-    have "sqrt 4 = sqrt (2^2)" by simp 
+      unfolding power2_eq_square by simp
+    have "sqrt 4 = sqrt (2^2)" by simp
     hence sqrt4: "sqrt 4 = 2" by (simp only: real_sqrt_abs)
     have th2: "2 *(y * sqrt ((sqrt (x * x + y * y) - x) * (sqrt (x * x + y * y) + x) / 4)) / \<bar>y\<bar> = y"
       using iffD2[OF real_sqrt_pow2_iff sum_power2_ge_zero[of x y]] y0
-      unfolding power2_eq_square 
+      unfolding power2_eq_square
       by (simp add: algebra_simps real_sqrt_divide sqrt4)
      from y0 xy have ?thesis  apply (simp add: csqrt_def power2_eq_square)
        apply (simp add: real_sqrt_sum_squares_mult_ge_zero[of x y] real_sqrt_pow2[OF th(1)[of x y], unfolded power2_eq_square] real_sqrt_pow2[OF th(2)[of x y], unfolded power2_eq_square] real_sqrt_mult[symmetric])
@@ -71,7 +71,7 @@
 lemma poly_bound_exists:
   shows "\<exists>m. m > 0 \<and> (\<forall>z. cmod z <= r \<longrightarrow> cmod (poly p z) \<le> m)"
 proof(induct p)
-  case 0 thus ?case by (rule exI[where x=1], simp) 
+  case 0 thus ?case by (rule exI[where x=1], simp)
 next
   case (pCons c cs)
   from pCons.hyps obtain m where m: "\<forall>z. cmod z \<le> r \<longrightarrow> cmod (poly cs z) \<le> m"
@@ -156,14 +156,14 @@
 proof-
   from ex bz obtain x Y where x: "P x" and Y: "\<And>x. P x \<Longrightarrow> x < Y"  by blast
   from ex have thx:"\<exists>x. x \<in> Collect P" by blast
-  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y" 
+  from bz have thY: "\<exists>Y. isUb UNIV (Collect P) Y"
     by(auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def order_le_less)
   from reals_complete[OF thx thY] obtain L where L: "isLub UNIV (Collect P) L"
     by blast
   from Y[OF x] have xY: "x < Y" .
-  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)  
-  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y" 
-    apply (clarsimp, atomize (full)) by auto 
+  from L have L': "\<forall>x. P x \<longrightarrow> x \<le> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
+  from Y have Y': "\<forall>x. P x \<longrightarrow> x \<le> Y"
+    apply (clarsimp, atomize (full)) by auto
   from L Y' have "L \<le> Y" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
   {fix y
     {fix z assume z: "P z" "y < z"
@@ -171,7 +171,7 @@
     moreover
     {assume yL: "y < L" "\<forall>z. P z \<longrightarrow> \<not> y < z"
       hence nox: "\<forall>z. P z \<longrightarrow> y \<ge> z" by auto
-      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def) 
+      from nox L have "y \<ge> L" by (auto simp add: isUb_def isLub_def setge_def setle_def leastP_def Ball_def)
       with yL(1) have False  by arith}
     ultimately have "(\<exists>x. P x \<and> y < x) \<longleftrightarrow> y < L" by blast}
   thus ?thesis by blast
@@ -190,7 +190,7 @@
     hence "abs (2*x) \<le> 1" "abs (2*y) \<le> 1" by simp_all
     hence "(abs (2 * x))^2 <= 1^2" "(abs (2 * y)) ^2 <= 1^2"
       by - (rule power_mono, simp, simp)+
-    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1" 
+    hence th0: "4*x^2 \<le> 1" "4*y^2 \<le> 1"
       by (simp_all  add: power2_abs power_mult_distrib)
     from add_mono[OF th0] xy have False by simp }
   thus ?thesis unfolding linorder_not_le[symmetric] by blast
@@ -216,21 +216,21 @@
   {assume o: "odd n"
     from b have b': "b^2 \<noteq> 0" unfolding power2_eq_square by simp
     have "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
-    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) = 
+    Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
     ((Re (inverse b))^2 + (Im (inverse b))^2) * \<bar>Im b * Im b + Re b * Re b\<bar>" by algebra
-    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2" 
+    also have "\<dots> = cmod (inverse b) ^2 * cmod b ^ 2"
       apply (simp add: cmod_def) using realpow_two_le_add_order[of "Re b" "Im b"]
       by (simp add: power2_eq_square)
-    finally 
+    finally
     have th0: "Im (inverse b) * (Im (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) +
     Re (inverse b) * (Re (inverse b) * \<bar>Im b * Im b + Re b * Re b\<bar>) =
-    1" 
+    1"
       apply (simp add: power2_eq_square norm_mult[symmetric] norm_inverse[symmetric])
       using right_inverse[OF b']
       by (simp add: power2_eq_square[symmetric] power_inverse[symmetric] algebra_simps)
     have th0: "cmod (complex_of_real (cmod b) / b) = 1"
       apply (simp add: complex_Re_mult cmod_def power2_eq_square Re_complex_of_real Im_complex_of_real divide_inverse algebra_simps )
-      by (simp add: real_sqrt_mult[symmetric] th0)        
+      by (simp add: real_sqrt_mult[symmetric] th0)
     from o have "\<exists>m. n = Suc (2*m)" by presburger+
     then obtain m where m: "n = Suc (2*m)" by blast
     from unimodular_reduce_norm[OF th0] o
@@ -246,7 +246,7 @@
     then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
     let ?w = "v / complex_of_real (root n (cmod b))"
     from odd_real_root_pow[OF o, of "cmod b"]
-    have th1: "?w ^ n = v^n / complex_of_real (cmod b)" 
+    have th1: "?w ^ n = v^n / complex_of_real (cmod b)"
       by (simp add: power_divide complex_of_real_power)
     have th2:"cmod (complex_of_real (cmod b) / b) = 1" using b by (simp add: norm_divide)
     hence th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0" by simp
@@ -257,7 +257,7 @@
       using b v by (simp add: th2)
 
     from mult_less_imp_less_left[OF th4 th3]
-    have "?P ?w n" unfolding th1 . 
+    have "?P ?w n" unfolding th1 .
     hence "\<exists>z. ?P z n" .. }
   ultimately show "\<exists>z. ?P z n" by blast
 qed
@@ -272,14 +272,14 @@
   assumes r: "\<forall>n. cmod (s n) \<le> r"
   shows "\<exists>f z. subseq f \<and> (\<forall>e >0. \<exists>N. \<forall>n \<ge> N. cmod (s (f n) - z) < e)"
 proof-
-  from seq_monosub[of "Re o s"] 
-  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))" 
+  from seq_monosub[of "Re o s"]
+  obtain f g where f: "subseq f" "monoseq (\<lambda>n. Re (s (f n)))"
     unfolding o_def by blast
-  from seq_monosub[of "Im o s o f"] 
-  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast  
+  from seq_monosub[of "Im o s o f"]
+  obtain g where g: "subseq g" "monoseq (\<lambda>n. Im (s(f(g n))))" unfolding o_def by blast
   let ?h = "f o g"
-  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith 
-  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>" 
+  from r[rule_format, of 0] have rp: "r \<ge> 0" using norm_ge_zero[of "s 0"] by arith
+  have th:"\<forall>n. r + 1 \<ge> \<bar> Re (s n)\<bar>"
   proof
     fix n
     from abs_Re_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
@@ -290,7 +290,7 @@
     apply (rule exI[where x= "r + 1"])
     using th rp apply simp
     using f(2) .
-  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>" 
+  have th:"\<forall>n. r + 1 \<ge> \<bar> Im (s n)\<bar>"
   proof
     fix n
     from abs_Im_le_cmod[of "s n"] r[rule_format, of n]  show "\<bar>Im (s n)\<bar> \<le> r + 1" by arith
@@ -303,17 +303,17 @@
     using th rp apply simp
     using g(2) .
 
-  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x" 
-    by blast 
-  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r" 
+  from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x"
+    by blast
+  hence  x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Re (s (f n)) - x \<bar> < r"
     unfolding LIMSEQ_def real_norm_def .
 
-  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y" 
-    by blast 
-  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r" 
+  from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y"
+    by blast
+  hence  y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar> Im (s (f (g n))) - y \<bar> < r"
     unfolding LIMSEQ_def real_norm_def .
   let ?w = "Complex x y"
-  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto 
+  from f(1) g(1) have hs: "subseq ?h" unfolding subseq_def by auto
   {fix e assume ep: "e > (0::real)"
     hence e2: "e/2 > 0" by simp
     from x[rule_format, OF e2] y[rule_format, OF e2]
@@ -321,16 +321,16 @@
     {fix n assume nN12: "n \<ge> N1 + N2"
       hence nN1: "g n \<ge> N1" and nN2: "n \<ge> N2" using seq_suble[OF g(1), of n] by arith+
       from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
-      have "cmod (s (?h n) - ?w) < e" 
+      have "cmod (s (?h n) - ?w) < e"
 	using metric_bound_lemma[of "s (f (g n))" ?w] by simp }
     hence "\<exists>N. \<forall>n\<ge>N. cmod (s (?h n) - ?w) < e" by blast }
-  with hs show ?thesis  by blast  
+  with hs show ?thesis  by blast
 qed
 
 text{* Polynomial is continuous. *}
 
 lemma poly_cont:
-  assumes ep: "e > 0" 
+  assumes ep: "e > 0"
   shows "\<exists>d >0. \<forall>w. 0 < cmod (w - z) \<and> cmod (w - z) < d \<longrightarrow> cmod (poly p w - poly p z) < e"
 proof-
   obtain q where q: "degree q = degree p" "\<And>x. poly q x = poly p (z + x)"
@@ -348,35 +348,35 @@
     case 0 thus ?case  using ep by auto
   next
     case (pCons c cs)
-    from poly_bound_exists[of 1 "cs"] 
+    from poly_bound_exists[of 1 "cs"]
     obtain m where m: "m > 0" "\<And>z. cmod z \<le> 1 \<Longrightarrow> cmod (poly cs z) \<le> m" by blast
     from ep m(1) have em0: "e/m > 0" by (simp add: field_simps)
     have one0: "1 > (0::real)"  by arith
-    from real_lbound_gt_zero[OF one0 em0] 
+    from real_lbound_gt_zero[OF one0 em0]
     obtain d where d: "d >0" "d < 1" "d < e / m" by blast
-    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e" 
+    from d(1,3) m(1) have dm: "d*m > 0" "d*m < e"
       by (simp_all add: field_simps real_mult_order)
-    show ?case 
+    show ?case
       proof(rule ex_forward[OF real_lbound_gt_zero[OF one0 em0]], clarsimp simp add: norm_mult)
 	fix d w
 	assume H: "d > 0" "d < 1" "d < e/m" "w\<noteq>z" "cmod (w-z) < d"
 	hence d1: "cmod (w-z) \<le> 1" "d \<ge> 0" by simp_all
 	from H(3) m(1) have dme: "d*m < e" by (simp add: field_simps)
-	from H have th: "cmod (w-z) \<le> d" by simp 
+	from H have th: "cmod (w-z) \<le> d" by simp
 	from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
 	show "cmod (w - z) * cmod (poly cs (w - z)) < e" by simp
-      qed  
+      qed
     qed
 qed
 
-text{* Hence a polynomial attains minimum on a closed disc 
+text{* Hence a polynomial attains minimum on a closed disc
   in the complex plane. *}
 lemma  poly_minimum_modulus_disc:
   "\<exists>z. \<forall>w. cmod w \<le> r \<longrightarrow> cmod (poly p z) \<le> cmod (poly p w)"
 proof-
   {assume "\<not> r \<ge> 0" hence ?thesis unfolding linorder_not_le
       apply -
-      apply (rule exI[where x=0]) 
+      apply (rule exI[where x=0])
       apply auto
       apply (subgoal_tac "cmod w < 0")
       apply simp
@@ -384,35 +384,35 @@
       done }
   moreover
   {assume rp: "r \<ge> 0"
-    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp 
+    from rp have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))" by simp
     hence mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"  by blast
     {fix x z
       assume H: "cmod z \<le> r" "cmod (poly p z) = - x" "\<not>x < 1"
       hence "- x < 0 " by arith
       with H(2) norm_ge_zero[of "poly p z"]  have False by simp }
     then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z" by blast
-    from real_sup_exists[OF mth1 mth2] obtain s where 
+    from real_sup_exists[OF mth1 mth2] obtain s where
       s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow>(y < s)" by blast
     let ?m = "-s"
     {fix y
-      from s[rule_format, of "-y"] have 
-    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" 
+      from s[rule_format, of "-y"] have
+    "(\<exists>z x. cmod z \<le> r \<and> -(- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y"
 	unfolding minus_less_iff[of y ] equation_minus_iff by blast }
     note s1 = this[unfolded minus_minus]
-    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m" 
+    from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
       by auto
     {fix n::nat
-      from s1[rule_format, of "?m + 1/real (Suc n)"] 
+      from s1[rule_format, of "?m + 1/real (Suc n)"]
       have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)"
 	by simp}
     hence th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
-    from choice[OF th] obtain g where 
-      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)" 
+    from choice[OF th] obtain g where
+      g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m+1 /real(Suc n)"
       by blast
-    from bolzano_weierstrass_complex_disc[OF g(1)] 
+    from bolzano_weierstrass_complex_disc[OF g(1)]
     obtain f z where fz: "subseq f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
-      by blast    
-    {fix w 
+      by blast
+    {fix w
       assume wr: "cmod w \<le> r"
       let ?e = "\<bar>cmod (poly p z) - ?m\<bar>"
       {assume e: "?e > 0"
@@ -423,8 +423,8 @@
 	  have "cmod(poly p w - poly p z) < ?e / 2"
 	    using d(2)[rule_format, of w] w e by (cases "w=z", simp_all)}
 	note th1 = this
-	
-	from fz(2)[rule_format, OF d(1)] obtain N1 where 
+
+	from fz(2)[rule_format, OF d(1)] obtain N1 where
 	  N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d" by blast
 	from reals_Archimedean2[of "2/?e"] obtain N2::nat where
 	  N2: "2/?e < real N2" by blast
@@ -434,13 +434,13 @@
 	have "a < e2 \<Longrightarrow> abs(b - m) < e2 \<Longrightarrow> 2 * e2 <= abs(b - m) + a
           ==> False" by arith}
       note th0 = this
-      have ath: 
+      have ath:
 	"\<And>m x e. m <= x \<Longrightarrow>  x < m + e ==> abs(x - m::real) < e" by arith
       from s1m[OF g(1)[rule_format]]
       have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
       from seq_suble[OF fz(1), of "N1+N2"]
       have th00: "real (Suc (N1+N2)) \<le> real (Suc (f (N1+N2)))" by simp
-      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"  
+      have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1+N2)) > 0"
 	using N2 by auto
       from frac_le[OF th000 th00] have th00: "?m +1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))" by simp
       from g(2)[rule_format, of "f (N1 + N2)"]
@@ -451,17 +451,17 @@
       with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
       have "?e/2 > 1/ real (Suc (N1 + N2))" by (simp add: inverse_eq_divide)
       with ath[OF th31 th32]
-      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith  
-      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c" 
+      have thc1:"\<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar>< ?e/2" by arith
+      have ath2: "\<And>(a::real) b c m. \<bar>a - b\<bar> <= c ==> \<bar>b - m\<bar> <= \<bar>a - m\<bar> + c"
 	by arith
       have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar>
-\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)" 
+\<le> cmod (poly p (g (f (N1 + N2))) - poly p z)"
 	by (simp add: norm_triangle_ineq3)
       from ath2[OF th22, of ?m]
       have thc2: "2*(?e/2) \<le> \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)" by simp
       from th0[OF th2 thc1 thc2] have False .}
       hence "?e = 0" by auto
-      then have "cmod (poly p z) = ?m" by simp  
+      then have "cmod (poly p z) = ?m" by simp
       with s1m[OF wr]
       have "cmod (poly p z) \<le> cmod (poly p w)" by simp }
     hence ?thesis by blast}
@@ -474,7 +474,7 @@
   apply (simp add: power2_eq_square[symmetric])
   done
 
-lemma cispi: "cis pi = -1" 
+lemma cispi: "cis pi = -1"
   unfolding cis_def
   by simp
 
@@ -491,7 +491,7 @@
   shows "\<exists>r. \<forall>z. r \<le> cmod z \<longrightarrow> d \<le> cmod (poly (pCons a p) z)"
 using ex
 proof(induct p arbitrary: a d)
-  case (pCons c cs a d) 
+  case (pCons c cs a d)
   {assume H: "cs \<noteq> 0"
     with pCons.hyps obtain r where r: "\<forall>z. r \<le> cmod z \<longrightarrow> d + cmod a \<le> cmod (poly (pCons c cs) z)" by blast
     let ?r = "1 + \<bar>r\<bar>"
@@ -504,8 +504,8 @@
       have th1: "d \<le> cmod(z * poly (pCons c cs) z) - cmod a"
 	unfolding norm_mult by (simp add: algebra_simps)
       from complex_mod_triangle_sub[of "z * poly (pCons c cs) z" a]
-      have th2: "cmod(z * poly (pCons c cs) z) - cmod a \<le> cmod (poly (pCons a (pCons c cs)) z)" 
-	by (simp add: diff_le_eq algebra_simps) 
+      have th2: "cmod(z * poly (pCons c cs) z) - cmod a \<le> cmod (poly (pCons a (pCons c cs)) z)"
+	by (simp add: diff_le_eq algebra_simps)
       from th1 th2 have "d \<le> cmod (poly (pCons a (pCons c cs)) z)"  by arith}
     hence ?case by blast}
   moreover
@@ -515,13 +515,13 @@
     {fix z
       assume h: "(\<bar>d\<bar> + cmod a) / cmod c \<le> cmod z"
       from c0 have "cmod c > 0" by simp
-      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)" 
+      from h c0 have th0: "\<bar>d\<bar> + cmod a \<le> cmod (z*c)"
 	by (simp add: field_simps norm_mult)
       have ath: "\<And>mzh mazh ma. mzh <= mazh + ma ==> abs(d) + ma <= mzh ==> d <= mazh" by arith
       from complex_mod_triangle_sub[of "z*c" a ]
       have th1: "cmod (z * c) \<le> cmod (a + z * c) + cmod a"
 	by (simp add: algebra_simps)
-      from ath[OF th1 th0] have "d \<le> cmod (poly (pCons a (pCons c cs)) z)" 
+      from ath[OF th1 th0] have "d \<le> cmod (poly (pCons a (pCons c cs)) z)"
         using cs0' by simp}
     then have ?case  by blast}
   ultimately show ?case by blast
@@ -531,15 +531,15 @@
 lemma poly_minimum_modulus:
   "\<exists>z.\<forall>w. cmod (poly p z) \<le> cmod (poly p w)"
 proof(induct p)
-  case (pCons c cs) 
+  case (pCons c cs)
   {assume cs0: "cs \<noteq> 0"
     from poly_infinity[OF cs0, of "cmod (poly (pCons c cs) 0)" c]
     obtain r where r: "\<And>z. r \<le> cmod z \<Longrightarrow> cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)" by blast
     have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>" by arith
-    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"] 
+    from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"]
     obtain v where v: "\<And>w. cmod w \<le> \<bar>r\<bar> \<Longrightarrow> cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)" by blast
     {fix z assume z: "r \<le> cmod z"
-      from v[of 0] r[OF z] 
+      from v[of 0] r[OF z]
       have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)"
 	by simp }
     note v0 = this
@@ -558,17 +558,17 @@
   unfolding constant_def psize_def
   apply (induct p, auto)
   done
- 
+
 lemma poly_replicate_append:
   "poly (monom 1 n * p) (x::'a::{recpower, comm_ring_1}) = x^n * poly p x"
   by (simp add: poly_monom)
 
-text {* Decomposition of polynomial, skipping zero coefficients 
+text {* Decomposition of polynomial, skipping zero coefficients
   after the first.  *}
 
 lemma poly_decompose_lemma:
  assumes nz: "\<not>(\<forall>z. z\<noteq>0 \<longrightarrow> poly p z = (0::'a::{recpower,idom}))"
-  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (psize q + k) = psize p \<and> 
+  shows "\<exists>k a q. a\<noteq>0 \<and> Suc (psize q + k) = psize p \<and>
                  (\<forall>z. poly p z = z^k * poly (pCons a q) z)"
 unfolding psize_def
 using nz
@@ -596,9 +596,9 @@
 lemma poly_decompose:
   assumes nc: "~constant(poly p)"
   shows "\<exists>k a q. a\<noteq>(0::'a::{recpower,idom}) \<and> k\<noteq>0 \<and>
-               psize q + k + 1 = psize p \<and> 
+               psize q + k + 1 = psize p \<and>
               (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
-using nc 
+using nc
 proof(induct p)
   case 0 thus ?case by (simp add: constant_def)
 next
@@ -608,8 +608,8 @@
       from C have "poly (pCons c cs) x = poly (pCons c cs) y" by (cases "x=0", auto)}
     with pCons.prems have False by (auto simp add: constant_def)}
   hence th: "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)" ..
-  from poly_decompose_lemma[OF th] 
-  show ?case 
+  from poly_decompose_lemma[OF th]
+  show ?case
     apply clarsimp
     apply (rule_tac x="k+1" in exI)
     apply (rule_tac x="a" in exI)
@@ -633,7 +633,7 @@
   let ?ths = "\<exists>z. ?p z = 0"
 
   from nonconstant_length[OF nc] have n2: "n\<ge> 2" by (simp add: n)
-  from poly_minimum_modulus obtain c where 
+  from poly_minimum_modulus obtain c where
     c: "\<forall>w. cmod (?p c) \<le> cmod (?p w)" by blast
   {assume pc: "?p c = 0" hence ?ths by blast}
   moreover
@@ -643,18 +643,18 @@
     {assume h: "constant (poly q)"
       from q(2) have th: "\<forall>x. poly q (x - c) = ?p x" by auto
       {fix x y
-	from th have "?p x = poly q (x - c)" by auto 
-	also have "\<dots> = poly q (y - c)" 
+	from th have "?p x = poly q (x - c)" by auto
+	also have "\<dots> = poly q (y - c)"
 	  using h unfolding constant_def by blast
 	also have "\<dots> = ?p y" using th by auto
 	finally have "?p x = ?p y" .}
       with nc have False unfolding constant_def by blast }
     hence qnc: "\<not> constant (poly q)" by blast
     from q(2) have pqc0: "?p c = poly q 0" by simp
-    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp 
+    from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)" by simp
     let ?a0 = "poly q 0"
-    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp 
-    from a00 
+    from pc0 pqc0 have a00: "?a0 \<noteq> 0" by simp
+    from a00
     have qr: "\<forall>z. poly q z = poly (smult (inverse ?a0) q) z * ?a0"
       by simp
     let ?r = "smult (inverse ?a0) q"
@@ -663,38 +663,38 @@
       by (simp add: expand_poly_eq)
     {assume h: "\<And>x y. poly ?r x = poly ?r y"
       {fix x y
-	from qr[rule_format, of x] 
+	from qr[rule_format, of x]
 	have "poly q x = poly ?r x * ?a0" by auto
 	also have "\<dots> = poly ?r y * ?a0" using h by simp
 	also have "\<dots> = poly q y" using qr[rule_format, of y] by simp
-	finally have "poly q x = poly q y" .} 
+	finally have "poly q x = poly q y" .}
       with qnc have False unfolding constant_def by blast}
     hence rnc: "\<not> constant (poly ?r)" unfolding constant_def by blast
     from qr[rule_format, of 0] a00  have r01: "poly ?r 0 = 1" by auto
-    {fix w 
+    {fix w
       have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
 	using qr[rule_format, of w] a00 by (simp add: divide_inverse mult_ac)
       also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
 	using a00 unfolding norm_divide by (simp add: field_simps)
       finally have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" .}
     note mrmq_eq = this
-    from poly_decompose[OF rnc] obtain k a s where 
-      kas: "a\<noteq>0" "k\<noteq>0" "psize s + k + 1 = psize ?r" 
+    from poly_decompose[OF rnc] obtain k a s where
+      kas: "a\<noteq>0" "k\<noteq>0" "psize s + k + 1 = psize ?r"
       "\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
     {assume "k + 1 = n"
       with kas(3) lgqr[symmetric] q(1) n[symmetric] have s0:"s=0" by auto
       {fix w
-	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)" 
+	have "cmod (poly ?r w) = cmod (1 + a * w ^ k)"
 	  using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)}
       note hth = this [symmetric]
-	from reduce_poly_simple[OF kas(1,2)] 
+	from reduce_poly_simple[OF kas(1,2)]
       have "\<exists>w. cmod (poly ?r w) < 1" unfolding hth by blast}
     moreover
     {assume kn: "k+1 \<noteq> n"
       from kn kas(3) q(1) n[symmetric] lgqr have k1n: "k + 1 < n" by simp
-      have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))" 
+      have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
 	unfolding constant_def poly_pCons poly_monom
-	using kas(1) apply simp 
+	using kas(1) apply simp
 	by (rule exI[where x=0], rule exI[where x=1], simp)
       from kas(1) kas(2) have th02: "k+1 = psize (pCons 1 (monom a (k - 1)))"
 	by (simp add: psize_def degree_monom_eq)
@@ -702,12 +702,12 @@
       obtain w where w: "1 + w^k * a = 0"
 	unfolding poly_pCons poly_monom
 	using kas(2) by (cases k, auto simp add: algebra_simps)
-      from poly_bound_exists[of "cmod w" s] obtain m where 
+      from poly_bound_exists[of "cmod w" s] obtain m where
 	m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
       have w0: "w\<noteq>0" using kas(2) w by (auto simp add: power_0_left)
       from w have "(1 + w ^ k * a) - 1 = 0 - 1" by simp
       then have wm1: "w^k * a = - 1" by simp
-      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)" 
+      have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
 	using norm_ge_zero[of w] w0 m(1)
 	  by (simp add: inverse_eq_divide zero_less_mult_iff)
       with real_down2[OF zero_less_one] obtain t where
@@ -717,20 +717,20 @@
       have "1 + ?w^k * (a + ?w * poly s ?w) = 1 + ?ct^k * (w^k * a) + ?w^k * ?w * poly s ?w" using kas(1) by (simp add: algebra_simps power_mult_distrib)
       also have "\<dots> = complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w"
 	unfolding wm1 by (simp)
-      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)" 
+      finally have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) = cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
 	apply -
 	apply (rule cong[OF refl[of cmod]])
 	apply assumption
 	done
-      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"] 
-      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp 
+      with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
+      have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)" unfolding norm_of_real by simp
       have ath: "\<And>x (t::real). 0\<le> x \<Longrightarrow> x < t \<Longrightarrow> t\<le>1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1" by arith
       have "t *cmod w \<le> 1 * cmod w" apply (rule mult_mono) using t(1,2) by auto
-      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult) 
+      then have tw: "cmod ?w \<le> cmod w" using t(1) by (simp add: norm_mult)
       from t inv0 have "t* (cmod w ^ (k + 1) * m) < 1"
 	by (simp add: inverse_eq_divide field_simps)
-      with zero_less_power[OF t(1), of k] 
-      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1" 
+      with zero_less_power[OF t(1), of k]
+      have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
 	apply - apply (rule mult_strict_left_mono) by simp_all
       have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k+1) * cmod (poly s ?w)))"  using w0 t(1)
 	by (simp add: algebra_simps power_mult_distrib norm_of_real norm_power norm_mult)
@@ -741,15 +741,15 @@
 	apply (rule mult_mono, simp_all add: norm_ge_zero)+
 	apply (simp add: zero_le_mult_iff zero_le_power)
 	done
-      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp 
-      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1" 
+      with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k" by simp
+      from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
 	by auto
       from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
-      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" . 
+      have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
       from th11 th12
-      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith 
-      then have "cmod (poly ?r ?w) < 1" 
-	unfolding kas(4)[rule_format, of ?w] r01 by simp 
+      have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"  by arith
+      then have "cmod (poly ?r ?w) < 1"
+	unfolding kas(4)[rule_format, of ?w] r01 by simp
       then have "\<exists>w. cmod (poly ?r w) < 1" by blast}
     ultimately have cr0_contr: "\<exists>w. cmod (poly ?r w) < 1" by blast
     from cr0_contr cq0 q(2)
@@ -769,18 +769,18 @@
   moreover
   {assume c0: "c\<noteq>0"
     {assume nc: "constant (poly (pCons c cs))"
-      from nc[unfolded constant_def, rule_format, of 0] 
-      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto 
+      from nc[unfolded constant_def, rule_format, of 0]
+      have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
       hence "cs = 0"
 	proof(induct cs)
 	  case (pCons d ds)
 	  {assume "d=0" hence ?case using pCons.prems pCons.hyps by simp}
 	  moreover
 	  {assume d0: "d\<noteq>0"
-	    from poly_bound_exists[of 1 ds] obtain m where 
+	    from poly_bound_exists[of 1 ds] obtain m where
 	      m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
 	    have dm: "cmod d / m > 0" using d0 m(1) by (simp add: field_simps)
-	    from real_down2[OF dm zero_less_one] obtain x where 
+	    from real_down2[OF dm zero_less_one] obtain x where
 	      x: "x > 0" "x < cmod d / m" "x < 1" by blast
 	    let ?x = "complex_of_real x"
 	    from x have cx: "?x \<noteq> 0"  "cmod ?x \<le> 1" by simp_all
@@ -792,12 +792,12 @@
 	    from x(2) m(1) have "x*m < cmod d" by (simp add: field_simps)
 	    with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d" by auto
 	    with cth  have ?case by blast}
-	  ultimately show ?case by blast 
+	  ultimately show ?case by blast
 	qed simp}
-      then have nc: "\<not> constant (poly (pCons c cs))" using pCons.prems c0 
+      then have nc: "\<not> constant (poly (pCons c cs))" using pCons.prems c0
 	by blast
       from fundamental_theorem_of_algebra[OF nc] have ?case .}
-  ultimately show ?case by blast  
+  ultimately show ?case by blast
 qed simp
 
 
@@ -814,15 +814,15 @@
   assume IH: "\<forall>m<n. \<forall>p q.
                  (\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longrightarrow>
                  degree p = m \<longrightarrow> m \<noteq> 0 \<longrightarrow> p dvd (q ^ m)"
-    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0" 
+    and pq0: "\<forall>x. poly p x = 0 \<longrightarrow> poly q x = 0"
     and dpn: "degree p = n" and n0: "n \<noteq> 0"
   from dpn n0 have pne: "p \<noteq> 0" by auto
   let ?ths = "p dvd (q ^ n)"
   {fix a assume a: "poly p a = 0"
     {assume oa: "order a p \<noteq> 0"
       let ?op = "order a p"
-      from pne have ap: "([:- a, 1:] ^ ?op) dvd p" 
-	"\<not> [:- a, 1:] ^ (Suc ?op) dvd p" using order by blast+ 
+      from pne have ap: "([:- a, 1:] ^ ?op) dvd p"
+	"\<not> [:- a, 1:] ^ (Suc ?op) dvd p" using order by blast+
       note oop = order_degree[OF pne, unfolded dpn]
       {assume q0: "q = 0"
 	hence ?ths using n0
@@ -892,7 +892,7 @@
   {assume exa: "\<not> (\<exists>a. poly p a = 0)"
     from fundamental_theorem_of_algebra_alt[of p] exa obtain c where
       ccs: "c\<noteq>0" "p = pCons c 0" by blast
-    
+
     then have pp: "\<And>x. poly p x =  c" by simp
     let ?w = "[:1/c:] * (q ^ n)"
     from ccs
@@ -903,7 +903,7 @@
 qed
 
 lemma nullstellensatz_univariate:
-  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> 
+  "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow>
     p dvd (q ^ (degree p)) \<or> (p = 0 \<and> q = 0)"
 proof-
   {assume pe: "p = 0"
@@ -933,7 +933,7 @@
 	{fix x assume h: "poly p x = 0" "poly q x \<noteq> 0"
 	  hence "poly (q ^ (Suc n)) x \<noteq> 0" by simp
 	  hence False using u h(1) by (simp only: poly_mult) simp}}
-	with n nullstellensatz_lemma[of p q "degree p"] dp 
+	with n nullstellensatz_lemma[of p q "degree p"] dp
 	have ?thesis by auto}
     ultimately have ?thesis by blast}
   ultimately show ?thesis by blast
@@ -966,13 +966,13 @@
 
 (* Arithmetic operations on multivariate polynomials.                        *)
 
-lemma mpoly_base_conv: 
+lemma mpoly_base_conv:
   "(0::complex) \<equiv> poly 0 x" "c \<equiv> poly [:c:] x" "x \<equiv> poly [:0,1:] x" by simp_all
 
-lemma mpoly_norm_conv: 
+lemma mpoly_norm_conv:
   "poly [:0:] (x::complex) \<equiv> poly 0 x" "poly [:poly 0 y:] x \<equiv> poly 0 x" by simp_all
 
-lemma mpoly_sub_conv: 
+lemma mpoly_sub_conv:
   "poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
   by (simp add: diff_def)
 
@@ -982,9 +982,9 @@
 
 lemma resolve_eq_raw:  "poly 0 x \<equiv> 0" "poly [:c:] x \<equiv> (c::complex)" by auto
 lemma  resolve_eq_then: "(P \<Longrightarrow> (Q \<equiv> Q1)) \<Longrightarrow> (\<not>P \<Longrightarrow> (Q \<equiv> Q2))
-  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast 
+  \<Longrightarrow> Q \<equiv> P \<and> Q1 \<or> \<not>P\<and> Q2" apply (atomize (full)) by blast
 
-lemma poly_divides_pad_rule: 
+lemma poly_divides_pad_rule:
   fixes p q :: "complex poly"
   assumes pq: "p dvd q"
   shows "p dvd (pCons (0::complex) q)"
@@ -994,7 +994,7 @@
   with pq show ?thesis by (rule dvd_trans)
 qed
 
-lemma poly_divides_pad_const_rule: 
+lemma poly_divides_pad_const_rule:
   fixes p q :: "complex poly"
   assumes pq: "p dvd q"
   shows "p dvd (smult a q)"
@@ -1005,12 +1005,12 @@
 qed
 
 
-lemma poly_divides_conv0:  
+lemma poly_divides_conv0:
   fixes p :: "complex poly"
   assumes lgpq: "degree q < degree p" and lq:"p \<noteq> 0"
   shows "p dvd q \<equiv> q = 0" (is "?lhs \<equiv> ?rhs")
 proof-
-  {assume r: ?rhs 
+  {assume r: ?rhs
     hence "q = p * 0" by simp
     hence ?lhs ..}
   moreover
@@ -1023,10 +1023,10 @@
         by (rule dvd_imp_degree_le)
       with lgpq have ?rhs by simp }
     ultimately have ?rhs by blast }
-  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast) 
+  ultimately show "?lhs \<equiv> ?rhs" by - (atomize (full), blast)
 qed
 
-lemma poly_divides_conv1: 
+lemma poly_divides_conv1:
   assumes a0: "a\<noteq> (0::complex)" and pp': "(p::complex poly) dvd p'"
   and qrp': "smult a q - p' \<equiv> r"
   shows "p dvd q \<equiv> p dvd (r::complex poly)" (is "?lhs \<equiv> ?rhs")
@@ -1046,7 +1046,7 @@
       by (simp add: algebra_simps mult_smult_right smult_smult)
     hence ?lhs ..}
   ultimately have "?lhs = ?rhs" by blast }
-thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast) 
+thus "?lhs \<equiv> ?rhs"  by - (atomize(full), blast)
 qed
 
 lemma basic_cqe_conv1:
@@ -1056,8 +1056,8 @@
   "(\<exists>x. poly 0 x = 0) \<equiv> True"
   "(\<exists>x. poly [:c:] x = 0) \<equiv> c = 0" by simp_all
 
-lemma basic_cqe_conv2: 
-  assumes l:"p \<noteq> 0" 
+lemma basic_cqe_conv2:
+  assumes l:"p \<noteq> 0"
   shows "(\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)) \<equiv> True"
 proof-
   {fix h t
@@ -1065,7 +1065,7 @@
     with l have False by simp}
   hence th: "\<not> (\<exists> h t. h\<noteq>0 \<and> t=0 \<and> pCons a (pCons b p) = pCons h t)"
     by blast
-  from fundamental_theorem_of_algebra_alt[OF th] 
+  from fundamental_theorem_of_algebra_alt[OF th]
   show "(\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)) \<equiv> True" by auto
 qed
 
@@ -1080,7 +1080,7 @@
 
 lemma basic_cqe_conv3:
   fixes p q :: "complex poly"
-  assumes l: "p \<noteq> 0" 
+  assumes l: "p \<noteq> 0"
   shows "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<equiv> \<not> ((pCons a p) dvd (q ^ (psize p)))"
 proof-
   from l have dp:"degree (pCons a p) = psize p" by (simp add: psize_def)
@@ -1108,7 +1108,7 @@
 lemma negate_negate_rule: "Trueprop P \<equiv> \<not> P \<equiv> False" by (atomize (full), auto)
 
 lemma complex_entire: "(z::complex) \<noteq> 0 \<and> w \<noteq> 0 \<equiv> z*w \<noteq> 0" by simp
-lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)" 
+lemma resolve_eq_ne: "(P \<equiv> True) \<equiv> (\<not>P \<equiv> False)" "(P \<equiv> False) \<equiv> (\<not>P \<equiv> True)"
   by (atomize (full)) simp_all
 lemma cqe_conv1: "poly 0 x = 0 \<longleftrightarrow> True"  by simp
 lemma cqe_conv2: "(p \<Longrightarrow> (q \<equiv> r)) \<equiv> ((p \<and> q) \<equiv> (p \<and> r))"  (is "?l \<equiv> ?r")
--- a/src/HOL/Library/OptionalSugar.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/OptionalSugar.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -4,13 +4,21 @@
 *)
 (*<*)
 theory OptionalSugar
-imports LaTeXsugar
+imports LaTeXsugar Complex_Main
 begin
 
-(* set *)
+(* hiding set *)
 translations
   "xs" <= "CONST set xs"
 
+(* hiding numeric conversions - embeddings only *)
+translations
+  "n" <= "CONST of_nat n"
+  "n" <= "CONST int n"
+  "n" <= "real n"
+  "n" <= "CONST real_of_nat n"
+  "n" <= "CONST real_of_int n"
+
 (* append *)
 syntax (latex output)
   "appendL" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "\<^raw:\isacharat>" 65)
--- a/src/HOL/Library/Permutations.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Permutations.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -18,7 +18,7 @@
 (* ------------------------------------------------------------------------- *)
 
 declare swap_self[simp]
-lemma swapid_sym: "Fun.swap a b id = Fun.swap b a id" 
+lemma swapid_sym: "Fun.swap a b id = Fun.swap b a id"
   by (auto simp add: expand_fun_eq swap_def fun_upd_def)
 lemma swap_id_refl: "Fun.swap a a id = id" by simp
 lemma swap_id_sym: "Fun.swap a b id = Fun.swap b a id"
@@ -45,18 +45,18 @@
 
 lemma permutes_image: assumes pS: "p permutes S" shows "p ` S = S"
   using pS
-  unfolding permutes_def 
-  apply - 
-  apply (rule set_ext) 
+  unfolding permutes_def
+  apply -
+  apply (rule set_ext)
   apply (simp add: image_iff)
   apply metis
   done
 
-lemma permutes_inj: "p permutes S ==> inj p " 
-  unfolding permutes_def inj_on_def by blast 
+lemma permutes_inj: "p permutes S ==> inj p "
+  unfolding permutes_def inj_on_def by blast
 
-lemma permutes_surj: "p permutes s ==> surj p" 
-  unfolding permutes_def surj_def by metis 
+lemma permutes_surj: "p permutes s ==> surj p"
+  unfolding permutes_def surj_def by metis
 
 lemma permutes_inv_o: assumes pS: "p permutes S"
   shows " p o inv p = id"
@@ -65,7 +65,7 @@
   unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+
 
 
-lemma permutes_inverses: 
+lemma permutes_inverses:
   fixes p :: "'a \<Rightarrow> 'a"
   assumes pS: "p permutes S"
   shows "p (inv p x) = x"
@@ -76,11 +76,11 @@
   unfolding permutes_def by blast
 
 lemma permutes_empty[simp]: "p permutes {} \<longleftrightarrow> p = id"
-  unfolding expand_fun_eq permutes_def apply simp by metis 
+  unfolding expand_fun_eq permutes_def apply simp by metis
 
 lemma permutes_sing[simp]: "p permutes {a} \<longleftrightarrow> p = id"
   unfolding expand_fun_eq permutes_def apply simp by metis
- 
+
 lemma permutes_univ: "p permutes UNIV \<longleftrightarrow> (\<forall>y. \<exists>!x. p x = y)"
   unfolding permutes_def by simp
 
@@ -104,13 +104,13 @@
 (* Group properties.                                                         *)
 (* ------------------------------------------------------------------------- *)
 
-lemma permutes_id: "id permutes S" unfolding permutes_def by simp 
+lemma permutes_id: "id permutes S" unfolding permutes_def by simp
 
 lemma permutes_compose: "p permutes S \<Longrightarrow> q permutes S ==> q o p permutes S"
   unfolding permutes_def o_def by metis
 
 lemma permutes_inv: assumes pS: "p permutes S" shows "inv p permutes S"
-  using pS unfolding permutes_def permutes_inv_eq[OF pS] by metis  
+  using pS unfolding permutes_def permutes_inv_eq[OF pS] by metis
 
 lemma permutes_inv_inv: assumes pS: "p permutes S" shows "inv (inv p) = p"
   unfolding expand_fun_eq permutes_inv_eq[OF pS] permutes_inv_eq[OF permutes_inv[OF pS]]
@@ -120,7 +120,7 @@
 (* The number of permutations on a finite set.                               *)
 (* ------------------------------------------------------------------------- *)
 
-lemma permutes_insert_lemma: 
+lemma permutes_insert_lemma:
   assumes pS: "p permutes (insert a S)"
   shows "Fun.swap a (p a) id o p permutes S"
   apply (rule permutes_superset[where S = "insert a S"])
@@ -134,17 +134,17 @@
         (\<lambda>(b,p). Fun.swap a b id o p) ` {(b,p). b \<in> insert a S \<and> p \<in> {p. p permutes S}}"
 proof-
 
-  {fix p 
+  {fix p
     {assume pS: "p permutes insert a S"
       let ?b = "p a"
       let ?q = "Fun.swap a (p a) id o p"
-      have th0: "p = Fun.swap a ?b id o ?q" unfolding expand_fun_eq o_assoc by simp 
-      have th1: "?b \<in> insert a S " unfolding permutes_in_image[OF pS] by simp 
+      have th0: "p = Fun.swap a ?b id o ?q" unfolding expand_fun_eq o_assoc by simp
+      have th1: "?b \<in> insert a S " unfolding permutes_in_image[OF pS] by simp
       from permutes_insert_lemma[OF pS] th0 th1
       have "\<exists> b q. p = Fun.swap a b id o q \<and> b \<in> insert a S \<and> q permutes S" by blast}
     moreover
     {fix b q assume bq: "p = Fun.swap a b id o q" "b \<in> insert a S" "q permutes S"
-      from permutes_subset[OF bq(3), of "insert a S"] 
+      from permutes_subset[OF bq(3), of "insert a S"]
       have qS: "q permutes insert a S" by auto
       have aS: "a \<in> insert a S" by simp
       from bq(1) permutes_compose[OF qS permutes_swap_id[OF aS bq(2)]]
@@ -168,8 +168,8 @@
     show "\<forall>n. ({} hassize n) \<longrightarrow> ({p. p permutes {}} hassize fact n)"
       by (simp add: hassize_def permutes_empty)
   next
-    fix x F 
-    assume fF: "finite F" and xF: "x \<notin> F" 
+    fix x F
+    assume fF: "finite F" and xF: "x \<notin> F"
       and H: "\<forall>n. (F hassize n) \<longrightarrow> ({p. p permutes F} hassize fact n)"
     {fix n assume H0: "insert x F hassize n"
       let ?xF = "{p. p permutes insert x F}"
@@ -180,7 +180,7 @@
       have xfgpF': "?xF = ?g ` ?pF'" .
       from hassize_insert[OF xF H0] have Fs: "F hassize (n - 1)" .
       from H Fs have pFs: "?pF hassize fact (n - 1)" by blast
-      hence pF'f: "finite ?pF'" using H0 unfolding hassize_def 
+      hence pF'f: "finite ?pF'" using H0 unfolding hassize_def
 	apply (simp only: Collect_split Collect_mem_eq)
 	apply (rule finite_cartesian_product)
 	apply simp_all
@@ -189,12 +189,12 @@
       have ginj: "inj_on ?g ?pF'"
       proof-
 	{
-	  fix b p c q assume bp: "(b,p) \<in> ?pF'" and cq: "(c,q) \<in> ?pF'" 
+	  fix b p c q assume bp: "(b,p) \<in> ?pF'" and cq: "(c,q) \<in> ?pF'"
 	    and eq: "?g (b,p) = ?g (c,q)"
 	  from bp cq have ths: "b \<in> insert x F" "c \<in> insert x F" "x \<in> insert x F" "p permutes F" "q permutes F" by auto
-	  from ths(4) xF eq have "b = ?g (b,p) x" unfolding permutes_def 
+	  from ths(4) xF eq have "b = ?g (b,p) x" unfolding permutes_def
 	    by (auto simp add: swap_def fun_upd_def expand_fun_eq)
-	  also have "\<dots> = ?g (c,q) x" using ths(5) xF eq  
+	  also have "\<dots> = ?g (c,q) x" using ths(5) xF eq
 	    by (auto simp add: swap_def fun_upd_def expand_fun_eq)
 	  also have "\<dots> = c"using ths(5) xF unfolding permutes_def
 	    by (auto simp add: swap_def fun_upd_def expand_fun_eq)
@@ -208,16 +208,16 @@
       qed
       from xF H0 have n0: "n \<noteq> 0 " by (auto simp add: hassize_def)
       hence "\<exists>m. n = Suc m" by arith
-      then obtain m where n[simp]: "n = Suc m" by blast 
-      from pFs H0 have xFc: "card ?xF = fact n" 
+      then obtain m where n[simp]: "n = Suc m" by blast
+      from pFs H0 have xFc: "card ?xF = fact n"
 	unfolding xfgpF' card_image[OF ginj] hassize_def
 	apply (simp only: Collect_split Collect_mem_eq card_cartesian_product)
 	by simp
-      from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" unfolding xfgpF' by simp 
+      from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" unfolding xfgpF' by simp
       have "?xF hassize fact n"
-	using xFf xFc 
+	using xFf xFc
 	unfolding hassize_def  xFf by blast }
-    thus "\<forall>n. (insert x F hassize n) \<longrightarrow> ({p. p permutes insert x F} hassize fact n)" 
+    thus "\<forall>n. (insert x F hassize n) \<longrightarrow> ({p. p permutes insert x F} hassize fact n)"
       by blast
   qed
   with Sn show ?thesis by blast
@@ -230,11 +230,11 @@
 (* Permutations of index set for iterated operations.                        *)
 (* ------------------------------------------------------------------------- *)
 
-lemma (in ab_semigroup_mult) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S" 
+lemma (in ab_semigroup_mult) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S"
   shows "fold_image times f z S = fold_image times (f o p) z S"
   using fold_image_reindex[OF fS subset_inj_on[OF permutes_inj[OF pS], of S, simplified], of f z]
   unfolding permutes_image[OF pS] .
-lemma (in ab_semigroup_add) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S" 
+lemma (in ab_semigroup_add) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S"
   shows "fold_image plus f z S = fold_image plus (f o p) z S"
 proof-
   interpret ab_semigroup_mult plus apply unfold_locales apply (simp add: add_assoc)
@@ -244,22 +244,22 @@
   unfolding permutes_image[OF pS] .
 qed
 
-lemma setsum_permute: assumes pS: "p permutes S" 
+lemma setsum_permute: assumes pS: "p permutes S"
   shows "setsum f S = setsum (f o p) S"
   unfolding setsum_def using fold_image_permute[of S p f 0] pS by clarsimp
 
-lemma setsum_permute_natseg:assumes pS: "p permutes {m .. n}" 
+lemma setsum_permute_natseg:assumes pS: "p permutes {m .. n}"
   shows "setsum f {m .. n} = setsum (f o p) {m .. n}"
-  using setsum_permute[OF pS, of f ] pS by blast 
+  using setsum_permute[OF pS, of f ] pS by blast
 
-lemma setprod_permute: assumes pS: "p permutes S" 
+lemma setprod_permute: assumes pS: "p permutes S"
   shows "setprod f S = setprod (f o p) S"
-  unfolding setprod_def 
+  unfolding setprod_def
   using ab_semigroup_mult_class.fold_image_permute[of S p f 1] pS by clarsimp
 
-lemma setprod_permute_natseg:assumes pS: "p permutes {m .. n}" 
+lemma setprod_permute_natseg:assumes pS: "p permutes {m .. n}"
   shows "setprod f {m .. n} = setprod (f o p) {m .. n}"
-  using setprod_permute[OF pS, of f ] pS by blast 
+  using setprod_permute[OF pS, of f ] pS by blast
 
 (* ------------------------------------------------------------------------- *)
 (* Various combinations of transpositions with 2, 1 and 0 common elements.   *)
@@ -298,16 +298,16 @@
 
 lemma permutation_swap_id: "permutation (Fun.swap a b id)"
   apply (cases "a=b", simp_all)
-  unfolding permutation_def using swapidseq_swap[of a b] by blast 
+  unfolding permutation_def using swapidseq_swap[of a b] by blast
 
 lemma swapidseq_comp_add: "swapidseq n p \<Longrightarrow> swapidseq m q ==> swapidseq (n + m) (p o q)"
   proof (induct n p arbitrary: m q rule: swapidseq.induct)
     case (id m q) thus ?case by simp
   next
-    case (comp_Suc n p a b m q) 
+    case (comp_Suc n p a b m q)
     have th: "Suc n + m = Suc (n + m)" by arith
-    show ?case unfolding th o_assoc[symmetric] 
-      apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems]  comp_Suc.hyps(3) by blast+ 
+    show ?case unfolding th o_assoc[symmetric]
+      apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems]  comp_Suc.hyps(3) by blast+
 qed
 
 lemma permutation_compose: "permutation p \<Longrightarrow> permutation q ==> permutation(p o q)"
@@ -321,17 +321,17 @@
 lemma swapidseq_inverse_exists: "swapidseq n p ==> \<exists>q. swapidseq n q \<and> p o q = id \<and> q o p = id"
 proof(induct n p rule: swapidseq.induct)
   case id  thus ?case by (rule exI[where x=id], simp)
-next 
+next
   case (comp_Suc n p a b)
   from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
   let ?q = "q o Fun.swap a b id"
   note H = comp_Suc.hyps
   from swapidseq_swap[of a b] H(3)  have th0: "swapidseq 1 (Fun.swap a b id)" by simp
-  from swapidseq_comp_add[OF q(1) th0] have th1:"swapidseq (Suc n) ?q" by simp 
+  from swapidseq_comp_add[OF q(1) th0] have th1:"swapidseq (Suc n) ?q" by simp
   have "Fun.swap a b id o p o ?q = Fun.swap a b id o (p o q) o Fun.swap a b id" by (simp add: o_assoc)
   also have "\<dots> = id" by (simp add: q(2))
   finally have th2: "Fun.swap a b id o p o ?q = id" .
-  have "?q \<circ> (Fun.swap a b id \<circ> p) = q \<circ> (Fun.swap a b id o Fun.swap a b id) \<circ> p" by (simp only: o_assoc) 
+  have "?q \<circ> (Fun.swap a b id \<circ> p) = q \<circ> (Fun.swap a b id o Fun.swap a b id) \<circ> p" by (simp only: o_assoc)
   hence "?q \<circ> (Fun.swap a b id \<circ> p) = id" by (simp add: q(3))
   with th1 th2 show ?case by blast
 qed
@@ -351,15 +351,15 @@
    (\<And>a b c d. a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> (a = c \<and> b = d \<or>  a = c \<and> b \<noteq> d \<or> a \<noteq> c \<and> b = d \<or> a \<noteq> c \<and> a \<noteq> d \<and> b \<noteq> c \<and> b \<noteq> d) ==> P a b c d)
    ==> (\<And>a b c d. a \<noteq> b --> c \<noteq> d \<longrightarrow>  P a b c d)" by metis
 
-lemma swap_general: "a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> Fun.swap a b id o Fun.swap c d id = id \<or> 
-  (\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id)" 
+lemma swap_general: "a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> Fun.swap a b id o Fun.swap c d id = id \<or>
+  (\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id)"
 proof-
   assume H: "a\<noteq>b" "c\<noteq>d"
-have "a \<noteq> b \<longrightarrow> c \<noteq> d \<longrightarrow> 
-(  Fun.swap a b id o Fun.swap c d id = id \<or> 
-  (\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id))" 
+have "a \<noteq> b \<longrightarrow> c \<noteq> d \<longrightarrow>
+(  Fun.swap a b id o Fun.swap c d id = id \<or>
+  (\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id))"
   apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d])
-  apply (simp_all only: swapid_sym) 
+  apply (simp_all only: swapid_sym)
   apply (case_tac "a = c \<and> b = d", clarsimp simp only: swapid_sym swap_id_idempotent)
   apply (case_tac "a = c \<and> b \<noteq> d")
   apply (rule disjI2)
@@ -379,7 +379,7 @@
   apply (rule_tac x="b" in exI)
   apply (clarsimp simp add: expand_fun_eq swap_def)
   done
-with H show ?thesis by metis 
+with H show ?thesis by metis
 qed
 
 lemma swapidseq_id_iff[simp]: "swapidseq 0 p \<longleftrightarrow> p = id"
@@ -411,12 +411,12 @@
     c d q m where cdqm: "Suc n = Suc m" "p = Fun.swap c d id o q" "swapidseq m q" "c \<noteq> d" "n = m"
     by auto
   {assume H: "Fun.swap a b id o Fun.swap c d id = id"
-    
-    have ?case apply (simp only: cdqm o_assoc H) 
+
+    have ?case apply (simp only: cdqm o_assoc H)
       by (simp add: cdqm)}
   moreover
   { fix x y z
-    assume H: "x\<noteq>a" "y\<noteq>a" "z \<noteq>a" "x \<noteq>y" 
+    assume H: "x\<noteq>a" "y\<noteq>a" "z \<noteq>a" "x \<noteq>y"
       "Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id"
     from H have az: "a \<noteq> z" by simp
 
@@ -430,23 +430,23 @@
     hence th1: "(Fun.swap a z id o q) a = a" unfolding th3 .
     from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az th1]
     have th2: "swapidseq (n - 1) (Fun.swap a z id o q)" "n \<noteq> 0" by blast+
-    have th: "Suc n - 1 = Suc (n - 1)" using th2(2) by auto 
+    have th: "Suc n - 1 = Suc (n - 1)" using th2(2) by auto
     have ?case unfolding cdqm(2) H o_assoc th
       apply (simp only: Suc_not_Zero simp_thms o_assoc[symmetric])
       apply (rule comp_Suc)
       using th2 H apply blast+
       done}
-  ultimately show ?case using swap_general[OF Suc.prems(2) cdqm(4)] by metis 
+  ultimately show ?case using swap_general[OF Suc.prems(2) cdqm(4)] by metis
 qed
 
-lemma swapidseq_identity_even: 
+lemma swapidseq_identity_even:
   assumes "swapidseq n (id :: 'a \<Rightarrow> 'a)" shows "even n"
   using `swapidseq n id`
 proof(induct n rule: nat_less_induct)
   fix n
   assume H: "\<forall>m<n. swapidseq m (id::'a \<Rightarrow> 'a) \<longrightarrow> even m" "swapidseq n (id :: 'a \<Rightarrow> 'a)"
-  {assume "n = 0" hence "even n" by arith} 
-  moreover 
+  {assume "n = 0" hence "even n" by arith}
+  moreover
   {fix a b :: 'a and q m
     assume h: "n = Suc m" "(id :: 'a \<Rightarrow> 'a) = Fun.swap a b id \<circ> q" "swapidseq m q" "a \<noteq> b"
     from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]]
@@ -462,13 +462,13 @@
 
 definition "evenperm p = even (SOME n. swapidseq n p)"
 
-lemma swapidseq_even_even: assumes 
+lemma swapidseq_even_even: assumes
   m: "swapidseq m p" and n: "swapidseq n p"
   shows "even m \<longleftrightarrow> even n"
 proof-
   from swapidseq_inverse_exists[OF n]
   obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
-  
+
   from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]]
   show ?thesis by arith
 qed
@@ -491,12 +491,12 @@
 apply (rule evenperm_unique[where n="if a = b then 0 else 1"])
 by (simp_all add: swapidseq_swap)
 
-lemma evenperm_comp: 
+lemma evenperm_comp:
   assumes p: "permutation p" and q:"permutation q"
   shows "evenperm (p o q) = (evenperm p = evenperm q)"
 proof-
-  from p q obtain 
-    n m where n: "swapidseq n p" and m: "swapidseq m q" 
+  from p q obtain
+    n m where n: "swapidseq n p" and m: "swapidseq m q"
     unfolding permutation_def by blast
   note nm =  swapidseq_comp_add[OF n m]
   have th: "even (n + m) = (even n \<longleftrightarrow> even m)" by arith
@@ -525,15 +525,15 @@
   apply metis
   done
 
-lemma permutation_bijective: 
-  assumes p: "permutation p" 
+lemma permutation_bijective:
+  assumes p: "permutation p"
   shows "bij p"
 proof-
   from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
-  from swapidseq_inverse_exists[OF n] obtain q where 
+  from swapidseq_inverse_exists[OF n] obtain q where
     q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
   thus ?thesis unfolding bij_iff  apply (auto simp add: expand_fun_eq) apply metis done
-qed  
+qed
 
 lemma permutation_finite_support: assumes p: "permutation p"
   shows "finite {x. p x \<noteq> x}"
@@ -555,7 +555,7 @@
 lemma bij_inv_eq_iff: "bij p ==> x = inv p y \<longleftrightarrow> p x = y"
   using surj_f_inv_f[of p] inv_f_f[of f] by (auto simp add: bij_def)
 
-lemma bij_swap_comp: 
+lemma bij_swap_comp:
   assumes bp: "bij p" shows "Fun.swap a b id o p = Fun.swap (inv p a) (inv p b) p"
   using surj_f_inv_f[OF bij_is_surj[OF bp]]
   by (simp add: expand_fun_eq swap_def bij_inv_eq_iff[OF bp])
@@ -563,12 +563,12 @@
 lemma bij_swap_ompose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id o p)"
 proof-
   assume H: "bij p"
-  show ?thesis 
+  show ?thesis
     unfolding bij_swap_comp[OF H] bij_swap_iff
     using H .
 qed
 
-lemma permutation_lemma: 
+lemma permutation_lemma:
   assumes fS: "finite S" and p: "bij p" and pS: "\<forall>x. x\<notin> S \<longrightarrow> p x = x"
   shows "permutation p"
 using fS p pS
@@ -580,9 +580,9 @@
   let ?q = "Fun.swap a (p a) id o ?r "
   have raa: "?r a = a" by (simp add: swap_def)
   from bij_swap_ompose_bij[OF insert(4)]
-  have br: "bij ?r"  . 
-  
-  from insert raa have th: "\<forall>x. x \<notin> F \<longrightarrow> ?r x = x"    
+  have br: "bij ?r"  .
+
+  from insert raa have th: "\<forall>x. x \<notin> F \<longrightarrow> ?r x = x"
     apply (clarsimp simp add: swap_def)
     apply (erule_tac x="x" in allE)
     apply auto
@@ -594,7 +594,7 @@
   thus ?case by (simp add: o_assoc)
 qed
 
-lemma permutation: "permutation p \<longleftrightarrow> bij p \<and> finite {x. p x \<noteq> x}" 
+lemma permutation: "permutation p \<longleftrightarrow> bij p \<and> finite {x. p x \<noteq> x}"
   (is "?lhs \<longleftrightarrow> ?b \<and> ?f")
 proof
   assume p: ?lhs
@@ -620,7 +620,7 @@
   finally have th0: "p o q o (inv q o inv p) = id" .
   have "inv q o inv p o (p o q) = inv q o (inv p o p) o q" by (simp add: o_assoc)
   also have "\<dots> = id" by (simp add: ps qs)
-  finally have th1: "inv q o inv p o (p o q) = id" . 
+  finally have th1: "inv q o inv p o (p o q) = id" .
   from inv_unique_comp[OF th0 th1] show ?thesis .
 qed
 
@@ -646,20 +646,20 @@
          ==> (\<And>p. p permutes S ==> P p)"
 proof(induct S rule: finite_induct)
   case empty thus ?case by auto
-next 
+next
   case (insert x F p)
   let ?r = "Fun.swap x (p x) id o p"
   let ?q = "Fun.swap x (p x) id o ?r"
   have qp: "?q = p" by (simp add: o_assoc)
   from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r" by blast
-  from permutes_in_image[OF insert.prems(3), of x] 
+  from permutes_in_image[OF insert.prems(3), of x]
   have pxF: "p x \<in> insert x F" by simp
   have xF: "x \<in> insert x F" by simp
   have rp: "permutation ?r"
-    unfolding permutation_permutes using insert.hyps(1) 
+    unfolding permutation_permutes using insert.hyps(1)
       permutes_insert_lemma[OF insert.prems(3)] by blast
-  from insert.prems(2)[OF xF pxF Pr Pr rp] 
-  show ?case  unfolding qp . 
+  from insert.prems(2)[OF xF pxF Pr Pr rp]
+  show ?case  unfolding qp .
 qed
 
 (* ------------------------------------------------------------------------- *)
@@ -668,7 +668,7 @@
 
 definition "sign p = (if evenperm p then (1::int) else -1)"
 
-lemma sign_nz: "sign p \<noteq> 0" by (simp add: sign_def) 
+lemma sign_nz: "sign p \<noteq> 0" by (simp add: sign_def)
 lemma sign_id: "sign id = 1" by (simp add: sign_def)
 lemma sign_inverse: "permutation p ==> sign (inv p) = sign p"
   by (simp add: sign_def evenperm_inv)
@@ -685,16 +685,16 @@
   assumes p: "p permutes (S::'a::wellorder set)" and le: "\<forall>i \<in> S.  p i <= i" shows "p = id"
 proof-
   {fix n
-    have "p n = n" 
+    have "p n = n"
       using p le
     proof(induct n arbitrary: S rule: less_induct)
-      fix n S assume H: "\<And>m S. \<lbrakk>m < n; p permutes S; \<forall>i\<in>S. p i \<le> i\<rbrakk> \<Longrightarrow> p m = m" 
+      fix n S assume H: "\<And>m S. \<lbrakk>m < n; p permutes S; \<forall>i\<in>S. p i \<le> i\<rbrakk> \<Longrightarrow> p m = m"
 	"p permutes S" "\<forall>i \<in>S. p i \<le> i"
       {assume "n \<notin> S"
 	with H(2) have "p n = n" unfolding permutes_def by metis}
       moreover
       {assume ns: "n \<in> S"
-	from H(3)  ns have "p n < n \<or> p n = n" by auto 
+	from H(3)  ns have "p n < n \<or> p n = n" by auto
 	moreover{assume h: "p n < n"
 	  from H h have "p (p n) = p n" by metis
 	  with permutes_inj[OF H(2)] have "p n = n" unfolding inj_on_def by blast
@@ -713,9 +713,9 @@
     with le have "p (inv p i) \<ge> inv p i" by blast
     with permutes_inverses[OF p] have "i \<ge> inv p i" by simp}
   then have th: "\<forall>i\<in>S. inv p i \<le> i"  by blast
-  from permutes_natset_le[OF permutes_inv[OF p] th] 
+  from permutes_natset_le[OF permutes_inv[OF p] th]
   have "inv p = inv id" by simp
-  then show ?thesis 
+  then show ?thesis
     apply (subst permutes_inv_inv[OF p, symmetric])
     apply (rule inv_unique_comp)
     apply simp_all
@@ -730,7 +730,7 @@
   apply auto
   done
 
-lemma image_compose_permutations_left: 
+lemma image_compose_permutations_left:
   assumes q: "q permutes S" shows "{q o p | p. p permutes S} = {p . p permutes S}"
 apply (rule set_ext)
 apply auto
@@ -759,7 +759,7 @@
 lemma setsum_permutations_inverse: "setsum f {p. p permutes S} = setsum (\<lambda>p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs")
 proof-
   let ?S = "{p . p permutes S}"
-have th0: "inj_on inv ?S" 
+have th0: "inj_on inv ?S"
 proof(auto simp add: inj_on_def)
   fix q r
   assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r"
@@ -828,16 +828,16 @@
     by (simp add: expand_fun_eq)
   have th1: "\<And>P Q. P \<times> Q = {(a,b). a \<in> P \<and> b \<in> Q}" by blast
   have th2: "\<And>P Q. P \<Longrightarrow> (P \<Longrightarrow> Q) \<Longrightarrow> P \<and> Q" by blast
-  show ?thesis 
-    unfolding permutes_insert    
+  show ?thesis
+    unfolding permutes_insert
     unfolding setsum_cartesian_product
     unfolding  th1[symmetric]
     unfolding th0
   proof(rule setsum_reindex)
     let ?f = "(\<lambda>(b, y). Fun.swap a b id \<circ> y)"
     let ?P = "{p. p permutes S}"
-    {fix b c p q assume b: "b \<in> insert a S" and c: "c \<in> insert a S" 
-      and p: "p permutes S" and q: "q permutes S" 
+    {fix b c p q assume b: "b \<in> insert a S" and c: "c \<in> insert a S"
+      and p: "p permutes S" and q: "q permutes S"
       and eq: "Fun.swap a b id o p = Fun.swap a c id o q"
       from p q aS have pa: "p a = a" and qa: "q a = a"
 	unfolding permutes_def by metis+
@@ -851,8 +851,8 @@
 	by (simp add: o_def)
       with bc have "b = c \<and> p = q" by blast
     }
-    
-    then show "inj_on ?f (insert a S \<times> ?P)" 
+
+    then show "inj_on ?f (insert a S \<times> ?P)"
       unfolding inj_on_def
       apply clarify by metis
   qed
--- a/src/HOL/Library/Pocklington.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Pocklington.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Library/Pocklington.thy
     ID:         $Id$
-    Author:     Amine Chaieb 
+    Author:     Amine Chaieb
 *)
 
 header {* Pocklington's Theorem for Primes *}
@@ -25,13 +25,13 @@
   shows "\<exists>q. x = y + n * q"
 using xyn xy unfolding modeq_def using nat_mod_eq_lemma by blast
 
-lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" 
+lemma nat_mod[algebra]: "[x = y] (mod n) \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)"
 unfolding modeq_def nat_mod_eq_iff ..
 
 (* Lemmas about previously defined terms.                                    *)
 
-lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)" 
-  (is "?lhs \<longleftrightarrow> ?rhs") 
+lemma prime: "prime p \<longleftrightarrow> p \<noteq> 0 \<and> p\<noteq>1 \<and> (\<forall>m. 0 < m \<and> m < p \<longrightarrow> coprime p m)"
+  (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   {assume "p=0 \<or> p=1" hence ?thesis using prime_0 prime_1 by (cases "p=0", simp_all)}
   moreover
@@ -40,7 +40,7 @@
       {fix m assume m: "m > 0" "m < p"
 	{assume "m=1" hence "coprime p m" by simp}
 	moreover
-	{assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2) 
+	{assume "p dvd m" hence "p \<le> m" using dvd_imp_le m by blast with m(2)
 	  have "coprime p m" by simp}
 	ultimately have "coprime p m" using prime_coprime[OF H, of m] by blast}
       hence ?rhs using p0 by auto}
@@ -70,11 +70,11 @@
 
 (* Congruences.                                                              *)
 
-lemma cong_mod_01[simp,presburger]: 
+lemma cong_mod_01[simp,presburger]:
   "[x = y] (mod 0) \<longleftrightarrow> x = y" "[x = y] (mod 1)" "[x = 0] (mod n) \<longleftrightarrow> n dvd x"
   by (simp_all add: modeq_def, presburger)
 
-lemma cong_sub_cases: 
+lemma cong_sub_cases:
   "[x = y] (mod n) \<longleftrightarrow> (if x <= y then [y - x = 0] (mod n) else [x - y = 0] (mod n))"
 apply (auto simp add: nat_mod)
 apply (rule_tac x="q2" in exI)
@@ -95,17 +95,17 @@
   {assume az: "a\<noteq>0"
     {assume xy: "x \<le> y" hence axy': "a*x \<le> a*y" by simp
       with axy cong_sub_cases[of "a*x" "a*y" n]  have "[a*(y - x) = 0] (mod n)"
-	by (simp only: if_True diff_mult_distrib2) 
-      hence th: "n dvd a*(y -x)" by simp 
+	by (simp only: if_True diff_mult_distrib2)
+      hence th: "n dvd a*(y -x)" by simp
       from coprime_divprod[OF th] an have "n dvd y - x"
 	by (simp add: coprime_commute)
       hence ?thesis using xy cong_sub_cases[of x y n] by simp}
     moreover
-    {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith 
+    {assume H: "\<not>x \<le> y" hence xy: "y \<le> x"  by arith
       from H az have axy': "\<not> a*x \<le> a*y" by auto
       with axy H cong_sub_cases[of "a*x" "a*y" n]  have "[a*(x - y) = 0] (mod n)"
-	by (simp only: if_False diff_mult_distrib2) 
-      hence th: "n dvd a*(x - y)" by simp 
+	by (simp only: if_False diff_mult_distrib2)
+      hence th: "n dvd a*(x - y)" by simp
       from coprime_divprod[OF th] an have "n dvd x - y"
 	by (simp add: coprime_commute)
       hence ?thesis using xy cong_sub_cases[of x y n] by simp}
@@ -121,7 +121,7 @@
 
 lemma eq_imp_cong: "a = b \<Longrightarrow> [a = b] (mod n)" by (simp add: cong_refl)
 
-lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)" 
+lemma cong_commute: "[x = y] (mod n) \<longleftrightarrow> [y = x] (mod n)"
   by (auto simp add: modeq_def)
 
 lemma cong_trans[trans]: "[x = y] (mod n) \<Longrightarrow> [y = z] (mod n) \<Longrightarrow> [x = z] (mod n)"
@@ -132,21 +132,21 @@
 proof-
   have "(x + y) mod n = (x mod n + y mod n) mod n"
     by (simp add: mod_add_left_eq[of x y n] mod_add_right_eq[of "x mod n" y n])
-  also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp 
+  also have "\<dots> = (x' mod n + y' mod n) mod n" using xx' yy' modeq_def by simp
   also have "\<dots> = (x' + y') mod n"
     by (simp add: mod_add_left_eq[of x' y' n] mod_add_right_eq[of "x' mod n" y' n])
-  finally show ?thesis unfolding modeq_def . 
+  finally show ?thesis unfolding modeq_def .
 qed
 
 lemma cong_mult: assumes xx': "[x = x'] (mod n)" and yy':"[y = y'] (mod n)"
   shows "[x * y = x' * y'] (mod n)"
 proof-
-  have "(x * y) mod n = (x mod n) * (y mod n) mod n"  
+  have "(x * y) mod n = (x mod n) * (y mod n) mod n"
     by (simp add: mod_mult_left_eq[of x y n] mod_mult_right_eq[of "x mod n" y n])
-  also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp  
+  also have "\<dots> = (x' mod n) * (y' mod n) mod n" using xx'[unfolded modeq_def] yy'[unfolded modeq_def] by simp
   also have "\<dots> = (x' * y') mod n"
     by (simp add: mod_mult_left_eq[of x' y' n] mod_mult_right_eq[of "x' mod n" y' n])
-  finally show ?thesis unfolding modeq_def . 
+  finally show ?thesis unfolding modeq_def .
 qed
 
 lemma cong_exp: "[x = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
@@ -155,19 +155,19 @@
   and yx: "y <= x" and yx': "y' <= x'"
   shows "[x - y = x' - y'] (mod n)"
 proof-
-  { fix x a x' a' y b y' b' 
+  { fix x a x' a' y b y' b'
     have "(x::nat) + a = x' + a' \<Longrightarrow> y + b = y' + b' \<Longrightarrow> y <= x \<Longrightarrow> y' <= x'
       \<Longrightarrow> (x - y) + (a + b') = (x' - y') + (a' + b)" by arith}
   note th = this
-  from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2" 
+  from xx' yy' obtain q1 q2 q1' q2' where q12: "x + n*q1 = x'+n*q2"
     and q12': "y + n*q1' = y'+n*q2'" unfolding nat_mod by blast+
   from th[OF q12 q12' yx yx']
-  have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')" 
+  have "(x - y) + n*(q1 + q2') = (x' - y') + n*(q2 + q1')"
     by (simp add: right_distrib)
   thus ?thesis unfolding nat_mod by blast
 qed
 
-lemma cong_mult_lcancel_eq: assumes an: "coprime a n" 
+lemma cong_mult_lcancel_eq: assumes an: "coprime a n"
   shows "[a * x = a * y] (mod n) \<longleftrightarrow> [x = y] (mod n)" (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs .
@@ -176,23 +176,23 @@
   from cong_mult_rcancel[OF an H'] show ?rhs  .
 qed
 
-lemma cong_mult_rcancel_eq: assumes an: "coprime a n" 
+lemma cong_mult_rcancel_eq: assumes an: "coprime a n"
   shows "[x * a = y * a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
 using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute)
 
-lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)" 
+lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \<longleftrightarrow> [x = y] (mod n)"
   by (simp add: nat_mod)
 
 lemma cong_add_rcancel_eq: "[x + a = y + a] (mod n) \<longleftrightarrow> [x = y] (mod n)"
   by (simp add: nat_mod)
 
-lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)" 
+lemma cong_add_rcancel: "[x + a = y + a] (mod n) \<Longrightarrow> [x = y] (mod n)"
   by (simp add: nat_mod)
 
 lemma cong_add_lcancel: "[a + x = a + y] (mod n) \<Longrightarrow> [x = y] (mod n)"
   by (simp add: nat_mod)
 
-lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)" 
+lemma cong_add_lcancel_eq_0: "[a + x = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
   by (simp add: nat_mod)
 
 lemma cong_add_rcancel_eq_0: "[x + a = a] (mod n) \<longleftrightarrow> [x = 0] (mod n)"
@@ -200,14 +200,14 @@
 
 lemma cong_imp_eq: assumes xn: "x < n" and yn: "y < n" and xy: "[x = y] (mod n)"
   shows "x = y"
-  using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] . 
+  using xy[unfolded modeq_def mod_less[OF xn] mod_less[OF yn]] .
 
 lemma cong_divides_modulus: "[x = y] (mod m) \<Longrightarrow> n dvd m ==> [x = y] (mod n)"
   apply (auto simp add: nat_mod dvd_def)
   apply (rule_tac x="k*q1" in exI)
   apply (rule_tac x="k*q2" in exI)
   by simp
-  
+
 lemma cong_0_divides: "[x = 0] (mod n) \<longleftrightarrow> n dvd x" by simp
 
 lemma cong_1_divides:"[x = 1] (mod n) ==> n dvd x - 1"
@@ -220,13 +220,13 @@
 apply (rule_tac x="k + q2 - q1" in exI, simp add: add_mult_distrib2 diff_mult_distrib2)
 done
 
-lemma cong_coprime: assumes xy: "[x = y] (mod n)" 
+lemma cong_coprime: assumes xy: "[x = y] (mod n)"
   shows "coprime n x \<longleftrightarrow> coprime n y"
 proof-
   {assume "n=0" hence ?thesis using xy by simp}
   moreover
   {assume nz: "n \<noteq> 0"
-  have "coprime n x \<longleftrightarrow> coprime (x mod n) n" 
+  have "coprime n x \<longleftrightarrow> coprime (x mod n) n"
     by (simp add: coprime_mod[OF nz, of x] coprime_commute[of n x])
   also have "\<dots> \<longleftrightarrow> coprime (y mod n) n" using xy[unfolded modeq_def] by simp
   also have "\<dots> \<longleftrightarrow> coprime y n" by (simp add: coprime_mod[OF nz, of y])
@@ -236,7 +236,7 @@
 
 lemma cong_mod: "~(n = 0) \<Longrightarrow> [a mod n = a] (mod n)" by (simp add: modeq_def)
 
-lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0) 
+lemma mod_mult_cong: "~(a = 0) \<Longrightarrow> ~(b = 0)
   \<Longrightarrow> [x mod (a * b) = y] (mod a) \<longleftrightarrow> [x = y] (mod a)"
   by (simp add: modeq_def mod_mult2_eq mod_add_left_eq)
 
@@ -258,10 +258,10 @@
 
 lemma cong_to_1: "[a = 1] (mod n) \<longleftrightarrow> a = 0 \<and> n = 1 \<or> (\<exists>m. a = 1 + m * n)"
 proof-
-  {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis 
+  {assume "n = 0 \<or> n = 1\<or> a = 0 \<or> a = 1" hence ?thesis
       apply (cases "n=0", simp_all add: cong_commute)
       apply (cases "n=1", simp_all add: cong_commute modeq_def)
-      apply arith 
+      apply arith
       by (cases "a=1", simp_all add: modeq_def cong_commute)}
   moreover
   {assume n: "n\<noteq>0" "n\<noteq>1" and a:"a\<noteq>0" "a \<noteq> 1" hence a': "a \<ge> 1" by simp
@@ -277,7 +277,7 @@
   {assume "a=0" hence ?thesis using an by (simp add: modeq_def)}
   moreover
   {assume az: "a\<noteq>0"
-  from bezout_add_strong[OF az, of n] 
+  from bezout_add_strong[OF az, of n]
   obtain d x y where dxy: "d dvd a" "d dvd n" "a*x = n*y + d" by blast
   from an[unfolded coprime, rule_format, of d] dxy(1,2) have d1: "d = 1" by blast
   hence "a*x*b = (n*y + 1)*b" using dxy(3) by simp
@@ -315,12 +315,12 @@
   from pa have ap: "coprime a p" by (simp add: coprime_commute)
   from prime_coprime[OF p, of x] dvd_imp_le[of p x] x0 xp have px:"coprime x p"
     by (auto simp add: coprime_commute)
-  from cong_solve_unique[OF px p01(1)] 
+  from cong_solve_unique[OF px p01(1)]
   obtain y where y: "y < p" "[x * y = a] (mod p)" "\<forall>z. z < p \<and> [x * z = a] (mod p) \<longrightarrow> z = y" by blast
   {assume y0: "y = 0"
     with y(2) have th: "p dvd a" by (simp add: cong_commute[of 0 a p])
     with p coprime_prime[OF pa, of p] have False by simp}
-  with y show ?thesis unfolding Ex1_def using neq0_conv by blast 
+  with y show ?thesis unfolding Ex1_def using neq0_conv by blast
 qed
 lemma cong_unique_inverse_prime:
   assumes p: "prime p" and x0: "0 < x" and xp: "x < p"
@@ -329,13 +329,13 @@
 
 (* Forms of the Chinese remainder theorem.                                   *)
 
-lemma cong_chinese: 
-  assumes ab: "coprime a b" and  xya: "[x = y] (mod a)" 
+lemma cong_chinese:
+  assumes ab: "coprime a b" and  xya: "[x = y] (mod a)"
   and xyb: "[x = y] (mod b)"
   shows "[x = y] (mod a*b)"
   using ab xya xyb
-  by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b] 
-    cong_sub_cases[of x y "a*b"]) 
+  by (simp add: cong_sub_cases[of x y a] cong_sub_cases[of x y b]
+    cong_sub_cases[of x y "a*b"])
 (cases "x \<le> y", simp_all add: divides_mul[of a _ b])
 
 lemma chinese_remainder_unique:
@@ -343,9 +343,9 @@
   shows "\<exists>!x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
 proof-
   from az bz have abpos: "a*b > 0" by simp
-  from chinese_remainder[OF ab az bz] obtain x q1 q2 where 
+  from chinese_remainder[OF ab az bz] obtain x q1 q2 where
     xq12: "x = m + q1 * a" "x = n + q2 * b" by blast
-  let ?w = "x mod (a*b)" 
+  let ?w = "x mod (a*b)"
   have wab: "?w < a*b" by (simp add: mod_less_divisor[OF abpos])
   from xq12(1) have "?w mod a = ((m + q1 * a) mod (a*b)) mod a" by simp
   also have "\<dots> = m mod a" apply (simp add: mod_mult2_eq)
@@ -361,19 +361,19 @@
   {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)"
     with th1 th2 have H': "[y = ?w] (mod a)" "[y = ?w] (mod b)"
       by (simp_all add: modeq_def)
-    from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab] 
+    from cong_chinese[OF ab H'] mod_less[OF H(1)] mod_less[OF wab]
     have "y = ?w" by (simp add: modeq_def)}
   with th1 th2 wab show ?thesis by blast
 qed
 
 lemma chinese_remainder_coprime_unique:
-  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0" 
+  assumes ab: "coprime a b" and az: "a \<noteq> 0" and bz: "b \<noteq> 0"
   and ma: "coprime m a" and nb: "coprime n b"
   shows "\<exists>!x. coprime x (a * b) \<and> x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
 proof-
   let ?P = "\<lambda>x. x < a * b \<and> [x = m] (mod a) \<and> [x = n] (mod b)"
   from chinese_remainder_unique[OF ab az bz]
-  obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)" 
+  obtain x where x: "x < a * b" "[x = m] (mod a)" "[x = n] (mod b)"
     "\<forall>y. ?P y \<longrightarrow> y = x" by blast
   from ma nb cong_coprime[OF x(2)] cong_coprime[OF x(3)]
   have "coprime x a" "coprime x b" by (simp_all add: coprime_commute)
@@ -396,7 +396,7 @@
 declare coprime_1[presburger]
 lemma phi_1[simp]: "\<phi> 1 = 1"
 proof-
-  {fix m 
+  {fix m
     have "0 < m \<and> m <= 1 \<and> coprime m 1 \<longleftrightarrow> m = 1" by presburger }
   thus ?thesis by (simp add: phi_def)
 qed
@@ -424,7 +424,7 @@
 lemma phi_another: assumes n: "n\<noteq>1"
   shows "\<phi> n = card {m. 0 < m \<and> m < n \<and> coprime m n }"
 proof-
-  {fix m 
+  {fix m
     from n have "0 < m \<and> m < n \<and> coprime m n \<longleftrightarrow> coprime m n \<and> m < n"
       by (cases "m=0", auto)}
   thus ?thesis unfolding phi_alt by auto
@@ -440,11 +440,11 @@
 lemma stupid[simp]: "{m. (0::nat) < m \<and> m < n} = {1..<n}"
   by auto
 
-lemma phi_limit_strong: assumes n: "n\<noteq>1" 
+lemma phi_limit_strong: assumes n: "n\<noteq>1"
   shows "\<phi>(n) \<le> n - 1"
 proof-
   show ?thesis
-    unfolding phi_another[OF n] finite_number_segment[of n, symmetric] 
+    unfolding phi_another[OF n] finite_number_segment[of n, symmetric]
     by (rule card_mono[of "{m. 0 < m \<and> m < n}" "{m. 0 < m \<and> m < n \<and> coprime m n}"], auto)
 qed
 
@@ -452,7 +452,7 @@
   shows "\<phi>(n) \<ge> 1"
 proof-
   let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
-  from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt 
+  from card_0_eq[of ?S] n have "\<phi> n \<noteq> 0" unfolding phi_alt
     apply auto
     apply (cases "n=1", simp_all)
     apply (rule exI[where x=1], simp)
@@ -466,10 +466,10 @@
 lemma phi_lowerbound_2: assumes n: "3 <= n" shows "2 <= \<phi> (n)"
 proof-
   let ?S = "{ m. 0 < m \<and> m <= n \<and> coprime m n }"
-  have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"] 
+  have inS: "{1, n - 1} \<subseteq> ?S" using n coprime_plus1[of "n - 1"]
     by (auto simp add: coprime_commute)
   from n have c2: "card {1, n - 1} = 2" by (auto simp add: card_insert_if)
-  from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis 
+  from card_mono[of ?S "{1, n - 1}", simplified inS c2] show ?thesis
     unfolding phi_def by auto
 qed
 
@@ -483,12 +483,12 @@
     let ?S' = "{m. 0 < m \<and> m < n \<and> coprime m n}"
     have fS':"finite ?S'" apply (rule finite_subset[of ?S' ?S]) by auto
     {assume H: "\<phi> n = n - 1 \<and> n\<noteq>0 \<and> n\<noteq>1"
-      hence ceq: "card ?S' = card ?S" 
+      hence ceq: "card ?S' = card ?S"
       using n finite_number_segment[of n] phi_another[OF n(2)] by simp
       {fix m assume m: "0 < m" "m < n" "\<not> coprime m n"
 	hence mS': "m \<notin> ?S'" by auto
 	have "insert m ?S' \<le> ?S" using m by auto
-	from m have "card (insert m ?S') \<le> card ?S" 
+	from m have "card (insert m ?S') \<le> card ?S"
 	  by - (rule card_mono[of ?S "insert m ?S'"], auto)
 	hence False
 	  unfolding card_insert_disjoint[of "?S'" m, OF fS' mS'] ceq
@@ -497,7 +497,7 @@
       hence "prime n" unfolding prime using n by (simp add: coprime_commute)}
     moreover
     {assume H: "prime n"
-      hence "?S = ?S'" unfolding prime using n 
+      hence "?S = ?S'" unfolding prime using n
 	by (auto simp add: coprime_commute)
       hence "card ?S = card ?S'" by simp
       hence "\<phi> n = n - 1" unfolding phi_another[OF n(2)] by simp}
@@ -510,7 +510,7 @@
 lemma phi_multiplicative: assumes ab: "coprime a b"
   shows "\<phi> (a * b) = \<phi> a * \<phi> b"
 proof-
-  {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1" 
+  {assume "a = 0 \<or> b = 0 \<or> a = 1 \<or> b = 1"
     hence ?thesis
       by (cases "a=0", simp, cases "b=0", simp, cases"a=1", simp_all) }
   moreover
@@ -524,11 +524,11 @@
 	hence x': "coprime x (a*b)" "x < a*b" by simp_all
 	hence xab: "coprime x a" "coprime x b" by (simp_all add: coprime_mul_eq)
 	from mod_less_divisor a b have xab':"x mod a < a" "x mod b < b" by auto
-	from xab xab' have "?f x \<in> (?S a \<times> ?S b)" 
+	from xab xab' have "?f x \<in> (?S a \<times> ?S b)"
 	  by (simp add: coprime_mod[OF a(1)] coprime_mod[OF b(1)])}
       moreover
       {fix x y assume x: "x \<in> ?S a" and y: "y \<in> ?S b"
-	hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all	
+	hence x': "coprime x a" "x < a" and y': "coprime y b" "y < b" by simp_all
 	from chinese_remainder_coprime_unique[OF ab a(1) b(1) x'(1) y'(1)]
 	obtain z where z: "coprime z (a * b)" "z < a * b" "[z = x] (mod a)"
 	  "[z = y] (mod b)" by blast
@@ -540,9 +540,9 @@
     have finj: "inj_on ?f (?S (a*b))"
       unfolding inj_on_def
     proof(clarify)
-      fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)" 
+      fix x y assume H: "coprime x (a * b)" "x < a * b" "coprime y (a * b)"
 	"y < a * b" "x mod a = y mod a" "x mod b = y mod b"
-      hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b" 
+      hence cp: "coprime x a" "coprime x b" "coprime y a" "coprime y b"
 	by (simp_all add: coprime_mul_eq)
       from chinese_remainder_coprime_unique[OF ab a(1) b(1) cp(3,4)] H
       show "x = y" unfolding modeq_def by blast
@@ -593,7 +593,7 @@
     have cardfS: "\<phi> n = card ?S" unfolding phi_alt ..
     {fix m assume m: "m \<in> ?S"
       hence "coprime m n" by simp
-      with coprime_mul[of n a m] an have "coprime (a*m) n" 
+      with coprime_mul[of n a m] an have "coprime (a*m) n"
 	by (simp add: coprime_commute)}
     hence Sn: "\<forall>m\<in> ?S. coprime (a*m) n " by blast
     from coprime_nproduct[OF fS, of n "\<lambda>m. m"] have nP:"coprime ?P n"
@@ -606,7 +606,7 @@
       hence hS: "?h ` ?S = ?S"by (auto simp add: image_iff)
       have "a\<noteq>0" using an n1 nz apply- apply (rule ccontr) by simp
       hence inj: "inj_on (op * a) ?S" unfolding inj_on_def by simp
-      
+
       have eq0: "fold_image op * (?h \<circ> op * a) 1 {m. coprime m n \<and> m < n} =
      fold_image op * (\<lambda>m. m) 1 {m. coprime m n \<and> m < n}"
       proof (rule fold_image_eq_general[where h="?h o (op * a)"])
@@ -614,12 +614,12 @@
       next
 	{fix y assume yS: "y \<in> ?S" hence y: "coprime y n" "y < n" by simp_all
 	  from cong_solve_unique[OF an nz, of y]
-	  obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast  
+	  obtain x where x:"x < n" "[a * x = y] (mod n)" "\<forall>z. z < n \<and> [a * z = y] (mod n) \<longrightarrow> z=x" by blast
 	  from cong_coprime[OF x(2)] y(1)
 	  have xm: "coprime x n" by (simp add: coprime_mul_eq coprime_commute)
 	  {fix z assume "z \<in> ?S" "(?h \<circ> op * a) z = y"
 	    hence z: "coprime z n" "z < n" "(?h \<circ> op * a) z = y" by simp_all
-	    from x(3)[rule_format, of z] z(2,3) have "z=x" 
+	    from x(3)[rule_format, of z] z(2,3) have "z=x"
 	      unfolding modeq_def mod_less[OF y(2)] by simp}
 	  with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
 	    unfolding modeq_def mod_less[OF y(2)] by auto }
@@ -630,7 +630,7 @@
 	  hence x: "coprime x n" "x < n" by simp_all
 	  with an have "coprime (a*x) n"
 	    by (simp add: coprime_mul_eq[of n a x] coprime_commute)
-	  hence "?h (a*x) \<in> ?S" using nz 
+	  hence "?h (a*x) \<in> ?S" using nz
 	    by (simp add: coprime_mod[OF nz] mod_less_divisor)}
 	thus " \<forall>x\<in>{m. coprime m n \<and> m < n}.
        ((\<lambda>m. m mod n) \<circ> op * a) x \<in> {m. coprime m n \<and> m < n} \<and>
@@ -643,7 +643,7 @@
       also have "[setprod (?h o (op * a)) ?S = ?P ] (mod n)"
 	using eq0 fS an by (simp add: setprod_def modeq_def o_def)
       finally show "[?P*a^ (\<phi> n) = ?P*1] (mod n)"
-	unfolding cardfS mult_commute[of ?P "a^ (card ?S)"] 
+	unfolding cardfS mult_commute[of ?P "a^ (card ?S)"]
 	  nproduct_cmul[OF fS, symmetric] mult_1_right by simp
     qed
     from cong_mult_lcancel[OF nP Paphi] have ?thesis . }
@@ -670,18 +670,18 @@
     from m obtain m' where m': "m = Suc m'" by (cases m, blast+)
     {fix d
       assume d: "d dvd a" "d dvd n"
-      from n have n1: "1 < n" by arith      
+      from n have n1: "1 < n" by arith
       from am mod_less[OF n1] have am1: "a^m mod n = 1" unfolding modeq_def by simp
       from dvd_mult2[OF d(1), of "a^m'"] have dam:"d dvd a^m" by (simp add: m')
       from dvd_mod_iff[OF d(2), of "a^m"] dam am1
       have "d = 1" by simp }
     hence ?thesis unfolding coprime by auto
   }
-  ultimately show ?thesis by blast 
+  ultimately show ?thesis by blast
 qed
 
 lemma lucas_weak:
-  assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)" 
+  assumes n: "n \<ge> 2" and an:"[a^(n - 1) = 1] (mod n)"
   and nm: "\<forall>m. 0 <m \<and> m < n - 1 \<longrightarrow> \<not> [a^m = 1] (mod n)"
   shows "prime n"
 proof-
@@ -696,7 +696,7 @@
   with phi_prime[of n] n1(1,2) show ?thesis by simp
 qed
 
-lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))" 
+lemma nat_exists_least_iff: "(\<exists>(n::nat). P n) \<longleftrightarrow> (\<exists>n. P n \<and> (\<forall>m < n. \<not> P m))"
   (is "?lhs \<longleftrightarrow> ?rhs")
 proof
   assume ?rhs thus ?lhs by blast
@@ -708,11 +708,11 @@
   with LeastI_ex[OF H] show ?rhs by blast
 qed
 
-lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))" 
+lemma nat_exists_least_iff': "(\<exists>(n::nat). P n) \<longleftrightarrow> (P (Least P) \<and> (\<forall>m < (Least P). \<not> P m))"
   (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   {assume ?rhs hence ?lhs by blast}
-  moreover 
+  moreover
   { assume H: ?lhs then obtain n where n: "P n" by blast
     let ?x = "Least P"
     {fix m assume m: "m < ?x"
@@ -720,13 +720,13 @@
     with LeastI_ex[OF H] have ?rhs by blast}
   ultimately show ?thesis by blast
 qed
- 
+
 lemma power_mod: "((x::nat) mod m)^n mod m = x^n mod m"
 proof(induct n)
   case 0 thus ?case by simp
 next
-  case (Suc n) 
-  have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m" 
+  case (Suc n)
+  have "(x mod m)^(Suc n) mod m = ((x mod m) * (((x mod m) ^ n) mod m)) mod m"
     by (simp add: mod_mult_right_eq[symmetric])
   also have "\<dots> = ((x mod m) * (x^n mod m)) mod m" using Suc.hyps by simp
   also have "\<dots> = x^(Suc n) mod m"
@@ -735,38 +735,38 @@
 qed
 
 lemma lucas:
-  assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)" 
+  assumes n2: "n \<ge> 2" and an1: "[a^(n - 1) = 1] (mod n)"
   and pn: "\<forall>p. prime p \<and> p dvd n - 1 \<longrightarrow> \<not> [a^((n - 1) div p) = 1] (mod n)"
   shows "prime n"
 proof-
   from n2 have n01: "n\<noteq>0" "n\<noteq>1" "n - 1 \<noteq> 0" by arith+
   from mod_less_divisor[of n 1] n01 have onen: "1 mod n = 1" by simp
-  from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1] 
+  from lucas_coprime_lemma[OF n01(3) an1] cong_coprime[OF an1]
   have an: "coprime a n" "coprime (a^(n - 1)) n" by (simp_all add: coprime_commute)
   {assume H0: "\<exists>m. 0 < m \<and> m < n - 1 \<and> [a ^ m = 1] (mod n)" (is "EX m. ?P m")
-    from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where 
+    from H0[unfolded nat_exists_least_iff[of ?P]] obtain m where
       m: "0 < m" "m < n - 1" "[a ^ m = 1] (mod n)" "\<forall>k <m. \<not>?P k" by blast
-    {assume nm1: "(n - 1) mod m > 0" 
-      from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast 
+    {assume nm1: "(n - 1) mod m > 0"
+      from mod_less_divisor[OF m(1)] have th0:"(n - 1) mod m < m" by blast
       let ?y = "a^ ((n - 1) div m * m)"
       note mdeq = mod_div_equality[of "(n - 1)" m]
-      from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]], 
+      from coprime_exp[OF an(1)[unfolded coprime_commute[of a n]],
 	of "(n - 1) div m * m"]
-      have yn: "coprime ?y n" by (simp add: coprime_commute) 
-      have "?y mod n = (a^m)^((n - 1) div m) mod n" 
+      have yn: "coprime ?y n" by (simp add: coprime_commute)
+      have "?y mod n = (a^m)^((n - 1) div m) mod n"
 	by (simp add: algebra_simps power_mult)
-      also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n" 
+      also have "\<dots> = (a^m mod n)^((n - 1) div m) mod n"
 	using power_mod[of "a^m" n "(n - 1) div m"] by simp
-      also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen 
+      also have "\<dots> = 1" using m(3)[unfolded modeq_def onen] onen
 	by (simp add: power_Suc0)
-      finally have th3: "?y mod n = 1"  . 
-      have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)" 
+      finally have th3: "?y mod n = 1"  .
+      have th2: "[?y * a ^ ((n - 1) mod m) = ?y* 1] (mod n)"
 	using an1[unfolded modeq_def onen] onen
 	  mod_div_equality[of "(n - 1)" m, symmetric]
 	by (simp add:power_add[symmetric] modeq_def th3 del: One_nat_def)
       from cong_mult_lcancel[of ?y n "a^((n - 1) mod m)" 1, OF yn th2]
-      have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)"  . 
-      from m(4)[rule_format, OF th0] nm1 
+      have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)"  .
+      from m(4)[rule_format, OF th0] nm1
 	less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1
       have False by blast }
     hence "(n - 1) mod m = 0" by auto
@@ -781,7 +781,7 @@
     also have "\<dots> = ((a^m)^(r div p)) mod n" by (simp add: power_mult)
     also have "\<dots> = ((a^m mod n)^(r div p)) mod n" using power_mod[of "a^m" "n" "r div p" ] ..
     also have "\<dots> = 1" using m(3) onen by (simp add: modeq_def power_Suc0)
-    finally have "[(a ^ ((n - 1) div p))= 1] (mod n)" 
+    finally have "[(a ^ ((n - 1) div p))= 1] (mod n)"
       using onen by (simp add: modeq_def)
     with pn[rule_format, OF th] have False by blast}
   hence th: "\<forall>m. 0 < m \<and> m < n - 1 \<longrightarrow> \<not> [a ^ m = 1] (mod n)" by blast
@@ -795,7 +795,7 @@
 (* This has the expected properties.                                         *)
 
 lemma coprime_ord:
-  assumes na: "coprime n a" 
+  assumes na: "coprime n a"
   shows "ord n a > 0 \<and> [a ^(ord n a) = 1] (mod n) \<and> (\<forall>m. 0 < m \<and> m < ord n a \<longrightarrow> \<not> [a^ m = 1] (mod n))"
 proof-
   let ?P = "\<lambda>d. 0 < d \<and> [a ^ d = 1] (mod n)"
@@ -803,12 +803,12 @@
   from na have o: "ord n a = Least ?P" by (simp add: ord_def)
   {assume "n=0 \<or> n=1" with na have "\<exists>m>0. ?P m" apply auto apply (rule exI[where x=1]) by (simp  add: modeq_def)}
   moreover
-  {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith 
+  {assume "n\<noteq>0 \<and> n\<noteq>1" hence n2:"n \<ge> 2" by arith
     from na have na': "coprime a n" by (simp add: coprime_commute)
     from phi_lowerbound_1[OF n2] fermat_little[OF na']
     have ex: "\<exists>m>0. ?P m" by - (rule exI[where x="\<phi> n"], auto) }
   ultimately have ex: "\<exists>m>0. ?P m" by blast
-  from nat_exists_least_iff'[of ?P] ex na show ?thesis 
+  from nat_exists_least_iff'[of ?P] ex na show ?thesis
     unfolding o[symmetric] by auto
 qed
 (* With the special value 0 for non-coprime case, it's more convenient.      *)
@@ -818,8 +818,8 @@
 using coprime_ord[of n a]
 by (blast, simp add: ord_def modeq_def)
 
-lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast 
-lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)" 
+lemma ord: "[a^(ord n a) = 1] (mod n)" using ord_works by blast
+lemma ord_minimal: "0 < m \<Longrightarrow> m < ord n a \<Longrightarrow> ~[a^m = 1] (mod n)"
   using ord_works by blast
 lemma ord_eq_0: "ord n a = 0 \<longleftrightarrow> ~coprime n a"
 by (cases "coprime n a", simp add: neq0_conv coprime_ord, simp add: neq0_conv ord_def)
@@ -831,20 +831,20 @@
   then obtain k where "d = ord n a * k" unfolding dvd_def by blast
   hence "[a ^ d = (a ^ (ord n a) mod n)^k] (mod n)"
     by (simp add : modeq_def power_mult power_mod)
-  also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)" 
-    using ord[of a n, unfolded modeq_def] 
+  also have "[(a ^ (ord n a) mod n)^k = 1] (mod n)"
+    using ord[of a n, unfolded modeq_def]
     by (simp add: modeq_def power_mod power_Suc0)
   finally  show ?lhs .
-next 
+next
   assume lh: ?lhs
   { assume H: "\<not> coprime n a"
     hence o: "ord n a = 0" by (simp add: ord_def)
     {assume d: "d=0" with o H have ?rhs by (simp add: modeq_def)}
     moreover
     {assume d0: "d\<noteq>0" then obtain d' where d': "d = Suc d'" by (cases d, auto)
-      from H[unfolded coprime] 
-      obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto 
-      from lh[unfolded nat_mod] 
+      from H[unfolded coprime]
+      obtain p where p: "p dvd n" "p dvd a" "p \<noteq> 1" by auto
+      from lh[unfolded nat_mod]
       obtain q1 q2 where q12:"a ^ d + n * q1 = 1 + n * q2" by blast
       hence "a ^ d + n * q1 - n * q2 = 1" by simp
       with nat_dvd_diff [OF dvd_add [OF divides_rexp[OF p(2), of d'] dvd_mult2[OF p(1), of q1]] dvd_mult2[OF p(1), of q2]] d' have "p dvd 1" by simp
@@ -856,13 +856,13 @@
     let ?o = "ord n a"
     let ?q = "d div ord n a"
     let ?r = "d mod ord n a"
-    from cong_exp[OF ord[of a n], of ?q] 
+    from cong_exp[OF ord[of a n], of ?q]
     have eqo: "[(a^?o)^?q = 1] (mod n)"  by (simp add: modeq_def power_Suc0)
     from H have onz: "?o \<noteq> 0" by (simp add: ord_eq_0)
     hence op: "?o > 0" by simp
     from mod_div_equality[of d "ord n a"] lh
     have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute)
-    hence "[(a^?o)^?q * (a^?r) = 1] (mod n)" 
+    hence "[(a^?o)^?q * (a^?r) = 1] (mod n)"
       by (simp add: modeq_def power_mult[symmetric] power_add[symmetric])
     hence th: "[a^?r = 1] (mod n)"
       using eqo mod_mult_left_eq[of "(a^?o)^?q" "a^?r" n]
@@ -870,9 +870,9 @@
       by (simp add: mod_mult_left_eq[symmetric])
     {assume r: "?r = 0" hence ?rhs by (simp add: dvd_eq_mod_eq_0)}
     moreover
-    {assume r: "?r \<noteq> 0" 
+    {assume r: "?r \<noteq> 0"
       with mod_less_divisor[OF op, of d] have r0o:"?r >0 \<and> ?r < ?o" by simp
-      from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th     
+      from conjunct2[OF ord_works[of a n], rule_format, OF r0o] th
       have ?rhs by blast}
     ultimately have ?rhs by blast}
   ultimately  show ?rhs by blast
@@ -880,18 +880,18 @@
 
 lemma order_divides_phi: "coprime n a \<Longrightarrow> ord n a dvd \<phi> n"
 using ord_divides fermat_little coprime_commute by simp
-lemma order_divides_expdiff: 
+lemma order_divides_expdiff:
   assumes na: "coprime n a"
   shows "[a^d = a^e] (mod n) \<longleftrightarrow> [d = e] (mod (ord n a))"
 proof-
-  {fix n a d e 
+  {fix n a d e
     assume na: "coprime n a" and ed: "(e::nat) \<le> d"
     hence "\<exists>c. d = e + c" by arith
     then obtain c where c: "d = e + c" by arith
     from na have an: "coprime a n" by (simp add: coprime_commute)
-    from coprime_exp[OF na, of e] 
+    from coprime_exp[OF na, of e]
     have aen: "coprime (a^e) n" by (simp add: coprime_commute)
-    from coprime_exp[OF na, of c] 
+    from coprime_exp[OF na, of c]
     have acn: "coprime (a^c) n" by (simp add: coprime_commute)
     have "[a^d = a^e] (mod n) \<longleftrightarrow> [a^(e + c) = a^(e + 0)] (mod n)"
       using c by simp
@@ -923,29 +923,29 @@
   moreover
   {assume n: "n\<noteq>0" "n\<noteq>1"
     {assume pn: "prime n"
-      
+
       from pn[unfolded prime_def] have "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
-	using n 
+	using n
 	apply (cases "n = 0 \<or> n=1",simp)
 	by (clarsimp, erule_tac x="p" in allE, auto)}
     moreover
     {assume H: "\<forall>p. prime p \<and> p dvd n \<longrightarrow> p = n"
       from n have n1: "n > 1" by arith
       {fix m assume m: "m dvd n" "m\<noteq>1"
-	from prime_factor[OF m(2)] obtain p where 
+	from prime_factor[OF m(2)] obtain p where
 	  p: "prime p" "p dvd m" by blast
-	from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast  
+	from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast
 	with p(2) have "n dvd m"  by simp
 	hence "m=n"  using dvd_anti_sym[OF m(1)] by simp }
       with n1 have "prime n"  unfolding prime_def by auto }
-    ultimately have ?thesis using n by blast} 
-  ultimately       show ?thesis by auto 
+    ultimately have ?thesis using n by blast}
+  ultimately       show ?thesis by auto
 qed
 
 lemma prime_divisor_sqrt:
   "prime n \<longleftrightarrow> n \<noteq> 1 \<and> (\<forall>d. d dvd n \<and> d^2 \<le> n \<longrightarrow> d = 1)"
 proof-
-  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 
+  {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1
     by (auto simp add: nat_power_eq_0_iff)}
   moreover
   {assume n: "n\<noteq>0" "n\<noteq>1"
@@ -953,7 +953,7 @@
     {fix d assume d: "d dvd n" "d^2 \<le> n" and H: "\<forall>m. m dvd n \<longrightarrow> m=1 \<or> m=n"
       from H d have d1n: "d = 1 \<or> d=n" by blast
       {assume dn: "d=n"
-	have "n^2 > n*1" using n 
+	have "n^2 > n*1" using n
 	  by (simp add: power2_eq_square mult_less_cancel1)
 	with dn d(2) have "d=1" by simp}
       with d1n have "d = 1" by blast  }
@@ -978,20 +978,20 @@
   ultimately show ?thesis by auto
 qed
 lemma prime_prime_factor_sqrt:
-  "prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)" 
+  "prime n \<longleftrightarrow> n \<noteq> 0 \<and> n \<noteq> 1 \<and> \<not> (\<exists>p. prime p \<and> p dvd n \<and> p^2 \<le> n)"
   (is "?lhs \<longleftrightarrow>?rhs")
 proof-
   {assume "n=0 \<or> n=1" hence ?thesis using prime_0 prime_1 by auto}
   moreover
   {assume n: "n\<noteq>0" "n\<noteq>1"
     {assume H: ?lhs
-      from H[unfolded prime_divisor_sqrt] n 
+      from H[unfolded prime_divisor_sqrt] n
       have ?rhs  apply clarsimp by (erule_tac x="p" in allE, simp add: prime_1)
     }
     moreover
     {assume H: ?rhs
       {fix d assume d: "d dvd n" "d^2 \<le> n" "d\<noteq>1"
-	from prime_factor[OF d(3)] 
+	from prime_factor[OF d(3)]
 	obtain p where p: "prime p" "p dvd d" by blast
 	from n have np: "n > 0" by arith
 	from d(1) n have "d \<noteq> 0" by - (rule ccontr, auto)
@@ -1012,7 +1012,7 @@
   shows "[p = 1] (mod q)"
 proof-
   from pp prime_0 prime_1 have p01: "p \<noteq> 0" "p \<noteq> 1" by - (rule ccontr, simp)+
-  from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def] 
+  from cong_1_divides[OF an, unfolded nqr, unfolded dvd_def]
   obtain k where k: "a ^ (q * r) - 1 = n*k" by blast
   from pn[unfolded dvd_def] obtain l where l: "n = p*l" by blast
   {assume a0: "a = 0"
@@ -1026,7 +1026,7 @@
   hence odq: "ord p (a^r) dvd q"
     unfolding ord_divides[symmetric] power_mult[symmetric] nat_mod  by blast
   from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast
-  {assume d1: "d \<noteq> 1" 
+  {assume d1: "d \<noteq> 1"
     from prime_factor[OF d1] obtain P where P: "prime P" "P dvd d" by blast
     from d dvd_mult[OF P(2), of "ord p (a^r)"] have Pq: "P dvd q" by simp
     from aq P(1) Pq have caP:"coprime (a^ ((n - 1) div P) - 1) n" by blast
@@ -1037,16 +1037,16 @@
     have "ord p (a^r) * t*r = r * ord p (a^r) * t" by algebra
     hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t"
       by (simp only: power_mult)
-    have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)" 
+    have "[((a ^ r) ^ ord p (a^r)) ^ t= 1^t] (mod p)"
       by (rule cong_exp, rule ord)
-    then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)" 
+    then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)"
       by (simp add: power_Suc0)
     from cong_1_divides[OF th] exps have pd0: "p dvd a^(ord p (a^r) * t*r) - 1" by simp
     from nqr s s' have "(n - 1) div P = ord p (a^r) * t*r" using P0 by simp
     with caP have "coprime (a^(ord p (a^r) * t*r) - 1) n" by simp
     with p01 pn pd0 have False unfolding coprime by auto}
-  hence d1: "d = 1" by blast 
-  hence o: "ord p (a^r) = q" using d by simp  
+  hence d1: "d = 1" by blast
+  hence o: "ord p (a^r) = q" using d by simp
   from pp phi_prime[of p] have phip: " \<phi> p = p - 1" by simp
   {fix d assume d: "d dvd p" "d dvd a" "d \<noteq> 1"
     from pp[unfolded prime_def] d have dp: "d = p" by blast
@@ -1054,14 +1054,14 @@
     with divides_rexp[OF d(2)[unfolded dp], of "n - 2"]
     have th0: "p dvd a ^ (n - 1)" by simp
     from n have n0: "n \<noteq> 0" by simp
-    from d(2) an n12[symmetric] have a0: "a \<noteq> 0" 
+    from d(2) an n12[symmetric] have a0: "a \<noteq> 0"
       by - (rule ccontr, simp add: modeq_def)
     have th1: "a^ (n - 1) \<noteq> 0" using n d(2) dp a0 by (auto simp add: neq0_conv)
-    from coprime_minus1[OF th1, unfolded coprime] 
+    from coprime_minus1[OF th1, unfolded coprime]
       dvd_trans[OF pn cong_1_divides[OF an]] th0 d(3) dp
     have False by auto}
-  hence cpa: "coprime p a" using coprime by auto 
-  from coprime_exp[OF cpa, of r] coprime_commute 
+  hence cpa: "coprime p a" using coprime by auto
+  from coprime_exp[OF cpa, of r] coprime_commute
   have arp: "coprime (a^r) p" by blast
   from fermat_little[OF arp, simplified ord_divides] o phip
   have "q dvd (p - 1)" by simp
@@ -1099,7 +1099,7 @@
   shows "prime n"
 proof-
   {fix p assume p: "prime p" "p dvd q"
-    from aq[rule_format] p obtain b where 
+    from aq[rule_format] p obtain b where
       b: "[a^((n - 1) div p) = b] (mod n)" "coprime (b - 1) n" by blast
     {assume a0: "a=0"
       from n an have "[0 = 1] (mod n)" unfolding a0 power_0_left by auto
@@ -1108,21 +1108,21 @@
     hence a1: "a \<ge> 1" by arith
     from one_le_power[OF a1] have ath: "1 \<le> a ^ ((n - 1) div p)" .
     {assume b0: "b = 0"
-      from p(2) nqr have "(n - 1) mod p = 0" 
+      from p(2) nqr have "(n - 1) mod p = 0"
 	apply (simp only: dvd_eq_mod_eq_0[symmetric]) by (rule dvd_mult2, simp)
-      with mod_div_equality[of "n - 1" p] 
-      have "(n - 1) div p * p= n - 1" by auto 
+      with mod_div_equality[of "n - 1" p]
+      have "(n - 1) div p * p= n - 1" by auto
       hence eq: "(a^((n - 1) div p))^p = a^(n - 1)"
 	by (simp only: power_mult[symmetric])
       from prime_ge_2[OF p(1)] have pS: "Suc (p - 1) = p" by arith
       from b(1) have d: "n dvd a^((n - 1) div p)" unfolding b0 cong_0_divides .
       from divides_rexp[OF d, of "p - 1"] pS eq cong_divides[OF an] n
       have False by simp}
-    then have b0: "b \<noteq> 0" ..  
-    hence b1: "b \<ge> 1" by arith    
+    then have b0: "b \<noteq> 0" ..
+    hence b1: "b \<ge> 1" by arith
     from cong_coprime[OF cong_sub[OF b(1) cong_refl[of 1] ath b1]] b(2) nqr
     have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute)}
-  hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n " 
+  hence th: "\<forall>p. prime p \<and> p dvd q \<longrightarrow> coprime (a ^ ((n - 1) div p) - 1) n "
     by blast
   from pocklington[OF n nqr sqr an th] show ?thesis .
 qed
@@ -1137,7 +1137,7 @@
 proof(induct n rule: nat_less_induct)
   fix n assume H: "\<forall>m<n. m \<noteq> 0 \<longrightarrow> (\<exists>ps. primefact ps m)" and n: "n\<noteq>0"
   let ?ths = "\<exists>ps. primefact ps n"
-  {assume "n = 1" 
+  {assume "n = 1"
     hence "primefact [] n" by (simp add: primefact_def)
     hence ?ths by blast }
   moreover
@@ -1154,7 +1154,7 @@
   ultimately show ?ths by blast
 qed
 
-lemma primefact_contains: 
+lemma primefact_contains:
   assumes pf: "primefact ps n" and p: "prime p" and pn: "p dvd n"
   shows "p \<in> set ps"
   using pf p pn
@@ -1162,14 +1162,14 @@
   case Nil thus ?case by (auto simp add: primefact_def)
 next
   case (Cons q qs p n)
-  from Cons.prems[unfolded primefact_def] 
+  from Cons.prems[unfolded primefact_def]
   have q: "prime q" "q * foldr op * qs 1 = n" "\<forall>p \<in>set qs. prime p"  and p: "prime p" "p dvd q * foldr op * qs 1" by simp_all
   {assume "p dvd q"
     with p(1) q(1) have "p = q" unfolding prime_def by auto
     hence ?case by simp}
   moreover
   { assume h: "p dvd foldr op * qs 1"
-    from q(3) have pqs: "primefact qs (foldr op * qs 1)" 
+    from q(3) have pqs: "primefact qs (foldr op * qs 1)"
       by (simp add: primefact_def)
     from Cons.hyps[OF pqs p(1) h] have ?case by simp}
   ultimately show ?case using prime_divprod[OF p] by blast
@@ -1180,13 +1180,13 @@
 (* Variant of Lucas theorem.                                                 *)
 
 lemma lucas_primefact:
-  assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)" 
-  and psn: "foldr op * ps 1 = n - 1" 
+  assumes n: "n \<ge> 2" and an: "[a^(n - 1) = 1] (mod n)"
+  and psn: "foldr op * ps 1 = n - 1"
   and psp: "list_all (\<lambda>p. prime p \<and> \<not> [a^((n - 1) div p) = 1] (mod n)) ps"
   shows "prime n"
 proof-
   {fix p assume p: "prime p" "p dvd n - 1" "[a ^ ((n - 1) div p) = 1] (mod n)"
-    from psn psp have psn1: "primefact ps (n - 1)" 
+    from psn psp have psn1: "primefact ps (n - 1)"
       by (auto simp add: list_all_iff primefact_variant)
     from p(3) primefact_contains[OF psn1 p(1,2)] psp
     have False by (induct ps, auto)}
@@ -1198,26 +1198,26 @@
 lemma mod_le: assumes n: "n \<noteq> (0::nat)" shows "m mod n \<le> m"
 proof-
     from mod_div_equality[of m n]
-    have "\<exists>x. x + m mod n = m" by blast 
+    have "\<exists>x. x + m mod n = m" by blast
     then show ?thesis by auto
 qed
-  
+
 
 lemma pocklington_primefact:
   assumes n: "n \<ge> 2" and qrn: "q*r = n - 1" and nq2: "n \<le> q^2"
-  and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q" 
+  and arnb: "(a^r) mod n = b" and psq: "foldr op * ps 1 = q"
   and bqn: "(b^q) mod n = 1"
   and psp: "list_all (\<lambda>p. prime p \<and> coprime ((b^(q div p)) mod n - 1) n) ps"
   shows "prime n"
 proof-
   from bqn psp qrn
   have bqn: "a ^ (n - 1) mod n = 1"
-    and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"  unfolding arnb[symmetric] power_mod 
+    and psp: "list_all (\<lambda>p. prime p \<and> coprime (a^(r *(q div p)) mod n - 1) n) ps"  unfolding arnb[symmetric] power_mod
     by (simp_all add: power_mult[symmetric] algebra_simps)
   from n  have n0: "n > 0" by arith
   from mod_div_equality[of "a^(n - 1)" n]
     mod_less_divisor[OF n0, of "a^(n - 1)"]
-  have an1: "[a ^ (n - 1) = 1] (mod n)" 
+  have an1: "[a ^ (n - 1) = 1] (mod n)"
     unfolding nat_mod bqn
     apply -
     apply (rule exI[where x="0"])
@@ -1226,11 +1226,11 @@
   {fix p assume p: "prime p" "p dvd q"
     from psp psq have pfpsq: "primefact ps q"
       by (auto simp add: primefact_variant list_all_iff)
-    from psp primefact_contains[OF pfpsq p] 
+    from psp primefact_contains[OF pfpsq p]
     have p': "coprime (a ^ (r * (q div p)) mod n - 1) n"
       by (simp add: list_all_iff)
     from prime_ge_2[OF p(1)] have p01: "p \<noteq> 0" "p \<noteq> 1" "p =Suc(p - 1)" by arith+
-    from div_mult1_eq[of r q p] p(2) 
+    from div_mult1_eq[of r q p] p(2)
     have eq1: "r* (q div p) = (n - 1) div p"
       unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute)
     have ath: "\<And>a (b::nat). a <= b \<Longrightarrow> a \<noteq> 0 ==> 1 <= a \<and> 1 <= b" by arith
@@ -1243,22 +1243,22 @@
       hence eq0: "(a^((n - 1) div p))^p = (n*s)^p" by simp
       from qrn[symmetric] have qn1: "q dvd n - 1" unfolding dvd_def by auto
       from dvd_trans[OF p(2) qn1] div_mod_equality'[of "n - 1" p]
-      have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)  
+      have npp: "(n - 1) div p * p = n - 1" by (simp add: dvd_eq_mod_eq_0)
       with eq0 have "a^ (n - 1) = (n*s)^p"
 	by (simp add: power_mult[symmetric])
       hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp
       also have "\<dots> = 0" by (simp add: mult_assoc)
       finally have False by simp }
-      then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto 
-    have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"  
-      unfolding modeq_def by simp 
+      then have th11: "a ^ ((n - 1) div p) mod n \<noteq> 0" by auto
+    have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)"
+      unfolding modeq_def by simp
     from cong_sub[OF th1 cong_refl[of 1]]  ath[OF th10 th11]
     have th: "[a ^ ((n - 1) div p) mod n - 1 = a ^ ((n - 1) div p) - 1] (mod n)"
-      by blast 
-    from cong_coprime[OF th] p'[unfolded eq1] 
+      by blast
+    from cong_coprime[OF th] p'[unfolded eq1]
     have "coprime (a ^ ((n - 1) div p) - 1) n" by (simp add: coprime_commute) }
   with pocklington[OF n qrn[symmetric] nq2 an1]
-  show ?thesis by blast    
+  show ?thesis by blast
 qed
 
 end
--- a/src/HOL/Library/Topology_Euclidean_Space.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Topology_Euclidean_Space.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -15,7 +15,7 @@
 subsection{* General notion of a topology *}
 
 definition "istopology L \<longleftrightarrow> {} \<in> L \<and> (\<forall>S \<in>L. \<forall>T \<in>L. S \<inter> T \<in> L) \<and> (\<forall>K. K \<subseteq>L \<longrightarrow> \<Union> K \<in> L)"
-typedef (open) 'a topology = "{L::('a set) set. istopology L}" 
+typedef (open) 'a topology = "{L::('a set) set. istopology L}"
   morphisms "openin" "topology"
   unfolding istopology_def by blast
 
@@ -72,10 +72,10 @@
   {assume ?lhs then have ?rhs by auto }
   moreover
   {assume H: ?rhs
-    then obtain t where t: "\<forall>x\<in>S. openin U (t x) \<and> x \<in> t x \<and> t x \<subseteq> S" 
+    then obtain t where t: "\<forall>x\<in>S. openin U (t x) \<and> x \<in> t x \<and> t x \<subseteq> S"
       unfolding Ball_def ex_simps(6)[symmetric] choice_iff by blast
     from t have th0: "\<forall>x\<in> t`S. openin U x" by auto
-    have "\<Union> t`S = S" using t by auto 
+    have "\<Union> t`S = S" using t by auto
     with openin_Union[OF th0] have "openin U S" by simp }
   ultimately show ?thesis by blast
 qed
@@ -86,7 +86,7 @@
 
 lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" by (metis closedin_def)
 lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def)
-lemma closedin_topspace[intro,simp]: 
+lemma closedin_topspace[intro,simp]:
   "closedin U (topspace U)" by (simp add: closedin_def)
 lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"
   by (auto simp add: Diff_Un closedin_def)
@@ -114,11 +114,11 @@
   have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S]  oS cT
     by (auto simp add: topspace_def openin_subset)
   then show ?thesis using oS cT by (auto simp add: closedin_def)
-qed  
+qed
 
 lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)"
 proof-
-  have "S - T = S \<inter> (topspace U - T)" using closedin_subset[of U S]  oS cT 
+  have "S - T = S \<inter> (topspace U - T)" using closedin_subset[of U S]  oS cT
     by (auto simp add: topspace_def )
   then show ?thesis using oS cT by (auto simp add: openin_closedin_eq)
 qed
@@ -136,27 +136,27 @@
     then have "A \<inter> B \<in> ?L" by blast}
   moreover
   {fix K assume K: "K \<subseteq> ?L"
-    have th0: "?L = (\<lambda>S. S \<inter> V) ` openin U " 
-      apply (rule set_ext) 
-      apply (simp add: Ball_def image_iff) 
+    have th0: "?L = (\<lambda>S. S \<inter> V) ` openin U "
+      apply (rule set_ext)
+      apply (simp add: Ball_def image_iff)
       by (metis mem_def)
     from K[unfolded th0 subset_image_iff]
     obtain Sk where Sk: "Sk \<subseteq> openin U" "K = (\<lambda>S. S \<inter> V) ` Sk" by blast
     have "\<Union>K = (\<Union>Sk) \<inter> V" using Sk by auto
     moreover have "openin U (\<Union> Sk)" using Sk by (auto simp add: subset_eq mem_def)
     ultimately have "\<Union>K \<in> ?L" by blast}
-  ultimately show ?thesis unfolding istopology_def by blast 
-qed
-
-lemma openin_subtopology: 
+  ultimately show ?thesis unfolding istopology_def by blast
+qed
+
+lemma openin_subtopology:
   "openin (subtopology U V) S \<longleftrightarrow> (\<exists> T. (openin U T) \<and> (S = T \<inter> V))"
-  unfolding subtopology_def topology_inverse'[OF istopology_subtopology] 
-  by (auto simp add: Collect_def) 
+  unfolding subtopology_def topology_inverse'[OF istopology_subtopology]
+  by (auto simp add: Collect_def)
 
 lemma topspace_subtopology: "topspace(subtopology U V) = topspace U \<inter> V"
   by (auto simp add: topspace_def openin_subtopology)
 
-lemma closedin_subtopology: 
+lemma closedin_subtopology:
   "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"
   unfolding closedin_def topspace_subtopology
   apply (simp add: openin_subtopology)
@@ -172,7 +172,7 @@
   apply (rule exI[where x="topspace U"])
   by auto
 
-lemma subtopology_superset: assumes UV: "topspace U \<subseteq> V" 
+lemma subtopology_superset: assumes UV: "topspace U \<subseteq> V"
   shows "subtopology U V = U"
 proof-
   {fix S
@@ -289,8 +289,8 @@
 definition "ball x e = {y. dist x y < e}"
 definition "cball x e = {y. dist x y \<le> e}"
 
-lemma mem_ball[simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e" by (simp add: ball_def) 
-lemma mem_cball[simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e" by (simp add: cball_def) 
+lemma mem_ball[simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e" by (simp add: ball_def)
+lemma mem_cball[simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e" by (simp add: cball_def)
 lemma mem_ball_0[simp]: "x \<in> ball 0 e \<longleftrightarrow> norm x < e" by (simp add: dist_def)
 lemma mem_cball_0[simp]: "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e" by (simp add: dist_def)
 lemma centre_in_cball[simp]: "x \<in> cball x e \<longleftrightarrow> 0\<le> e"  by simp
@@ -301,14 +301,14 @@
   by (simp add: expand_set_eq) arith
 
 lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s"
-  by (simp add: expand_set_eq) 
+  by (simp add: expand_set_eq)
 
 subsection{* Topological properties of open balls *}
 
-lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b" 
-  "(a::real) - b < 0 \<longleftrightarrow> a < b" 
+lemma diff_less_iff: "(a::real) - b > 0 \<longleftrightarrow> a > b"
+  "(a::real) - b < 0 \<longleftrightarrow> a < b"
   "a - b < c \<longleftrightarrow> a < c +b" "a - b > c \<longleftrightarrow> a > c +b" by arith+
-lemma diff_le_iff: "(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" "(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b" 
+lemma diff_le_iff: "(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" "(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b"
   "a - b \<le> c \<longleftrightarrow> a \<le> c +b" "a - b \<ge> c \<longleftrightarrow> a \<ge> c +b"  by arith+
 
 lemma open_ball[intro, simp]: "open (ball x e)"
@@ -343,9 +343,9 @@
   by (auto simp add: openin_subtopology open_openin[symmetric])
 
 lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)"
-  by (auto simp add: openin_open) 
-
-lemma open_openin_trans[trans]: 
+  by (auto simp add: openin_open)
+
+lemma open_openin_trans[trans]:
  "open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T"
   by (metis Int_absorb1  openin_open_Int)
 
@@ -367,7 +367,7 @@
 lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
   by (auto simp add: closedin_closed)
 
-lemma openin_euclidean_subtopology_iff: "openin (subtopology euclidean U) S 
+lemma openin_euclidean_subtopology_iff: "openin (subtopology euclidean U) S
   \<longleftrightarrow> S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
   {assume ?lhs hence ?rhs unfolding openin_subtopology open_openin[symmetric]
@@ -389,22 +389,22 @@
       then obtain x where "x\<in>S" and x:"y \<in> ball x (d x)" by auto
       assume "y\<in>U"
       hence "y\<in>S" using d[OF `x\<in>S`] and x by(auto simp add: dist_sym) }
-    ultimately have "S = ?T \<inter> U" by blast 
+    ultimately have "S = ?T \<inter> U" by blast
     with oT have ?lhs unfolding openin_subtopology open_openin[symmetric] by blast}
   ultimately show ?thesis by blast
 qed
 
 text{* These "transitivity" results are handy too. *}
 
-lemma openin_trans[trans]: "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T 
+lemma openin_trans[trans]: "openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T
   \<Longrightarrow> openin (subtopology euclidean U) S"
   unfolding open_openin openin_open by blast
 
 lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S"
   by (auto simp add: openin_open intro: openin_trans)
 
-lemma closedin_trans[trans]: 
- "closedin (subtopology euclidean T) S \<Longrightarrow> 
+lemma closedin_trans[trans]:
+ "closedin (subtopology euclidean T) S \<Longrightarrow>
            closedin (subtopology euclidean U) T
            ==> closedin (subtopology euclidean U) S"
   by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc)
@@ -415,10 +415,10 @@
 subsection{* Connectedness *}
 
 definition "connected S \<longleftrightarrow>
-  ~(\<exists>e1 e2. open e1 \<and> open e2 \<and> S \<subseteq> (e1 \<union> e2) \<and> (e1 \<inter> e2 \<inter> S = {}) 
+  ~(\<exists>e1 e2. open e1 \<and> open e2 \<and> S \<subseteq> (e1 \<union> e2) \<and> (e1 \<inter> e2 \<inter> S = {})
   \<and> ~(e1 \<inter> S = {}) \<and> ~(e2 \<inter> S = {}))"
 
-lemma connected_local: 
+lemma connected_local:
  "connected S \<longleftrightarrow> ~(\<exists>e1 e2.
                  openin (subtopology euclidean S) e1 \<and>
                  openin (subtopology euclidean S) e2 \<and>
@@ -430,7 +430,7 @@
 
 lemma exists_diff: "(\<exists>S. P(UNIV - S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
-  
+
   {assume "?lhs" hence ?rhs by blast }
   moreover
   {fix S assume H: "P S"
@@ -443,10 +443,10 @@
         (\<forall>T. openin (subtopology euclidean S) T \<and>
             closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
 proof-
-  have " \<not> connected S \<longleftrightarrow> (\<exists>e1 e2. open e1 \<and> open (UNIV - e2) \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})" 
-    unfolding connected_def openin_open closedin_closed 
+  have " \<not> connected S \<longleftrightarrow> (\<exists>e1 e2. open e1 \<and> open (UNIV - e2) \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
+    unfolding connected_def openin_open closedin_closed
     apply (subst exists_diff) by blast
-  hence th0: "connected S \<longleftrightarrow> \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})" 
+  hence th0: "connected S \<longleftrightarrow> \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (UNIV - e2) \<and> e1 \<inter> (UNIV - e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (UNIV - e2) \<inter> S \<noteq> {})"
     (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") apply (simp add: closed_def) by metis
 
   have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
@@ -465,7 +465,7 @@
 
 subsection{* Hausdorff and other separation properties *}
 
-lemma hausdorff: 
+lemma hausdorff:
   assumes xy: "x \<noteq> y"
   shows "\<exists>U V. open U \<and> open V \<and> x\<in> U \<and> y \<in> V \<and> (U \<inter> V = {})" (is "\<exists>U V. ?P U V")
 proof-
@@ -479,7 +479,7 @@
 qed
 
 lemma separation_t2: "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {})"
-  using hausdorff[of x y] by blast 
+  using hausdorff[of x y] by blast
 
 lemma separation_t1: "x \<noteq> y \<longleftrightarrow> (\<exists>U V. open U \<and> open V \<and> x \<in>U \<and> y\<notin> U \<and> x\<notin>V \<and> y\<in>V)"
   using separation_t2[of x y] by blast
@@ -515,7 +515,7 @@
   {
     fix e::real assume ep: "e>0"
     from vector_choose_size[of "e/2"] ep have "\<exists>(c:: real ^'n). norm c = e/2" by auto
-    then obtain c ::"real^'n" where c: "norm c = e/2" by blast 
+    then obtain c ::"real^'n" where c: "norm c = e/2" by blast
     let ?x = "x + c"
     have "?x \<noteq> x" using c ep by (auto simp add: norm_eq_0_imp)
     moreover have "dist ?x x < e" using c ep apply simp by norm
@@ -536,7 +536,7 @@
 proof-
   let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
   let ?O = "{x::real^'n. \<forall>i\<in>?U. x$i\<ge>0}"
-  {fix x:: "real^'n" and i::nat assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e" and i: "i \<in> ?U" 
+  {fix x:: "real^'n" and i::nat assume H: "\<forall>e>0. \<exists>x'\<in>?O. x' \<noteq> x \<and> dist x' x < e" and i: "i \<in> ?U"
     and xi: "x$i < 0"
     from xi have th0: "-x$i > 0" by arith
     from H[rule_format, OF th0] obtain x' where x': "x' \<in>?O" "x' \<noteq> x" "dist x' x < -x $ i" by blast
@@ -548,7 +548,7 @@
       have th2: "\<bar>dist x x'\<bar> \<ge> \<bar>(x' - x)$i\<bar>" using  component_le_norm[OF i, of "x'-x"]
 	apply (simp add: dist_def) by norm
       from th[OF th1 th2] x'(3) have False by (simp add: dist_sym dist_pos_le) }
-  then show ?thesis unfolding closed_limpt islimpt_approachable 
+  then show ?thesis unfolding closed_limpt islimpt_approachable
     unfolding not_le[symmetric] by blast
 qed
 
@@ -556,7 +556,7 @@
 proof(induct rule: finite_induct[OF fS])
   case 1 thus ?case apply auto by ferrack
 next
-  case (2 x F) 
+  case (2 x F)
   from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x" by blast
   {assume "x = a" hence ?case using d by auto  }
   moreover
@@ -569,7 +569,7 @@
 qed
 
 lemma islimpt_finite: assumes fS: "finite S" shows "\<not> a islimpt S"
-  unfolding islimpt_approachable 
+  unfolding islimpt_approachable
   using finite_set_avoid[OF fS, of a] by (metis dist_sym  not_le)
 
 lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T"
@@ -582,10 +582,10 @@
   apply auto
   done
 
-lemma discrete_imp_closed: 
+lemma discrete_imp_closed:
   assumes e: "0 < e" and d: "\<forall>x \<in> S. \<forall>y \<in> S. norm(y - x) < e \<longrightarrow> y = x"
   shows "closed S"
-proof-  
+proof-
   {fix x assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"
     from e have e2: "e/2 > 0" by arith
     from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y\<noteq>x" "dist y x < e/2" by blast
@@ -593,7 +593,7 @@
     from e2 y(2) have mp: "?m > 0" by (simp add: dist_nz[THEN sym])
     from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z\<noteq>x" "dist z x < ?m" by blast
     have th: "norm (z - y) < e" using z y by norm
-    from d[rule_format, OF y(1) z(1) th] y z 
+    from d[rule_format, OF y(1) z(1) th] y z
     have False by (auto simp add: dist_sym)}
   then show ?thesis by (metis islimpt_approachable closed_limpt)
 qed
@@ -614,7 +614,7 @@
   apply (subst open_subopen) by blast
 
 lemma interior_interior[simp]: "interior(interior S) = interior S" by (metis interior_eq open_interior)
-lemma interior_subset: "interior S \<subseteq> S" by (auto simp add: interior_def) 
+lemma interior_subset: "interior S \<subseteq> S" by (auto simp add: interior_def)
 lemma subset_interior: "S \<subseteq> T ==> (interior S) \<subseteq> (interior T)" by (auto simp add: interior_def)
 lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T ==> T \<subseteq> (interior S)" by (auto simp add: interior_def)
 lemma interior_unique: "T \<subseteq> S \<Longrightarrow> open T  \<Longrightarrow> (\<forall>T'. T' \<subseteq> S \<and> open T' \<longrightarrow> T' \<subseteq> T) \<Longrightarrow> interior S = T"
@@ -638,16 +638,16 @@
   {fix d::real assume d: "d>0"
     let ?m = "min d e / 2"
     have mde2: "?m \<ge> 0" using e(1) d(1) by arith
-    from vector_choose_dist[OF mde2, of x] 
+    from vector_choose_dist[OF mde2, of x]
     obtain y where y: "dist x y = ?m" by blast
     have th: "dist x y < e" "dist x y < d" unfolding y using e(1) d(1) by arith+
-    have "\<exists>x'\<in>S. x'\<noteq> x \<and> dist x' x < d" 
+    have "\<exists>x'\<in>S. x'\<noteq> x \<and> dist x' x < d"
       apply (rule bexI[where x=y])
       using e th y by (auto simp add: dist_sym)}
   then show ?thesis unfolding islimpt_approachable by blast
 qed
 
-lemma interior_closed_Un_empty_interior: 
+lemma interior_closed_Un_empty_interior:
   assumes cS: "closed S" and iT: "interior T = {}"
   shows "interior(S \<union> T) = interior S"
 proof-
@@ -655,17 +655,17 @@
     by (rule subset_interior, blast)
   moreover
   {fix x e assume e: "e > 0" "\<forall>x' \<in> ball x e. x'\<in>(S\<union>T)"
-    {fix y assume y: "y \<in> ball x e" 
+    {fix y assume y: "y \<in> ball x e"
       {fix d::real assume d: "d > 0"
 	let ?k = "min d (e - dist x y)"
 	have kp: "?k > 0" using d e(1) y[unfolded mem_ball] by norm
-	have "?k/2 \<ge> 0" using kp by simp 
+	have "?k/2 \<ge> 0" using kp by simp
 	then obtain w where w: "dist y w = ?k/ 2" by (metis vector_choose_dist)
-	from iT[unfolded expand_set_eq mem_interior] 
+	from iT[unfolded expand_set_eq mem_interior]
 	have "\<not> ball w (?k/4) \<subseteq> T" using kp by (auto simp add: Arith_Tools.less_divide_eq_number_of1)
 	then obtain z where z: "dist w z < ?k/4" "z \<notin> T" by (auto simp add: subset_eq)
 	have "z \<notin> T \<and> z\<noteq> y \<and> dist z y < d \<and> dist x z < e" using z apply simp
-	  using w e(1) d apply (auto simp only: dist_sym)  
+	  using w e(1) d apply (auto simp only: dist_sym)
 	  apply (auto simp add: min_def cong del: if_weak_cong)
 	  apply (cases "d \<le> e - dist x y", auto simp add: ring_simps cong del: if_weak_cong)
 	  apply norm
@@ -681,7 +681,7 @@
       then have "y\<in>S" by (metis islimpt_approachable cS closed_limpt) }
     then have "x \<in> interior S" unfolding mem_interior using e(1) by blast}
   hence "interior (S\<union>T) \<subseteq> interior S" unfolding mem_interior Ball_def subset_eq by blast
-  ultimately show ?thesis by blast 
+  ultimately show ?thesis by blast
 qed
 
 
@@ -891,7 +891,7 @@
 	using open_ball[of a e] dist_refl[of a] `e > 0`
 	by (auto, erule_tac x = "ball a e" in allE, auto)
     }
-    ultimately have ?rhse by auto 
+    ultimately have ?rhse by auto
   }
   thus ?rhs by auto
 next
@@ -916,7 +916,7 @@
   ultimately show ?lhs unfolding frontier_closures using closure_def[of "UNIV - S"] by auto
 qed
 
-lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S" 
+lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"
   by (metis frontier_def closure_closed Diff_subset)
 
 lemma frontier_empty: "frontier {} = {}"
@@ -931,26 +931,26 @@
   thus ?thesis using frontier_subset_closed[of S] by auto
 qed
 
-lemma frontier_complement: "frontier(UNIV - S) = frontier S" 
+lemma frontier_complement: "frontier(UNIV - S) = frontier S"
   by (auto simp add: frontier_def closure_complement interior_complement)
 
 lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
-  using frontier_complement frontier_subset_eq[of "UNIV - S"] 
+  using frontier_complement frontier_subset_eq[of "UNIV - S"]
   unfolding open_closed by auto
 
 subsection{* A variant of nets (Slightly non-standard but good for our purposes). *}
 
-typedef (open) 'a net = 
-  "{g :: 'a \<Rightarrow> 'a \<Rightarrow> bool. \<forall>x y. (\<forall>z. g z x \<longrightarrow> g z y) \<or> (\<forall>z. g z y \<longrightarrow> g z x)}" 
+typedef (open) 'a net =
+  "{g :: 'a \<Rightarrow> 'a \<Rightarrow> bool. \<forall>x y. (\<forall>z. g z x \<longrightarrow> g z y) \<or> (\<forall>z. g z y \<longrightarrow> g z x)}"
   morphisms "netord" "mknet" by blast
 lemma net: "(\<forall>z. netord n z x \<longrightarrow> netord n z y) \<or> (\<forall>z. netord n z y \<longrightarrow> netord n z x)"
   using netord[of n] by auto
 
 lemma oldnet: "netord n x x \<Longrightarrow> netord n y y \<Longrightarrow>
-  \<exists>z. netord n z z \<and> (\<forall>w. netord n w z \<longrightarrow> netord n w x \<and> netord n w y)" 
+  \<exists>z. netord n z z \<and> (\<forall>w. netord n w z \<longrightarrow> netord n w x \<and> netord n w y)"
   by (metis net)
 
-lemma net_dilemma: 
+lemma net_dilemma:
  "\<exists>a. (\<exists>x. netord net x a) \<and> (\<forall>x. netord net x a \<longrightarrow> P x) \<Longrightarrow>
          \<exists>b. (\<exists>x. netord net x b) \<and> (\<forall>x. netord net x b \<longrightarrow> Q x)
          \<Longrightarrow> \<exists>c. (\<exists>x. netord net x c) \<and> (\<forall>x. netord net x c \<longrightarrow> P x \<and> Q x)"
@@ -962,10 +962,10 @@
 definition "at_infinity = mknet(\<lambda>x y. norm x \<ge> norm y)"
 definition "sequentially = mknet(\<lambda>(m::nat) n. m >= n)"
 
-definition within :: "'a net \<Rightarrow> 'a set \<Rightarrow> 'a net" (infixr "within" 70) where 
+definition within :: "'a net \<Rightarrow> 'a set \<Rightarrow> 'a net" (infixr "within" 70) where
   within_def: "net within S = mknet (\<lambda>x y. netord net x y \<and> x \<in> S)"
 
-definition indirection :: "real ^'n \<Rightarrow> real ^'n \<Rightarrow> (real ^'n) net" (infixr "indirection" 70) where 
+definition indirection :: "real ^'n \<Rightarrow> real ^'n \<Rightarrow> (real ^'n) net" (infixr "indirection" 70) where
   indirection_def: "a indirection v = (at a) within {b. \<exists>c\<ge>0. b - a = c*s v}"
 
 text{* Prove That They are all nets. *}
@@ -973,8 +973,8 @@
 lemma mknet_inverse': "netord (mknet r) = r \<longleftrightarrow> (\<forall>x y. (\<forall>z. r z x \<longrightarrow> r z y) \<or> (\<forall>z. r z y \<longrightarrow> r z x))"
   using mknet_inverse[of r] apply (auto simp add: netord_inverse) by (metis net)
 
-method_setup net = {* 
- let 
+method_setup net = {*
+ let
   val ss1 = HOL_basic_ss addsimps [@{thm expand_fun_eq} RS sym]
   val ss2 = HOL_basic_ss addsimps [@{thm mknet_inverse'}]
   fun tac ths = ObjectLogic.full_atomize_tac THEN' Simplifier.simp_tac (ss1 addsimps ths) THEN' Simplifier.asm_full_simp_tac ss2
@@ -986,7 +986,7 @@
   apply (net at_def)
   by (metis dist_sym real_le_linear real_le_trans)
 
-lemma at_infinity: 
+lemma at_infinity:
  "\<And>x y. netord at_infinity x y \<longleftrightarrow> norm x >= norm y"
   apply (net at_infinity_def)
   apply (metis real_le_linear real_le_trans)
@@ -1016,24 +1016,24 @@
 subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *}
 
 
-definition "trivial_limit (net:: 'a net) \<longleftrightarrow> 
+definition "trivial_limit (net:: 'a net) \<longleftrightarrow>
   (\<forall>(a::'a) b. a = b) \<or> (\<exists>(a::'a) b. a \<noteq> b \<and> (\<forall>x. ~(netord (net) x a) \<and> ~(netord(net) x b)))"
 
 
 lemma trivial_limit_within: "trivial_limit (at (a::real^'n) within S) \<longleftrightarrow> ~(a islimpt S)"
 proof-
-  {assume "\<forall>(a::real^'n) b. a = b" hence "\<not> a islimpt S" 
+  {assume "\<forall>(a::real^'n) b. a = b" hence "\<not> a islimpt S"
       apply (simp add: islimpt_approachable_le)
       by (rule exI[where x=1], auto)}
   moreover
   {fix b c assume bc: "b \<noteq> c" "\<forall>x. \<not> netord (at a within S) x b \<and> \<not> netord (at a within S) x c"
     have "dist a b > 0 \<or> dist a c > 0" using bc by (auto simp add: within at dist_nz[THEN sym])
-    then have "\<not> a islimpt S" 
+    then have "\<not> a islimpt S"
       using bc
       unfolding within at dist_nz islimpt_approachable_le
       by(auto simp add: dist_triangle dist_sym dist_eq_0[THEN sym]) }
   moreover
-  {assume "\<not> a islimpt S" 
+  {assume "\<not> a islimpt S"
     then obtain e where e: "e > 0" "\<forall>x' \<in> S. x' \<noteq> a \<longrightarrow> dist x' a > e"
       unfolding islimpt_approachable_le by (auto simp add: not_le)
     from e vector_choose_dist[of e a] obtain b where b: "dist a b = e" by auto
@@ -1041,13 +1041,13 @@
     moreover have "\<forall>x. \<not> ((0 < dist x a \<and> dist x a \<le> dist a a) \<and> x \<in> S) \<and>
                  \<not> ((0 < dist x a \<and> dist x a \<le> dist b a) \<and> x \<in> S)"
       using e(2) b by (auto simp add: dist_refl dist_sym)
-    ultimately have "trivial_limit (at a within S)"  unfolding trivial_limit_def within at  
+    ultimately have "trivial_limit (at a within S)"  unfolding trivial_limit_def within at
       by blast}
-  ultimately show ?thesis unfolding trivial_limit_def by blast    
+  ultimately show ?thesis unfolding trivial_limit_def by blast
 qed
 
 lemma trivial_limit_at: "~(trivial_limit (at a))"
-  apply (subst within_UNIV[symmetric]) 
+  apply (subst within_UNIV[symmetric])
   by (simp add: trivial_limit_within islimpt_UNIV)
 
 lemma trivial_limit_at_infinity: "~(trivial_limit (at_infinity :: ('a::{norm,zero_neq_one}) net))"
@@ -1055,7 +1055,7 @@
   by (metis order_refl zero_neq_one)
 
 lemma trivial_limit_sequentially:  "~(trivial_limit sequentially)"
-  by (auto simp add: trivial_limit_def sequentially) 
+  by (auto simp add: trivial_limit_def sequentially)
 
 subsection{* Some property holds "sufficiently close" to the limit point. *}
 
@@ -1128,7 +1128,7 @@
 lemma eventually_and: " eventually (\<lambda>x. P x \<and> Q x) net \<longleftrightarrow> eventually P net \<and> eventually Q net"
   apply (simp add: eventually_def)
   apply (cases "trivial_limit net")
-  using net_dilemma[of net P Q] by auto 
+  using net_dilemma[of net P Q] by auto
 
 lemma eventually_mono: "(\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually P net  \<Longrightarrow> eventually Q net"
   by (metis eventually_def)
@@ -1155,7 +1155,7 @@
   text{* Notation Lim to avoid collition with lim defined in analysis *}
 definition "Lim net f = (THE l. (f ---> l) net)"
 
-lemma Lim: 
+lemma Lim:
  "(f ---> l) net \<longleftrightarrow>
         trivial_limit net \<or>
         (\<forall>e>0. \<exists>y. (\<exists>x. netord net x y) \<and>
@@ -1181,7 +1181,7 @@
   "(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x::real^'n. norm x >= b \<longrightarrow> dist (f x) l < e)"
   by (auto simp add: tendsto_def eventually_at_infinity)
 
-lemma Lim_sequentially: 
+lemma Lim_sequentially:
  "(S ---> l) sequentially \<longleftrightarrow>
           (\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (S n) l < e)"
   by (auto simp add: tendsto_def eventually_sequentially)
@@ -1210,7 +1210,7 @@
   thus ?thesis unfolding Lim_within by auto
 qed
 
-lemma Lim_Un_univ: 
+lemma Lim_Un_univ:
  "(f ---> l) (at x within S) \<Longrightarrow> (f ---> l) (at x within T) \<Longrightarrow>  S \<union> T = (UNIV::(real^'n) set)
         ==> (f ---> l) (at x)"
   by (metis Lim_Un within_UNIV)
@@ -1236,18 +1236,18 @@
   assume ?rhs
   { fix e::real assume "e>0"
     then obtain d where "d>0" and d:"\<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e" using `?rhs` unfolding Lim_at by auto
-    hence "\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e" using `d>0` by auto 
+    hence "\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e" using `d>0` by auto
   }
   thus ?lhs using Lim_at_within[of f l a S] by (auto simp add: Lim_at)
 qed
 
 text{* Another limit point characterization. *}
 
-lemma islimpt_sequential: 
+lemma islimpt_sequential:
  "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S -{x}) \<and> (f ---> x) sequentially)" (is "?lhs = ?rhs")
 proof
   assume ?lhs
-  then obtain f where f:"\<forall>y. y>0 \<longrightarrow> f y \<in> S \<and> f y \<noteq> x \<and> dist (f y) x < y" 
+  then obtain f where f:"\<forall>y. y>0 \<longrightarrow> f y \<in> S \<and> f y \<noteq> x \<and> dist (f y) x < y"
     unfolding islimpt_approachable using choice[of "\<lambda>e y. e>0 \<longrightarrow> y\<in>S \<and> y\<noteq>x \<and> dist y x < e"] by auto
   { fix n::nat
     have "f (inverse (real n + 1)) \<in> S - {x}" using f by auto
@@ -1262,10 +1262,10 @@
   }
   hence " ((\<lambda>n. f (inverse (real n + 1))) ---> x) sequentially"
     unfolding Lim_sequentially using f by auto
-  ultimately show ?rhs apply (rule_tac x="(\<lambda>n::nat. f (inverse (real n + 1)))" in exI) by auto 
+  ultimately show ?rhs apply (rule_tac x="(\<lambda>n::nat. f (inverse (real n + 1)))" in exI) by auto
 next
   assume ?rhs
-  then obtain f::"nat\<Rightarrow>real^'a"  where f:"(\<forall>n. f n \<in> S - {x})" "(\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f n) x < e)" unfolding Lim_sequentially by auto 
+  then obtain f::"nat\<Rightarrow>real^'a"  where f:"(\<forall>n. f n \<in> S - {x})" "(\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f n) x < e)" unfolding Lim_sequentially by auto
   { fix e::real assume "e>0"
     then obtain N where "dist (f N) x < e" using f(2) by auto
     moreover have "f N\<in>S" "f N \<noteq> x" using f(1) by auto
@@ -1277,7 +1277,7 @@
 text{* Basic arithmetical combining theorems for limits. *}
 
 lemma Lim_linear: fixes f :: "('a \<Rightarrow> real^'n)" and h :: "(real^'n \<Rightarrow> real^'m)"
-  assumes "(f ---> l) net" "linear h" 
+  assumes "(f ---> l) net" "linear h"
   shows "((\<lambda>x. h (f x)) ---> h l) net"
 proof (cases "trivial_limit net")
   case True
@@ -1320,7 +1320,7 @@
  "(f ---> l) net \<Longrightarrow> (g ---> m) net \<Longrightarrow> ((\<lambda>x. f(x) + g(x)) ---> l + m) net"
 proof-
   assume as:"(f ---> l) net" "(g ---> m) net"
-  { fix e::real 
+  { fix e::real
     assume "e>0"
     hence *:"eventually (\<lambda>x. dist (f x) l < e/2) net"
             "eventually (\<lambda>x. dist (g x) m < e/2) net" using as
@@ -1353,9 +1353,9 @@
 lemma Lim_null_norm: "(f ---> 0) net \<longleftrightarrow> ((\<lambda>x. vec1(norm(f x))) ---> 0) net"
   by (simp add: Lim dist_def norm_vec1)
 
-lemma Lim_null_comparison: 
+lemma Lim_null_comparison:
   assumes "eventually (\<lambda>x. norm(f x) <= g x) net" "((\<lambda>x. vec1(g x)) ---> 0) net"
-  shows "(f ---> 0) net" 
+  shows "(f ---> 0) net"
 proof(simp add: tendsto_def, rule+)
   fix e::real assume "0<e"
   { fix x
@@ -1363,7 +1363,7 @@
     hence "dist (f x) 0 < e"  unfolding vec_def using dist_vec1[of "g x" "0"]
       by (vector dist_def norm_vec1 dist_refl real_vector_norm_def dot_def vec1_def)
   }
-  thus "eventually (\<lambda>x. dist (f x) 0 < e) net" 
+  thus "eventually (\<lambda>x. dist (f x) 0 < e) net"
     using eventually_and[of "\<lambda>x. norm(f x) <= g x" "\<lambda>x. dist (vec1 (g x)) 0 < e" net]
     using eventually_mono[of "(\<lambda>x. norm (f x) \<le> g x \<and> dist (vec1 (g x)) 0 < e)" "(\<lambda>x. dist (f x) 0 < e)" net]
     using assms `e>0` unfolding tendsto_def by auto
@@ -1381,7 +1381,7 @@
   apply (rule component_le_norm)
   by auto
 
-lemma Lim_transform_bound: 
+lemma Lim_transform_bound:
   assumes "eventually (\<lambda>n. norm(f n) <= norm(g n)) net"  "(g ---> 0) net"
   shows "(f ---> 0) net"
 proof(simp add: tendsto_def, rule+)
@@ -1416,7 +1416,7 @@
 
 text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
 
-lemma Lim_norm_ubound: 
+lemma Lim_norm_ubound:
   assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) <= e) net"
   shows "norm(l) <= e"
 proof-
@@ -1466,24 +1466,24 @@
   thus ?thesis using assms using Lim_sub[of f l net f l'] by simp
 qed
 
-lemma tendsto_Lim: 
+lemma tendsto_Lim:
  "~(trivial_limit (net::('b::zero_neq_one net))) \<Longrightarrow> (f ---> l) net ==> Lim net f = l"
   unfolding Lim_def using Lim_unique[of net f] by auto
 
 text{* Limit under bilinear function (surprisingly tedious, but important) *}
 
 lemma norm_bound_lemma:
-  "0 < e \<Longrightarrow> \<exists>d>0. \<forall>(x'::real^'b) y'::real^'a. norm(x' - (x::real^'b)) < d \<and> norm(y' - y) < d \<longrightarrow> norm(x') * norm(y' - y) + norm(x' - x) * norm(y) < e" 
-proof- 
+  "0 < e \<Longrightarrow> \<exists>d>0. \<forall>(x'::real^'b) y'::real^'a. norm(x' - (x::real^'b)) < d \<and> norm(y' - y) < d \<longrightarrow> norm(x') * norm(y' - y) + norm(x' - x) * norm(y) < e"
+proof-
   assume e: "0 < e"
-  have th1: "(2 * norm x + 2 * norm y + 2) > 0" using norm_ge_zero[of x] norm_ge_zero[of y] by norm 
+  have th1: "(2 * norm x + 2 * norm y + 2) > 0" using norm_ge_zero[of x] norm_ge_zero[of y] by norm
   hence th0: "0 < e / (2 * norm x + 2 * norm y + 2)"  using `e>0` using divide_pos_pos by auto
-  moreover 
+  moreover
   { fix x' y'
     assume h: "norm (x' - x) < 1" "norm (x' - x) < e / (2 * norm x + 2 * norm y + 2)"
       "norm (y' - y) < 1" "norm (y' - y) < e / (2 * norm x + 2 * norm y + 2)"
     have th: "\<And>a b (c::real). a \<ge> 0 \<Longrightarrow> c \<ge> 0 \<Longrightarrow> a + (b + c) < e ==> b < e " by arith
-    from h have thx: "norm (x' - x) * norm y < e / 2" 
+    from h have thx: "norm (x' - x) * norm y < e / 2"
       using th0 th1 apply (simp add: field_simps)
       apply (rule th) defer defer apply assumption
       by (simp_all add: norm_ge_zero zero_le_mult_iff)
@@ -1492,7 +1492,7 @@
       using h(1) using norm_triangle_ineq2[of x' x] by auto
     hence *:"norm x' < 1 + norm x"  by auto
 
-    have thy: "norm (y' - y) * norm x' < e / (2 * norm x + 2 * norm y + 2) * (1 + norm x)" 
+    have thy: "norm (y' - y) * norm x' < e / (2 * norm x + 2 * norm y + 2) * (1 + norm x)"
       using mult_strict_mono'[OF h(4) * norm_ge_zero norm_ge_zero] by auto
     also have "\<dots> \<le> e/2" apply simp unfolding divide_le_eq
       using th1 th0 `e>0` apply auto
@@ -1503,7 +1503,7 @@
   ultimately show ?thesis apply(rule_tac x="min 1 (e / 2 / (norm x + norm y + 1))" in exI) by auto
 qed
 
-lemma Lim_bilinear: 
+lemma Lim_bilinear:
   fixes net :: "'a net" and h:: "real ^'m \<Rightarrow> real ^'n \<Rightarrow> real ^'p"
   assumes "(f ---> l) net" and "(g ---> m) net" and "bilinear h"
   shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"
@@ -1525,7 +1525,7 @@
 	using d[THEN spec[where x="f x"], THEN spec[where x="g x"]] unfolding dist_def  by auto
       have "norm (h (f x) (g x - m)) + norm (h (f x - l) m) \<le> B * norm (f x) * norm (g x - m) + B * norm (f x - l) * norm m"
 	using B[THEN spec[where x="f x"], THEN spec[where x="g x - m"]]
-	using B[THEN spec[where x="f x - l"], THEN spec[where x="m"]] by auto 
+	using B[THEN spec[where x="f x - l"], THEN spec[where x="m"]] by auto
       also have "\<dots> < e" using ** and `B>0` by(auto simp add: field_simps)
       finally have "dist (h (f x) (g x)) (h l m) < e" unfolding dist_def and * using norm_triangle_lt by auto
     }
@@ -1548,10 +1548,10 @@
   { fix e::real assume "e>0"
     with `?lhs` obtain d where d:"d>0" "\<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e" unfolding Lim_at by auto
     { fix x::"real^'a" assume "0 < dist x 0 \<and> dist x 0 < d"
-      hence "dist (f (a + x)) l < e" using d 
+      hence "dist (f (a + x)) l < e" using d
       apply(erule_tac x="x+a" in allE) by(auto simp add: comm_monoid_add.mult_commute dist_def dist_sym)
     }
-    hence "\<exists>d>0. \<forall>x. 0 < dist x 0 \<and> dist x 0 < d \<longrightarrow> dist (f (a + x)) l < e" using d(1) by auto 
+    hence "\<exists>d>0. \<forall>x. 0 < dist x 0 \<and> dist x 0 < d \<longrightarrow> dist (f (a + x)) l < e" using d(1) by auto
   }
   thus "?rhs" unfolding Lim_at by auto
 next
@@ -1563,7 +1563,7 @@
       hence "dist (f x) l < e" using d apply(erule_tac x="x-a" in allE)
 	by(auto simp add: comm_monoid_add.mult_commute dist_def dist_sym)
     }
-    hence "\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e" using d(1) by auto 
+    hence "\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e" using d(1) by auto
   }
   thus "?lhs" unfolding Lim_at by auto
 qed
@@ -1578,7 +1578,7 @@
   { fix x assume "x \<noteq> a"
     then obtain y where y:"dist y a \<le> dist a a \<and> 0 < dist y a \<and> y \<in> S \<or> dist y a \<le> dist x a \<and> 0 < dist y a \<and> y \<in> S" using assms unfolding trivial_limit_def within at by blast
     assume "\<forall>y. \<not> netord (at a within S) y x"
-    hence "x = a" using y unfolding within at by (auto simp add: dist_refl dist_nz) 
+    hence "x = a" using y unfolding within at by (auto simp add: dist_refl dist_nz)
   }
   moreover
   have "\<forall>y. \<not> netord (at a within S) y a"  using assms unfolding trivial_limit_def within at by (auto simp add: dist_refl)
@@ -1602,7 +1602,7 @@
 lemma Lim_transform_eventually:  "eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net ==> (g ---> l) net"
   using Lim_eventually[of "\<lambda>x. f x - g x" 0 net] Lim_transform[of f g net l] by auto
 
-lemma Lim_transform_within: 
+lemma Lim_transform_within:
   assumes "0 < d" "(\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x')"
           "(f ---> l) (at x within S)"
   shows   "(g ---> l) (at x within S)"
@@ -1611,7 +1611,7 @@
   thus ?thesis using Lim_transform[of f g "at x within S" l] using assms(3) by auto
 qed
 
-lemma Lim_transform_at: "0 < d \<Longrightarrow> (\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x') \<Longrightarrow> 
+lemma Lim_transform_at: "0 < d \<Longrightarrow> (\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x') \<Longrightarrow>
   (f ---> l) (at x) ==> (g ---> l) (at x)"
   apply (subst within_UNIV[symmetric])
   using Lim_transform_within[of d UNIV x f g l]
@@ -1619,51 +1619,51 @@
 
 text{* Common case assuming being away from some crucial point like 0. *}
 
-lemma Lim_transform_away_within: 
+lemma Lim_transform_away_within:
   fixes f:: "real ^'m \<Rightarrow> real ^'n"
-  assumes "a\<noteq>b" "\<forall>x\<in> S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" 
+  assumes "a\<noteq>b" "\<forall>x\<in> S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
   and "(f ---> l) (at a within S)"
   shows "(g ---> l) (at a within S)"
 proof-
-  have "\<forall>x'\<in>S. 0 < dist x' a \<and> dist x' a < dist a b \<longrightarrow> f x' = g x'" using assms(2) 
-    apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_sym dist_refl) 
+  have "\<forall>x'\<in>S. 0 < dist x' a \<and> dist x' a < dist a b \<longrightarrow> f x' = g x'" using assms(2)
+    apply auto apply(erule_tac x=x' in ballE) by (auto simp add: dist_sym dist_refl)
   thus ?thesis using Lim_transform_within[of "dist a b" S a f g l] using assms(1,3) unfolding dist_nz by auto
 qed
 
-lemma Lim_transform_away_at: 
+lemma Lim_transform_away_at:
   fixes f:: "real ^'m \<Rightarrow> real ^'n"
-  assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" 
+  assumes ab: "a\<noteq>b" and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"
   and fl: "(f ---> l) (at a)"
   shows "(g ---> l) (at a)"
-  using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl 
+  using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl
   by (auto simp add: within_UNIV)
 
 text{* Alternatively, within an open set. *}
 
-lemma Lim_transform_within_open: 
+lemma Lim_transform_within_open:
   fixes f:: "real ^'m \<Rightarrow> real ^'n"
   assumes "open S"  "a \<in> S"  "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"  "(f ---> l) (at a)"
   shows "(g ---> l) (at a)"
 proof-
   from assms(1,2) obtain e::real where "e>0" and e:"ball a e \<subseteq> S" unfolding open_contains_ball by auto
-  hence "\<forall>x'. 0 < dist x' a \<and> dist x' a < e \<longrightarrow> f x' = g x'" using assms(3) 
+  hence "\<forall>x'. 0 < dist x' a \<and> dist x' a < e \<longrightarrow> f x' = g x'" using assms(3)
     unfolding ball_def subset_eq apply auto apply(erule_tac x=x' in allE) apply(erule_tac x=x' in ballE) by(auto simp add: dist_refl dist_sym)
   thus ?thesis using Lim_transform_at[of e a f g l] `e>0` assms(4) by auto
 qed
 
 text{* A congruence rule allowing us to transform limits assuming not at point. *}
 
-lemma Lim_cong_within[cong add]: 
+lemma Lim_cong_within[cong add]:
  "(\<And>x. x \<noteq> a \<Longrightarrow> f x = g x) ==> ((\<lambda>x. f x) ---> l) (at a within S) \<longleftrightarrow> ((g ---> l) (at a within S))"
   by (simp add: Lim_within dist_nz[symmetric])
 
-lemma Lim_cong_at[cong add]: 
+lemma Lim_cong_at[cong add]:
  "(\<And>x. x \<noteq> a ==> f x = g x) ==> (((\<lambda>x. f x) ---> l) (at a) \<longleftrightarrow> ((g ---> l) (at a)))"
   by (simp add: Lim_at dist_nz[symmetric])
 
 text{* Useful lemmas on closure and set of possible sequential limits.*}
 
-lemma closure_sequential: 
+lemma closure_sequential:
  "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)" (is "?lhs = ?rhs")
 proof
   assume "?lhs" moreover
@@ -1679,7 +1679,7 @@
   thus "?lhs" unfolding closure_def unfolding islimpt_sequential by auto
 qed
 
-lemma closed_sequential_limits: 
+lemma closed_sequential_limits:
  "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"
   unfolding closed_limpt
   by (metis closure_sequential closure_closed closed_limpt islimpt_sequential mem_delete)
@@ -1731,7 +1731,7 @@
   }
   thus ?thesis unfolding closed_sequential_limits by auto
 qed
-  
+
 lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0.  cball x e \<subseteq> S)"
 proof-
   { fix x and e::real assume "x\<in>S" "e>0" "ball x e \<subseteq> S"
@@ -1755,7 +1755,7 @@
   assume "?lhs"
   { assume "e \<le> 0"
     hence *:"ball x e = {}" using ball_eq_empty[of x e] by auto
-    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto 
+    have False using `?lhs` unfolding * using islimpt_EMPTY[of y] by auto
   }
   hence "e > 0" by (metis dlo_simps(3))
   moreover
@@ -1774,7 +1774,7 @@
 	case False
 
 	have "dist x (y - (d / (2 * dist y x)) *s (y - x))
-	      = norm (x - y + (d / (2 * norm (y - x))) *s (y - x))"  
+	      = norm (x - y + (d / (2 * norm (y - x))) *s (y - x))"
 	  unfolding mem_cball mem_ball dist_def diff_diff_eq2 diff_add_eq[THEN sym] by auto
 	also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"
 	  using vector_sadd_rdistrib[of "- 1" "d / (2 * norm (y - x))", THEN sym, of "y - x"]
@@ -1788,8 +1788,8 @@
 
 	moreover
 
-	have "(d / (2*dist y x)) *s (y - x) \<noteq> 0" 
-	  using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding vector_mul_eq_0 by (auto simp add: dist_sym dist_refl) 
+	have "(d / (2*dist y x)) *s (y - x) \<noteq> 0"
+	  using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding vector_mul_eq_0 by (auto simp add: dist_sym dist_refl)
 	moreover
 	have "dist (y - (d / (2 * dist y x)) *s (y - x)) y < d" unfolding dist_def apply simp unfolding norm_minus_cancel norm_mul
 	  using `d>0` `x\<noteq>y`[unfolded dist_nz] dist_sym[of x y]
@@ -1803,7 +1803,7 @@
 	case True
 	obtain z where **:"dist y z = (min e d) / 2" using vector_choose_dist[of "(min e d) / 2" y]
 	  using `d > 0` `e>0` by (auto simp add: dist_refl)
-	show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" 
+	show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
 	  apply(rule_tac x=z in bexI) unfolding `x=y` dist_sym dist_refl dist_nz using **  `d > 0` `e>0` by auto
       next
 	case False thus "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"
@@ -1817,7 +1817,7 @@
   apply (simp add: closure_def islimpt_ball expand_set_eq)
   by arith
 
-lemma interior_cball: "interior(cball x e) = ball x e" 
+lemma interior_cball: "interior(cball x e) = ball x e"
 proof(cases "e\<ge>0")
   case False note cs = this
   from cs have "ball x e = {}" using ball_empty[of e x] by auto moreover
@@ -1825,14 +1825,14 @@
     hence False unfolding mem_cball using dist_nz[of x y] cs by (auto simp add: dist_refl)  }
   hence "cball x e = {}" by auto
   hence "interior (cball x e) = {}" using interior_empty by auto
-  ultimately show ?thesis by blast 
+  ultimately show ?thesis by blast
 next
   case True note cs = this
   have "ball x e \<subseteq> cball x e" using ball_subset_cball by auto moreover
   { fix S y assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"
     then obtain d where "d>0" and d:"\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" unfolding open_def by blast
-    
-    then obtain xa where xa:"dist y xa = d / 2" using vector_choose_dist[of "d/2" y] by auto 
+
+    then obtain xa where xa:"dist y xa = d / 2" using vector_choose_dist[of "d/2" y] by auto
     hence xa_y:"xa \<noteq> y" using dist_nz[of y xa] using `d>0` by auto
     have "xa\<in>S" using d[THEN spec[where x=xa]] using xa apply(auto simp add: dist_sym) unfolding dist_nz[THEN sym] using xa_y by auto
     hence xa_cball:"xa \<in> cball x e" using as(1) by auto
@@ -1844,7 +1844,7 @@
     next
       case False
       have "dist (y + (d / 2 / dist y x) *s (y - x)) y < d" unfolding dist_def
-	using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto 
+	using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto
       hence *:"y + (d / 2 / dist y x) *s (y - x) \<in> cball x e" using d as(1)[unfolded subset_eq] by blast
       have "y - x \<noteq> 0" using `x \<noteq> y` by auto
       hence **:"d / (2 * norm (y - x)) > 0" unfolding zero_less_norm_iff[THEN sym]
@@ -1861,10 +1861,10 @@
     qed  }
   hence "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e" by auto
   ultimately show ?thesis using interior_unique[of "ball x e" "cball x e"] using open_ball[of x e] by auto
-qed 
+qed
 
 lemma frontier_ball: "0 < e ==> frontier(ball a e) = {x. dist a x = e}"
-  apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le) 
+  apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le)
   apply (simp add: expand_set_eq)
   by arith
 
@@ -1886,7 +1886,7 @@
     hence "e = 0" using as apply(erule_tac x=y in allE) by (auto simp add: dist_pos_le dist_refl)
   }
   thus ?thesis unfolding expand_set_eq mem_cball by (auto simp add: dist_refl dist_nz dist_le_0)
-qed  
+qed
 
 lemma cball_sing:  "e = 0 ==> cball x e = {x}" by (simp add: cball_eq_sing)
 
@@ -1934,7 +1934,7 @@
     then obtain xa where xa:"\<forall>n. xa n \<in> S"  "(xa ---> x) sequentially" unfolding closure_sequential by auto
     moreover have "\<exists>y. \<exists>x. netord sequentially x y" using trivial_limit_sequentially unfolding trivial_limit_def by blast
     hence "\<exists>y. (\<exists>x. netord sequentially x y) \<and> (\<forall>x. netord sequentially x y \<longrightarrow> norm (xa x) \<le> a)" unfolding sequentially_def using a xa(1) by auto
-    ultimately have "norm x \<le> a" using Lim_norm_ubound[of sequentially xa x a] trivial_limit_sequentially unfolding eventually_def by auto 
+    ultimately have "norm x \<le> a" using Lim_norm_ubound[of sequentially xa x a] trivial_limit_sequentially unfolding eventually_def by auto
   }
   thus ?thesis unfolding bounded_def by auto
 qed
@@ -1955,13 +1955,13 @@
     hence "bounded (insert x F)" unfolding bounded_def by(auto intro!: add exI[of _ "max a (norm x)"])
   }
   thus ?thesis using finite_induct[of S bounded]  using bounded_empty assms by auto
-qed 
+qed
 
 lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"
   apply (auto simp add: bounded_def)
   by (rule_tac x="max a aa" in exI, auto)
 
-lemma bounded_Union[intro]: "finite F \<Longrightarrow> (\<forall>S\<in>F. bounded S) \<Longrightarrow> bounded(\<Union>F)" 
+lemma bounded_Union[intro]: "finite F \<Longrightarrow> (\<forall>S\<in>F. bounded S) \<Longrightarrow> bounded(\<Union>F)"
   by (induct rule: finite_induct[of F], auto)
 
 lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x <= b)"
@@ -1988,9 +1988,9 @@
   then show "\<exists>(x::real^'n). b < norm x"  by blast
 qed
 
-lemma bounded_linear_image: 
+lemma bounded_linear_image:
   fixes f :: "real^'m \<Rightarrow> real^'n"
-  assumes "bounded S" "linear f" 
+  assumes "bounded S" "linear f"
   shows "bounded(f ` S)"
 proof-
   from assms(1) obtain b where b:"b>0" "\<forall>x\<in>S. norm x \<le> b" unfolding bounded_pos by auto
@@ -2014,7 +2014,7 @@
   { fix x assume "x\<in>S"
     hence "norm (a + x) \<le> b + norm a" using norm_triangle_ineq[of a x] b by auto
   }
-  thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"] 
+  thus ?thesis unfolding bounded_pos using norm_ge_zero[of a] b(1) using add_strict_increasing[of b 0 "norm a"]
     by (auto intro!: add exI[of _ "b + norm a"])
 qed
 
@@ -2024,7 +2024,7 @@
 lemma bounded_vec1: "bounded(vec1 ` S) \<longleftrightarrow>  (\<exists>a. \<forall>x\<in>S. abs x <= a)"
   by (simp add: bounded_def forall_vec1 norm_vec1 vec1_in_image_vec1)
 
-lemma bounded_has_rsup: assumes "bounded(vec1 ` S)" "S \<noteq> {}" 
+lemma bounded_has_rsup: assumes "bounded(vec1 ` S)" "S \<noteq> {}"
   shows "\<forall>x\<in>S. x <= rsup S" and "\<forall>b. (\<forall>x\<in>S. x <= b) \<longrightarrow> rsup S <= b"
 proof
   fix x assume "x\<in>S"
@@ -2051,7 +2051,7 @@
   moreover
   have **:"isUb UNIV ?S (max x (rsup S))" unfolding isUb_def setle_def using * by auto
   { fix y assume as:"isUb UNIV (insert x S) y"
-    hence "max x (rsup S) \<le> y" unfolding isUb_def using rsup_le[OF `S\<noteq>{}`] 
+    hence "max x (rsup S) \<le> y" unfolding isUb_def using rsup_le[OF `S\<noteq>{}`]
       unfolding setle_def by auto  }
   hence "max x (rsup S) <=* isUb UNIV (insert x S)" unfolding setge_def Ball_def mem_def by auto
   hence "isLub UNIV ?S (max x (rsup S))" using ** isLubI2[of UNIV ?S "max x (rsup S)"] unfolding Collect_def by auto
@@ -2098,7 +2098,7 @@
   moreover
   have **:"isLb UNIV ?S (min x (rinf S))" unfolding isLb_def setge_def using * by auto
   { fix y assume as:"isLb UNIV (insert x S) y"
-    hence "min x (rinf S) \<ge> y" unfolding isLb_def using rinf_ge[OF `S\<noteq>{}`] 
+    hence "min x (rinf S) \<ge> y" unfolding isLb_def using rinf_ge[OF `S\<noteq>{}`]
       unfolding setge_def by auto  }
   hence "isLb UNIV (insert x S) *<= min x (rinf S)" unfolding setle_def Ball_def mem_def by auto
   hence "isGlb UNIV ?S (min x (rinf S))" using ** isGlbI2[of UNIV ?S "min x (rinf S)"] unfolding Collect_def by auto
@@ -2110,8 +2110,8 @@
 
 subsection{* Compactness (the definition is the one based on convegent subsequences). *}
 
-definition "compact S \<longleftrightarrow> 
-   (\<forall>(f::nat \<Rightarrow> real^'n). (\<forall>n. f n \<in> S) \<longrightarrow> 
+definition "compact S \<longleftrightarrow>
+   (\<forall>(f::nat \<Rightarrow> real^'n). (\<forall>n. f n \<in> S) \<longrightarrow>
        (\<exists>l\<in>S. \<exists>r. (\<forall>m n. m < n \<longrightarrow> r m < r n) \<and> ((f o r) ---> l) sequentially))"
 
 lemma monotone_bigger: fixes r::"nat\<Rightarrow>nat"
@@ -2149,11 +2149,11 @@
 proof-
   have "isUb UNIV (range s) b" using assms(2) and abs_le_D1 unfolding isUb_def and setle_def by auto
   then obtain t where t:"isLub UNIV (range s) t" using reals_complete[of "range s" ] by auto
-  { fix e::real assume "e>0" and as:"\<forall>N. \<exists>n\<ge>N. \<not> \<bar>s n - t\<bar> < e" 
+  { fix e::real assume "e>0" and as:"\<forall>N. \<exists>n\<ge>N. \<not> \<bar>s n - t\<bar> < e"
     { fix n::nat
       obtain N where "N\<ge>n" and n:"\<bar>s N - t\<bar> \<ge> e" using as[THEN spec[where x=n]] by auto
       have "t \<ge> s N" using isLub_isUb[OF t, unfolded isUb_def setle_def] by auto
-      with n have "s N \<le> t - e" using `e>0` by auto 
+      with n have "s N \<le> t - e" using `e>0` by auto
       hence "s n \<le> t - e" using assms(1)[THEN spec[where x=n], THEN spec[where x=N]] using `n\<le>N` by auto  }
     hence "isUb UNIV (range s) (t - e)" unfolding isUb_def and setle_def by auto
     hence False using isLub_le_isUb[OF t, of "t - e"] and `e>0` by auto  }
@@ -2165,20 +2165,20 @@
   shows "\<exists>l. \<forall>e::real>0. \<exists>N. \<forall>n\<ge>N. abs(s n - l) < e"
   using convergent_bounded_increasing[of s b] assms using convergent_bounded_increasing[of "\<lambda>n. - s n" b]
   apply auto unfolding minus_add_distrib[THEN sym, unfolded diff_minus[THEN sym]]
-  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto 
-
-lemma compact_real_lemma: 
+  unfolding abs_minus_cancel by(rule_tac x="-l" in exI)auto
+
+lemma compact_real_lemma:
  assumes "\<forall>n::nat. abs(s n) \<le> b"
   shows "\<exists>l r. (\<forall>m n::nat. m < n --> r m < r n) \<and>
            (\<forall>e>0::real. \<exists>N. \<forall>n\<ge>N. (abs(s (r n) - l) < e))"
 proof-
-  obtain r where r:"\<forall>m n::nat. m < n \<longrightarrow> r m < r n" 
-    "(\<forall>m n. m \<le> n \<longrightarrow> s (r m) \<le> s (r n)) \<or> (\<forall>m n. m \<le> n \<longrightarrow> s (r n) \<le> s (r m))" 
+  obtain r where r:"\<forall>m n::nat. m < n \<longrightarrow> r m < r n"
+    "(\<forall>m n. m \<le> n \<longrightarrow> s (r m) \<le> s (r n)) \<or> (\<forall>m n. m \<le> n \<longrightarrow> s (r n) \<le> s (r m))"
     using seq_monosub[of s] by (auto simp add: subseq_def monoseq_def)
   thus ?thesis using convergent_bounded_monotone[of "s o r" b] and assms by auto
 qed
 
-lemma compact_lemma: 
+lemma compact_lemma:
   assumes "bounded s" and "\<forall>n. (x::nat \<Rightarrow>real^'a) n \<in> s"
   shows "\<forall>d\<in>{1.. dimindex(UNIV::'a set)}.
         \<exists>l::(real^'a). \<exists> r. (\<forall>n m::nat. m < n --> r m < r n) \<and>
@@ -2209,7 +2209,7 @@
 	obtain l1::"real^'a" and r1 where r1:"\<forall>n m::nat. m < n \<longrightarrow> r1 m < r1 n" and lr1:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>i\<in>{1..d}. \<bar>x (r1 n) $ i - l1 $ i\<bar> < e"
 	  using Suc(1)[OF d] by auto
 	obtain l2 r2 where r2:"\<forall>m n::nat. m < n \<longrightarrow> r2 m < r2 n" and lr2:"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>(x \<circ> r1) (r2 n) $ (Suc d) - l2\<bar> < e"
-	  using b'[OF Suc(2)] and compact_real_lemma[of "\<lambda>i. ((x \<circ> r1) i)$(Suc d)" b] by auto 
+	  using b'[OF Suc(2)] and compact_real_lemma[of "\<lambda>i. ((x \<circ> r1) i)$(Suc d)" b] by auto
 	def r \<equiv> "r1 \<circ> r2" have r:"\<forall>m n. m < n \<longrightarrow> r m < r n" unfolding r_def o_def using r1 and r2 by auto
 	moreover
 	def l \<equiv> "(\<chi> i. if i = Suc d then l2 else l1$i)::real^'a"
@@ -2251,13 +2251,13 @@
 	  hence "\<bar>((f \<circ> r) n - l) $ i\<bar> < e / real_of_nat ?d" using `n\<ge>N` using N[THEN spec[where x=n]]
 	    apply auto apply(erule_tac x=i in ballE) unfolding vector_minus_component[OF i] by auto  }
 	ultimately have "(\<Sum>i = 1..?d. \<bar>((f \<circ> r) n - l) $ i\<bar>)
-	  < (\<Sum>i = 1..?d. e / real_of_nat ?d)" 
+	  < (\<Sum>i = 1..?d. e / real_of_nat ?d)"
 	  using setsum_strict_mono[of "{1..?d}" "\<lambda>i. \<bar>((f \<circ> r) n - l) $ i\<bar>" "\<lambda>i. e / (real_of_nat ?d)"] by auto
-	hence "(\<Sum>i = 1..?d. \<bar>((f \<circ> r) n - l) $ i\<bar>) < e" unfolding setsum_constant using dimindex_nonzero[of "UNIV::'a set"] by auto 
+	hence "(\<Sum>i = 1..?d. \<bar>((f \<circ> r) n - l) $ i\<bar>) < e" unfolding setsum_constant using dimindex_nonzero[of "UNIV::'a set"] by auto
 	hence "dist ((f \<circ> r) n) l < e" unfolding dist_def using norm_le_l1[of "(f \<circ> r) n - l"] by auto  }
       hence "\<exists>N. \<forall>n\<ge>N. dist ((f \<circ> r) n) l < e" by auto  }
     hence *:"((f \<circ> r) ---> l) sequentially" unfolding Lim_sequentially by auto
-    moreover have "l\<in>s" 
+    moreover have "l\<in>s"
       using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="f \<circ> r"], THEN spec[where x=l]] and * and as by auto
     ultimately have "\<exists>l\<in>s. \<exists>r. (\<forall>m n. m < n \<longrightarrow> r m < r n) \<and> ((f \<circ> r) ---> l) sequentially" using r by auto  }
   thus ?thesis unfolding compact_def by auto
@@ -2298,7 +2298,7 @@
     by blast
 qed
 
-lemma convergent_imp_cauchy: 
+lemma convergent_imp_cauchy:
  "(s ---> l) sequentially ==> cauchy s"
 proof(simp only: cauchy_def, rule, rule)
   fix e::real assume "e>0" "(s ---> l) sequentially"
@@ -2326,18 +2326,18 @@
 lemma compact_imp_complete: assumes "compact s" shows "complete s"
 proof-
   { fix f assume as: "(\<forall>n::nat. f n \<in> s)" "cauchy f"
-    from as(1) obtain l r where lr: "l\<in>s" "(\<forall>m n. m < n \<longrightarrow> r m < r n)" "((f \<circ> r) ---> l) sequentially" using assms unfolding compact_def by blast 
+    from as(1) obtain l r where lr: "l\<in>s" "(\<forall>m n. m < n \<longrightarrow> r m < r n)" "((f \<circ> r) ---> l) sequentially" using assms unfolding compact_def by blast
 
     { fix n :: nat have lr':"n \<le> r n"
     proof (induct n)
       show "0 \<le> r 0" using lr(2) by blast
-    next fix na assume "na \<le> r na" moreover have "na < Suc na \<longrightarrow> r na < r (Suc na)" using lr(2) by blast 
+    next fix na assume "na \<le> r na" moreover have "na < Suc na \<longrightarrow> r na < r (Suc na)" using lr(2) by blast
       ultimately show "Suc na \<le> r (Suc na)" by auto
     qed } note lr' = this
 
     { fix e::real assume "e>0"
       from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" unfolding cauchy_def using `e>0` apply (erule_tac x="e/2" in allE) by auto
-      from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto 
+      from lr(3)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" using `e>0` by auto
       { fix n::nat assume n:"n \<ge> max N M"
 	have "dist ((f \<circ> r) n) l < e/2" using n M by auto
 	moreover have "r n \<ge> N" using lr'[of n] n by auto
@@ -2348,7 +2348,7 @@
   thus ?thesis unfolding complete_def by auto
 qed
 
-lemma complete_univ: 
+lemma complete_univ:
  "complete UNIV"
 proof(simp add: complete_def, rule, rule)
   fix f::"nat \<Rightarrow> real^'n" assume "cauchy f"
@@ -2394,13 +2394,13 @@
   "helper_1 s e n = (SOME y::real^'n. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))"
 declare helper_1.simps[simp del]
 
-lemma compact_imp_totally_bounded: 
+lemma compact_imp_totally_bounded:
   assumes "compact s"
   shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"
 proof(rule, rule, rule ccontr)
   fix e::real assume "e>0" and assm:"\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>(\<lambda>x. ball x e) ` k)"
   def x \<equiv> "helper_1 s e"
-  { fix n 
+  { fix n
     have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"
     proof(induct_tac rule:nat_less_induct)
       fix n  def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"
@@ -2429,12 +2429,12 @@
 proof(rule ccontr)
   assume "\<not> (\<exists>e>0. \<forall>x\<in>s. \<exists>b\<in>t. ball x e \<subseteq> b)"
   hence cont:"\<forall>e>0. \<exists>x\<in>s. \<forall>xa\<in>t. \<not> (ball x e \<subseteq> xa)" by auto
-  { fix n::nat 
+  { fix n::nat
     have "1 / real (n + 1) > 0" by auto
     hence "\<exists>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> (ball x (inverse (real (n+1))) \<subseteq> xa))" using cont unfolding Bex_def by auto }
   hence "\<forall>n::nat. \<exists>x. x \<in> s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)" by auto
   then obtain f where f:"\<forall>n::nat. f n \<in> s \<and> (\<forall>xa\<in>t. \<not> ball (f n) (inverse (real (n + 1))) \<subseteq> xa)"
-    using choice[of "\<lambda>n::nat. \<lambda>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)"] by auto 
+    using choice[of "\<lambda>n::nat. \<lambda>x. x\<in>s \<and> (\<forall>xa\<in>t. \<not> ball x (inverse (real (n + 1))) \<subseteq> xa)"] by auto
 
   then obtain l r where l:"l\<in>s" and r:"\<forall>m n. m < n \<longrightarrow> r m < r n" and lr:"((f \<circ> r) ---> l) sequentially"
     using assms(1)[unfolded compact_def, THEN spec[where x=f]] by auto
@@ -2445,10 +2445,10 @@
 
   then obtain N1 where N1:"\<forall>n\<ge>N1. dist ((f \<circ> r) n) l < e / 2"
     using lr[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
-  
+
   obtain N2::nat where N2:"N2>0" "inverse (real N2) < e /2" using real_arch_inv[of "e/2"] and `e>0` by auto
   have N2':"inverse (real (r (N1 + N2) +1 )) < e/2"
-    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2 
+    apply(rule order_less_trans) apply(rule less_imp_inverse_less) using N2
     using monotone_bigger[OF r, of "N1 + N2"] by auto
 
   def x \<equiv> "(f (r (N1 + N2)))"
@@ -2456,10 +2456,10 @@
     using f[THEN spec[where x="r (N1 + N2)"]] using `b\<in>t` by auto
   have "\<exists>y\<in>ball x (inverse (real (r (N1 + N2) + 1))). y\<notin>b" apply(rule ccontr) using x by auto
   then obtain y where y:"y \<in> ball x (inverse (real (r (N1 + N2) + 1)))" "y \<notin> b" by auto
-  
+
   have "dist x l < e/2" using N1 unfolding x_def o_def by auto
   hence "dist y l < e" using y N2' using dist_triangle[of y l x]by (auto simp add:dist_sym)
-    
+
   thus False using e and `y\<notin>b` by auto
 qed
 
@@ -2487,13 +2487,13 @@
 
 subsection{* Bolzano-Weierstrass property. *}
 
-lemma heine_borel_imp_bolzano_weierstrass: 
+lemma heine_borel_imp_bolzano_weierstrass:
   assumes "\<forall>f. (\<forall>t \<in> f. open t) \<and> s \<subseteq> (\<Union> f) --> (\<exists>f'. f' \<subseteq> f \<and> finite f' \<and> s \<subseteq> (\<Union> f'))"
           "infinite t"  "t \<subseteq> s"
   shows "\<exists>x \<in> s. x islimpt t"
 proof(rule ccontr)
   assume "\<not> (\<exists>x \<in> s. x islimpt t)"
-  then obtain f where f:"\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)" unfolding islimpt_def 
+  then obtain f where f:"\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)" unfolding islimpt_def
     using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"] by auto
   obtain g where g:"g\<subseteq>{t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
     using assms(1)[THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]] using f by auto
@@ -2508,13 +2508,13 @@
     then obtain y where "y\<in>s" "h = f y" using g'[THEN bspec[where x=h]] by auto
     hence "y = x" using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] by auto
     hence False using `f x \<notin> g` `h\<in>g` unfolding `h = f y` by auto  }
-  hence "f ` t \<subseteq> g" by auto 
-  ultimately show False using g(2) using finite_subset by auto 
+  hence "f ` t \<subseteq> g" by auto
+  ultimately show False using g(2) using finite_subset by auto
 qed
 
 subsection{* Complete the chain of compactness variants. *}
 
-primrec helper_2::"(real \<Rightarrow> real^'n) \<Rightarrow> nat \<Rightarrow> real ^'n" where 
+primrec helper_2::"(real \<Rightarrow> real^'n) \<Rightarrow> nat \<Rightarrow> real ^'n" where
   "helper_2 beyond 0 = beyond 0" |
   "helper_2 beyond (Suc n) = beyond (norm (helper_2 beyond n) + 1 )"
 
@@ -2533,8 +2533,8 @@
     proof(induct n)
       case 0 thus ?case by auto
     next
-      case (Suc n) 
-      have *:"norm (x n) + 1 < norm (x (Suc n))" unfolding x_def and helper_2.simps  
+      case (Suc n)
+      have *:"norm (x n) + 1 < norm (x (Suc n))" unfolding x_def and helper_2.simps
 	using beyond(2)[of "norm (helper_2 beyond n) + 1"] by auto
       thus ?case proof(cases "m < n")
 	case True thus ?thesis using Suc and * by auto
@@ -2546,7 +2546,7 @@
   { fix m n ::nat assume "m\<noteq>n"
     have "1 < dist (x m) (x n)"
     proof(cases "m<n")
-      case True 
+      case True
       hence "1 < norm (x n) - norm (x m)" using *[of m n] by auto
       thus ?thesis unfolding dist_sym[of "x m" "x n"] unfolding dist_def using norm_triangle_sub[of "x n" "x m"] by auto
     next
@@ -2561,10 +2561,10 @@
   { fix n::nat
     have "x n \<in> s"
     proof(cases "n = 0")
-      case True thus ?thesis unfolding x_def using beyond by auto 
+      case True thus ?thesis unfolding x_def using beyond by auto
     next
       case False then obtain z where "n = Suc z" using not0_implies_Suc by auto
-      thus ?thesis unfolding x_def using beyond by auto 
+      thus ?thesis unfolding x_def using beyond by auto
     qed  }
   ultimately have "infinite (range x) \<and> range x \<subseteq> s" unfolding x_def using range_inj_infinite[of "helper_2 beyond"] using beyond(1) by auto
 
@@ -2575,7 +2575,7 @@
   show False using y and z and dist_triangle_half_l[of "x y" l 1 "x z"] and **[of y z] by auto
 qed
 
-lemma sequence_infinite_lemma: 
+lemma sequence_infinite_lemma:
   assumes "\<forall>n::nat. (f n  \<noteq> l)"  "(f ---> l) sequentially"
   shows "infinite {y::real^'a. (\<exists> n. y = f n)}"
 proof(rule ccontr)
@@ -2586,10 +2586,10 @@
   have "0 < Min ?A" using assms(1) unfolding dist_nz unfolding Min_gr_iff[OF **] by auto
   then obtain N where "dist (f N) l < Min ?A" using assms(2)[unfolded Lim_sequentially, THEN spec[where x="Min ?A"]] by auto
   moreover have "dist (f N) l \<in> ?A" by auto
-  ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto  
-qed
-
-lemma sequence_unique_limpt: 
+  ultimately show False using Min_le[OF **(1), of "dist (f N) l"] by auto
+qed
+
+lemma sequence_unique_limpt:
   assumes "\<forall>n::nat. (f n \<noteq> l)"  "(f ---> l) sequentially"  "l' islimpt {y.  (\<exists>n. y = f n)}"
   shows "l' = l"
 proof(rule ccontr)
@@ -2599,14 +2599,14 @@
     using assms(2)[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
   def d \<equiv> "Min (insert (e/2) ((\<lambda>n. if dist (f n) l' = 0 then e/2 else dist (f n) l') ` {0 .. N}))"
   have "d>0" using `e>0` unfolding d_def e_def using dist_pos_le[of _ l', unfolded order_le_less] by auto
-  obtain k where k:"f k \<noteq> l'"  "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto 
+  obtain k where k:"f k \<noteq> l'"  "dist (f k) l' < d" using `d>0` and assms(3)[unfolded islimpt_approachable, THEN spec[where x="d"]] by auto
   have "k\<ge>N" using k(1)[unfolded dist_nz] using k(2)[unfolded d_def]
     by force
   hence "dist l' l < e" using N[THEN spec[where x=k]] using k(2)[unfolded d_def] and dist_triangle_half_r[of "f k" l' e l] by auto
   thus False unfolding e_def by auto
 qed
 
-lemma bolzano_weierstrass_imp_closed: 
+lemma bolzano_weierstrass_imp_closed:
   assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)"
   shows "closed s"
 proof-
@@ -2636,7 +2636,7 @@
   thus ?lhs using bolzano_weierstrass_imp_bounded[of s] bolzano_weierstrass_imp_closed[of s] bounded_closed_imp_compact[of s] by blast
 qed
 
-lemma compact_eq_bolzano_weierstrass: 
+lemma compact_eq_bolzano_weierstrass:
         "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" (is "?lhs = ?rhs")
 proof
   assume ?lhs thus ?rhs unfolding compact_eq_heine_borel using heine_borel_imp_bolzano_weierstrass[of s] by auto
@@ -2644,7 +2644,7 @@
   assume ?rhs thus ?lhs using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed bounded_closed_imp_compact by auto
 qed
 
-lemma compact_eq_bounded_closed: 
+lemma compact_eq_bounded_closed:
  "compact s \<longleftrightarrow> bounded s \<and> closed s"  (is "?lhs = ?rhs")
 proof
   assume ?lhs thus ?rhs unfolding compact_eq_bolzano_weierstrass using bolzano_weierstrass_imp_bounded bolzano_weierstrass_imp_closed by auto
@@ -2652,46 +2652,46 @@
   assume ?rhs thus ?lhs using bounded_closed_imp_compact by auto
 qed
 
-lemma compact_imp_bounded: 
+lemma compact_imp_bounded:
  "compact s ==> bounded s"
   unfolding compact_eq_bounded_closed
   by simp
 
-lemma compact_imp_closed: 
+lemma compact_imp_closed:
  "compact s ==> closed s"
   unfolding compact_eq_bounded_closed
   by simp
 
 text{* In particular, some common special cases. *}
 
-lemma compact_empty[simp]: 
+lemma compact_empty[simp]:
  "compact {}"
   unfolding compact_def
   by simp
 
   (* FIXME : Rename *)
-lemma compact_union[intro]: 
+lemma compact_union[intro]:
  "compact s \<Longrightarrow> compact t ==> compact (s \<union> t)"
   unfolding compact_eq_bounded_closed
   using bounded_Un[of s t]
   using closed_Un[of s t]
   by simp
 
-lemma compact_inter[intro]: 
+lemma compact_inter[intro]:
  "compact s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
   unfolding compact_eq_bounded_closed
   using bounded_Int[of s t]
   using closed_Int[of s t]
   by simp
 
-lemma compact_inter_closed[intro]: 
+lemma compact_inter_closed[intro]:
  "compact s \<Longrightarrow> closed t ==> compact (s \<inter> t)"
   unfolding compact_eq_bounded_closed
   using closed_Int[of s t]
   using bounded_subset[of "s \<inter> t" s]
   by blast
 
-lemma closed_inter_compact[intro]: 
+lemma closed_inter_compact[intro]:
  "closed s \<Longrightarrow> compact t ==> compact (s \<inter> t)"
 proof-
   assume "closed s" "compact t"
@@ -2702,58 +2702,58 @@
     by auto
 qed
 
-lemma finite_imp_closed: 
+lemma finite_imp_closed:
  "finite s ==> closed s"
 proof-
   assume "finite s" hence "\<not>( \<exists>t. t \<subseteq> s \<and> infinite t)" using finite_subset by auto
   thus ?thesis using bolzano_weierstrass_imp_closed[of s] by auto
 qed
 
-lemma finite_imp_compact: 
+lemma finite_imp_compact:
  "finite s ==> compact s"
   unfolding compact_eq_bounded_closed
   using finite_imp_closed finite_imp_bounded
   by blast
 
-lemma compact_sing[simp]: 
+lemma compact_sing[simp]:
  "compact {a}"
   using finite_imp_compact[of "{a}"]
   by blast
 
-lemma closed_sing[simp]: 
+lemma closed_sing[simp]:
  "closed {a}"
   using compact_eq_bounded_closed compact_sing[of a]
   by blast
 
-lemma compact_cball[simp]: 
+lemma compact_cball[simp]:
  "compact(cball x e)"
   using compact_eq_bounded_closed bounded_cball closed_cball
   by blast
 
-lemma compact_frontier_bounded[intro]: 
+lemma compact_frontier_bounded[intro]:
  "bounded s ==> compact(frontier s)"
   unfolding frontier_def
-  using compact_eq_bounded_closed 
+  using compact_eq_bounded_closed
   by blast
 
-lemma compact_frontier[intro]: 
+lemma compact_frontier[intro]:
  "compact s ==> compact (frontier s)"
   using compact_eq_bounded_closed compact_frontier_bounded
   by blast
 
-lemma frontier_subset_compact: 
+lemma frontier_subset_compact:
  "compact s ==> frontier s \<subseteq> s"
   using frontier_subset_closed compact_eq_bounded_closed
   by blast
 
-lemma open_delete: 
+lemma open_delete:
  "open s ==> open(s - {x})"
   using open_diff[of s "{x}"] closed_sing
   by blast
 
 text{* Finite intersection property. I could make it an equivalence in fact. *}
 
-lemma compact_imp_fip: 
+lemma compact_imp_fip:
   assumes "compact s"  "\<forall>t \<in> f. closed t"
         "\<forall>f'. finite f' \<and> f' \<subseteq> f --> (s \<inter> (\<Inter> f') \<noteq> {})"
   shows "s \<inter> (\<Inter> f) \<noteq> {}"
@@ -2769,14 +2769,14 @@
 
 subsection{* Bounded closed nest property (proof does not use Heine-Borel).            *}
 
-lemma bounded_closed_nest: 
+lemma bounded_closed_nest:
   assumes "\<forall>n. closed(s n)" "\<forall>n. (s n \<noteq> {})"
   "(\<forall>m n. m \<le> n --> s n \<subseteq> s m)"  "bounded(s 0)"
   shows "\<exists> a::real^'a. \<forall>n::nat. a \<in> s(n)"
 proof-
   from assms(2) obtain x where x:"\<forall>n::nat. x n \<in> s n" using choice[of "\<lambda>n x. x\<in> s n"] by auto
   from assms(4,1) have *:"compact (s 0)" using bounded_closed_imp_compact[of "s 0"] by auto
-  
+
   then obtain l r where lr:"l\<in>s 0" "\<forall>m n. m < n \<longrightarrow> r m < r n" "((x \<circ> r) ---> l) sequentially"
     unfolding compact_def apply(erule_tac x=x in allE)  using x using assms(3) by blast
 
@@ -2789,22 +2789,22 @@
       hence "(x \<circ> r) (max N n) \<in> s n"
 	using x apply(erule_tac x=n in allE)
 	using x apply(erule_tac x="r (max N n)" in allE)
-	using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto 
+	using assms(3) apply(erule_tac x=n in allE)apply( erule_tac x="r (max N n)" in allE) by auto
       ultimately have "\<exists>y\<in>s n. dist y l < e" by auto
     }
     hence "l \<in> s n" using closed_approachable[of "s n" l] assms(1) by blast
   }
-  thus ?thesis by auto  
+  thus ?thesis by auto
 qed
 
 text{* Decreasing case does not even need compactness, just completeness.        *}
 
-lemma decreasing_closed_nest: 
+lemma decreasing_closed_nest:
   assumes "\<forall>n. closed(s n)"
           "\<forall>n. (s n \<noteq> {})"
           "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
           "\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y \<in> (s n). dist x y < e"
-  shows "\<exists>a::real^'a. \<forall>n::nat. a \<in> s n" 
+  shows "\<exists>a::real^'a. \<forall>n::nat. a \<in> s n"
 proof-
   have "\<forall>n. \<exists> x. x\<in>s n" using assms(2) by auto
   hence "\<exists>t. \<forall>n. t n \<in> s n" using choice[of "\<lambda> n x. x \<in> s n"] by auto
@@ -2832,7 +2832,7 @@
 
 text{* Strengthen it to the intersection actually being a singleton.             *}
 
-lemma decreasing_closed_nest_sing: 
+lemma decreasing_closed_nest_sing:
   assumes "\<forall>n. closed(s n)"
           "\<forall>n. s n \<noteq> {}"
           "\<forall>m n. m \<le> n --> s n \<subseteq> s m"
@@ -2844,7 +2844,7 @@
     { fix e::real assume "e>0"
       hence "dist a b < e" using assms(4 )using b using a by blast
     }
-    hence "dist a b = 0" by (metis dist_eq_0 dist_nz real_less_def)  
+    hence "dist a b = 0" by (metis dist_eq_0 dist_nz real_less_def)
   }
   with a have "\<Inter>{t. \<exists>n. t = s n} = {a}"  unfolding dist_eq_0 by auto
   thus ?thesis by auto
@@ -2863,7 +2863,7 @@
     { fix n m::nat and x::"'b" assume "N \<le> m \<and> N \<le> n \<and> P x"
       hence "dist (s m x) (s n x) < e"
 	using N[THEN spec[where x=m], THEN spec[where x=x]]
-	using N[THEN spec[where x=n], THEN spec[where x=x]] 
+	using N[THEN spec[where x=n], THEN spec[where x=x]]
 	using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  }
     hence "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x  --> dist (s m x) (s n x) < e"  by auto  }
   thus ?rhs by auto
@@ -2881,11 +2881,11 @@
       fix n::nat assume "n\<ge>N"
       hence "dist(s n x)(l x) < e"  using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]
 	using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] by (auto simp add: dist_sym)  }
-    hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }  
+    hence "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" by auto }
   thus ?lhs by auto
 qed
 
-lemma uniformly_cauchy_imp_uniformly_convergent: 
+lemma uniformly_cauchy_imp_uniformly_convergent:
   assumes "\<forall>e>0.\<exists>N. \<forall>m (n::nat) x. N \<le> m \<and> N \<le> n \<and> P x --> dist(s m x)(s n x) < e"
           "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n --> dist(s n x)(l x) < e)"
   shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x --> dist(s n x)(l x) < e"
@@ -2903,7 +2903,7 @@
 
 definition "continuous net f \<longleftrightarrow> (f ---> f(netlimit net)) net"
 
-lemma continuous_trivial_limit: 
+lemma continuous_trivial_limit:
  "trivial_limit net ==> continuous net f"
   unfolding continuous_def tendsto_def eventually_def by auto
 
@@ -2917,7 +2917,7 @@
 lemma continuous_at: "continuous (at x) f \<longleftrightarrow> (f ---> f(x)) (at x)" using within_UNIV[of x]
   using continuous_within[of x UNIV f] by auto
 
-lemma continuous_at_within: 
+lemma continuous_at_within:
   assumes "continuous (at x) f"  shows "continuous (at x within s) f"
 proof(cases "x islimpt s")
   case True show ?thesis using assms unfolding continuous_def and netlimit_at
@@ -2932,8 +2932,8 @@
 
 lemma continuous_within_eps_delta:
   "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s.  dist x' x < d --> dist (f x') (f x) < e)"
-  unfolding continuous_within and Lim_within 
-  apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto 
+  unfolding continuous_within and Lim_within
+  apply auto unfolding dist_nz[THEN sym] apply(auto elim!:allE) apply(rule_tac x=d in exI) by auto
 
 lemma continuous_at_eps_delta: "continuous (at x) f \<longleftrightarrow>  (\<forall>e>0. \<exists>d>0.
                            \<forall>x'. dist x' x < d --> dist(f x')(f x) < e)"
@@ -2942,7 +2942,7 @@
 
 text{* Versions in terms of open balls. *}
 
-lemma continuous_within_ball: 
+lemma continuous_within_ball:
  "continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
                             f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")
 proof
@@ -2957,7 +2957,7 @@
     hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_sym)  }
   thus ?rhs by auto
 next
-  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq 
+  assume ?rhs thus ?lhs unfolding continuous_within Lim_within ball_def subset_eq
     apply (auto simp add: dist_sym) apply(erule_tac x=e in allE) by auto
 qed
 
@@ -2967,7 +2967,7 @@
   assume ?lhs thus ?rhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
     apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x=xa in allE) apply (auto simp add: dist_refl dist_sym dist_nz)
     unfolding dist_nz[THEN sym] by (auto simp add: dist_refl)
-next 
+next
   assume ?rhs thus ?lhs unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball
     apply auto apply(erule_tac x=e in allE) apply auto apply(rule_tac x=d in exI) apply auto apply(erule_tac x="f xa" in allE) by (auto simp add: dist_refl dist_sym dist_nz)
 qed
@@ -2983,11 +2983,11 @@
 
 text{* Some simple consequential lemmas. *}
 
-lemma uniformly_continuous_imp_continuous: 
+lemma uniformly_continuous_imp_continuous:
  " uniformly_continuous_on s f ==> continuous_on s f"
   unfolding uniformly_continuous_on_def continuous_on_def by blast
 
-lemma continuous_at_imp_continuous_within: 
+lemma continuous_at_imp_continuous_within:
  "continuous (at x) f ==> continuous (at x within s) f"
   unfolding continuous_within continuous_at using Lim_at_within by auto
 
@@ -3000,12 +3000,12 @@
   { fix x' assume "\<not> 0 < dist x' x"
     hence "x=x'"
       using dist_nz[of x' x] by auto
-    hence "dist (f x') (f x) < e" using dist_refl[of "f x'"] `e>0` by auto 
+    hence "dist (f x') (f x) < e" using dist_refl[of "f x'"] `e>0` by auto
   }
   thus "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using d by auto
 qed
 
-lemma continuous_on_eq_continuous_within: 
+lemma continuous_on_eq_continuous_within:
  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)" (is "?lhs = ?rhs")
 proof
   assume ?rhs
@@ -3015,7 +3015,7 @@
     then obtain d where "d>0" and d:"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" by auto
     { fix x' assume as:"x'\<in>s" "dist x' x < d"
       hence "dist (f x') (f x) < e" using dist_refl[of "f x'"] `e>0` d `x'\<in>s` dist_eq_0[of x' x] dist_pos_le[of x' x] as(2) by (metis dist_eq_0 dist_nz) }
-    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `d>0` by (auto simp add: dist_refl) 
+    hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `d>0` by (auto simp add: dist_refl)
   }
   thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto
 next
@@ -3023,37 +3023,37 @@
   thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast
 qed
 
-lemma continuous_on: 
+lemma continuous_on:
  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. (f ---> f(x)) (at x within s))"
   by (auto simp add: continuous_on_eq_continuous_within continuous_within)
 
-lemma continuous_on_eq_continuous_at: 
+lemma continuous_on_eq_continuous_at:
  "open s ==> (continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x) f))"
   by (auto simp add: continuous_on continuous_at Lim_within_open)
 
-lemma continuous_within_subset: 
+lemma continuous_within_subset:
  "continuous (at x within s) f \<Longrightarrow> t \<subseteq> s
              ==> continuous (at x within t) f"
   unfolding continuous_within by(metis Lim_within_subset)
 
-lemma continuous_on_subset: 
+lemma continuous_on_subset:
  "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
   unfolding continuous_on by (metis subset_eq Lim_within_subset)
 
-lemma continuous_on_interior: 
+lemma continuous_on_interior:
  "continuous_on s f \<Longrightarrow> x \<in> interior s ==> continuous (at x) f"
 unfolding interior_def
 apply simp
 by (meson continuous_on_eq_continuous_at continuous_on_subset)
 
-lemma continuous_on_eq: 
+lemma continuous_on_eq:
  "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f
            ==> continuous_on s g"
   by (simp add: continuous_on_def)
 
 text{* Characterization of various kinds of continuity in terms of sequences.  *}
 
-lemma continuous_within_sequentially: 
+lemma continuous_within_sequentially:
  "continuous (at a within s) f \<longleftrightarrow>
                 (\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially
                      --> ((f o x) ---> f a) sequentially)" (is "?lhs = ?rhs")
@@ -3066,7 +3066,7 @@
     hence "\<exists>N. \<forall>n\<ge>N. dist ((f \<circ> x) n) (f a) < e"
       apply(rule_tac  x=N in exI) using N d  apply auto using x(1)
       apply(erule_tac x=n in allE) apply(erule_tac x=n in allE)
-      apply(erule_tac x="x n" in ballE)  apply auto unfolding dist_nz[THEN sym] apply auto unfolding dist_refl using `e>0` by auto 
+      apply(erule_tac x="x n" in ballE)  apply auto unfolding dist_nz[THEN sym] apply auto unfolding dist_refl using `e>0` by auto
   }
   thus ?rhs unfolding continuous_within unfolding Lim_sequentially by simp
 next
@@ -3093,12 +3093,12 @@
   thus ?lhs  unfolding continuous_within unfolding Lim_within unfolding Lim_sequentially by blast
 qed
 
-lemma continuous_at_sequentially: 
+lemma continuous_at_sequentially:
  "continuous (at a) f \<longleftrightarrow> (\<forall>x. (x ---> a) sequentially
                   --> ((f o x) ---> f a) sequentially)"
   using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto
 
-lemma continuous_on_sequentially: 
+lemma continuous_on_sequentially:
  "continuous_on s f \<longleftrightarrow>  (\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially
                     --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs")
 proof
@@ -3107,13 +3107,13 @@
   assume ?lhs thus ?rhs unfolding continuous_on_eq_continuous_within using continuous_within_sequentially[of _ s f] by auto
 qed
 
-lemma uniformly_continuous_on_sequentially: 
+lemma uniformly_continuous_on_sequentially:
  "uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>
                     ((\<lambda>n. x n - y n) ---> 0) sequentially
                     \<longrightarrow> ((\<lambda>n. f(x n) - f(y n)) ---> 0) sequentially)" (is "?lhs = ?rhs")
 proof
   assume ?lhs
-  { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. x n - y n) ---> 0) sequentially"    
+  { fix x y assume x:"\<forall>n. x n \<in> s" and y:"\<forall>n. y n \<in> s" and xy:"((\<lambda>n. x n - y n) ---> 0) sequentially"
     { fix e::real assume "e>0"
       then obtain d where "d>0" and d:"\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e"
 	using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto
@@ -3131,7 +3131,7 @@
     then obtain e where "e>0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e" unfolding uniformly_continuous_on_def by auto
     then obtain fa where fa:"\<forall>x.  0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e"
       using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"] unfolding Bex_def
-      by (auto simp add: dist_sym) 
+      by (auto simp add: dist_sym)
     def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"
     def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"
     have xyn:"\<forall>n. x n \<in> s \<and> y n \<in> s" and xy0:"\<forall>n. dist (x n) (y n) < inverse (real n + 1)" and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"
@@ -3152,7 +3152,7 @@
 
 text{* The usual transformation theorems. *}
 
-lemma continuous_transform_within: 
+lemma continuous_transform_within:
   assumes "0 < d" "x \<in> s" "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'"
           "continuous (at x within s) f"
   shows "continuous (at x within s) g"
@@ -3163,8 +3163,8 @@
       hence "dist (f x') (g x) < e" using assms(2,3) apply(erule_tac x=x in ballE) unfolding dist_refl using d' by auto  }
     hence "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
     hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto  }
-  hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto 
-  thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast 
+  hence "(f ---> g x) (at x within s)" unfolding Lim_within using assms(1) by auto
+  thus ?thesis unfolding continuous_within using Lim_transform_within[of d s x f g "g x"] using assms by blast
 qed
 
 lemma continuous_transform_at:
@@ -3180,8 +3180,8 @@
     hence "\<forall>xa. 0 < dist xa x \<and> dist xa x < (min d d') \<longrightarrow> dist (f xa) (g x) < e" by blast
     hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (g x) < e" using `d>0` `d'>0` by(rule_tac x="min d d'" in exI)auto
   }
-  hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto 
-  thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast 
+  hence "(f ---> g x) (at x)" unfolding Lim_at using assms(1) by auto
+  thus ?thesis unfolding continuous_at using Lim_transform_at[of d x f g "g x"] using assms by blast
 qed
 
 text{* Combination results for pointwise continuity. *}
@@ -3189,55 +3189,55 @@
 lemma continuous_const: "continuous net (\<lambda>x::'a::zero_neq_one. c)"
   by(auto simp add: continuous_def Lim_const)
 
-lemma continuous_cmul: 
+lemma continuous_cmul:
  "continuous net f ==> continuous net (\<lambda>x. c *s f x)"
  by(auto simp add: continuous_def Lim_cmul)
-  
-lemma continuous_neg: 
+
+lemma continuous_neg:
  "continuous net f ==> continuous net (\<lambda>x. -(f x))"
  by(auto simp add: continuous_def Lim_neg)
 
-lemma continuous_add: 
+lemma continuous_add:
  "continuous net f \<Longrightarrow> continuous net g
            ==> continuous net (\<lambda>x. f x + g x)"
  by(auto simp add: continuous_def Lim_add)
-  
-lemma continuous_sub: 
+
+lemma continuous_sub:
  "continuous net f \<Longrightarrow> continuous net g
            ==> continuous net (\<lambda>x. f(x) - g(x))"
  by(auto simp add: continuous_def Lim_sub)
-  
+
 text{* Same thing for setwise continuity. *}
 
-lemma continuous_on_const: 
+lemma continuous_on_const:
  "continuous_on s (\<lambda>x. c)"
   unfolding continuous_on_eq_continuous_within using continuous_const by blast
 
-lemma continuous_on_cmul: 
+lemma continuous_on_cmul:
  "continuous_on s f ==>  continuous_on s (\<lambda>x. c *s (f x))"
   unfolding continuous_on_eq_continuous_within using continuous_cmul by blast
 
-lemma continuous_on_neg: 
+lemma continuous_on_neg:
  "continuous_on s f ==> continuous_on s (\<lambda>x. -(f x))"
   unfolding continuous_on_eq_continuous_within using continuous_neg by blast
 
-lemma continuous_on_add: 
+lemma continuous_on_add:
  "continuous_on s f \<Longrightarrow> continuous_on s g
            ==> continuous_on s (\<lambda>x. f x + g x)"
   unfolding continuous_on_eq_continuous_within using continuous_add by blast
 
-lemma continuous_on_sub: 
+lemma continuous_on_sub:
  "continuous_on s f \<Longrightarrow> continuous_on s g
            ==> continuous_on s (\<lambda>x. f(x) - g(x))"
   unfolding continuous_on_eq_continuous_within using continuous_sub by blast
 
 text{* Same thing for uniform continuity, using sequential formulations. *}
 
-lemma uniformly_continuous_on_const: 
+lemma uniformly_continuous_on_const:
  "uniformly_continuous_on s (\<lambda>x. c)"
   unfolding uniformly_continuous_on_sequentially using Lim_const[of 0] by auto
 
-lemma uniformly_continuous_on_cmul: 
+lemma uniformly_continuous_on_cmul:
   assumes "uniformly_continuous_on s f"
   shows "uniformly_continuous_on s (\<lambda>x. c *s f(x))"
 proof-
@@ -3249,12 +3249,12 @@
   thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
 qed
 
-lemma uniformly_continuous_on_neg: 
+lemma uniformly_continuous_on_neg:
  "uniformly_continuous_on s f
          ==> uniformly_continuous_on s (\<lambda>x. -(f x))"
   using uniformly_continuous_on_cmul[of s f "-1"] unfolding pth_3 by auto
-  
-lemma uniformly_continuous_on_add: 
+
+lemma uniformly_continuous_on_add:
   assumes "uniformly_continuous_on s f" "uniformly_continuous_on s g"
   shows "uniformly_continuous_on s (\<lambda>x. f(x) + g(x) ::real^'n)"
 proof-
@@ -3267,7 +3267,7 @@
   thus ?thesis using assms unfolding uniformly_continuous_on_sequentially by auto
 qed
 
-lemma uniformly_continuous_on_sub: 
+lemma uniformly_continuous_on_sub:
  "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s g
            ==> uniformly_continuous_on s  (\<lambda>x. f x - g x)"
   unfolding ab_diff_minus
@@ -3276,51 +3276,51 @@
 
 text{* Identity function is continuous in every sense. *}
 
-lemma continuous_within_id: 
+lemma continuous_within_id:
  "continuous (at a within s) (\<lambda>x. x)"
   unfolding continuous_within Lim_within by auto
-  
-lemma continuous_at_id: 
+
+lemma continuous_at_id:
  "continuous (at a) (\<lambda>x. x)"
   unfolding continuous_at Lim_at by auto
-  
-lemma continuous_on_id: 
+
+lemma continuous_on_id:
  "continuous_on s (\<lambda>x. x)"
   unfolding continuous_on Lim_within by auto
 
-lemma uniformly_continuous_on_id: 
+lemma uniformly_continuous_on_id:
  "uniformly_continuous_on s (\<lambda>x. x)"
   unfolding uniformly_continuous_on_def by auto
-  
+
 text{* Continuity of all kinds is preserved under composition. *}
 
-lemma continuous_within_compose: 
+lemma continuous_within_compose:
   assumes "continuous (at x within s) f"   "continuous (at (f x) within f ` s) g"
   shows "continuous (at x within s) (g o f)"
 proof-
   { fix e::real assume "e>0"
     with assms(2)[unfolded continuous_within Lim_within] obtain d  where "d>0" and d:"\<forall>xa\<in>f ` s. 0 < dist xa (f x) \<and> dist xa (f x) < d \<longrightarrow> dist (g xa) (g (f x)) < e" by auto
     from assms(1)[unfolded continuous_within Lim_within] obtain d' where "d'>0" and d':"\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d' \<longrightarrow> dist (f xa) (f x) < d" using `d>0` by auto
-    { fix y assume as:"y\<in>s"  "0 < dist y x"  "dist y x < d'" 
+    { fix y assume as:"y\<in>s"  "0 < dist y x"  "dist y x < d'"
       hence "dist (f y) (f x) < d" using d'[THEN bspec[where x=y]] by (auto simp add:dist_sym)
       hence "dist (g (f y)) (g (f x)) < e" using as(1) d[THEN bspec[where x="f y"]] unfolding dist_nz[THEN sym] using `e>0` by (auto simp add: dist_refl)   }
     hence "\<exists>d>0. \<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (g (f xa)) (g (f x)) < e" using `d'>0` by auto  }
   thus ?thesis unfolding continuous_within Lim_within by auto
 qed
 
-lemma continuous_at_compose: 
-  assumes "continuous (at x) f"  "continuous (at (f x)) g"  
+lemma continuous_at_compose:
+  assumes "continuous (at x) f"  "continuous (at (f x)) g"
   shows "continuous (at x) (g o f)"
 proof-
   have " continuous (at (f x) within range f) g" using assms(2) using continuous_within_subset[of "f x" UNIV g "range f", unfolded within_UNIV] by auto
   thus ?thesis using assms(1) using continuous_within_compose[of x UNIV f g, unfolded within_UNIV] by auto
 qed
 
-lemma continuous_on_compose: 
+lemma continuous_on_compose:
  "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)"
   unfolding continuous_on_eq_continuous_within using continuous_within_compose[of _ s f g] by auto
 
-lemma uniformly_continuous_on_compose: 
+lemma uniformly_continuous_on_compose:
   assumes "uniformly_continuous_on s f"  "uniformly_continuous_on (f ` s) g"
   shows "uniformly_continuous_on s (g o f)"
 proof-
@@ -3333,7 +3333,7 @@
 
 text{* Continuity in terms of open preimages. *}
 
-lemma continuous_at_open: 
+lemma continuous_at_open:
  "continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))" (is "?lhs = ?rhs")
 proof
   assume ?lhs
@@ -3341,7 +3341,7 @@
     then obtain e where "e>0" and e:"ball (f x) e \<subseteq> t" unfolding open_contains_ball by auto
 
     obtain d where "d>0" and d:"\<forall>y. 0 < dist y x \<and> dist y x < d \<longrightarrow> dist (f y) (f x) < e" using `e>0` using `?lhs`[unfolded continuous_at Lim_at open_def] by auto
-    
+
     have "open (ball x d)" using open_ball by auto
     moreover have "x \<in> ball x d" unfolding centre_in_ball using `d>0` by simp
     moreover
@@ -3349,8 +3349,8 @@
 	using e[unfolded subset_eq Ball_def mem_ball, THEN spec[where x="f x'"]]    d[THEN spec[where x=x']]
 	unfolding mem_ball apply (auto simp add: dist_sym)
 	unfolding dist_nz[THEN sym] using as(2) by auto  }
-    hence "\<forall>x'\<in>ball x d. f x' \<in> t" by auto 
-    ultimately have "\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x'\<in>s. f x' \<in> t)" 
+    hence "\<forall>x'\<in>ball x d. f x' \<in> t" by auto
+    ultimately have "\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x'\<in>s. f x' \<in> t)"
       apply(rule_tac x="ball x d" in exI) by simp  }
   thus ?rhs by auto
 next
@@ -3363,19 +3363,19 @@
       hence "dist (f y) (f x) < e" using d[unfolded subset_eq Ball_def mem_ball, THEN spec[where x=y]]
 	using s(3)[THEN bspec[where x=y], unfolded mem_ball] by (auto simp add: dist_sym)  }
     hence "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `d>0` by auto  }
-  thus ?lhs unfolding continuous_at Lim_at by auto 
-qed
-
-lemma continuous_on_open: 
+  thus ?lhs unfolding continuous_at Lim_at by auto
+qed
+
+lemma continuous_on_open:
  "continuous_on s f \<longleftrightarrow>
         (\<forall>t. openin (subtopology euclidean (f ` s)) t
             --> openin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
 proof
-  assume ?lhs 
+  assume ?lhs
   { fix t assume as:"openin (subtopology euclidean (f ` s)) t"
     have "{x \<in> s. f x \<in> t} \<subseteq> s" using as[unfolded openin_euclidean_subtopology_iff] by auto
-    moreover 
-    { fix x assume as':"x\<in>{x \<in> s. f x \<in> t}" 
+    moreover
+    { fix x assume as':"x\<in>{x \<in> s. f x \<in> t}"
       then obtain e where e: "e>0" "\<forall>x'\<in>f ` s. dist x' (f x) < e \<longrightarrow> x' \<in> t" using as[unfolded openin_euclidean_subtopology_iff, THEN conjunct2, THEN bspec[where x="f x"]] by auto
       from this(1) obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" using `?lhs`[unfolded continuous_on Lim_within, THEN bspec[where x=x]] using as' by auto
       have "\<exists>e>0. \<forall>x'\<in>s. dist x' x < e \<longrightarrow> x' \<in> {x \<in> s. f x \<in> t}" using d e unfolding dist_nz[THEN sym] by (rule_tac x=d in exI, auto simp add: dist_refl)  }
@@ -3383,11 +3383,11 @@
   thus ?rhs unfolding continuous_on Lim_within using openin by auto
 next
   assume ?rhs
-  { fix e::real and x assume "x\<in>s" "e>0" 
+  { fix e::real and x assume "x\<in>s" "e>0"
     { fix xa x' assume "dist (f xa) (f x) < e" "xa \<in> s" "x' \<in> s" "dist (f xa) (f x') < e - dist (f xa) (f x)"
-      hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"] 
+      hence "dist (f x') (f x) < e" using dist_triangle[of "f x'" "f x" "f xa"]
 	by (auto simp add: dist_sym)  }
-    hence "ball (f x) e \<inter> f ` s \<subseteq> f ` s \<and> (\<forall>xa\<in>ball (f x) e \<inter> f ` s. \<exists>ea>0. \<forall>x'\<in>f ` s. dist x' xa < ea \<longrightarrow> x' \<in> ball (f x) e \<inter> f ` s)" apply auto 
+    hence "ball (f x) e \<inter> f ` s \<subseteq> f ` s \<and> (\<forall>xa\<in>ball (f x) e \<inter> f ` s. \<exists>ea>0. \<forall>x'\<in>f ` s. dist x' xa < ea \<longrightarrow> x' \<in> ball (f x) e \<inter> f ` s)" apply auto
       apply(rule_tac x="e - dist (f xa) (f x)" in exI) using `e>0` by (auto simp add: dist_sym)
     hence "\<forall>xa\<in>{xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}. \<exists>ea>0. \<forall>x'\<in>s. dist x' xa < ea \<longrightarrow> x' \<in> {xa \<in> s. f xa \<in> ball (f x) e \<inter> f ` s}"
       using `?rhs`[unfolded openin_euclidean_subtopology_iff, THEN spec[where x="ball (f x) e \<inter> f ` s"]] by auto
@@ -3399,7 +3399,7 @@
 (* Similarly in terms of closed sets.                                        *)
 (* ------------------------------------------------------------------------- *)
 
-lemma continuous_on_closed: 
+lemma continuous_on_closed:
  "continuous_on s f \<longleftrightarrow>  (\<forall>t. closedin (subtopology euclidean (f ` s)) t  --> closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
 proof
   assume ?lhs
@@ -3423,7 +3423,7 @@
 
 text{* Half-global and completely global cases.                                  *}
 
-lemma continuous_open_in_preimage: 
+lemma continuous_open_in_preimage:
   assumes "continuous_on s f"  "open t"
   shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
 proof-
@@ -3433,7 +3433,7 @@
   thus ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]] using * by auto
 qed
 
-lemma continuous_closed_in_preimage: 
+lemma continuous_closed_in_preimage:
   assumes "continuous_on s f"  "closed t"
   shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
 proof-
@@ -3444,50 +3444,50 @@
     using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]] using * by auto
 qed
 
-lemma continuous_open_preimage: 
+lemma continuous_open_preimage:
   assumes "continuous_on s f" "open s" "open t"
   shows "open {x \<in> s. f x \<in> t}"
 proof-
-  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T" 
+  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
     using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto
   thus ?thesis using open_inter[of s T, OF assms(2)] by auto
 qed
 
-lemma continuous_closed_preimage: 
+lemma continuous_closed_preimage:
   assumes "continuous_on s f" "closed s" "closed t"
   shows "closed {x \<in> s. f x \<in> t}"
 proof-
-  obtain T where T: "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T" 
+  obtain T where T: "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
     using continuous_closed_in_preimage[OF assms(1,3)] unfolding closedin_closed by auto
   thus ?thesis using closed_Int[of s T, OF assms(2)] by auto
 qed
 
-lemma continuous_open_preimage_univ: 
+lemma continuous_open_preimage_univ:
  "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
   using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto
 
-lemma continuous_closed_preimage_univ: 
+lemma continuous_closed_preimage_univ:
  "(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s ==> closed {x. f x \<in> s}"
   using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto
 
 text{* Equality of continuous functions on closure and related results.          *}
 
-lemma continuous_closed_in_preimage_constant: 
+lemma continuous_closed_in_preimage_constant:
  "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
   using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto
 
-lemma continuous_closed_preimage_constant: 
+lemma continuous_closed_preimage_constant:
  "continuous_on s f \<Longrightarrow> closed s ==> closed {x \<in> s. f x = a}"
   using continuous_closed_preimage[of s f "{a}"] closed_sing by auto
 
-lemma continuous_constant_on_closure: 
+lemma continuous_constant_on_closure:
   assumes "continuous_on (closure s) f"
           "\<forall>x \<in> s. f x = a"
   shows "\<forall>x \<in> (closure s). f x = a"
-    using continuous_closed_preimage_constant[of "closure s" f a] 
+    using continuous_closed_preimage_constant[of "closure s" f a]
     assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset unfolding subset_eq by auto
 
-lemma image_closure_subset: 
+lemma image_closure_subset:
   assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
   shows "f ` (closure s) \<subseteq> t"
 proof-
@@ -3499,7 +3499,7 @@
   thus ?thesis by auto
 qed
 
-lemma continuous_on_closure_norm_le: 
+lemma continuous_on_closure_norm_le:
   assumes "continuous_on (closure s) f"  "\<forall>y \<in> s. norm(f y) \<le> b"  "x \<in> (closure s)"
   shows "norm(f x) \<le> b"
 proof-
@@ -3511,7 +3511,7 @@
 
 text{* Making a continuous function avoid some value in a neighbourhood.         *}
 
-lemma continuous_within_avoid: 
+lemma continuous_within_avoid:
   assumes "continuous (at x within s) f"  "x \<in> s"  "f x \<noteq> a"
   shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"
 proof-
@@ -3523,44 +3523,44 @@
   thus ?thesis using `d>0` by auto
 qed
 
-lemma continuous_at_avoid: 
+lemma continuous_at_avoid:
   assumes "continuous (at x) f"  "f x \<noteq> a"
   shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
 using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto
 
-lemma continuous_on_avoid: 
+lemma continuous_on_avoid:
   assumes "continuous_on s f"  "x \<in> s"  "f x \<noteq> a"
   shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"
 using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], OF assms(2)]  continuous_within_avoid[of x s f a]  assms(2,3) by auto
 
-lemma continuous_on_open_avoid: 
+lemma continuous_on_open_avoid:
   assumes "continuous_on s f"  "open s"  "x \<in> s"  "f x \<noteq> a"
   shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"
 using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]  continuous_at_avoid[of x f a]  assms(3,4) by auto
 
 text{* Proving a function is constant by proving open-ness of level set.         *}
 
-lemma continuous_levelset_open_in_cases: 
+lemma continuous_levelset_open_in_cases:
  "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
         openin (subtopology euclidean s) {x \<in> s. f x = a}
         ==> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
 unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
 
-lemma continuous_levelset_open_in: 
+lemma continuous_levelset_open_in:
  "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
         openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
         (\<exists>x \<in> s. f x = a)  ==> (\<forall>x \<in> s. f x = a)"
 using continuous_levelset_open_in_cases[of s f ]
 by meson
 
-lemma continuous_levelset_open: 
+lemma continuous_levelset_open:
   assumes "connected s"  "continuous_on s f"  "open {x \<in> s. f x = a}"  "\<exists>x \<in> s.  f x = a"
   shows "\<forall>x \<in> s. f x = a"
 using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by auto
 
 text{* Some arithmetical combinations (more to prove).                           *}
 
-lemma open_scaling[intro]: 
+lemma open_scaling[intro]:
   assumes "c \<noteq> 0"  "open s"
   shows "open((\<lambda>x. c *s x) ` s)"
 proof-
@@ -3569,7 +3569,7 @@
     have "e * abs c > 0" using assms(1)[unfolded zero_less_abs_iff[THEN sym]] using real_mult_order[OF `e>0`] by auto
     moreover
     { fix y assume "dist y (c *s x) < e * \<bar>c\<bar>"
-      hence "norm ((1 / c) *s y - x) < e" unfolding dist_def 
+      hence "norm ((1 / c) *s y - x) < e" unfolding dist_def
 	using norm_mul[of c "(1 / c) *s y - x", unfolded vector_ssub_ldistrib, unfolded vector_smult_assoc] assms(1)
 	  mult_less_imp_less_left[of "abs c" "norm ((1 / c) *s y - x)" e, unfolded real_mult_commute[of "abs c" e]] assms(1)[unfolded zero_less_abs_iff[THEN sym]] by simp
       hence "y \<in> op *s c ` s" using rev_image_eqI[of "(1 / c) *s y" s y "op *s c"]  e[THEN spec[where x="(1 / c) *s y"]]  assms(1) unfolding dist_def vector_smult_assoc by auto  }
@@ -3577,10 +3577,10 @@
   thus ?thesis unfolding open_def by auto
 qed
 
-lemma open_negations: 
+lemma open_negations:
  "open s ==> open ((\<lambda> x. -x) ` s)" unfolding pth_3 by auto
-  
-lemma open_translation: 
+
+lemma open_translation:
   assumes "open s"  shows "open((\<lambda>x. a + x) ` s)"
 proof-
   { fix x have "continuous (at x) (\<lambda>x. x - a)" using continuous_sub[of "at x" "\<lambda>x. x" "\<lambda>x. a"] continuous_at_id[of x] continuous_const[of "at x" a] by auto  }
@@ -3588,7 +3588,7 @@
   ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s] using assms by auto
 qed
 
-lemma open_affinity: 
+lemma open_affinity:
   assumes "open s"  "c \<noteq> 0"
   shows "open ((\<lambda>x. a + c *s x) ` s)"
 proof-
@@ -3599,12 +3599,12 @@
 
 lemma interior_translation: "interior ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (interior s)"
 proof (rule set_ext, rule)
-  fix x assume "x \<in> interior (op + a ` s)" 
+  fix x assume "x \<in> interior (op + a ` s)"
   then obtain e where "e>0" and e:"ball x e \<subseteq> op + a ` s" unfolding mem_interior by auto
   hence "ball (x - a) e \<subseteq> s" unfolding subset_eq Ball_def mem_ball dist_def apply auto apply(erule_tac x="a + xa" in allE) unfolding ab_group_add_class.diff_diff_eq[THEN sym] by auto
   thus "x \<in> op + a ` interior s" unfolding image_iff apply(rule_tac x="x - a" in bexI) unfolding mem_interior using `e > 0` by auto
 next
-  fix x assume "x \<in> op + a ` interior s" 
+  fix x assume "x \<in> op + a ` interior s"
   then obtain y e where "e>0" and e:"ball y e \<subseteq> s" and y:"x = a + y" unfolding image_iff Bex_def mem_interior by auto
   { fix z have *:"a + y - z = y + a - z" by auto
     assume "z\<in>ball x e"
@@ -3616,7 +3616,7 @@
 
 subsection {* Preservation of compactness and connectedness under continuous function.  *}
 
-lemma compact_continuous_image: 
+lemma compact_continuous_image:
   assumes "continuous_on s f"  "compact s"
   shows "compact(f ` s)"
 proof-
@@ -3629,10 +3629,10 @@
       { fix n::nat assume "n\<ge>N" hence "dist ((x \<circ> r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto  }
       hence "\<exists>N. \<forall>n\<ge>N. dist ((x \<circ> r) n) (f l) < e" by auto  }
     hence "\<exists>l\<in>f ` s. \<exists>r. (\<forall>m n. m < n \<longrightarrow> r m < r n) \<and> ((x \<circ> r) ---> l) sequentially" unfolding Lim_sequentially using r lr `l\<in>s` by auto  }
-  thus ?thesis unfolding compact_def by auto 
-qed
-
-lemma connected_continuous_image: 
+  thus ?thesis unfolding compact_def by auto
+qed
+
+lemma connected_continuous_image:
   assumes "continuous_on s f"  "connected s"
   shows "connected(f ` s)"
 proof-
@@ -3641,14 +3641,14 @@
       using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]
       using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]
       using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
-    hence False using as(1,2) 
+    hence False using as(1,2)
       using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto }
   thus ?thesis unfolding connected_clopen by auto
 qed
 
 text{* Continuity implies uniform continuity on a compact domain.                *}
 
-lemma compact_uniformly_continuous: 
+lemma compact_uniformly_continuous:
   assumes "continuous_on s f"  "compact s"
   shows "uniformly_continuous_on s f"
 proof-
@@ -3656,7 +3656,7 @@
       hence "\<forall>xa. \<exists>y. 0 < xa \<longrightarrow> (y > 0 \<and> (\<forall>x'\<in>s. dist x' x < y \<longrightarrow> dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto
       hence "\<exists>fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)" using choice[of "\<lambda>e d. e>0 \<longrightarrow> d>0 \<and>(\<forall>x'\<in>s. (dist x' x < d \<longrightarrow> dist (f x') (f x) < e))"] by auto  }
     then have "\<forall>x\<in>s. \<exists>y. \<forall>xa. 0 < xa \<longrightarrow> (\<forall>x'\<in>s. y xa > 0 \<and> (dist x' x < y xa \<longrightarrow> dist (f x') (f x) < xa))" by auto
-    then obtain d where d:"\<forall>e>0. \<forall>x\<in>s. \<forall>x'\<in>s. d x e > 0 \<and> (dist x' x < d x e \<longrightarrow> dist (f x') (f x) < e)"  
+    then obtain d where d:"\<forall>e>0. \<forall>x\<in>s. \<forall>x'\<in>s. d x e > 0 \<and> (dist x' x < d x e \<longrightarrow> dist (f x') (f x) < e)"
       using bchoice[of s "\<lambda>x fa. \<forall>xa>0. \<forall>x'\<in>s. fa xa > 0 \<and> (dist x' x < fa xa \<longrightarrow> dist (f x') (f x) < xa)"] by blast
 
   { fix e::real assume "e>0"
@@ -3672,11 +3672,11 @@
       hence "x\<in>ball z (d z (e / 2))" using `ea>0` unfolding subset_eq by auto
       hence "dist (f z) (f x) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `x\<in>s` and `z\<in>s`
 	by (auto  simp add: dist_sym)
-      moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq] 
+      moreover have "y\<in>ball z (d z (e / 2))" using as and `ea>0` and z[unfolded subset_eq]
 	by (auto simp add: dist_sym)
-      hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s` 
+      hence "dist (f z) (f y) < e / 2" using d[THEN spec[where x="e/2"]] and `e>0` and `y\<in>s` and `z\<in>s`
 	by (auto  simp add: dist_sym)
-      ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"] 
+      ultimately have "dist (f y) (f x) < e" using dist_triangle_half_r[of "f z" "f x" e "f y"]
 	by (auto simp add: dist_sym)  }
     then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `ea>0` by auto  }
   thus ?thesis unfolding uniformly_continuous_on_def by auto
@@ -3684,16 +3684,16 @@
 
 text{* Continuity of inverse function on compact domain. *}
 
-lemma continuous_on_inverse: 
+lemma continuous_on_inverse:
   assumes "continuous_on s f"  "compact s"  "\<forall>x \<in> s. g (f x) = x"
   shows "continuous_on (f ` s) g"
 proof-
   have *:"g ` f ` s = s" using assms(3) by (auto simp add: image_iff)
   { fix t assume t:"closedin (subtopology euclidean (g ` f ` s)) t"
     then obtain T where T: "closed T" "t = s \<inter> T" unfolding closedin_closed unfolding * by auto
-    have "continuous_on (s \<inter> T) f" using continuous_on_subset[OF assms(1), of "s \<inter> t"] 
+    have "continuous_on (s \<inter> T) f" using continuous_on_subset[OF assms(1), of "s \<inter> t"]
       unfolding T(2) and Int_left_absorb by auto
-    moreover have "compact (s \<inter> T)" 
+    moreover have "compact (s \<inter> T)"
       using assms(2) unfolding compact_eq_bounded_closed
       using bounded_subset[of s "s \<inter> T"] and T(1) by auto
     ultimately have "closed (f ` t)" using T(1) unfolding T(2)
@@ -3706,7 +3706,7 @@
 
 subsection{* A uniformly convergent limit of continuous functions is continuous.       *}
 
-lemma continuous_uniform_limit: 
+lemma continuous_uniform_limit:
   assumes "\<not> (trivial_limit net)"  "eventually (\<lambda>n. continuous_on s (f n)) net"
   "\<forall>e>0. eventually (\<lambda>n. \<forall>x \<in> s. norm(f n x - g x) < e) net"
   shows "continuous_on s g"
@@ -3720,7 +3720,7 @@
       using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
     { fix y assume "y\<in>s" "dist y x < d"
       hence "dist (f n y) (f n x) < e / 3" using d[THEN bspec[where x=y]] by auto
-      hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"] 
+      hence "norm (f n y - g x) < 2 * e / 3" using norm_triangle_lt[of "f n y - f n x" "f n x - g x" "2*e/3"]
 	using n(1)[THEN bspec[where x=x], OF `x\<in>s`] unfolding dist_def unfolding ab_group_add_class.ab_diff_minus by auto
       hence "dist (g y) (g x) < e" unfolding dist_def using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
 	unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
@@ -3738,37 +3738,37 @@
     { fix x::"real^'a" assume "norm x < e / B"
       hence "B * norm x < e" using `B>0` using mult_strict_right_mono[of "norm x" " e / B" B] unfolding real_mult_commute by auto
       hence "norm (f x) < e" using B[THEN spec[where x=x]] `B>0` using order_le_less_trans[of "norm (f x)" "B * norm x" e] by auto   }
-    moreover have "e / B > 0" using `e>0` `B>0` divide_pos_pos by auto 
+    moreover have "e / B > 0" using `e>0` `B>0` divide_pos_pos by auto
     ultimately have "\<exists>d>0. \<forall>x. 0 < dist x 0 \<and> dist x 0 < d \<longrightarrow> dist (f x) 0 < e" unfolding dist_def by auto  }
   thus ?thesis unfolding Lim_at by auto
 qed
 
-lemma linear_continuous_at: 
+lemma linear_continuous_at:
   assumes "linear f"  shows "continuous (at a) f"
   unfolding continuous_at Lim_at_zero[of f "f a" a] using linear_lim_0[OF assms]
-  unfolding Lim_null[of "\<lambda>x. f (a + x)"] unfolding linear_sub[OF assms, THEN sym] by auto 
-
-lemma linear_continuous_within: 
+  unfolding Lim_null[of "\<lambda>x. f (a + x)"] unfolding linear_sub[OF assms, THEN sym] by auto
+
+lemma linear_continuous_within:
  "linear f ==> continuous (at x within s) f"
   using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto
 
-lemma linear_continuous_on: 
+lemma linear_continuous_on:
  "linear f ==> continuous_on s f"
   using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto
 
 text{* Also bilinear functions, in composition form.                             *}
 
-lemma bilinear_continuous_at_compose: 
+lemma bilinear_continuous_at_compose:
  "continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bilinear h
         ==> continuous (at x) (\<lambda>x. h (f x) (g x))"
   unfolding continuous_at using Lim_bilinear[of f "f x" "(at x)" g "g x" h] by auto
-  
-lemma bilinear_continuous_within_compose: 
+
+lemma bilinear_continuous_within_compose:
  "continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bilinear h
         ==> continuous (at x within s) (\<lambda>x. h (f x) (g x))"
   unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
-  
-lemma bilinear_continuous_on_compose: 
+
+lemma bilinear_continuous_on_compose:
  "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bilinear h
              ==> continuous_on s (\<lambda>x. h (f x) (g x))"
   unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto
@@ -3777,43 +3777,43 @@
 subsection{* Topological stuff lifted from and dropped to R                            *}
 
 
-lemma open_vec1: 
+lemma open_vec1:
  "open(vec1 ` s) \<longleftrightarrow>
         (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)" (is "?lhs = ?rhs")
-  unfolding open_def apply simp unfolding forall_vec1 dist_vec1 vec1_in_image_vec1 by simp 
-  
-lemma islimpt_approachable_vec1: 
+  unfolding open_def apply simp unfolding forall_vec1 dist_vec1 vec1_in_image_vec1 by simp
+
+lemma islimpt_approachable_vec1:
  "(vec1 x) islimpt (vec1 ` s) \<longleftrightarrow>
          (\<forall>e>0.  \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)"
   by (auto simp add: islimpt_approachable dist_vec1 vec1_eq)
 
-lemma closed_vec1: 
+lemma closed_vec1:
  "closed (vec1 ` s) \<longleftrightarrow>
         (\<forall>x. (\<forall>e>0.  \<exists>x' \<in> s. x' \<noteq> x \<and> abs(x' - x) < e)
             --> x \<in> s)"
   unfolding closed_limpt islimpt_approachable forall_vec1 apply simp
   unfolding dist_vec1 vec1_in_image_vec1 abs_minus_commute by auto
 
-lemma continuous_at_vec1_range: 
+lemma continuous_at_vec1_range:
  "continuous (at x) (vec1 o f) \<longleftrightarrow> (\<forall>e>0. \<exists>d>0.
         \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"
   unfolding continuous_at unfolding Lim_at apply simp unfolding dist_vec1 unfolding dist_nz[THEN sym] unfolding dist_def apply auto
   apply(erule_tac x=e in allE) apply auto apply (rule_tac x=d in exI) apply auto apply (erule_tac x=x' in allE) apply auto
   apply(erule_tac x=e in allE) by auto
 
-lemma continuous_on_vec1_range: 
+lemma continuous_on_vec1_range:
  " continuous_on s (vec1 o f) \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d --> abs(f x' - f x) < e))"
   unfolding continuous_on_def apply (simp del: dist_sym) unfolding dist_vec1 unfolding dist_def ..
-  
-lemma continuous_at_vec1_norm: 
+
+lemma continuous_at_vec1_norm:
  "\<forall>x. continuous (at x) (vec1 o norm)"
   unfolding continuous_at_vec1_range using real_abs_sub_norm order_le_less_trans by blast
 
 lemma continuous_on_vec1_norm:
  "\<forall>s. continuous_on s (vec1 o norm)"
-unfolding continuous_on_vec1_range norm_vec1[THEN sym] by (metis norm_vec1 order_le_less_trans real_abs_sub_norm) 
-
-lemma continuous_at_vec1_component: 
+unfolding continuous_on_vec1_range norm_vec1[THEN sym] by (metis norm_vec1 order_le_less_trans real_abs_sub_norm)
+
+lemma continuous_at_vec1_component:
   assumes "1 \<le> i" "i \<le> dimindex(UNIV::('a set))"
   shows "continuous (at (a::real^'a)) (\<lambda> x. vec1(x$i))"
 proof-
@@ -3822,7 +3822,7 @@
   thus ?thesis unfolding continuous_at tendsto_def eventually_at dist_vec1 by auto
 qed
 
-lemma continuous_on_vec1_component: 
+lemma continuous_on_vec1_component:
   assumes "i \<in> {1..dimindex (UNIV::'a set)}"  shows "continuous_on s (\<lambda> x::real^'a. vec1(x$i))"
 proof-
   { fix e::real and x xa assume "x\<in>s" "e>0" "xa\<in>s" "0 < norm (xa - x) \<and> norm (xa - x) < e"
@@ -3830,15 +3830,15 @@
   thus ?thesis unfolding continuous_on Lim_within dist_vec1 unfolding dist_def by auto
 qed
 
-lemma continuous_at_vec1_infnorm: 
+lemma continuous_at_vec1_infnorm:
  "continuous (at x) (vec1 o infnorm)"
-  unfolding continuous_at Lim_at o_def unfolding dist_vec1 unfolding dist_def 
+  unfolding continuous_at Lim_at o_def unfolding dist_vec1 unfolding dist_def
   apply auto apply (rule_tac x=e in exI) apply auto
   using order_trans[OF real_abs_sub_infnorm infnorm_le_norm, of _ x] by (metis xt1(7))
 
 text{* Hence some handy theorems on distance, diameter etc. of/from a set.       *}
 
-lemma compact_attains_sup: 
+lemma compact_attains_sup:
   assumes "compact (vec1 ` s)"  "s \<noteq> {}"
   shows "\<exists>x \<in> s. \<forall>y \<in> s. y \<le> x"
 proof-
@@ -3851,7 +3851,7 @@
     apply(rule_tac x="rsup s" in bexI) by auto
 qed
 
-lemma compact_attains_inf: 
+lemma compact_attains_inf:
   assumes "compact (vec1 ` s)" "s \<noteq> {}"  shows "\<exists>x \<in> s. \<forall>y \<in> s. x \<le> y"
 proof-
   from assms(1) have a:"bounded (vec1 ` s)" "closed (vec1 ` s)" unfolding compact_eq_bounded_closed by auto
@@ -3862,25 +3862,25 @@
     { fix x assume "x \<in> s"
       hence *:"abs (x - rinf s) = x - rinf s" using as(1)[THEN bspec[where x=x]] by auto
       have "rinf s + e \<le> x" using as(4)[THEN bspec[where x=x]] using as(2) `x\<in>s` unfolding * by auto }
-    hence "isLb UNIV s (rinf s + e)" unfolding isLb_def and setge_def by auto 
+    hence "isLb UNIV s (rinf s + e)" unfolding isLb_def and setge_def by auto
     ultimately have False using isGlb_le_isLb[of UNIV s "rinf s" "rinf s + e"] using `e>0` by auto  }
   thus ?thesis using bounded_has_rinf(1)[OF a(1) assms(2)] using a(2)[unfolded closed_vec1, THEN spec[where x="rinf s"]]
     apply(rule_tac x="rinf s" in bexI) by auto
 qed
 
-lemma continuous_attains_sup: 
+lemma continuous_attains_sup:
  "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s (vec1 o f)
         ==> (\<exists>x \<in> s. \<forall>y \<in> s.  f y \<le> f x)"
   using compact_attains_sup[of "f ` s"]
   using compact_continuous_image[of s "vec1 \<circ> f"] unfolding image_compose by auto
 
-lemma continuous_attains_inf: 
- "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s (vec1 o f) 
+lemma continuous_attains_inf:
+ "compact s \<Longrightarrow> s \<noteq> {} \<Longrightarrow> continuous_on s (vec1 o f)
         ==> (\<exists>x \<in> s. \<forall>y \<in> s. f x \<le> f y)"
   using compact_attains_inf[of "f ` s"]
   using compact_continuous_image[of s "vec1 \<circ> f"] unfolding image_compose by auto
 
-lemma distance_attains_sup: 
+lemma distance_attains_sup:
   assumes "compact s" "s \<noteq> {}"
   shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a y \<le> dist a x"
 proof-
@@ -3896,7 +3896,7 @@
 
 text{* For *minimal* distance, we only need closure, not compactness.            *}
 
-lemma distance_attains_inf: 
+lemma distance_attains_inf:
   assumes "closed s"  "s \<noteq> {}"
   shows "\<exists>x \<in> s. \<forall>y \<in> s. dist a x \<le> dist a y"
 proof-
@@ -3911,16 +3911,16 @@
       hence "\<bar>norm (x' - a) - norm (x - a)\<bar> < e"
 	using real_abs_sub_norm[of "x' - a" "x - a"]  by auto  }
     hence "\<exists>d>0. \<forall>x'\<in>?B. norm (x' - x) < d \<longrightarrow> \<bar>dist x' a - dist x a\<bar> < e" using `e>0` unfolding dist_def by auto }
-  hence "continuous_on (cball a (dist b a) \<inter> s) (vec1 \<circ> dist a)" unfolding continuous_on_vec1_range 
+  hence "continuous_on (cball a (dist b a) \<inter> s) (vec1 \<circ> dist a)" unfolding continuous_on_vec1_range
     by (auto  simp add: dist_sym)
   moreover have "compact ?B" using compact_cball[of a "dist b a"] unfolding compact_eq_bounded_closed using bounded_Int and closed_Int and assms(1) by auto
   ultimately obtain x where "x\<in>cball a (dist b a) \<inter> s" "\<forall>y\<in>cball a (dist b a) \<inter> s. dist a x \<le> dist a y" using continuous_attains_inf[of ?B "dist a"] by fastsimp
-  thus ?thesis by fastsimp     
+  thus ?thesis by fastsimp
 qed
 
 subsection{* We can now extend limit compositions to consider the scalar multiplier.   *}
 
-lemma Lim_mul: 
+lemma Lim_mul:
   assumes "((vec1 o c) ---> vec1 d) net"  "(f ---> l) net"
   shows "((\<lambda>x. c(x) *s f x) ---> (d *s l)) net"
 proof-
@@ -3930,31 +3930,31 @@
   thus ?thesis using Lim_bilinear[OF assms, of "\<lambda>x y. (dest_vec1 x) *s y"] by auto
 qed
 
-lemma Lim_vmul: 
+lemma Lim_vmul:
  "((vec1 o c) ---> vec1 d) net ==> ((\<lambda>x. c(x) *s v) ---> d *s v) net"
   using Lim_mul[of c d net "\<lambda>x. v" v] using Lim_const[of v] by auto
 
-lemma continuous_vmul: 
+lemma continuous_vmul:
  "continuous net (vec1 o c) ==> continuous net (\<lambda>x. c(x) *s v)"
   unfolding continuous_def using Lim_vmul[of c] by auto
-  
-lemma continuous_mul: 
+
+lemma continuous_mul:
  "continuous net (vec1 o c) \<Longrightarrow> continuous net f
              ==> continuous net (\<lambda>x. c(x) *s f x) "
   unfolding continuous_def using Lim_mul[of c] by auto
 
-lemma continuous_on_vmul: 
+lemma continuous_on_vmul:
  "continuous_on s (vec1 o c) ==> continuous_on s (\<lambda>x. c(x) *s v)"
   unfolding continuous_on_eq_continuous_within using continuous_vmul[of _ c] by auto
 
-lemma continuous_on_mul: 
+lemma continuous_on_mul:
  "continuous_on s (vec1 o c) \<Longrightarrow> continuous_on s f
              ==> continuous_on s (\<lambda>x. c(x) *s f x)"
   unfolding continuous_on_eq_continuous_within using continuous_mul[of _ c] by auto
 
 text{* And so we have continuity of inverse.                                     *}
 
-lemma Lim_inv: 
+lemma Lim_inv:
   assumes "((vec1 o f) ---> vec1 l) (net::'a net)"  "l \<noteq> 0"
   shows "((vec1 o inverse o f) ---> vec1(inverse l)) net"
 proof(cases "trivial_limit net")
@@ -3974,7 +3974,7 @@
       have "\<bar>f x\<bar> * 2 \<ge> \<bar>l\<bar>" using * by (auto simp del: Arith_Tools.less_divide_eq_number_of1)
       hence "\<bar>f x\<bar> * 2 * \<bar>l\<bar>  \<ge> \<bar>l\<bar> * \<bar>l\<bar>" unfolding mult_le_cancel_right by auto
       hence "\<bar>f x * l\<bar> * 2  \<ge> \<bar>l\<bar>^2" unfolding real_mult_commute and power2_eq_square by auto
-      hence ***:"inverse \<bar>f x * l\<bar> \<le> inverse (l\<twosuperior> / 2)" using fxl0 
+      hence ***:"inverse \<bar>f x * l\<bar> \<le> inverse (l\<twosuperior> / 2)" using fxl0
 	using le_imp_inverse_le[of "l^2 / 2" "\<bar>f x * l\<bar>"]  by auto
 
       have "dist ((vec1 \<circ> inverse \<circ> f) x) (vec1 (inverse l)) < e" unfolding o_def unfolding dist_vec1
@@ -3983,19 +3983,19 @@
 	unfolding real_divide_def[THEN sym]
 	unfolding divide_divide_eq_left
 	unfolding nonzero_abs_divide[OF fxl0]
-	using mult_less_le_imp_less[OF **, of "inverse \<bar>f x * l\<bar>", of "inverse (l^2 / 2)"] using *** using fx0 `l\<noteq>0` 
+	using mult_less_le_imp_less[OF **, of "inverse \<bar>f x * l\<bar>", of "inverse (l^2 / 2)"] using *** using fx0 `l\<noteq>0`
 	unfolding inverse_eq_divide using `e>0` by auto   }
     hence "(\<exists>y. (\<exists>x. netord net x y) \<and> (\<forall>x. netord net x y \<longrightarrow> dist ((vec1 \<circ> inverse \<circ> f) x) (vec1 (inverse l)) < e))"
       using y1 by auto  }
   thus ?thesis unfolding tendsto_def eventually_def by auto
 qed
 
-lemma continuous_inv: 
+lemma continuous_inv:
  "continuous net (vec1 o f) \<Longrightarrow> f(netlimit net) \<noteq> 0
            ==> continuous net (vec1 o inverse o f)"
   unfolding continuous_def using Lim_inv by auto
-  
-lemma continuous_at_within_inv: 
+
+lemma continuous_at_within_inv:
   assumes "continuous (at a within s) (vec1 o f)" "f a \<noteq> 0"
   shows "continuous (at a within s) (vec1 o inverse o f)"
 proof(cases "trivial_limit (at a within s)")
@@ -4005,7 +4005,7 @@
   thus ?thesis using netlimit_within[OF cs] assms(2) continuous_inv[OF assms(1)] by auto
 qed
 
-lemma continuous_at_inv: 
+lemma continuous_at_inv:
  "continuous (at a) (vec1 o f) \<Longrightarrow> f a \<noteq> 0
          ==> continuous (at a) (vec1 o inverse o f) "
   using within_UNIV[THEN sym, of a] using continuous_at_within_inv[of a UNIV] by auto
@@ -4018,12 +4018,12 @@
 proof-
   obtain a b where ab:"\<forall>x\<in>s. norm x \<le> a" "\<forall>x\<in>t. norm x \<le> b" using assms[unfolded bounded_def] by auto
   { fix x y assume "x\<in>s" "y\<in>t"
-    hence "norm x \<le> a" "norm y \<le> b" using ab by auto 
+    hence "norm x \<le> a" "norm y \<le> b" using ab by auto
     hence "norm (pastecart x y) \<le> a + b" using norm_pastecart[of x y] by auto }
   thus ?thesis unfolding bounded_def by auto
 qed
 
-lemma closed_pastecart: 
+lemma closed_pastecart:
   assumes "closed s"  "closed t"
   shows "closed {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
 proof-
@@ -4036,7 +4036,7 @@
 	hence "dist (fstcart (x n)) (fstcart l) < e" "dist (sndcart (x n)) (sndcart l) < e"
 	  using N[THEN spec[where x=n]] dist_fstcart[of "x n" l] dist_sndcart[of "x n" l] by auto   }
       hence "\<exists>N. \<forall>n\<ge>N. dist (fstcart (x n)) (fstcart l) < e" "\<exists>N. \<forall>n\<ge>N. dist (sndcart (x n)) (sndcart l) < e" by auto  }
-    ultimately have "fstcart l \<in> s" "sndcart l \<in> t" 
+    ultimately have "fstcart l \<in> s" "sndcart l \<in> t"
       using assms(1)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. fstcart (x n)"], THEN spec[where x="fstcart l"]]
       using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. sndcart (x n)"], THEN spec[where x="sndcart l"]]
       unfolding Lim_sequentially by auto
@@ -4044,13 +4044,13 @@
   thus ?thesis unfolding closed_sequential_limits by auto
 qed
 
-lemma compact_pastecart: 
+lemma compact_pastecart:
  "compact s \<Longrightarrow> compact t ==> compact {pastecart x y | x y . x \<in> s \<and> y \<in> t}"
   unfolding compact_eq_bounded_closed using bounded_pastecart[of s t] closed_pastecart[of s t] by auto
 
 text{* Hence some useful properties follow quite easily.                         *}
 
-lemma compact_scaling: 
+lemma compact_scaling:
   assumes "compact s"  shows "compact ((\<lambda>x. c *s x) ` s)"
 proof-
   let ?f = "\<lambda>x. c *s x"
@@ -4059,27 +4059,27 @@
     using linear_continuous_at[OF *] assms by auto
 qed
 
-lemma compact_negations: 
+lemma compact_negations:
   assumes "compact s"  shows "compact ((\<lambda>x. -x) ` s)"
 proof-
   have "uminus ` s = (\<lambda>x. -1 *s x) ` s" apply auto unfolding image_iff pth_3 by auto
   thus ?thesis using compact_scaling[OF assms, of "-1"] by auto
 qed
 
-lemma compact_sums: 
+lemma compact_sums:
   assumes "compact s"  "compact t"  shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
 proof-
-  have *:"{x + y | x y. x \<in> s \<and> y \<in> t} =(\<lambda>z. fstcart z + sndcart z) ` {pastecart x y | x y.  x \<in> s \<and> y \<in> t}"    
+  have *:"{x + y | x y. x \<in> s \<and> y \<in> t} =(\<lambda>z. fstcart z + sndcart z) ` {pastecart x y | x y.  x \<in> s \<and> y \<in> t}"
     apply auto unfolding image_iff apply(rule_tac x="pastecart xa y" in bexI) unfolding fstcart_pastecart sndcart_pastecart by auto
   have "linear (\<lambda>z::real^('a, 'a) finite_sum. fstcart z + sndcart z)" unfolding linear_def
     unfolding fstcart_add sndcart_add apply auto
     unfolding vector_add_ldistrib fstcart_cmul[THEN sym] sndcart_cmul[THEN sym] by auto
   hence "continuous_on {pastecart x y |x y. x \<in> s \<and> y \<in> t} (\<lambda>z. fstcart z + sndcart z)"
-    using continuous_at_imp_continuous_on linear_continuous_at by auto 
+    using continuous_at_imp_continuous_on linear_continuous_at by auto
   thus ?thesis unfolding * using compact_continuous_image compact_pastecart[OF assms] by auto
 qed
 
-lemma compact_differences: 
+lemma compact_differences:
   assumes "compact s" "compact t"  shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
 proof-
   have "{x - y | x y::real^'a. x\<in>s \<and> y \<in> t} =  {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
@@ -4087,14 +4087,14 @@
   thus ?thesis using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
 qed
 
-lemma compact_translation: 
+lemma compact_translation:
   assumes "compact s"  shows "compact ((\<lambda>x. a + x) ` s)"
 proof-
   have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s" by auto
   thus ?thesis using compact_sums[OF assms compact_sing[of a]] by auto
 qed
 
-lemma compact_affinity: 
+lemma compact_affinity:
  assumes "compact s"  shows "compact ((\<lambda>x. a + c *s x) ` s)"
 proof-
   have "op + a ` op *s c ` s = (\<lambda>x. a + c *s x) ` s" by auto
@@ -4103,30 +4103,30 @@
 
 text{* Hence we get the following.                                               *}
 
-lemma compact_sup_maxdistance: 
+lemma compact_sup_maxdistance:
   assumes "compact s"  "s \<noteq> {}"
   shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. norm(u - v) \<le> norm(x - y)"
 proof-
   have "{x - y | x y . x\<in>s \<and> y\<in>s} \<noteq> {}" using `s \<noteq> {}` by auto
   then obtain x where x:"x\<in>{x - y |x y. x \<in> s \<and> y \<in> s}"  "\<forall>y\<in>{x - y |x y. x \<in> s \<and> y \<in> s}. norm y \<le> norm x"
     using compact_differences[OF assms(1) assms(1)]
-    using distance_attains_sup[unfolded dist_def, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by(auto simp add: norm_minus_cancel) 
+    using distance_attains_sup[unfolded dist_def, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by(auto simp add: norm_minus_cancel)
   from x(1) obtain a b where "a\<in>s" "b\<in>s" "x = a - b" by auto
-  thus ?thesis using x(2)[unfolded `x = a - b`] by blast    
+  thus ?thesis using x(2)[unfolded `x = a - b`] by blast
 qed
 
 text{* We can state this in terms of diameter of a set.                          *}
 
 definition "diameter s = (if s = {} then 0::real else rsup {norm(x - y) | x y. x \<in> s \<and> y \<in> s})"
 
-lemma diameter_bounded: 
-  assumes "bounded s" 
+lemma diameter_bounded:
+  assumes "bounded s"
   shows "\<forall>x\<in>s. \<forall>y\<in>s. norm(x - y) \<le> diameter s"
         "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)"
 proof-
   let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
   obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_def] by auto
-  { fix x y assume "x \<in> s" "y \<in> s" 
+  { fix x y assume "x \<in> s" "y \<in> s"
     hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps)  }
   note * = this
   { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
@@ -4134,7 +4134,7 @@
     have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s` isLubD1[OF lub] unfolding setle_def by auto  }
   moreover
   { fix d::real assume "d>0" "d < diameter s"
-    hence "s\<noteq>{}" unfolding diameter_def by auto 
+    hence "s\<noteq>{}" unfolding diameter_def by auto
     hence lub:"isLub UNIV ?D (rsup ?D)" using * rsup[of ?D] unfolding setle_def by auto
     have "\<exists>d' \<in> ?D. d' > d"
     proof(rule ccontr)
@@ -4148,11 +4148,11 @@
         "\<forall>d>0. d < diameter s --> (\<exists>x\<in>s. \<exists>y\<in>s. norm(x - y) > d)" by auto
 qed
 
-lemma diameter_bounded_bound: 
+lemma diameter_bounded_bound:
  "bounded s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s ==> norm(x - y) \<le> diameter s"
   using diameter_bounded by blast
 
-lemma diameter_compact_attained: 
+lemma diameter_compact_attained:
   assumes "compact s"  "s \<noteq> {}"
   shows "\<exists>x\<in>s. \<exists>y\<in>s. (norm(x - y) = diameter s)"
 proof-
@@ -4160,18 +4160,18 @@
   then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. norm (u - v) \<le> norm (x - y)" using compact_sup_maxdistance[OF assms] by auto
   hence "diameter s \<le> norm (x - y)" using rsup_le[of "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}" "norm (x - y)"]
     unfolding setle_def and diameter_def by auto
-  thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto 
+  thus ?thesis using diameter_bounded(1)[OF b, THEN bspec[where x=x], THEN bspec[where x=y], OF xys] and xys by auto
 qed
 
 text{* Related results with closure as the conclusion.                           *}
 
-lemma closed_scaling: 
+lemma closed_scaling:
   assumes "closed s" shows "closed ((\<lambda>x. c *s x) ` s)"
 proof(cases "s={}")
-  case True thus ?thesis by auto 
+  case True thus ?thesis by auto
 next
   case False
-  show ?thesis 
+  show ?thesis
   proof(cases "c=0")
     have *:"(\<lambda>x. 0) ` s = {0}" using `s\<noteq>{}` by auto
     case True thus ?thesis apply auto unfolding * using closed_sing by auto
@@ -4179,10 +4179,10 @@
     case False
     { fix x l assume as:"\<forall>n::nat. x n \<in> op *s c ` s"  "(x ---> l) sequentially"
       { fix n::nat have "(1 / c) *s x n \<in> s" using as(1)[THEN spec[where x=n]] using `c\<noteq>0` by (auto simp add: vector_smult_assoc) }
-      moreover 
+      moreover
       { fix e::real assume "e>0"
 	hence "0 < e *\<bar>c\<bar>"  using `c\<noteq>0` mult_pos_pos[of e "abs c"] by auto
-	then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>" using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto 
+	then obtain N where "\<forall>n\<ge>N. dist (x n) l < e * \<bar>c\<bar>" using as(2)[unfolded Lim_sequentially, THEN spec[where x="e * abs c"]] by auto
 	hence "\<exists>N. \<forall>n\<ge>N. dist ((1 / c) *s x n) ((1 / c) *s l) < e" unfolding dist_def unfolding vector_ssub_ldistrib[THEN sym] norm_mul
 	  using mult_imp_div_pos_less[of "abs c" _ e] `c\<noteq>0` by auto  }
       hence "((\<lambda>n. (1 / c) *s x n) ---> (1 / c) *s l) sequentially" unfolding Lim_sequentially by auto
@@ -4192,11 +4192,11 @@
   qed
 qed
 
-lemma closed_negations: 
+lemma closed_negations:
   assumes "closed s"  shows "closed ((\<lambda>x. -x) ` s)"
   using closed_scaling[OF assms, of "-1"] unfolding  pth_3 by auto
 
-lemma compact_closed_sums: 
+lemma compact_closed_sums:
   assumes "compact s"  "closed t"  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
 proof-
   let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
@@ -4208,23 +4208,23 @@
     have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"
       using Lim_sub[OF lim_subsequence[OF r as(2)] lr] and f(1) unfolding o_def by auto
     hence "l - l' \<in> t"
-      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda> n. snd (f (r n))"], THEN spec[where x="l - l'"]] 
+      using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda> n. snd (f (r n))"], THEN spec[where x="l - l'"]]
       using f(3) by auto
     hence "l \<in> ?S" using `l' \<in> s` apply auto apply(rule_tac x=l' in exI) apply(rule_tac x="l - l'" in exI) by auto
   }
   thus ?thesis unfolding closed_sequential_limits by auto
 qed
 
-lemma closed_compact_sums: 
-  assumes "closed s"  "compact t"  
+lemma closed_compact_sums:
+  assumes "closed s"  "compact t"
   shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
 proof-
-  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}" apply auto 
+  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}" apply auto
     apply(rule_tac x=y in exI) apply auto apply(rule_tac x=y in exI) by auto
   thus ?thesis using compact_closed_sums[OF assms(2,1)] by simp
 qed
 
-lemma compact_closed_differences: 
+lemma compact_closed_differences:
   assumes "compact s"  "closed t"
   shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
 proof-
@@ -4233,42 +4233,42 @@
   thus ?thesis using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto
 qed
 
-lemma closed_compact_differences: 
+lemma closed_compact_differences:
   assumes "closed s" "compact t"
   shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
 proof-
-  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}" 
+  have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
     apply auto apply(rule_tac x=xa in exI) apply auto apply(rule_tac x=xa in exI) by auto
  thus ?thesis using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp
 qed
 
-lemma closed_translation: 
+lemma closed_translation:
   assumes "closed s"  shows "closed ((\<lambda>x. a + x) ` s)"
 proof-
   have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
   thus ?thesis using compact_closed_sums[OF compact_sing[of a] assms] by auto
 qed
 
-lemma translation_UNIV: 
+lemma translation_UNIV:
  "range (\<lambda>x::real^'a. a + x) = UNIV"
   apply (auto simp add: image_iff) apply(rule_tac x="x - a" in exI) by auto
 
 lemma translation_diff: "(\<lambda>x::real^'a. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)" by auto
 
-lemma closure_translation: 
+lemma closure_translation:
  "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"
 proof-
   have *:"op + a ` (UNIV - s) = UNIV - op + a ` s"  apply auto unfolding image_iff apply(rule_tac x="x - a" in bexI) by auto
   show ?thesis unfolding closure_interior translation_diff translation_UNIV using interior_translation[of a "UNIV - s"] unfolding * by auto
 qed
 
-lemma frontier_translation: 
+lemma frontier_translation:
  "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"
   unfolding frontier_def translation_diff interior_translation closure_translation by auto
 
 subsection{* Separation between points and sets.                                       *}
 
-lemma separate_point_closed: 
+lemma separate_point_closed:
  "closed s \<Longrightarrow> a \<notin> s  ==> (\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x)"
 proof(cases "s = {}")
   case True
@@ -4280,28 +4280,28 @@
   with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s` by blast
 qed
 
-lemma separate_compact_closed: 
+lemma separate_compact_closed:
   assumes "compact s" and "closed t" and "s \<inter> t = {}"
   shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
 proof-
   have "0 \<notin> {x - y |x y. x \<in> s \<and> y \<in> t}" using assms(3) by auto
-  then obtain d where "d>0" and d:"\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. d \<le> dist 0 x" 
+  then obtain d where "d>0" and d:"\<forall>x\<in>{x - y |x y. x \<in> s \<and> y \<in> t}. d \<le> dist 0 x"
     using separate_point_closed[OF compact_closed_differences[OF assms(1,2)], of 0] by auto
   { fix x y assume "x\<in>s" "y\<in>t"
     hence "x - y \<in> {x - y |x y. x \<in> s \<and> y \<in> t}" by auto
-    hence "d \<le> dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_sym 
+    hence "d \<le> dist (x - y) 0" using d[THEN bspec[where x="x - y"]] using dist_sym
       by (auto  simp add: dist_sym)
     hence "d \<le> dist x y" unfolding dist_def by auto  }
   thus ?thesis using `d>0` by auto
 qed
 
-lemma separate_closed_compact: 
+lemma separate_closed_compact:
   assumes "closed s" and "compact t" and "s \<inter> t = {}"
   shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"
 proof-
   have *:"t \<inter> s = {}" using assms(3) by auto
   show ?thesis using separate_compact_closed[OF assms(2,1) *]
-    apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE) 
+    apply auto apply(rule_tac x=d in exI) apply auto apply (erule_tac x=y in ballE)
     by (auto simp add: dist_sym)
 qed
 
@@ -4331,7 +4331,7 @@
     let ?x = "(1/2) *s (a + b)"
     { fix i assume i:"i\<in>dimset a"
       hence "a$i < b$i" using as[THEN bspec[where x=i]] by auto
-      hence "a$i < ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i < b$i" 
+      hence "a$i < ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i < b$i"
 	unfolding vector_smult_component[OF i] and vector_add_component[OF i]
 	by (auto simp add: Arith_Tools.less_divide_eq_number_of1)  }
     hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
@@ -4346,7 +4346,7 @@
     let ?x = "(1/2) *s (a + b)"
     { fix i assume i:"i\<in>dimset a"
       hence "a$i \<le> b$i" using as[THEN bspec[where x=i]] by auto
-      hence "a$i \<le> ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i \<le> b$i" 
+      hence "a$i \<le> ((1/2) *s (a+b)) $ i" "((1/2) *s (a+b)) $ i \<le> b$i"
 	unfolding vector_smult_component[OF i] and vector_add_component[OF i]
 	by (auto simp add: Arith_Tools.less_divide_eq_number_of1)  }
     hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
@@ -4359,7 +4359,7 @@
   unfolding interval_eq_empty[of a b] by auto
 
 lemma subset_interval_imp: fixes a :: "real^'n" shows
- "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and 
+ "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c .. d} \<subseteq> {a .. b}" and
  "(\<forall>i \<in> dimset a. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> {c .. d} \<subseteq> {a<..<b}" and
  "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a .. b}" and
  "(\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> {c<..<d} \<subseteq> {a<..<b}"
@@ -4401,7 +4401,7 @@
 
 lemma subset_interval: fixes a :: "real^'n" shows
  "{c .. d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i \<le> d$i) --> (\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1) and
- "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i \<le> d$i) --> (\<forall>i \<in> dimset a. a$i < c$i \<and> d$i < b$i)" (is ?th2) and 
+ "{c .. d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i \<le> d$i) --> (\<forall>i \<in> dimset a. a$i < c$i \<and> d$i < b$i)" (is ?th2) and
  "{c<..<d} \<subseteq> {a .. b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i < d$i) --> (\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3) and
  "{c<..<d} \<subseteq> {a<..<b} \<longleftrightarrow> (\<forall>i \<in> dimset a. c$i < d$i) --> (\<forall>i \<in> dimset a. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
 proof-
@@ -4445,12 +4445,12 @@
     ultimately
     have "a$i \<le> c$i \<and> d$i \<le> b$i" by auto
   } note part1 = this
-  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+ 
+  thus ?th3 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+
   { assume as:"{c<..<d} \<subseteq> {a<..<b}" "\<forall>i \<in> dimset a. c$i < d$i"
     fix i assume i:"i \<in> dimset a"
     from as(1) have "{c<..<d} \<subseteq> {a..b}" using interval_open_subset_closed[of a b] by auto
     hence "a$i \<le> c$i \<and> d$i \<le> b$i" using part1 and as(2) and i by auto  } note * = this
-  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+ 
+  thus ?th4 unfolding subset_eq and Ball_def and mem_interval apply auto by(erule_tac x=xa in allE, simp)+
 qed
 
 lemma disjoint_interval: fixes a::"real^'n" shows
@@ -4462,13 +4462,13 @@
   let ?z = "(\<chi> i. ((max (a$i) (c$i)) + (min (b$i) (d$i))) / 2)::real^'n"
   show ?th1 ?th2 ?th3 ?th4
   unfolding expand_set_eq and Int_iff and empty_iff and mem_interval and ball_conj_distrib[THEN sym] and eq_False
-  by (auto simp add: Cart_lambda_beta' Arith_Tools.less_divide_eq_number_of1 intro!: bexI elim!: allE[where x="?z"]) 
+  by (auto simp add: Cart_lambda_beta' Arith_Tools.less_divide_eq_number_of1 intro!: bexI elim!: allE[where x="?z"])
 qed
 
 lemma inter_interval: fixes a :: "'a::linorder^'n" shows
  "{a .. b} \<inter> {c .. d} =  {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
   unfolding expand_set_eq and Int_iff and mem_interval
-  by (auto simp add: Cart_lambda_beta' Arith_Tools.less_divide_eq_number_of1 intro!: bexI) 
+  by (auto simp add: Cart_lambda_beta' Arith_Tools.less_divide_eq_number_of1 intro!: bexI)
 
 (* Moved interval_open_subset_closed a bit upwards *)
 
@@ -4483,7 +4483,7 @@
       hence "\<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i"
 	using x[unfolded mem_interval, THEN bspec[where x=i]]
 	using open_interval_lemma[of "a$i" "x$i" "b$i"] by auto  }
- 
+
     hence "\<forall>i\<in>dimset x. \<exists>d>0. \<forall>x'. abs (x' - (x$i)) < d \<longrightarrow> a$i < x' \<and> x' < b$i" by auto
     then obtain d where d:"\<forall>i\<in>dimset x. 0 < d i \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d i \<longrightarrow> a $ i < x' \<and> x' < b $ i)"
       using bchoice[of "dimset x" "\<lambda>i d. d>0 \<and> (\<forall>x'. \<bar>x' - x $ i\<bar> < d \<longrightarrow> a $ i < x' \<and> x' < b $ i)"] by auto
@@ -4494,9 +4494,9 @@
     moreover
     { fix x' assume as:"dist x' x < ?d"
       { fix i assume i:"i \<in> dimset x"
-	have "\<bar>x'$i - x $ i\<bar> < d i" 
+	have "\<bar>x'$i - x $ i\<bar> < d i"
 	  using norm_bound_component_lt[OF as[unfolded dist_def], THEN bspec[where x=i], OF i]
-	  unfolding vector_minus_component[OF i] and Min_gr_iff[OF **] using i by auto 
+	  unfolding vector_minus_component[OF i] and Min_gr_iff[OF **] using i by auto
 	hence "a $ i < x' $ i" "x' $ i < b $ i" using d[THEN bspec[where x=i], OF i] by auto  }
       hence "a < x' \<and> x' < b" unfolding vector_less_def by auto  }
     ultimately have "\<exists>e>0. \<forall>x'. dist x' x < e \<longrightarrow> x' \<in> {a<..<b}" by auto
@@ -4518,7 +4518,7 @@
       hence False unfolding mem_interval and dist_def
 	using component_le_norm[OF i, of "y-x", unfolded vector_minus_component[OF i]] and i and xb by(auto elim!: ballE[where x=i])
     } hence "x$i \<le> b$i" by(rule ccontr)auto
-    ultimately 
+    ultimately
     have "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" by auto }
   thus ?thesis unfolding closed_limpt islimpt_approachable mem_interval by auto
 qed
@@ -4541,8 +4541,8 @@
 	using e[THEN spec[where x="x - (e/2) *s basis i"]]
 	and   e[THEN spec[where x="x + (e/2) *s basis i"]]
 	unfolding mem_interval using i by auto
-      hence "a $ i < x $ i" and "x $ i < b $ i" 
-	unfolding vector_minus_component[OF i] and vector_add_component[OF i] 
+      hence "a $ i < x $ i" and "x $ i < b $ i"
+	unfolding vector_minus_component[OF i] and vector_add_component[OF i]
 	unfolding vector_smult_component[OF i] and basis_component[OF i] using `e>0` by auto   }
     hence "x \<in> {a<..<b}" unfolding mem_interval by auto  }
   thus "?L \<subseteq> ?R" unfolding interior_def and subset_eq by auto
@@ -4578,11 +4578,11 @@
 
 lemma open_interval_midpoint: fixes a :: "real^'n"
   assumes "{a<..<b} \<noteq> {}" shows "((1/2) *s (a + b)) \<in> {a<..<b}"
-proof- 
+proof-
   { fix i assume i:"i\<in>dimset a"
-    hence "a $ i < ((1 / 2) *s (a + b)) $ i \<and> ((1 / 2) *s (a + b)) $ i < b $ i" 
+    hence "a $ i < ((1 / 2) *s (a + b)) $ i \<and> ((1 / 2) *s (a + b)) $ i < b $ i"
       using assms[unfolded interval_ne_empty, THEN bspec[where x=i]]
-      unfolding vector_smult_component[OF i] and vector_add_component[OF i] 
+      unfolding vector_smult_component[OF i] and vector_add_component[OF i]
       by(auto simp add: Arith_Tools.less_divide_eq_number_of1)  }
   thus ?thesis unfolding mem_interval by auto
 qed
@@ -4609,17 +4609,17 @@
   thus ?thesis unfolding mem_interval by auto
 qed
 
-lemma closure_open_interval: fixes a :: "real^'n" 
-  assumes "{a<..<b} \<noteq> {}" 
+lemma closure_open_interval: fixes a :: "real^'n"
+  assumes "{a<..<b} \<noteq> {}"
   shows "closure {a<..<b} = {a .. b}"
 proof-
   have ab:"a < b" using assms[unfolded interval_ne_empty] unfolding vector_less_def by auto
   let ?c = "(1 / 2) *s (a + b)"
-  { fix x assume as:"x \<in> {a .. b}" 
+  { fix x assume as:"x \<in> {a .. b}"
     def f == "\<lambda>n::nat. x + (inverse (real n + 1)) *s (?c - x)"
     { fix n assume fn:"f n < b \<longrightarrow> a < f n \<longrightarrow> f n = x" and xc:"x \<noteq> ?c"
       have *:"0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" unfolding inverse_le_1_iff by auto
-      have "inverse (real n + 1) *s (1 / 2) *s (a + b) + (1 - inverse (real n + 1)) *s x = 
+      have "inverse (real n + 1) *s (1 / 2) *s (a + b) + (1 - inverse (real n + 1)) *s x =
 	x + inverse (real n + 1) *s ((1 / 2) *s (a + b) - x)" by (auto simp add: vector_ssub_ldistrib vector_add_ldistrib field_simps vector_sadd_rdistrib[THEN sym])
       hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
       hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib)  }
@@ -4633,7 +4633,7 @@
       hence "((vec1 \<circ> (\<lambda>n. inverse (real n + 1))) ---> vec1 0) sequentially"
 	unfolding Lim_sequentially by(auto simp add: dist_vec1)
       hence "(f ---> x) sequentially" unfolding f_def
-	using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *s ((1 / 2) *s (a + b) - x)" 0 sequentially x] 
+	using Lim_add[OF Lim_const, of "\<lambda>n::nat. (inverse (real n + 1)) *s ((1 / 2) *s (a + b) - x)" 0 sequentially x]
 	using Lim_vmul[of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *s (a + b) - x)"] by auto  }
     ultimately have "x \<in> closure {a<..<b}"
       using as and open_interval_midpoint[OF assms] unfolding closure_def unfolding islimpt_sequential by(cases "x=?c")auto  }
@@ -4645,7 +4645,7 @@
 proof-
   obtain b where "b>0" and b:"\<forall>x\<in>s. norm x \<le> b" using assms[unfolded bounded_pos] by auto
   def a \<equiv> "(\<chi> i. b+1)::real^'n"
-  { fix x assume "x\<in>s" 
+  { fix x assume "x\<in>s"
     fix i assume i:"i\<in>dimset a"
     have "(-a)$i < x$i" and "x$i < a$i" using b[THEN bspec[where x=x], OF `x\<in>s`] and component_le_norm[OF i, of x]
       unfolding vector_uminus_component[OF i] and a_def and Cart_lambda_beta'[OF i] by auto
@@ -4653,26 +4653,26 @@
   thus ?thesis by(auto intro: exI[where x=a] simp add: vector_less_def)
 qed
 
-lemma bounded_subset_open_interval: 
+lemma bounded_subset_open_interval:
   "bounded s ==> (\<exists>a b. s \<subseteq> {a<..<b})"
   by(metis bounded_subset_open_interval_symmetric)
 
-lemma bounded_subset_closed_interval_symmetric: 
+lemma bounded_subset_closed_interval_symmetric:
   assumes "bounded s" shows "\<exists>a. s \<subseteq> {-a .. a}"
 proof-
   obtain a where "s \<subseteq> {- a<..<a}" using bounded_subset_open_interval_symmetric[OF assms] by auto
   thus ?thesis using interval_open_subset_closed[of "-a" a] by auto
 qed
 
-lemma bounded_subset_closed_interval: 
+lemma bounded_subset_closed_interval:
   "bounded s ==> (\<exists>a b. s \<subseteq> {a .. b})"
   using bounded_subset_closed_interval_symmetric[of s] by auto
 
-lemma frontier_closed_interval: 
+lemma frontier_closed_interval:
  "frontier {a .. b} = {a .. b} - {a<..<b}"
   unfolding frontier_def unfolding interior_closed_interval and closure_closed[OF closed_interval] ..
 
-lemma frontier_open_interval: 
+lemma frontier_open_interval:
  "frontier {a<..<b} = (if {a<..<b} = {} then {} else {a .. b} - {a<..<b})"
 proof(cases "{a<..<b} = {}")
   case True thus ?thesis using frontier_empty by auto
@@ -4737,7 +4737,7 @@
   unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and dim1 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 dim1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto 
+  unfolding expand_set_eq apply simp unfolding vector_less_def and vector_less_eq_def and dim1 and dest_vec1_eq[THEN sym] and dest_vec1_def by auto
 
 (* Some stuff for half-infinite intervals too; FIXME: notation?  *)
 
@@ -4778,12 +4778,12 @@
   apply(erule_tac x=i in ballE)+ apply simp+
   done
 
-lemma is_interval_empty: 
+lemma is_interval_empty:
  "is_interval {}"
   unfolding is_interval_def
   by simp
 
-lemma is_interval_univ: 
+lemma is_interval_univ:
  "is_interval UNIV"
   unfolding is_interval_def
   by simp
@@ -4801,22 +4801,22 @@
     then obtain x y where x:"netord net x y" and y:"\<forall>x. netord net x y \<longrightarrow> dist l (f x) < e / norm a" apply(erule_tac x="e / norm a" in allE) apply auto using False using norm_ge_zero[of a] apply auto
       using divide_pos_pos[of e "norm a"] by auto
     { fix z assume "netord net z y" hence "dist l (f z) < e / norm a" using y by blast
-      hence "norm a * norm (l - f z) < e" unfolding dist_def and 
-	pos_less_divide_eq[OF False[unfolded vec_0 zero_less_norm_iff[of a, THEN sym]]] and real_mult_commute by auto 
+      hence "norm a * norm (l - f z) < e" unfolding dist_def and
+	pos_less_divide_eq[OF False[unfolded vec_0 zero_less_norm_iff[of a, THEN sym]]] and real_mult_commute by auto
       hence "\<bar>a \<bullet> l - a \<bullet> f z\<bar> < e" using order_le_less_trans[OF norm_cauchy_schwarz_abs[of a "l - f z"], of e] unfolding dot_rsub[symmetric] by auto  }
     hence "\<exists>y. (\<exists>x. netord net x y) \<and> (\<forall>x. netord net x y \<longrightarrow> \<bar>a \<bullet> l - a \<bullet> f x\<bar> < e)" using x by auto  }
-  thus ?thesis using assms unfolding Lim apply (auto simp add: dist_sym) 
-    unfolding dist_vec1 by auto  
-qed
-
-lemma continuous_at_vec1_dot: 
+  thus ?thesis using assms unfolding Lim apply (auto simp add: dist_sym)
+    unfolding dist_vec1 by auto
+qed
+
+lemma continuous_at_vec1_dot:
  "continuous (at x) (vec1 o (\<lambda>y. a \<bullet> y))"
 proof-
   have "((\<lambda>x. x) ---> x) (at x)" unfolding Lim_at by auto
   thus ?thesis unfolding continuous_at and o_def using Lim_vec1_dot[of "\<lambda>x. x" x "at x" a] by auto
 qed
 
-lemma continuous_on_vec1_dot: 
+lemma continuous_on_vec1_dot:
  "continuous_on s (vec1 o (\<lambda>y. a \<bullet> y)) "
   using continuous_at_imp_continuous_on[of s "vec1 o (\<lambda>y. a \<bullet> y)"]
   using continuous_at_vec1_dot
@@ -4828,12 +4828,12 @@
   have *:"{x \<in> UNIV. (vec1 \<circ> op \<bullet> a) x \<in> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b}} = {x. a \<bullet> x \<le> b}" by auto
   let ?T = "{x::real^1. (\<forall>i\<in>dimset x. x$i \<le> (vec1 b)$i)}"
   have "closed ?T" using closed_interval_left[of "vec1 b"] by simp
-  moreover have "vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> ?T" unfolding dim1 
-    unfolding image_def apply auto unfolding vec1_component[unfolded One_nat_def] by auto 
-  ultimately have "\<exists>T. closed T \<and> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> T" by auto 
-  hence "closedin euclidean {x \<in> UNIV. (vec1 \<circ> op \<bullet> a) x \<in> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b}}" 
+  moreover have "vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> ?T" unfolding dim1
+    unfolding image_def apply auto unfolding vec1_component[unfolded One_nat_def] by auto
+  ultimately have "\<exists>T. closed T \<and> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b} = range (vec1 \<circ> op \<bullet> a) \<inter> T" by auto
+  hence "closedin euclidean {x \<in> UNIV. (vec1 \<circ> op \<bullet> a) x \<in> vec1 ` {r. \<exists>x. a \<bullet> x = r \<and> r \<le> b}}"
     using continuous_on_vec1_dot[of UNIV a, unfolded continuous_on_closed subtopology_UNIV] unfolding closedin_closed
-    by (auto elim!: allE[where x="vec1 ` {r. (\<exists>x. a \<bullet> x = r \<and> r \<le> b)}"]) 
+    by (auto elim!: allE[where x="vec1 ` {r. (\<exists>x. a \<bullet> x = r \<and> r \<le> b)}"])
   thus ?thesis unfolding closed_closedin[THEN sym] and * by auto
 qed
 
@@ -4850,7 +4850,7 @@
   assumes "i \<in> {1 .. dimindex (UNIV::'n set)}" shows "closed {x::real^'n. x$i \<le> a}"
   using closed_halfspace_le[of "(basis i)::real^'n" a] unfolding dot_basis[OF assms] by auto
 
-lemma closed_halfspace_component_ge: 
+lemma closed_halfspace_component_ge:
   assumes "i \<in> {1 .. dimindex (UNIV::'n set)}" shows "closed {x::real^'n. x$i \<ge> a}"
   using closed_halfspace_ge[of a "(basis i)::real^'n"] unfolding dot_basis[OF assms] by auto
 
@@ -4868,11 +4868,11 @@
   thus ?thesis using closed_halfspace_le[unfolded closed_def, of a b] by auto
 qed
 
-lemma open_halfspace_component_lt: 
+lemma open_halfspace_component_lt:
   assumes "i \<in> {1 .. dimindex(UNIV::'n set)}" shows "open {x::real^'n. x$i < a}"
   using open_halfspace_lt[of "(basis i)::real^'n" a] unfolding dot_basis[OF assms] by auto
 
-lemma open_halfspace_component_gt: 
+lemma open_halfspace_component_gt:
   assumes "i \<in> {1 .. dimindex(UNIV::'n set)}" shows "open {x::real^'n. x$i  > a}"
   using open_halfspace_gt[of a "(basis i)::real^'n"] unfolding dot_basis[OF assms] by auto
 
@@ -4914,14 +4914,14 @@
 
 text{* Limits relative to a union.                                               *}
 
-lemma Lim_within_union: 
+lemma Lim_within_union:
  "(f ---> l) (at x within (s \<union> t)) \<longleftrightarrow>
   (f ---> l) (at x within s) \<and> (f ---> l) (at x within t)"
   unfolding Lim_within apply auto apply blast apply blast
     apply(erule_tac x=e in allE)+ apply auto
     apply(rule_tac x="min d da" in exI) by auto
 
-lemma continuous_on_union: 
+lemma continuous_on_union:
   assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
   shows "continuous_on (s \<union> t) f"
   using assms unfolding continuous_on unfolding Lim_within_union
@@ -4938,7 +4938,7 @@
   moreover
   have "\<forall>x\<in>t. g x = (if P x then f x else g x)" using assms(5) by auto
   hence "continuous_on t ?h" using continuous_on_eq[of t g ?h] using assms(4) by auto
-  ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto 
+  ultimately show ?thesis using continuous_on_union[OF assms(1,2), of ?h] by auto
 qed
 
 
@@ -4971,7 +4971,7 @@
   { fix m::nat
     have "\<And> n. n\<ge>m \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)"
       apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq)  }
-  hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" by auto 
+  hence "\<forall>m n. m \<le> n \<longrightarrow> dest_vec1 (s m) \<le> dest_vec1 (s n)" by auto
   then obtain l where "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<bar>dest_vec1 (s n) - l\<bar> < e" using convergent_bounded_monotone[OF a] by auto
   thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI)
     unfolding dist_def unfolding abs_dest_vec1 and dest_vec1_sub by auto
@@ -4994,20 +4994,20 @@
   apply(rule_tac x = "(\<lambda>x::real^'b.x)" in exI)
   by blast
 
-lemma homeomorphic_sym: 
+lemma homeomorphic_sym:
  "s homeomorphic t \<longleftrightarrow> t homeomorphic s"
 unfolding homeomorphic_def
 unfolding homeomorphism_def
 by blast
 
-lemma homeomorphic_trans: 
+lemma homeomorphic_trans:
   assumes "s homeomorphic t" "t homeomorphic u" shows "s homeomorphic u"
 proof-
   obtain f1 g1 where fg1:"\<forall>x\<in>s. g1 (f1 x) = x"  "f1 ` s = t" "continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1"
     using assms(1) unfolding homeomorphic_def homeomorphism_def by auto
   obtain f2 g2 where fg2:"\<forall>x\<in>t. g2 (f2 x) = x"  "f2 ` t = u" "continuous_on t f2" "\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2"
     using assms(2) unfolding homeomorphic_def homeomorphism_def by auto
-  
+
   { fix x assume "x\<in>s" hence "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x" using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) by auto }
   moreover have "(f2 \<circ> f1) ` s = u" using fg1(2) fg2(2) by auto
   moreover have "continuous_on s (f2 \<circ> f1)" using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto
@@ -5017,29 +5017,29 @@
   ultimately show ?thesis unfolding homeomorphic_def homeomorphism_def apply(rule_tac x="f2 \<circ> f1" in exI) apply(rule_tac x="g1 \<circ> g2" in exI) by auto
 qed
 
-lemma homeomorphic_minimal: 
+lemma homeomorphic_minimal:
  "s homeomorphic t \<longleftrightarrow>
     (\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>
            (\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>
            continuous_on s f \<and> continuous_on t g)"
 unfolding homeomorphic_def homeomorphism_def
 apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI)
-apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto 
+apply auto apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply auto
 unfolding image_iff
-apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE) 
+apply(erule_tac x="g x" in ballE) apply(erule_tac x="x" in ballE)
 apply auto apply(rule_tac x="g x" in bexI) apply auto
-apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE) 
+apply(erule_tac x="f x" in ballE) apply(erule_tac x="x" in ballE)
 apply auto apply(rule_tac x="f x" in bexI) by auto
 
 subsection{* Relatively weak hypotheses if a set is compact.                           *}
 
 definition "inv_on f s = (\<lambda>x. SOME y. y\<in>s \<and> f y = x)"
 
-lemma assumes "inj_on f s" "x\<in>s" 
-  shows "inv_on f s (f x) = x" 
+lemma assumes "inj_on f s" "x\<in>s"
+  shows "inv_on f s (f x) = x"
  using assms unfolding inj_on_def inv_on_def by auto
 
-lemma homeomorphism_compact: 
+lemma homeomorphism_compact:
   assumes "compact s" "continuous_on s f"  "f ` s = t"  "inj_on f s"
   shows "\<exists>g. homeomorphism s t f g"
 proof-
@@ -5053,33 +5053,33 @@
   moreover
   { fix x
     have "x\<in>s \<Longrightarrow> x \<in> g ` t" using g[THEN bspec[where x=x]] unfolding image_iff using assms(3) by(auto intro!: bexI[where x="f x"])
-    moreover 
+    moreover
     { assume "x\<in>g ` t"
       then obtain y where y:"y\<in>t" "g y = x" by auto
       then obtain x' where x':"x'\<in>s" "f x' = y" using assms(3) by auto
       hence "x \<in> s" unfolding g_def using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] unfolding y(2)[THEN sym] and g_def by auto }
     ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" by auto  }
   hence "g ` t = s" by auto
-  ultimately 
-  show ?thesis unfolding homeomorphism_def homeomorphic_def 
-    apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto 
-qed
-
-lemma homeomorphic_compact: 
+  ultimately
+  show ?thesis unfolding homeomorphism_def homeomorphic_def
+    apply(rule_tac x=g in exI) using g and assms(3) and continuous_on_inverse[OF assms(2,1), of g, unfolded assms(3)] and assms(2) by auto
+qed
+
+lemma homeomorphic_compact:
  "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s
           \<Longrightarrow> s homeomorphic t"
   unfolding homeomorphic_def by(metis homeomorphism_compact)
 
 text{* Preservation of topological properties.                                   *}
 
-lemma homeomorphic_compactness: 
+lemma homeomorphic_compactness:
  "s homeomorphic t ==> (compact s \<longleftrightarrow> compact t)"
 unfolding homeomorphic_def homeomorphism_def
 by (metis compact_continuous_image)
 
 text{* Results on translation, scaling etc.                                      *}
 
-lemma homeomorphic_scaling: 
+lemma homeomorphic_scaling:
   assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. c *s x) ` s)"
   unfolding homeomorphic_minimal
   apply(rule_tac x="\<lambda>x. c *s x" in exI)
@@ -5087,14 +5087,14 @@
   apply auto unfolding vector_smult_assoc using assms apply auto
   using continuous_on_cmul[OF continuous_on_id] by auto
 
-lemma homeomorphic_translation: 
+lemma homeomorphic_translation:
  "s homeomorphic ((\<lambda>x. a + x) ` s)"
   unfolding homeomorphic_minimal
   apply(rule_tac x="\<lambda>x. a + x" in exI)
   apply(rule_tac x="\<lambda>x. -a + x" in exI)
   using continuous_on_add[OF continuous_on_const continuous_on_id] by auto
 
-lemma homeomorphic_affinity: 
+lemma homeomorphic_affinity:
   assumes "c \<noteq> 0"  shows "s homeomorphic ((\<lambda>x. a + c *s x) ` s)"
 proof-
   have *:"op + a ` op *s c ` s = (\<lambda>x. a + c *s x) ` s" by auto
@@ -5115,7 +5115,7 @@
     apply(rule_tac x="\<lambda>x. a + (d/e) *s (x - b)" in exI)
     apply (auto simp add: dist_sym) unfolding dist_def and vector_smult_assoc using assms apply auto
     unfolding norm_minus_cancel and norm_mul
-    using continuous_on_add[OF continuous_on_const continuous_on_cmul[OF continuous_on_sub[OF continuous_on_id continuous_on_const]]] 
+    using continuous_on_add[OF continuous_on_const continuous_on_cmul[OF continuous_on_sub[OF continuous_on_id continuous_on_const]]]
     apply (auto simp add: dist_sym)
     using pos_less_divide_eq[OF *(1), THEN sym] unfolding real_mult_commute[of _ "\<bar>e / d\<bar>"]
     using pos_less_divide_eq[OF *(2), THEN sym] unfolding real_mult_commute[of _ "\<bar>d / e\<bar>"]
@@ -5127,7 +5127,7 @@
     apply(rule_tac x="\<lambda>x. a + (d/e) *s (x - b)" in exI)
     apply (auto simp add: dist_sym) unfolding dist_def and vector_smult_assoc using assms apply auto
     unfolding norm_minus_cancel and norm_mul
-    using continuous_on_add[OF continuous_on_const continuous_on_cmul[OF continuous_on_sub[OF continuous_on_id continuous_on_const]]] 
+    using continuous_on_add[OF continuous_on_const continuous_on_cmul[OF continuous_on_sub[OF continuous_on_id continuous_on_const]]]
     apply auto
     using pos_le_divide_eq[OF *(1), THEN sym] unfolding real_mult_commute[of _ "\<bar>e / d\<bar>"]
     using pos_le_divide_eq[OF *(2), THEN sym] unfolding real_mult_commute[of _ "\<bar>d / e\<bar>"]
@@ -5136,25 +5136,25 @@
 
 text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
 
-lemma cauchy_isometric: 
+lemma cauchy_isometric:
   assumes e:"0 < e" and s:"subspace s" and f:"linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and xs:"\<forall>n::nat. x n \<in> s" and cf:"cauchy(f o x)"
   shows "cauchy x"
 proof-
   { fix d::real assume "d>0"
-    then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d" 
+    then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"
       using cf[unfolded cauchy o_def dist_def, THEN spec[where x="e*d"]] and e and mult_pos_pos[of e d] by auto
     { fix n assume "n\<ge>N"
       hence "norm (f (x n - x N)) < e * d" using N[THEN spec[where x=n]] unfolding linear_sub[OF f, THEN sym] by auto
       moreover have "e * norm (x n - x N) \<le> norm (f (x n - x N))"
 	using subspace_sub[OF s, of "x n" "x N"] using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]
 	using normf[THEN bspec[where x="x n - x N"]] by auto
-      ultimately have "norm (x n - x N) < d" using `e>0` 
+      ultimately have "norm (x n - x N) < d" using `e>0`
 	using mult_left_less_imp_less[of e "norm (x n - x N)" d] by auto   }
     hence "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto }
   thus ?thesis unfolding cauchy and dist_def by auto
 qed
 
-lemma complete_isometric_image: 
+lemma complete_isometric_image:
   assumes "0 < e" and s:"subspace s" and f:"linear f" and normf:"\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" and cs:"complete s"
   shows "complete(f ` s)"
 proof-
@@ -5164,10 +5164,10 @@
     hence x:"\<forall>n. x n \<in> s"  "\<forall>n. g n = f (x n)" by auto
     hence "f \<circ> x = g" unfolding expand_fun_eq by auto
     then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"
-      using cs[unfolded complete_def, THEN spec[where x="x"]] 
+      using cs[unfolded complete_def, THEN spec[where x="x"]]
       using cauchy_isometric[OF `0<e` s f normf] and cfg and x(1) by auto
     hence "\<exists>l\<in>f ` s. (g ---> l) sequentially"
-      using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] 
+      using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l]
       unfolding `f \<circ> x = g` by auto  }
   thus ?thesis unfolding complete_def by auto
 qed
@@ -5201,7 +5201,7 @@
   moreover have "?S \<noteq> {}" using a by auto
   ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto
   then obtain b where "b\<in>s" and ba:"norm b = norm a" and b:"\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)" unfolding *[THEN sym] unfolding image_iff by auto
-  
+
   let ?e = "norm (f b) / norm b"
   have "norm b > 0" using ba and a and norm_ge_zero by auto
   moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF `b\<in>s`] using `norm b >0` unfolding zero_less_norm_iff by auto
@@ -5213,40 +5213,40 @@
       case True thus "norm (f b) / norm b * norm x \<le> norm (f x)" by auto
     next
       case False
-      hence *:"0 < norm a / norm x" using `a\<noteq>0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos) 
+      hence *:"0 < norm a / norm x" using `a\<noteq>0` unfolding zero_less_norm_iff[THEN sym] by(simp only: divide_pos_pos)
       have "\<forall>c. \<forall>x\<in>s. c *s x \<in> s" using s[unfolded subspace_def] by auto
       hence "(norm a / norm x) *s x \<in> {x \<in> s. norm x = norm a}" using `x\<in>s` and `x\<noteq>0` by auto
-      thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *s x"]] 
+      thus "norm (f b) / norm b * norm x \<le> norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *s x"]]
 	unfolding linear_cmul[OF f(1)] and norm_mul and ba using `x\<noteq>0` `a\<noteq>0`
 	by (auto simp add: real_mult_commute pos_le_divide_eq pos_divide_le_eq)
     qed }
-  ultimately 
+  ultimately
   show ?thesis by auto
 qed
 
-lemma closed_injective_image_subspace: 
+lemma closed_injective_image_subspace:
   assumes "subspace s" "linear f" "\<forall>x\<in>s. f x = 0 --> x = 0" "closed s"
   shows "closed(f ` s)"
 proof-
   obtain e where "e>0" and e:"\<forall>x\<in>s. e * norm x \<le> norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto
   show ?thesis using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)
-    unfolding complete_eq_closed[THEN sym] by auto 
+    unfolding complete_eq_closed[THEN sym] by auto
 qed
 
 subsection{* Some properties of a canonical subspace.                                  *}
 
-lemma subspace_substandard: 
+lemma subspace_substandard:
  "subspace {x::real^'n. (\<forall>i \<in> dimset x. d < i \<longrightarrow> x$i = 0)}"
   unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE)
 
-lemma closed_substandard: 
+lemma closed_substandard:
  "closed {x::real^'n. \<forall>i \<in> dimset x. d < i --> x$i = 0}" (is "closed ?A")
 proof-
   let ?D = "{Suc d..dimindex(UNIV::('n set))}"
   let ?Bs = "{{x::real^'n. basis i \<bullet> x = 0}| i. i \<in> ?D}"
   { fix x
     { assume "x\<in>?A"
-      hence x:"\<forall>i\<in>?D. d < i \<longrightarrow> x $ i = 0" by auto 
+      hence x:"\<forall>i\<in>?D. d < i \<longrightarrow> x $ i = 0" by auto
       hence "x\<in> \<Inter> ?Bs" by(auto simp add: dot_basis x) }
     moreover
     { assume x:"x\<in>\<Inter>?Bs"
@@ -5266,7 +5266,7 @@
 proof-
   let ?D = "{1..dimindex (UNIV::'n set)}"
   let ?B = "(basis::nat\<Rightarrow>real^'n) ` {1..d}"
-    
+
     let ?bas = "basis::nat \<Rightarrow> real^'n"
 
   have "?B \<subseteq> ?A" by (auto simp add: basis_component)
@@ -5279,7 +5279,7 @@
       thus ?case using subspace_0[OF subspace_span[of "{}"]] by auto
     next
       case (Suc n)
-      hence *:"\<forall>i\<in>?D. Suc n < i \<longrightarrow> x $ i = 0" by auto 
+      hence *:"\<forall>i\<in>?D. Suc n < i \<longrightarrow> x $ i = 0" by auto
       have **:"{1..n} \<subseteq> {1..Suc n}" by auto
       def y \<equiv> "x - x$(Suc n) *s basis (Suc n)"
       have y:"x = y + (x$Suc n) *s basis (Suc n)" unfolding y_def by auto
@@ -5291,9 +5291,9 @@
 	using span_mono[of "?bas ` {1..n}" "?bas ` {1..Suc n}"]
 	using image_mono[OF **, of basis] by auto
       moreover
-      have "basis (Suc n) \<in> span (?bas ` {1..Suc n})" by(rule span_superset, auto)  
+      have "basis (Suc n) \<in> span (?bas ` {1..Suc n})" by(rule span_superset, auto)
       hence "x$(Suc n) *s basis (Suc n) \<in> span (?bas ` {1..Suc n})" using span_mul by auto
-      ultimately 
+      ultimately
       have "y + x$(Suc n) *s basis (Suc n) \<in> span (?bas ` {1..Suc n})"
 	using span_add by auto
       thus ?case using y by auto
@@ -5328,66 +5328,66 @@
     by(erule_tac x=0 in ballE) auto
   moreover have "closed ?t" using closed_substandard by auto
   moreover have "subspace ?t" using subspace_substandard by auto
-  ultimately show ?thesis using closed_injective_image_subspace[of ?t f] 
+  ultimately show ?thesis using closed_injective_image_subspace[of ?t f]
     unfolding f(2) using f(1) by auto
 qed
 
-lemma complete_subspace: 
+lemma complete_subspace:
   "subspace s ==> complete s"
   using complete_eq_closed closed_subspace
   by auto
 
-lemma dim_closure: 
+lemma dim_closure:
  "dim(closure s) = dim s" (is "?dc = ?d")
 proof-
   have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]
-    using closed_subspace[OF subspace_span, of s] 
+    using closed_subspace[OF subspace_span, of s]
     using dim_subset[of "closure s" "span s"] unfolding dim_span by auto
   thus ?thesis using dim_subset[OF closure_subset, of s] by auto
 qed
 
 text{* Affine transformations of intervals.                                      *}
 
-lemma affinity_inverses: 
-  assumes m0: "m \<noteq> (0::'a::field)" 
+lemma affinity_inverses:
+  assumes m0: "m \<noteq> (0::'a::field)"
   shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
   "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
   using m0
 apply (auto simp add: expand_fun_eq vector_add_ldistrib vector_smult_assoc)
 by (simp add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
 
-lemma real_affinity_le: 
+lemma real_affinity_le:
  "0 < (m::'a::ordered_field) ==> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"
   by (simp add: field_simps inverse_eq_divide)
 
-lemma real_le_affinity: 
+lemma real_le_affinity:
  "0 < (m::'a::ordered_field) ==> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)"
   by (simp add: field_simps inverse_eq_divide)
 
-lemma real_affinity_lt: 
+lemma real_affinity_lt:
  "0 < (m::'a::ordered_field) ==> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))"
   by (simp add: field_simps inverse_eq_divide)
 
-lemma real_lt_affinity: 
+lemma real_lt_affinity:
  "0 < (m::'a::ordered_field) ==> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)"
   by (simp add: field_simps inverse_eq_divide)
 
-lemma real_affinity_eq: 
+lemma real_affinity_eq:
  "(m::'a::ordered_field) \<noteq> 0 ==> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))"
   by (simp add: field_simps inverse_eq_divide)
 
-lemma real_eq_affinity: 
+lemma real_eq_affinity:
  "(m::'a::ordered_field) \<noteq> 0 ==> (y = m * x + c  \<longleftrightarrow> inverse(m) * y + -(c / m) = x)"
   by (simp add: field_simps inverse_eq_divide)
 
-lemma vector_affinity_eq: 
-  assumes m0: "(m::'a::field) \<noteq> 0" 
+lemma vector_affinity_eq:
+  assumes m0: "(m::'a::field) \<noteq> 0"
   shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
 proof
   assume h: "m *s x + c = y"
   hence "m *s x = y - c" by (simp add: ring_simps)
   hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
-  then show "x = inverse m *s y + - (inverse m *s c)" 
+  then show "x = inverse m *s y + - (inverse m *s c)"
     using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
 next
   assume h: "x = inverse m *s y + - (inverse m *s c)"
@@ -5395,7 +5395,7 @@
     using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
 qed
 
-lemma vector_eq_affinity: 
+lemma vector_eq_affinity:
  "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
   using vector_affinity_eq[where m=m and x=x and y=y and c=c]
   by metis
@@ -5415,18 +5415,18 @@
   case False
   { fix y assume "a \<le> y" "y \<le> b" "m > 0"
     hence "m *s a + c \<le> m *s y + c"  "m *s y + c \<le> m *s b + c"
-      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component) 
+      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component)
   } moreover
   { fix y assume "a \<le> y" "y \<le> b" "m < 0"
-    hence "m *s b + c \<le> m *s y + c"  "m *s y + c \<le> m *s a + c" 
-      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE) 
+    hence "m *s b + c \<le> m *s y + c"  "m *s y + c \<le> m *s a + c"
+      unfolding vector_less_eq_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
   } moreover
   { fix y assume "m > 0"  "m *s a + c \<le> y"  "y \<le> m *s b + c"
     hence "y \<in> (\<lambda>x. m *s x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval vector_less_eq_def
       apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
 	intro!: exI[where x="(1 / m) *s (y - c)"])
-      by(auto elim!: ballE simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute) 
+      by(auto elim!: ballE simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute)
   } moreover
   { fix y assume "m *s b + c \<le> y" "y \<le> m *s a + c" "m < 0"
     hence "y \<in> (\<lambda>x. m *s x + c) ` {a..b}"
@@ -5440,12 +5440,12 @@
 
 subsection{* Banach fixed point theorem (not really topological...)                    *}
 
-lemma banach_fix: 
+lemma banach_fix:
   assumes s:"complete s" "s \<noteq> {}" and c:"0 \<le> c" "c < 1" and f:"(f ` s) \<subseteq> s" and
           lipschitz:"\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"
   shows "\<exists>! x\<in>s. (f x = x)"
 proof-
-  have "1 - c > 0" using c by auto 
+  have "1 - c > 0" using c by auto
 
   from s(2) obtain z0 where "z0 \<in> s" by auto
   def z \<equiv> "\<lambda> n::nat. fun_pow n f z0"
@@ -5486,7 +5486,7 @@
       also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
 	unfolding power_add by (auto simp add: ring_simps)
       also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
-	using c by (auto simp add: ring_simps dist_pos_le) 
+	using c by (auto simp add: ring_simps dist_pos_le)
       finally show ?case by auto
     qed
   } note cf_z2 = this
@@ -5495,11 +5495,11 @@
     proof(cases "d = 0")
       case True
       hence "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (auto simp add: pos_prod_le[OF `1 - c > 0`] dist_le_0)
-      thus ?thesis using `e>0` by auto  
+      thus ?thesis using `e>0` by auto
     next
       case False hence "d>0" unfolding d_def using dist_pos_le[of "z 0" "z 1"]
-	by (metis False d_def real_less_def) 
-      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0` 
+	by (metis False d_def real_less_def)
+      hence "0 < e * (1 - c) / d" using `e>0` and `1-c>0`
 	using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"] by auto
       then obtain N where N:"c ^ N < e * (1 - c) / d" using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto
       { fix m n::nat assume "m>n" and as:"m\<ge>N" "n\<ge>N"
@@ -5510,8 +5510,8 @@
 	  using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]
 	  using `0 < 1 - c` by auto
 
-	have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" 
-	  using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`] 
+	have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"
+	  using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`]
 	  by (auto simp add: real_mult_commute dist_sym)
 	also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"
 	  using mult_right_mono[OF * order_less_imp_le[OF **]]
@@ -5529,23 +5529,23 @@
 	next
 	  case False thus ?thesis using as and *[of n m] *[of m n] unfolding nat_neq_iff by (auto simp add: dist_sym)
 	qed }
-      thus ?thesis by auto 
+      thus ?thesis by auto
     qed
   }
   hence "cauchy z" unfolding cauchy_def by auto
   then obtain x where "x\<in>s" and x:"(z ---> x) sequentially" using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto
-  
+
   def e \<equiv> "dist (f x) x"
   have "e = 0" proof(rule ccontr)
     assume "e \<noteq> 0" hence "e>0" unfolding e_def using dist_pos_le[of "f x" x]
-      by (metis dist_eq_0 dist_nz dist_sym e_def) 
+      by (metis dist_eq_0 dist_nz dist_sym e_def)
     then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2"
       using x[unfolded Lim_sequentially, THEN spec[where x="e/2"]] by auto
     hence N':"dist (z N) x < e / 2" by auto
 
-    have *:"c * dist (z N) x \<le> dist (z N) x" unfolding mult_le_cancel_right2 
+    have *:"c * dist (z N) x \<le> dist (z N) x" unfolding mult_le_cancel_right2
       using dist_pos_le[of "z N" x] and c
-      by (metis dist_eq_0 dist_nz dist_sym order_less_asym real_less_def) 
+      by (metis dist_eq_0 dist_nz dist_sym order_less_asym real_less_def)
     have "dist (f (z N)) (f x) \<le> c * dist (z N) x" using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]
       using z_in_s[of N] `x\<in>s` using c by auto
     also have "\<dots> < e / 2" using N' and c using * by auto
@@ -5559,15 +5559,15 @@
     hence "dist x y \<le> c * dist x y" using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]
       using `x\<in>s` and `f x = x` by auto
     hence "dist x y = 0" unfolding mult_le_cancel_right1
-      using c and dist_pos_le[of x y] by auto 
+      using c and dist_pos_le[of x y] by auto
     hence "y = x" unfolding dist_eq_0 by auto
   }
-  ultimately show ?thesis unfolding Bex1_def using `x\<in>s` by blast+ 
+  ultimately show ?thesis unfolding Bex1_def using `x\<in>s` by blast+
 qed
 
 subsection{* Edelstein fixed point theorem.                                            *}
 
-lemma edelstein_fix: 
+lemma edelstein_fix:
   assumes s:"compact s" "s \<noteq> {}" and gs:"(g ` s) \<subseteq> s"
       and dist:"\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y"
   shows "\<exists>! x::real^'a\<in>s. g x = x"
@@ -5575,10 +5575,10 @@
   obtain x where "x\<in>s" using s(2) by auto
   case False hence g:"\<forall>x\<in>s. g x = x" by auto
   { fix y assume "y\<in>s"
-    hence "x = y" using `x\<in>s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]] 
+    hence "x = y" using `x\<in>s` and dist[THEN bspec[where x=x], THEN bspec[where x=y]]
       unfolding g[THEN bspec[where x=x], OF `x\<in>s`]
       unfolding g[THEN bspec[where x=y], OF `y\<in>s`] by auto  }
-  thus ?thesis unfolding Bex1_def using `x\<in>s` and g by blast+ 
+  thus ?thesis unfolding Bex1_def using `x\<in>s` and g by blast+
 next
   case True
   then obtain x where [simp]:"x\<in>s" and "g x \<noteq> x" by auto
@@ -5606,7 +5606,7 @@
       case (Suc n)
       thus ?case proof(cases "m\<le>n")
 	case True thus ?thesis using Suc(1)
-	  using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto 
+	  using dist'[OF fs fs, OF `w\<in>s` `z\<in>s`, of n n] by auto
       next
 	case False hence mn:"m = Suc n" using Suc(2) by simp
 	show ?thesis unfolding mn  by auto
@@ -5634,7 +5634,7 @@
     { fix x y ::"real^'a"
       have "dist (-x) (-y) = dist x y" unfolding dist_def
 	using norm_minus_cancel[of "x - y"] by (auto simp add: uminus_add_conv_diff) } note ** = this
-    
+
     { assume as:"dist a b > dist (f n x) (f n y)"
       then obtain Na Nb where "\<forall>m\<ge>Na. dist (f (r m) x) a < (dist a b - dist (f n x) (f n y)) / 2"
 	and "\<forall>m\<ge>Nb. dist (f (r m) y) b < (dist a b - dist (f n x) (f n y)) / 2"
@@ -5662,7 +5662,7 @@
       using lima limb unfolding Lim_sequentially
       apply (auto elim!: allE[where x="e/2"]) apply(rule_tac x="r (max N Na)" in exI) unfolding h_def by fastsimp
     then obtain n where n:"dist (f n x) a < e/2 \<and> dist (f n y) b < e/2" by auto
-    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a" 
+    have "dist (f (Suc n) x) (g a) \<le> dist (f n x) a"
       using dist[THEN bspec[where x="f n x"], THEN bspec[where x="a"]] and fs by auto
     moreover have "dist (f (Suc n) y) (g b) \<le> dist (f n y) b"
       using dist[THEN bspec[where x="f n y"], THEN bspec[where x="b"]] and fs by auto
@@ -5679,7 +5679,7 @@
   hence "((sndcart \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
     apply (rule allE[where x="\<lambda>n. (fstcart \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
     using lima unfolding h_def o_def using fs[OF `x\<in>s`] by (auto simp add: y_def)
-  hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"] 
+  hence "g a = a" using Lim_unique[OF trivial_limit_sequentially limb, of "g a"]
     unfolding `a=b` and o_assoc by auto
   moreover
   { fix x assume "x\<in>s" "g x = x" "x\<noteq>a"
@@ -5688,4 +5688,4 @@
   ultimately show "\<exists>!x\<in>s. g x = x" unfolding Bex1_def using `a\<in>s` by blast
 qed
 
-end 
\ No newline at end of file
+end
--- a/src/HOL/Library/Univ_Poly.thy	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Library/Univ_Poly.thy	Thu Mar 12 23:01:25 2009 +0100
@@ -19,7 +19,7 @@
 
 text{*addition*}
 
-primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65) 
+primrec (in semiring_0) padd :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixl "+++" 65)
 where
   padd_Nil:  "[] +++ l2 = l2"
 | padd_Cons: "(h#t) +++ l2 = (if l2 = [] then h#t
@@ -79,7 +79,7 @@
                       ~ (([-a, 1] %^ (Suc n)) divides p))"
 
      --{*degree of a polynomial*}
-definition (in semiring_0) degree :: "'a list => nat" where 
+definition (in semiring_0) degree :: "'a list => nat" where
   "degree p = length (pnormalize p) - 1"
 
      --{*squarefree polynomials --- NB with respect to real roots only.*}
@@ -140,12 +140,12 @@
 proof(induct p1 arbitrary: p2)
   case Nil thus ?case by simp
 next
-  case (Cons a as p2) thus ?case 
+  case (Cons a as p2) thus ?case
     by (cases p2, simp_all  add: add_ac right_distrib)
 qed
 
 lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x"
-apply (induct "p") 
+apply (induct "p")
 apply (case_tac [2] "x=zero")
 apply (auto simp add: right_distrib mult_ac)
 done
@@ -163,7 +163,7 @@
   case Nil thus ?case by simp
 next
   case (Cons a as p2)
-  thus ?case by (cases as, 
+  thus ?case by (cases as,
     simp_all add: poly_cmult poly_add left_distrib right_distrib mult_ac)
 qed
 
@@ -213,11 +213,11 @@
   thus ?case by blast
 next
   case (Cons  x xs)
-  {fix h 
-    from Cons.hyps[rule_format, of x] 
+  {fix h
+    from Cons.hyps[rule_format, of x]
     obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
-    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)" 
-      using qr by(cases q, simp_all add: algebra_simps diff_def[symmetric] 
+    have "h#x#xs = [a*r + h] +++ [-a, 1] *** (r#q)"
+      using qr by(cases q, simp_all add: algebra_simps diff_def[symmetric]
 	minus_mult_left[symmetric] right_minus)
     hence "\<exists>q r. h#x#xs = [r] +++ [-a, 1] *** q" by blast}
   thus ?case by blast
@@ -232,11 +232,11 @@
   {assume p: "p = []" hence ?thesis by simp}
   moreover
   {fix x xs assume p: "p = x#xs"
-    {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0" 
+    {fix q assume "p = [-a, 1] *** q" hence "poly p a = 0"
 	by (simp add: poly_add poly_cmult minus_mult_left[symmetric])}
     moreover
     {assume p0: "poly p a = 0"
-      from poly_linear_rem[of x xs a] obtain q r 
+      from poly_linear_rem[of x xs a] obtain q r
       where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast
       have "r = 0" using p0 by (simp only: p qr poly_mult poly_add) simp
       hence "\<exists>q. p = [- a, 1] *** q" using p qr  apply - apply (rule exI[where x=q])apply auto apply (cases q) apply auto done}
@@ -266,7 +266,7 @@
 lemma (in semiring_0) poly_root_mult_length[simp]: "length([a,b] *** p) = Suc (length p)"
 by (simp add: poly_add_length)
 
-lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: 
+lemma (in idom) poly_mult_not_eq_poly_Nil[simp]:
  "poly (p *** q) x \<noteq> poly [] x \<longleftrightarrow> poly p x \<noteq> poly [] x \<and> poly q x \<noteq> poly [] x"
 by (auto simp add: poly_mult)
 
@@ -280,31 +280,31 @@
 
 text{*A nontrivial polynomial of degree n has no more than n roots*}
 lemma (in idom) poly_roots_index_lemma:
-   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n" 
+   assumes p: "poly p x \<noteq> poly [] x" and n: "length p = n"
   shows "\<exists>i. \<forall>x. poly p x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)"
   using p n
 proof(induct n arbitrary: p x)
-  case 0 thus ?case by simp 
+  case 0 thus ?case by simp
 next
   case (Suc n p x)
   {assume C: "\<And>i. \<exists>x. poly p x = 0 \<and> (\<forall>m\<le>Suc n. x \<noteq> i m)"
     from Suc.prems have p0: "poly p x \<noteq> 0" "p\<noteq> []" by auto
-    from p0(1)[unfolded poly_linear_divides[of p x]] 
+    from p0(1)[unfolded poly_linear_divides[of p x]]
     have "\<forall>q. p \<noteq> [- x, 1] *** q" by blast
     from C obtain a where a: "poly p a = 0" by blast
-    from a[unfolded poly_linear_divides[of p a]] p0(2) 
+    from a[unfolded poly_linear_divides[of p a]] p0(2)
     obtain q where q: "p = [-a, 1] *** q" by blast
     have lg: "length q = n" using q Suc.prems(2) by simp
-    from q p0 have qx: "poly q x \<noteq> poly [] x" 
+    from q p0 have qx: "poly q x \<noteq> poly [] x"
       by (auto simp add: poly_mult poly_add poly_cmult)
-    from Suc.hyps[OF qx lg] obtain i where 
+    from Suc.hyps[OF qx lg] obtain i where
       i: "\<forall>x. poly q x = 0 \<longrightarrow> (\<exists>m\<le>n. x = i m)" by blast
     let ?i = "\<lambda>m. if m = Suc n then a else i m"
-    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m" 
+    from C[of ?i] obtain y where y: "poly p y = 0" "\<forall>m\<le> Suc n. y \<noteq> ?i m"
       by blast
-    from y have "y = a \<or> poly q y = 0" 
+    from y have "y = a \<or> poly q y = 0"
       by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps)
-    with i[rule_format, of y] y(1) y(2) have False apply auto 
+    with i[rule_format, of y] y(1) y(2) have False apply auto
       apply (erule_tac x="m" in allE)
       apply auto
       done}
@@ -320,7 +320,7 @@
       \<exists>N i. \<forall>x. (poly p x = 0) --> (\<exists>n. (n::nat) < N & x = i n)"
 apply (drule poly_roots_index_length, safe)
 apply (rule_tac x = "Suc (length p)" in exI)
-apply (rule_tac x = i in exI) 
+apply (rule_tac x = i in exI)
 apply (simp add: less_Suc_eq_le)
 done
 
@@ -344,8 +344,8 @@
 apply (erule_tac x="x" in allE, clarsimp)
 by (case_tac "n=length p", auto simp add: order_le_less)
 
-lemma (in ring_char_0) UNIV_ring_char_0_infinte: 
-  "\<not> (finite (UNIV:: 'a set))" 
+lemma (in ring_char_0) UNIV_ring_char_0_infinte:
+  "\<not> (finite (UNIV:: 'a set))"
 proof
   assume F: "finite (UNIV :: 'a set)"
   have "finite (UNIV :: nat set)"
@@ -357,7 +357,7 @@
   with infinite_UNIV_nat show False ..
 qed
 
-lemma (in idom_char_0) poly_roots_finite: "(poly p \<noteq> poly []) = 
+lemma (in idom_char_0) poly_roots_finite: "(poly p \<noteq> poly []) =
   finite {x. poly p x = 0}"
 proof
   assume H: "poly p \<noteq> poly []"
@@ -370,7 +370,7 @@
     using finite_subset
   proof-
     fix x i
-    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}" 
+    assume F: "\<not> finite {x. poly p x = (0\<Colon>'a)}"
       and P: "\<forall>x. poly p x = (0\<Colon>'a) \<longrightarrow> x \<in> set i"
     let ?M= "{x. poly p x = (0\<Colon>'a)}"
     from P have "?M \<subseteq> set i" by auto
@@ -378,12 +378,12 @@
   qed
 next
   assume F: "finite {x. poly p x = (0\<Colon>'a)}"
-  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto  
+  show "poly p \<noteq> poly []" using F UNIV_ring_char_0_infinte by auto
 qed
 
 text{*Entirety and Cancellation for polynomials*}
 
-lemma (in idom_char_0) poly_entire_lemma2: 
+lemma (in idom_char_0) poly_entire_lemma2:
   assumes p0: "poly p \<noteq> poly []" and q0: "poly q \<noteq> poly []"
   shows "poly (p***q) \<noteq> poly []"
 proof-
@@ -392,7 +392,7 @@
   with p0 q0 show ?thesis  unfolding poly_roots_finite by auto
 qed
 
-lemma (in idom_char_0) poly_entire: 
+lemma (in idom_char_0) poly_entire:
   "poly (p *** q) = poly [] \<longleftrightarrow> poly p = poly [] \<or> poly q = poly []"
 using poly_entire_lemma2[of p q]
 by (auto simp add: expand_fun_eq poly_mult)
@@ -420,8 +420,8 @@
 
 lemma (in recpower_idom) poly_exp_eq_zero[simp]:
      "(poly (p %^ n) = poly []) = (poly p = poly [] & n \<noteq> 0)"
-apply (simp only: fun_eq add: all_simps [symmetric]) 
-apply (rule arg_cong [where f = All]) 
+apply (simp only: fun_eq add: all_simps [symmetric])
+apply (rule arg_cong [where f = All])
 apply (rule ext)
 apply (induct n)
 apply (auto simp add: poly_exp poly_mult)
@@ -435,7 +435,7 @@
 apply (subst add_commute)
 apply (subst add_assoc)
 apply simp
-done 
+done
 
 lemma (in recpower_idom) poly_exp_prime_eq_zero: "(poly ([a, 1] %^ n) \<noteq> poly [])"
 by auto
@@ -445,8 +445,8 @@
 lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] ==> h = 0 & poly t = poly []"
 apply(simp add: fun_eq)
 apply (case_tac "h = zero")
-apply (drule_tac [2] x = zero in spec, auto) 
-apply (cases "poly t = poly []", simp) 
+apply (drule_tac [2] x = zero in spec, auto)
+apply (cases "poly t = poly []", simp)
 proof-
   fix x
   assume H: "\<forall>x. x = (0\<Colon>'a) \<or> poly t x = (0\<Colon>'a)"  and pnz: "poly t \<noteq> poly []"
@@ -563,17 +563,17 @@
   {assume p0: "poly p a = 0"
     from Suc.prems have h: "length p = Suc n" "poly p \<noteq> poly []" by auto
     hence pN: "p \<noteq> []" by auto
-    from p0[unfolded poly_linear_divides] pN  obtain q where 
+    from p0[unfolded poly_linear_divides] pN  obtain q where
       q: "p = [-a, 1] *** q" by blast
-    from q h p0 have qh: "length q = n" "poly q \<noteq> poly []" 
+    from q h p0 have qh: "length q = n" "poly q \<noteq> poly []"
       apply -
       apply simp
       apply (simp only: fun_eq)
       apply (rule ccontr)
       apply (simp add: fun_eq poly_add poly_cmult minus_mult_left[symmetric])
       done
-    from Suc.hyps[OF qh] obtain m r where 
-      mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0" by blast    
+    from Suc.hyps[OF qh] obtain m r where
+      mr: "q = mulexp m [-a,1] r" "poly r a \<noteq> 0" by blast
     from mr q have "p = mulexp (Suc m) [-a,1] r \<and> poly r a \<noteq> 0" by simp
     hence ?case by blast}
   moreover
@@ -600,7 +600,7 @@
 
 (* FIXME: Tidy up *)
 
-lemma (in recpower_semiring_1) 
+lemma (in recpower_semiring_1)
   zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)"
   by (induct n, simp_all add: power_Suc)
 
@@ -614,15 +614,15 @@
 from lp p0
 show ?thesis
 apply -
-apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)  
+apply (drule poly_order_exists_lemma [where a=a], assumption, clarify)
 apply (rule_tac x = n in exI, safe)
 apply (unfold divides_def)
 apply (rule_tac x = q in exI)
 apply (induct_tac "n", simp)
 apply (simp (no_asm_simp) add: poly_add poly_cmult poly_mult right_distrib mult_ac)
 apply safe
-apply (subgoal_tac "?poly (?mulexp n [uminus a, one] q) \<noteq> ?poly (pmult (?pexp [uminus a, one] (Suc n)) qa)") 
-apply simp 
+apply (subgoal_tac "?poly (?mulexp n [uminus a, one] q) \<noteq> ?poly (pmult (?pexp [uminus a, one] (Suc n)) qa)")
+apply simp
 apply (induct_tac "n")
 apply (simp del: pmult_Cons pexp_Suc)
 apply (erule_tac Q = "?poly q a = zero" in contrapos_np)
@@ -670,7 +670,7 @@
 lemma (in recpower_idom_char_0) order_unique: "[| poly p \<noteq> poly []; ([-a, 1] %^ n) divides p;
          ~(([-a, 1] %^ (Suc n)) divides p)
       |] ==> (n = order a p)"
-by (insert order [of a n p], auto) 
+by (insert order [of a n p], auto)
 
 lemma (in recpower_idom_char_0) order_unique_lemma: "(poly p \<noteq> poly [] & ([-a, 1] %^ n) divides p &
          ~(([-a, 1] %^ (Suc n)) divides p))
@@ -695,7 +695,7 @@
 lemma (in recpower_idom_char_0) order_root: "(poly p a = 0) = ((poly p = poly []) | order a p \<noteq> 0)"
 proof-
   let ?poly = poly
-  show ?thesis 
+  show ?thesis
 apply (case_tac "?poly p = ?poly []", auto)
 apply (simp add: poly_linear_divides del: pmult_Cons, safe)
 apply (drule_tac [!] a = a in order2)
@@ -709,7 +709,7 @@
 lemma (in recpower_idom_char_0) order_divides: "(([-a, 1] %^ n) divides p) = ((poly p = poly []) | n \<le> order a p)"
 proof-
   let ?poly = poly
-  show ?thesis 
+  show ?thesis
 apply (case_tac "?poly p = ?poly []", auto)
 apply (simp add: divides_def fun_eq poly_mult)
 apply (rule_tac x = "[]" in exI)
@@ -753,14 +753,14 @@
 apply (simp add: fun_eq poly_exp_add poly_mult mult_ac del: pmult_Cons)
 done
 
-lemma (in recpower_idom_char_0) order_mult: 
+lemma (in recpower_idom_char_0) order_mult:
   assumes pq0: "poly (p *** q) \<noteq> poly []"
   shows "order a (p *** q) = order a p + order a q"
 proof-
   let ?order = order
   let ?divides = "op divides"
   let ?poly = poly
-from pq0 
+from pq0
 show ?thesis
 apply (cut_tac a = a and p = "pmult p q" and n = "?order a p + ?order a q" in order)
 apply (auto simp add: poly_entire simp del: pmult_Cons)
@@ -821,12 +821,12 @@
      "list_all (%c. c = 0) p \<longleftrightarrow>  pnormalize p = []"
 by (induct "p", auto)
 
-lemma (in idom_char_0) degree_zero: 
+lemma (in idom_char_0) degree_zero:
   assumes pN: "poly p = poly []" shows"degree p = 0"
 proof-
   let ?pn = pnormalize
   from pN
-  show ?thesis 
+  show ?thesis
     apply (simp add: degree_def)
     apply (case_tac "?pn p = []")
     apply (auto simp add: poly_zero lemma_degree_zero )
@@ -835,10 +835,10 @@
 
 lemma (in semiring_0) pnormalize_sing: "(pnormalize [x] = [x]) \<longleftrightarrow> x \<noteq> 0" by simp
 lemma (in semiring_0) pnormalize_pair: "y \<noteq> 0 \<longleftrightarrow> (pnormalize [x, y] = [x, y])" by simp
-lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)" 
+lemma (in semiring_0) pnormal_cons: "pnormal p \<Longrightarrow> pnormal (c#p)"
   unfolding pnormal_def by simp
 lemma (in semiring_0) pnormal_tail: "p\<noteq>[] \<Longrightarrow> pnormal (c#p) \<Longrightarrow> pnormal p"
-  unfolding pnormal_def 
+  unfolding pnormal_def
   apply (cases "pnormalize p = []", auto)
   by (cases "c = 0", auto)
 
@@ -846,7 +846,7 @@
 lemma (in semiring_0) pnormal_last_nonzero: "pnormal p ==> last p \<noteq> 0"
 proof(induct p)
   case Nil thus ?case by (simp add: pnormal_def)
-next 
+next
   case (Cons a as) thus ?case
     apply (simp add: pnormal_def)
     apply (cases "pnormalize as = []", simp_all)
@@ -877,12 +877,12 @@
   hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
     unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
   hence "c = d \<and> (\<forall>x. poly (cs +++ -- ds) x = 0)"
-    unfolding poly_zero[symmetric] by simp 
+    unfolding poly_zero[symmetric] by simp
   thus ?rhs  by (simp add: poly_minus poly_add algebra_simps expand_fun_eq)
 next
   assume ?rhs then show ?lhs by(simp add:expand_fun_eq)
 qed
-  
+
 lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \<Longrightarrow> pnormalize p = pnormalize q"
 proof(induct q arbitrary: p)
   case Nil thus ?case by (simp only: poly_zero lemma_degree_zero) simp
@@ -892,7 +892,7 @@
   proof(induct p)
     case Nil
     hence "poly [] = poly (c#cs)" by blast
-    then have "poly (c#cs) = poly [] " by simp 
+    then have "poly (c#cs) = poly [] " by simp
     thus ?case by (simp only: poly_zero lemma_degree_zero) simp
   next
     case (Cons d ds)
@@ -913,7 +913,7 @@
 
 lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \<le> length p" by (induct p, auto)
 
-lemma (in semiring_0) last_linear_mul_lemma: 
+lemma (in semiring_0) last_linear_mul_lemma:
   "last ((a %* p) +++ (x#(b %* p))) = (if p=[] then x else b*last p)"
 
 apply (induct p arbitrary: a x b, auto)
@@ -948,14 +948,14 @@
 proof-
   from p have pnz: "pnormalize p \<noteq> []"
     unfolding poly_zero lemma_degree_zero .
-  
+
   from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz]
   have l0: "last ([a, 1] *** pnormalize p) \<noteq> 0" by simp
   from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a]
     pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz
- 
+
 
-  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" 
+  have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1"
     by (auto simp add: poly_length_mult)
 
   have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)"
@@ -980,26 +980,26 @@
     by (simp add: mult_ac add_ac right_distrib)
   note deq = degree_unique[OF eq]
   {assume p: "poly p = poly []"
-    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []" 
+    with eq have eq': "poly ([a,1] %^(Suc n) *** p) = poly []"
       by - (rule ext,simp add: poly_mult poly_cmult poly_add)
     from degree_unique[OF eq'] p have ?case by (simp add: degree_def)}
   moreover
   {assume p: "poly p \<noteq> poly []"
     from p have ap: "poly ([a,1] *** p) \<noteq> poly []"
-      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto 
+      using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto
     have eq: "poly ([a,1] %^(Suc n) *** p) = poly ([a,1]%^n *** ([a,1] *** p))"
      by (rule ext, simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps)
    from ap have ap': "(poly ([a,1] *** p) = poly []) = False" by blast
    have  th0: "degree ([a,1]%^n *** ([a,1] *** p)) = degree ([a,1] *** p) + n"
      apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap')
      by simp
-    
+
    from degree_unique[OF eq] ap p th0 linear_mul_degree[OF p, of a]
    have ?case by (auto simp del: poly.simps)}
   ultimately show ?case by blast
 qed
 
-lemma (in recpower_idom_char_0) order_degree: 
+lemma (in recpower_idom_char_0) order_degree:
   assumes p0: "poly p \<noteq> poly []"
   shows "order a p \<le> degree p"
 proof-
@@ -1007,7 +1007,7 @@
   obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast
   {assume "poly q = poly []"
     with q p0 have False by (simp add: poly_mult poly_entire)}
-  with degree_unique[OF q, unfolded linear_pow_mul_degree] 
+  with degree_unique[OF q, unfolded linear_pow_mul_degree]
   show ?thesis by auto
 qed
 
--- a/src/HOL/Nominal/nominal_primrec.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Nominal/nominal_primrec.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -12,6 +12,10 @@
     (binding * typ option * mixfix) list ->
     (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> local_theory -> Proof.state
+  val add_primrec_cmd: string list option -> string option ->
+    (binding * string option * mixfix) list ->
+    (binding * string option * mixfix) list ->
+    (Attrib.binding * string) list -> local_theory -> Proof.state
 end;
 
 structure NominalPrimrec : NOMINAL_PRIMREC =
@@ -36,10 +40,10 @@
       (fn Free (v, _) => insert (op =) v | _ => I) body []))
   in (curry subst_bounds (map2 (curry Free) vs' Ts |> rev) body) end;
 
-fun process_eqn lthy is_fixed spec rec_fns = 
+fun process_eqn lthy is_fixed spec rec_fns =
   let
     val eq = unquantify spec;
-    val (lhs, rhs) = 
+    val (lhs, rhs) =
       HOLogic.dest_eq (HOLogic.dest_Trueprop (Logic.strip_imp_concl eq))
       handle TERM _ => raise RecError "not a proper equation";
 
@@ -67,7 +71,7 @@
     fun check_vars _ [] = ()
       | check_vars s vars = raise RecError (s ^ commas_quote (map fst vars))
   in
-    if length middle > 1 then 
+    if length middle > 1 then
       raise RecError "more than one non-variable in pattern"
     else
      (check_vars "repeated variable names in pattern: " (duplicates (op =) lfrees);
@@ -159,7 +163,7 @@
               val (rhs', (fnames'', fnss'')) = subst (map2 (fn (x, y) => fn z =>
                 (Free x, (body_index y, Free z))) recs subs) rhs (fnames', fnss')
                   handle RecError s => primrec_eq_err lthy s eq
-            in (fnames'', fnss'', 
+            in (fnames'', fnss'',
                 (list_abs_free (cargs' @ subs, rhs'))::fns)
             end)
 
@@ -172,7 +176,7 @@
             val SOME (_, _, eqns' as (_, (ls, _, rs, _, _)) :: _) =
               AList.lookup (op =) eqns fname;
             val (fnames', fnss', fns) = fold_rev (trans eqns') constrs
-              ((i, fname)::fnames, fnss, []) 
+              ((i, fname)::fnames, fnss, [])
           in
             (fnames', (i, (fname, ls, rs, fns))::fnss')
           end
@@ -235,15 +239,9 @@
 
 local
 
-fun prepare_spec prep_spec ctxt raw_fixes raw_spec =
-  let
-    val ((fixes, spec), _) = prep_spec
-      raw_fixes (map (single o apsnd single) raw_spec) ctxt
-  in (fixes, map (apsnd the_single) spec) end;
-
 fun gen_primrec set_group prep_spec prep_term invs fctxt raw_fixes raw_params raw_spec lthy =
   let
-    val (fixes', spec) = prepare_spec prep_spec lthy (raw_fixes @ raw_params) raw_spec;
+    val (fixes', spec) = fst (prep_spec (raw_fixes @ raw_params) raw_spec lthy);
     val fixes = List.take (fixes', length raw_fixes);
     val (names_atts, spec') = split_list spec;
     val eqns' = map unquantify spec'
@@ -261,12 +259,12 @@
        then () else primrec_err param_err);
     val tnames = distinct (op =) (map (#1 o snd) eqns);
     val dts = find_dts dt_info tnames tnames;
-    val main_fns = 
+    val main_fns =
       map (fn (tname, {index, ...}) =>
-        (index, 
+        (index,
           (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns))
       dts;
-    val {descr, rec_names, rec_rewrites, ...} = 
+    val {descr, rec_names, rec_rewrites, ...} =
       if null dts then
         primrec_err ("datatypes " ^ commas_quote tnames ^ "\nare not mutually recursive")
       else snd (hd dts);
@@ -388,15 +386,15 @@
 
 in
 
-val add_primrec = gen_primrec false Specification.check_specification (K I);
-val add_primrec_cmd = gen_primrec true Specification.read_specification Syntax.read_term;
+val add_primrec = gen_primrec false Specification.check_spec (K I);
+val add_primrec_cmd = gen_primrec true Specification.read_spec Syntax.read_term;
 
 end;
 
 
 (* outer syntax *)
 
-local structure P = OuterParse and K = OuterKeyword in
+local structure P = OuterParse in
 
 val freshness_context = P.reserved "freshness_context";
 val invariant = P.reserved "invariant";
@@ -408,28 +406,16 @@
     (Scan.repeat1 (unless_flag P.term) >> SOME) -- Scan.optional parser1 NONE ||
   (parser1 >> pair NONE);
 val options =
-  Scan.optional (P.$$$ "(" |-- P.!!!
-    (parser2 --| P.$$$ ")")) (NONE, NONE);
-
-fun pipe_error t = P.!!! (Scan.fail_with (K
-  (cat_lines ["Equations must be separated by " ^ quote "|", quote t])));
-
-val statement = SpecParse.opt_thm_name ":" -- P.prop --| Scan.ahead
-  ((P.term :-- pipe_error) || Scan.succeed ("",""));
-
-val statements = P.enum1 "|" statement;
-
-val primrec_decl = P.opt_target -- options --
-  P.fixes -- P.for_fixes --| P.$$$ "where" -- statements;
+  Scan.optional (P.$$$ "(" |-- P.!!! (parser2 --| P.$$$ ")")) (NONE, NONE);
 
 val _ =
-  OuterSyntax.command "nominal_primrec" "define primitive recursive functions on nominal datatypes" K.thy_goal
-    (primrec_decl >> (fn ((((opt_target, (invs, fctxt)), raw_fixes), raw_params), raw_spec) =>
-      Toplevel.print o Toplevel.local_theory_to_proof opt_target
-        (add_primrec_cmd invs fctxt raw_fixes raw_params raw_spec)));
+  OuterSyntax.local_theory_to_proof "nominal_primrec"
+    "define primitive recursive functions on nominal datatypes" OuterKeyword.thy_goal
+    (options -- P.fixes -- P.for_fixes -- SpecParse.where_alt_specs
+      >> (fn ((((invs, fctxt), fixes), params), specs) =>
+        add_primrec_cmd invs fctxt fixes params specs));
 
 end;
 
-
 end;
 
--- a/src/HOL/Statespace/state_space.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Statespace/state_space.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -419,7 +419,7 @@
         val expr = ([(suffix valuetypesN name, 
                      (("",false),Expression.Positional (map SOME pars)))],[]);
       in
-        prove_interpretation_in (ALLGOALS o solve_tac o Assumption.prems_of)
+        prove_interpretation_in (ALLGOALS o solve_tac o Assumption.all_prems_of)
           (suffix valuetypesN name, expr) thy
       end;
 
--- a/src/HOL/Tools/function_package/fundef_common.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_common.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -359,7 +359,7 @@
                          >> (fn sow => (map (snd o snd) sow, map (apsnd fst) sow))
 in
   fun fundef_parser default_cfg = 
-      config_parser default_cfg -- P.fixes --| P.$$$ "where" -- flags_statements
+      config_parser default_cfg -- P.fixes --| P.where_ -- flags_statements
 end
 
 
--- a/src/HOL/Tools/function_package/fundef_package.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Tools/function_package/fundef_package.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -93,13 +93,12 @@
     end
 
 
-fun gen_add_fundef is_external prep default_constraint fixspec eqnss config flags lthy =
+fun gen_add_fundef is_external prep default_constraint fixspec eqns config flags lthy =
     let
       val constrn_fxs = map (fn (b, T, mx) => (b, SOME (the_default default_constraint T), mx))
-      val ((fixes0, spec0), ctxt') = 
-        prep (constrn_fxs fixspec) (map (single o apsnd single) eqnss) lthy
+      val ((fixes0, spec0), ctxt') = prep (constrn_fxs fixspec) eqns lthy
       val fixes = map (apfst (apfst Binding.name_of)) fixes0;
-      val spec = map (apfst (apfst Binding.name_of)) spec0;
+      val spec = map (fn ((b, atts), prop) => ((Binding.name_of b, atts), [prop])) spec0;
       val (eqs, post, sort_cont, cnames) = FundefCommon.get_preproc lthy config flags ctxt' fixes spec
 
       val defname = mk_defname fixes
@@ -163,9 +162,8 @@
   |> LocalTheory.set_group (serial_string ())
   |> setup_termination_proof term_opt;
 
-val add_fundef = gen_add_fundef true Specification.read_specification "_::type"
-val add_fundef_i = 
-  gen_add_fundef false Specification.check_specification (TypeInfer.anyT HOLogic.typeS)
+val add_fundef = gen_add_fundef true Specification.read_spec "_::type"
+val add_fundef_i = gen_add_fundef false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
 
 
 (* Datatype hook to declare datatype congs as "fundef_congs" *)
--- a/src/HOL/Tools/inductive_package.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Tools/inductive_package.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -842,11 +842,10 @@
 
 fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos int lthy =
   let
-    val ((vars, specs), _) = lthy |> ProofContext.set_mode ProofContext.mode_abbrev
-      |> Specification.read_specification
-          (cnames_syn @ pnames_syn) (map (fn (a, s) => [(a, [s])]) intro_srcs);
+    val ((vars, intrs), _) = lthy
+      |> ProofContext.set_mode ProofContext.mode_abbrev
+      |> Specification.read_spec (cnames_syn @ pnames_syn) intro_srcs;
     val (cs, ps) = chop (length cnames_syn) vars;
-    val intrs = map (apsnd the_single) specs;
     val monos = Attrib.eval_thms lthy raw_monos;
     val flags = {quiet_mode = false, verbose = verbose, kind = Thm.theoremK,
       alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false,
@@ -946,22 +945,12 @@
 
 val _ = OuterKeyword.keyword "monos";
 
-(* FIXME eliminate *)
-fun flatten_specification specs = specs |> maps
-  (fn (a, (concl, [])) => concl |> map
-        (fn ((b, atts), [B]) =>
-              if Binding.is_empty a then ((b, atts), B)
-              else if Binding.is_empty b then ((a, atts), B)
-              else error "Illegal nested case names"
-          | ((b, _), _) => error "Illegal simultaneous specification")
-    | (a, _) => error ("Illegal local specification parameters for " ^ quote (Binding.str_of a)));
-
 fun gen_ind_decl mk_def coind =
   P.fixes -- P.for_fixes --
-  Scan.optional (P.$$$ "where" |-- P.!!! SpecParse.specification) [] --
+  Scan.optional SpecParse.where_alt_specs [] --
   Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) []
   >> (fn (((preds, params), specs), monos) =>
-      (snd oo gen_add_inductive mk_def true coind preds params (flatten_specification specs) monos));
+      (snd oo gen_add_inductive mk_def true coind preds params specs monos));
 
 val ind_decl = gen_ind_decl add_ind_def;
 
@@ -971,7 +960,7 @@
 val _ =
   OuterSyntax.local_theory "inductive_cases"
     "create simplified instances of elimination rules (improper)" K.thy_script
-    (P.and_list1 SpecParse.spec >> (fn specs => snd o inductive_cases specs));
+    (P.and_list1 SpecParse.specs >> (snd oo inductive_cases));
 
 end;
 
--- a/src/HOL/Tools/primrec_package.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOL/Tools/primrec_package.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -9,6 +9,8 @@
 sig
   val add_primrec: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> local_theory -> thm list * local_theory
+  val add_primrec_cmd: (binding * string option * mixfix) list ->
+    (Attrib.binding * string) list -> local_theory -> thm list * local_theory
   val add_primrec_global: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> theory -> thm list * theory
   val add_primrec_overloaded: (string * (string * typ) * bool) list ->
@@ -213,12 +215,6 @@
 
 local
 
-fun prepare_spec prep_spec ctxt raw_fixes raw_spec =
-  let
-    val ((fixes, spec), _) = prep_spec
-      raw_fixes (map (single o apsnd single) raw_spec) ctxt
-  in (fixes, map (apsnd the_single) spec) end;
-
 fun prove_spec ctxt names rec_rewrites defs eqs =
   let
     val rewrites = map mk_meta_eq rec_rewrites @ map (snd o snd) defs;
@@ -228,7 +224,7 @@
 
 fun gen_primrec set_group prep_spec raw_fixes raw_spec lthy =
   let
-    val (fixes, spec) = prepare_spec prep_spec lthy raw_fixes raw_spec;
+    val (fixes, spec) = fst (prep_spec raw_fixes raw_spec lthy);
     val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v
       orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes) o snd) spec [];
     val tnames = distinct (op =) (map (#1 o snd) eqns);
@@ -268,8 +264,8 @@
 
 in
 
-val add_primrec = gen_primrec false Specification.check_specification;
-val add_primrec_cmd = gen_primrec true Specification.read_specification;
+val add_primrec = gen_primrec false Specification.check_spec;
+val add_primrec_cmd = gen_primrec true Specification.read_spec;
 
 end;
 
@@ -300,24 +296,16 @@
 val old_primrec_decl =
   opt_unchecked_name -- Scan.repeat1 ((SpecParse.opt_thm_name ":" >> apfst Binding.name_of) -- P.prop);
 
-fun pipe_error t = P.!!! (Scan.fail_with (K
-  (cat_lines ["Equations must be separated by " ^ quote "|", quote t])));
-
-val statement = SpecParse.opt_thm_name ":" -- P.prop --| Scan.ahead
-  ((P.term :-- pipe_error) || Scan.succeed ("",""));
-
-val statements = P.enum1 "|" statement;
-
-val primrec_decl = P.opt_target -- P.fixes --| P.$$$ "where" -- statements;
+val primrec_decl = P.opt_target -- P.fixes -- SpecParse.where_alt_specs;
 
 val _ =
   OuterSyntax.command "primrec" "define primitive recursive functions on datatypes" K.thy_decl
-    ((primrec_decl >> (fn ((opt_target, raw_fixes), raw_spec) =>
-      Toplevel.local_theory opt_target (add_primrec_cmd raw_fixes raw_spec #> snd)))
+    ((primrec_decl >> (fn ((opt_target, fixes), specs) =>
+      Toplevel.local_theory opt_target (add_primrec_cmd fixes specs #> snd)))
     || (old_primrec_decl >> (fn ((unchecked, alt_name), eqns) =>
       Toplevel.theory (snd o
-        (if unchecked then OldPrimrecPackage.add_primrec_unchecked else OldPrimrecPackage.add_primrec) alt_name
-          (map P.triple_swap eqns)))));
+        (if unchecked then OldPrimrecPackage.add_primrec_unchecked else OldPrimrecPackage.add_primrec)
+          alt_name (map P.triple_swap eqns)))));
 
 end;
 
--- a/src/HOLCF/Tools/domain/domain_axioms.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOLCF/Tools/domain/domain_axioms.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -108,7 +108,16 @@
     [take_def, finite_def])
 end; (* let (calc_axioms) *)
 
-fun infer_props thy = map (apsnd (FixrecPackage.legacy_infer_prop thy));
+
+(* legacy type inference *)
+
+fun legacy_infer_term thy t =
+  singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t);
+
+fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t);
+
+fun infer_props thy = map (apsnd (legacy_infer_prop thy));
+
 
 fun add_axioms_i x = snd o PureThy.add_axioms (map (Thm.no_attributes o apfst Binding.name) x);
 fun add_axioms_infer axms thy = add_axioms_i (infer_props thy axms) thy;
--- a/src/HOLCF/Tools/fixrec_package.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/HOLCF/Tools/fixrec_package.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -6,17 +6,12 @@
 
 signature FIXREC_PACKAGE =
 sig
-  val legacy_infer_term: theory -> term -> term
-  val legacy_infer_prop: theory -> term -> term
-
-  val add_fixrec: bool -> (binding * string option * mixfix) list
+  val add_fixrec: bool -> (binding * typ option * mixfix) list
+    -> (Attrib.binding * term) list -> local_theory -> local_theory
+  val add_fixrec_cmd: bool -> (binding * string option * mixfix) list
     -> (Attrib.binding * string) list -> local_theory -> local_theory
-
-  val add_fixrec_i: bool -> (binding * typ option * mixfix) list
-    -> (Attrib.binding * term) list -> local_theory -> local_theory
-
-  val add_fixpat: Attrib.binding * string list -> theory -> theory
-  val add_fixpat_i: Thm.binding * term list -> theory -> theory
+  val add_fixpat: Thm.binding * term list -> theory -> theory
+  val add_fixpat_cmd: Attrib.binding * string list -> theory -> theory
   val add_matchers: (string * string) list -> theory -> theory
   val setup: theory -> theory
 end;
@@ -24,14 +19,6 @@
 structure FixrecPackage: FIXREC_PACKAGE =
 struct
 
-(* legacy type inference *)
-(* used by the domain package *)
-fun legacy_infer_term thy t =
-  singleton (Syntax.check_terms (ProofContext.init thy)) (Sign.intern_term thy t);
-
-fun legacy_infer_prop thy t = legacy_infer_term thy (TypeInfer.constrain propT t);
-
-
 val fix_eq2 = @{thm fix_eq2};
 val def_fix_ind = @{thm def_fix_ind};
 
@@ -341,20 +328,9 @@
 local
 (* code adapted from HOL/Tools/primrec_package.ML *)
 
-fun prepare_spec prep_spec ctxt raw_fixes raw_spec =
-  let
-    val ((fixes, spec), _) = prep_spec
-      raw_fixes (map (single o apsnd single) raw_spec) ctxt
-  in (fixes, map (apsnd the_single) spec) end;
-
 fun gen_fixrec
   (set_group : bool)
-  (prep_spec : (binding * 'a option * mixfix) list ->
-       (Attrib.binding * 'b list) list list ->
-      Proof.context ->
-   (((binding * typ) * mixfix) list * (Attrib.binding * term list) list)
-    * Proof.context
-  )
+  prep_spec
   (strict : bool)
   raw_fixes
   raw_spec
@@ -362,7 +338,7 @@
   let
     val (fixes : ((binding * typ) * mixfix) list,
          spec : (Attrib.binding * term) list) =
-          prepare_spec prep_spec lthy raw_fixes raw_spec;
+          fst (prep_spec raw_fixes raw_spec lthy);
     val chead_of_spec =
       chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd;
     fun name_of (Free (n, _)) = n
@@ -405,8 +381,8 @@
 
 in
 
-val add_fixrec_i = gen_fixrec false Specification.check_specification;
-val add_fixrec = gen_fixrec true Specification.read_specification;
+val add_fixrec = gen_fixrec false Specification.check_spec;
+val add_fixrec_cmd = gen_fixrec true Specification.read_spec;
 
 end; (* local *)
 
@@ -434,8 +410,8 @@
     (snd o PureThy.add_thmss [((name, simps), atts)]) thy
   end;
 
-val add_fixpat = gen_add_fixpat Syntax.read_term_global Attrib.attribute;
-val add_fixpat_i = gen_add_fixpat Sign.cert_term (K I);
+val add_fixpat = gen_add_fixpat Sign.cert_term (K I);
+val add_fixpat_cmd = gen_add_fixpat Syntax.read_term_global Attrib.attribute;
 
 
 (*************************************************************************)
@@ -444,43 +420,14 @@
 
 local structure P = OuterParse and K = OuterKeyword in
 
-(* bool parser *)
-val fixrec_strict = P.opt_keyword "permissive" >> not;
-
-fun pipe_error t = P.!!! (Scan.fail_with (K
-  (cat_lines ["Equations must be separated by " ^ quote "|", quote t])));
-
-(* (Attrib.binding * string) parser *)
-val statement = SpecParse.opt_thm_name ":" -- P.prop --| Scan.ahead
-  ((P.term :-- pipe_error) || Scan.succeed ("",""));
-
-(* ((Attrib.binding * string) list) parser *)
-val statements = P.enum1 "|" statement;
-
-(* (((xstring option * bool) * (Binding.binding * string option * Mixfix.mixfix) list)
-   * (Attrib.binding * string) list) parser *)
-val fixrec_decl =
-  P.opt_target -- fixrec_strict -- P.fixes --| P.$$$ "where" -- statements;
+val _ = OuterSyntax.local_theory "fixrec" "define recursive functions (HOLCF)" K.thy_decl
+  ((P.opt_keyword "permissive" >> not) -- P.fixes -- SpecParse.where_alt_specs
+    >> (fn ((strict, fixes), specs) => add_fixrec_cmd strict fixes specs));
 
-(* this builds a parser for a new keyword, fixrec, whose functionality 
-is defined by add_fixrec *)
-val _ =
-  let
-    val desc = "define recursive functions (HOLCF)";
-    fun fixrec (((opt_target, strict), raw_fixes), raw_spec) =
-      Toplevel.local_theory opt_target (add_fixrec strict raw_fixes raw_spec);
-  in
-    OuterSyntax.command "fixrec" desc K.thy_decl (fixrec_decl >> fixrec)
-  end;
-
-(* fixpat parser *)
-val fixpat_decl = SpecParse.opt_thm_name ":" -- Scan.repeat1 P.prop;
-
-val _ =
-  OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
-    (fixpat_decl >> (Toplevel.theory o add_fixpat));
+val _ = OuterSyntax.command "fixpat" "define rewrites for fixrec functions" K.thy_decl
+  (SpecParse.specs >> (Toplevel.theory o add_fixpat_cmd));
   
-end; (* local structure *)
+end;
 
 val setup = FixrecMatchData.init;
 
--- a/src/Pure/General/binding.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/General/binding.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -14,7 +14,7 @@
   val make: bstring * Position.T -> binding
   val pos_of: binding -> Position.T
   val name: bstring -> binding
-  val name_of: binding -> string
+  val name_of: binding -> bstring
   val map_name: (bstring -> bstring) -> binding -> binding
   val prefix_name: string -> binding -> binding
   val suffix_name: string -> binding -> binding
@@ -22,8 +22,8 @@
   val empty: binding
   val is_empty: binding -> bool
   val qualify: bool -> string -> binding -> binding
-  val qualified_name: bstring -> binding
-  val qualified_name_of: binding -> bstring
+  val qualified_name: string -> binding
+  val qualified_name_of: binding -> string
   val prefix_of: binding -> (string * bool) list
   val map_prefix: ((string * bool) list -> (string * bool) list) -> binding -> binding
   val prefix: bool -> string -> binding -> binding
--- a/src/Pure/General/name_space.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/General/name_space.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -36,13 +36,13 @@
   val root_path: naming -> naming
   val parent_path: naming -> naming
   val no_base_names: naming -> naming
-  val sticky_prefix: string -> naming -> naming
+  val mandatory_path: string -> naming -> naming
   type 'a table = T * 'a Symtab.table
-  val bind: naming -> binding * 'a -> 'a table -> string * 'a table       (*exception Symtab.DUP*)
+  val define: naming -> binding * 'a -> 'a table -> string * 'a table (*exception Symtab.DUP*)
   val empty_table: 'a table
-  val merge_tables: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table  (*exception Symtab.DUP*)
+  val merge_tables: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table (*exception Symtab.DUP*)
   val join_tables: (string -> 'a * 'a -> 'a) (*exception Symtab.DUP/Symtab.SAME*) ->
-    'a table * 'a table -> 'a table                                       (*exception Symtab.DUP*)
+    'a table * 'a table -> 'a table (*exception Symtab.DUP*)
   val dest_table: 'a table -> (string * 'a) list
   val extern_table: 'a table -> (xstring * 'a) list
 end;
@@ -156,8 +156,8 @@
 fun merge (NameSpace (tab1, xtab1), NameSpace (tab2, xtab2)) =
   let
     val tab' = (tab1, tab2) |> Symtab.join
-      (K (fn names as ((names1, names1'), (names2, names2')) =>
-        if pointer_eq names then raise Symtab.SAME
+      (K (fn ((names1, names1'), (names2, names2')) =>
+        if pointer_eq (names1, names2) andalso pointer_eq (names1', names2') then raise Symtab.SAME
         else (Library.merge (op =) (names1, names2), Library.merge (op =) (names1', names2'))));
     val xtab' = (xtab1, xtab2) |> Symtab.join
       (K (fn xnames =>
@@ -187,22 +187,22 @@
 val default_naming = make_naming ([], false);
 
 fun add_path elems = map_naming (fn (path, no_base_names) =>
-  (path @ map (rpair false) (Long_Name.explode elems), no_base_names));
+  (path @ [(elems, false)], no_base_names));
 
 val root_path = map_naming (fn (_, no_base_names) => ([], no_base_names));
 
 val parent_path = map_naming (fn (path, no_base_names) =>
   (perhaps (try (#1 o split_last)) path, no_base_names));
 
-fun sticky_prefix elems = map_naming (fn (path, no_base_names) =>
-  (path @ map (rpair true) (Long_Name.explode elems), no_base_names));
+fun mandatory_path elems = map_naming (fn (path, no_base_names) =>
+  (path @ [(elems, true)], no_base_names));
 
 val no_base_names = map_naming (fn (path, _) => (path, true));
 
 
 (* full name *)
 
-fun err_bad binding = error ("Bad name declaration " ^ quote (Binding.str_of binding));
+fun err_bad binding = error ("Bad name binding " ^ quote (Binding.str_of binding));
 
 fun name_spec (Naming {path, ...}) binding =
   let
@@ -230,7 +230,6 @@
   | mandatory_prefixes1 ((x, true) :: xs) = map (cons x) (mandatory_prefixes1 xs)
   | mandatory_prefixes1 ((x, false) :: xs) = map (cons x) (mandatory_prefixes xs);
 
-fun mandatory_suffixes1 xs = map rev (mandatory_prefixes1 (rev xs));
 fun mandatory_suffixes xs = map rev (mandatory_prefixes (rev xs));
 
 fun accesses (naming as Naming {no_base_names, ...}) binding =
@@ -264,7 +263,7 @@
 
 type 'a table = T * 'a Symtab.table;
 
-fun bind naming (binding, x) (space, tab) =
+fun define naming (binding, x) (space, tab) =
   let val (name, space') = declare naming binding space
   in (name, (space', Symtab.update_new (name, x) tab)) end;
 
--- a/src/Pure/Isar/args.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/args.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -227,7 +227,7 @@
 val bang_facts = Scan.peek (fn context =>
   P.position ($$$ "!") >> (fn (_, pos) =>
     (warning ("use of prems in proof method" ^ Position.str_of pos);
-      Assumption.prems_of (Context.proof_of context))) || Scan.succeed []);
+      Assumption.all_prems_of (Context.proof_of context))) || Scan.succeed []);
 
 val from_to =
   P.nat -- ($$$ "-" |-- P.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) ||
--- a/src/Pure/Isar/attrib.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/attrib.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -146,7 +146,7 @@
   let
     val new_attrs =
       raw_attrs |> map (fn (name, att, comment) => (Binding.name name, ((att, comment), stamp ())));
-    fun add attrs = fold (snd oo NameSpace.bind (Sign.naming_of thy)) new_attrs attrs
+    fun add attrs = fold (snd oo NameSpace.define (Sign.naming_of thy)) new_attrs attrs
       handle Symtab.DUP dup => error ("Duplicate declaration of attributes " ^ quote dup);
   in Attributes.map add thy end;
 
--- a/src/Pure/Isar/expression.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/expression.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -674,8 +674,7 @@
             |> def_pred aname parms defs' exts exts';
           val (_, thy'') =
             thy'
-            |> Sign.add_path (Binding.name_of aname)
-            |> Sign.no_base_names
+            |> Sign.mandatory_path (Binding.name_of aname)
             |> PureThy.note_thmss Thm.internalK
               [((Binding.name introN, []), [([intro], [Locale.unfold_attrib])])]
             ||> Sign.restore_naming thy';
@@ -689,8 +688,7 @@
             |> def_pred pname parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred);
           val (_, thy'''') =
             thy'''
-            |> Sign.add_path (Binding.name_of pname)
-            |> Sign.no_base_names
+            |> Sign.mandatory_path (Binding.name_of pname)
             |> PureThy.note_thmss Thm.internalK
                  [((Binding.name introN, []), [([intro], [Locale.intro_attrib])]),
                   ((Binding.name axiomsN, []),
--- a/src/Pure/Isar/isar_cmd.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -6,7 +6,8 @@
 
 signature ISAR_CMD =
 sig
-  val generic_setup: string * Position.T -> theory -> theory
+  val global_setup: string * Position.T -> theory -> theory
+  val local_setup: string * Position.T -> Proof.context -> Proof.context
   val parse_ast_translation: bool * (string * Position.T) -> theory -> theory
   val parse_translation: bool * (string * Position.T) -> theory -> theory
   val print_translation: bool * (string * Position.T) -> theory -> theory
@@ -86,12 +87,16 @@
 
 (** theory declarations **)
 
-(* generic_setup *)
+(* generic setup *)
 
-fun generic_setup (txt, pos) =
+fun global_setup (txt, pos) =
   ML_Context.expression pos "val setup: theory -> theory" "Context.map_theory setup" txt
   |> Context.theory_map;
 
+fun local_setup (txt, pos) =
+  ML_Context.expression pos "val setup: local_theory -> local_theory" "Context.map_proof setup" txt
+  |> Context.proof_map;
+
 
 (* translation functions *)
 
--- a/src/Pure/Isar/isar_syn.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/isar_syn.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -185,9 +185,11 @@
 
 (* axioms and definitions *)
 
+val named_spec = SpecParse.thm_name ":" -- P.prop >> (fn ((x, y), z) => ((x, z), y));
+
 val _ =
   OuterSyntax.command "axioms" "state arbitrary propositions (axiomatic!)" K.thy_decl
-    (Scan.repeat1 SpecParse.spec_name >> (Toplevel.theory o IsarCmd.add_axioms));
+    (Scan.repeat1 named_spec >> (Toplevel.theory o IsarCmd.add_axioms));
 
 val opt_unchecked_overloaded =
   Scan.optional (P.$$$ "(" |-- P.!!!
@@ -196,7 +198,7 @@
 
 val _ =
   OuterSyntax.command "defs" "define constants" K.thy_decl
-    (opt_unchecked_overloaded -- Scan.repeat1 SpecParse.spec_name
+    (opt_unchecked_overloaded -- Scan.repeat1 named_spec
       >> (Toplevel.theory o IsarCmd.add_defs));
 
 
@@ -254,7 +256,7 @@
 val _ =
   OuterSyntax.command "axiomatization" "axiomatic constant specification" K.thy_decl
     (Scan.optional P.fixes [] --
-      Scan.optional (P.where_ |-- P.!!! (P.and_list1 SpecParse.spec)) []
+      Scan.optional (P.where_ |-- P.!!! (P.and_list1 SpecParse.specs)) []
     >> (fn (x, y) => Toplevel.theory (#2 o Specification.axiomatization_cmd x y)));
 
 
@@ -318,8 +320,12 @@
     (P.ML_source >> (Toplevel.no_timing oo IsarCmd.ml_diag false));
 
 val _ =
-  OuterSyntax.command "setup" "apply ML theory setup" (K.tag_ml K.thy_decl)
-    (P.ML_source >> (Toplevel.theory o IsarCmd.generic_setup));
+  OuterSyntax.command "setup" "ML theory setup" (K.tag_ml K.thy_decl)
+    (P.ML_source >> (Toplevel.theory o IsarCmd.global_setup));
+
+val _ =
+  OuterSyntax.local_theory "local_setup" "ML local theory setup" (K.tag_ml K.thy_decl)
+    (P.ML_source >> IsarCmd.local_setup);
 
 val _ =
   OuterSyntax.command "method_setup" "define proof method in ML" (K.tag_ml K.thy_decl)
--- a/src/Pure/Isar/local_defs.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/local_defs.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -135,7 +135,7 @@
 fun export inner outer =    (*beware of closure sizes*)
   let
     val exp = Assumption.export false inner outer;
-    val prems = Assumption.prems_of inner;
+    val prems = Assumption.all_prems_of inner;
   in fn th =>
     let
       val th' = exp th;
--- a/src/Pure/Isar/local_theory.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/local_theory.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -138,12 +138,12 @@
 
 fun full_naming lthy =
   Sign.naming_of (ProofContext.theory_of lthy)
-  |> NameSpace.sticky_prefix (#theory_prefix (get_lthy lthy));
+  |> NameSpace.mandatory_path (#theory_prefix (get_lthy lthy));
 
 fun full_name naming = NameSpace.full_name (full_naming naming);
 
 fun theory_result f lthy = lthy |> raw_theory_result (fn thy => thy
-  |> Sign.sticky_prefix (#theory_prefix (get_lthy lthy))
+  |> Sign.mandatory_path (#theory_prefix (get_lthy lthy))
   |> f
   ||> Sign.restore_naming thy);
 
--- a/src/Pure/Isar/locale.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/locale.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -161,7 +161,7 @@
   | NONE => error ("Unknown locale " ^ quote name);
 
 fun register_locale binding parameters spec intros axioms decls notes dependencies thy =
-  thy |> LocalesData.map (NameSpace.bind (Sign.naming_of thy) (binding,
+  thy |> LocalesData.map (NameSpace.define (Sign.naming_of thy) (binding,
     mk_locale ((parameters, spec, intros, axioms), ((decls, notes), dependencies))) #> snd);
 
 fun change_locale name =
--- a/src/Pure/Isar/method.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/method.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -394,7 +394,7 @@
     val new_meths = raw_meths |> map (fn (name, f, comment) =>
       (Binding.name name, ((f, comment), stamp ())));
 
-    fun add meths = fold (snd oo NameSpace.bind (Sign.naming_of thy)) new_meths meths
+    fun add meths = fold (snd oo NameSpace.define (Sign.naming_of thy)) new_meths meths
       handle Symtab.DUP dup => error ("Duplicate declaration of method " ^ quote dup);
   in Methods.map add thy end;
 
--- a/src/Pure/Isar/proof_context.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/proof_context.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -97,7 +97,7 @@
   val get_thms: Proof.context -> xstring -> thm list
   val get_thm: Proof.context -> xstring -> thm
   val add_path: string -> Proof.context -> Proof.context
-  val sticky_prefix: string -> Proof.context -> Proof.context
+  val mandatory_path: string -> Proof.context -> Proof.context
   val restore_naming: Proof.context -> Proof.context -> Proof.context
   val reset_naming: Proof.context -> Proof.context
   val note_thmss: string -> (Thm.binding * (Facts.ref * attribute list) list) list ->
@@ -301,7 +301,7 @@
   in Display.pretty_thm_aux (Syntax.pp_global thy) true false [] th end;
 
 fun pretty_thm ctxt th =
-  let val asms = map Thm.term_of (Assumption.assms_of ctxt)
+  let val asms = map Thm.term_of (Assumption.all_assms_of ctxt)
   in Display.pretty_thm_aux (Syntax.pp ctxt) false true asms th end;
 
 fun pretty_thms ctxt [th] = pretty_thm ctxt th
@@ -945,7 +945,7 @@
 (* name space operations *)
 
 val add_path        = map_naming o NameSpace.add_path;
-val sticky_prefix   = map_naming o NameSpace.sticky_prefix;
+val mandatory_path  = map_naming o NameSpace.mandatory_path;
 val restore_naming  = map_naming o K o naming_of;
 val reset_naming    = map_naming (K local_naming);
 
@@ -1370,7 +1370,7 @@
           Pretty.commas (map prt_fix fixes))];
 
       (*prems*)
-      val prems = Assumption.prems_of ctxt;
+      val prems = Assumption.all_prems_of ctxt;
       val len = length prems;
       val suppressed = len - ! prems_limit;
       val prt_prems = if null prems then []
--- a/src/Pure/Isar/spec_parse.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/spec_parse.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -13,10 +13,10 @@
   val opt_attribs: Attrib.src list parser
   val thm_name: string -> Attrib.binding parser
   val opt_thm_name: string -> Attrib.binding parser
-  val spec: (Attrib.binding * string list) parser
-  val named_spec: (Attrib.binding * string list) parser
-  val spec_name: ((binding * string) * Attrib.src list) parser
-  val spec_opt_name: ((binding * string) * Attrib.src list) parser
+  val spec: (Attrib.binding * string) parser
+  val specs: (Attrib.binding * string list) parser
+  val alt_specs: (Attrib.binding * string) list parser
+  val where_alt_specs: (Attrib.binding * string) list parser
   val xthm: (Facts.ref * Attrib.src list) parser
   val xthms1: (Facts.ref * Attrib.src list) list parser
   val name_facts: (Attrib.binding * (Facts.ref * Attrib.src list) list) list parser
@@ -30,8 +30,6 @@
   val statement: (Attrib.binding * (string * string list) list) list parser
   val general_statement: (Element.context list * Element.statement) parser
   val statement_keyword: string parser
-  val specification:
-    (binding * ((Attrib.binding * string list) list * (binding * string option) list)) list parser
 end;
 
 structure SpecParse: SPEC_PARSE =
@@ -54,11 +52,13 @@
   Scan.optional ((P.binding -- opt_attribs || attribs >> pair Binding.empty) --| P.$$$ s)
     Attrib.empty_binding;
 
-val spec = opt_thm_name ":" -- Scan.repeat1 P.prop;
-val named_spec = thm_name ":" -- Scan.repeat1 P.prop;
+val spec = opt_thm_name ":" -- P.prop;
+val specs = opt_thm_name ":" -- Scan.repeat1 P.prop;
 
-val spec_name = thm_name ":" -- P.prop >> (fn ((x, y), z) => ((x, z), y));
-val spec_opt_name = opt_thm_name ":" -- P.prop >> (fn ((x, y), z) => ((x, z), y));
+val alt_specs =
+  P.enum1 "|" (spec --| Scan.option (Scan.ahead (P.name || P.$$$ "[") -- P.!!! (P.$$$ "|")));
+
+val where_alt_specs = P.where_ |-- P.!!! alt_specs;
 
 val xthm =
   P.$$$ "[" |-- attribs --| P.$$$ "]" >> pair (Facts.named "") ||
@@ -143,9 +143,4 @@
 
 val statement_keyword = P.$$$ "obtains" || P.$$$ "shows";
 
-
-(* specifications *)
-
-val specification = P.enum1 "|" (P.parbinding -- (P.and_list1 spec -- P.for_simple_fixes));
-
 end;
--- a/src/Pure/Isar/specification.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/specification.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -8,16 +8,22 @@
 signature SPECIFICATION =
 sig
   val print_consts: local_theory -> (string * typ -> bool) -> (string * typ) list -> unit
+  val check_spec:
+    (binding * typ option * mixfix) list -> (Attrib.binding * term) list -> Proof.context ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context
+  val read_spec:
+    (binding * string option * mixfix) list -> (Attrib.binding * string) list -> Proof.context ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context
+  val check_free_spec:
+    (binding * typ option * mixfix) list -> (Attrib.binding * term) list -> Proof.context ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context
+  val read_free_spec:
+    (binding * string option * mixfix) list -> (Attrib.binding * string) list -> Proof.context ->
+    (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context
   val check_specification: (binding * typ option * mixfix) list ->
-    (Attrib.binding * term list) list list -> Proof.context ->
+    (Attrib.binding * term list) list -> Proof.context ->
     (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
   val read_specification: (binding * string option * mixfix) list ->
-    (Attrib.binding * string list) list list -> Proof.context ->
-    (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val check_free_specification: (binding * typ option * mixfix) list ->
-    (Attrib.binding * term list) list -> Proof.context ->
-    (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
-  val read_free_specification: (binding * string option * mixfix) list ->
     (Attrib.binding * string list) list -> Proof.context ->
     (((binding * typ) * mixfix) list * (Attrib.binding * term list) list) * Proof.context
   val axiomatization: (binding * typ option * mixfix) list ->
@@ -97,7 +103,7 @@
       in fold_rev close bounds A end;
   in map close_form As end;
 
-fun prep_spec prep_vars parse_prop prep_att do_close raw_vars raw_specss ctxt =
+fun prepare prep_vars parse_prop prep_att do_close raw_vars raw_specss ctxt =
   let
     val thy = ProofContext.theory_of ctxt;
 
@@ -122,15 +128,30 @@
     val name_atts = map (fn ((name, atts), _) => (name, map (prep_att thy) atts)) (flat raw_specss);
   in ((params, name_atts ~~ specs), specs_ctxt) end;
 
-fun read_spec x = prep_spec ProofContext.read_vars Syntax.parse_prop Attrib.intern_src x;
-fun check_spec x = prep_spec ProofContext.cert_vars (K I) (K I) x;
+
+fun single_spec (a, prop) = [(a, [prop])];
+fun the_spec (a, [prop]) = (a, prop);
+
+fun prep_spec prep_vars parse_prop prep_att do_close vars specs =
+  prepare prep_vars parse_prop prep_att do_close
+    vars (map single_spec specs) #>> apsnd (map the_spec);
+
+fun prep_specification prep_vars parse_prop prep_att vars specs =
+  prepare prep_vars parse_prop prep_att true [specs];
 
 in
 
-fun read_specification x = read_spec true x;
-fun check_specification x = check_spec true x;
-fun read_free_specification vars specs = read_spec false vars [specs];
-fun check_free_specification vars specs = check_spec false vars [specs];
+fun check_spec x = prep_spec ProofContext.cert_vars (K I) (K I) true x;
+fun read_spec x = prep_spec ProofContext.read_vars Syntax.parse_prop Attrib.intern_src true x;
+
+fun check_free_spec x = prep_spec ProofContext.cert_vars (K I) (K I) false x;
+fun read_free_spec x = prep_spec ProofContext.read_vars Syntax.parse_prop Attrib.intern_src false x;
+
+fun check_specification vars specs =
+  prepare ProofContext.cert_vars (K I) (K I) true vars [specs];
+
+fun read_specification vars specs =
+  prepare ProofContext.read_vars Syntax.parse_prop Attrib.intern_src true vars [specs];
 
 end;
 
@@ -139,7 +160,7 @@
 
 fun gen_axioms do_print prep raw_vars raw_specs thy =
   let
-    val ((vars, specs), _) = prep raw_vars [raw_specs] (ProofContext.init thy);
+    val ((vars, specs), _) = prep raw_vars raw_specs (ProofContext.init thy);
     val xs = map (fn ((b, T), _) => (Binding.name_of b, T)) vars;
 
     (*consts*)
@@ -147,12 +168,16 @@
     val subst = Term.subst_atomic (map Free xs ~~ consts);
 
     (*axioms*)
-    val (axioms, axioms_thy) = consts_thy |> fold_map (fn ((b, atts), props) =>
-        fold_map Thm.add_axiom  (* FIXME proper use of binding!? *)
-          ((map o apfst) Binding.name (PureThy.name_multi (Binding.name_of b) (map subst props)))
-        #>> (fn ths => ((b, atts), [(map Drule.standard' ths, [])]))) specs;
+    val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), props) =>
+        fold_map Thm.add_axiom
+          (map (apfst (fn a => Binding.map_name (K a) b))
+            (PureThy.name_multi (Binding.name_of b) (map subst props)))
+        #>> (fn ths => ((b, atts), [(map Drule.standard' ths, [])])));
+
+    (*facts*)
     val (facts, thy') = axioms_thy |> PureThy.note_thmss Thm.axiomK
       (Attrib.map_facts (Attrib.attribute_i axioms_thy) axioms);
+
     val _ =
       if not do_print then ()
       else print_consts (ProofContext.init thy') (K false) xs;
@@ -164,10 +189,9 @@
 
 (* definition *)
 
-fun gen_def do_print prep (raw_var, (raw_a, raw_prop)) lthy =
+fun gen_def do_print prep (raw_var, raw_spec) lthy =
   let
-    val (vars, [((raw_name, atts), [prop])]) =
-      fst (prep (the_list raw_var) [(raw_a, [raw_prop])] lthy);
+    val (vars, [((raw_name, atts), prop)]) = fst (prep (the_list raw_var) [raw_spec] lthy);
     val (((x, T), rhs), prove) = LocalDefs.derived_def lthy true prop;
     val var as (b, _) =
       (case vars of
@@ -193,16 +217,16 @@
       else print_consts lthy3 (member (op =) (Term.add_frees lhs' [])) [(x, T)];
   in ((lhs, (def_name, th')), lthy3) end;
 
-val definition = gen_def false check_free_specification;
-val definition_cmd = gen_def true read_free_specification;
+val definition = gen_def false check_free_spec;
+val definition_cmd = gen_def true read_free_spec;
 
 
 (* abbreviation *)
 
 fun gen_abbrev do_print prep mode (raw_var, raw_prop) lthy =
   let
-    val ((vars, [(_, [prop])]), _) =
-      prep (the_list raw_var) [(("", []), [raw_prop])]
+    val ((vars, [(_, prop)]), _) =
+      prep (the_list raw_var) [(Attrib.empty_binding, raw_prop)]
         (lthy |> ProofContext.set_mode ProofContext.mode_abbrev);
     val ((x, T), rhs) = LocalDefs.abs_def (#2 (LocalDefs.cert_def lthy prop));
     val var =
@@ -223,8 +247,8 @@
     val _ = if not do_print then () else print_consts lthy' (K false) [(x, T)];
   in lthy' end;
 
-val abbreviation = gen_abbrev false check_free_specification;
-val abbreviation_cmd = gen_abbrev true read_free_specification;
+val abbreviation = gen_abbrev false check_free_spec;
+val abbreviation_cmd = gen_abbrev true read_free_spec;
 
 
 (* notation *)
@@ -256,15 +280,12 @@
 
 local
 
-fun subtract_prems ctxt1 ctxt2 =
-  Library.drop (length (Assumption.prems_of ctxt1), Assumption.prems_of ctxt2);
-
 fun prep_statement prep_att prep_stmt elems concl ctxt =
   (case concl of
     Element.Shows shows =>
       let
         val (propp, elems_ctxt) = prep_stmt elems (map snd shows) ctxt;
-        val prems = subtract_prems ctxt elems_ctxt;
+        val prems = Assumption.local_prems_of elems_ctxt ctxt;
         val stmt = Attrib.map_specs prep_att (map fst shows ~~ propp);
         val goal_ctxt = fold (fold (Variable.auto_fixes o fst)) propp elems_ctxt;
       in ((prems, stmt, []), goal_ctxt) end
@@ -300,7 +321,7 @@
 
         val atts = map (Attrib.internal o K)
           [RuleCases.consumes (~ (length obtains)), RuleCases.case_names case_names];
-        val prems = subtract_prems ctxt elems_ctxt;
+        val prems = Assumption.local_prems_of elems_ctxt ctxt;
         val stmt = [((Binding.empty, atts), [(thesis, [])])];
 
         val (facts, goal_ctxt) = elems_ctxt
--- a/src/Pure/Isar/theory_target.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/theory_target.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -52,7 +52,7 @@
     val fixes =
       map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) (#1 (ProofContext.inferred_fixes ctxt));
     val assumes =
-      map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])])) (Assumption.assms_of ctxt);
+      map (fn A => (Attrib.empty_binding, [(Thm.term_of A, [])])) (Assumption.all_assms_of ctxt);
     val elems =
       (if null fixes then [] else [Element.Fixes fixes]) @
       (if null assumes then [] else [Element.Assumes assumes]);
@@ -109,7 +109,7 @@
     val th = Goal.norm_result raw_th;
     val (defs, th') = LocalDefs.export ctxt thy_ctxt th;
     val concl_conv = MetaSimplifier.rewrite true defs (Thm.cprop_of th);
-    val assms = map (MetaSimplifier.rewrite_rule defs o Thm.assume) (Assumption.assms_of ctxt);
+    val assms = map (MetaSimplifier.rewrite_rule defs o Thm.assume) (Assumption.all_assms_of ctxt);
     val nprems = Thm.nprems_of th' - Thm.nprems_of th;
 
     (*export fixes*)
--- a/src/Pure/Isar/toplevel.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Isar/toplevel.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -294,9 +294,10 @@
 
 fun debugging f x =
   if ! debug then
-    (case exception_trace (fn () => SOME (f x) handle UNDEF => NONE) of
-      SOME y => y
-    | NONE => raise UNDEF)
+    Exn.release (exception_trace (fn () =>
+      Exn.Result (f x) handle
+        exn as UNDEF => Exn.Exn exn
+      | exn as EXCURSION_FAIL _ => Exn.Exn exn))
   else f x;
 
 fun toplevel_error f x =
--- a/src/Pure/Tools/find_theorems.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/Tools/find_theorems.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -336,7 +336,7 @@
 
 fun find_theorems ctxt opt_goal rem_dups raw_criteria =
   let
-    val add_prems = Seq.hd o (TRY (Method.insert_tac (Assumption.prems_of ctxt) 1));
+    val add_prems = Seq.hd o (TRY (Method.insert_tac (Assumption.all_prems_of ctxt) 1));
     val opt_goal' = Option.map add_prems opt_goal;
 
     val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
--- a/src/Pure/assumption.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/assumption.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -1,7 +1,7 @@
 (*  Title:      Pure/assumption.ML
     Author:     Makarius
 
-Local assumptions, parameterized by export rules.
+Context assumptions, parameterized by export rules.
 *)
 
 signature ASSUMPTION =
@@ -10,12 +10,13 @@
   val assume_export: export
   val presume_export: export
   val assume: cterm -> thm
-  val assms_of: Proof.context -> cterm list
-  val prems_of: Proof.context -> thm list
+  val all_assms_of: Proof.context -> cterm list
+  val all_prems_of: Proof.context -> thm list
   val extra_hyps: Proof.context -> thm -> term list
+  val local_assms_of: Proof.context -> Proof.context -> cterm list
+  val local_prems_of: Proof.context -> Proof.context -> thm list
   val add_assms: export -> cterm list -> Proof.context -> thm list * Proof.context
   val add_assumes: cterm list -> Proof.context -> thm list * Proof.context
-  val add_view: Proof.context -> cterm list -> Proof.context -> Proof.context
   val export: bool -> Proof.context -> Proof.context -> thm -> thm
   val export_term: Proof.context -> Proof.context -> term -> term
   val export_morphism: Proof.context -> Proof.context -> morphism
@@ -54,7 +55,7 @@
 (** local context data **)
 
 datatype data = Data of
- {assms: (export * cterm list) list,    (*assumes and views: A ==> _*)
+ {assms: (export * cterm list) list,    (*assumes: A ==> _*)
   prems: thm list};                     (*prems: A |- A*)
 
 fun make_data (assms, prems) = Data {assms = assms, prems = prems};
@@ -68,18 +69,31 @@
 fun map_data f = Data.map (fn Data {assms, prems} => make_data (f (assms, prems)));
 fun rep_data ctxt = Data.get ctxt |> (fn Data args => args);
 
-val assumptions_of = #assms o rep_data;
-val assms_of = maps #2 o assumptions_of;
-val prems_of = #prems o rep_data;
+
+(* all assumptions *)
+
+val all_assumptions_of = #assms o rep_data;
+val all_assms_of = maps #2 o all_assumptions_of;
+val all_prems_of = #prems o rep_data;
 
-fun extra_hyps ctxt th = subtract (op aconv) (map Thm.term_of (assms_of ctxt)) (Thm.hyps_of th);
+fun extra_hyps ctxt th =
+  subtract (op aconv) (map Thm.term_of (all_assms_of ctxt)) (Thm.hyps_of th);
+
+(*named prems -- legacy feature*)
+val _ = Context.>>
+  (Context.map_theory (PureThy.add_thms_dynamic (Binding.name "prems",
+    fn Context.Theory _ => [] | Context.Proof ctxt => all_prems_of ctxt)));
 
 
-(* named prems -- legacy feature *)
+(* local assumptions *)
+
+fun local_assumptions_of inner outer =
+  Library.drop (length (all_assumptions_of outer), all_assumptions_of inner);
 
-val _ = Context.>>
-  (Context.map_theory (PureThy.add_thms_dynamic (Binding.name "prems",
-    fn Context.Theory _ => [] | Context.Proof ctxt => prems_of ctxt)));
+val local_assms_of = maps #2 oo local_assumptions_of;
+
+fun local_prems_of inner outer =
+  Library.drop (length (all_prems_of outer), all_prems_of inner);
 
 
 (* add assumptions *)
@@ -92,27 +106,16 @@
 
 val add_assumes = add_assms assume_export;
 
-fun add_view outer view = map_data (fn (asms, prems) =>
-  let
-    val (asms1, asms2) = chop (length (assumptions_of outer)) asms;
-    val asms' = asms1 @ [(assume_export, view)] @ asms2;
-  in (asms', prems) end);
-
 
 (* export *)
 
-fun diff_assms inner outer =
-  Library.drop (length (assumptions_of outer), assumptions_of inner);
-
 fun export is_goal inner outer =
-  let val asms = diff_assms inner outer in
-    MetaSimplifier.norm_hhf_protect
-    #> fold_rev (fn (e, As) => #1 (e is_goal As)) asms
-    #> MetaSimplifier.norm_hhf_protect
-  end;
+  MetaSimplifier.norm_hhf_protect #>
+  fold_rev (fn (e, As) => #1 (e is_goal As)) (local_assumptions_of inner outer) #>
+  MetaSimplifier.norm_hhf_protect;
 
 fun export_term inner outer =
-  fold_rev (fn (e, As) => #2 (e false As)) (diff_assms inner outer);
+  fold_rev (fn (e, As) => #2 (e false As)) (local_assumptions_of inner outer);
 
 fun export_morphism inner outer =
   let
--- a/src/Pure/axclass.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/axclass.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -370,8 +370,7 @@
     val T' = Type.strip_sorts T;
   in
     thy
-    |> Sign.sticky_prefix name_inst
-    |> Sign.no_base_names
+    |> Sign.mandatory_path name_inst
     |> Sign.declare_const [] ((Binding.name c', T'), NoSyn)
     |-> (fn const' as Const (c'', _) =>
       Thm.add_def false true
@@ -481,8 +480,7 @@
     val class_triv = Thm.class_triv def_thy class;
     val ([(_, [intro]), (_, classrel), (_, axioms)], facts_thy) =
       def_thy
-      |> Sign.add_path (Binding.name_of bconst)
-      |> Sign.no_base_names
+      |> Sign.mandatory_path (Binding.name_of bconst)
       |> PureThy.note_thmss ""
         [((Binding.name introN, []), [([Drule.standard raw_intro], [])]),
          ((Binding.name superN, []), [(map Drule.standard raw_classrel, [])]),
--- a/src/Pure/consts.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/consts.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -215,7 +215,7 @@
 fun err_dup_const c =
   error ("Duplicate declaration of constant " ^ quote c);
 
-fun extend_decls naming decl tab = NameSpace.bind naming decl tab
+fun extend_decls naming decl tab = NameSpace.define naming decl tab
   handle Symtab.DUP c => err_dup_const c;
 
 
--- a/src/Pure/goal.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/goal.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -110,7 +110,7 @@
     val cert = Thm.cterm_of thy;
     val certT = Thm.ctyp_of thy;
 
-    val assms = Assumption.assms_of ctxt;
+    val assms = Assumption.all_assms_of ctxt;
     val As = map Thm.term_of assms;
 
     val xs = map Free (fold Term.add_frees (prop :: As) []);
--- a/src/Pure/sign.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/sign.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -124,7 +124,7 @@
   val add_path: string -> theory -> theory
   val root_path: theory -> theory
   val parent_path: theory -> theory
-  val sticky_prefix: string -> theory -> theory
+  val mandatory_path: string -> theory -> theory
   val no_base_names: theory -> theory
   val local_path: theory -> theory
   val restore_naming: theory -> theory -> theory
@@ -619,7 +619,7 @@
 val add_path        = map_naming o NameSpace.add_path;
 val root_path       = map_naming NameSpace.root_path;
 val parent_path     = map_naming NameSpace.parent_path;
-val sticky_prefix   = map_naming o NameSpace.sticky_prefix;
+val mandatory_path  = map_naming o NameSpace.mandatory_path;
 val no_base_names   = map_naming NameSpace.no_base_names;
 
 fun local_path thy = thy |> root_path |> add_path (Context.theory_name thy);
--- a/src/Pure/simplifier.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/simplifier.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -222,7 +222,7 @@
         val simproc' = morph_simproc phi simproc;
       in
         Simprocs.map (fn simprocs =>
-          NameSpace.bind naming (b', simproc') simprocs |> snd
+          NameSpace.define naming (b', simproc') simprocs |> snd
             handle Symtab.DUP dup => err_dup_simproc dup)
         #> map_ss (fn ss => ss addsimprocs [simproc'])
       end)
--- a/src/Pure/theory.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/theory.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -175,7 +175,7 @@
 fun gen_add_axioms prep_axm raw_axms thy = thy |> map_axioms (fn axioms =>
   let
     val axms = map (apsnd Logic.varify o prep_axm thy) raw_axms;
-    val axioms' = fold (snd oo NameSpace.bind (Sign.naming_of thy)) axms axioms
+    val axioms' = fold (snd oo NameSpace.define (Sign.naming_of thy)) axms axioms
       handle Symtab.DUP dup => err_dup_axm dup;
   in axioms' end);
 
--- a/src/Pure/thm.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Pure/thm.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -1700,7 +1700,7 @@
 fun add_oracle (b, oracle) thy =
   let
     val naming = Sign.naming_of thy;
-    val (name, tab') = NameSpace.bind naming (b, serial ()) (Oracles.get thy)
+    val (name, tab') = NameSpace.define naming (b, serial ()) (Oracles.get thy)
       handle Symtab.DUP _ => err_dup_ora (Binding.str_of b);
     val thy' = Oracles.put tab' thy;
   in ((name, invoke_oracle (Theory.check_thy thy') name oracle), thy') end;
--- a/src/Tools/quickcheck.ML	Thu Mar 12 18:01:27 2009 +0100
+++ b/src/Tools/quickcheck.ML	Thu Mar 12 23:01:25 2009 +0100
@@ -144,7 +144,7 @@
 fun test_goal_auto int state =
   let
     val ctxt = Proof.context_of state;
-    val assms = map term_of (Assumption.assms_of ctxt);
+    val assms = map term_of (Assumption.all_assms_of ctxt);
     val Test_Params { size, iterations, default_type } =
       (snd o Data.get o Proof.theory_of) state;
     fun test () =