merge
authorblanchet
Mon, 26 Apr 2010 21:25:32 +0200
changeset 36403 9a4baad039c4
parent 36367 49c7dee21a7f (diff)
parent 36402 1b20805974c7 (current diff)
child 36404 cb3dc64f13d7
merge
src/HOL/IsaMakefile
src/HOL/Tools/ATP_Manager/atp_manager.ML
src/HOL/Tools/ATP_Manager/atp_minimal.ML
src/HOL/Tools/ATP_Manager/atp_wrapper.ML
--- a/NEWS	Mon Apr 26 21:20:43 2010 +0200
+++ b/NEWS	Mon Apr 26 21:25:32 2010 +0200
@@ -6,6 +6,12 @@
 
 *** General ***
 
+* Schematic theorem statements need to be explicitly markup as such,
+via commands 'schematic_lemma', 'schematic_theorem',
+'schematic_corollary'.  Thus the relevance of the proof is made
+syntactically clear, which impacts performance in a parallel or
+asynchronous interactive environment.  Minor INCOMPATIBILITY.
+
 * Authentic syntax for *all* logical entities (type classes, type
 constructors, term constants): provides simple and robust
 correspondence between formal entities and concrete syntax.  Within
@@ -112,6 +118,17 @@
 
 *** HOL ***
 
+* Abstract algebra:
+  * classes division_ring_inverse_zero, field_inverse_zero, linordered_field_inverse_zero
+    include rule inverse 0 = 0 -- subsumes former division_by_zero class.
+  * numerous lemmas have been ported from field to division_ring;
+  * dropped theorem group group_simps, use algebra_simps instead;
+  * dropped theorem group ring_simps, use field_simps instead;
+  * proper theorem collection field_simps subsumes former theorem groups field_eq_simps and field_simps;
+  * dropped lemma eq_minus_self_iff which is a duplicate for equal_neg_zero.
+
+  INCOMPATIBILITY.
+
 * Library theory 'RBT' renamed to 'RBT_Impl'; new library theory 'RBT'
 provides abstract red-black tree type which is backed by RBT_Impl
 as implementation.  INCOMPATIBILTY.
--- a/doc-src/IsarImplementation/Thy/Logic.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/doc-src/IsarImplementation/Thy/Logic.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -128,8 +128,7 @@
   @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\
   @{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\
   @{index_ML Sign.add_types: "(binding * int * mixfix) list -> theory -> theory"} \\
-  @{index_ML Sign.add_tyabbrs_i: "
-  (binding * string list * typ * mixfix) list -> theory -> theory"} \\
+  @{index_ML Sign.add_type_abbrev: "binding * string list * typ -> theory -> theory"} \\
   @{index_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\
   @{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\
   @{index_ML Sign.primitive_arity: "arity -> theory -> theory"} \\
@@ -168,9 +167,9 @@
   type constructors @{text "\<kappa>"} with @{text "k"} arguments and
   optional mixfix syntax.
 
-  \item @{ML Sign.add_tyabbrs_i}~@{text "[(\<kappa>, \<^vec>\<alpha>, \<tau>, mx), \<dots>]"}
-  defines a new type abbreviation @{text "(\<^vec>\<alpha>)\<kappa> = \<tau>"} with
-  optional mixfix syntax.
+  \item @{ML Sign.add_type_abbrev}~@{text "(\<kappa>, \<^vec>\<alpha>,
+  \<tau>)"} defines a new type abbreviation @{text
+  "(\<^vec>\<alpha>)\<kappa> = \<tau>"}.
 
   \item @{ML Sign.primitive_class}~@{text "(c, [c\<^isub>1, \<dots>,
   c\<^isub>n])"} declares a new class @{text "c"}, together with class
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex	Mon Apr 26 21:20:43 2010 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex	Mon Apr 26 21:25:32 2010 +0200
@@ -139,8 +139,7 @@
   \indexdef{}{ML}{Sign.subsort}\verb|Sign.subsort: theory -> sort * sort -> bool| \\
   \indexdef{}{ML}{Sign.of\_sort}\verb|Sign.of_sort: theory -> typ * sort -> bool| \\
   \indexdef{}{ML}{Sign.add\_types}\verb|Sign.add_types: (binding * int * mixfix) list -> theory -> theory| \\
-  \indexdef{}{ML}{Sign.add\_tyabbrs\_i}\verb|Sign.add_tyabbrs_i: |\isasep\isanewline%
-\verb|  (binding * string list * typ * mixfix) list -> theory -> theory| \\
+  \indexdef{}{ML}{Sign.add\_type\_abbrev}\verb|Sign.add_type_abbrev: binding * string list * typ -> theory -> theory| \\
   \indexdef{}{ML}{Sign.primitive\_class}\verb|Sign.primitive_class: binding * class list -> theory -> theory| \\
   \indexdef{}{ML}{Sign.primitive\_classrel}\verb|Sign.primitive_classrel: class * class -> theory -> theory| \\
   \indexdef{}{ML}{Sign.primitive\_arity}\verb|Sign.primitive_arity: arity -> theory -> theory| \\
@@ -176,9 +175,7 @@
   type constructors \isa{{\isasymkappa}} with \isa{k} arguments and
   optional mixfix syntax.
 
-  \item \verb|Sign.add_tyabbrs_i|~\isa{{\isacharbrackleft}{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec {\isasymalpha}{\isacharcomma}\ {\isasymtau}{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isasymdots}{\isacharbrackright}}
-  defines a new type abbreviation \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}} with
-  optional mixfix syntax.
+  \item \verb|Sign.add_type_abbrev|~\isa{{\isacharparenleft}{\isasymkappa}{\isacharcomma}\ \isactrlvec {\isasymalpha}{\isacharcomma}\ {\isasymtau}{\isacharparenright}} defines a new type abbreviation \isa{{\isacharparenleft}\isactrlvec {\isasymalpha}{\isacharparenright}{\isasymkappa}\ {\isacharequal}\ {\isasymtau}}.
 
   \item \verb|Sign.primitive_class|~\isa{{\isacharparenleft}c{\isacharcomma}\ {\isacharbrackleft}c\isactrlisub {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ c\isactrlisub n{\isacharbrackright}{\isacharparenright}} declares a new class \isa{c}, together with class
   relations \isa{c\ {\isasymsubseteq}\ c\isactrlisub i}, for \isa{i\ {\isacharequal}\ {\isadigit{1}}{\isacharcomma}\ {\isasymdots}{\isacharcomma}\ n}.
--- a/doc-src/IsarRef/Thy/Proof.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/doc-src/IsarRef/Thy/Proof.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -366,6 +366,9 @@
     @{command_def "lemma"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
     @{command_def "theorem"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
     @{command_def "corollary"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
+    @{command_def "schematic_lemma"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
+    @{command_def "schematic_theorem"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
+    @{command_def "schematic_corollary"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
     @{command_def "have"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
     @{command_def "show"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
     @{command_def "hence"} & : & @{text "proof(state) \<rightarrow> proof(prove)"} \\
@@ -406,7 +409,8 @@
   and assumptions, cf.\ \secref{sec:obtain}).
 
   \begin{rail}
-    ('lemma' | 'theorem' | 'corollary') target? (goal | longgoal)
+    ('lemma' | 'theorem' | 'corollary' |
+     'schematic\_lemma' | 'schematic\_theorem' | 'schematic\_corollary') target? (goal | longgoal)
     ;
     ('have' | 'show' | 'hence' | 'thus') goal
     ;
@@ -438,6 +442,18 @@
   "lemma"}~@{text "a: \<phi>"}, but the facts are internally marked as
   being of a different kind.  This discrimination acts like a formal
   comment.
+
+  \item @{command "schematic_lemma"}, @{command "schematic_theorem"},
+  @{command "schematic_corollary"} are similar to @{command "lemma"},
+  @{command "theorem"}, @{command "corollary"}, respectively but allow
+  the statement to contain unbound schematic variables.
+
+  Under normal circumstances, an Isar proof text needs to specify
+  claims explicitly.  Schematic goals are more like goals in Prolog,
+  where certain results are synthesized in the course of reasoning.
+  With schematic statements, the inherent compositionality of Isar
+  proofs is lost, which also impacts performance, because proof
+  checking is forced into sequential mode.
   
   \item @{command "have"}~@{text "a: \<phi>"} claims a local goal,
   eventually resulting in a fact within the current logical context.
@@ -487,26 +503,6 @@
   context introductions, (2) the resulting rule is annotated
   accordingly to support symbolic case splits when used with the
   @{method_ref cases} method (cf.\ \secref{sec:cases-induct}).
-
-  \medskip
-
-  \begin{warn}
-    Isabelle/Isar suffers theory-level goal statements to contain
-    \emph{unbound schematic variables}, although this does not conform
-    to the aim of human-readable proof documents!  The main problem
-    with schematic goals is that the actual outcome is usually hard to
-    predict, depending on the behavior of the proof methods applied
-    during the course of reasoning.  Note that most semi-automated
-    methods heavily depend on several kinds of implicit rule
-    declarations within the current theory context.  As this would
-    also result in non-compositional checking of sub-proofs,
-    \emph{local goals} are not allowed to be schematic at all.
-    Nevertheless, schematic goals do have their use in Prolog-style
-    interactive synthesis of proven results, usually by stepwise
-    refinement via emulation of traditional Isabelle tactic scripts
-    (see also \secref{sec:tactic-commands}).  In any case, users
-    should know what they are doing.
-  \end{warn}
 *}
 
 
--- a/doc-src/IsarRef/Thy/document/Proof.tex	Mon Apr 26 21:20:43 2010 +0200
+++ b/doc-src/IsarRef/Thy/document/Proof.tex	Mon Apr 26 21:25:32 2010 +0200
@@ -385,6 +385,9 @@
     \indexdef{}{command}{lemma}\hypertarget{command.lemma}{\hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
     \indexdef{}{command}{theorem}\hypertarget{command.theorem}{\hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
     \indexdef{}{command}{corollary}\hypertarget{command.corollary}{\hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
+    \indexdef{}{command}{schematic\_lemma}\hypertarget{command.schematic-lemma}{\hyperlink{command.schematic-lemma}{\mbox{\isa{\isacommand{schematic{\isacharunderscore}lemma}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
+    \indexdef{}{command}{schematic\_theorem}\hypertarget{command.schematic-theorem}{\hyperlink{command.schematic-theorem}{\mbox{\isa{\isacommand{schematic{\isacharunderscore}theorem}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
+    \indexdef{}{command}{schematic\_corollary}\hypertarget{command.schematic-corollary}{\hyperlink{command.schematic-corollary}{\mbox{\isa{\isacommand{schematic{\isacharunderscore}corollary}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
     \indexdef{}{command}{have}\hypertarget{command.have}{\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}} & : & \isa{{\isachardoublequote}proof{\isacharparenleft}state{\isacharparenright}\ {\isacharbar}\ proof{\isacharparenleft}chain{\isacharparenright}\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
     \indexdef{}{command}{show}\hypertarget{command.show}{\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}} & : & \isa{{\isachardoublequote}proof{\isacharparenleft}state{\isacharparenright}\ {\isacharbar}\ proof{\isacharparenleft}chain{\isacharparenright}\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
     \indexdef{}{command}{hence}\hypertarget{command.hence}{\hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}}} & : & \isa{{\isachardoublequote}proof{\isacharparenleft}state{\isacharparenright}\ {\isasymrightarrow}\ proof{\isacharparenleft}prove{\isacharparenright}{\isachardoublequote}} \\
@@ -425,7 +428,8 @@
   and assumptions, cf.\ \secref{sec:obtain}).
 
   \begin{rail}
-    ('lemma' | 'theorem' | 'corollary') target? (goal | longgoal)
+    ('lemma' | 'theorem' | 'corollary' |
+     'schematic\_lemma' | 'schematic\_theorem' | 'schematic\_corollary') target? (goal | longgoal)
     ;
     ('have' | 'show' | 'hence' | 'thus') goal
     ;
@@ -454,6 +458,18 @@
   \item \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}~\isa{{\isachardoublequote}a{\isacharcolon}\ {\isasymphi}{\isachardoublequote}} and \hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}~\isa{{\isachardoublequote}a{\isacharcolon}\ {\isasymphi}{\isachardoublequote}} are essentially the same as \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}~\isa{{\isachardoublequote}a{\isacharcolon}\ {\isasymphi}{\isachardoublequote}}, but the facts are internally marked as
   being of a different kind.  This discrimination acts like a formal
   comment.
+
+  \item \hyperlink{command.schematic-lemma}{\mbox{\isa{\isacommand{schematic{\isacharunderscore}lemma}}}}, \hyperlink{command.schematic-theorem}{\mbox{\isa{\isacommand{schematic{\isacharunderscore}theorem}}}},
+  \hyperlink{command.schematic-corollary}{\mbox{\isa{\isacommand{schematic{\isacharunderscore}corollary}}}} are similar to \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}},
+  \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}, \hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}, respectively but allow
+  the statement to contain unbound schematic variables.
+
+  Under normal circumstances, an Isar proof text needs to specify
+  claims explicitly.  Schematic goals are more like goals in Prolog,
+  where certain results are synthesized in the course of reasoning.
+  With schematic statements, the inherent compositionality of Isar
+  proofs is lost, which also impacts performance, because proof
+  checking is forced into sequential mode.
   
   \item \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isachardoublequote}a{\isacharcolon}\ {\isasymphi}{\isachardoublequote}} claims a local goal,
   eventually resulting in a fact within the current logical context.
@@ -499,27 +515,7 @@
   meaning: (1) during the of this claim they refer to the the local
   context introductions, (2) the resulting rule is annotated
   accordingly to support symbolic case splits when used with the
-  \indexref{}{method}{cases}\hyperlink{method.cases}{\mbox{\isa{cases}}} method (cf.\ \secref{sec:cases-induct}).
-
-  \medskip
-
-  \begin{warn}
-    Isabelle/Isar suffers theory-level goal statements to contain
-    \emph{unbound schematic variables}, although this does not conform
-    to the aim of human-readable proof documents!  The main problem
-    with schematic goals is that the actual outcome is usually hard to
-    predict, depending on the behavior of the proof methods applied
-    during the course of reasoning.  Note that most semi-automated
-    methods heavily depend on several kinds of implicit rule
-    declarations within the current theory context.  As this would
-    also result in non-compositional checking of sub-proofs,
-    \emph{local goals} are not allowed to be schematic at all.
-    Nevertheless, schematic goals do have their use in Prolog-style
-    interactive synthesis of proven results, usually by stepwise
-    refinement via emulation of traditional Isabelle tactic scripts
-    (see also \secref{sec:tactic-commands}).  In any case, users
-    should know what they are doing.
-  \end{warn}%
+  \indexref{}{method}{cases}\hyperlink{method.cases}{\mbox{\isa{cases}}} method (cf.\ \secref{sec:cases-induct}).%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/etc/isar-keywords-ZF.el	Mon Apr 26 21:20:43 2010 +0200
+++ b/etc/isar-keywords-ZF.el	Mon Apr 26 21:25:32 2010 +0200
@@ -162,6 +162,9 @@
     "realizers"
     "remove_thy"
     "rep_datatype"
+    "schematic_corollary"
+    "schematic_lemma"
+    "schematic_theorem"
     "sect"
     "section"
     "setup"
@@ -425,6 +428,9 @@
     "instance"
     "interpretation"
     "lemma"
+    "schematic_corollary"
+    "schematic_lemma"
+    "schematic_theorem"
     "subclass"
     "sublocale"
     "theorem"))
--- a/etc/isar-keywords.el	Mon Apr 26 21:20:43 2010 +0200
+++ b/etc/isar-keywords.el	Mon Apr 26 21:25:32 2010 +0200
@@ -220,6 +220,9 @@
     "remove_thy"
     "rep_datatype"
     "repdef"
+    "schematic_corollary"
+    "schematic_lemma"
+    "schematic_theorem"
     "sect"
     "section"
     "setup"
@@ -570,6 +573,9 @@
     "quotient_type"
     "recdef_tc"
     "rep_datatype"
+    "schematic_corollary"
+    "schematic_lemma"
+    "schematic_theorem"
     "specification"
     "subclass"
     "sublocale"
--- a/lib/scripts/keywords	Mon Apr 26 21:20:43 2010 +0200
+++ b/lib/scripts/keywords	Mon Apr 26 21:25:32 2010 +0200
@@ -39,6 +39,7 @@
     elsif (m/^Outer syntax command:\s*"(.*)"\s*\((.*)\)/) {
       my $name = $1;
       my $kind = $2;
+      if ($kind eq "theory-schematic-goal") { $kind = "theory-goal"; }
       &set_keyword($name, $kind);
     }
   }
--- a/src/CCL/Fix.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/CCL/Fix.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -98,7 +98,7 @@
 
 (* All fixed points are lam-expressions *)
 
-lemma idgenfp_lam: "idgen(d) = d ==> d = lam x. ?f(x)"
+schematic_lemma idgenfp_lam: "idgen(d) = d ==> d = lam x. ?f(x)"
   apply (unfold idgen_def)
   apply (erule ssubst)
   apply (rule refl)
@@ -130,7 +130,7 @@
   apply simp
   done
 
-lemma po_eta_lemma: "idgen(d) = d ==> d = lam x. ?f(x)"
+schematic_lemma po_eta_lemma: "idgen(d) = d ==> d = lam x. ?f(x)"
   apply (unfold idgen_def)
   apply (erule sym)
   done
--- a/src/CCL/Wfd.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/CCL/Wfd.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -573,29 +573,29 @@
 
 subsection {* Factorial *}
 
-lemma
+schematic_lemma
   "letrec f n be ncase(n,succ(zero),%x. nrec(n,zero,%y g. nrec(f(x),g,%z h. succ(h))))  
    in f(succ(succ(zero))) ---> ?a"
   by eval
 
-lemma
+schematic_lemma
   "letrec f n be ncase(n,succ(zero),%x. nrec(n,zero,%y g. nrec(f(x),g,%z h. succ(h))))  
    in f(succ(succ(succ(zero)))) ---> ?a"
   by eval
 
 subsection {* Less Than Or Equal *}
 
-lemma
+schematic_lemma
   "letrec f p be split(p,%m n. ncase(m,true,%x. ncase(n,false,%y. f(<x,y>))))
    in f(<succ(zero), succ(zero)>) ---> ?a"
   by eval
 
-lemma
+schematic_lemma
   "letrec f p be split(p,%m n. ncase(m,true,%x. ncase(n,false,%y. f(<x,y>))))
    in f(<succ(zero), succ(succ(succ(succ(zero))))>) ---> ?a"
   by eval
 
-lemma
+schematic_lemma
   "letrec f p be split(p,%m n. ncase(m,true,%x. ncase(n,false,%y. f(<x,y>))))
    in f(<succ(succ(succ(succ(succ(zero))))), succ(succ(succ(succ(zero))))>) ---> ?a"
   by eval
@@ -603,12 +603,12 @@
 
 subsection {* Reverse *}
 
-lemma
+schematic_lemma
   "letrec id l be lcase(l,[],%x xs. x$id(xs))  
    in id(zero$succ(zero)$[]) ---> ?a"
   by eval
 
-lemma
+schematic_lemma
   "letrec rev l be lcase(l,[],%x xs. lrec(rev(xs),x$[],%y ys g. y$g))  
    in rev(zero$succ(zero)$(succ((lam x. x)`succ(zero)))$([])) ---> ?a"
   by eval
--- a/src/CTT/Arith.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/CTT/Arith.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -256,7 +256,8 @@
 (*Addition is the inverse of subtraction: if b<=x then b#+(x-b) = x.
   An example of induction over a quantified formula (a product).
   Uses rewriting with a quantified, implicative inductive hypothesis.*)
-lemma add_diff_inverse_lemma: "b:N ==> ?a : PROD x:N. Eq(N, b-x, 0) --> Eq(N, b #+ (x-b), x)"
+schematic_lemma add_diff_inverse_lemma:
+  "b:N ==> ?a : PROD x:N. Eq(N, b-x, 0) --> Eq(N, b #+ (x-b), x)"
 apply (tactic {* NE_tac @{context} "b" 1 *})
 (*strip one "universal quantifier" but not the "implication"*)
 apply (rule_tac [3] intr_rls)
@@ -324,7 +325,7 @@
 done
 
 (*If a+b=0 then a=0.   Surprisingly tedious*)
-lemma add_eq0_lemma: "[| a:N;  b:N |] ==> ?c : PROD u: Eq(N,a#+b,0) .  Eq(N,a,0)"
+schematic_lemma add_eq0_lemma: "[| a:N;  b:N |] ==> ?c : PROD u: Eq(N,a#+b,0) .  Eq(N,a,0)"
 apply (tactic {* NE_tac @{context} "a" 1 *})
 apply (rule_tac [3] replace_type)
 apply (tactic "arith_rew_tac []")
@@ -344,7 +345,7 @@
 done
 
 (*Here is a lemma to infer a-b=0 and b-a=0 from a|-|b=0, below. *)
-lemma absdiff_eq0_lem:
+schematic_lemma absdiff_eq0_lem:
     "[| a:N;  b:N;  a |-| b = 0 : N |] ==>
      ?a : SUM v: Eq(N, a-b, 0) . Eq(N, b-a, 0)"
 apply (unfold absdiff_def)
--- a/src/CTT/ex/Elimination.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/CTT/ex/Elimination.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -14,41 +14,41 @@
 
 text "This finds the functions fst and snd!"
 
-lemma [folded basic_defs]: "A type ==> ?a : (A*A) --> A"
+schematic_lemma [folded basic_defs]: "A type ==> ?a : (A*A) --> A"
 apply (tactic {* pc_tac [] 1 *})
 done
 
-lemma [folded basic_defs]: "A type ==> ?a : (A*A) --> A"
+schematic_lemma [folded basic_defs]: "A type ==> ?a : (A*A) --> A"
 apply (tactic {* pc_tac [] 1 *})
 back
 done
 
 text "Double negation of the Excluded Middle"
-lemma "A type ==> ?a : ((A + (A-->F)) --> F) --> F"
+schematic_lemma "A type ==> ?a : ((A + (A-->F)) --> F) --> F"
 apply (tactic "intr_tac []")
 apply (rule ProdE)
 apply assumption
 apply (tactic "pc_tac [] 1")
 done
 
-lemma "[| A type;  B type |] ==> ?a : (A*B) --> (B*A)"
+schematic_lemma "[| A type;  B type |] ==> ?a : (A*B) --> (B*A)"
 apply (tactic "pc_tac [] 1")
 done
 (*The sequent version (ITT) could produce an interesting alternative
   by backtracking.  No longer.*)
 
 text "Binary sums and products"
-lemma "[| A type; B type; C type |] ==> ?a : (A+B --> C) --> (A-->C) * (B-->C)"
+schematic_lemma "[| A type; B type; C type |] ==> ?a : (A+B --> C) --> (A-->C) * (B-->C)"
 apply (tactic "pc_tac [] 1")
 done
 
 (*A distributive law*)
-lemma "[| A type;  B type;  C type |] ==> ?a : A * (B+C)  -->  (A*B + A*C)"
+schematic_lemma "[| A type;  B type;  C type |] ==> ?a : A * (B+C)  -->  (A*B + A*C)"
 apply (tactic "pc_tac [] 1")
 done
 
 (*more general version, same proof*)
-lemma
+schematic_lemma
   assumes "A type"
     and "!!x. x:A ==> B(x) type"
     and "!!x. x:A ==> C(x) type"
@@ -57,12 +57,12 @@
 done
 
 text "Construction of the currying functional"
-lemma "[| A type;  B type;  C type |] ==> ?a : (A*B --> C) --> (A--> (B-->C))"
+schematic_lemma "[| A type;  B type;  C type |] ==> ?a : (A*B --> C) --> (A--> (B-->C))"
 apply (tactic "pc_tac [] 1")
 done
 
 (*more general goal with same proof*)
-lemma
+schematic_lemma
   assumes "A type"
     and "!!x. x:A ==> B(x) type"
     and "!!z. z: (SUM x:A. B(x)) ==> C(z) type"
@@ -72,12 +72,12 @@
 done
 
 text "Martin-Lof (1984), page 48: axiom of sum-elimination (uncurry)"
-lemma "[| A type;  B type;  C type |] ==> ?a : (A --> (B-->C)) --> (A*B --> C)"
+schematic_lemma "[| A type;  B type;  C type |] ==> ?a : (A --> (B-->C)) --> (A*B --> C)"
 apply (tactic "pc_tac [] 1")
 done
 
 (*more general goal with same proof*)
-lemma
+schematic_lemma
   assumes "A type"
     and "!!x. x:A ==> B(x) type"
     and "!!z. z: (SUM x:A . B(x)) ==> C(z) type"
@@ -87,12 +87,12 @@
 done
 
 text "Function application"
-lemma "[| A type;  B type |] ==> ?a : ((A --> B) * A) --> B"
+schematic_lemma "[| A type;  B type |] ==> ?a : ((A --> B) * A) --> B"
 apply (tactic "pc_tac [] 1")
 done
 
 text "Basic test of quantifier reasoning"
-lemma
+schematic_lemma
   assumes "A type"
     and "B type"
     and "!!x y.[| x:A;  y:B |] ==> C(x,y) type"
@@ -103,7 +103,7 @@
 done
 
 text "Martin-Lof (1984) pages 36-7: the combinator S"
-lemma
+schematic_lemma
   assumes "A type"
     and "!!x. x:A ==> B(x) type"
     and "!!x y.[| x:A; y:B(x) |] ==> C(x,y) type"
@@ -113,7 +113,7 @@
 done
 
 text "Martin-Lof (1984) page 58: the axiom of disjunction elimination"
-lemma
+schematic_lemma
   assumes "A type"
     and "B type"
     and "!!z. z: A+B ==> C(z) type"
@@ -123,7 +123,7 @@
 done
 
 (*towards AXIOM OF CHOICE*)
-lemma [folded basic_defs]:
+schematic_lemma [folded basic_defs]:
   "[| A type; B type; C type |] ==> ?a : (A --> B*C) --> (A-->B) * (A-->C)"
 apply (tactic "pc_tac [] 1")
 done
@@ -131,7 +131,7 @@
 
 (*Martin-Lof (1984) page 50*)
 text "AXIOM OF CHOICE!  Delicate use of elimination rules"
-lemma
+schematic_lemma
   assumes "A type"
     and "!!x. x:A ==> B(x) type"
     and "!!x y.[| x:A;  y:B(x) |] ==> C(x,y) type"
@@ -149,7 +149,7 @@
 done
 
 text "Axiom of choice.  Proof without fst, snd.  Harder still!"
-lemma [folded basic_defs]:
+schematic_lemma [folded basic_defs]:
   assumes "A type"
     and "!!x. x:A ==> B(x) type"
     and "!!x y.[| x:A;  y:B(x) |] ==> C(x,y) type"
@@ -174,7 +174,7 @@
 text "Example of sequent_style deduction"
 (*When splitting z:A*B, the assumption C(z) is affected;  ?a becomes
     lam u. split(u,%v w.split(v,%x y.lam z. <x,<y,z>>) ` w)     *)
-lemma
+schematic_lemma
   assumes "A type"
     and "B type"
     and "!!z. z:A*B ==> C(z) type"
--- a/src/CTT/ex/Synthesis.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/CTT/ex/Synthesis.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -10,7 +10,7 @@
 begin
 
 text "discovery of predecessor function"
-lemma "?a : SUM pred:?A . Eq(N, pred`0, 0)
+schematic_lemma "?a : SUM pred:?A . Eq(N, pred`0, 0)
                   *  (PROD n:N. Eq(N, pred ` succ(n), n))"
 apply (tactic "intr_tac []")
 apply (tactic eqintr_tac)
@@ -20,7 +20,7 @@
 done
 
 text "the function fst as an element of a function type"
-lemma [folded basic_defs]:
+schematic_lemma [folded basic_defs]:
   "A type ==> ?a: SUM f:?B . PROD i:A. PROD j:A. Eq(A, f ` <i,j>, i)"
 apply (tactic "intr_tac []")
 apply (tactic eqintr_tac)
@@ -34,7 +34,7 @@
 text "An interesting use of the eliminator, when"
 (*The early implementation of unification caused non-rigid path in occur check
   See following example.*)
-lemma "?a : PROD i:N. Eq(?A, ?b(inl(i)), <0    ,   i>)
+schematic_lemma "?a : PROD i:N. Eq(?A, ?b(inl(i)), <0    ,   i>)
                    * Eq(?A, ?b(inr(i)), <succ(0), i>)"
 apply (tactic "intr_tac []")
 apply (tactic eqintr_tac)
@@ -46,13 +46,13 @@
  This prevents the cycle in the first unification (no longer needed).
  Requires flex-flex to preserve the dependence.
  Simpler still: make ?A into a constant type N*N.*)
-lemma "?a : PROD i:N. Eq(?A(i), ?b(inl(i)), <0   ,   i>)
+schematic_lemma "?a : PROD i:N. Eq(?A(i), ?b(inl(i)), <0   ,   i>)
                   *  Eq(?A(i), ?b(inr(i)), <succ(0),i>)"
 oops
 
 text "A tricky combination of when and split"
 (*Now handled easily, but caused great problems once*)
-lemma [folded basic_defs]:
+schematic_lemma [folded basic_defs]:
   "?a : PROD i:N. PROD j:N. Eq(?A, ?b(inl(<i,j>)), i)
                            *  Eq(?A, ?b(inr(<i,j>)), j)"
 apply (tactic "intr_tac []")
@@ -65,18 +65,18 @@
 done
 
 (*similar but allows the type to depend on i and j*)
-lemma "?a : PROD i:N. PROD j:N. Eq(?A(i,j), ?b(inl(<i,j>)), i)
+schematic_lemma "?a : PROD i:N. PROD j:N. Eq(?A(i,j), ?b(inl(<i,j>)), i)
                           *   Eq(?A(i,j), ?b(inr(<i,j>)), j)"
 oops
 
 (*similar but specifying the type N simplifies the unification problems*)
-lemma "?a : PROD i:N. PROD j:N. Eq(N, ?b(inl(<i,j>)), i)
+schematic_lemma "?a : PROD i:N. PROD j:N. Eq(N, ?b(inl(<i,j>)), i)
                           *   Eq(N, ?b(inr(<i,j>)), j)"
 oops
 
 
 text "Deriving the addition operator"
-lemma [folded arith_defs]:
+schematic_lemma [folded arith_defs]:
   "?c : PROD n:N. Eq(N, ?f(0,n), n)
                   *  (PROD m:N. Eq(N, ?f(succ(m), n), succ(?f(m,n))))"
 apply (tactic "intr_tac []")
@@ -86,7 +86,7 @@
 done
 
 text "The addition function -- using explicit lambdas"
-lemma [folded arith_defs]:
+schematic_lemma [folded arith_defs]:
   "?c : SUM plus : ?A .
          PROD x:N. Eq(N, plus`0`x, x)
                 *  (PROD y:N. Eq(N, plus`succ(y)`x, succ(plus`y`x)))"
--- a/src/CTT/ex/Typechecking.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/CTT/ex/Typechecking.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -11,18 +11,18 @@
 
 subsection {* Single-step proofs: verifying that a type is well-formed *}
 
-lemma "?A type"
+schematic_lemma "?A type"
 apply (rule form_rls)
 done
 
-lemma "?A type"
+schematic_lemma "?A type"
 apply (rule form_rls)
 back
 apply (rule form_rls)
 apply (rule form_rls)
 done
 
-lemma "PROD z:?A . N + ?B(z) type"
+schematic_lemma "PROD z:?A . N + ?B(z) type"
 apply (rule form_rls)
 apply (rule form_rls)
 apply (rule form_rls)
@@ -37,30 +37,30 @@
 apply (tactic form_tac)
 done
 
-lemma "<0, succ(0)> : ?A"
+schematic_lemma "<0, succ(0)> : ?A"
 apply (tactic "intr_tac []")
 done
 
-lemma "PROD w:N . Eq(?A,w,w) type"
+schematic_lemma "PROD w:N . Eq(?A,w,w) type"
 apply (tactic "typechk_tac []")
 done
 
-lemma "PROD x:N . PROD y:N . Eq(?A,x,y) type"
+schematic_lemma "PROD x:N . PROD y:N . Eq(?A,x,y) type"
 apply (tactic "typechk_tac []")
 done
 
 text "typechecking an application of fst"
-lemma "(lam u. split(u, %v w. v)) ` <0, succ(0)> : ?A"
+schematic_lemma "(lam u. split(u, %v w. v)) ` <0, succ(0)> : ?A"
 apply (tactic "typechk_tac []")
 done
 
 text "typechecking the predecessor function"
-lemma "lam n. rec(n, 0, %x y. x) : ?A"
+schematic_lemma "lam n. rec(n, 0, %x y. x) : ?A"
 apply (tactic "typechk_tac []")
 done
 
 text "typechecking the addition function"
-lemma "lam n. lam m. rec(n, m, %x y. succ(y)) : ?A"
+schematic_lemma "lam n. lam m. rec(n, m, %x y. succ(y)) : ?A"
 apply (tactic "typechk_tac []")
 done
 
@@ -68,18 +68,18 @@
   For concreteness, every type variable left over is forced to be N*)
 ML {* val N_tac = TRYALL (rtac (thm "NF")) *}
 
-lemma "lam w. <w,w> : ?A"
+schematic_lemma "lam w. <w,w> : ?A"
 apply (tactic "typechk_tac []")
 apply (tactic N_tac)
 done
 
-lemma "lam x. lam y. x : ?A"
+schematic_lemma "lam x. lam y. x : ?A"
 apply (tactic "typechk_tac []")
 apply (tactic N_tac)
 done
 
 text "typechecking fst (as a function object)"
-lemma "lam i. split(i, %j k. j) : ?A"
+schematic_lemma "lam i. split(i, %j k. j) : ?A"
 apply (tactic "typechk_tac []")
 apply (tactic N_tac)
 done
--- a/src/Cube/Example.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Cube/Example.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -30,98 +30,98 @@
 
 subsection {* Simple types *}
 
-lemma "A:* |- A->A : ?T"
+schematic_lemma "A:* |- A->A : ?T"
   by (depth_solve rules)
 
-lemma "A:* |- Lam a:A. a : ?T"
+schematic_lemma "A:* |- Lam a:A. a : ?T"
   by (depth_solve rules)
 
-lemma "A:* B:* b:B |- Lam x:A. b : ?T"
+schematic_lemma "A:* B:* b:B |- Lam x:A. b : ?T"
   by (depth_solve rules)
 
-lemma "A:* b:A |- (Lam a:A. a)^b: ?T"
+schematic_lemma "A:* b:A |- (Lam a:A. a)^b: ?T"
   by (depth_solve rules)
 
-lemma "A:* B:* c:A b:B |- (Lam x:A. b)^ c: ?T"
+schematic_lemma "A:* B:* c:A b:B |- (Lam x:A. b)^ c: ?T"
   by (depth_solve rules)
 
-lemma "A:* B:* |- Lam a:A. Lam b:B. a : ?T"
+schematic_lemma "A:* B:* |- Lam a:A. Lam b:B. a : ?T"
   by (depth_solve rules)
 
 
 subsection {* Second-order types *}
 
-lemma (in L2) "|- Lam A:*. Lam a:A. a : ?T"
+schematic_lemma (in L2) "|- Lam A:*. Lam a:A. a : ?T"
   by (depth_solve rules)
 
-lemma (in L2) "A:* |- (Lam B:*.Lam b:B. b)^A : ?T"
+schematic_lemma (in L2) "A:* |- (Lam B:*.Lam b:B. b)^A : ?T"
   by (depth_solve rules)
 
-lemma (in L2) "A:* b:A |- (Lam B:*.Lam b:B. b) ^ A ^ b: ?T"
+schematic_lemma (in L2) "A:* b:A |- (Lam B:*.Lam b:B. b) ^ A ^ b: ?T"
   by (depth_solve rules)
 
-lemma (in L2) "|- Lam B:*.Lam a:(Pi A:*.A).a ^ ((Pi A:*.A)->B) ^ a: ?T"
+schematic_lemma (in L2) "|- Lam B:*.Lam a:(Pi A:*.A).a ^ ((Pi A:*.A)->B) ^ a: ?T"
   by (depth_solve rules)
 
 
 subsection {* Weakly higher-order propositional logic *}
 
-lemma (in Lomega) "|- Lam A:*.A->A : ?T"
+schematic_lemma (in Lomega) "|- Lam A:*.A->A : ?T"
   by (depth_solve rules)
 
-lemma (in Lomega) "B:* |- (Lam A:*.A->A) ^ B : ?T"
+schematic_lemma (in Lomega) "B:* |- (Lam A:*.A->A) ^ B : ?T"
   by (depth_solve rules)
 
-lemma (in Lomega) "B:* b:B |- (Lam y:B. b): ?T"
+schematic_lemma (in Lomega) "B:* b:B |- (Lam y:B. b): ?T"
   by (depth_solve rules)
 
-lemma (in Lomega) "A:* F:*->* |- F^(F^A): ?T"
+schematic_lemma (in Lomega) "A:* F:*->* |- F^(F^A): ?T"
   by (depth_solve rules)
 
-lemma (in Lomega) "A:* |- Lam F:*->*.F^(F^A): ?T"
+schematic_lemma (in Lomega) "A:* |- Lam F:*->*.F^(F^A): ?T"
   by (depth_solve rules)
 
 
 subsection {* LP *}
 
-lemma (in LP) "A:* |- A -> * : ?T"
+schematic_lemma (in LP) "A:* |- A -> * : ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->* a:A |- P^a: ?T"
+schematic_lemma (in LP) "A:* P:A->* a:A |- P^a: ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->A->* a:A |- Pi a:A. P^a^a: ?T"
+schematic_lemma (in LP) "A:* P:A->A->* a:A |- Pi a:A. P^a^a: ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->* Q:A->* |- Pi a:A. P^a -> Q^a: ?T"
+schematic_lemma (in LP) "A:* P:A->* Q:A->* |- Pi a:A. P^a -> Q^a: ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->* |- Pi a:A. P^a -> P^a: ?T"
+schematic_lemma (in LP) "A:* P:A->* |- Pi a:A. P^a -> P^a: ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->* |- Lam a:A. Lam x:P^a. x: ?T"
+schematic_lemma (in LP) "A:* P:A->* |- Lam a:A. Lam x:P^a. x: ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->* Q:* |- (Pi a:A. P^a->Q) -> (Pi a:A. P^a) -> Q : ?T"
+schematic_lemma (in LP) "A:* P:A->* Q:* |- (Pi a:A. P^a->Q) -> (Pi a:A. P^a) -> Q : ?T"
   by (depth_solve rules)
 
-lemma (in LP) "A:* P:A->* Q:* a0:A |-
+schematic_lemma (in LP) "A:* P:A->* Q:* a0:A |-
         Lam x:Pi a:A. P^a->Q. Lam y:Pi a:A. P^a. x^a0^(y^a0): ?T"
   by (depth_solve rules)
 
 
 subsection {* Omega-order types *}
 
-lemma (in L2) "A:* B:* |- Pi C:*.(A->B->C)->C : ?T"
+schematic_lemma (in L2) "A:* B:* |- Pi C:*.(A->B->C)->C : ?T"
   by (depth_solve rules)
 
-lemma (in Lomega2) "|- Lam A:*.Lam B:*.Pi C:*.(A->B->C)->C : ?T"
+schematic_lemma (in Lomega2) "|- Lam A:*.Lam B:*.Pi C:*.(A->B->C)->C : ?T"
   by (depth_solve rules)
 
-lemma (in Lomega2) "|- Lam A:*.Lam B:*.Lam x:A. Lam y:B. x : ?T"
+schematic_lemma (in Lomega2) "|- Lam A:*.Lam B:*.Lam x:A. Lam y:B. x : ?T"
   by (depth_solve rules)
 
-lemma (in Lomega2) "A:* B:* |- ?p : (A->B) -> ((B->Pi P:*.P)->(A->Pi P:*.P))"
+schematic_lemma (in Lomega2) "A:* B:* |- ?p : (A->B) -> ((B->Pi P:*.P)->(A->Pi P:*.P))"
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
@@ -145,14 +145,14 @@
 
 subsection {* Second-order Predicate Logic *}
 
-lemma (in LP2) "A:* P:A->* |- Lam a:A. P^a->(Pi A:*.A) : ?T"
+schematic_lemma (in LP2) "A:* P:A->* |- Lam a:A. P^a->(Pi A:*.A) : ?T"
   by (depth_solve rules)
 
-lemma (in LP2) "A:* P:A->A->* |-
+schematic_lemma (in LP2) "A:* P:A->A->* |-
     (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P : ?T"
   by (depth_solve rules)
 
-lemma (in LP2) "A:* P:A->A->* |-
+schematic_lemma (in LP2) "A:* P:A->A->* |-
     ?p: (Pi a:A. Pi b:A. P^a^b->P^b^a->Pi P:*.P) -> Pi a:A. P^a^a->Pi P:*.P"
   -- {* Antisymmetry implies irreflexivity: *}
   apply (strip_asms rules)
@@ -174,22 +174,22 @@
 
 subsection {* LPomega *}
 
-lemma (in LPomega) "A:* |- Lam P:A->A->*.Lam a:A. P^a^a : ?T"
+schematic_lemma (in LPomega) "A:* |- Lam P:A->A->*.Lam a:A. P^a^a : ?T"
   by (depth_solve rules)
 
-lemma (in LPomega) "|- Lam A:*.Lam P:A->A->*.Lam a:A. P^a^a : ?T"
+schematic_lemma (in LPomega) "|- Lam A:*.Lam P:A->A->*.Lam a:A. P^a^a : ?T"
   by (depth_solve rules)
 
 
 subsection {* Constructions *}
 
-lemma (in CC) "|- Lam A:*.Lam P:A->*.Lam a:A. P^a->Pi P:*.P: ?T"
+schematic_lemma (in CC) "|- Lam A:*.Lam P:A->*.Lam a:A. P^a->Pi P:*.P: ?T"
   by (depth_solve rules)
 
-lemma (in CC) "|- Lam A:*.Lam P:A->*.Pi a:A. P^a: ?T"
+schematic_lemma (in CC) "|- Lam A:*.Lam P:A->*.Pi a:A. P^a: ?T"
   by (depth_solve rules)
 
-lemma (in CC) "A:* P:A->* a:A |- ?p : (Pi a:A. P^a)->P^a"
+schematic_lemma (in CC) "A:* P:A->* a:A |- ?p : (Pi a:A. P^a)->P^a"
   apply (strip_asms rules)
   apply (rule lam_ss)
     apply (depth_solve1 rules)
@@ -201,15 +201,15 @@
 
 subsection {* Some random examples *}
 
-lemma (in LP2) "A:* c:A f:A->A |-
+schematic_lemma (in LP2) "A:* c:A f:A->A |-
     Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
   by (depth_solve rules)
 
-lemma (in CC) "Lam A:*.Lam c:A. Lam f:A->A.
+schematic_lemma (in CC) "Lam A:*.Lam c:A. Lam f:A->A.
     Lam a:A. Pi P:A->*.P^c -> (Pi x:A. P^x->P^(f^x)) -> P^a : ?T"
   by (depth_solve rules)
 
-lemma (in LP2)
+schematic_lemma (in LP2)
   "A:* a:A b:A |- ?p: (Pi P:A->*.P^a->P^b) -> (Pi P:A->*.P^b->P^a)"
   -- {* Symmetry of Leibnitz equality *}
   apply (strip_asms rules)
--- a/src/FOL/ex/Classical.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOL/ex/Classical.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -363,7 +363,7 @@
 
 text{*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
   fast DISCOVERS who killed Agatha. *}
-lemma "lives(agatha) & lives(butler) & lives(charles) &  
+schematic_lemma "lives(agatha) & lives(butler) & lives(charles) &  
    (killed(agatha,agatha) | killed(butler,agatha) | killed(charles,agatha)) &  
    (\<forall>x y. killed(x,y) --> hates(x,y) & ~richer(x,y)) &  
    (\<forall>x. hates(agatha,x) --> ~hates(charles,x)) &  
--- a/src/FOL/ex/Prolog.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOL/ex/Prolog.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -22,18 +22,18 @@
   revNil:  "rev(Nil,Nil)"
   revCons: "[| rev(xs,ys);  app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)"
 
-lemma "app(a:b:c:Nil, d:e:Nil, ?x)"
+schematic_lemma "app(a:b:c:Nil, d:e:Nil, ?x)"
 apply (rule appNil appCons)
 apply (rule appNil appCons)
 apply (rule appNil appCons)
 apply (rule appNil appCons)
 done
 
-lemma "app(?x, c:d:Nil, a:b:c:d:Nil)"
+schematic_lemma "app(?x, c:d:Nil, a:b:c:d:Nil)"
 apply (rule appNil appCons)+
 done
 
-lemma "app(?x, ?y, a:b:c:d:Nil)"
+schematic_lemma "app(?x, ?y, a:b:c:d:Nil)"
 apply (rule appNil appCons)+
 back
 back
@@ -46,15 +46,15 @@
 
 lemmas rules = appNil appCons revNil revCons
 
-lemma "rev(a:b:c:d:Nil, ?x)"
+schematic_lemma "rev(a:b:c:d:Nil, ?x)"
 apply (rule rules)+
 done
 
-lemma "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"
+schematic_lemma "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:Nil, ?w)"
 apply (rule rules)+
 done
 
-lemma "rev(?x, a:b:c:Nil)"
+schematic_lemma "rev(?x, a:b:c:Nil)"
 apply (rule rules)+  -- {* does not solve it directly! *}
 back
 back
@@ -65,22 +65,22 @@
 val prolog_tac = DEPTH_FIRST (has_fewer_prems 1) (resolve_tac (@{thms rules}) 1)
 *}
 
-lemma "rev(?x, a:b:c:Nil)"
+schematic_lemma "rev(?x, a:b:c:Nil)"
 apply (tactic prolog_tac)
 done
 
-lemma "rev(a:?x:c:?y:Nil, d:?z:b:?u)"
+schematic_lemma "rev(a:?x:c:?y:Nil, d:?z:b:?u)"
 apply (tactic prolog_tac)
 done
 
 (*rev([a..p], ?w) requires 153 inferences *)
-lemma "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"
+schematic_lemma "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"
 apply (tactic {* DEPTH_SOLVE (resolve_tac ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
 done
 
 (*?x has 16, ?y has 32;  rev(?y,?w) requires 561 (rather large) inferences
   total inferences = 2 + 1 + 17 + 561 = 581*)
-lemma "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"
+schematic_lemma "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"
 apply (tactic {* DEPTH_SOLVE (resolve_tac ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
 done
 
--- a/src/FOL/ex/Quantifiers_Cla.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOL/ex/Quantifiers_Cla.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -58,11 +58,11 @@
   apply fast?
   oops
 
-lemma "P(?a) --> (ALL x. P(x))"
+schematic_lemma "P(?a) --> (ALL x. P(x))"
   apply fast?
   oops
 
-lemma "(P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
+schematic_lemma "(P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
   apply fast?
   oops
 
@@ -76,7 +76,7 @@
 lemma "(P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
   by fast
 
-lemma "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
+schematic_lemma "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
   by fast
 
 lemma "(ALL x. Q(x))  -->  (EX x. Q(x))"
--- a/src/FOL/ex/Quantifiers_Int.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOL/ex/Quantifiers_Int.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -58,11 +58,11 @@
   apply (tactic "IntPr.fast_tac 1")?
   oops
 
-lemma "P(?a) --> (ALL x. P(x))"
+schematic_lemma "P(?a) --> (ALL x. P(x))"
   apply (tactic "IntPr.fast_tac 1")?
   oops
 
-lemma "(P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
+schematic_lemma "(P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
   apply (tactic "IntPr.fast_tac 1")?
   oops
 
@@ -76,7 +76,7 @@
 lemma "(P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
   by (tactic "IntPr.fast_tac 1")
 
-lemma "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
+schematic_lemma "(ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
   by (tactic "IntPr.fast_tac 1")
 
 lemma "(ALL x. Q(x))  -->  (EX x. Q(x))"
--- a/src/FOLP/FOLP.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/FOLP.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -19,7 +19,7 @@
 
 (*** Classical introduction rules for | and EX ***)
 
-lemma disjCI:
+schematic_lemma disjCI:
   assumes "!!x. x:~Q ==> f(x):P"
   shows "?p : P|Q"
   apply (rule classical)
@@ -28,7 +28,7 @@
   done
 
 (*introduction rule involving only EX*)
-lemma ex_classical:
+schematic_lemma ex_classical:
   assumes "!!u. u:~(EX x. P(x)) ==> f(u):P(a)"
   shows "?p : EX x. P(x)"
   apply (rule classical)
@@ -36,7 +36,7 @@
   done
 
 (*version of above, simplifying ~EX to ALL~ *)
-lemma exCI:
+schematic_lemma exCI:
   assumes "!!u. u:ALL x. ~P(x) ==> f(u):P(a)"
   shows "?p : EX x. P(x)"
   apply (rule ex_classical)
@@ -45,7 +45,7 @@
   apply (erule exI)
   done
 
-lemma excluded_middle: "?p : ~P | P"
+schematic_lemma excluded_middle: "?p : ~P | P"
   apply (rule disjCI)
   apply assumption
   done
@@ -54,7 +54,7 @@
 (*** Special elimination rules *)
 
 (*Classical implies (-->) elimination. *)
-lemma impCE:
+schematic_lemma impCE:
   assumes major: "p:P-->Q"
     and r1: "!!x. x:~P ==> f(x):R"
     and r2: "!!y. y:Q ==> g(y):R"
@@ -65,7 +65,7 @@
   done
 
 (*Double negation law*)
-lemma notnotD: "p:~~P ==> ?p : P"
+schematic_lemma notnotD: "p:~~P ==> ?p : P"
   apply (rule classical)
   apply (erule notE)
   apply assumption
@@ -76,7 +76,7 @@
 
 (*Classical <-> elimination.  Proof substitutes P=Q in
     ~P ==> ~Q    and    P ==> Q  *)
-lemma iffCE:
+schematic_lemma iffCE:
   assumes major: "p:P<->Q"
     and r1: "!!x y.[| x:P; y:Q |] ==> f(x,y):R"
     and r2: "!!x y.[| x:~P; y:~Q |] ==> g(x,y):R"
@@ -91,7 +91,7 @@
 
 
 (*Should be used as swap since ~P becomes redundant*)
-lemma swap:
+schematic_lemma swap:
   assumes major: "p:~P"
     and r: "!!x. x:~Q ==> f(x):P"
   shows "?p : Q"
@@ -136,7 +136,7 @@
     addSEs [@{thm exE}, @{thm ex1E}] addEs [@{thm all_dupE}];
 *}
 
-lemma cla_rews:
+schematic_lemma cla_rews:
   "?p1 : P | ~P"
   "?p2 : ~P | P"
   "?p3 : ~ ~ P <-> P"
--- a/src/FOLP/IFOLP.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/IFOLP.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -159,7 +159,7 @@
 
 (*** Sequent-style elimination rules for & --> and ALL ***)
 
-lemma conjE:
+schematic_lemma conjE:
   assumes "p:P&Q"
     and "!!x y.[| x:P; y:Q |] ==> f(x,y):R"
   shows "?a:R"
@@ -168,7 +168,7 @@
   apply (rule conjunct2 [OF assms(1)])
   done
 
-lemma impE:
+schematic_lemma impE:
   assumes "p:P-->Q"
     and "q:P"
     and "!!x. x:Q ==> r(x):R"
@@ -176,7 +176,7 @@
   apply (rule assms mp)+
   done
 
-lemma allE:
+schematic_lemma allE:
   assumes "p:ALL x. P(x)"
     and "!!y. y:P(x) ==> q(y):R"
   shows "?p:R"
@@ -184,7 +184,7 @@
   done
 
 (*Duplicates the quantifier; for use with eresolve_tac*)
-lemma all_dupE:
+schematic_lemma all_dupE:
   assumes "p:ALL x. P(x)"
     and "!!y z.[| y:P(x); z:ALL x. P(x) |] ==> q(y,z):R"
   shows "?p:R"
@@ -194,21 +194,21 @@
 
 (*** Negation rules, which translate between ~P and P-->False ***)
 
-lemma notI:
+schematic_lemma notI:
   assumes "!!x. x:P ==> q(x):False"
   shows "?p:~P"
   unfolding not_def
   apply (assumption | rule assms impI)+
   done
 
-lemma notE: "p:~P \<Longrightarrow> q:P \<Longrightarrow> ?p:R"
+schematic_lemma notE: "p:~P \<Longrightarrow> q:P \<Longrightarrow> ?p:R"
   unfolding not_def
   apply (drule (1) mp)
   apply (erule FalseE)
   done
 
 (*This is useful with the special implication rules for each kind of P. *)
-lemma not_to_imp:
+schematic_lemma not_to_imp:
   assumes "p:~P"
     and "!!x. x:(P-->False) ==> q(x):Q"
   shows "?p:Q"
@@ -217,13 +217,13 @@
 
 (* For substitution int an assumption P, reduce Q to P-->Q, substitute into
    this implication, then apply impI to move P back into the assumptions.*)
-lemma rev_mp: "[| p:P;  q:P --> Q |] ==> ?p:Q"
+schematic_lemma rev_mp: "[| p:P;  q:P --> Q |] ==> ?p:Q"
   apply (assumption | rule mp)+
   done
 
 
 (*Contrapositive of an inference rule*)
-lemma contrapos:
+schematic_lemma contrapos:
   assumes major: "p:~Q"
     and minor: "!!y. y:P==>q(y):Q"
   shows "?a:~P"
@@ -271,7 +271,7 @@
 
 (*** If-and-only-if ***)
 
-lemma iffI:
+schematic_lemma iffI:
   assumes "!!x. x:P ==> q(x):Q"
     and "!!x. x:Q ==> r(x):P"
   shows "?p:P<->Q"
@@ -282,7 +282,7 @@
 
 (*Observe use of rewrite_rule to unfold "<->" in meta-assumptions (prems) *)
   
-lemma iffE:
+schematic_lemma iffE:
   assumes "p:P <-> Q"
     and "!!x y.[| x:P-->Q; y:Q-->P |] ==> q(x,y):R"
   shows "?p:R"
@@ -294,28 +294,28 @@
 
 (* Destruct rules for <-> similar to Modus Ponens *)
 
-lemma iffD1: "[| p:P <-> Q; q:P |] ==> ?p:Q"
+schematic_lemma iffD1: "[| p:P <-> Q; q:P |] ==> ?p:Q"
   unfolding iff_def
   apply (rule conjunct1 [THEN mp], assumption+)
   done
 
-lemma iffD2: "[| p:P <-> Q; q:Q |] ==> ?p:P"
+schematic_lemma iffD2: "[| p:P <-> Q; q:Q |] ==> ?p:P"
   unfolding iff_def
   apply (rule conjunct2 [THEN mp], assumption+)
   done
 
-lemma iff_refl: "?p:P <-> P"
+schematic_lemma iff_refl: "?p:P <-> P"
   apply (rule iffI)
    apply assumption+
   done
 
-lemma iff_sym: "p:Q <-> P ==> ?p:P <-> Q"
+schematic_lemma iff_sym: "p:Q <-> P ==> ?p:P <-> Q"
   apply (erule iffE)
   apply (rule iffI)
    apply (erule (1) mp)+
   done
 
-lemma iff_trans: "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R"
+schematic_lemma iff_trans: "[| p:P <-> Q; q:Q<-> R |] ==> ?p:P <-> R"
   apply (rule iffI)
    apply (assumption | erule iffE | erule (1) impE)+
   done
@@ -326,7 +326,7 @@
  do NOT mean the same thing.  The parser treats EX!x y.P(x,y) as sequential.
 ***)
 
-lemma ex1I:
+schematic_lemma ex1I:
   assumes "p:P(a)"
     and "!!x u. u:P(x) ==> f(u) : x=a"
   shows "?p:EX! x. P(x)"
@@ -334,7 +334,7 @@
   apply (assumption | rule assms exI conjI allI impI)+
   done
 
-lemma ex1E:
+schematic_lemma ex1E:
   assumes "p:EX! x. P(x)"
     and "!!x u v. [| u:P(x);  v:ALL y. P(y) --> y=x |] ==> f(x,u,v):R"
   shows "?a : R"
@@ -353,7 +353,7 @@
     REPEAT1 (eresolve_tac [asm_rl, @{thm mp}] i)
 *}
 
-lemma conj_cong:
+schematic_lemma conj_cong:
   assumes "p:P <-> P'"
     and "!!x. x:P' ==> q(x):Q <-> Q'"
   shows "?p:(P&Q) <-> (P'&Q')"
@@ -362,12 +362,12 @@
     erule iffE conjE mp | tactic {* iff_tac @{thms assms} 1 *})+
   done
 
-lemma disj_cong:
+schematic_lemma disj_cong:
   "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P|Q) <-> (P'|Q')"
   apply (erule iffE disjE disjI1 disjI2 | assumption | rule iffI | tactic {* mp_tac 1 *})+
   done
 
-lemma imp_cong:
+schematic_lemma imp_cong:
   assumes "p:P <-> P'"
     and "!!x. x:P' ==> q(x):Q <-> Q'"
   shows "?p:(P-->Q) <-> (P'-->Q')"
@@ -376,24 +376,24 @@
     tactic {* iff_tac @{thms assms} 1 *})+
   done
 
-lemma iff_cong:
+schematic_lemma iff_cong:
   "[| p:P <-> P'; q:Q <-> Q' |] ==> ?p:(P<->Q) <-> (P'<->Q')"
   apply (erule iffE | assumption | rule iffI | tactic {* mp_tac 1 *})+
   done
 
-lemma not_cong:
+schematic_lemma not_cong:
   "p:P <-> P' ==> ?p:~P <-> ~P'"
   apply (assumption | rule iffI notI | tactic {* mp_tac 1 *} | erule iffE notE)+
   done
 
-lemma all_cong:
+schematic_lemma all_cong:
   assumes "!!x. f(x):P(x) <-> Q(x)"
   shows "?p:(ALL x. P(x)) <-> (ALL x. Q(x))"
   apply (assumption | rule iffI allI | tactic {* mp_tac 1 *} | erule allE |
     tactic {* iff_tac @{thms assms} 1 *})+
   done
 
-lemma ex_cong:
+schematic_lemma ex_cong:
   assumes "!!x. f(x):P(x) <-> Q(x)"
   shows "?p:(EX x. P(x)) <-> (EX x. Q(x))"
   apply (erule exE | assumption | rule iffI exI | tactic {* mp_tac 1 *} |
@@ -413,7 +413,7 @@
 
 lemmas refl = ieqI
 
-lemma subst:
+schematic_lemma subst:
   assumes prem1: "p:a=b"
     and prem2: "q:P(a)"
   shows "?p : P(b)"
@@ -423,17 +423,17 @@
   apply assumption
   done
 
-lemma sym: "q:a=b ==> ?c:b=a"
+schematic_lemma sym: "q:a=b ==> ?c:b=a"
   apply (erule subst)
   apply (rule refl)
   done
 
-lemma trans: "[| p:a=b;  q:b=c |] ==> ?d:a=c"
+schematic_lemma trans: "[| p:a=b;  q:b=c |] ==> ?d:a=c"
   apply (erule (1) subst)
   done
 
 (** ~ b=a ==> ~ a=b **)
-lemma not_sym: "p:~ b=a ==> ?q:~ a=b"
+schematic_lemma not_sym: "p:~ b=a ==> ?q:~ a=b"
   apply (erule contrapos)
   apply (erule sym)
   done
@@ -442,7 +442,7 @@
 lemmas ssubst = sym [THEN subst, standard]
 
 (*A special case of ex1E that would otherwise need quantifier expansion*)
-lemma ex1_equalsE: "[| p:EX! x. P(x);  q:P(a);  r:P(b) |] ==> ?d:a=b"
+schematic_lemma ex1_equalsE: "[| p:EX! x. P(x);  q:P(a);  r:P(b) |] ==> ?d:a=b"
   apply (erule ex1E)
   apply (rule trans)
    apply (rule_tac [2] sym)
@@ -451,17 +451,17 @@
 
 (** Polymorphic congruence rules **)
 
-lemma subst_context: "[| p:a=b |]  ==>  ?d:t(a)=t(b)"
+schematic_lemma subst_context: "[| p:a=b |]  ==>  ?d:t(a)=t(b)"
   apply (erule ssubst)
   apply (rule refl)
   done
 
-lemma subst_context2: "[| p:a=b;  q:c=d |]  ==>  ?p:t(a,c)=t(b,d)"
+schematic_lemma subst_context2: "[| p:a=b;  q:c=d |]  ==>  ?p:t(a,c)=t(b,d)"
   apply (erule ssubst)+
   apply (rule refl)
   done
 
-lemma subst_context3: "[| p:a=b;  q:c=d;  r:e=f |]  ==>  ?p:t(a,c,e)=t(b,d,f)"
+schematic_lemma subst_context3: "[| p:a=b;  q:c=d;  r:e=f |]  ==>  ?p:t(a,c,e)=t(b,d,f)"
   apply (erule ssubst)+
   apply (rule refl)
   done
@@ -470,7 +470,7 @@
         a = b
         |   |
         c = d   *)
-lemma box_equals: "[| p:a=b;  q:a=c;  r:b=d |] ==> ?p:c=d"
+schematic_lemma box_equals: "[| p:a=b;  q:a=c;  r:b=d |] ==> ?p:c=d"
   apply (rule trans)
    apply (rule trans)
     apply (rule sym)
@@ -478,7 +478,7 @@
   done
 
 (*Dual of box_equals: for proving equalities backwards*)
-lemma simp_equals: "[| p:a=c;  q:b=d;  r:c=d |] ==> ?p:a=b"
+schematic_lemma simp_equals: "[| p:a=c;  q:b=d;  r:c=d |] ==> ?p:a=b"
   apply (rule trans)
    apply (rule trans)
     apply (assumption | rule sym)+
@@ -486,17 +486,17 @@
 
 (** Congruence rules for predicate letters **)
 
-lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')"
+schematic_lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')"
   apply (rule iffI)
    apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
   done
 
-lemma pred2_cong: "[| p:a=a';  q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')"
+schematic_lemma pred2_cong: "[| p:a=a';  q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')"
   apply (rule iffI)
    apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
   done
 
-lemma pred3_cong: "[| p:a=a';  q:b=b';  r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')"
+schematic_lemma pred3_cong: "[| p:a=a';  q:b=b';  r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')"
   apply (rule iffI)
    apply (tactic {* DEPTH_SOLVE (atac 1 ORELSE eresolve_tac [@{thm subst}, @{thm ssubst}] 1) *})
   done
@@ -514,14 +514,14 @@
    R. Dyckhoff, Contraction-free sequent calculi for intuitionistic logic
     (preprint, University of St Andrews, 1991)  ***)
 
-lemma conj_impE:
+schematic_lemma conj_impE:
   assumes major: "p:(P&Q)-->S"
     and minor: "!!x. x:P-->(Q-->S) ==> q(x):R"
   shows "?p:R"
   apply (assumption | rule conjI impI major [THEN mp] minor)+
   done
 
-lemma disj_impE:
+schematic_lemma disj_impE:
   assumes major: "p:(P|Q)-->S"
     and minor: "!!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R"
   shows "?p:R"
@@ -532,7 +532,7 @@
 
 (*Simplifies the implication.  Classical version is stronger.
   Still UNSAFE since Q must be provable -- backtracking needed.  *)
-lemma imp_impE:
+schematic_lemma imp_impE:
   assumes major: "p:(P-->Q)-->S"
     and r1: "!!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q"
     and r2: "!!x. x:S ==> r(x):R"
@@ -542,7 +542,7 @@
 
 (*Simplifies the implication.  Classical version is stronger.
   Still UNSAFE since ~P must be provable -- backtracking needed.  *)
-lemma not_impE:
+schematic_lemma not_impE:
   assumes major: "p:~P --> S"
     and r1: "!!y. y:P ==> q(y):False"
     and r2: "!!y. y:S ==> r(y):R"
@@ -551,7 +551,7 @@
   done
 
 (*Simplifies the implication.   UNSAFE.  *)
-lemma iff_impE:
+schematic_lemma iff_impE:
   assumes major: "p:(P<->Q)-->S"
     and r1: "!!x y.[| x:P; y:Q-->S |] ==> q(x,y):Q"
     and r2: "!!x y.[| x:Q; y:P-->S |] ==> r(x,y):P"
@@ -561,7 +561,7 @@
   done
 
 (*What if (ALL x.~~P(x)) --> ~~(ALL x.P(x)) is an assumption? UNSAFE*)
-lemma all_impE:
+schematic_lemma all_impE:
   assumes major: "p:(ALL x. P(x))-->S"
     and r1: "!!x. q:P(x)"
     and r2: "!!y. y:S ==> r(y):R"
@@ -570,7 +570,7 @@
   done
 
 (*Unsafe: (EX x.P(x))-->S  is equivalent to  ALL x.P(x)-->S.  *)
-lemma ex_impE:
+schematic_lemma ex_impE:
   assumes major: "p:(EX x. P(x))-->S"
     and r: "!!y. y:P(a)-->S ==> q(y):R"
   shows "?p:R"
@@ -578,7 +578,7 @@
   done
 
 
-lemma rev_cut_eq:
+schematic_lemma rev_cut_eq:
   assumes "p:a=b"
     and "!!x. x:a=b ==> f(x):R"
   shows "?p:R"
@@ -619,7 +619,7 @@
 
 (*** Rewrite rules ***)
 
-lemma conj_rews:
+schematic_lemma conj_rews:
   "?p1 : P & True <-> P"
   "?p2 : True & P <-> P"
   "?p3 : P & False <-> False"
@@ -631,7 +631,7 @@
   apply (tactic {* fn st => IntPr.fast_tac 1 st *})+
   done
 
-lemma disj_rews:
+schematic_lemma disj_rews:
   "?p1 : P | True <-> True"
   "?p2 : True | P <-> True"
   "?p3 : P | False <-> P"
@@ -641,13 +641,13 @@
   apply (tactic {* IntPr.fast_tac 1 *})+
   done
 
-lemma not_rews:
+schematic_lemma not_rews:
   "?p1 : ~ False <-> True"
   "?p2 : ~ True <-> False"
   apply (tactic {* IntPr.fast_tac 1 *})+
   done
 
-lemma imp_rews:
+schematic_lemma imp_rews:
   "?p1 : (P --> False) <-> ~P"
   "?p2 : (P --> True) <-> True"
   "?p3 : (False --> P) <-> True"
@@ -657,7 +657,7 @@
   apply (tactic {* IntPr.fast_tac 1 *})+
   done
 
-lemma iff_rews:
+schematic_lemma iff_rews:
   "?p1 : (True <-> P) <-> P"
   "?p2 : (P <-> True) <-> P"
   "?p3 : (P <-> P) <-> True"
@@ -666,14 +666,14 @@
   apply (tactic {* IntPr.fast_tac 1 *})+
   done
 
-lemma quant_rews:
+schematic_lemma quant_rews:
   "?p1 : (ALL x. P) <-> P"
   "?p2 : (EX x. P) <-> P"
   apply (tactic {* IntPr.fast_tac 1 *})+
   done
 
 (*These are NOT supplied by default!*)
-lemma distrib_rews1:
+schematic_lemma distrib_rews1:
   "?p1 : ~(P|Q) <-> ~P & ~Q"
   "?p2 : P & (Q | R) <-> P&Q | P&R"
   "?p3 : (Q | R) & P <-> Q&P | R&P"
@@ -681,7 +681,7 @@
   apply (tactic {* IntPr.fast_tac 1 *})+
   done
 
-lemma distrib_rews2:
+schematic_lemma distrib_rews2:
   "?p1 : ~(EX x. NORM(P(x))) <-> (ALL x. ~NORM(P(x)))"
   "?p2 : ((EX x. NORM(P(x))) --> Q) <-> (ALL x. NORM(P(x)) --> Q)"
   "?p3 : (EX x. NORM(P(x))) & NORM(Q) <-> (EX x. NORM(P(x)) & NORM(Q))"
@@ -691,11 +691,11 @@
 
 lemmas distrib_rews = distrib_rews1 distrib_rews2
 
-lemma P_Imp_P_iff_T: "p:P ==> ?p:(P <-> True)"
+schematic_lemma P_Imp_P_iff_T: "p:P ==> ?p:(P <-> True)"
   apply (tactic {* IntPr.fast_tac 1 *})
   done
 
-lemma not_P_imp_P_iff_F: "p:~P ==> ?p:(P <-> False)"
+schematic_lemma not_P_imp_P_iff_F: "p:~P ==> ?p:(P <-> False)"
   apply (tactic {* IntPr.fast_tac 1 *})
   done
 
--- a/src/FOLP/ex/Classical.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Classical.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -9,14 +9,14 @@
 imports FOLP
 begin
 
-lemma "?p : (P --> Q | R) --> (P-->Q) | (P-->R)"
+schematic_lemma "?p : (P --> Q | R) --> (P-->Q) | (P-->R)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*If and only if*)
-lemma "?p : (P<->Q) <-> (Q<->P)"
+schematic_lemma "?p : (P<->Q) <-> (Q<->P)"
   by (tactic "fast_tac FOLP_cs 1")
 
-lemma "?p : ~ (P <-> ~P)"
+schematic_lemma "?p : ~ (P <-> ~P)"
   by (tactic "fast_tac FOLP_cs 1")
 
 
@@ -32,138 +32,138 @@
 
 text "Pelletier's examples"
 (*1*)
-lemma "?p : (P-->Q)  <->  (~Q --> ~P)"
+schematic_lemma "?p : (P-->Q)  <->  (~Q --> ~P)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*2*)
-lemma "?p : ~ ~ P  <->  P"
+schematic_lemma "?p : ~ ~ P  <->  P"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*3*)
-lemma "?p : ~(P-->Q) --> (Q-->P)"
+schematic_lemma "?p : ~(P-->Q) --> (Q-->P)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*4*)
-lemma "?p : (~P-->Q)  <->  (~Q --> P)"
+schematic_lemma "?p : (~P-->Q)  <->  (~Q --> P)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*5*)
-lemma "?p : ((P|Q)-->(P|R)) --> (P|(Q-->R))"
+schematic_lemma "?p : ((P|Q)-->(P|R)) --> (P|(Q-->R))"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*6*)
-lemma "?p : P | ~ P"
+schematic_lemma "?p : P | ~ P"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*7*)
-lemma "?p : P | ~ ~ ~ P"
+schematic_lemma "?p : P | ~ ~ ~ P"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*8.  Peirce's law*)
-lemma "?p : ((P-->Q) --> P)  -->  P"
+schematic_lemma "?p : ((P-->Q) --> P)  -->  P"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*9*)
-lemma "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
+schematic_lemma "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*10*)
-lemma "?p : (Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P<->Q)"
+schematic_lemma "?p : (Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P<->Q)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*11.  Proved in each direction (incorrectly, says Pelletier!!)  *)
-lemma "?p : P<->P"
+schematic_lemma "?p : P<->P"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*12.  "Dijkstra's law"*)
-lemma "?p : ((P <-> Q) <-> R)  <->  (P <-> (Q <-> R))"
+schematic_lemma "?p : ((P <-> Q) <-> R)  <->  (P <-> (Q <-> R))"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*13.  Distributive law*)
-lemma "?p : P | (Q & R)  <-> (P | Q) & (P | R)"
+schematic_lemma "?p : P | (Q & R)  <-> (P | Q) & (P | R)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*14*)
-lemma "?p : (P <-> Q) <-> ((Q | ~P) & (~Q|P))"
+schematic_lemma "?p : (P <-> Q) <-> ((Q | ~P) & (~Q|P))"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*15*)
-lemma "?p : (P --> Q) <-> (~P | Q)"
+schematic_lemma "?p : (P --> Q) <-> (~P | Q)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*16*)
-lemma "?p : (P-->Q) | (Q-->P)"
+schematic_lemma "?p : (P-->Q) | (Q-->P)"
   by (tactic "fast_tac FOLP_cs 1")
 
 (*17*)
-lemma "?p : ((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))"
+schematic_lemma "?p : ((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))"
   by (tactic "fast_tac FOLP_cs 1")
 
 
 text "Classical Logic: examples with quantifiers"
 
-lemma "?p : (ALL x. P(x) & Q(x)) <-> (ALL x. P(x))  &  (ALL x. Q(x))"
+schematic_lemma "?p : (ALL x. P(x) & Q(x)) <-> (ALL x. P(x))  &  (ALL x. Q(x))"
   by (tactic "fast_tac FOLP_cs 1")
 
-lemma "?p : (EX x. P-->Q(x))  <->  (P --> (EX x. Q(x)))"
+schematic_lemma "?p : (EX x. P-->Q(x))  <->  (P --> (EX x. Q(x)))"
   by (tactic "fast_tac FOLP_cs 1")
 
-lemma "?p : (EX x. P(x)-->Q)  <->  (ALL x. P(x)) --> Q"
+schematic_lemma "?p : (EX x. P(x)-->Q)  <->  (ALL x. P(x)) --> Q"
   by (tactic "fast_tac FOLP_cs 1")
 
-lemma "?p : (ALL x. P(x)) | Q  <->  (ALL x. P(x) | Q)"
+schematic_lemma "?p : (ALL x. P(x)) | Q  <->  (ALL x. P(x) | Q)"
   by (tactic "fast_tac FOLP_cs 1")
 
 
 text "Problems requiring quantifier duplication"
 
 (*Needs multiple instantiation of ALL.*)
-lemma "?p : (ALL x. P(x)-->P(f(x)))  &  P(d)-->P(f(f(f(d))))"
+schematic_lemma "?p : (ALL x. P(x)-->P(f(x)))  &  P(d)-->P(f(f(f(d))))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 (*Needs double instantiation of the quantifier*)
-lemma "?p : EX x. P(x) --> P(a) & P(b)"
+schematic_lemma "?p : EX x. P(x) --> P(a) & P(b)"
   by (tactic "best_tac FOLP_dup_cs 1")
 
-lemma "?p : EX z. P(z) --> (ALL x. P(x))"
+schematic_lemma "?p : EX z. P(z) --> (ALL x. P(x))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 
 text "Hard examples with quantifiers"
 
 text "Problem 18"
-lemma "?p : EX y. ALL x. P(y)-->P(x)"
+schematic_lemma "?p : EX y. ALL x. P(y)-->P(x)"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 19"
-lemma "?p : EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"
+schematic_lemma "?p : EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 20"
-lemma "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))      
+schematic_lemma "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))      
     --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 21"
-lemma "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))";
+schematic_lemma "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> (EX x. P<->Q(x))";
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 22"
-lemma "?p : (ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
+schematic_lemma "?p : (ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 23"
-lemma "?p : (ALL x. P | Q(x))  <->  (P | (ALL x. Q(x)))"
+schematic_lemma "?p : (ALL x. P | Q(x))  <->  (P | (ALL x. Q(x)))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 24"
-lemma "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &   
+schematic_lemma "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &   
      (~(EX x. P(x)) --> (EX x. Q(x))) & (ALL x. Q(x)|R(x) --> S(x))   
     --> (EX x. P(x)&R(x))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 25"
-lemma "?p : (EX x. P(x)) &  
+schematic_lemma "?p : (EX x. P(x)) &  
        (ALL x. L(x) --> ~ (M(x) & R(x))) &  
        (ALL x. P(x) --> (M(x) & L(x))) &   
        ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))  
@@ -171,13 +171,13 @@
   oops
 
 text "Problem 26"
-lemma "?u : ((EX x. p(x)) <-> (EX x. q(x))) &   
+schematic_lemma "?u : ((EX x. p(x)) <-> (EX x. q(x))) &   
      (ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y)))   
   --> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))";
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 27"
-lemma "?p : (EX x. P(x) & ~Q(x)) &    
+schematic_lemma "?p : (EX x. P(x) & ~Q(x)) &    
               (ALL x. P(x) --> R(x)) &    
               (ALL x. M(x) & L(x) --> P(x)) &    
               ((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x)))   
@@ -185,49 +185,49 @@
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 28.  AMENDED"
-lemma "?p : (ALL x. P(x) --> (ALL x. Q(x))) &    
+schematic_lemma "?p : (ALL x. P(x) --> (ALL x. Q(x))) &    
         ((ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) &   
         ((EX x. S(x)) --> (ALL x. L(x) --> M(x)))   
     --> (ALL x. P(x) & L(x) --> M(x))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 29.  Essentially the same as Principia Mathematica *11.71"
-lemma "?p : (EX x. P(x)) & (EX y. Q(y))   
+schematic_lemma "?p : (EX x. P(x)) & (EX y. Q(y))   
     --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y))   <->      
          (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 30"
-lemma "?p : (ALL x. P(x) | Q(x) --> ~ R(x)) &  
+schematic_lemma "?p : (ALL x. P(x) | Q(x) --> ~ R(x)) &  
         (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))   
     --> (ALL x. S(x))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 31"
-lemma "?p : ~(EX x. P(x) & (Q(x) | R(x))) &  
+schematic_lemma "?p : ~(EX x. P(x) & (Q(x) | R(x))) &  
         (EX x. L(x) & P(x)) &  
         (ALL x. ~ R(x) --> M(x))   
     --> (EX x. L(x) & M(x))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 32"
-lemma "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) &  
+schematic_lemma "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) &  
         (ALL x. S(x) & R(x) --> L(x)) &  
         (ALL x. M(x) --> R(x))   
     --> (ALL x. P(x) & M(x) --> L(x))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 33"
-lemma "?p : (ALL x. P(a) & (P(x)-->P(b))-->P(c))  <->     
+schematic_lemma "?p : (ALL x. P(a) & (P(x)-->P(b))-->P(c))  <->     
      (ALL x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 35"
-lemma "?p : EX x y. P(x,y) -->  (ALL u v. P(u,v))"
+schematic_lemma "?p : EX x y. P(x,y) -->  (ALL u v. P(u,v))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 36"
-lemma
+schematic_lemma
 "?p : (ALL x. EX y. J(x,y)) &  
       (ALL x. EX y. G(x,y)) &  
       (ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z)))    
@@ -235,7 +235,7 @@
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 37"
-lemma "?p : (ALL z. EX w. ALL x. EX y.  
+schematic_lemma "?p : (ALL z. EX w. ALL x. EX y.  
            (P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u. Q(u,w)))) &  
         (ALL x z. ~P(x,z) --> (EX y. Q(y,z))) &  
         ((EX x y. Q(x,y)) --> (ALL x. R(x,x)))   
@@ -243,21 +243,21 @@
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 39"
-lemma "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
+schematic_lemma "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 40.  AMENDED"
-lemma "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) -->   
+schematic_lemma "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) -->   
               ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 41"
-lemma "?p : (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x))   
+schematic_lemma "?p : (ALL z. EX y. ALL x. f(x,y) <-> f(x,z) & ~ f(x,x))   
           --> ~ (EX z. ALL x. f(x,z))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 44"
-lemma "?p : (ALL x. f(x) -->                                     
+schematic_lemma "?p : (ALL x. f(x) -->                                     
               (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y))))  &        
               (EX x. j(x) & (ALL y. g(y) --> h(x,y)))                    
               --> (EX x. j(x) & ~f(x))"
@@ -266,37 +266,37 @@
 text "Problems (mainly) involving equality or functions"
 
 text "Problem 48"
-lemma "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"
+schematic_lemma "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 50"
 (*What has this to do with equality?*)
-lemma "?p : (ALL x. P(a,x) | (ALL y. P(x,y))) --> (EX x. ALL y. P(x,y))"
+schematic_lemma "?p : (ALL x. P(a,x) | (ALL y. P(x,y))) --> (EX x. ALL y. P(x,y))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 56"
-lemma
+schematic_lemma
  "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 57"
-lemma
+schematic_lemma
 "?p : P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &  
       (ALL x y z. P(x,y) & P(y,z) --> P(x,z))    -->   P(f(a,b), f(a,c))"
   by (tactic "fast_tac FOLP_cs 1")
 
 text "Problem 58  NOT PROVED AUTOMATICALLY"
-lemma
+schematic_lemma
   notes f_cong = subst_context [where t = f]
   shows "?p : (ALL x y. f(x)=g(y)) --> (ALL x y. f(f(x))=f(g(y)))"
   by (tactic {* fast_tac (FOLP_cs addSIs [@{thm f_cong}]) 1 *})
 
 text "Problem 59"
-lemma "?p : (ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"
+schematic_lemma "?p : (ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"
   by (tactic "best_tac FOLP_dup_cs 1")
 
 text "Problem 60"
-lemma "?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
+schematic_lemma "?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
   by (tactic "fast_tac FOLP_cs 1")
 
 end
--- a/src/FOLP/ex/Foundation.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Foundation.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -9,7 +9,7 @@
 imports IFOLP
 begin
 
-lemma "?p : A&B  --> (C-->A&C)"
+schematic_lemma "?p : A&B  --> (C-->A&C)"
 apply (rule impI)
 apply (rule impI)
 apply (rule conjI)
@@ -19,7 +19,7 @@
 done
 
 text {*A form of conj-elimination*}
-lemma
+schematic_lemma
   assumes "p : A & B"
     and "!!x y. x : A ==> y : B ==> f(x, y) : C"
   shows "?p : C"
@@ -30,7 +30,7 @@
 apply (rule prems)
 done
 
-lemma
+schematic_lemma
   assumes "!!A x. x : ~ ~A ==> cla(x) : A"
   shows "?p : B | ~B"
 apply (rule prems)
@@ -48,7 +48,7 @@
 apply assumption
 done
 
-lemma
+schematic_lemma
   assumes "!!A x. x : ~ ~A ==> cla(x) : A"
   shows "?p : B | ~B"
 apply (rule prems)
@@ -63,7 +63,7 @@
 done
 
 
-lemma
+schematic_lemma
   assumes "p : A | ~A"
     and "q : ~ ~A"
   shows "?p : A"
@@ -79,7 +79,7 @@
 
 subsection "Examples with quantifiers"
 
-lemma
+schematic_lemma
   assumes "p : ALL z. G(z)"
   shows "?p : ALL z. G(z)|H(z)"
 apply (rule allI)
@@ -87,20 +87,20 @@
 apply (rule prems [THEN spec])
 done
 
-lemma "?p : ALL x. EX y. x=y"
+schematic_lemma "?p : ALL x. EX y. x=y"
 apply (rule allI)
 apply (rule exI)
 apply (rule refl)
 done
 
-lemma "?p : EX y. ALL x. x=y"
+schematic_lemma "?p : EX y. ALL x. x=y"
 apply (rule exI)
 apply (rule allI)
 apply (rule refl)?
 oops
 
 text {* Parallel lifting example. *}
-lemma "?p : EX u. ALL x. EX v. ALL y. EX w. P(u,x,v,y,w)"
+schematic_lemma "?p : EX u. ALL x. EX v. ALL y. EX w. P(u,x,v,y,w)"
 apply (rule exI allI)
 apply (rule exI allI)
 apply (rule exI allI)
@@ -108,7 +108,7 @@
 apply (rule exI allI)
 oops
 
-lemma
+schematic_lemma
   assumes "p : (EX z. F(z)) & B"
   shows "?p : EX z. F(z) & B"
 apply (rule conjE)
@@ -122,7 +122,7 @@
 done
 
 text {* A bigger demonstration of quantifiers -- not in the paper. *}
-lemma "?p : (EX y. ALL x. Q(x,y)) -->  (ALL x. EX y. Q(x,y))"
+schematic_lemma "?p : (EX y. ALL x. Q(x,y)) -->  (ALL x. EX y. Q(x,y))"
 apply (rule impI)
 apply (rule allI)
 apply (rule exE, assumption)
--- a/src/FOLP/ex/If.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/If.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -5,14 +5,14 @@
 definition "if" :: "[o,o,o]=>o" where
   "if(P,Q,R) == P&Q | ~P&R"
 
-lemma ifI:
+schematic_lemma ifI:
   assumes "!!x. x : P ==> f(x) : Q"  "!!x. x : ~P ==> g(x) : R"
   shows "?p : if(P,Q,R)"
 apply (unfold if_def)
 apply (tactic {* fast_tac (FOLP_cs addIs @{thms assms}) 1 *})
 done
 
-lemma ifE:
+schematic_lemma ifE:
   assumes 1: "p : if(P,Q,R)" and
     2: "!!x y. [| x : P; y : Q |] ==> f(x, y) : S" and
     3: "!!x y. [| x : ~P; y : R |] ==> g(x, y) : S"
@@ -22,7 +22,7 @@
 apply (tactic {* fast_tac (FOLP_cs addIs [@{thm 2}, @{thm 3}]) 1 *})
 done
 
-lemma if_commute: "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
+schematic_lemma if_commute: "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
 apply (rule iffI)
 apply (erule ifE)
 apply (erule ifE)
@@ -32,11 +32,11 @@
 
 ML {* val if_cs = FOLP_cs addSIs [@{thm ifI}] addSEs [@{thm ifE}] *}
 
-lemma if_commute: "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
+schematic_lemma if_commute: "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
 apply (tactic {* fast_tac if_cs 1 *})
 done
 
-lemma nested_ifs: "?p : if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
+schematic_lemma nested_ifs: "?p : if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
 apply (tactic {* fast_tac if_cs 1 *})
 done
 
--- a/src/FOLP/ex/Intro.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Intro.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -13,7 +13,7 @@
 
 subsubsection {* Some simple backward proofs *}
 
-lemma mythm: "?p : P|P --> P"
+schematic_lemma mythm: "?p : P|P --> P"
 apply (rule impI)
 apply (rule disjE)
 prefer 3 apply (assumption)
@@ -21,7 +21,7 @@
 apply assumption
 done
 
-lemma "?p : (P & Q) | R --> (P | R)"
+schematic_lemma "?p : (P & Q) | R --> (P | R)"
 apply (rule impI)
 apply (erule disjE)
 apply (drule conjunct1)
@@ -31,7 +31,7 @@
 done
 
 (*Correct version, delaying use of "spec" until last*)
-lemma "?p : (ALL x y. P(x,y)) --> (ALL z w. P(w,z))"
+schematic_lemma "?p : (ALL x y. P(x,y)) --> (ALL z w. P(w,z))"
 apply (rule impI)
 apply (rule allI)
 apply (rule allI)
@@ -43,13 +43,13 @@
 
 subsubsection {* Demonstration of @{text "fast"} *}
 
-lemma "?p : (EX y. ALL x. J(y,x) <-> ~J(x,x))
+schematic_lemma "?p : (EX y. ALL x. J(y,x) <-> ~J(x,x))
         -->  ~ (ALL x. EX y. ALL z. J(z,y) <-> ~ J(z,x))"
 apply (tactic {* fast_tac FOLP_cs 1 *})
 done
 
 
-lemma "?p : ALL x. P(x,f(x)) <->
+schematic_lemma "?p : ALL x. P(x,f(x)) <->
         (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
 apply (tactic {* fast_tac FOLP_cs 1 *})
 done
@@ -57,7 +57,7 @@
 
 subsubsection {* Derivation of conjunction elimination rule *}
 
-lemma
+schematic_lemma
   assumes major: "p : P&Q"
     and minor: "!!x y. [| x : P; y : Q |] ==> f(x, y) : R"
   shows "?p : R"
@@ -71,7 +71,7 @@
 
 text {* Derivation of negation introduction *}
 
-lemma
+schematic_lemma
   assumes "!!x. x : P ==> f(x) : False"
   shows "?p : ~ P"
 apply (unfold not_def)
@@ -80,7 +80,7 @@
 apply assumption
 done
 
-lemma
+schematic_lemma
   assumes major: "p : ~P"
     and minor: "q : P"
   shows "?p : R"
@@ -91,7 +91,7 @@
 done
 
 text {* Alternative proof of the result above *}
-lemma
+schematic_lemma
   assumes major: "p : ~P"
     and minor: "q : P"
   shows "?p : R"
--- a/src/FOLP/ex/Intuitionistic.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Intuitionistic.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -30,39 +30,39 @@
 imports IFOLP
 begin
 
-lemma "?p : ~~(P&Q) <-> ~~P & ~~Q"
+schematic_lemma "?p : ~~(P&Q) <-> ~~P & ~~Q"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ~~~P <-> ~P"
+schematic_lemma "?p : ~~~P <-> ~P"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ~~((P --> Q | R)  -->  (P-->Q) | (P-->R))"
+schematic_lemma "?p : ~~((P --> Q | R)  -->  (P-->Q) | (P-->R))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P<->Q) <-> (Q<->P)"
+schematic_lemma "?p : (P<->Q) <-> (Q<->P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 subsection {* Lemmas for the propositional double-negation translation *}
 
-lemma "?p : P --> ~~P"
+schematic_lemma "?p : P --> ~~P"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ~~(~~P --> P)"
+schematic_lemma "?p : ~~(~~P --> P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ~~P & ~~(P --> Q) --> ~~Q"
+schematic_lemma "?p : ~~P & ~~(P --> Q) --> ~~Q"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 subsection {* The following are classically but not constructively valid *}
 
 (*The attempt to prove them terminates quickly!*)
-lemma "?p : ((P-->Q) --> P)  -->  P"
+schematic_lemma "?p : ((P-->Q) --> P)  -->  P"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : (P&Q-->R)  -->  (P-->R) | (Q-->R)"
+schematic_lemma "?p : (P&Q-->R)  -->  (P-->R) | (Q-->R)"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
@@ -70,74 +70,74 @@
 subsection {* Intuitionistic FOL: propositional problems based on Pelletier *}
 
 text "Problem ~~1"
-lemma "?p : ~~((P-->Q)  <->  (~Q --> ~P))"
+schematic_lemma "?p : ~~((P-->Q)  <->  (~Q --> ~P))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~2"
-lemma "?p : ~~(~~P  <->  P)"
+schematic_lemma "?p : ~~(~~P  <->  P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem 3"
-lemma "?p : ~(P-->Q) --> (Q-->P)"
+schematic_lemma "?p : ~(P-->Q) --> (Q-->P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~4"
-lemma "?p : ~~((~P-->Q)  <->  (~Q --> P))"
+schematic_lemma "?p : ~~((~P-->Q)  <->  (~Q --> P))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~5"
-lemma "?p : ~~((P|Q-->P|R) --> P|(Q-->R))"
+schematic_lemma "?p : ~~((P|Q-->P|R) --> P|(Q-->R))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~6"
-lemma "?p : ~~(P | ~P)"
+schematic_lemma "?p : ~~(P | ~P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~7"
-lemma "?p : ~~(P | ~~~P)"
+schematic_lemma "?p : ~~(P | ~~~P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~8.  Peirce's law"
-lemma "?p : ~~(((P-->Q) --> P)  -->  P)"
+schematic_lemma "?p : ~~(((P-->Q) --> P)  -->  P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem 9"
-lemma "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
+schematic_lemma "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem 10"
-lemma "?p : (Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
+schematic_lemma "?p : (Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "11.  Proved in each direction (incorrectly, says Pelletier!!) "
-lemma "?p : P<->P"
+schematic_lemma "?p : P<->P"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~12.  Dijkstra's law  "
-lemma "?p : ~~(((P <-> Q) <-> R)  <->  (P <-> (Q <-> R)))"
+schematic_lemma "?p : ~~(((P <-> Q) <-> R)  <->  (P <-> (Q <-> R)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ((P <-> Q) <-> R)  -->  ~~(P <-> (Q <-> R))"
+schematic_lemma "?p : ((P <-> Q) <-> R)  -->  ~~(P <-> (Q <-> R))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem 13.  Distributive law"
-lemma "?p : P | (Q & R)  <-> (P | Q) & (P | R)"
+schematic_lemma "?p : P | (Q & R)  <-> (P | Q) & (P | R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~14"
-lemma "?p : ~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
+schematic_lemma "?p : ~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~15"
-lemma "?p : ~~((P --> Q) <-> (~P | Q))"
+schematic_lemma "?p : ~~((P --> Q) <-> (~P | Q))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~16"
-lemma "?p : ~~((P-->Q) | (Q-->P))"
+schematic_lemma "?p : ~~((P-->Q) | (Q-->P))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~17"
-lemma "?p : ~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
+schematic_lemma "?p : ~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
   by (tactic {* IntPr.fast_tac 1 *})  -- slow
 
 
@@ -145,43 +145,43 @@
 
 text "The converse is classical in the following implications..."
 
-lemma "?p : (EX x. P(x)-->Q)  -->  (ALL x. P(x)) --> Q"
+schematic_lemma "?p : (EX x. P(x)-->Q)  -->  (ALL x. P(x)) --> Q"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
+schematic_lemma "?p : ((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ((ALL x. ~P(x))-->Q)  -->  ~ (ALL x. ~ (P(x)|Q))"
+schematic_lemma "?p : ((ALL x. ~P(x))-->Q)  -->  ~ (ALL x. ~ (P(x)|Q))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (ALL x. P(x)) | Q  -->  (ALL x. P(x) | Q)"
+schematic_lemma "?p : (ALL x. P(x)) | Q  -->  (ALL x. P(x) | Q)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
+schematic_lemma "?p : (EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "The following are not constructively valid!"
 text "The attempt to prove them terminates quickly!"
 
-lemma "?p : ((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
+schematic_lemma "?p : ((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : (P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
+schematic_lemma "?p : (P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : (ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
+schematic_lemma "?p : (ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : (ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
+schematic_lemma "?p : (ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
 (*Classically but not intuitionistically valid.  Proved by a bug in 1986!*)
-lemma "?p : EX x. Q(x) --> (ALL x. Q(x))"
+schematic_lemma "?p : EX x. Q(x) --> (ALL x. Q(x))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
@@ -194,32 +194,32 @@
 *}
 
 text "Problem ~~18"
-lemma "?p : ~~(EX y. ALL x. P(y)-->P(x))" oops
+schematic_lemma "?p : ~~(EX y. ALL x. P(y)-->P(x))" oops
 (*NOT PROVED*)
 
 text "Problem ~~19"
-lemma "?p : ~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))" oops
+schematic_lemma "?p : ~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))" oops
 (*NOT PROVED*)
 
 text "Problem 20"
-lemma "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))      
+schematic_lemma "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))      
     --> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem 21"
-lemma "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))" oops
+schematic_lemma "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))" oops
 (*NOT PROVED*)
 
 text "Problem 22"
-lemma "?p : (ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
+schematic_lemma "?p : (ALL x. P <-> Q(x))  -->  (P <-> (ALL x. Q(x)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem ~~23"
-lemma "?p : ~~ ((ALL x. P | Q(x))  <->  (P | (ALL x. Q(x))))"
+schematic_lemma "?p : ~~ ((ALL x. P | Q(x))  <->  (P | (ALL x. Q(x))))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 text "Problem 24"
-lemma "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &   
+schematic_lemma "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &   
      (~(EX x. P(x)) --> (EX x. Q(x))) & (ALL x. Q(x)|R(x) --> S(x))   
     --> ~~(EX x. P(x)&R(x))"
 (*Not clear why fast_tac, best_tac, ASTAR and ITER_DEEPEN all take forever*)
@@ -230,7 +230,7 @@
   done
 
 text "Problem 25"
-lemma "?p : (EX x. P(x)) &   
+schematic_lemma "?p : (EX x. P(x)) &   
         (ALL x. L(x) --> ~ (M(x) & R(x))) &   
         (ALL x. P(x) --> (M(x) & L(x))) &    
         ((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))   
@@ -238,69 +238,69 @@
   by (tactic "IntPr.best_tac 1")
 
 text "Problem 29.  Essentially the same as Principia Mathematica *11.71"
-lemma "?p : (EX x. P(x)) & (EX y. Q(y))   
+schematic_lemma "?p : (EX x. P(x)) & (EX y. Q(y))   
     --> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y))   <->      
          (ALL x y. P(x) & Q(y) --> R(x) & S(y)))"
   by (tactic "IntPr.fast_tac 1")
 
 text "Problem ~~30"
-lemma "?p : (ALL x. (P(x) | Q(x)) --> ~ R(x)) &  
+schematic_lemma "?p : (ALL x. (P(x) | Q(x)) --> ~ R(x)) &  
         (ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))   
     --> (ALL x. ~~S(x))"
   by (tactic "IntPr.fast_tac 1")
 
 text "Problem 31"
-lemma "?p : ~(EX x. P(x) & (Q(x) | R(x))) &  
+schematic_lemma "?p : ~(EX x. P(x) & (Q(x) | R(x))) &  
         (EX x. L(x) & P(x)) &  
         (ALL x. ~ R(x) --> M(x))   
     --> (EX x. L(x) & M(x))"
   by (tactic "IntPr.fast_tac 1")
 
 text "Problem 32"
-lemma "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) &  
+schematic_lemma "?p : (ALL x. P(x) & (Q(x)|R(x))-->S(x)) &  
         (ALL x. S(x) & R(x) --> L(x)) &  
         (ALL x. M(x) --> R(x))   
     --> (ALL x. P(x) & M(x) --> L(x))"
   by (tactic "IntPr.best_tac 1") -- slow
 
 text "Problem 39"
-lemma "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
+schematic_lemma "?p : ~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
   by (tactic "IntPr.best_tac 1")
 
 text "Problem 40.  AMENDED"
-lemma "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) -->   
+schematic_lemma "?p : (EX y. ALL x. F(x,y) <-> F(x,x)) -->   
               ~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"
   by (tactic "IntPr.best_tac 1") -- slow
 
 text "Problem 44"
-lemma "?p : (ALL x. f(x) -->                                    
+schematic_lemma "?p : (ALL x. f(x) -->                                    
               (EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y))))  &        
               (EX x. j(x) & (ALL y. g(y) --> h(x,y)))                    
               --> (EX x. j(x) & ~f(x))"
   by (tactic "IntPr.best_tac 1")
 
 text "Problem 48"
-lemma "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"
+schematic_lemma "?p : (a=b | c=d) & (a=c | b=d) --> a=d | b=c"
   by (tactic "IntPr.best_tac 1")
 
 text "Problem 51"
-lemma
+schematic_lemma
     "?p : (EX z w. ALL x y. P(x,y) <->  (x=z & y=w)) -->   
      (EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"
   by (tactic "IntPr.best_tac 1") -- {*60 seconds*}
 
 text "Problem 56"
-lemma "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
+schematic_lemma "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
   by (tactic "IntPr.best_tac 1")
 
 text "Problem 57"
-lemma
+schematic_lemma
     "?p : P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &  
      (ALL x y z. P(x,y) & P(y,z) --> P(x,z))    -->   P(f(a,b), f(a,c))"
   by (tactic "IntPr.best_tac 1")
 
 text "Problem 60"
-lemma "?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
+schematic_lemma "?p : ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
   by (tactic "IntPr.best_tac 1")
 
 end
--- a/src/FOLP/ex/Nat.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Nat.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -44,7 +44,7 @@
 
 subsection {* Proofs about the natural numbers *}
 
-lemma Suc_n_not_n: "?p : ~ (Suc(k) = k)"
+schematic_lemma Suc_n_not_n: "?p : ~ (Suc(k) = k)"
 apply (rule_tac n = k in induct)
 apply (rule notI)
 apply (erule Suc_neq_0)
@@ -53,7 +53,7 @@
 apply (erule Suc_inject)
 done
 
-lemma "?p : (k+m)+n = k+(m+n)"
+schematic_lemma "?p : (k+m)+n = k+(m+n)"
 apply (rule induct)
 back
 back
@@ -63,23 +63,23 @@
 back
 oops
 
-lemma add_0 [simp]: "?p : 0+n = n"
+schematic_lemma add_0 [simp]: "?p : 0+n = n"
 apply (unfold add_def)
 apply (rule rec_0)
 done
 
-lemma add_Suc [simp]: "?p : Suc(m)+n = Suc(m+n)"
+schematic_lemma add_Suc [simp]: "?p : Suc(m)+n = Suc(m+n)"
 apply (unfold add_def)
 apply (rule rec_Suc)
 done
 
 
-lemma Suc_cong: "p : x = y \<Longrightarrow> ?p : Suc(x) = Suc(y)"
+schematic_lemma Suc_cong: "p : x = y \<Longrightarrow> ?p : Suc(x) = Suc(y)"
   apply (erule subst)
   apply (rule refl)
   done
 
-lemma Plus_cong: "[| p : a = x;  q: b = y |] ==> ?p : a + b = x + y"
+schematic_lemma Plus_cong: "[| p : a = x;  q: b = y |] ==> ?p : a + b = x + y"
   apply (erule subst, erule subst, rule refl)
   done
 
@@ -89,19 +89,19 @@
   val add_ss = FOLP_ss addcongs @{thms nat_congs} addrews [@{thm add_0}, @{thm add_Suc}]
 *}
 
-lemma add_assoc: "?p : (k+m)+n = k+(m+n)"
+schematic_lemma add_assoc: "?p : (k+m)+n = k+(m+n)"
 apply (rule_tac n = k in induct)
 apply (tactic {* SIMP_TAC add_ss 1 *})
 apply (tactic {* ASM_SIMP_TAC add_ss 1 *})
 done
 
-lemma add_0_right: "?p : m+0 = m"
+schematic_lemma add_0_right: "?p : m+0 = m"
 apply (rule_tac n = m in induct)
 apply (tactic {* SIMP_TAC add_ss 1 *})
 apply (tactic {* ASM_SIMP_TAC add_ss 1 *})
 done
 
-lemma add_Suc_right: "?p : m+Suc(n) = Suc(m+n)"
+schematic_lemma add_Suc_right: "?p : m+Suc(n) = Suc(m+n)"
 apply (rule_tac n = m in induct)
 apply (tactic {* ALLGOALS (ASM_SIMP_TAC add_ss) *})
 done
--- a/src/FOLP/ex/Propositional_Cla.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Propositional_Cla.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -11,103 +11,103 @@
 
 
 text "commutative laws of & and | "
-lemma "?p : P & Q  -->  Q & P"
+schematic_lemma "?p : P & Q  -->  Q & P"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : P | Q  -->  Q | P"
+schematic_lemma "?p : P | Q  -->  Q | P"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "associative laws of & and | "
-lemma "?p : (P & Q) & R  -->  P & (Q & R)"
+schematic_lemma "?p : (P & Q) & R  -->  P & (Q & R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (P | Q) | R  -->  P | (Q | R)"
+schematic_lemma "?p : (P | Q) | R  -->  P | (Q | R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "distributive laws of & and | "
-lemma "?p : (P & Q) | R  --> (P | R) & (Q | R)"
+schematic_lemma "?p : (P & Q) | R  --> (P | R) & (Q | R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (P | R) & (Q | R)  --> (P & Q) | R"
+schematic_lemma "?p : (P | R) & (Q | R)  --> (P & Q) | R"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (P | Q) & R  --> (P & R) | (Q & R)"
+schematic_lemma "?p : (P | Q) & R  --> (P & R) | (Q & R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
-lemma "?p : (P & R) | (Q & R)  --> (P | Q) & R"
+schematic_lemma "?p : (P & R) | (Q & R)  --> (P | Q) & R"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "Laws involving implication"
 
-lemma "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
+schematic_lemma "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (P & Q --> R) <-> (P--> (Q-->R))"
+schematic_lemma "?p : (P & Q --> R) <-> (P--> (Q-->R))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
+schematic_lemma "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
+schematic_lemma "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (P --> Q & R) <-> (P-->Q)  &  (P-->R)"
+schematic_lemma "?p : (P --> Q & R) <-> (P-->Q)  &  (P-->R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "Propositions-as-types"
 
 (*The combinator K*)
-lemma "?p : P --> (Q --> P)"
+schematic_lemma "?p : P --> (Q --> P)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 (*The combinator S*)
-lemma "?p : (P-->Q-->R)  --> (P-->Q) --> (P-->R)"
+schematic_lemma "?p : (P-->Q-->R)  --> (P-->Q) --> (P-->R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 (*Converse is classical*)
-lemma "?p : (P-->Q) | (P-->R)  -->  (P --> Q | R)"
+schematic_lemma "?p : (P-->Q) | (P-->R)  -->  (P --> Q | R)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (P-->Q)  -->  (~Q --> ~P)"
+schematic_lemma "?p : (P-->Q)  -->  (~Q --> ~P)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "Schwichtenberg's examples (via T. Nipkow)"
 
-lemma stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
+schematic_lemma stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)  
+schematic_lemma stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)  
               --> ((P --> Q) --> P) --> P"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)  
+schematic_lemma peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)  
                --> (((P --> Q) --> R) --> P --> Q) --> P --> Q"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
   
-lemma peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
+schematic_lemma peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
+schematic_lemma mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
+schematic_lemma mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)  
+schematic_lemma tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)  
           --> (((P8 --> P2) --> P9) --> P3 --> P10)  
           --> (P1 --> P8) --> P6 --> P7  
           --> (((P3 --> P2) --> P9) --> P4)  
           --> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)  
+schematic_lemma tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)  
      --> (((P3 --> P2) --> P9) --> P4)  
      --> (((P6 --> P1) --> P2) --> P9)  
      --> (((P7 --> P1) --> P10) --> P4 --> P5)  
--- a/src/FOLP/ex/Propositional_Int.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Propositional_Int.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -11,103 +11,103 @@
 
 
 text "commutative laws of & and | "
-lemma "?p : P & Q  -->  Q & P"
+schematic_lemma "?p : P & Q  -->  Q & P"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : P | Q  -->  Q | P"
+schematic_lemma "?p : P | Q  -->  Q | P"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "associative laws of & and | "
-lemma "?p : (P & Q) & R  -->  P & (Q & R)"
+schematic_lemma "?p : (P & Q) & R  -->  P & (Q & R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P | Q) | R  -->  P | (Q | R)"
+schematic_lemma "?p : (P | Q) | R  -->  P | (Q | R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "distributive laws of & and | "
-lemma "?p : (P & Q) | R  --> (P | R) & (Q | R)"
+schematic_lemma "?p : (P & Q) | R  --> (P | R) & (Q | R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P | R) & (Q | R)  --> (P & Q) | R"
+schematic_lemma "?p : (P | R) & (Q | R)  --> (P & Q) | R"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P | Q) & R  --> (P & R) | (Q & R)"
+schematic_lemma "?p : (P | Q) & R  --> (P & R) | (Q & R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
-lemma "?p : (P & R) | (Q & R)  --> (P | Q) & R"
+schematic_lemma "?p : (P & R) | (Q & R)  --> (P | Q) & R"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "Laws involving implication"
 
-lemma "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
+schematic_lemma "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P & Q --> R) <-> (P--> (Q-->R))"
+schematic_lemma "?p : (P & Q --> R) <-> (P--> (Q-->R))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
+schematic_lemma "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
+schematic_lemma "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P --> Q & R) <-> (P-->Q)  &  (P-->R)"
+schematic_lemma "?p : (P --> Q & R) <-> (P-->Q)  &  (P-->R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "Propositions-as-types"
 
 (*The combinator K*)
-lemma "?p : P --> (Q --> P)"
+schematic_lemma "?p : P --> (Q --> P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 (*The combinator S*)
-lemma "?p : (P-->Q-->R)  --> (P-->Q) --> (P-->R)"
+schematic_lemma "?p : (P-->Q-->R)  --> (P-->Q) --> (P-->R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 (*Converse is classical*)
-lemma "?p : (P-->Q) | (P-->R)  -->  (P --> Q | R)"
+schematic_lemma "?p : (P-->Q) | (P-->R)  -->  (P --> Q | R)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (P-->Q)  -->  (~Q --> ~P)"
+schematic_lemma "?p : (P-->Q)  -->  (~Q --> ~P)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "Schwichtenberg's examples (via T. Nipkow)"
 
-lemma stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
+schematic_lemma stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)  
+schematic_lemma stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)  
               --> ((P --> Q) --> P) --> P"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)  
+schematic_lemma peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)  
                --> (((P --> Q) --> R) --> P --> Q) --> P --> Q"
   by (tactic {* IntPr.fast_tac 1 *})
   
-lemma peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
+schematic_lemma peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
+schematic_lemma mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
+schematic_lemma mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)  
+schematic_lemma tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)  
           --> (((P8 --> P2) --> P9) --> P3 --> P10)  
           --> (P1 --> P8) --> P6 --> P7  
           --> (((P3 --> P2) --> P9) --> P4)  
           --> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)  
+schematic_lemma tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)  
      --> (((P3 --> P2) --> P9) --> P4)  
      --> (((P6 --> P1) --> P2) --> P9)  
      --> (((P7 --> P1) --> P10) --> P4 --> P5)  
--- a/src/FOLP/ex/Quantifiers_Cla.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Quantifiers_Cla.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -10,92 +10,92 @@
 imports FOLP
 begin
 
-lemma "?p : (ALL x y. P(x,y))  -->  (ALL y x. P(x,y))"
+schematic_lemma "?p : (ALL x y. P(x,y))  -->  (ALL y x. P(x,y))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))"
+schematic_lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 (*Converse is false*)
-lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
+schematic_lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (ALL x. P-->Q(x))  <->  (P--> (ALL x. Q(x)))"
+schematic_lemma "?p : (ALL x. P-->Q(x))  <->  (P--> (ALL x. Q(x)))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
-lemma "?p : (ALL x. P(x)-->Q)  <->  ((EX x. P(x)) --> Q)"
+schematic_lemma "?p : (ALL x. P(x)-->Q)  <->  ((EX x. P(x)) --> Q)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "Some harder ones"
 
-lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
+schematic_lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 (*Converse is false*)
-lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x))  &  (EX x. Q(x))"
+schematic_lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x))  &  (EX x. Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "Basic test of quantifier reasoning"
 (*TRUE*)
-lemma "?p : (EX y. ALL x. Q(x,y)) -->  (ALL x. EX y. Q(x,y))"
+schematic_lemma "?p : (EX y. ALL x. Q(x,y)) -->  (ALL x. EX y. Q(x,y))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
+schematic_lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "The following should fail, as they are false!"
 
-lemma "?p : (ALL x. EX y. Q(x,y))  -->  (EX y. ALL x. Q(x,y))"
+schematic_lemma "?p : (ALL x. EX y. Q(x,y))  -->  (EX y. ALL x. Q(x,y))"
   apply (tactic {* Cla.fast_tac FOLP_cs 1 *})?
   oops
 
-lemma "?p : (EX x. Q(x))  -->  (ALL x. Q(x))"
+schematic_lemma "?p : (EX x. Q(x))  -->  (ALL x. Q(x))"
   apply (tactic {* Cla.fast_tac FOLP_cs 1 *})?
   oops
 
-lemma "?p : P(?a) --> (ALL x. P(x))"
+schematic_lemma "?p : P(?a) --> (ALL x. P(x))"
   apply (tactic {* Cla.fast_tac FOLP_cs 1 *})?
   oops
 
-lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
+schematic_lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
   apply (tactic {* Cla.fast_tac FOLP_cs 1 *})?
   oops
 
 
 text "Back to things that are provable..."
 
-lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
+schematic_lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 (*An example of why exI should be delayed as long as possible*)
-lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
+schematic_lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
+schematic_lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
-lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
+schematic_lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 
 text "Some slow ones"
 
 (*Principia Mathematica *11.53  *)
-lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
+schematic_lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 (*Principia Mathematica *11.55  *)
-lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"
+schematic_lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 (*Principia Mathematica *11.61  *)
-lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"
+schematic_lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"
   by (tactic {* Cla.fast_tac FOLP_cs 1 *})
 
 end
--- a/src/FOLP/ex/Quantifiers_Int.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/FOLP/ex/Quantifiers_Int.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -10,92 +10,92 @@
 imports IFOLP
 begin
 
-lemma "?p : (ALL x y. P(x,y))  -->  (ALL y x. P(x,y))"
+schematic_lemma "?p : (ALL x y. P(x,y))  -->  (ALL y x. P(x,y))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))"
+schematic_lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 (*Converse is false*)
-lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
+schematic_lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (ALL x. P-->Q(x))  <->  (P--> (ALL x. Q(x)))"
+schematic_lemma "?p : (ALL x. P-->Q(x))  <->  (P--> (ALL x. Q(x)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
-lemma "?p : (ALL x. P(x)-->Q)  <->  ((EX x. P(x)) --> Q)"
+schematic_lemma "?p : (ALL x. P(x)-->Q)  <->  ((EX x. P(x)) --> Q)"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "Some harder ones"
 
-lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
+schematic_lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 (*Converse is false*)
-lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x))  &  (EX x. Q(x))"
+schematic_lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x))  &  (EX x. Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "Basic test of quantifier reasoning"
 (*TRUE*)
-lemma "?p : (EX y. ALL x. Q(x,y)) -->  (ALL x. EX y. Q(x,y))"
+schematic_lemma "?p : (EX y. ALL x. Q(x,y)) -->  (ALL x. EX y. Q(x,y))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
+schematic_lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "The following should fail, as they are false!"
 
-lemma "?p : (ALL x. EX y. Q(x,y))  -->  (EX y. ALL x. Q(x,y))"
+schematic_lemma "?p : (ALL x. EX y. Q(x,y))  -->  (EX y. ALL x. Q(x,y))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : (EX x. Q(x))  -->  (ALL x. Q(x))"
+schematic_lemma "?p : (EX x. Q(x))  -->  (ALL x. Q(x))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : P(?a) --> (ALL x. P(x))"
+schematic_lemma "?p : P(?a) --> (ALL x. P(x))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
-lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
+schematic_lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
   apply (tactic {* IntPr.fast_tac 1 *})?
   oops
 
 
 text "Back to things that are provable..."
 
-lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
+schematic_lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 (*An example of why exI should be delayed as long as possible*)
-lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
+schematic_lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
+schematic_lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
   by (tactic {* IntPr.fast_tac 1 *})
 
-lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
+schematic_lemma "?p : (ALL x. Q(x))  -->  (EX x. Q(x))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 
 text "Some slow ones"
 
 (*Principia Mathematica *11.53  *)
-lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
+schematic_lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 (*Principia Mathematica *11.55  *)
-lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"
+schematic_lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 (*Principia Mathematica *11.61  *)
-lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"
+schematic_lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"
   by (tactic {* IntPr.fast_tac 1 *})
 
 end
--- a/src/HOL/Bali/DeclConcepts.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Bali/DeclConcepts.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -1390,7 +1390,7 @@
 "accimethds G pack I
   \<equiv> if G\<turnstile>Iface I accessible_in pack 
        then imethds G I
-       else \<lambda> k. {}"
+       else (\<lambda> k. {})"
 text {* only returns imethds if the interface is accessible *}
 
 definition methd :: "prog \<Rightarrow> qtname  \<Rightarrow> (sig,qtname \<times> methd) table" where
--- a/src/HOL/Bali/Example.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Bali/Example.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -1070,7 +1070,7 @@
 
 section "well-typedness"
 
-lemma wt_test: "\<lparr>prg=tprg,cls=Main,lcl=empty(VName e\<mapsto>Class Base)\<rparr>\<turnstile>test ?pTs\<Colon>\<surd>"
+schematic_lemma wt_test: "\<lparr>prg=tprg,cls=Main,lcl=empty(VName e\<mapsto>Class Base)\<rparr>\<turnstile>test ?pTs\<Colon>\<surd>"
 apply (unfold test_def arr_viewed_from_def)
 (* ?pTs = [Class Base] *)
 apply (rule wtIs (* ;; *))
@@ -1122,7 +1122,7 @@
 
 section "definite assignment"
 
-lemma da_test: "\<lparr>prg=tprg,cls=Main,lcl=empty(VName e\<mapsto>Class Base)\<rparr>
+schematic_lemma da_test: "\<lparr>prg=tprg,cls=Main,lcl=empty(VName e\<mapsto>Class Base)\<rparr>
                   \<turnstile>{} \<guillemotright>\<langle>test ?pTs\<rangle>\<guillemotright> \<lparr>nrm={VName e},brk=\<lambda> l. UNIV\<rparr>"
 apply (unfold test_def arr_viewed_from_def)
 apply (rule da.Comp)
@@ -1241,7 +1241,7 @@
 
 
 declare Pair_eq [simp del]
-lemma exec_test: 
+schematic_lemma exec_test: 
 "\<lbrakk>the (new_Addr (heap  s1)) = a;  
   the (new_Addr (heap ?s2)) = b;  
   the (new_Addr (heap ?s3)) = c\<rbrakk> \<Longrightarrow>  
--- a/src/HOL/Bali/WellType.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Bali/WellType.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -94,7 +94,7 @@
 "accObjectmheads G S T
    \<equiv> if G\<turnstile>RefT T accessible_in (pid S)
         then Objectmheads G S
-        else \<lambda>sig. {}"
+        else (\<lambda>sig. {})"
 primrec
 "mheads G S  NullT     = (\<lambda>sig. {})"
 "mheads G S (IfaceT I) = (\<lambda>sig. (\<lambda>(I,h).(IfaceT I,h)) 
--- a/src/HOL/Big_Operators.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Big_Operators.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -555,7 +555,7 @@
 qed
 
 lemma setsum_mono2:
-fixes f :: "'a \<Rightarrow> 'b :: {ordered_ab_semigroup_add_imp_le,comm_monoid_add}"
+fixes f :: "'a \<Rightarrow> 'b :: ordered_comm_monoid_add"
 assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
 shows "setsum f A \<le> setsum f B"
 proof -
@@ -1030,15 +1030,15 @@
 lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
   (setprod f (A - {a}) :: 'a :: {field}) =
   (if a:A then setprod f A / f a else setprod f A)"
-by (erule finite_induct) (auto simp add: insert_Diff_if)
+  by (erule finite_induct) (auto simp add: insert_Diff_if)
 
 lemma setprod_inversef: 
-  fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+  fixes f :: "'b \<Rightarrow> 'a::{field,division_ring_inverse_zero}"
   shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
 by (erule finite_induct) auto
 
 lemma setprod_dividef:
-  fixes f :: "'b \<Rightarrow> 'a::{field,division_by_zero}"
+  fixes f :: "'b \<Rightarrow> 'a::{field,division_ring_inverse_zero}"
   shows "finite A
     ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
 apply (subgoal_tac
@@ -1140,7 +1140,7 @@
       using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
       by simp
     then have ?thesis using a cA
-      by (simp add: fA1 ring_simps cong add: setprod_cong cong del: if_weak_cong)}
+      by (simp add: fA1 field_simps cong add: setprod_cong cong del: if_weak_cong)}
   ultimately show ?thesis by blast
 qed
 
--- a/src/HOL/Boogie/Tools/boogie_commands.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -99,7 +99,7 @@
     |> (fn ctxt1 => ctxt1
     |> prepare
     |-> (fn us => fn ctxt2 => ctxt2
-    |> Proof.theorem_i NONE (fn thmss => fn ctxt =>
+    |> Proof.theorem NONE (fn thmss => fn ctxt =>
          let val export = map (finish ctxt1) o ProofContext.export ctxt ctxt2
          in after_qed (map export thmss) ctxt end) [map (rpair []) us]))
   end
@@ -188,7 +188,7 @@
 
   fun prove thy meth vc =
     ProofContext.init thy
-    |> Proof.theorem_i NONE (K I) [[(Boogie_VCs.prop_of_vc vc, [])]]
+    |> Proof.theorem NONE (K I) [[(Boogie_VCs.prop_of_vc vc, [])]]
     |> Proof.apply meth
     |> Seq.hd
     |> Proof.global_done_proof
--- a/src/HOL/Complete_Lattice.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Complete_Lattice.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -98,9 +98,9 @@
 
 syntax
   "_SUP1"     :: "pttrns => 'b => 'b"           ("(3SUP _./ _)" [0, 10] 10)
-  "_SUP"      :: "pttrn => 'a set => 'b => 'b"  ("(3SUP _:_./ _)" [0, 10] 10)
+  "_SUP"      :: "pttrn => 'a set => 'b => 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
   "_INF1"     :: "pttrns => 'b => 'b"           ("(3INF _./ _)" [0, 10] 10)
-  "_INF"      :: "pttrn => 'a set => 'b => 'b"  ("(3INF _:_./ _)" [0, 10] 10)
+  "_INF"      :: "pttrn => 'a set => 'b => 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
 
 translations
   "SUP x y. B"   == "SUP x. SUP y. B"
@@ -295,15 +295,15 @@
 
 syntax
   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3UN _./ _)" [0, 10] 10)
-  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3UN _:_./ _)" [0, 10] 10)
+  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3UN _:_./ _)" [0, 0, 10] 10)
 
 syntax (xsymbols)
   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>_./ _)" [0, 10] 10)
-  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>_\<in>_./ _)" [0, 10] 10)
+  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>_\<in>_./ _)" [0, 0, 10] 10)
 
 syntax (latex output)
   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
-  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
+  "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
 
 translations
   "UN x y. B"   == "UN x. UN y. B"
@@ -531,15 +531,15 @@
 
 syntax
   "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3INT _./ _)" [0, 10] 10)
-  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 10] 10)
+  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 0, 10] 10)
 
 syntax (xsymbols)
   "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
-  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 10] 10)
+  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 0, 10] 10)
 
 syntax (latex output)
   "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
-  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
+  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
 
 translations
   "INT x y. B"  == "INT x. INT y. B"
--- a/src/HOL/Complex.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Complex.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -99,7 +99,7 @@
 
 subsection {* Multiplication and Division *}
 
-instantiation complex :: "{field, division_by_zero}"
+instantiation complex :: "{field, division_ring_inverse_zero}"
 begin
 
 definition
--- a/src/HOL/Decision_Procs/Decision_Procs.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Decision_Procs/Decision_Procs.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -8,4 +8,4 @@
   "ex/Commutative_Ring_Ex" "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex"
 begin
 
-end
\ No newline at end of file
+end
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -27,7 +27,7 @@
   "tmsize (CNP n c a) = 3 + polysize c + tmsize a "
 
   (* Semantics of terms tm *)
-consts Itm :: "'a::{ring_char_0,division_by_zero,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
+consts Itm :: "'a::{ring_char_0,division_ring_inverse_zero,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
 primrec
   "Itm vs bs (CP c) = (Ipoly vs c)"
   "Itm vs bs (Bound n) = bs!n"
@@ -239,7 +239,7 @@
 lemma tmadd[simp]: "Itm vs bs (tmadd (t,s)) = Itm vs bs (Add t s)"
 apply (induct t s rule: tmadd.induct, simp_all add: Let_def)
 apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \<le> n2", simp_all)
-apply (case_tac "n1 = n2", simp_all add: ring_simps)
+apply (case_tac "n1 = n2", simp_all add: field_simps)
 apply (simp only: right_distrib[symmetric]) 
 by (auto simp del: polyadd simp add: polyadd[symmetric])
 
@@ -259,7 +259,7 @@
   "tmmul t = (\<lambda> i. Mul i t)"
 
 lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)"
-by (induct t arbitrary: i rule: tmmul.induct, simp_all add: ring_simps)
+by (induct t arbitrary: i rule: tmmul.induct, simp_all add: field_simps)
 
 lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
 by (induct t arbitrary: i rule: tmmul.induct, auto )
@@ -270,7 +270,7 @@
 by (induct t arbitrary: i rule: tmmul.induct, auto simp add: Let_def)
 
 lemma tmmul_allpolys_npoly[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
 
 definition tmneg :: "tm \<Rightarrow> tm" where
@@ -296,7 +296,7 @@
 using tmneg_def by simp
 lemma [simp]: "isnpoly (C (-1,1))" unfolding isnpoly_def by simp
 lemma tmneg_allpolys_npoly[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)" 
   unfolding tmneg_def by auto
 
@@ -310,7 +310,7 @@
 lemma tmsub_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmsub t s )"
 using tmsub_def by simp
 lemma tmsub_allpolys_npoly[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)" 
   unfolding tmsub_def by (simp add: isnpoly_def)
 
@@ -324,8 +324,8 @@
   "simptm (CNP n c t) = (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p ), simptm t))"
 
 lemma polynate_stupid: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
-  shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a::{ring_char_0,division_by_zero, field})" 
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
+  shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a::{ring_char_0,division_ring_inverse_zero, field})" 
 apply (subst polynate[symmetric])
 apply simp
 done
@@ -345,7 +345,7 @@
 lemma [simp]: "isnpoly 0\<^sub>p" and [simp]: "isnpoly (C(1,1))" 
   by (simp_all add: isnpoly_def)
 lemma simptm_allpolys_npoly[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows "allpolys isnpoly (simptm p)"
   by (induct p rule: simptm.induct, auto simp add: Let_def)
 
@@ -369,14 +369,14 @@
   "tmbound 0 (snd (split0 t)) \<and> (Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t)"
   apply (induct t rule: split0.induct)
   apply simp
-  apply (simp add: Let_def split_def ring_simps)
-  apply (simp add: Let_def split_def ring_simps)
-  apply (simp add: Let_def split_def ring_simps)
-  apply (simp add: Let_def split_def ring_simps)
-  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def field_simps)
+  apply (simp add: Let_def split_def field_simps)
+  apply (simp add: Let_def split_def field_simps)
+  apply (simp add: Let_def split_def field_simps)
+  apply (simp add: Let_def split_def field_simps)
   apply (simp add: Let_def split_def mult_assoc right_distrib[symmetric])
-  apply (simp add: Let_def split_def ring_simps)
-  apply (simp add: Let_def split_def ring_simps)
+  apply (simp add: Let_def split_def field_simps)
+  apply (simp add: Let_def split_def field_simps)
   done
 
 lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
@@ -387,7 +387,7 @@
 qed
 
 lemma split0_nb0: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows "split0 t = (c',t') \<Longrightarrow>  tmbound 0 t'"
 proof-
   fix c' t'
@@ -395,7 +395,7 @@
   with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'" by simp
 qed
 
-lemma split0_nb0'[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+lemma split0_nb0'[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows "tmbound0 (snd (split0 t))"
   using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"] by (simp add: tmbound0_tmbound_iff)
 
@@ -418,7 +418,7 @@
 lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
 by (induct p rule: split0.induct, auto simp  add: isnpoly_def Let_def split_def split0_stupid)
 
-lemma isnpoly_fst_split0:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+lemma isnpoly_fst_split0:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero, field})"
   shows 
   "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
   by (induct p rule: split0.induct, 
@@ -447,7 +447,7 @@
 by (induct p rule: fmsize.induct) simp_all
 
   (* Semantics of formulae (fm) *)
-consts Ifm ::"'a::{division_by_zero,linordered_field} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
+consts Ifm ::"'a::{division_ring_inverse_zero,linordered_field} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
 primrec
   "Ifm vs bs T = True"
   "Ifm vs bs F = False"
@@ -969,24 +969,24 @@
 definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))"
 definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))"
 
-lemma simplt_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simplt_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "islin (simplt t)"
   unfolding simplt_def 
   using split0_nb0'
 by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
   
-lemma simple_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simple_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "islin (simple t)"
   unfolding simple_def 
   using split0_nb0'
 by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
-lemma simpeq_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpeq_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "islin (simpeq t)"
   unfolding simpeq_def 
   using split0_nb0'
 by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
 
-lemma simpneq_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpneq_islin[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "islin (simpneq t)"
   unfolding simpneq_def 
   using split0_nb0'
@@ -994,7 +994,7 @@
 
 lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
   by (cases "split0 s", auto)
-lemma split0_npoly:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma split0_npoly:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   and n: "allpolys isnpoly t"
   shows "isnpoly (fst (split0 t))" and "allpolys isnpoly (snd (split0 t))"
   using n
@@ -1083,7 +1083,7 @@
   apply (case_tac poly, auto)
   done
 
-lemma simplt_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simplt_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
   using split0 [of "simptm t" vs bs]
 proof(simp add: simplt_def Let_def split_def)
@@ -1100,7 +1100,7 @@
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def lt_nb)
 qed
 
-lemma simple_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simple_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
   using split0 [of "simptm t" vs bs]
 proof(simp add: simple_def Let_def split_def)
@@ -1117,7 +1117,7 @@
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def le_nb)
 qed
 
-lemma simpeq_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpeq_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
   using split0 [of "simptm t" vs bs]
 proof(simp add: simpeq_def Let_def split_def)
@@ -1134,7 +1134,7 @@
        fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpeq_def Let_def split_def eq_nb)
 qed
 
-lemma simpneq_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpneq_nb[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
   using split0 [of "simptm t" vs bs]
 proof(simp add: simpneq_def Let_def split_def)
@@ -1267,7 +1267,7 @@
 lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
 by(induct p arbitrary: bs rule: simpfm.induct, auto)
 
-lemma simpfm_bound0:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpfm_bound0:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "bound0 p \<Longrightarrow> bound0 (simpfm p)"
 by (induct p rule: simpfm.induct, auto)
 
@@ -1296,7 +1296,7 @@
 lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)" by (simp add: disj_def)
 lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)" by (simp add: conj_def)
 
-lemma   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "qfree p \<Longrightarrow> islin (simpfm p)" 
   apply (induct p rule: simpfm.induct)
   apply (simp_all add: conj_lin disj_lin)
@@ -1698,11 +1698,11 @@
   {assume c: "?N c > 0"
       from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]  
       have px': "x < - ?Nt x s / ?N c" 
-        by (auto simp add: not_less ring_simps) 
+        by (auto simp add: not_less field_simps) 
     {assume y: "y < - ?Nt x s / ?N c" 
       hence "y * ?N c < - ?Nt x s"
         by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
-      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
     {assume y: "y > -?Nt x s / ?N c" 
@@ -1715,11 +1715,11 @@
   {assume c: "?N c < 0"
       from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]  
       have px': "x > - ?Nt x s / ?N c" 
-        by (auto simp add: not_less ring_simps) 
+        by (auto simp add: not_less field_simps) 
     {assume y: "y > - ?Nt x s / ?N c" 
       hence "y * ?N c < - ?Nt x s"
         by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
-      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
     {assume y: "y < -?Nt x s / ?N c" 
@@ -1743,11 +1743,11 @@
   moreover
   {assume c: "?N c > 0"
       from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"]  
-      have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less ring_simps) 
+      have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less field_simps) 
     {assume y: "y < - ?Nt x s / ?N c" 
       hence "y * ?N c < - ?Nt x s"
         by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
-      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
     {assume y: "y > -?Nt x s / ?N c" 
@@ -1759,11 +1759,11 @@
   moreover
   {assume c: "?N c < 0"
       from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"]  
-      have px': "x >= - ?Nt x s / ?N c" by (simp add: ring_simps) 
+      have px': "x >= - ?Nt x s / ?N c" by (simp add: field_simps) 
     {assume y: "y > - ?Nt x s / ?N c" 
       hence "y * ?N c < - ?Nt x s"
         by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
-      hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+      hence "?N c * y + ?Nt x s < 0" by (simp add: field_simps)
       hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
     moreover
     {assume y: "y < -?Nt x s / ?N c" 
@@ -1787,7 +1787,7 @@
   moreover
   {assume c: "?N c > 0" hence cnz: "?N c \<noteq> 0" by simp
     from px eq_divide_eq[of "x" "-?Nt x s" "?N c"]  cnz
-    have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+    have px': "x = - ?Nt x s / ?N c" by (simp add: field_simps)
     {assume y: "y < -?Nt x s / ?N c" 
       with ly have eu: "l < - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
@@ -1802,7 +1802,7 @@
   moreover
   {assume c: "?N c < 0" hence cnz: "?N c \<noteq> 0" by simp
     from px eq_divide_eq[of "x" "-?Nt x s" "?N c"]  cnz
-    have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+    have px': "x = - ?Nt x s / ?N c" by (simp add: field_simps)
     {assume y: "y < -?Nt x s / ?N c" 
       with ly have eu: "l < - ?Nt x s / ?N c" by auto
       with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
@@ -1829,7 +1829,7 @@
   moreover
   {assume c: "?N c \<noteq> 0"
     from yne c eq_divide_eq[of "y" "- ?Nt x s" "?N c"] have ?case
-      by (simp add: ring_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) }
+      by (simp add: field_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) }
   ultimately show ?case by blast
 qed (auto simp add: nth_pos2 tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"] bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"])
 
@@ -1844,7 +1844,7 @@
 
 lemma half_sum_eq: "(u + u) / (1+1) = (u::'a::{linordered_field})" 
 proof-
-  have "(u + u) = (1 + 1) * u" by (simp add: ring_simps)
+  have "(u + u) = (1 + 1) * u" by (simp add: field_simps)
   hence "(u + u) / (1+1) = (1 + 1)*u / (1 + 1)" by simp
   with nonzero_mult_divide_cancel_left[OF one_plus_one_nonzero, of u] show ?thesis by simp
 qed
@@ -1987,7 +1987,7 @@
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) = 0" 
       using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r= 0"
-      by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+      by (simp add: field_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
     
     also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r = 0" using d by simp 
     finally have ?thesis using c d 
@@ -2003,7 +2003,7 @@
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) = 0" 
       using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r= 0"
-      by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+      by (simp add: field_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
     also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r = 0" using c by simp 
     finally have ?thesis using c d 
       apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two)
@@ -2014,19 +2014,19 @@
   {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
     have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r = 0" 
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
     also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) =0 "
       using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r =0" 
-      using nonzero_mult_divide_cancel_left[OF dc] c d
-      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+      using nonzero_mult_divide_cancel_left [OF dc] c d
+      by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
     finally  have ?thesis using c d 
-      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex ring_simps)
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex field_simps)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
-      apply (simp add: ring_simps)
+      apply (simp add: field_simps)
       done }
   ultimately show ?thesis by blast
 qed
@@ -2075,7 +2075,7 @@
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) \<noteq> 0" 
       using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r\<noteq> 0"
-      by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+      by (simp add: field_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
     
     also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r \<noteq> 0" using d by simp 
     finally have ?thesis using c d 
@@ -2091,7 +2091,7 @@
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) \<noteq> 0" 
       using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
     also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r \<noteq> 0"
-      by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+      by (simp add: field_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
     also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r \<noteq> 0" using c by simp 
     finally have ?thesis using c d 
       apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two)
@@ -2102,7 +2102,7 @@
   {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
     have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r \<noteq> 0" 
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2110,11 +2110,11 @@
       using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r \<noteq> 0" 
       using nonzero_mult_divide_cancel_left[OF dc] c d
-      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
     finally  have ?thesis using c d 
-      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex ring_simps)
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex field_simps)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
-      apply (simp add: ring_simps)
+      apply (simp add: field_simps)
       done }
   ultimately show ?thesis by blast
 qed
@@ -2169,7 +2169,7 @@
     from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
     have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0" 
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2178,11 +2178,11 @@
       using dc' dc'' mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r < 0" 
       using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
-      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
     finally  have ?thesis using dc c d  nc nd dc'
-      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) 
     apply (simp only: one_add_one_is_two[symmetric] of_int_add)
-    by (simp add: ring_simps order_less_not_sym[OF dc])}
+    by (simp add: field_simps order_less_not_sym[OF dc])}
   moreover
   {assume dc: "?c*?d < 0" 
 
@@ -2191,7 +2191,7 @@
     hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
     have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0" 
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2201,78 +2201,78 @@
       using dc' order_less_not_sym[OF dc'] mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
     also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r < 0" 
       using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
-      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
     finally  have ?thesis using dc c d  nc nd
-      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) 
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
-      by (simp add: ring_simps order_less_not_sym[OF dc]) }
+      by (simp add: field_simps order_less_not_sym[OF dc]) }
   moreover
   {assume c: "?c > 0" and d: "?d=0"  
     from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
     from c have c': "(1 + 1)*?c \<noteq> 0" by simp
-    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) < 0"
       using c mult_less_cancel_left_disj[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
     also have "\<dots> \<longleftrightarrow> - ?a*?t+  (1 + 1)*?c *?r < 0" 
       using nonzero_mult_divide_cancel_left[OF c'] c
-      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using c order_less_not_sym[OF c] less_imp_neq[OF c]
-      by (simp add: ring_simps )  }
+      by (simp add: field_simps )  }
   moreover
   {assume c: "?c < 0" and d: "?d=0"  hence c': "(1 + 1)*?c \<noteq> 0" by simp
     from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
-    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) > 0"
       using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
     also have "\<dots> \<longleftrightarrow> ?a*?t -  (1 + 1)*?c *?r < 0" 
       using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
-        by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+        by (simp add: algebra_simps diff_divide_distrib del:  left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using c order_less_not_sym[OF c] less_imp_neq[OF c]
-      by (simp add: ring_simps )    }
+      by (simp add: field_simps )    }
   moreover
   moreover
   {assume c: "?c = 0" and d: "?d>0"  
     from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
     from d have d': "(1 + 1)*?d \<noteq> 0" by simp
-    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) < 0"
       using d mult_less_cancel_left_disj[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
     also have "\<dots> \<longleftrightarrow> - ?a*?s+  (1 + 1)*?d *?r < 0" 
       using nonzero_mult_divide_cancel_left[OF d'] d
-      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using d order_less_not_sym[OF d] less_imp_neq[OF d]
-      by (simp add: ring_simps )  }
+      by (simp add: field_simps)  }
   moreover
   {assume c: "?c = 0" and d: "?d<0"  hence d': "(1 + 1)*?d \<noteq> 0" by simp
     from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
-    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) > 0"
       using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
     also have "\<dots> \<longleftrightarrow> ?a*?s -  (1 + 1)*?d *?r < 0" 
       using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
-        by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+        by (simp add: algebra_simps diff_divide_distrib del:  left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using d order_less_not_sym[OF d] less_imp_neq[OF d]
-      by (simp add: ring_simps )    }
+      by (simp add: field_simps )    }
 ultimately show ?thesis by blast
 qed
 
@@ -2325,7 +2325,7 @@
     from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
     have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0" 
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2334,11 +2334,11 @@
       using dc' dc'' mult_le_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r <= 0" 
       using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
-      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
     finally  have ?thesis using dc c d  nc nd dc'
-      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) 
     apply (simp only: one_add_one_is_two[symmetric] of_int_add)
-    by (simp add: ring_simps order_less_not_sym[OF dc])}
+    by (simp add: field_simps order_less_not_sym[OF dc])}
   moreover
   {assume dc: "?c*?d < 0" 
 
@@ -2347,7 +2347,7 @@
     hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
     from add_frac_eq[OF c d, of "- ?t" "- ?s"]
     have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" 
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0" 
       by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
@@ -2357,78 +2357,78 @@
       using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
     also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r <= 0" 
       using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
-      by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib del: left_distrib)
     finally  have ?thesis using dc c d  nc nd
-      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) 
+      apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) 
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
-      by (simp add: ring_simps order_less_not_sym[OF dc]) }
+      by (simp add: field_simps order_less_not_sym[OF dc]) }
   moreover
   {assume c: "?c > 0" and d: "?d=0"  
     from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
     from c have c': "(1 + 1)*?c \<noteq> 0" by simp
-    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) <= 0"
       using c mult_le_cancel_left[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
     also have "\<dots> \<longleftrightarrow> - ?a*?t+  (1 + 1)*?c *?r <= 0" 
       using nonzero_mult_divide_cancel_left[OF c'] c
-      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using c order_less_not_sym[OF c] less_imp_neq[OF c]
-      by (simp add: ring_simps )  }
+      by (simp add: field_simps )  }
   moreover
   {assume c: "?c < 0" and d: "?d=0"  hence c': "(1 + 1)*?c \<noteq> 0" by simp
     from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
-    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: ring_simps)
+    from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) >= 0"
       using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
     also have "\<dots> \<longleftrightarrow> ?a*?t -  (1 + 1)*?c *?r <= 0" 
       using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
-        by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+        by (simp add: algebra_simps diff_divide_distrib del:  left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using c order_less_not_sym[OF c] less_imp_neq[OF c]
-      by (simp add: ring_simps )    }
+      by (simp add: field_simps )    }
   moreover
   moreover
   {assume c: "?c = 0" and d: "?d>0"  
     from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
     from d have d': "(1 + 1)*?d \<noteq> 0" by simp
-    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) <= 0"
       using d mult_le_cancel_left[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
     also have "\<dots> \<longleftrightarrow> - ?a*?s+  (1 + 1)*?d *?r <= 0" 
       using nonzero_mult_divide_cancel_left[OF d'] d
-      by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+      by (simp add: algebra_simps diff_divide_distrib less_le del: left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using d order_less_not_sym[OF d] less_imp_neq[OF d]
-      by (simp add: ring_simps )  }
+      by (simp add: field_simps )  }
   moreover
   {assume c: "?c = 0" and d: "?d<0"  hence d': "(1 + 1)*?d \<noteq> 0" by simp
     from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
-    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: ring_simps)
+    from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)"  by (simp add: field_simps)
     have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
     also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
     also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) >= 0"
       using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
     also have "\<dots> \<longleftrightarrow> ?a*?s -  (1 + 1)*?d *?r <= 0" 
       using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
-        by (simp add: ring_simps diff_divide_distrib del:  left_distrib)
+        by (simp add: algebra_simps diff_divide_distrib del:  left_distrib)
     finally have ?thesis using c d nc nd 
-      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+      apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm)
       apply (simp only: one_add_one_is_two[symmetric] of_int_add)
       using d order_less_not_sym[OF d] less_imp_neq[OF d]
-      by (simp add: ring_simps )    }
+      by (simp add: field_simps )    }
 ultimately show ?thesis by blast
 qed
 
@@ -2519,7 +2519,7 @@
 lemma remdps_set[simp]: "set (remdps xs) = set xs"
   by (induct xs rule: remdps.induct, auto)
 
-lemma simpfm_lin:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma simpfm_lin:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "qfree p \<Longrightarrow> islin (simpfm p)"
   by (induct p rule: simpfm.induct, auto simp add: conj_lin disj_lin)
 
@@ -2551,7 +2551,7 @@
   {fix c t d s assume ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
     from U_l ctU dsU have norm: "isnpoly c" "isnpoly d" by auto
     from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t]
-    have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: ring_simps)}
+    have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: field_simps)}
   hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (msubst ?q (x, y)) \<longleftrightarrow> ?I (msubst ?q (y, x))" by clarsimp
   {fix x assume xUp: "x \<in> set ?Up" 
     then  obtain c t d s where ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U" 
@@ -2616,7 +2616,7 @@
     let ?s = "Itm vs (x # bs) s"
     let ?t = "Itm vs (x # bs) t"
     have eq2: "\<And>(x::'a). x + x = (1 + 1) * x"
-      by  (simp add: ring_simps)
+      by  (simp add: field_simps)
     {assume "?c = 0 \<and> ?d = 0"
       with ct have ?D by simp}
     moreover
@@ -2747,12 +2747,12 @@
 using lp tnb
 by (induct p c t rule: msubstpos.induct, auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
 
-lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
   shows "bound0 (msubstneg p c t)"
 using lp tnb
 by (induct p c t rule: msubstneg.induct, auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
 
-lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
   shows "bound0 (msubst2 p c t)"
 using lp tnb
 by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
@@ -2899,14 +2899,14 @@
         by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
       from msubst2[OF lq norm2(1) z(1), of x bs] 
         msubst2[OF lq norm2(2) z(2), of x bs] H 
-      show ?rhs by (simp add: ring_simps)
+      show ?rhs by (simp add: field_simps)
     next
       assume H: ?rhs
       hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" 
         by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
       from msubst2[OF lq norm2(1) z(1), of x bs] 
         msubst2[OF lq norm2(2) z(2), of x bs] H 
-      show ?lhs by (simp add: ring_simps)
+      show ?lhs by (simp add: field_simps)
     qed}
   hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
     by clarsimp
@@ -3156,54 +3156,54 @@
 *} "Parametric QE for linear Arithmetic over fields, Version 2"
 
 
-lemma "\<exists>(x::'a::{division_by_zero,linordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
-  apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "y::'a::{division_by_zero,linordered_field,number_ring}")
-  apply (simp add: ring_simps)
+lemma "\<exists>(x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+  apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "y::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
+  apply (simp add: field_simps)
   apply (rule spec[where x=y])
-  apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "z::'a::{division_by_zero,linordered_field,number_ring}")
+  apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "z::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
   by simp
 
 text{* Collins/Jones Problem *}
 (*
-lemma "\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+lemma "\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
 proof-
-  have "(\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
-by (simp add: ring_simps)
+  have "(\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: field_simps)
 have "?rhs"
 
-  apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "a::'a::{division_by_zero,linordered_field,number_ring}" "b::'a::{division_by_zero,linordered_field,number_ring}")
-  apply (simp add: ring_simps)
+  apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "a::'a::{linordered_field, division_ring_inverse_zero, number_ring}" "b::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
+  apply (simp add: field_simps)
 oops
 *)
 (*
-lemma "ALL (x::'a::{division_by_zero,linordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
-apply (frpar type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "t::'a::{division_by_zero,linordered_field,number_ring}")
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "t::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
 oops
 *)
 
-lemma "\<exists>(x::'a::{division_by_zero,linordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
-  apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "y::'a::{division_by_zero,linordered_field,number_ring}")
-  apply (simp add: ring_simps)
+lemma "\<exists>(x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+  apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "y::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
+  apply (simp add: field_simps)
   apply (rule spec[where x=y])
-  apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "z::'a::{division_by_zero,linordered_field,number_ring}")
+  apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "z::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
   by simp
 
 text{* Collins/Jones Problem *}
 
 (*
-lemma "\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+lemma "\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
 proof-
-  have "(\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,linordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
-by (simp add: ring_simps)
+  have "(\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: field_simps)
 have "?rhs"
-  apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "a::'a::{division_by_zero,linordered_field,number_ring}" "b::'a::{division_by_zero,linordered_field,number_ring}")
+  apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "a::'a::{linordered_field, division_ring_inverse_zero, number_ring}" "b::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
   apply simp
 oops
 *)
 
 (*
-lemma "ALL (x::'a::{division_by_zero,linordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
-apply (frpar2 type: "'a::{division_by_zero,linordered_field,number_ring}" pars: "t::'a::{division_by_zero,linordered_field,number_ring}")
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar2 type: "'a::{linordered_field, division_ring_inverse_zero, number_ring}" pars: "t::'a::{linordered_field, division_ring_inverse_zero, number_ring}")
 apply (simp add: field_simps linorder_neq_iff[symmetric])
 apply ferrack
 oops
--- a/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -283,11 +283,11 @@
 apply (drule_tac x = "%m. if m = Suc n then a else i m" in spec, safe)
 apply (drule poly_mult_eq_zero_disj [THEN iffD1], safe)
 apply (drule_tac x = "Suc (length q)" in spec)
-apply (auto simp add: ring_simps)
+apply (auto simp add: field_simps)
 apply (drule_tac x = xa in spec)
-apply (clarsimp simp add: ring_simps)
+apply (clarsimp simp add: field_simps)
 apply (drule_tac x = m in spec)
-apply (auto simp add:ring_simps)
+apply (auto simp add:field_simps)
 done
 lemmas poly_roots_index_lemma1 = conjI [THEN poly_roots_index_lemma0, standard]
 
@@ -327,7 +327,7 @@
 apply (drule_tac x = "a#i" in spec)
 apply (auto simp only: poly_mult List.list.size)
 apply (drule_tac x = xa in spec)
-apply (clarsimp simp add: ring_simps)
+apply (clarsimp simp add: field_simps)
 done
 
 lemmas poly_roots_index_lemma2 = conjI [THEN poly_roots_index_lemma, standard]
@@ -413,7 +413,7 @@
 by (auto intro!: ext)
 
 lemma poly_add_minus_zero_iff: "(poly (p +++ -- q) = poly []) = (poly p = poly q)"
-by (auto simp add: ring_simps poly_add poly_minus_def fun_eq poly_cmult)
+by (auto simp add: field_simps poly_add poly_minus_def fun_eq poly_cmult)
 
 lemma poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
 by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult right_distrib)
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -230,7 +230,7 @@
 
 subsection{* Semantics of the polynomial representation *}
 
-consts Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{ring_char_0,power,division_by_zero,field}"
+consts Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{ring_char_0,power,division_ring_inverse_zero,field}"
 primrec
   "Ipoly bs (C c) = INum c"
   "Ipoly bs (Bound n) = bs!n"
@@ -241,7 +241,7 @@
   "Ipoly bs (Pw t n) = (Ipoly bs t) ^ n"
   "Ipoly bs (CN c n p) = (Ipoly bs c) + (bs!n)*(Ipoly bs p)"
 abbreviation
-  Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{ring_char_0,power,division_by_zero,field}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
+  Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{ring_char_0,power,division_ring_inverse_zero,field}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
   where "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<equiv> Ipoly bs p"
 
 lemma Ipoly_CInt: "Ipoly bs (C (i,1)) = of_int i" 
@@ -322,7 +322,7 @@
 qed auto
 
 lemma polyadd[simp]: "Ipoly bs (polyadd (p,q)) = (Ipoly bs p) + (Ipoly bs q)"
-by (induct p q rule: polyadd.induct, auto simp add: Let_def ring_simps right_distrib[symmetric] simp del: right_distrib)
+by (induct p q rule: polyadd.induct, auto simp add: Let_def field_simps right_distrib[symmetric] simp del: right_distrib)
 
 lemma polyadd_norm: "\<lbrakk> isnpoly p ; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polyadd(p,q))"
   using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp
@@ -394,7 +394,7 @@
 qed simp_all
 
 lemma polymul_properties:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> min n0 n1"
   shows "isnpolyh (p *\<^sub>p q) (min n0 n1)" 
   and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p)" 
@@ -565,22 +565,22 @@
 qed auto
 
 lemma polymul[simp]: "Ipoly bs (p *\<^sub>p q) = (Ipoly bs p) * (Ipoly bs q)"
-by(induct p q rule: polymul.induct, auto simp add: ring_simps)
+by(induct p q rule: polymul.induct, auto simp add: field_simps)
 
 lemma polymul_normh: 
-    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+    assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (p *\<^sub>p q) (min n0 n1)"
   using polymul_properties(1)  by blast
 lemma polymul_eq0_iff: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p) "
   using polymul_properties(2)  by blast
 lemma polymul_degreen:  
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 else degreen p m + degreen q m)"
   using polymul_properties(3) by blast
 lemma polymul_norm:   
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polymul (p,q))"
   using polymul_normh[of "p" "0" "q" "0"] isnpoly_def by simp
 
@@ -591,7 +591,7 @@
   by (induct p arbitrary: n0, auto)
 
 lemma monic_eqI: assumes np: "isnpolyh p n0" 
-  shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{ring_char_0,power,division_by_zero,field})"
+  shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{ring_char_0,power,division_ring_inverse_zero,field})"
   unfolding monic_def Let_def
 proof(cases "headconst p = 0\<^sub>N", simp_all add: headconst_zero[OF np])
   let ?h = "headconst p"
@@ -629,13 +629,13 @@
 
 lemma polysub_norm: "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polysub(p,q))"
   using polyadd_norm polyneg_norm by (simp add: polysub_def) 
-lemma polysub_same_0[simp]:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma polysub_same_0[simp]:   assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "isnpolyh p n0 \<Longrightarrow> polysub (p, p) = 0\<^sub>p"
 unfolding polysub_def split_def fst_conv snd_conv
 by (induct p arbitrary: n0,auto simp add: Let_def Nsub0[simplified Nsub_def])
 
 lemma polysub_0: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p -\<^sub>p q = 0\<^sub>p) = (p = q)"
   unfolding polysub_def split_def fst_conv snd_conv
   apply (induct p q arbitrary: n0 n1 rule:polyadd.induct, simp_all add: Nsub0[simplified Nsub_def])
@@ -657,7 +657,7 @@
   done
 
 text{* polypow is a power function and preserves normal forms *}
-lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{ring_char_0,division_by_zero,field})) ^ n"
+lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{field, division_ring_inverse_zero, ring_char_0})) ^ n"
 proof(induct n rule: polypow.induct)
   case 1 thus ?case by simp
 next
@@ -688,7 +688,7 @@
 qed
 
 lemma polypow_normh: 
-    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+    assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "isnpolyh p n \<Longrightarrow> isnpolyh (polypow k p) n"
 proof (induct k arbitrary: n rule: polypow.induct)
   case (2 k n)
@@ -701,17 +701,17 @@
 qed auto 
 
 lemma polypow_norm:   
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)"
   by (simp add: polypow_normh isnpoly_def)
 
 text{* Finally the whole normalization*}
 
-lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{ring_char_0,division_by_zero,field})"
+lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field, division_ring_inverse_zero, ring_char_0})"
 by (induct p rule:polynate.induct, auto)
 
 lemma polynate_norm[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "isnpoly (polynate p)"
   by (induct p rule: polynate.induct, simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm) (simp_all add: isnpoly_def)
 
@@ -736,29 +736,29 @@
   shows "isnpolyh (funpow k f p) n"
   using f np by (induct k arbitrary: p, auto)
 
-lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (Mul (Pw (Bound 0) n) p)"
+lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {field, division_ring_inverse_zero, ring_char_0}) = Ipoly bs (Mul (Pw (Bound 0) n) p)"
   by (induct n arbitrary: p, simp_all add: shift1_isnpoly shift1 power_Suc )
 
 lemma shift1_isnpolyh: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> isnpolyh (shift1 p) 0"
   using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by (simp add: shift1_def)
 
 lemma funpow_shift1_1: 
-  "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)"
+  "(Ipoly bs (funpow n shift1 p) :: 'a :: {field, division_ring_inverse_zero, ring_char_0}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)"
   by (simp add: funpow_shift1)
 
 lemma poly_cmul[simp]: "Ipoly bs (poly_cmul c p) = Ipoly bs (Mul (C c) p)"
-by (induct p  arbitrary: n0 rule: poly_cmul.induct, auto simp add: ring_simps)
+by (induct p  arbitrary: n0 rule: poly_cmul.induct, auto simp add: field_simps)
 
 lemma behead:
   assumes np: "isnpolyh p n"
-  shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {ring_char_0,division_by_zero,field})"
+  shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {field, division_ring_inverse_zero, ring_char_0})"
   using np
 proof (induct p arbitrary: n rule: behead.induct)
   case (1 c p n) hence pn: "isnpolyh p n" by simp
   from prems(2)[OF pn] 
   have th:"Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = Ipoly bs p" . 
   then show ?case using "1.hyps" apply (simp add: Let_def,cases "behead p = 0\<^sub>p")
-    by (simp_all add: th[symmetric] ring_simps power_Suc)
+    by (simp_all add: th[symmetric] field_simps power_Suc)
 qed (auto simp add: Let_def)
 
 lemma behead_isnpolyh:
@@ -981,7 +981,7 @@
   by (simp add: head_eq_headn0)
 
 lemma isnpolyh_zero_iff: 
-  assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{ring_char_0,power,division_by_zero,field})"
+  assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{ring_char_0,power,division_ring_inverse_zero,field})"
   shows "p = 0\<^sub>p"
 using nq eq
 proof (induct "maxindex p" arbitrary: p n0 rule: less_induct)
@@ -1033,7 +1033,7 @@
 
 lemma isnpolyh_unique:  
   assumes np:"isnpolyh p n0" and nq: "isnpolyh q n1"
-  shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{ring_char_0,power,division_by_zero,field})) \<longleftrightarrow>  p = q"
+  shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{ring_char_0,power,division_ring_inverse_zero,field})) \<longleftrightarrow>  p = q"
 proof(auto)
   assume H: "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a)= \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>"
   hence "\<forall>bs.\<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup>= (0::'a)" by simp
@@ -1046,50 +1046,50 @@
 
 text{* consequenses of unicity on the algorithms for polynomial normalization *}
 
-lemma polyadd_commute:   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+lemma polyadd_commute:   assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" and nq: "isnpolyh q n1" shows "p +\<^sub>p q = q +\<^sub>p p"
   using isnpolyh_unique[OF polyadd_normh[OF np nq] polyadd_normh[OF nq np]] by simp
 
 lemma zero_normh: "isnpolyh 0\<^sub>p n" by simp
 lemma one_normh: "isnpolyh 1\<^sub>p n" by simp
 lemma polyadd_0[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" shows "p +\<^sub>p 0\<^sub>p = p" and "0\<^sub>p +\<^sub>p p = p"
   using isnpolyh_unique[OF polyadd_normh[OF np zero_normh] np] 
     isnpolyh_unique[OF polyadd_normh[OF zero_normh np] np] by simp_all
 
 lemma polymul_1[simp]: 
-    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+    assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" shows "p *\<^sub>p 1\<^sub>p = p" and "1\<^sub>p *\<^sub>p p = p"
   using isnpolyh_unique[OF polymul_normh[OF np one_normh] np] 
     isnpolyh_unique[OF polymul_normh[OF one_normh np] np] by simp_all
 lemma polymul_0[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" shows "p *\<^sub>p 0\<^sub>p = 0\<^sub>p" and "0\<^sub>p *\<^sub>p p = 0\<^sub>p"
   using isnpolyh_unique[OF polymul_normh[OF np zero_normh] zero_normh] 
     isnpolyh_unique[OF polymul_normh[OF zero_normh np] zero_normh] by simp_all
 
 lemma polymul_commute: 
-    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+    assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np:"isnpolyh p n0" and nq: "isnpolyh q n1"
   shows "p *\<^sub>p q = q *\<^sub>p p"
-using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{ring_char_0,power,division_by_zero,field}"] by simp
+using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{ring_char_0,power,division_ring_inverse_zero,field}"] by simp
 
 declare polyneg_polyneg[simp]
   
 lemma isnpolyh_polynate_id[simp]: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np:"isnpolyh p n0" shows "polynate p = p"
-  using isnpolyh_unique[where ?'a= "'a::{ring_char_0,division_by_zero,field}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{ring_char_0,division_by_zero,field}"] by simp
+  using isnpolyh_unique[where ?'a= "'a::{field, division_ring_inverse_zero, ring_char_0}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{field, division_ring_inverse_zero, ring_char_0}"] by simp
 
 lemma polynate_idempotent[simp]: 
-    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+    assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "polynate (polynate p) = polynate p"
   using isnpolyh_polynate_id[OF polynate_norm[of p, unfolded isnpoly_def]] .
 
 lemma poly_nate_polypoly': "poly_nate bs p = polypoly' bs (polynate p)"
   unfolding poly_nate_def polypoly'_def ..
-lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{ring_char_0,division_by_zero,field}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)"
+lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{field, division_ring_inverse_zero, ring_char_0}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)"
   using polypoly'_poly[OF polynate_norm[unfolded isnpoly_def], symmetric, of bs p]
   unfolding poly_nate_polypoly' by (auto intro: ext)
 
@@ -1116,7 +1116,7 @@
 qed
 
 lemma degree_polysub_samehead: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and h: "head p = head q" 
   and d: "degree p = degree q"
   shows "degree (p -\<^sub>p q) < degree p \<or> (p -\<^sub>p q = 0\<^sub>p)"
@@ -1226,7 +1226,7 @@
 done
 
 lemma polymul_head_polyeq: 
-   assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+   assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "\<lbrakk>isnpolyh p n0; isnpolyh q n1 ; p \<noteq> 0\<^sub>p ; q \<noteq> 0\<^sub>p \<rbrakk> \<Longrightarrow> head (p *\<^sub>p q) = head p *\<^sub>p head q"
 proof (induct p q arbitrary: n0 n1 rule: polymul.induct)
   case (2 a b c' n' p' n0 n1)
@@ -1300,7 +1300,7 @@
   by (induct p arbitrary: n0 rule: polyneg.induct, auto)
 
 lemma degree_polymul:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" and nq: "isnpolyh q n1"
   shows "degree (p *\<^sub>p q) \<le> degree p + degree q"
   using polymul_degreen[OF np nq, where m="0"]  degree_eq_degreen0 by simp
@@ -1344,7 +1344,7 @@
 qed
 
 lemma polydivide_aux_properties:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" and ns: "isnpolyh s n1"
   and ap: "head p = a" and ndp: "degree p = n" and pnz: "p \<noteq> 0\<^sub>p"
   shows "polydivide_aux_dom (a,n,p,k,s) \<and> 
@@ -1415,19 +1415,19 @@
             from polyadd_normh[OF polymul_normh[OF np 
               polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr']
             have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp 
-            from asp have "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) = 
+            from asp have "\<forall> (bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) = 
               Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
-            hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) = 
+            hence " \<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) = 
               Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r" 
-              by (simp add: ring_simps)
-            hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+              by (simp add: field_simps)
+            hence " \<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
               Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 1\<^sub>p *\<^sub>p p) 
               + Ipoly bs p * Ipoly bs q + Ipoly bs r"
               by (auto simp only: funpow_shift1_1) 
-            hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+            hence "\<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
               Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 1\<^sub>p) 
-              + Ipoly bs q) + Ipoly bs r" by (simp add: ring_simps)
-            hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+              + Ipoly bs q) + Ipoly bs r" by (simp add: field_simps)
+            hence "\<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
               Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp
             with isnpolyh_unique[OF nakks' nqr']
             have "a ^\<^sub>p (k' - k) *\<^sub>p s = 
@@ -1444,9 +1444,9 @@
             apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
           have dom: ?dths apply (rule polydivide_aux_real_domintros) 
             using ba dn' domsp by simp_all
-          from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{ring_char_0,division_by_zero,field}"]
-          have " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs ?p'" by simp
-          hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
+          from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{field, division_ring_inverse_zero, ring_char_0}"]
+          have " \<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs s = Ipoly bs ?p'" by simp
+          hence "\<forall>(bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
             by (simp only: funpow_shift1_1) simp
           hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast
           {assume h1: "polydivide_aux (a,n,p,k,s) = (k',r)"
@@ -1501,17 +1501,17 @@
               and dr: "degree r = 0 \<or> degree r < degree p"
               and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto
             from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith
-            {fix bs:: "'a::{ring_char_0,division_by_zero,field} list"
+            {fix bs:: "'a::{field, division_ring_inverse_zero, ring_char_0} list"
               
             from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
             have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
             hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r"
-              by (simp add: ring_simps power_Suc)
+              by (simp add: field_simps power_Suc)
             hence "Ipoly bs a ^ (k' - k)  * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
               by (simp add:kk'' funpow_shift1_1[where n="degree s - n" and p="p"])
             hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
-              by (simp add: ring_simps)}
-            hence ieq:"\<forall>(bs :: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
+              by (simp add: field_simps)}
+            hence ieq:"\<forall>(bs :: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = 
               Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto 
             let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)"
             from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]]
@@ -1532,17 +1532,17 @@
             apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
           have dom: ?dths using sz ba dn' domsp 
             by - (rule polydivide_aux_real_domintros, simp_all)
-          {fix bs :: "'a::{ring_char_0,division_by_zero,field} list"
+          {fix bs :: "'a::{field, division_ring_inverse_zero, ring_char_0} list"
             from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
           have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp
           hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p" 
             by (simp add: funpow_shift1_1[where n="degree s - n" and p="p"])
           hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp
         }
-        hence hth: "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
+        hence hth: "\<forall> (bs:: 'a::{field, division_ring_inverse_zero, ring_char_0} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
           from hth
           have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)" 
-            using isnpolyh_unique[where ?'a = "'a::{ring_char_0,division_by_zero,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns] 
+            using isnpolyh_unique[where ?'a = "'a::{field, division_ring_inverse_zero, ring_char_0}", OF polymul_normh[OF head_isnpolyh[OF np] ns] 
                     polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]],
               simplified ap] by simp
           {assume h1: "polydivide_aux (a,n,p,k,s) = (k', r)"
@@ -1566,7 +1566,7 @@
 qed
 
 lemma polydivide_properties: 
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \<noteq> 0\<^sub>p"
   shows "(\<exists> k r. polydivide s p = (k,r) \<and> (\<exists>nr. isnpolyh r nr) \<and> (degree r = 0 \<or> degree r < degree p) 
   \<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow k (head p)) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))"
@@ -1698,11 +1698,11 @@
 definition "swapnorm n m t = polynate (swap n m t)"
 
 lemma swapnorm: assumes nbs: "n < length bs" and mbs: "m < length bs"
-  shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{ring_char_0,division_by_zero,field})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
+  shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{field, division_ring_inverse_zero, ring_char_0})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
   using swap[OF prems] swapnorm_def by simp
 
 lemma swapnorm_isnpoly[simp]: 
-    assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+    assumes "SORT_CONSTRAINT('a::{field, division_ring_inverse_zero, ring_char_0})"
   shows "isnpoly (swapnorm n m p)"
   unfolding swapnorm_def by simp
 
--- a/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Decision_Procs/ex/Dense_Linear_Order_Ex.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -7,147 +7,147 @@
 begin
 
 lemma
-  "\<exists>(y::'a::{linordered_field,number_ring, division_by_zero}) <2. x + 3* y < 0 \<and> x - y >0"
+  "\<exists>(y::'a::{linordered_field, division_ring_inverse_zero, number_ring}) <2. x + 3* y < 0 \<and> x - y >0"
   by ferrack
 
-lemma "~ (ALL x (y::'a::{linordered_field,number_ring, division_by_zero}). x < y --> 10*x < 11*y)"
+lemma "~ (ALL x (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}). x < y --> 10*x < 11*y)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. x < y --> (10*(x + 5*y + -1) < 60*y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x < y --> (10*(x + 5*y + -1) < 60*y)"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. x ~= y --> x < y"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x ~= y --> x < y"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (x ~= y & 10*x ~= 9*y & 10*x < y) --> x < y"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (x ~= y & 10*x ~= 9*y & 10*x < y) --> x < y"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (x ~= y & 5*x <= y) --> 500*x <= 100*y"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (x ~= y & 5*x <= y) --> 500*x <= 100*y"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX (y::'a::{linordered_field,number_ring, division_by_zero}). 4*x + 3*y <= 0 & 4*x + 3*y >= -1)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}). 4*x + 3*y <= 0 & 4*x + 3*y >= -1)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) < 0. (EX (y::'a::{linordered_field,number_ring, division_by_zero}) > 0. 7*x + y > 0 & x - y <= 9)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) < 0. (EX (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}) > 0. 7*x + y > 0 & x - y <= 9)"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). (0 < x & x < 1) --> (ALL y > 1. x + y ~= 1)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (0 < x & x < 1) --> (ALL y > 1. x + y ~= 1)"
   by ferrack
 
-lemma "EX x. (ALL (y::'a::{linordered_field,number_ring, division_by_zero}). y < 2 -->  2*(y - x) \<le> 0 )"
+lemma "EX x. (ALL (y::'a::{linordered_field, division_ring_inverse_zero, number_ring}). y < 2 -->  2*(y - x) \<le> 0 )"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). x < 10 | x > 20 | (EX y. y>= 0 & y <= 10 & x+y = 20)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). x < 10 | x > 20 | (EX y. y>= 0 & y <= 10 & x+y = 20)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. x + y < z --> y >= z --> x < 0"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. x + y < z --> y >= z --> x < 0"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z. x + 7*y < 5* z & 5*y >= 7*z & x < 0"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. x + 7*y < 5* z & 5*y >= 7*z & x < 0"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. abs (x + y) <= z --> (abs z = z)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. abs (x + y) <= z --> (abs z = z)"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z. x + 7*y - 5* z < 0 & 5*y + 7*z + 3*x < 0"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. x + 7*y - 5* z < 0 & 5*y + 7*z + 3*x < 0"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. (abs (5*x+3*y+z) <= 5*x+3*y+z & abs (5*x+3*y+z) >= - (5*x+3*y+z)) | (abs (5*x+3*y+z) >= 5*x+3*y+z & abs (5*x+3*y+z) <= - (5*x+3*y+z))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. (abs (5*x+3*y+z) <= 5*x+3*y+z & abs (5*x+3*y+z) >= - (5*x+3*y+z)) | (abs (5*x+3*y+z) >= 5*x+3*y+z & abs (5*x+3*y+z) <= - (5*x+3*y+z))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. x < y --> (EX z>0. x+z = y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x < y --> (EX z>0. x+z = y)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. x < y --> (EX z>0. x+z = y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. x < y --> (EX z>0. x+z = y)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX z>0. abs (x - y) <= z )"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX z>0. abs (x - y) <= z )"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (ALL z>=0. abs (3*x+7*y) <= 2*z + 1)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (ALL z>=0. abs (3*x+7*y) <= 2*z + 1)"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (ALL z<0. (z < x --> z <= y) & (z > y --> z >= x))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero})>0. (ALL y. (EX z. 13* abs z \<noteq> abs (12*y - x) & 5*x - 3*(abs y) <= 7*z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring})>0. (ALL y. (EX z. 13* abs z \<noteq> abs (12*y - x) & 5*x - 3*(abs y) <= 7*z))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). abs (4*x + 17) < 4 & (ALL y . abs (x*34 - 34*y - 9) \<noteq> 0 \<longrightarrow> (EX z. 5*x - 3*abs y <= 7*z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). abs (4*x + 17) < 4 & (ALL y . abs (x*34 - 34*y - 9) \<noteq> 0 \<longrightarrow> (EX z. 5*x - 3*abs y <= 7*z))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y > abs (23*x - 9). (ALL z > abs (3*y - 19* abs x). x+z > 2*y))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y > abs (23*x - 9). (ALL z > abs (3*y - 19* abs x). x+z > 2*y))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y< abs (3*x - 1). (ALL z >= (3*abs x - 1). abs (12*x - 13*y + 19*z) > abs (23*x) ))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y< abs (3*x - 1). (ALL z >= (3*abs x - 1). abs (12*x - 13*y + 19*z) > abs (23*x) ))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). abs x < 100 & (ALL y > x. (EX z<2*y - x. 5*x - 3*y <= 7*z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). abs x < 100 & (ALL y > x. (EX z<2*y - x. 5*x - 3*y <= 7*z))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. 7*x<3*y --> 5*y < 7*z --> z < 2*w --> 7*(2*w-x) > 2*y"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. 7*x<3*y --> 5*y < 7*z --> z < 2*w --> 7*(2*w-x) > 2*y"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. 5*x + 3*z - 17*w + abs (y - 8*x + z) <= 89"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. 5*x + 3*z - 17*w + abs (y - 8*x + z) <= 89"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. 5*x + 3*z - 17*w + 7* (y - 8*x + z) <= max y (7*z - x + w)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. 5*x + 3*z - 17*w + 7* (y - 8*x + z) <= max y (7*z - x + w)"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y z. (EX w >= (x+y+z). w <= abs x + abs y + abs z)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. (EX w >= (x+y+z). w <= abs x + abs y + abs z)"
   by ferrack
 
-lemma "~(ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y z w. 3* x + z*4 = 3*y & x + y < z & x> w & 3*x < w + y))"
+lemma "~(ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y z w. 3* x + z*4 = 3*y & x + y < z & x> w & 3*x < w + y))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX z w. abs (x-y) = (z-w) & z*1234 < 233*x & w ~= y)"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX z w. abs (x-y) = (z-w) & z*1234 < 233*x & w ~= y)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z. (ALL w >= abs (x+y+z). w >= abs x + abs y + abs z)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z. (ALL w >= abs (x+y+z). w >= abs x + abs y + abs z)"
   by ferrack
 
-lemma "EX z. (ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX w >= (x+y+z). w <= abs x + abs y + abs z))"
+lemma "EX z. (ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX w >= (x+y+z). w <= abs x + abs y + abs z))"
   by ferrack
 
-lemma "EX z. (ALL (x::'a::{linordered_field,number_ring, division_by_zero}) < abs z. (EX y w. x< y & x < z & x> w & 3*x < w + y))"
+lemma "EX z. (ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) < abs z. (EX y w. x< y & x < z & x> w & 3*x < w + y))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}) y. (EX z. (ALL w. abs (x-y) = abs (z-w) --> z < x & w ~= y))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y. (EX z. (ALL w. abs (x-y) = abs (z-w) --> z < x & w ~= y))"
   by ferrack
 
-lemma "EX y. (ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)))"
+lemma "EX y. (ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) z. (ALL w >= 13*x - 4*z. (EX y. w >= abs x + abs y + z))"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) z. (ALL w >= 13*x - 4*z. (EX y. w >= abs x + abs y + z))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). (ALL y < x. (EX z > (x+y).
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (ALL y < x. (EX z > (x+y).
   (ALL w. 5*w + 10*x - z >= y --> w + 7*x + 3*z >= 2*y)))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}). (ALL y. (EX z > y.
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (ALL y. (EX z > y.
   (ALL w . w < 13 --> w + 10*x - z >= y --> 5*w + 7*x + 13*z >= 2*y)))"
   by ferrack
 
-lemma "EX (x::'a::{linordered_field,number_ring, division_by_zero}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
+lemma "EX (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}) y z w. min (5*x + 3*z) (17*w) + 5* abs (y - 8*x + z) <= max y (7*z - x + w)"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (y - x) < w)))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (y - x) < w)))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (x + z) < w - y)))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y. (ALL z>19. y <= x + z & (EX w. abs (x + z) < w - y)))"
   by ferrack
 
-lemma "ALL (x::'a::{linordered_field,number_ring, division_by_zero}). (EX y. abs y ~= abs x & (ALL z> max x y. (EX w. w ~= y & w ~= z & 3*w - z >= x + y)))"
+lemma "ALL (x::'a::{linordered_field, division_ring_inverse_zero, number_ring}). (EX y. abs y ~= abs x & (ALL z> max x y. (EX w. w ~= y & w ~= z & 3*w - z >= x + y)))"
   by ferrack
 
 end
--- a/src/HOL/Decision_Procs/mir_tac.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Decision_Procs/mir_tac.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -33,7 +33,7 @@
              @{thm "real_of_nat_number_of"},
              @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"},
              @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"},
-             @{thm "Fields.divide_zero"}, 
+             @{thm "divide_zero"}, 
              @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, 
              @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
              @{thm "diff_def"}, @{thm "minus_divide_left"}]
--- a/src/HOL/Fields.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Fields.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -31,34 +31,6 @@
 
 subclass idom ..
 
-lemma right_inverse_eq: "b \<noteq> 0 \<Longrightarrow> a / b = 1 \<longleftrightarrow> a = b"
-proof
-  assume neq: "b \<noteq> 0"
-  {
-    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
-    also assume "a / b = 1"
-    finally show "a = b" by simp
-  next
-    assume "a = b"
-    with neq show "a / b = 1" by (simp add: divide_inverse)
-  }
-qed
-
-lemma nonzero_inverse_eq_divide: "a \<noteq> 0 \<Longrightarrow> inverse a = 1 / a"
-by (simp add: divide_inverse)
-
-lemma divide_self [simp]: "a \<noteq> 0 \<Longrightarrow> a / a = 1"
-by (simp add: divide_inverse)
-
-lemma divide_zero_left [simp]: "0 / a = 0"
-by (simp add: divide_inverse)
-
-lemma inverse_eq_divide: "inverse a = 1 / a"
-by (simp add: divide_inverse)
-
-lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
-by (simp add: divide_inverse algebra_simps)
-
 text{*There is no slick version using division by zero.*}
 lemma inverse_add:
   "[| a \<noteq> 0;  b \<noteq> 0 |]
@@ -80,14 +52,8 @@
   "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
 by (simp add: mult_commute [of _ c])
 
-lemma divide_1 [simp]: "a / 1 = a"
-by (simp add: divide_inverse)
-
-lemma times_divide_eq_right: "a * (b / c) = (a * b) / c"
-by (simp add: divide_inverse mult_assoc)
-
-lemma times_divide_eq_left: "(b / c) * a = (b * a) / c"
-by (simp add: divide_inverse mult_ac)
+lemma times_divide_eq_left [simp]: "(b / c) * a = (b * a) / c"
+  by (simp add: divide_inverse mult_ac)
 
 text {* These are later declared as simp rules. *}
 lemmas times_divide_eq [no_atp] = times_divide_eq_right times_divide_eq_left
@@ -108,7 +74,7 @@
 
 lemma nonzero_mult_divide_cancel_right [simp, no_atp]:
   "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
-using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
+  using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
 
 lemma nonzero_mult_divide_cancel_left [simp, no_atp]:
   "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
@@ -130,242 +96,184 @@
   "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
 using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
 
-lemma minus_divide_left: "- (a / b) = (-a) / b"
-by (simp add: divide_inverse)
-
-lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
-by (simp add: divide_inverse nonzero_inverse_minus_eq)
-
-lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
-by (simp add: divide_inverse nonzero_inverse_minus_eq)
-
-lemma divide_minus_left [simp, no_atp]: "(-a) / b = - (a / b)"
-by (simp add: divide_inverse)
-
-lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
-by (simp add: diff_minus add_divide_distrib)
+lemma add_divide_eq_iff [field_simps]:
+  "z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
+  by (simp add: add_divide_distrib)
 
-lemma add_divide_eq_iff:
-  "z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
-by (simp add: add_divide_distrib)
-
-lemma divide_add_eq_iff:
+lemma divide_add_eq_iff [field_simps]:
   "z \<noteq> 0 \<Longrightarrow> x / z + y = (x + z * y) / z"
-by (simp add: add_divide_distrib)
-
-lemma diff_divide_eq_iff:
-  "z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
-by (simp add: diff_divide_distrib)
-
-lemma divide_diff_eq_iff:
-  "z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
-by (simp add: diff_divide_distrib)
+  by (simp add: add_divide_distrib)
 
-lemma nonzero_eq_divide_eq: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
-proof -
-  assume [simp]: "c \<noteq> 0"
-  have "a = b / c \<longleftrightarrow> a * c = (b / c) * c" by simp
-  also have "... \<longleftrightarrow> a * c = b" by (simp add: divide_inverse mult_assoc)
-  finally show ?thesis .
-qed
-
-lemma nonzero_divide_eq_eq: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
-proof -
-  assume [simp]: "c \<noteq> 0"
-  have "b / c = a \<longleftrightarrow> (b / c) * c = a * c" by simp
-  also have "... \<longleftrightarrow> b = a * c" by (simp add: divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
-lemma divide_eq_imp: "c \<noteq> 0 \<Longrightarrow> b = a * c \<Longrightarrow> b / c = a"
-by simp
+lemma diff_divide_eq_iff [field_simps]:
+  "z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
+  by (simp add: diff_divide_distrib)
 
-lemma eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
-by (erule subst, simp)
-
-lemmas field_eq_simps[no_atp] = algebra_simps
-  (* pull / out*)
-  add_divide_eq_iff divide_add_eq_iff
-  diff_divide_eq_iff divide_diff_eq_iff
-  (* multiply eqn *)
-  nonzero_eq_divide_eq nonzero_divide_eq_eq
-(* is added later:
-  times_divide_eq_left times_divide_eq_right
-*)
-
-text{*An example:*}
-lemma "\<lbrakk>a\<noteq>b; c\<noteq>d; e\<noteq>f\<rbrakk> \<Longrightarrow> ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1"
-apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
- apply(simp add:field_eq_simps)
-apply(simp)
-done
+lemma divide_diff_eq_iff [field_simps]:
+  "z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
+  by (simp add: diff_divide_distrib)
 
 lemma diff_frac_eq:
   "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> x / y - w / z = (x * z - w * y) / (y * z)"
-by (simp add: field_eq_simps times_divide_eq)
+  by (simp add: field_simps)
 
 lemma frac_eq_eq:
   "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (x / y = w / z) = (x * z = w * y)"
-by (simp add: field_eq_simps times_divide_eq)
+  by (simp add: field_simps)
 
 end
 
-class division_by_zero = zero + inverse +
-  assumes inverse_zero [simp]: "inverse 0 = 0"
-
-lemma divide_zero [simp]:
-  "a / 0 = (0::'a::{field,division_by_zero})"
-by (simp add: divide_inverse)
-
-lemma divide_self_if [simp]:
-  "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
-by simp
-
-class linordered_field = field + linordered_idom
-
-lemma inverse_nonzero_iff_nonzero [simp]:
-   "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))"
-by (force dest: inverse_zero_imp_zero) 
+class field_inverse_zero = field +
+  assumes field_inverse_zero: "inverse 0 = 0"
+begin
 
-lemma inverse_minus_eq [simp]:
-   "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})"
-proof cases
-  assume "a=0" thus ?thesis by simp
-next
-  assume "a\<noteq>0" 
-  thus ?thesis by (simp add: nonzero_inverse_minus_eq)
-qed
+subclass division_ring_inverse_zero proof
+qed (fact field_inverse_zero)
 
-lemma inverse_eq_imp_eq:
-  "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})"
-apply (cases "a=0 | b=0") 
- apply (force dest!: inverse_zero_imp_zero
-              simp add: eq_commute [of "0::'a"])
-apply (force dest!: nonzero_inverse_eq_imp_eq) 
-done
-
-lemma inverse_eq_iff_eq [simp]:
-  "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))"
-by (force dest!: inverse_eq_imp_eq)
-
-lemma inverse_inverse_eq [simp]:
-     "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a"
-  proof cases
-    assume "a=0" thus ?thesis by simp
-  next
-    assume "a\<noteq>0" 
-    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
-  qed
+end
 
 text{*This version builds in division by zero while also re-orienting
       the right-hand side.*}
 lemma inverse_mult_distrib [simp]:
-     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
+     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_ring_inverse_zero})"
   proof cases
     assume "a \<noteq> 0 & b \<noteq> 0" 
-    thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute)
+    thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_ac)
   next
     assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
     thus ?thesis by force
   qed
 
 lemma inverse_divide [simp]:
-  "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
-by (simp add: divide_inverse mult_commute)
+  "inverse (a/b) = b / (a::'a::{field,division_ring_inverse_zero})"
+  by (simp add: divide_inverse mult_commute)
 
 
-subsection {* Calculations with fractions *}
+text {* Calculations with fractions *}
 
 text{* There is a whole bunch of simp-rules just for class @{text
 field} but none for class @{text field} and @{text nonzero_divides}
 because the latter are covered by a simproc. *}
 
 lemma mult_divide_mult_cancel_left:
-  "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
+  "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_ring_inverse_zero})"
 apply (cases "b = 0")
 apply simp_all
 done
 
 lemma mult_divide_mult_cancel_right:
-  "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
+  "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_ring_inverse_zero})"
 apply (cases "b = 0")
 apply simp_all
 done
 
 lemma divide_divide_eq_right [simp,no_atp]:
-  "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
+  "a / (b/c) = (a*c) / (b::'a::{field,division_ring_inverse_zero})"
 by (simp add: divide_inverse mult_ac)
 
 lemma divide_divide_eq_left [simp,no_atp]:
-  "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
+  "(a / b) / (c::'a::{field,division_ring_inverse_zero}) = a / (b*c)"
 by (simp add: divide_inverse mult_assoc)
 
 
-subsubsection{*Special Cancellation Simprules for Division*}
+text {*Special Cancellation Simprules for Division*}
 
 lemma mult_divide_mult_cancel_left_if[simp,no_atp]:
-fixes c :: "'a :: {field,division_by_zero}"
+fixes c :: "'a :: {field,division_ring_inverse_zero}"
 shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)"
 by (simp add: mult_divide_mult_cancel_left)
 
 
-subsection {* Division and Unary Minus *}
+text {* Division and Unary Minus *}
 
-lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
+lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_ring_inverse_zero})"
 by (simp add: divide_inverse)
 
 lemma divide_minus_right [simp, no_atp]:
-  "a / -(b::'a::{field,division_by_zero}) = -(a / b)"
+  "a / -(b::'a::{field,division_ring_inverse_zero}) = -(a / b)"
 by (simp add: divide_inverse)
 
 lemma minus_divide_divide:
-  "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
+  "(-a)/(-b) = a / (b::'a::{field,division_ring_inverse_zero})"
 apply (cases "b=0", simp) 
 apply (simp add: nonzero_minus_divide_divide) 
 done
 
 lemma eq_divide_eq:
-  "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
+  "((a::'a::{field,division_ring_inverse_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
 by (simp add: nonzero_eq_divide_eq)
 
 lemma divide_eq_eq:
-  "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
+  "(b/c = (a::'a::{field,division_ring_inverse_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
 by (force simp add: nonzero_divide_eq_eq)
 
+lemma inverse_eq_1_iff [simp]:
+  "(inverse x = 1) = (x = (1::'a::{field,division_ring_inverse_zero}))"
+by (insert inverse_eq_iff_eq [of x 1], simp) 
 
-subsection {* Ordered Fields *}
+lemma divide_eq_0_iff [simp,no_atp]:
+     "(a/b = 0) = (a=0 | b=(0::'a::{field,division_ring_inverse_zero}))"
+by (simp add: divide_inverse)
+
+lemma divide_cancel_right [simp,no_atp]:
+     "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_ring_inverse_zero}))"
+apply (cases "c=0", simp)
+apply (simp add: divide_inverse)
+done
+
+lemma divide_cancel_left [simp,no_atp]:
+     "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_ring_inverse_zero}))" 
+apply (cases "c=0", simp)
+apply (simp add: divide_inverse)
+done
+
+lemma divide_eq_1_iff [simp,no_atp]:
+     "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_ring_inverse_zero}))"
+apply (cases "b=0", simp)
+apply (simp add: right_inverse_eq)
+done
+
+lemma one_eq_divide_iff [simp,no_atp]:
+     "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_ring_inverse_zero}))"
+by (simp add: eq_commute [of 1])
+
+
+text {* Ordered Fields *}
+
+class linordered_field = field + linordered_idom
+begin
 
 lemma positive_imp_inverse_positive: 
-assumes a_gt_0: "0 < a"  shows "0 < inverse (a::'a::linordered_field)"
+  assumes a_gt_0: "0 < a" 
+  shows "0 < inverse a"
 proof -
   have "0 < a * inverse a" 
-    by (simp add: a_gt_0 [THEN order_less_imp_not_eq2])
+    by (simp add: a_gt_0 [THEN less_imp_not_eq2])
   thus "0 < inverse a" 
-    by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff)
+    by (simp add: a_gt_0 [THEN less_not_sym] zero_less_mult_iff)
 qed
 
 lemma negative_imp_inverse_negative:
-  "a < 0 ==> inverse a < (0::'a::linordered_field)"
-by (insert positive_imp_inverse_positive [of "-a"], 
-    simp add: nonzero_inverse_minus_eq order_less_imp_not_eq)
+  "a < 0 \<Longrightarrow> inverse a < 0"
+  by (insert positive_imp_inverse_positive [of "-a"], 
+    simp add: nonzero_inverse_minus_eq less_imp_not_eq)
 
 lemma inverse_le_imp_le:
-assumes invle: "inverse a \<le> inverse b" and apos:  "0 < a"
-shows "b \<le> (a::'a::linordered_field)"
+  assumes invle: "inverse a \<le> inverse b" and apos: "0 < a"
+  shows "b \<le> a"
 proof (rule classical)
   assume "~ b \<le> a"
   hence "a < b"  by (simp add: linorder_not_le)
-  hence bpos: "0 < b"  by (blast intro: apos order_less_trans)
+  hence bpos: "0 < b"  by (blast intro: apos less_trans)
   hence "a * inverse a \<le> a * inverse b"
-    by (simp add: apos invle order_less_imp_le mult_left_mono)
+    by (simp add: apos invle less_imp_le mult_left_mono)
   hence "(a * inverse a) * b \<le> (a * inverse b) * b"
-    by (simp add: bpos order_less_imp_le mult_right_mono)
-  thus "b \<le> a"  by (simp add: mult_assoc apos bpos order_less_imp_not_eq2)
+    by (simp add: bpos less_imp_le mult_right_mono)
+  thus "b \<le> a"  by (simp add: mult_assoc apos bpos less_imp_not_eq2)
 qed
 
 lemma inverse_positive_imp_positive:
-assumes inv_gt_0: "0 < inverse a" and nz: "a \<noteq> 0"
-shows "0 < (a::'a::linordered_field)"
+  assumes inv_gt_0: "0 < inverse a" and nz: "a \<noteq> 0"
+  shows "0 < a"
 proof -
   have "0 < inverse (inverse a)"
     using inv_gt_0 by (rule positive_imp_inverse_positive)
@@ -373,36 +281,17 @@
     using nz by (simp add: nonzero_inverse_inverse_eq)
 qed
 
-lemma inverse_positive_iff_positive [simp]:
-  "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))"
-apply (cases "a = 0", simp)
-apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
-done
-
 lemma inverse_negative_imp_negative:
-assumes inv_less_0: "inverse a < 0" and nz:  "a \<noteq> 0"
-shows "a < (0::'a::linordered_field)"
+  assumes inv_less_0: "inverse a < 0" and nz: "a \<noteq> 0"
+  shows "a < 0"
 proof -
   have "inverse (inverse a) < 0"
     using inv_less_0 by (rule negative_imp_inverse_negative)
   thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq)
 qed
 
-lemma inverse_negative_iff_negative [simp]:
-  "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))"
-apply (cases "a = 0", simp)
-apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
-done
-
-lemma inverse_nonnegative_iff_nonnegative [simp]:
-  "(0 \<le> inverse a) = (0 \<le> (a::'a::{linordered_field,division_by_zero}))"
-by (simp add: linorder_not_less [symmetric])
-
-lemma inverse_nonpositive_iff_nonpositive [simp]:
-  "(inverse a \<le> 0) = (a \<le> (0::'a::{linordered_field,division_by_zero}))"
-by (simp add: linorder_not_less [symmetric])
-
-lemma linordered_field_no_lb: "\<forall> x. \<exists>y. y < (x::'a::linordered_field)"
+lemma linordered_field_no_lb:
+  "\<forall>x. \<exists>y. y < x"
 proof
   fix x::'a
   have m1: "- (1::'a) < 0" by simp
@@ -411,7 +300,8 @@
   thus "\<exists>y. y < x" by blast
 qed
 
-lemma linordered_field_no_ub: "\<forall> x. \<exists>y. y > (x::'a::linordered_field)"
+lemma linordered_field_no_ub:
+  "\<forall> x. \<exists>y. y > x"
 proof
   fix x::'a
   have m1: " (1::'a) > 0" by simp
@@ -420,88 +310,379 @@
   thus "\<exists>y. y > x" by blast
 qed
 
-subsection{*Anti-Monotonicity of @{term inverse}*}
-
 lemma less_imp_inverse_less:
-assumes less: "a < b" and apos:  "0 < a"
-shows "inverse b < inverse (a::'a::linordered_field)"
+  assumes less: "a < b" and apos:  "0 < a"
+  shows "inverse b < inverse a"
 proof (rule ccontr)
   assume "~ inverse b < inverse a"
-  hence "inverse a \<le> inverse b" by (simp add: linorder_not_less)
+  hence "inverse a \<le> inverse b" by simp
   hence "~ (a < b)"
-    by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
+    by (simp add: not_less inverse_le_imp_le [OF _ apos])
   thus False by (rule notE [OF _ less])
 qed
 
 lemma inverse_less_imp_less:
-  "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::linordered_field)"
-apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"])
+  "inverse a < inverse b \<Longrightarrow> 0 < a \<Longrightarrow> b < a"
+apply (simp add: less_le [of "inverse a"] less_le [of "b"])
 apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) 
 done
 
 text{*Both premises are essential. Consider -1 and 1.*}
 lemma inverse_less_iff_less [simp,no_atp]:
-  "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
-by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) 
+  "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
+  by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) 
 
 lemma le_imp_inverse_le:
-  "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
-by (force simp add: order_le_less less_imp_inverse_less)
+  "a \<le> b \<Longrightarrow> 0 < a \<Longrightarrow> inverse b \<le> inverse a"
+  by (force simp add: le_less less_imp_inverse_less)
 
 lemma inverse_le_iff_le [simp,no_atp]:
- "[|0 < a; 0 < b|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
-by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) 
+  "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
+  by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) 
 
 
 text{*These results refer to both operands being negative.  The opposite-sign
 case is trivial, since inverse preserves signs.*}
 lemma inverse_le_imp_le_neg:
-  "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::linordered_field)"
+  "inverse a \<le> inverse b \<Longrightarrow> b < 0 \<Longrightarrow> b \<le> a"
 apply (rule classical) 
 apply (subgoal_tac "a < 0") 
- prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) 
+ prefer 2 apply force
 apply (insert inverse_le_imp_le [of "-b" "-a"])
-apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
+apply (simp add: nonzero_inverse_minus_eq) 
 done
 
 lemma less_imp_inverse_less_neg:
-   "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::linordered_field)"
+   "a < b \<Longrightarrow> b < 0 \<Longrightarrow> inverse b < inverse a"
 apply (subgoal_tac "a < 0") 
- prefer 2 apply (blast intro: order_less_trans) 
+ prefer 2 apply (blast intro: less_trans) 
 apply (insert less_imp_inverse_less [of "-b" "-a"])
-apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
+apply (simp add: nonzero_inverse_minus_eq) 
 done
 
 lemma inverse_less_imp_less_neg:
-   "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::linordered_field)"
+   "inverse a < inverse b \<Longrightarrow> b < 0 \<Longrightarrow> b < a"
 apply (rule classical) 
 apply (subgoal_tac "a < 0") 
  prefer 2
- apply (force simp add: linorder_not_less intro: order_le_less_trans) 
+ apply force
 apply (insert inverse_less_imp_less [of "-b" "-a"])
-apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
+apply (simp add: nonzero_inverse_minus_eq) 
 done
 
 lemma inverse_less_iff_less_neg [simp,no_atp]:
-  "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
+  "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a < inverse b \<longleftrightarrow> b < a"
 apply (insert inverse_less_iff_less [of "-b" "-a"])
 apply (simp del: inverse_less_iff_less 
-            add: order_less_imp_not_eq nonzero_inverse_minus_eq)
+            add: nonzero_inverse_minus_eq)
 done
 
 lemma le_imp_inverse_le_neg:
-  "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
-by (force simp add: order_le_less less_imp_inverse_less_neg)
+  "a \<le> b \<Longrightarrow> b < 0 ==> inverse b \<le> inverse a"
+  by (force simp add: le_less less_imp_inverse_less_neg)
 
 lemma inverse_le_iff_le_neg [simp,no_atp]:
- "[|a < 0; b < 0|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
-by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
+  "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> inverse a \<le> inverse b \<longleftrightarrow> b \<le> a"
+  by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
+
+lemma pos_le_divide_eq [field_simps]: "0 < c ==> (a \<le> b/c) = (a*c \<le> b)"
+proof -
+  assume less: "0<c"
+  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
+    by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (a*c \<le> b)"
+    by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma neg_le_divide_eq [field_simps]: "c < 0 ==> (a \<le> b/c) = (b \<le> a*c)"
+proof -
+  assume less: "c<0"
+  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
+    by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (b \<le> a*c)"
+    by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma pos_less_divide_eq [field_simps]:
+     "0 < c ==> (a < b/c) = (a*c < b)"
+proof -
+  assume less: "0<c"
+  hence "(a < b/c) = (a*c < (b/c)*c)"
+    by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (a*c < b)"
+    by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma neg_less_divide_eq [field_simps]:
+ "c < 0 ==> (a < b/c) = (b < a*c)"
+proof -
+  assume less: "c<0"
+  hence "(a < b/c) = ((b/c)*c < a*c)"
+    by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (b < a*c)"
+    by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma pos_divide_less_eq [field_simps]:
+     "0 < c ==> (b/c < a) = (b < a*c)"
+proof -
+  assume less: "0<c"
+  hence "(b/c < a) = ((b/c)*c < a*c)"
+    by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (b < a*c)"
+    by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma neg_divide_less_eq [field_simps]:
+ "c < 0 ==> (b/c < a) = (a*c < b)"
+proof -
+  assume less: "c<0"
+  hence "(b/c < a) = (a*c < (b/c)*c)"
+    by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (a*c < b)"
+    by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma pos_divide_le_eq [field_simps]: "0 < c ==> (b/c \<le> a) = (b \<le> a*c)"
+proof -
+  assume less: "0<c"
+  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
+    by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (b \<le> a*c)"
+    by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma neg_divide_le_eq [field_simps]: "c < 0 ==> (b/c \<le> a) = (a*c \<le> b)"
+proof -
+  assume less: "c<0"
+  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
+    by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq)
+  also have "... = (a*c \<le> b)"
+    by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
+of positivity/negativity needed for @{text field_simps}. Have not added @{text
+sign_simps} to @{text field_simps} because the former can lead to case
+explosions. *}
+
+lemmas sign_simps [no_atp] = algebra_simps
+  zero_less_mult_iff mult_less_0_iff
+
+lemmas (in -) sign_simps [no_atp] = algebra_simps
+  zero_less_mult_iff mult_less_0_iff
+
+(* Only works once linear arithmetic is installed:
+text{*An example:*}
+lemma fixes a b c d e f :: "'a::linordered_field"
+shows "\<lbrakk>a>b; c<d; e<f; 0 < u \<rbrakk> \<Longrightarrow>
+ ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) <
+ ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u"
+apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0")
+ prefer 2 apply(simp add:sign_simps)
+apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0")
+ prefer 2 apply(simp add:sign_simps)
+apply(simp add:field_simps)
+done
+*)
+
+lemma divide_pos_pos:
+  "0 < x ==> 0 < y ==> 0 < x / y"
+by(simp add:field_simps)
+
+lemma divide_nonneg_pos:
+  "0 <= x ==> 0 < y ==> 0 <= x / y"
+by(simp add:field_simps)
+
+lemma divide_neg_pos:
+  "x < 0 ==> 0 < y ==> x / y < 0"
+by(simp add:field_simps)
+
+lemma divide_nonpos_pos:
+  "x <= 0 ==> 0 < y ==> x / y <= 0"
+by(simp add:field_simps)
+
+lemma divide_pos_neg:
+  "0 < x ==> y < 0 ==> x / y < 0"
+by(simp add:field_simps)
+
+lemma divide_nonneg_neg:
+  "0 <= x ==> y < 0 ==> x / y <= 0" 
+by(simp add:field_simps)
+
+lemma divide_neg_neg:
+  "x < 0 ==> y < 0 ==> 0 < x / y"
+by(simp add:field_simps)
+
+lemma divide_nonpos_neg:
+  "x <= 0 ==> y < 0 ==> 0 <= x / y"
+by(simp add:field_simps)
+
+lemma divide_strict_right_mono:
+     "[|a < b; 0 < c|] ==> a / c < b / c"
+by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono 
+              positive_imp_inverse_positive)
 
 
-subsection{*Inverses and the Number One*}
+lemma divide_strict_right_mono_neg:
+     "[|b < a; c < 0|] ==> a / c < b / c"
+apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
+apply (simp add: less_imp_not_eq nonzero_minus_divide_right [symmetric])
+done
+
+text{*The last premise ensures that @{term a} and @{term b} 
+      have the same sign*}
+lemma divide_strict_left_mono:
+  "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / b"
+by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono)
+
+lemma divide_left_mono:
+  "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / b"
+by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono)
+
+lemma divide_strict_left_mono_neg:
+  "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / b"
+by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg)
+
+lemma mult_imp_div_pos_le: "0 < y ==> x <= z * y ==>
+    x / y <= z"
+by (subst pos_divide_le_eq, assumption+)
+
+lemma mult_imp_le_div_pos: "0 < y ==> z * y <= x ==>
+    z <= x / y"
+by(simp add:field_simps)
+
+lemma mult_imp_div_pos_less: "0 < y ==> x < z * y ==>
+    x / y < z"
+by(simp add:field_simps)
+
+lemma mult_imp_less_div_pos: "0 < y ==> z * y < x ==>
+    z < x / y"
+by(simp add:field_simps)
+
+lemma frac_le: "0 <= x ==> 
+    x <= y ==> 0 < w ==> w <= z  ==> x / z <= y / w"
+  apply (rule mult_imp_div_pos_le)
+  apply simp
+  apply (subst times_divide_eq_left)
+  apply (rule mult_imp_le_div_pos, assumption)
+  apply (rule mult_mono)
+  apply simp_all
+done
+
+lemma frac_less: "0 <= x ==> 
+    x < y ==> 0 < w ==> w <= z  ==> x / z < y / w"
+  apply (rule mult_imp_div_pos_less)
+  apply simp
+  apply (subst times_divide_eq_left)
+  apply (rule mult_imp_less_div_pos, assumption)
+  apply (erule mult_less_le_imp_less)
+  apply simp_all
+done
+
+lemma frac_less2: "0 < x ==> 
+    x <= y ==> 0 < w ==> w < z  ==> x / z < y / w"
+  apply (rule mult_imp_div_pos_less)
+  apply simp_all
+  apply (rule mult_imp_less_div_pos, assumption)
+  apply (erule mult_le_less_imp_less)
+  apply simp_all
+done
+
+text{*It's not obvious whether these should be simprules or not. 
+  Their effect is to gather terms into one big fraction, like
+  a*b*c / x*y*z. The rationale for that is unclear, but many proofs 
+  seem to need them.*}
+
+lemma less_half_sum: "a < b ==> a < (a+b) / (1+1)"
+by (simp add: field_simps zero_less_two)
+
+lemma gt_half_sum: "a < b ==> (a+b)/(1+1) < b"
+by (simp add: field_simps zero_less_two)
+
+subclass dense_linorder
+proof
+  fix x y :: 'a
+  from less_add_one show "\<exists>y. x < y" .. 
+  from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono)
+  then have "x - 1 < x + 1 - 1" by (simp only: diff_minus [symmetric])
+  then have "x - 1 < x" by (simp add: algebra_simps)
+  then show "\<exists>y. y < x" ..
+  show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
+qed
+
+lemma nonzero_abs_inverse:
+     "a \<noteq> 0 ==> \<bar>inverse a\<bar> = inverse \<bar>a\<bar>"
+apply (auto simp add: neq_iff abs_if nonzero_inverse_minus_eq 
+                      negative_imp_inverse_negative)
+apply (blast intro: positive_imp_inverse_positive elim: less_asym) 
+done
+
+lemma nonzero_abs_divide:
+     "b \<noteq> 0 ==> \<bar>a / b\<bar> = \<bar>a\<bar> / \<bar>b\<bar>"
+  by (simp add: divide_inverse abs_mult nonzero_abs_inverse) 
+
+lemma field_le_epsilon:
+  assumes e: "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
+  shows "x \<le> y"
+proof (rule dense_le)
+  fix t assume "t < x"
+  hence "0 < x - t" by (simp add: less_diff_eq)
+  from e [OF this] have "x + 0 \<le> x + (y - t)" by (simp add: algebra_simps)
+  then have "0 \<le> y - t" by (simp only: add_le_cancel_left)
+  then show "t \<le> y" by (simp add: algebra_simps)
+qed
+
+end
+
+class linordered_field_inverse_zero = linordered_field +
+  assumes linordered_field_inverse_zero: "inverse 0 = 0"
+begin
+
+subclass field_inverse_zero proof
+qed (fact linordered_field_inverse_zero)
+
+end
+
+lemma le_divide_eq:
+  "(a \<le> b/c) = 
+   (if 0 < c then a*c \<le> b
+             else if c < 0 then b \<le> a*c
+             else  a \<le> (0::'a::{linordered_field,division_ring_inverse_zero}))"
+apply (cases "c=0", simp) 
+apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) 
+done
+
+lemma inverse_positive_iff_positive [simp]:
+  "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_ring_inverse_zero}))"
+apply (cases "a = 0", simp)
+apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
+done
+
+lemma inverse_negative_iff_negative [simp]:
+  "(inverse a < 0) = (a < (0::'a::{linordered_field,division_ring_inverse_zero}))"
+apply (cases "a = 0", simp)
+apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
+done
+
+lemma inverse_nonnegative_iff_nonnegative [simp]:
+  "(0 \<le> inverse a) = (0 \<le> (a::'a::{linordered_field,division_ring_inverse_zero}))"
+by (simp add: linorder_not_less [symmetric])
+
+lemma inverse_nonpositive_iff_nonpositive [simp]:
+  "(inverse a \<le> 0) = (a \<le> (0::'a::{linordered_field,division_ring_inverse_zero}))"
+by (simp add: linorder_not_less [symmetric])
 
 lemma one_less_inverse_iff:
-  "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))"
+  "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_ring_inverse_zero}))"
 proof cases
   assume "0 < x"
     with inverse_less_iff_less [OF zero_less_one, of x]
@@ -518,278 +699,78 @@
   with notless show ?thesis by simp
 qed
 
-lemma inverse_eq_1_iff [simp]:
-  "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
-by (insert inverse_eq_iff_eq [of x 1], simp) 
-
 lemma one_le_inverse_iff:
-  "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_by_zero}))"
+  "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_ring_inverse_zero}))"
 by (force simp add: order_le_less one_less_inverse_iff)
 
 lemma inverse_less_1_iff:
-  "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_by_zero}))"
+  "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_ring_inverse_zero}))"
 by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) 
 
 lemma inverse_le_1_iff:
-  "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{linordered_field,division_by_zero}))"
+  "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{linordered_field,division_ring_inverse_zero}))"
 by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) 
 
-
-subsection{*Simplification of Inequalities Involving Literal Divisors*}
-
-lemma pos_le_divide_eq: "0 < (c::'a::linordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
-proof -
-  assume less: "0<c"
-  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
-    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
-  also have "... = (a*c \<le> b)"
-    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
-lemma neg_le_divide_eq: "c < (0::'a::linordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
-proof -
-  assume less: "c<0"
-  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
-    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
-  also have "... = (b \<le> a*c)"
-    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
-lemma le_divide_eq:
-  "(a \<le> b/c) = 
-   (if 0 < c then a*c \<le> b
-             else if c < 0 then b \<le> a*c
-             else  a \<le> (0::'a::{linordered_field,division_by_zero}))"
-apply (cases "c=0", simp) 
-apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) 
-done
-
-lemma pos_divide_le_eq: "0 < (c::'a::linordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
-proof -
-  assume less: "0<c"
-  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
-    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
-  also have "... = (b \<le> a*c)"
-    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
-lemma neg_divide_le_eq: "c < (0::'a::linordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
-proof -
-  assume less: "c<0"
-  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
-    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
-  also have "... = (a*c \<le> b)"
-    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
 lemma divide_le_eq:
   "(b/c \<le> a) = 
    (if 0 < c then b \<le> a*c
              else if c < 0 then a*c \<le> b
-             else 0 \<le> (a::'a::{linordered_field,division_by_zero}))"
+             else 0 \<le> (a::'a::{linordered_field,division_ring_inverse_zero}))"
 apply (cases "c=0", simp) 
 apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) 
 done
 
-lemma pos_less_divide_eq:
-     "0 < (c::'a::linordered_field) ==> (a < b/c) = (a*c < b)"
-proof -
-  assume less: "0<c"
-  hence "(a < b/c) = (a*c < (b/c)*c)"
-    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
-  also have "... = (a*c < b)"
-    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
-lemma neg_less_divide_eq:
- "c < (0::'a::linordered_field) ==> (a < b/c) = (b < a*c)"
-proof -
-  assume less: "c<0"
-  hence "(a < b/c) = ((b/c)*c < a*c)"
-    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
-  also have "... = (b < a*c)"
-    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
 lemma less_divide_eq:
   "(a < b/c) = 
    (if 0 < c then a*c < b
              else if c < 0 then b < a*c
-             else  a < (0::'a::{linordered_field,division_by_zero}))"
+             else  a < (0::'a::{linordered_field,division_ring_inverse_zero}))"
 apply (cases "c=0", simp) 
 apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) 
 done
 
-lemma pos_divide_less_eq:
-     "0 < (c::'a::linordered_field) ==> (b/c < a) = (b < a*c)"
-proof -
-  assume less: "0<c"
-  hence "(b/c < a) = ((b/c)*c < a*c)"
-    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
-  also have "... = (b < a*c)"
-    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
-lemma neg_divide_less_eq:
- "c < (0::'a::linordered_field) ==> (b/c < a) = (a*c < b)"
-proof -
-  assume less: "c<0"
-  hence "(b/c < a) = (a*c < (b/c)*c)"
-    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
-  also have "... = (a*c < b)"
-    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
-  finally show ?thesis .
-qed
-
 lemma divide_less_eq:
   "(b/c < a) = 
    (if 0 < c then b < a*c
              else if c < 0 then a*c < b
-             else 0 < (a::'a::{linordered_field,division_by_zero}))"
+             else 0 < (a::'a::{linordered_field,division_ring_inverse_zero}))"
 apply (cases "c=0", simp) 
 apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) 
 done
 
-
-subsection{*Field simplification*}
-
-text{* Lemmas @{text field_simps} multiply with denominators in in(equations)
-if they can be proved to be non-zero (for equations) or positive/negative
-(for inequations). Can be too aggressive and is therefore separate from the
-more benign @{text algebra_simps}. *}
-
-lemmas field_simps[no_atp] = field_eq_simps
-  (* multiply ineqn *)
-  pos_divide_less_eq neg_divide_less_eq
-  pos_less_divide_eq neg_less_divide_eq
-  pos_divide_le_eq neg_divide_le_eq
-  pos_le_divide_eq neg_le_divide_eq
-
-text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
-of positivity/negativity needed for @{text field_simps}. Have not added @{text
-sign_simps} to @{text field_simps} because the former can lead to case
-explosions. *}
-
-lemmas sign_simps[no_atp] = group_simps
-  zero_less_mult_iff  mult_less_0_iff
-
-(* Only works once linear arithmetic is installed:
-text{*An example:*}
-lemma fixes a b c d e f :: "'a::linordered_field"
-shows "\<lbrakk>a>b; c<d; e<f; 0 < u \<rbrakk> \<Longrightarrow>
- ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) <
- ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u"
-apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0")
- prefer 2 apply(simp add:sign_simps)
-apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0")
- prefer 2 apply(simp add:sign_simps)
-apply(simp add:field_simps)
-done
-*)
-
-
-subsection{*Division and Signs*}
+text {*Division and Signs*}
 
 lemma zero_less_divide_iff:
-     "((0::'a::{linordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
+     "((0::'a::{linordered_field,division_ring_inverse_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
 by (simp add: divide_inverse zero_less_mult_iff)
 
 lemma divide_less_0_iff:
-     "(a/b < (0::'a::{linordered_field,division_by_zero})) = 
+     "(a/b < (0::'a::{linordered_field,division_ring_inverse_zero})) = 
       (0 < a & b < 0 | a < 0 & 0 < b)"
 by (simp add: divide_inverse mult_less_0_iff)
 
 lemma zero_le_divide_iff:
-     "((0::'a::{linordered_field,division_by_zero}) \<le> a/b) =
+     "((0::'a::{linordered_field,division_ring_inverse_zero}) \<le> a/b) =
       (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
 by (simp add: divide_inverse zero_le_mult_iff)
 
 lemma divide_le_0_iff:
-     "(a/b \<le> (0::'a::{linordered_field,division_by_zero})) =
+     "(a/b \<le> (0::'a::{linordered_field,division_ring_inverse_zero})) =
       (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
 by (simp add: divide_inverse mult_le_0_iff)
 
-lemma divide_eq_0_iff [simp,no_atp]:
-     "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
-by (simp add: divide_inverse)
-
-lemma divide_pos_pos:
-  "0 < (x::'a::linordered_field) ==> 0 < y ==> 0 < x / y"
-by(simp add:field_simps)
-
-
-lemma divide_nonneg_pos:
-  "0 <= (x::'a::linordered_field) ==> 0 < y ==> 0 <= x / y"
-by(simp add:field_simps)
-
-lemma divide_neg_pos:
-  "(x::'a::linordered_field) < 0 ==> 0 < y ==> x / y < 0"
-by(simp add:field_simps)
-
-lemma divide_nonpos_pos:
-  "(x::'a::linordered_field) <= 0 ==> 0 < y ==> x / y <= 0"
-by(simp add:field_simps)
-
-lemma divide_pos_neg:
-  "0 < (x::'a::linordered_field) ==> y < 0 ==> x / y < 0"
-by(simp add:field_simps)
-
-lemma divide_nonneg_neg:
-  "0 <= (x::'a::linordered_field) ==> y < 0 ==> x / y <= 0" 
-by(simp add:field_simps)
-
-lemma divide_neg_neg:
-  "(x::'a::linordered_field) < 0 ==> y < 0 ==> 0 < x / y"
-by(simp add:field_simps)
-
-lemma divide_nonpos_neg:
-  "(x::'a::linordered_field) <= 0 ==> y < 0 ==> 0 <= x / y"
-by(simp add:field_simps)
-
-
-subsection{*Cancellation Laws for Division*}
-
-lemma divide_cancel_right [simp,no_atp]:
-     "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
-apply (cases "c=0", simp)
-apply (simp add: divide_inverse)
-done
-
-lemma divide_cancel_left [simp,no_atp]:
-     "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" 
-apply (cases "c=0", simp)
-apply (simp add: divide_inverse)
-done
-
-
-subsection {* Division and the Number One *}
+text {* Division and the Number One *}
 
 text{*Simplify expressions equated with 1*}
-lemma divide_eq_1_iff [simp,no_atp]:
-     "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
-apply (cases "b=0", simp)
-apply (simp add: right_inverse_eq)
-done
-
-lemma one_eq_divide_iff [simp,no_atp]:
-     "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
-by (simp add: eq_commute [of 1])
 
 lemma zero_eq_1_divide_iff [simp,no_atp]:
-     "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)"
+     "((0::'a::{linordered_field,division_ring_inverse_zero}) = 1/a) = (a = 0)"
 apply (cases "a=0", simp)
 apply (auto simp add: nonzero_eq_divide_eq)
 done
 
 lemma one_divide_eq_0_iff [simp,no_atp]:
-     "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)"
+     "(1/a = (0::'a::{linordered_field,division_ring_inverse_zero})) = (a = 0)"
 apply (cases "a=0", simp)
 apply (insert zero_neq_one [THEN not_sym])
 apply (auto simp add: nonzero_divide_eq_eq)
@@ -806,247 +787,120 @@
 declare zero_le_divide_1_iff [simp,no_atp]
 declare divide_le_0_1_iff [simp,no_atp]
 
-
-subsection {* Ordering Rules for Division *}
-
-lemma divide_strict_right_mono:
-     "[|a < b; 0 < c|] ==> a / c < b / (c::'a::linordered_field)"
-by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono 
-              positive_imp_inverse_positive)
-
 lemma divide_right_mono:
-     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_by_zero})"
+     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_ring_inverse_zero})"
 by (force simp add: divide_strict_right_mono order_le_less)
 
-lemma divide_right_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b 
+lemma divide_right_mono_neg: "(a::'a::{linordered_field,division_ring_inverse_zero}) <= b 
     ==> c <= 0 ==> b / c <= a / c"
 apply (drule divide_right_mono [of _ _ "- c"])
 apply auto
 done
 
-lemma divide_strict_right_mono_neg:
-     "[|b < a; c < 0|] ==> a / c < b / (c::'a::linordered_field)"
-apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
-apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric])
-done
-
-text{*The last premise ensures that @{term a} and @{term b} 
-      have the same sign*}
-lemma divide_strict_left_mono:
-  "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
-by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono)
-
-lemma divide_left_mono:
-  "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::linordered_field)"
-by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono)
-
-lemma divide_left_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b 
+lemma divide_left_mono_neg: "(a::'a::{linordered_field,division_ring_inverse_zero}) <= b 
     ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b"
   apply (drule divide_left_mono [of _ _ "- c"])
   apply (auto simp add: mult_commute)
 done
 
-lemma divide_strict_left_mono_neg:
-  "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
-by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg)
 
 
 text{*Simplify quotients that are compared with the value 1.*}
 
 lemma le_divide_eq_1 [no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
 by (auto simp add: le_divide_eq)
 
 lemma divide_le_eq_1 [no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
 by (auto simp add: divide_le_eq)
 
 lemma less_divide_eq_1 [no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
 by (auto simp add: less_divide_eq)
 
 lemma divide_less_eq_1 [no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
 by (auto simp add: divide_less_eq)
 
 
-subsection{*Conditional Simplification Rules: No Case Splits*}
+text {*Conditional Simplification Rules: No Case Splits*}
 
 lemma le_divide_eq_1_pos [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
 by (auto simp add: le_divide_eq)
 
 lemma le_divide_eq_1_neg [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
 by (auto simp add: le_divide_eq)
 
 lemma divide_le_eq_1_pos [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
 by (auto simp add: divide_le_eq)
 
 lemma divide_le_eq_1_neg [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
 by (auto simp add: divide_le_eq)
 
 lemma less_divide_eq_1_pos [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
 by (auto simp add: less_divide_eq)
 
 lemma less_divide_eq_1_neg [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
 by (auto simp add: less_divide_eq)
 
 lemma divide_less_eq_1_pos [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
 by (auto simp add: divide_less_eq)
 
 lemma divide_less_eq_1_neg [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
 by (auto simp add: divide_less_eq)
 
 lemma eq_divide_eq_1 [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "(1 = b/a) = ((a \<noteq> 0 & a = b))"
 by (auto simp add: eq_divide_eq)
 
 lemma divide_eq_eq_1 [simp,no_atp]:
-  fixes a :: "'a :: {linordered_field,division_by_zero}"
+  fixes a :: "'a :: {linordered_field,division_ring_inverse_zero}"
   shows "(b/a = 1) = ((a \<noteq> 0 & a = b))"
 by (auto simp add: divide_eq_eq)
 
-
-subsection {* Reasoning about inequalities with division *}
-
-lemma mult_imp_div_pos_le: "0 < (y::'a::linordered_field) ==> x <= z * y ==>
-    x / y <= z"
-by (subst pos_divide_le_eq, assumption+)
-
-lemma mult_imp_le_div_pos: "0 < (y::'a::linordered_field) ==> z * y <= x ==>
-    z <= x / y"
-by(simp add:field_simps)
-
-lemma mult_imp_div_pos_less: "0 < (y::'a::linordered_field) ==> x < z * y ==>
-    x / y < z"
-by(simp add:field_simps)
-
-lemma mult_imp_less_div_pos: "0 < (y::'a::linordered_field) ==> z * y < x ==>
-    z < x / y"
-by(simp add:field_simps)
-
-lemma frac_le: "(0::'a::linordered_field) <= x ==> 
-    x <= y ==> 0 < w ==> w <= z  ==> x / z <= y / w"
-  apply (rule mult_imp_div_pos_le)
-  apply simp
-  apply (subst times_divide_eq_left)
-  apply (rule mult_imp_le_div_pos, assumption)
-  apply (rule mult_mono)
-  apply simp_all
-done
-
-lemma frac_less: "(0::'a::linordered_field) <= x ==> 
-    x < y ==> 0 < w ==> w <= z  ==> x / z < y / w"
-  apply (rule mult_imp_div_pos_less)
-  apply simp
-  apply (subst times_divide_eq_left)
-  apply (rule mult_imp_less_div_pos, assumption)
-  apply (erule mult_less_le_imp_less)
-  apply simp_all
-done
-
-lemma frac_less2: "(0::'a::linordered_field) < x ==> 
-    x <= y ==> 0 < w ==> w < z  ==> x / z < y / w"
-  apply (rule mult_imp_div_pos_less)
-  apply simp_all
-  apply (subst times_divide_eq_left)
-  apply (rule mult_imp_less_div_pos, assumption)
-  apply (erule mult_le_less_imp_less)
-  apply simp_all
-done
-
-text{*It's not obvious whether these should be simprules or not. 
-  Their effect is to gather terms into one big fraction, like
-  a*b*c / x*y*z. The rationale for that is unclear, but many proofs 
-  seem to need them.*}
-
-declare times_divide_eq [simp]
-
-
-subsection {* Ordered Fields are Dense *}
-
-lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::linordered_field)"
-by (simp add: field_simps zero_less_two)
-
-lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::linordered_field) < b"
-by (simp add: field_simps zero_less_two)
-
-instance linordered_field < dense_linorder
-proof
-  fix x y :: 'a
-  have "x < x + 1" by simp
-  then show "\<exists>y. x < y" .. 
-  have "x - 1 < x" by simp
-  then show "\<exists>y. y < x" ..
-  show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
-qed
-
-
-subsection {* Absolute Value *}
-
-lemma nonzero_abs_inverse:
-     "a \<noteq> 0 ==> abs (inverse (a::'a::linordered_field)) = inverse (abs a)"
-apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq 
-                      negative_imp_inverse_negative)
-apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) 
-done
-
 lemma abs_inverse [simp]:
-     "abs (inverse (a::'a::{linordered_field,division_by_zero})) = 
-      inverse (abs a)"
+     "\<bar>inverse (a::'a::{linordered_field,division_ring_inverse_zero})\<bar> = 
+      inverse \<bar>a\<bar>"
 apply (cases "a=0", simp) 
 apply (simp add: nonzero_abs_inverse) 
 done
 
-lemma nonzero_abs_divide:
-     "b \<noteq> 0 ==> abs (a / (b::'a::linordered_field)) = abs a / abs b"
-by (simp add: divide_inverse abs_mult nonzero_abs_inverse) 
-
 lemma abs_divide [simp]:
-     "abs (a / (b::'a::{linordered_field,division_by_zero})) = abs a / abs b"
+     "\<bar>a / (b::'a::{linordered_field,division_ring_inverse_zero})\<bar> = \<bar>a\<bar> / \<bar>b\<bar>"
 apply (cases "b=0", simp) 
 apply (simp add: nonzero_abs_divide) 
 done
 
-lemma abs_div_pos: "(0::'a::{division_by_zero,linordered_field}) < y ==> 
-    abs x / y = abs (x / y)"
+lemma abs_div_pos: "(0::'a::{linordered_field,division_ring_inverse_zero}) < y ==> 
+    \<bar>x\<bar> / y = \<bar>x / y\<bar>"
   apply (subst abs_divide)
   apply (simp add: order_less_imp_le)
 done
 
-lemma field_le_epsilon:
-  fixes x y :: "'a\<Colon>linordered_field"
-  assumes e: "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
-  shows "x \<le> y"
-proof (rule dense_le)
-  fix t assume "t < x"
-  hence "0 < x - t" by (simp add: less_diff_eq)
-  from e[OF this]
-  show "t \<le> y" by (simp add: field_simps)
-qed
-
 lemma field_le_mult_one_interval:
-  fixes x :: "'a\<Colon>{linordered_field,division_by_zero}"
+  fixes x :: "'a\<Colon>{linordered_field,division_ring_inverse_zero}"
   assumes *: "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
   shows "x \<le> y"
 proof (cases "0 < x")
--- a/src/HOL/GCD.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/GCD.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -1034,11 +1034,11 @@
     thus "fst (bezw m n) * int m + snd (bezw m n) * int n = int (gcd m n)"
       apply (simp add: bezw_non_0 gcd_non_0_nat)
       apply (erule subst)
-      apply (simp add: ring_simps)
+      apply (simp add: field_simps)
       apply (subst mod_div_equality [of m n, symmetric])
       (* applying simp here undoes the last substitution!
          what is procedure cancel_div_mod? *)
-      apply (simp only: ring_simps zadd_int [symmetric]
+      apply (simp only: field_simps zadd_int [symmetric]
         zmult_int [symmetric])
       done
 qed
@@ -1389,7 +1389,7 @@
   show "lcm (lcm n m) p = lcm n (lcm m p)"
     by (rule lcm_unique_nat [THEN iffD1]) (metis dvd.order_trans lcm_unique_nat)
   show "lcm m n = lcm n m"
-    by (simp add: lcm_nat_def gcd_commute_nat ring_simps)
+    by (simp add: lcm_nat_def gcd_commute_nat field_simps)
 qed
 
 interpretation lcm_int!: abel_semigroup "lcm :: int \<Rightarrow> int \<Rightarrow> int"
--- a/src/HOL/Groebner_Basis.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Groebner_Basis.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -474,20 +474,20 @@
   fieldgb "op +" "op *" "op ^" "0::'a::{field,number_ring}" "1" "op -" "uminus" "op /" "inverse" apply (unfold_locales) by (simp_all add: divide_inverse)
 
 lemma divide_Numeral1: "(x::'a::{field,number_ring}) / Numeral1 = x" by simp
-lemma divide_Numeral0: "(x::'a::{field,number_ring, division_by_zero}) / Numeral0 = 0"
+lemma divide_Numeral0: "(x::'a::{field,number_ring, division_ring_inverse_zero}) / Numeral0 = 0"
   by simp
-lemma mult_frac_frac: "((x::'a::{field,division_by_zero}) / y) * (z / w) = (x*z) / (y*w)"
+lemma mult_frac_frac: "((x::'a::{field,division_ring_inverse_zero}) / y) * (z / w) = (x*z) / (y*w)"
   by simp
-lemma mult_frac_num: "((x::'a::{field, division_by_zero}) / y) * z  = (x*z) / y"
+lemma mult_frac_num: "((x::'a::{field, division_ring_inverse_zero}) / y) * z  = (x*z) / y"
   by simp
-lemma mult_num_frac: "((x::'a::{field, division_by_zero}) / y) * z  = (x*z) / y"
+lemma mult_num_frac: "((x::'a::{field, division_ring_inverse_zero}) / y) * z  = (x*z) / y"
   by simp
 
 lemma Numeral1_eq1_nat: "(1::nat) = Numeral1" by simp
 
-lemma add_frac_num: "y\<noteq> 0 \<Longrightarrow> (x::'a::{field, division_by_zero}) / y + z = (x + z*y) / y"
+lemma add_frac_num: "y\<noteq> 0 \<Longrightarrow> (x::'a::{field, division_ring_inverse_zero}) / y + z = (x + z*y) / y"
   by (simp add: add_divide_distrib)
-lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::{field, division_by_zero}) / y = (x + z*y) / y"
+lemma add_num_frac: "y\<noteq> 0 \<Longrightarrow> z + (x::'a::{field, division_ring_inverse_zero}) / y = (x + z*y) / y"
   by (simp add: add_divide_distrib)
 
 ML {*
@@ -627,7 +627,7 @@
 
 val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
            @{thm "divide_Numeral1"},
-           @{thm "Fields.divide_zero"}, @{thm "divide_Numeral0"},
+           @{thm "divide_zero"}, @{thm "divide_Numeral0"},
            @{thm "divide_divide_eq_left"}, @{thm "mult_frac_frac"},
            @{thm "mult_num_frac"}, @{thm "mult_frac_num"},
            @{thm "mult_frac_frac"}, @{thm "times_divide_eq_right"},
--- a/src/HOL/Groups.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Groups.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -12,13 +12,13 @@
 subsection {* Fact collections *}
 
 ML {*
-structure Algebra_Simps = Named_Thms(
-  val name = "algebra_simps"
-  val description = "algebra simplification rules"
+structure Ac_Simps = Named_Thms(
+  val name = "ac_simps"
+  val description = "associativity and commutativity simplification rules"
 )
 *}
 
-setup Algebra_Simps.setup
+setup Ac_Simps.setup
 
 text{* The rewrites accumulated in @{text algebra_simps} deal with the
 classical algebraic structures of groups, rings and family. They simplify
@@ -29,15 +29,28 @@
 Of course it also works for fields, but it knows nothing about multiplicative
 inverses or division. This is catered for by @{text field_simps}. *}
 
-
 ML {*
-structure Ac_Simps = Named_Thms(
-  val name = "ac_simps"
-  val description = "associativity and commutativity simplification rules"
+structure Algebra_Simps = Named_Thms(
+  val name = "algebra_simps"
+  val description = "algebra simplification rules"
 )
 *}
 
-setup Ac_Simps.setup
+setup Algebra_Simps.setup
+
+text{* Lemmas @{text field_simps} multiply with denominators in (in)equations
+if they can be proved to be non-zero (for equations) or positive/negative
+(for inequations). Can be too aggressive and is therefore separate from the
+more benign @{text algebra_simps}. *}
+
+ML {*
+structure Field_Simps = Named_Thms(
+  val name = "field_simps"
+  val description = "algebra simplification rules for fields"
+)
+*}
+
+setup Field_Simps.setup
 
 
 subsection {* Abstract structures *}
@@ -139,13 +152,13 @@
 subsection {* Semigroups and Monoids *}
 
 class semigroup_add = plus +
-  assumes add_assoc [algebra_simps]: "(a + b) + c = a + (b + c)"
+  assumes add_assoc [algebra_simps, field_simps]: "(a + b) + c = a + (b + c)"
 
 sublocale semigroup_add < add!: semigroup plus proof
 qed (fact add_assoc)
 
 class ab_semigroup_add = semigroup_add +
-  assumes add_commute [algebra_simps]: "a + b = b + a"
+  assumes add_commute [algebra_simps, field_simps]: "a + b = b + a"
 
 sublocale ab_semigroup_add < add!: abel_semigroup plus proof
 qed (fact add_commute)
@@ -153,7 +166,7 @@
 context ab_semigroup_add
 begin
 
-lemmas add_left_commute [algebra_simps] = add.left_commute
+lemmas add_left_commute [algebra_simps, field_simps] = add.left_commute
 
 theorems add_ac = add_assoc add_commute add_left_commute
 
@@ -162,13 +175,13 @@
 theorems add_ac = add_assoc add_commute add_left_commute
 
 class semigroup_mult = times +
-  assumes mult_assoc [algebra_simps]: "(a * b) * c = a * (b * c)"
+  assumes mult_assoc [algebra_simps, field_simps]: "(a * b) * c = a * (b * c)"
 
 sublocale semigroup_mult < mult!: semigroup times proof
 qed (fact mult_assoc)
 
 class ab_semigroup_mult = semigroup_mult +
-  assumes mult_commute [algebra_simps]: "a * b = b * a"
+  assumes mult_commute [algebra_simps, field_simps]: "a * b = b * a"
 
 sublocale ab_semigroup_mult < mult!: abel_semigroup times proof
 qed (fact mult_commute)
@@ -176,7 +189,7 @@
 context ab_semigroup_mult
 begin
 
-lemmas mult_left_commute [algebra_simps] = mult.left_commute
+lemmas mult_left_commute [algebra_simps, field_simps] = mult.left_commute
 
 theorems mult_ac = mult_assoc mult_commute mult_left_commute
 
@@ -371,7 +384,7 @@
 lemma add_diff_cancel: "a + b - b = a"
 by (simp add: diff_minus add_assoc)
 
-declare diff_minus[symmetric, algebra_simps]
+declare diff_minus[symmetric, algebra_simps, field_simps]
 
 lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
 proof
@@ -402,7 +415,7 @@
   then show "b = c" by simp
 qed
 
-lemma uminus_add_conv_diff[algebra_simps]:
+lemma uminus_add_conv_diff[algebra_simps, field_simps]:
   "- a + b = b - a"
 by (simp add:diff_minus add_commute)
 
@@ -414,22 +427,22 @@
   "- (a - b) = b - a"
 by (simp add: diff_minus add_commute)
 
-lemma add_diff_eq[algebra_simps]: "a + (b - c) = (a + b) - c"
+lemma add_diff_eq[algebra_simps, field_simps]: "a + (b - c) = (a + b) - c"
 by (simp add: diff_minus add_ac)
 
-lemma diff_add_eq[algebra_simps]: "(a - b) + c = (a + c) - b"
+lemma diff_add_eq[algebra_simps, field_simps]: "(a - b) + c = (a + c) - b"
 by (simp add: diff_minus add_ac)
 
-lemma diff_eq_eq[algebra_simps]: "a - b = c \<longleftrightarrow> a = c + b"
+lemma diff_eq_eq[algebra_simps, field_simps]: "a - b = c \<longleftrightarrow> a = c + b"
 by (auto simp add: diff_minus add_assoc)
 
-lemma eq_diff_eq[algebra_simps]: "a = c - b \<longleftrightarrow> a + b = c"
+lemma eq_diff_eq[algebra_simps, field_simps]: "a = c - b \<longleftrightarrow> a + b = c"
 by (auto simp add: diff_minus add_assoc)
 
-lemma diff_diff_eq[algebra_simps]: "(a - b) - c = a - (b + c)"
+lemma diff_diff_eq[algebra_simps, field_simps]: "(a - b) - c = a - (b + c)"
 by (simp add: diff_minus add_ac)
 
-lemma diff_diff_eq2[algebra_simps]: "a - (b - c) = (a + c) - b"
+lemma diff_diff_eq2[algebra_simps, field_simps]: "a - (b - c) = (a + c) - b"
 by (simp add: diff_minus add_ac)
 
 lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
@@ -750,35 +763,29 @@
   finally show ?thesis .
 qed
 
-lemma diff_less_eq[algebra_simps]: "a - b < c \<longleftrightarrow> a < c + b"
+lemma diff_less_eq[algebra_simps, field_simps]: "a - b < c \<longleftrightarrow> a < c + b"
 apply (subst less_iff_diff_less_0 [of a])
 apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
 apply (simp add: diff_minus add_ac)
 done
 
-lemma less_diff_eq[algebra_simps]: "a < c - b \<longleftrightarrow> a + b < c"
-apply (subst less_iff_diff_less_0 [of "plus a b"])
+lemma less_diff_eq[algebra_simps, field_simps]: "a < c - b \<longleftrightarrow> a + b < c"
+apply (subst less_iff_diff_less_0 [of "a + b"])
 apply (subst less_iff_diff_less_0 [of a])
 apply (simp add: diff_minus add_ac)
 done
 
-lemma diff_le_eq[algebra_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
+lemma diff_le_eq[algebra_simps, field_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
 by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
 
-lemma le_diff_eq[algebra_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
+lemma le_diff_eq[algebra_simps, field_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
 by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
 
 lemma le_iff_diff_le_0: "a \<le> b \<longleftrightarrow> a - b \<le> 0"
 by (simp add: algebra_simps)
 
-text{*Legacy - use @{text algebra_simps} *}
-lemmas group_simps[no_atp] = algebra_simps
-
 end
 
-text{*Legacy - use @{text algebra_simps} *}
-lemmas group_simps[no_atp] = algebra_simps
-
 class linordered_ab_semigroup_add =
   linorder + ordered_ab_semigroup_add
 
@@ -966,27 +973,26 @@
 
 end
 
--- {* FIXME localize the following *}
+context ordered_comm_monoid_add
+begin
 
 lemma add_increasing:
-  fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
-  shows  "[|0\<le>a; b\<le>c|] ==> b \<le> a + c"
-by (insert add_mono [of 0 a b c], simp)
+  "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
+  by (insert add_mono [of 0 a b c], simp)
 
 lemma add_increasing2:
-  fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
-  shows  "[|0\<le>c; b\<le>a|] ==> b \<le> a + c"
-by (simp add:add_increasing add_commute[of a])
+  "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
+  by (simp add: add_increasing add_commute [of a])
 
 lemma add_strict_increasing:
-  fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
-  shows "[|0<a; b\<le>c|] ==> b < a + c"
-by (insert add_less_le_mono [of 0 a b c], simp)
+  "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
+  by (insert add_less_le_mono [of 0 a b c], simp)
 
 lemma add_strict_increasing2:
-  fixes c :: "'a::{ordered_ab_semigroup_add_imp_le, comm_monoid_add}"
-  shows "[|0\<le>a; b<c|] ==> b < a + c"
-by (insert add_le_less_mono [of 0 a b c], simp)
+  "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
+  by (insert add_le_less_mono [of 0 a b c], simp)
+
+end
 
 class abs =
   fixes abs :: "'a \<Rightarrow> 'a"
@@ -1036,7 +1042,7 @@
 
 lemma abs_idempotent [simp]: "\<bar>\<bar>a\<bar>\<bar> = \<bar>a\<bar>"
 by (rule antisym)
-   (auto intro!: abs_ge_self abs_leI order_trans [of "uminus (abs a)" zero "abs a"])
+   (auto intro!: abs_ge_self abs_leI order_trans [of "- \<bar>a\<bar>" 0 "\<bar>a\<bar>"])
 
 lemma abs_eq_0 [simp]: "\<bar>a\<bar> = 0 \<longleftrightarrow> a = 0"
 proof -
@@ -1045,7 +1051,7 @@
     assume zero: "\<bar>a\<bar> = 0"
     with abs_ge_self show "a \<le> 0" by auto
     from zero have "\<bar>-a\<bar> = 0" by simp
-    with abs_ge_self [of "uminus a"] have "- a \<le> 0" by auto
+    with abs_ge_self [of "- a"] have "- a \<le> 0" by auto
     with neg_le_0_iff_le show "0 \<le> a" by auto
   qed
   then show ?thesis by auto
@@ -1114,32 +1120,31 @@
 by (insert abs_ge_self, blast intro: order_trans)
 
 lemma abs_le_D2: "\<bar>a\<bar> \<le> b \<Longrightarrow> - a \<le> b"
-by (insert abs_le_D1 [of "uminus a"], simp)
+by (insert abs_le_D1 [of "- a"], simp)
 
 lemma abs_le_iff: "\<bar>a\<bar> \<le> b \<longleftrightarrow> a \<le> b \<and> - a \<le> b"
 by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2)
 
 lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
-  apply (simp add: algebra_simps)
-  apply (subgoal_tac "abs a = abs (plus b (minus a b))")
-  apply (erule ssubst)
-  apply (rule abs_triangle_ineq)
-  apply (rule arg_cong[of _ _ abs])
-  apply (simp add: algebra_simps)
-done
+proof -
+  have "\<bar>a\<bar> = \<bar>b + (a - b)\<bar>"
+    by (simp add: algebra_simps add_diff_cancel)
+  then have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>a - b\<bar>"
+    by (simp add: abs_triangle_ineq)
+  then show ?thesis
+    by (simp add: algebra_simps)
+qed
+
+lemma abs_triangle_ineq2_sym: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>b - a\<bar>"
+  by (simp only: abs_minus_commute [of b] abs_triangle_ineq2)
 
 lemma abs_triangle_ineq3: "\<bar>\<bar>a\<bar> - \<bar>b\<bar>\<bar> \<le> \<bar>a - b\<bar>"
-  apply (subst abs_le_iff)
-  apply auto
-  apply (rule abs_triangle_ineq2)
-  apply (subst abs_minus_commute)
-  apply (rule abs_triangle_ineq2)
-done
+  by (simp add: abs_le_iff abs_triangle_ineq2 abs_triangle_ineq2_sym)
 
 lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
 proof -
-  have "abs(a - b) = abs(a + - b)" by (subst diff_minus, rule refl)
-  also have "... <= abs a + abs (- b)" by (rule abs_triangle_ineq)
+  have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (subst diff_minus, rule refl)
+  also have "... \<le> \<bar>a\<bar> + \<bar>- b\<bar>" by (rule abs_triangle_ineq)
   finally show ?thesis by simp
 qed
 
--- a/src/HOL/HOL.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/HOL.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -30,7 +30,6 @@
   ("~~/src/Tools/induct_tacs.ML")
   ("Tools/recfun_codegen.ML")
   "~~/src/Tools/more_conv.ML"
-  "~~/src/HOL/Tools/Sledgehammer/named_thm_set.ML"
 begin
 
 setup {* Intuitionistic.method_setup @{binding iprover} *}
@@ -74,7 +73,7 @@
 local
 
 consts
-  If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" 10)
+  If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10)
 
 
 subsubsection {* Additional concrete syntax *}
@@ -119,7 +118,7 @@
   "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
   ""            :: "letbind => letbinds"                 ("_")
   "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
-  "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
+  "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" [0, 10] 10)
 
   "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
   "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
@@ -801,7 +800,7 @@
 *}
 
 ML {*
-structure No_ATPs = Named_Thm_Set
+structure No_ATPs = Named_Thms
 (
   val name = "no_atp"
   val description = "theorems that should be filtered out by Sledgehammer"
@@ -1870,7 +1869,7 @@
 proof
   assume "PROP ?ofclass"
   show "PROP ?eq"
-    by (tactic {* ALLGOALS (rtac (Drule.unconstrainTs @{thm equals_eq})) *}) 
+    by (tactic {* ALLGOALS (rtac (Thm.unconstrain_allTs @{thm equals_eq})) *}) 
       (fact `PROP ?ofclass`)
 next
   assume "PROP ?eq"
--- a/src/HOL/Import/HOL/real.imp	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Import/HOL/real.imp	Mon Apr 26 21:25:32 2010 +0200
@@ -251,7 +251,7 @@
   "REAL_INV_INV" > "Rings.inverse_inverse_eq"
   "REAL_INV_EQ_0" > "Rings.inverse_nonzero_iff_nonzero"
   "REAL_INV_1OVER" > "Rings.inverse_eq_divide"
-  "REAL_INV_0" > "Rings.division_by_zero_class.inverse_zero"
+  "REAL_INV_0" > "Rings.division_ring_inverse_zero_class.inverse_zero"
   "REAL_INVINV" > "Rings.nonzero_inverse_inverse_eq"
   "REAL_INV1" > "Rings.inverse_1"
   "REAL_INJ" > "RealDef.real_of_nat_inject"
--- a/src/HOL/Import/HOL/realax.imp	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Import/HOL/realax.imp	Mon Apr 26 21:25:32 2010 +0200
@@ -101,7 +101,7 @@
   "REAL_LT_MUL" > "Rings.mult_pos_pos"
   "REAL_LT_IADD" > "Groups.add_strict_left_mono"
   "REAL_LDISTRIB" > "Rings.ring_eq_simps_2"
-  "REAL_INV_0" > "Rings.division_by_zero_class.inverse_zero"
+  "REAL_INV_0" > "Rings.division_ring_inverse_zero_class.inverse_zero"
   "REAL_ADD_SYM" > "Finite_Set.AC_add.f.AC_2"
   "REAL_ADD_LINV" > "HOL4Compat.REAL_ADD_LINV"
   "REAL_ADD_LID" > "Finite_Set.AC_add.f_e.left_ident"
--- a/src/HOL/Int.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Int.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -1995,15 +1995,15 @@
 text{*Division By @{text "-1"}*}
 
 lemma divide_minus1 [simp]:
-     "x/-1 = -(x::'a::{field,division_by_zero,number_ring})"
+     "x/-1 = -(x::'a::{field,division_ring_inverse_zero,number_ring})"
 by simp
 
 lemma minus1_divide [simp]:
-     "-1 / (x::'a::{field,division_by_zero,number_ring}) = - (1/x)"
+     "-1 / (x::'a::{field,division_ring_inverse_zero,number_ring}) = - (1/x)"
 by (simp add: divide_inverse)
 
 lemma half_gt_zero_iff:
-     "(0 < r/2) = (0 < (r::'a::{linordered_field,division_by_zero,number_ring}))"
+     "(0 < r/2) = (0 < (r::'a::{linordered_field,division_ring_inverse_zero,number_ring}))"
 by auto
 
 lemmas half_gt_zero [simp] = half_gt_zero_iff [THEN iffD2, standard]
--- a/src/HOL/IsaMakefile	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/IsaMakefile	Mon Apr 26 21:25:32 2010 +0200
@@ -136,7 +136,6 @@
   $(SRC)/Tools/random_word.ML \
   $(SRC)/Tools/value.ML \
   HOL.thy \
-  Tools/Sledgehammer/named_thm_set.ML \
   Tools/hologic.ML \
   Tools/recfun_codegen.ML \
   Tools/simpdata.ML
@@ -1080,6 +1079,8 @@
 
 $(OUT)/HOL-Multivariate_Analysis: $(OUT)/HOL-SMT	\
   Multivariate_Analysis/ROOT.ML				\
+  Multivariate_Analysis/document/root.tex		\
+  Multivariate_Analysis/L2_Norm.thy			\
   Multivariate_Analysis/Multivariate_Analysis.thy	\
   Multivariate_Analysis/Determinants.thy		\
   Multivariate_Analysis/Finite_Cartesian_Product.thy	\
--- a/src/HOL/Lambda/Type.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Lambda/Type.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -83,10 +83,10 @@
 
 subsection {* Some examples *}
 
-lemma "e \<turnstile> Abs (Abs (Abs (Var 1 \<degree> (Var 2 \<degree> Var 1 \<degree> Var 0)))) : ?T"
+schematic_lemma "e \<turnstile> Abs (Abs (Abs (Var 1 \<degree> (Var 2 \<degree> Var 1 \<degree> Var 0)))) : ?T"
   by force
 
-lemma "e \<turnstile> Abs (Abs (Abs (Var 2 \<degree> Var 0 \<degree> (Var 1 \<degree> Var 0)))) : ?T"
+schematic_lemma "e \<turnstile> Abs (Abs (Abs (Var 2 \<degree> Var 0 \<degree> (Var 1 \<degree> Var 0)))) : ?T"
   by force
 
 
--- a/src/HOL/Lattices.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Lattices.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -365,13 +365,9 @@
 
 subsection {* Bounded lattices and boolean algebras *}
 
-class bounded_lattice = lattice + top + bot
+class bounded_lattice_bot = lattice + bot
 begin
 
-lemma dual_bounded_lattice:
-  "bounded_lattice (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) \<top> \<bottom>"
-  by unfold_locales (auto simp add: less_le_not_le)
-
 lemma inf_bot_left [simp]:
   "\<bottom> \<sqinter> x = \<bottom>"
   by (rule inf_absorb1) simp
@@ -380,6 +376,23 @@
   "x \<sqinter> \<bottom> = \<bottom>"
   by (rule inf_absorb2) simp
 
+lemma sup_bot_left [simp]:
+  "\<bottom> \<squnion> x = x"
+  by (rule sup_absorb2) simp
+
+lemma sup_bot_right [simp]:
+  "x \<squnion> \<bottom> = x"
+  by (rule sup_absorb1) simp
+
+lemma sup_eq_bot_iff [simp]:
+  "x \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
+  by (simp add: eq_iff)
+
+end
+
+class bounded_lattice_top = lattice + top
+begin
+
 lemma sup_top_left [simp]:
   "\<top> \<squnion> x = \<top>"
   by (rule sup_absorb1) simp
@@ -396,21 +409,18 @@
   "x \<sqinter> \<top> = x"
   by (rule inf_absorb1) simp
 
-lemma sup_bot_left [simp]:
-  "\<bottom> \<squnion> x = x"
-  by (rule sup_absorb2) simp
-
-lemma sup_bot_right [simp]:
-  "x \<squnion> \<bottom> = x"
-  by (rule sup_absorb1) simp
-
 lemma inf_eq_top_iff [simp]:
   "x \<sqinter> y = \<top> \<longleftrightarrow> x = \<top> \<and> y = \<top>"
   by (simp add: eq_iff)
 
-lemma sup_eq_bot_iff [simp]:
-  "x \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>"
-  by (simp add: eq_iff)
+end
+
+class bounded_lattice = bounded_lattice_bot + bounded_lattice_top
+begin
+
+lemma dual_bounded_lattice:
+  "bounded_lattice (op \<ge>) (op >) (op \<squnion>) (op \<sqinter>) \<top> \<bottom>"
+  by unfold_locales (auto simp add: less_le_not_le)
 
 end
 
--- a/src/HOL/Library/Abstract_Rat.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Abstract_Rat.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -184,7 +184,7 @@
 
 lemma isnormNum_unique[simp]: 
   assumes na: "isnormNum x" and nb: "isnormNum y" 
-  shows "((INum x ::'a::{ring_char_0,field, division_by_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
+  shows "((INum x ::'a::{ring_char_0,field, division_ring_inverse_zero}) = INum y) = (x = y)" (is "?lhs = ?rhs")
 proof
   have "\<exists> a b a' b'. x = (a,b) \<and> y = (a',b')" by auto
   then obtain a b a' b' where xy[simp]: "x = (a,b)" "y=(a',b')" by blast
@@ -217,7 +217,7 @@
 qed
 
 
-lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> (INum x = (0::'a::{ring_char_0, field,division_by_zero})) = (x = 0\<^sub>N)"
+lemma isnormNum0[simp]: "isnormNum x \<Longrightarrow> (INum x = (0::'a::{ring_char_0, field,division_ring_inverse_zero})) = (x = 0\<^sub>N)"
   unfolding INum_int(2)[symmetric]
   by (rule isnormNum_unique, simp_all)
 
@@ -245,7 +245,7 @@
 done
 
 
-lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{ring_char_0,field, division_by_zero})"
+lemma normNum[simp]: "INum (normNum x) = (INum x :: 'a::{ring_char_0,field, division_ring_inverse_zero})"
 proof-
   have "\<exists> a b. x = (a,b)" by auto
   then obtain a b where x[simp]: "x = (a,b)" by blast
@@ -260,7 +260,7 @@
   ultimately show ?thesis by blast
 qed
 
-lemma INum_normNum_iff: "(INum x ::'a::{field, division_by_zero, ring_char_0}) = INum y \<longleftrightarrow> normNum x = normNum y" (is "?lhs = ?rhs")
+lemma INum_normNum_iff: "(INum x ::'a::{field, division_ring_inverse_zero, ring_char_0}) = INum y \<longleftrightarrow> normNum x = normNum y" (is "?lhs = ?rhs")
 proof -
   have "normNum x = normNum y \<longleftrightarrow> (INum (normNum x) :: 'a) = INum (normNum y)"
     by (simp del: normNum)
@@ -268,7 +268,7 @@
   finally show ?thesis by simp
 qed
 
-lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {ring_char_0,division_by_zero,field})"
+lemma Nadd[simp]: "INum (x +\<^sub>N y) = INum x + (INum y :: 'a :: {ring_char_0,division_ring_inverse_zero,field})"
 proof-
 let ?z = "0:: 'a"
   have " \<exists> a b. x = (a,b)" " \<exists> a' b'. y = (a',b')" by auto
@@ -300,7 +300,7 @@
   ultimately show ?thesis by blast
 qed
 
-lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {ring_char_0,division_by_zero,field}) "
+lemma Nmul[simp]: "INum (x *\<^sub>N y) = INum x * (INum y:: 'a :: {ring_char_0,division_ring_inverse_zero,field}) "
 proof-
   let ?z = "0::'a"
   have " \<exists> a b. x = (a,b)" " \<exists> a' b'. y = (a',b')" by auto
@@ -323,16 +323,16 @@
 lemma Nneg[simp]: "INum (~\<^sub>N x) = - (INum x ::'a:: field)"
   by (simp add: Nneg_def split_def INum_def)
 
-lemma Nsub[simp]: shows "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {ring_char_0,division_by_zero,field})"
+lemma Nsub[simp]: shows "INum (x -\<^sub>N y) = INum x - (INum y:: 'a :: {ring_char_0,division_ring_inverse_zero,field})"
 by (simp add: Nsub_def split_def)
 
-lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: {division_by_zero,field}) / (INum x)"
+lemma Ninv[simp]: "INum (Ninv x) = (1::'a :: {division_ring_inverse_zero,field}) / (INum x)"
   by (simp add: Ninv_def INum_def split_def)
 
-lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {ring_char_0, division_by_zero,field})" by (simp add: Ndiv_def)
+lemma Ndiv[simp]: "INum (x \<div>\<^sub>N y) = INum x / (INum y ::'a :: {ring_char_0, division_ring_inverse_zero,field})" by (simp add: Ndiv_def)
 
 lemma Nlt0_iff[simp]: assumes nx: "isnormNum x" 
-  shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field})< 0) = 0>\<^sub>N x "
+  shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field})< 0) = 0>\<^sub>N x "
 proof-
   have " \<exists> a b. x = (a,b)" by simp
   then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -345,7 +345,7 @@
 qed
 
 lemma Nle0_iff[simp]:assumes nx: "isnormNum x" 
-  shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
+  shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field}) \<le> 0) = 0\<ge>\<^sub>N x"
 proof-
   have " \<exists> a b. x = (a,b)" by simp
   then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -357,7 +357,7 @@
   ultimately show ?thesis by blast
 qed
 
-lemma Ngt0_iff[simp]:assumes nx: "isnormNum x" shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field})> 0) = 0<\<^sub>N x"
+lemma Ngt0_iff[simp]:assumes nx: "isnormNum x" shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field})> 0) = 0<\<^sub>N x"
 proof-
   have " \<exists> a b. x = (a,b)" by simp
   then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -369,7 +369,7 @@
   ultimately show ?thesis by blast
 qed
 lemma Nge0_iff[simp]:assumes nx: "isnormNum x" 
-  shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field}) \<ge> 0) = 0\<le>\<^sub>N x"
+  shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field}) \<ge> 0) = 0\<le>\<^sub>N x"
 proof-
   have " \<exists> a b. x = (a,b)" by simp
   then obtain a b where x[simp]:"x = (a,b)" by blast
@@ -382,7 +382,7 @@
 qed
 
 lemma Nlt_iff[simp]: assumes nx: "isnormNum x" and ny: "isnormNum y"
-  shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field}) < INum y) = (x <\<^sub>N y)"
+  shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field}) < INum y) = (x <\<^sub>N y)"
 proof-
   let ?z = "0::'a"
   have "((INum x ::'a) < INum y) = (INum (x -\<^sub>N y) < ?z)" using nx ny by simp
@@ -391,7 +391,7 @@
 qed
 
 lemma Nle_iff[simp]: assumes nx: "isnormNum x" and ny: "isnormNum y"
-  shows "((INum x :: 'a :: {ring_char_0,division_by_zero,linordered_field})\<le> INum y) = (x \<le>\<^sub>N y)"
+  shows "((INum x :: 'a :: {ring_char_0,division_ring_inverse_zero,linordered_field})\<le> INum y) = (x \<le>\<^sub>N y)"
 proof-
   have "((INum x ::'a) \<le> INum y) = (INum (x -\<^sub>N y) \<le> (0::'a))" using nx ny by simp
   also have "\<dots> = (0\<ge>\<^sub>N (x -\<^sub>N y))" using Nle0_iff[OF Nsub_normN[OF ny]] by simp
@@ -399,7 +399,7 @@
 qed
 
 lemma Nadd_commute:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "x +\<^sub>N y = y +\<^sub>N x"
 proof-
   have n: "isnormNum (x +\<^sub>N y)" "isnormNum (y +\<^sub>N x)" by simp_all
@@ -408,7 +408,7 @@
 qed
 
 lemma [simp]:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "(0, b) +\<^sub>N y = normNum y"
     and "(a, 0) +\<^sub>N y = normNum y" 
     and "x +\<^sub>N (0, b) = normNum x"
@@ -420,7 +420,7 @@
   done
 
 lemma normNum_nilpotent_aux[simp]:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   assumes nx: "isnormNum x" 
   shows "normNum x = x"
 proof-
@@ -432,7 +432,7 @@
 qed
 
 lemma normNum_nilpotent[simp]:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "normNum (normNum x) = normNum x"
   by simp
 
@@ -440,11 +440,11 @@
   by (simp_all add: normNum_def)
 
 lemma normNum_Nadd:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "normNum (x +\<^sub>N y) = x +\<^sub>N y" by simp
 
 lemma Nadd_normNum1[simp]:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "normNum x +\<^sub>N y = x +\<^sub>N y"
 proof-
   have n: "isnormNum (normNum x +\<^sub>N y)" "isnormNum (x +\<^sub>N y)" by simp_all
@@ -454,7 +454,7 @@
 qed
 
 lemma Nadd_normNum2[simp]:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "x +\<^sub>N normNum y = x +\<^sub>N y"
 proof-
   have n: "isnormNum (x +\<^sub>N normNum y)" "isnormNum (x +\<^sub>N y)" by simp_all
@@ -464,7 +464,7 @@
 qed
 
 lemma Nadd_assoc:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   shows "x +\<^sub>N y +\<^sub>N z = x +\<^sub>N (y +\<^sub>N z)"
 proof-
   have n: "isnormNum (x +\<^sub>N y +\<^sub>N z)" "isnormNum (x +\<^sub>N (y +\<^sub>N z))" by simp_all
@@ -476,7 +476,7 @@
   by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute)
 
 lemma Nmul_assoc:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   assumes nx: "isnormNum x" and ny:"isnormNum y" and nz:"isnormNum z"
   shows "x *\<^sub>N y *\<^sub>N z = x *\<^sub>N (y *\<^sub>N z)"
 proof-
@@ -487,7 +487,7 @@
 qed
 
 lemma Nsub0:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   assumes x: "isnormNum x" and y:"isnormNum y" shows "(x -\<^sub>N y = 0\<^sub>N) = (x = y)"
 proof-
   { fix h :: 'a
@@ -502,7 +502,7 @@
   by (simp_all add: Nmul_def Let_def split_def)
 
 lemma Nmul_eq0[simp]:
-  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+  assumes "SORT_CONSTRAINT('a::{ring_char_0,division_ring_inverse_zero,field})"
   assumes nx:"isnormNum x" and ny: "isnormNum y"
   shows "(x*\<^sub>N y = 0\<^sub>N) = (x = 0\<^sub>N \<or> y = 0\<^sub>N)"
 proof-
--- a/src/HOL/Library/Binomial.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Binomial.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -236,10 +236,10 @@
     have th1: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) =
       (\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
       apply (rule setprod_reindex_cong[where f = "Suc"])
-      using n0 by (auto simp add: expand_fun_eq ring_simps)
+      using n0 by (auto simp add: expand_fun_eq field_simps)
     have ?thesis apply (simp add: pochhammer_def)
     unfolding setprod_insert[OF th0, unfolded eq]
-    using th1 by (simp add: ring_simps)}
+    using th1 by (simp add: field_simps)}
 ultimately show ?thesis by blast
 qed
 
@@ -378,10 +378,10 @@
       by simp
     from n h th0 
     have "fact k * fact (n - k) * (n choose k) = k * (fact h * fact (m - h) * (m choose h)) +  (m - h) * (fact k * fact (m - k) * (m choose k))"
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     also have "\<dots> = (k + (m - h)) * fact m"
       using H[rule_format, OF mn hm'] H[rule_format, OF mn km]
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     finally have ?ths using h n km by simp}
   moreover have "n=0 \<or> k = 0 \<or> k = n \<or> (EX m h. n=Suc m \<and> k = Suc h \<and> h < m)" using kn by presburger
   ultimately show ?ths by blast
@@ -391,7 +391,7 @@
   assumes kn: "k \<le> n" 
   shows "(of_nat (n choose k) :: 'a::field_char_0) = of_nat (fact n) / (of_nat (fact k) * of_nat (fact (n - k)))"
   using binomial_fact_lemma[OF kn]
-  by (simp add: field_simps of_nat_mult[symmetric])
+  by (simp add: field_simps of_nat_mult [symmetric])
 
 lemma binomial_gbinomial: "of_nat (n choose k) = of_nat n gchoose k"
 proof-
@@ -414,7 +414,7 @@
       apply clarsimp
       apply (presburger)
       apply presburger
-      by (simp add: expand_fun_eq ring_simps of_nat_add[symmetric] del: of_nat_add)
+      by (simp add: expand_fun_eq field_simps of_nat_add[symmetric] del: of_nat_add)
     have th0: "finite {1..n - Suc h}" "finite {n - h .. n}" 
 "{1..n - Suc h} \<inter> {n - h .. n} = {}" and eq3: "{1..n - Suc h} \<union> {n - h .. n} = {1..n}" using h kn by auto
     from eq[symmetric]
@@ -451,7 +451,7 @@
     pochhammer_Suc fact_Suc of_nat_mult right_diff_distrib power_Suc
     by (simp add:  field_simps del: of_nat_Suc)
   also have "\<dots> = ?l" unfolding gbinomial_pochhammer
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   finally show ?thesis ..
 qed
 
@@ -482,17 +482,17 @@
 
     have "of_nat (fact (Suc k)) * (a gchoose k + (a gchoose (Suc k))) = ((a gchoose Suc h) * of_nat (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0\<Colon>nat..Suc h}. a - of_nat i)" 
       unfolding h
-      apply (simp add: ring_simps del: fact_Suc)
+      apply (simp add: field_simps del: fact_Suc)
       unfolding gbinomial_mult_fact'
       apply (subst fact_Suc)
       unfolding of_nat_mult 
       apply (subst mult_commute)
       unfolding mult_assoc
       unfolding gbinomial_mult_fact
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     also have "\<dots> = (\<Prod>i\<in>{0..h}. a - of_nat i) * (a + 1)"
       unfolding gbinomial_mult_fact' setprod_nat_ivl_Suc
-      by (simp add: ring_simps h)
+      by (simp add: field_simps h)
     also have "\<dots> = (\<Prod>i\<in>{0..k}. (a + 1) - of_nat i)"
       using eq0
       unfolding h  setprod_nat_ivl_1_Suc
--- a/src/HOL/Library/Bit.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Bit.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -49,7 +49,7 @@
 
 subsection {* Type @{typ bit} forms a field *}
 
-instantiation bit :: "{field, division_by_zero}"
+instantiation bit :: "{field, division_ring_inverse_zero}"
 begin
 
 definition plus_bit_def:
--- a/src/HOL/Library/Formal_Power_Series.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -388,6 +388,8 @@
   then show "a*b \<noteq> 0" unfolding fps_nonzero_nth by blast
 qed
 
+instance fps :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors ..
+
 instance fps :: (idom) idom ..
 
 instantiation fps :: (comm_ring_1) number_ring
@@ -586,7 +588,7 @@
   from k have "real k > - log y x" by simp
   then have "ln y * real k > - ln x" unfolding log_def
     using ln_gt_zero_iff[OF yp] y1
-    by (simp add: minus_divide_left field_simps del:minus_divide_left[symmetric] )
+    by (simp add: minus_divide_left field_simps del:minus_divide_left[symmetric])
   then have "ln y * real k + ln x > 0" by simp
   then have "exp (real k * ln y + ln x) > exp 0"
     by (simp add: mult_ac)
@@ -594,7 +596,7 @@
     unfolding exp_zero exp_add exp_real_of_nat_mult
     exp_ln[OF xp] exp_ln[OF yp] by simp
   then have "x > (1/y)^k" using yp 
-    by (simp add: field_simps nonzero_power_divide )
+    by (simp add: field_simps nonzero_power_divide)
   then show ?thesis using kp by blast
 qed
 lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)" by (simp add: X_def)
@@ -649,8 +651,7 @@
 
 declare setsum_cong[fundef_cong]
 
-
-instantiation fps :: ("{comm_monoid_add,inverse, times, uminus}") inverse
+instantiation fps :: ("{comm_monoid_add, inverse, times, uminus}") inverse
 begin
 
 fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a" where
@@ -658,9 +659,12 @@
 | "natfun_inverse f n = - inverse (f$0) * setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
 
 definition fps_inverse_def:
-  "inverse f = (if f$0 = 0 then 0 else Abs_fps (natfun_inverse f))"
+  "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]:
@@ -671,10 +675,7 @@
   apply (auto simp add: expand_fps_eq fps_inverse_def)
   by (case_tac n, auto)
 
-instance fps :: ("{comm_monoid_add,inverse, times, uminus}")  division_by_zero
-  by default (rule fps_inverse_zero)
-
-lemma inverse_mult_eq_1[intro]: assumes f0: "f$0 \<noteq> (0::'a::field)"
+lemma inverse_mult_eq_1 [intro]: assumes f0: "f$0 \<noteq> (0::'a::field)"
   shows "inverse f * f = 1"
 proof-
   have c: "inverse f * f = f * inverse f" by (simp add: mult_commute)
@@ -692,7 +693,7 @@
     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"
-      by (simp add: ring_simps)
+      by (simp add: field_simps)
     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
@@ -765,7 +766,7 @@
 lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n+1)" by (simp add: fps_deriv_def)
 
 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)
+  unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: field_simps)
 
 lemma fps_deriv_mult[simp]:
   fixes f :: "('a :: comm_ring_1) fps"
@@ -816,7 +817,7 @@
       unfolding s0 s1
       unfolding setsum_addf[symmetric] setsum_right_distrib
       apply (rule setsum_cong2)
-      by (auto simp add: of_nat_diff ring_simps)
+      by (auto simp add: of_nat_diff field_simps)
     finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .}
   then show ?thesis unfolding fps_eq_iff by auto
 qed
@@ -877,7 +878,7 @@
 proof-
   have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0" by simp
   also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f-g)$0)" unfolding fps_deriv_eq_0_iff ..
-  finally show ?thesis by (simp add: ring_simps)
+  finally show ?thesis by (simp add: field_simps)
 qed
 
 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)"
@@ -928,7 +929,7 @@
 qed
 
 lemma fps_deriv_maclauren_0: "(fps_nth_deriv k (f:: ('a::comm_semiring_1) fps)) $ 0 = of_nat (fact k) * f$(k)"
-  by (induct k arbitrary: f) (auto simp add: ring_simps of_nat_mult)
+  by (induct k arbitrary: f) (auto simp add: field_simps of_nat_mult)
 
 subsection {* Powers*}
 
@@ -942,7 +943,7 @@
   case (Suc n)
   note h = Suc.hyps[OF `a$0 = 1`]
   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)
+    using h `a$0 = 1`  fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: field_simps)
 qed
 
 lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
@@ -1004,7 +1005,7 @@
   case 0 thus ?case by (simp add: power_0)
 next
   case (Suc n)
-  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: ring_simps power_Suc)
+  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: field_simps power_Suc)
   also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}" by (simp add: fps_mult_nth)
   also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
     apply (rule setsum_mono_zero_right)
@@ -1044,8 +1045,8 @@
 qed
 
 lemma fps_deriv_power: "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a:: comm_ring_1) * fps_deriv a * a ^ (n - 1)"
-  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)
+  apply (induct n, auto simp add: power_Suc field_simps fps_const_add[symmetric] simp del: fps_const_add)
+  by (case_tac n, auto simp add: power_Suc field_simps)
 
 lemma fps_inverse_deriv:
   fixes a:: "('a :: field) fps"
@@ -1059,11 +1060,11 @@
   with inverse_mult_eq_1[OF a0]
   have "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) = 0"
     unfolding power2_eq_square
-    apply (simp add: ring_simps)
+    apply (simp add: field_simps)
     by (simp add: mult_assoc[symmetric])
   hence "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * inverse a ^ 2 = 0 - fps_deriv a * inverse a ^ 2"
     by simp
-  then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: ring_simps)
+  then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: field_simps)
 qed
 
 lemma fps_inverse_mult:
@@ -1083,7 +1084,7 @@
     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)
+      by (simp add: field_simps)
     then have ?thesis using inverse_mult_eq_1[OF a0] inverse_mult_eq_1[OF b0] by simp}
 ultimately show ?thesis by blast
 qed
@@ -1104,7 +1105,7 @@
   assumes a0: "b$0 \<noteq> 0"
   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])
+  by (simp add: fps_divide_def field_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
 
 
 lemma fps_inverse_gp': "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
@@ -1120,7 +1121,7 @@
 proof-
   have eq: "(1 + X) * ?r = 1"
     unfolding minus_one_power_iff
-    by (auto simp add: ring_simps fps_eq_iff)
+    by (auto simp add: field_simps fps_eq_iff)
   show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
 qed
 
@@ -1184,7 +1185,7 @@
     next
       case (Suc k)
       note th = Suc.hyps[symmetric]
-      have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps)
+      have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: field_simps)
       also  have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
         using th
         unfolding fps_sub_nth by simp
@@ -1208,10 +1209,10 @@
 definition "XD = op * X o fps_deriv"
 
 lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: ('a::comm_ring_1) fps)"
-  by (simp add: XD_def ring_simps)
+  by (simp add: XD_def field_simps)
 
 lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
-  by (simp add: XD_def ring_simps)
+  by (simp add: XD_def field_simps)
 
 lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) = fps_const c * XD a + fps_const d * XD (b :: ('a::comm_ring_1) fps)"
   by simp
@@ -1225,7 +1226,7 @@
 
 lemma fps_mult_XD_shift:
   "(XD ^^ k) (a:: ('a::{comm_ring_1}) fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
-  by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff ring_simps del: One_nat_def)
+  by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff field_simps del: One_nat_def)
 
 subsubsection{* Rule 3 is trivial and is given by @{text fps_times_def}*}
 subsubsection{* Rule 5 --- summation and "division" by (1 - X)*}
@@ -1687,7 +1688,7 @@
         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 )
+          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"
           unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
@@ -2017,11 +2018,11 @@
 proof-
   {fix n
     have "(fps_deriv (a oo b))$n = setsum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
-      by (simp add: fps_compose_def ring_simps setsum_right_distrib del: of_nat_Suc)
+      by (simp add: fps_compose_def field_simps setsum_right_distrib del: of_nat_Suc)
     also have "\<dots> = setsum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
-      by (simp add: ring_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
+      by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
   also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
-    unfolding fps_mult_left_const_nth  by (simp add: ring_simps)
+    unfolding fps_mult_left_const_nth  by (simp add: field_simps)
   also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
     unfolding fps_mult_nth ..
   also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
@@ -2169,7 +2170,7 @@
   by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum_0')
 
 lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
-  by (simp add: fps_eq_iff fps_compose_nth ring_simps setsum_addf)
+  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_addf)
 
 lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\<lambda>i. f i oo a) S"
 proof-
@@ -2211,7 +2212,7 @@
     apply (simp add: fps_mult_nth setsum_right_distrib)
     apply (subst setsum_commute)
     apply (rule setsum_cong2)
-    by (auto simp add: ring_simps)
+    by (auto simp add: field_simps)
   also have "\<dots> = ?l"
     apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
     apply (rule setsum_cong2)
@@ -2311,7 +2312,7 @@
 qed
 
 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])
+  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
 
 lemma fps_compose_sub_distrib:
   shows "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
@@ -2534,7 +2535,7 @@
   {fix n
     have "?l$n = ?r $ n"
   apply (auto simp add: E_def field_simps power_Suc[symmetric]simp del: fact_Suc of_nat_Suc power_Suc)
-  by (simp add: of_nat_mult ring_simps)}
+  by (simp add: of_nat_mult field_simps)}
 then show ?thesis by (simp add: fps_eq_iff)
 qed
 
@@ -2552,7 +2553,7 @@
       using fact_gt_zero_nat
       apply (simp add: field_simps del: of_nat_Suc fact_Suc)
       apply (drule sym)
-      by (simp add: ring_simps of_nat_mult power_Suc)}
+      by (simp add: field_simps of_nat_mult power_Suc)}
   note th' = this
   have ?rhs
     by (auto simp add: fps_eq_iff fps_const_mult_left E_def intro : th')}
@@ -2569,7 +2570,7 @@
 lemma E_add_mult: "E (a + b) = E (a::'a::field_char_0) * E b" (is "?l = ?r")
 proof-
   have "fps_deriv (?r) = fps_const (a+b) * ?r"
-    by (simp add: fps_const_add[symmetric] ring_simps del: fps_const_add)
+    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
   then have "?r = ?l" apply (simp only: E_unique_ODE)
     by (simp add: fps_mult_nth E_def)
   then show ?thesis ..
@@ -2617,13 +2618,13 @@
   (is "inverse ?l = ?r")
 proof-
   have th: "?l * ?r = 1"
-    by (auto simp add: ring_simps fps_eq_iff minus_one_power_iff)
+    by (auto simp add: field_simps fps_eq_iff minus_one_power_iff)
   have th': "?l $ 0 \<noteq> 0" by (simp add: )
   from fps_inverse_unique[OF th' th] show ?thesis .
 qed
 
 lemma E_power_mult: "(E (c::'a::field_char_0))^n = E (of_nat n * c)"
-  by (induct n, auto simp add: ring_simps E_add_mult power_Suc)
+  by (induct n, auto simp add: field_simps E_add_mult power_Suc)
 
 lemma radical_E:
   assumes r: "r (Suc k) 1 = 1" 
@@ -2648,7 +2649,7 @@
 text{* The generalized binomial theorem as a  consequence of @{thm E_add_mult} *}
 
 lemma gbinomial_theorem: 
-  "((a::'a::{field_char_0, division_by_zero})+b) ^ n = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
+  "((a::'a::{field_char_0, division_ring_inverse_zero})+b) ^ n = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))"
 proof-
   from E_add_mult[of a b] 
   have "(E (a + b)) $ n = (E a * E b)$n" by simp
@@ -2693,9 +2694,9 @@
   have b0: "?b $ 0 = 0" by simp
   have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
   have "fps_deriv (E a - 1) oo fps_inv (E a - 1) = (fps_const a * (E a - 1) + fps_const a) oo fps_inv (E a - 1)"
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   also have "\<dots> = fps_const a * (X + 1)" apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1])
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   finally have eq: "fps_deriv (E a - 1) oo fps_inv (E a - 1) = fps_const a * (X + 1)" .
   from fps_inv_deriv[OF b0 b1, unfolded eq]
   have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
@@ -2742,7 +2743,7 @@
   have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
   also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
     apply (simp only: fps_divide_def  mult_assoc[symmetric] inverse_mult_eq_1[OF x10])
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
   finally have eq: "?l = ?r \<longleftrightarrow> ?lhs" by simp
   moreover
   {assume h: "?l = ?r" 
@@ -2751,7 +2752,7 @@
       
       from lrn 
       have "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" 
-        apply (simp add: ring_simps del: of_nat_Suc)
+        apply (simp add: field_simps del: of_nat_Suc)
         by (cases n, simp_all add: field_simps del: of_nat_Suc)
     }
     note th0 = this
@@ -2763,22 +2764,22 @@
         thus ?case unfolding th0
           apply (simp add: field_simps del: of_nat_Suc)
           unfolding mult_assoc[symmetric] gbinomial_mult_1
-          by (simp add: ring_simps)
+          by (simp add: field_simps)
       qed}
     note th1 = this
     have ?rhs
       apply (simp add: fps_eq_iff)
       apply (subst th1)
-      by (simp add: ring_simps)}
+      by (simp add: field_simps)}
   moreover
   {assume h: ?rhs
   have th00:"\<And>x y. x * (a$0 * y) = a$0 * (x*y)" by (simp add: mult_commute)
     have "?l = ?r" 
       apply (subst h)
       apply (subst (2) h)
-      apply (clarsimp simp add: fps_eq_iff ring_simps)
+      apply (clarsimp simp add: fps_eq_iff field_simps)
       unfolding mult_assoc[symmetric] th00 gbinomial_mult_1
-      by (simp add: ring_simps gbinomial_mult_1)}
+      by (simp add: field_simps gbinomial_mult_1)}
   ultimately show ?thesis by blast
 qed
 
@@ -2797,9 +2798,9 @@
   have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
   also have "\<dots> = inverse (1 + X) * (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
     unfolding fps_binomial_deriv
-    by (simp add: fps_divide_def ring_simps)
+    by (simp add: fps_divide_def field_simps)
   also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
-    by (simp add: ring_simps fps_divide_def fps_const_add[symmetric] del: fps_const_add)
+    by (simp add: field_simps fps_divide_def fps_const_add[symmetric] del: fps_const_add)
   finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + X)"
     by (simp add: fps_divide_def)
   have "?P = fps_const (?P$0) * ?b (c + d)"
@@ -3017,7 +3018,7 @@
       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
-        by (simp add: fps_cos_def ring_simps power_Suc )}
+        by (simp add: fps_cos_def field_simps power_Suc )}
     then show "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
 qed
@@ -3043,7 +3044,7 @@
       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
-        by (simp add: fps_sin_def ring_simps power_Suc)}
+        by (simp add: fps_sin_def field_simps power_Suc)}
     then show "?lhs $ n = ?rhs $ n"
       by (cases "even n", simp_all add: fps_deriv_def fps_sin_def
         fps_cos_def)
@@ -3054,7 +3055,7 @@
 proof-
   have "fps_deriv ?lhs = 0"
     apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv power_Suc)
-    by (simp add: ring_simps fps_const_neg[symmetric] del: fps_const_neg)
+    by (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
   then have "?lhs = fps_const (?lhs $ 0)"
     unfolding fps_deriv_eq_0_iff .
   also have "\<dots> = 1"
@@ -3176,7 +3177,7 @@
   have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
   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)
+    apply (simp add: fps_tan_def fps_divide_deriv[OF th0] fps_sin_deriv fps_cos_deriv add: fps_const_neg[symmetric] field_simps power2_eq_square del: fps_const_neg)
     unfolding right_distrib[symmetric]
     by simp
 qed
@@ -3251,7 +3252,7 @@
 subsection {* Hypergeometric series *}
 
 
-definition "F as bs (c::'a::{field_char_0, division_by_zero}) = Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
+definition "F as bs (c::'a::{field_char_0, division_ring_inverse_zero}) = Abs_fps (%n. (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
 
 lemma F_nth[simp]: "F as bs c $ n =  (foldl (%r a. r* pochhammer a n) 1 as * c^n)/ (foldl (%r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
   by (simp add: F_def)
@@ -3320,9 +3321,9 @@
   by (simp add: fps_eq_iff fps_integral_def)
 
 lemma F_minus_nat: 
-  "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, division_by_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k /
+  "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0, division_ring_inverse_zero}) $ k = (if k <= n then pochhammer (- of_nat n) k * c ^ k /
     (pochhammer (- of_nat (n + m)) k * of_nat (fact k)) else 0)"
-  "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, division_by_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k /
+  "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0, division_ring_inverse_zero}) $ k = (if k <= m then pochhammer (- of_nat m) k * c ^ k /
     (pochhammer (- of_nat (m + n)) k * of_nat (fact k)) else 0)"
   by (auto simp add: pochhammer_eq_0_iff)
 
@@ -3345,6 +3346,6 @@
 
   shows "foldr (%c r. f c o r) cs (%c. g c a) c0 $ n = 
   foldr (%c r. (k c + of_nat n) * r) cs (g c0 a $ n)"
-  by (induct cs arbitrary: c0, auto simp add: algebra_simps f )
+  by (induct cs arbitrary: c0, auto simp add: algebra_simps f)
 
 end
--- a/src/HOL/Library/Fraction_Field.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Fraction_Field.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -232,7 +232,7 @@
 thm mult_eq_0_iff
 end
 
-instantiation fract :: (idom) "{field, division_by_zero}"
+instantiation fract :: (idom) field
 begin
 
 definition
@@ -256,9 +256,6 @@
   by (simp add: divide_fract_def)
 
 instance proof
-  show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
-    (simp add: fract_collapse)
-next
   fix q :: "'a fract"
   assume "q \<noteq> 0"
   then show "inverse q * q = 1" apply (cases q rule: Fract_cases_nonzero)
@@ -270,5 +267,262 @@
 
 end
 
+instance fract :: (idom) division_ring_inverse_zero
+proof
+  show "inverse 0 = (0:: 'a fract)" by (simp add: fract_expand)
+    (simp add: fract_collapse)
+qed
 
-end
\ No newline at end of file
+
+subsubsection {* The ordered field of fractions over an ordered idom *}
+
+lemma le_congruent2:
+  "(\<lambda>x y::'a \<times> 'a::linordered_idom.
+    {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})
+    respects2 fractrel"
+proof (clarsimp simp add: congruent2_def)
+  fix a b a' b' c d c' d' :: 'a
+  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+  assume eq1: "a * b' = a' * b"
+  assume eq2: "c * d' = c' * d"
+
+  let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
+  {
+    fix a b c d x :: 'a assume x: "x \<noteq> 0"
+    have "?le a b c d = ?le (a * x) (b * x) c d"
+    proof -
+      from x have "0 < x * x" by (auto simp add: zero_less_mult_iff)
+      hence "?le a b c d =
+          ((a * d) * (b * d) * (x * x) \<le> (c * b) * (b * d) * (x * x))"
+        by (simp add: mult_le_cancel_right)
+      also have "... = ?le (a * x) (b * x) c d"
+        by (simp add: mult_ac)
+      finally show ?thesis .
+    qed
+  } note le_factor = this
+
+  let ?D = "b * d" and ?D' = "b' * d'"
+  from neq have D: "?D \<noteq> 0" by simp
+  from neq have "?D' \<noteq> 0" by simp
+  hence "?le a b c d = ?le (a * ?D') (b * ?D') c d"
+    by (rule le_factor)
+  also have "... = ((a * b') * ?D * ?D' * d * d' \<le> (c * d') * ?D * ?D' * b * b')"
+    by (simp add: mult_ac)
+  also have "... = ((a' * b) * ?D * ?D' * d * d' \<le> (c' * d) * ?D * ?D' * b * b')"
+    by (simp only: eq1 eq2)
+  also have "... = ?le (a' * ?D) (b' * ?D) c' d'"
+    by (simp add: mult_ac)
+  also from D have "... = ?le a' b' c' d'"
+    by (rule le_factor [symmetric])
+  finally show "?le a b c d = ?le a' b' c' d'" .
+qed
+
+instantiation fract :: (linordered_idom) linorder
+begin
+
+definition
+  le_fract_def [code del]:
+   "q \<le> r \<longleftrightarrow> contents (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
+      {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})"
+
+definition
+  less_fract_def [code del]: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
+
+lemma le_fract [simp]:
+  assumes "b \<noteq> 0" and "d \<noteq> 0"
+  shows "Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
+by (simp add: Fract_def le_fract_def le_congruent2 UN_fractrel2 assms)
+
+lemma less_fract [simp]:
+  assumes "b \<noteq> 0" and "d \<noteq> 0"
+  shows "Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
+by (simp add: less_fract_def less_le_not_le mult_ac assms)
+
+instance proof
+  fix q r s :: "'a fract"
+  assume "q \<le> r" and "r \<le> s" thus "q \<le> s"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: 'a
+    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
+    assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract e f"
+    show "Fract a b \<le> Fract e f"
+    proof -
+      from neq obtain bb: "0 < b * b" and dd: "0 < d * d" and ff: "0 < f * f"
+        by (auto simp add: zero_less_mult_iff linorder_neq_iff)
+      have "(a * d) * (b * d) * (f * f) \<le> (c * b) * (b * d) * (f * f)"
+      proof -
+        from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+          by simp
+        with ff show ?thesis by (simp add: mult_le_cancel_right)
+      qed
+      also have "... = (c * f) * (d * f) * (b * b)"
+        by (simp only: mult_ac)
+      also have "... \<le> (e * d) * (d * f) * (b * b)"
+      proof -
+        from neq 2 have "(c * f) * (d * f) \<le> (e * d) * (d * f)"
+          by simp
+        with bb show ?thesis by (simp add: mult_le_cancel_right)
+      qed
+      finally have "(a * f) * (b * f) * (d * d) \<le> e * b * (b * f) * (d * d)"
+        by (simp only: mult_ac)
+      with dd have "(a * f) * (b * f) \<le> (e * b) * (b * f)"
+        by (simp add: mult_le_cancel_right)
+      with neq show ?thesis by simp
+    qed
+  qed
+next
+  fix q r :: "'a fract"
+  assume "q \<le> r" and "r \<le> q" thus "q = r"
+  proof (induct q, induct r)
+    fix a b c d :: 'a
+    assume neq: "b \<noteq> 0"  "d \<noteq> 0"
+    assume 1: "Fract a b \<le> Fract c d" and 2: "Fract c d \<le> Fract a b"
+    show "Fract a b = Fract c d"
+    proof -
+      from neq 1 have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+        by simp
+      also have "... \<le> (a * d) * (b * d)"
+      proof -
+        from neq 2 have "(c * b) * (d * b) \<le> (a * d) * (d * b)"
+          by simp
+        thus ?thesis by (simp only: mult_ac)
+      qed
+      finally have "(a * d) * (b * d) = (c * b) * (b * d)" .
+      moreover from neq have "b * d \<noteq> 0" by simp
+      ultimately have "a * d = c * b" by simp
+      with neq show ?thesis by (simp add: eq_fract)
+    qed
+  qed
+next
+  fix q r :: "'a fract"
+  show "q \<le> q"
+    by (induct q) simp
+  show "(q < r) = (q \<le> r \<and> \<not> r \<le> q)"
+    by (simp only: less_fract_def)
+  show "q \<le> r \<or> r \<le> q"
+    by (induct q, induct r)
+       (simp add: mult_commute, rule linorder_linear)
+qed
+
+end
+
+instantiation fract :: (linordered_idom) "{distrib_lattice, abs_if, sgn_if}"
+begin
+
+definition
+  abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
+
+definition
+  sgn_fract_def:
+    "sgn (q::'a fract) = (if q=0 then 0 else if 0<q then 1 else - 1)"
+
+theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
+  by (auto simp add: abs_fract_def Zero_fract_def le_less
+      eq_fract zero_less_mult_iff mult_less_0_iff split: abs_split)
+
+definition
+  inf_fract_def:
+    "(inf \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
+
+definition
+  sup_fract_def:
+    "(sup \<Colon> 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
+
+instance by intro_classes
+  (auto simp add: abs_fract_def sgn_fract_def
+    min_max.sup_inf_distrib1 inf_fract_def sup_fract_def)
+
+end
+
+instance fract :: (linordered_idom) linordered_field
+proof
+  fix q r s :: "'a fract"
+  show "q \<le> r ==> s + q \<le> s + r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: 'a
+    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
+    assume le: "Fract a b \<le> Fract c d"
+    show "Fract e f + Fract a b \<le> Fract e f + Fract c d"
+    proof -
+      let ?F = "f * f" from neq have F: "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      from neq le have "(a * d) * (b * d) \<le> (c * b) * (b * d)"
+        by simp
+      with F have "(a * d) * (b * d) * ?F * ?F \<le> (c * b) * (b * d) * ?F * ?F"
+        by (simp add: mult_le_cancel_right)
+      with neq show ?thesis by (simp add: field_simps)
+    qed
+  qed
+  show "q < r ==> 0 < s ==> s * q < s * r"
+  proof (induct q, induct r, induct s)
+    fix a b c d e f :: 'a
+    assume neq: "b \<noteq> 0"  "d \<noteq> 0"  "f \<noteq> 0"
+    assume le: "Fract a b < Fract c d"
+    assume gt: "0 < Fract e f"
+    show "Fract e f * Fract a b < Fract e f * Fract c d"
+    proof -
+      let ?E = "e * f" and ?F = "f * f"
+      from neq gt have "0 < ?E"
+        by (auto simp add: Zero_fract_def order_less_le eq_fract)
+      moreover from neq have "0 < ?F"
+        by (auto simp add: zero_less_mult_iff)
+      moreover from neq le have "(a * d) * (b * d) < (c * b) * (b * d)"
+        by simp
+      ultimately have "(a * d) * (b * d) * ?E * ?F < (c * b) * (b * d) * ?E * ?F"
+        by (simp add: mult_less_cancel_right)
+      with neq show ?thesis
+        by (simp add: mult_ac)
+    qed
+  qed
+qed
+
+lemma fract_induct_pos [case_names Fract]:
+  fixes P :: "'a::linordered_idom fract \<Rightarrow> bool"
+  assumes step: "\<And>a b. 0 < b \<Longrightarrow> P (Fract a b)"
+  shows "P q"
+proof (cases q)
+  have step': "\<And>a b. b < 0 \<Longrightarrow> P (Fract a b)"
+  proof -
+    fix a::'a and b::'a
+    assume b: "b < 0"
+    hence "0 < -b" by simp
+    hence "P (Fract (-a) (-b))" by (rule step)
+    thus "P (Fract a b)" by (simp add: order_less_imp_not_eq [OF b])
+  qed
+  case (Fract a b)
+  thus "P q" by (force simp add: linorder_neq_iff step step')
+qed
+
+lemma zero_less_Fract_iff:
+  "0 < b \<Longrightarrow> 0 < Fract a b \<longleftrightarrow> 0 < a"
+  by (auto simp add: Zero_fract_def zero_less_mult_iff)
+
+lemma Fract_less_zero_iff:
+  "0 < b \<Longrightarrow> Fract a b < 0 \<longleftrightarrow> a < 0"
+  by (auto simp add: Zero_fract_def mult_less_0_iff)
+
+lemma zero_le_Fract_iff:
+  "0 < b \<Longrightarrow> 0 \<le> Fract a b \<longleftrightarrow> 0 \<le> a"
+  by (auto simp add: Zero_fract_def zero_le_mult_iff)
+
+lemma Fract_le_zero_iff:
+  "0 < b \<Longrightarrow> Fract a b \<le> 0 \<longleftrightarrow> a \<le> 0"
+  by (auto simp add: Zero_fract_def mult_le_0_iff)
+
+lemma one_less_Fract_iff:
+  "0 < b \<Longrightarrow> 1 < Fract a b \<longleftrightarrow> b < a"
+  by (auto simp add: One_fract_def mult_less_cancel_right_disj)
+
+lemma Fract_less_one_iff:
+  "0 < b \<Longrightarrow> Fract a b < 1 \<longleftrightarrow> a < b"
+  by (auto simp add: One_fract_def mult_less_cancel_right_disj)
+
+lemma one_le_Fract_iff:
+  "0 < b \<Longrightarrow> 1 \<le> Fract a b \<longleftrightarrow> b \<le> a"
+  by (auto simp add: One_fract_def mult_le_cancel_right)
+
+lemma Fract_le_one_iff:
+  "0 < b \<Longrightarrow> Fract a b \<le> 1 \<longleftrightarrow> a \<le> b"
+  by (auto simp add: One_fract_def mult_le_cancel_right)
+
+end
--- a/src/HOL/Library/FrechetDeriv.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/FrechetDeriv.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -385,7 +385,7 @@
   fixes x :: "'a::{real_normed_algebra,comm_ring_1}"
   shows "FDERIV (\<lambda>x. x ^ Suc n) x :> (\<lambda>h. (1 + of_nat n) * x ^ n * h)"
  apply (induct n)
-  apply (simp add: power_Suc FDERIV_ident)
+  apply (simp add: FDERIV_ident)
  apply (drule FDERIV_mult [OF FDERIV_ident])
  apply (simp only: of_nat_Suc left_distrib mult_1_left)
  apply (simp only: power_Suc right_distrib add_ac mult_ac)
--- a/src/HOL/Library/Numeral_Type.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Numeral_Type.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -213,7 +213,7 @@
 
 lemma comm_ring_1: "OFCLASS('a, comm_ring_1_class)"
 apply (intro_classes, unfold definitions)
-apply (simp_all add: Rep_simps zmod_simps ring_simps)
+apply (simp_all add: Rep_simps zmod_simps field_simps)
 done
 
 end
--- a/src/HOL/Library/Permutations.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Permutations.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -5,7 +5,7 @@
 header {* Permutations, both general and specifically on finite sets.*}
 
 theory Permutations
-imports Finite_Cartesian_Product Parity Fact
+imports Parity Fact
 begin
 
 definition permutes (infixr "permutes" 41) where
@@ -96,7 +96,7 @@
 
 lemma permutes_superset:
   "p permutes S \<Longrightarrow> (\<forall>x \<in> S - T. p x = x) \<Longrightarrow> p permutes T"
-by (simp add: Ball_def permutes_def Diff_iff) metis
+by (simp add: Ball_def permutes_def) metis
 
 (* ------------------------------------------------------------------------- *)
 (* Group properties.                                                         *)
@@ -125,7 +125,7 @@
   apply (rule permutes_compose[OF pS])
   apply (rule permutes_swap_id, simp)
   using permutes_in_image[OF pS, of a] apply simp
-  apply (auto simp add: Ball_def Diff_iff swap_def)
+  apply (auto simp add: Ball_def swap_def)
   done
 
 lemma permutes_insert: "{p. p permutes (insert a S)} =
@@ -154,7 +154,7 @@
 lemma card_permutations: assumes Sn: "card S = n" and fS: "finite S"
   shows "card {p. p permutes S} = fact n"
 using fS Sn proof (induct arbitrary: n)
-  case empty thus ?case by (simp add: permutes_empty)
+  case empty thus ?case by simp
 next
   case (insert x F)
   { fix n assume H0: "card (insert x F) = n"
--- a/src/HOL/Library/Polynomial.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Polynomial.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -1093,10 +1093,10 @@
 apply (cases "r = 0")
 apply (cases "r' = 0")
 apply (simp add: pdivmod_rel_def)
-apply (simp add: pdivmod_rel_def ring_simps degree_mult_eq)
+apply (simp add: pdivmod_rel_def field_simps degree_mult_eq)
 apply (cases "r' = 0")
 apply (simp add: pdivmod_rel_def degree_mult_eq)
-apply (simp add: pdivmod_rel_def ring_simps)
+apply (simp add: pdivmod_rel_def field_simps)
 apply (simp add: degree_mult_eq degree_add_less)
 done
 
--- a/src/HOL/Library/Product_Vector.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Product_Vector.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -49,21 +49,37 @@
   "open (S :: ('a \<times> 'b) set) \<longleftrightarrow>
     (\<forall>x\<in>S. \<exists>A B. open A \<and> open B \<and> x \<in> A \<times> B \<and> A \<times> B \<subseteq> S)"
 
+lemma open_prod_elim:
+  assumes "open S" and "x \<in> S"
+  obtains A B where "open A" and "open B" and "x \<in> A \<times> B" and "A \<times> B \<subseteq> S"
+using assms unfolding open_prod_def by fast
+
+lemma open_prod_intro:
+  assumes "\<And>x. x \<in> S \<Longrightarrow> \<exists>A B. open A \<and> open B \<and> x \<in> A \<times> B \<and> A \<times> B \<subseteq> S"
+  shows "open S"
+using assms unfolding open_prod_def by fast
+
 instance proof
   show "open (UNIV :: ('a \<times> 'b) set)"
     unfolding open_prod_def by auto
 next
   fix S T :: "('a \<times> 'b) set"
-  assume "open S" "open T" thus "open (S \<inter> T)"
-    unfolding open_prod_def
-    apply clarify
-    apply (drule (1) bspec)+
-    apply (clarify, rename_tac Sa Ta Sb Tb)
-    apply (rule_tac x="Sa \<inter> Ta" in exI)
-    apply (rule_tac x="Sb \<inter> Tb" in exI)
-    apply (simp add: open_Int)
-    apply fast
-    done
+  assume "open S" "open T"
+  show "open (S \<inter> T)"
+  proof (rule open_prod_intro)
+    fix x assume x: "x \<in> S \<inter> T"
+    from x have "x \<in> S" by simp
+    obtain Sa Sb where A: "open Sa" "open Sb" "x \<in> Sa \<times> Sb" "Sa \<times> Sb \<subseteq> S"
+      using `open S` and `x \<in> S` by (rule open_prod_elim)
+    from x have "x \<in> T" by simp
+    obtain Ta Tb where B: "open Ta" "open Tb" "x \<in> Ta \<times> Tb" "Ta \<times> Tb \<subseteq> T"
+      using `open T` and `x \<in> T` by (rule open_prod_elim)
+    let ?A = "Sa \<inter> Ta" and ?B = "Sb \<inter> Tb"
+    have "open ?A \<and> open ?B \<and> x \<in> ?A \<times> ?B \<and> ?A \<times> ?B \<subseteq> S \<inter> T"
+      using A B by (auto simp add: open_Int)
+    thus "\<exists>A B. open A \<and> open B \<and> x \<in> A \<times> B \<and> A \<times> B \<subseteq> S \<inter> T"
+      by fast
+  qed
 next
   fix K :: "('a \<times> 'b) set set"
   assume "\<forall>S\<in>K. open S" thus "open (\<Union>K)"
@@ -151,6 +167,12 @@
 lemma dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\<twosuperior> + (dist b d)\<twosuperior>)"
   unfolding dist_prod_def by simp
 
+lemma dist_fst_le: "dist (fst x) (fst y) \<le> dist x y"
+unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge1)
+
+lemma dist_snd_le: "dist (snd x) (snd y) \<le> dist x y"
+unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge2)
+
 instance proof
   fix x y :: "'a \<times> 'b"
   show "dist x y = 0 \<longleftrightarrow> x = y"
@@ -168,23 +190,32 @@
   fix S :: "('a \<times> 'b) set"
   show "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
   proof
-    assume "open S" thus "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S"
-    unfolding open_prod_def open_dist
-    apply safe
-    apply (drule (1) bspec)
-    apply clarify
-    apply (drule (1) bspec)+
-    apply (clarify, rename_tac r s)
-    apply (rule_tac x="min r s" in exI, simp)
-    apply (clarify, rename_tac c d)
-    apply (erule subsetD)
-    apply (simp add: dist_Pair_Pair)
-    apply (rule conjI)
-    apply (drule spec, erule mp)
-    apply (erule le_less_trans [OF real_sqrt_sum_squares_ge1])
-    apply (drule spec, erule mp)
-    apply (erule le_less_trans [OF real_sqrt_sum_squares_ge2])
-    done
+    assume "open S" show "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S"
+    proof
+      fix x assume "x \<in> S"
+      obtain A B where "open A" "open B" "x \<in> A \<times> B" "A \<times> B \<subseteq> S"
+        using `open S` and `x \<in> S` by (rule open_prod_elim)
+      obtain r where r: "0 < r" "\<forall>y. dist y (fst x) < r \<longrightarrow> y \<in> A"
+        using `open A` and `x \<in> A \<times> B` unfolding open_dist by auto
+      obtain s where s: "0 < s" "\<forall>y. dist y (snd x) < s \<longrightarrow> y \<in> B"
+        using `open B` and `x \<in> A \<times> B` unfolding open_dist by auto
+      let ?e = "min r s"
+      have "0 < ?e \<and> (\<forall>y. dist y x < ?e \<longrightarrow> y \<in> S)"
+      proof (intro allI impI conjI)
+        show "0 < min r s" by (simp add: r(1) s(1))
+      next
+        fix y assume "dist y x < min r s"
+        hence "dist y x < r" and "dist y x < s"
+          by simp_all
+        hence "dist (fst y) (fst x) < r" and "dist (snd y) (snd x) < s"
+          by (auto intro: le_less_trans dist_fst_le dist_snd_le)
+        hence "fst y \<in> A" and "snd y \<in> B"
+          by (simp_all add: r(2) s(2))
+        hence "y \<in> A \<times> B" by (induct y, simp)
+        with `A \<times> B \<subseteq> S` show "y \<in> S" ..
+      qed
+      thus "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" ..
+    qed
   next
     assume "\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S" thus "open S"
     unfolding open_prod_def open_dist
@@ -223,12 +254,6 @@
 
 subsection {* Continuity of operations *}
 
-lemma dist_fst_le: "dist (fst x) (fst y) \<le> dist x y"
-unfolding dist_prod_def by simp
-
-lemma dist_snd_le: "dist (snd x) (snd y) \<le> dist x y"
-unfolding dist_prod_def by simp
-
 lemma tendsto_fst [tendsto_intros]:
   assumes "(f ---> a) net"
   shows "((\<lambda>x. fst (f x)) ---> fst a) net"
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -1282,9 +1282,9 @@
   fun simple_cterm_ord t u = Term_Ord.fast_term_ord (term_of t, term_of u) = LESS
  val concl = Thm.dest_arg o cprop_of
  val shuffle1 =
-   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps) })
+   fconv_rule (rewr_conv @{lemma "(a + x == y) == (x == y - (a::real))" by (atomize (full)) (simp add: field_simps) })
  val shuffle2 =
-    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: ring_simps)})
+    fconv_rule (rewr_conv @{lemma "(x + a == y) ==  (x == y - (a::real))" by (atomize (full)) (simp add: field_simps)})
  fun substitutable_monomial fvs tm = case term_of tm of
     Free(_,@{typ real}) => if not (member (op aconvc) fvs tm) then (Rat.one,tm)
                            else raise Failure "substitutable_monomial"
--- a/src/HOL/Limits.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Limits.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -11,62 +11,59 @@
 subsection {* Nets *}
 
 text {*
-  A net is now defined as a filter base.
-  The definition also allows non-proper filter bases.
+  A net is now defined simply as a filter.
+  The definition also allows non-proper filters.
 *}
 
+locale is_filter =
+  fixes net :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
+  assumes True: "net (\<lambda>x. True)"
+  assumes conj: "net (\<lambda>x. P x) \<Longrightarrow> net (\<lambda>x. Q x) \<Longrightarrow> net (\<lambda>x. P x \<and> Q x)"
+  assumes mono: "\<forall>x. P x \<longrightarrow> Q x \<Longrightarrow> net (\<lambda>x. P x) \<Longrightarrow> net (\<lambda>x. Q x)"
+
 typedef (open) 'a net =
-  "{net :: 'a set set. (\<exists>A. A \<in> net)
-    \<and> (\<forall>A\<in>net. \<forall>B\<in>net. \<exists>C\<in>net. C \<subseteq> A \<and> C \<subseteq> B)}"
+  "{net :: ('a \<Rightarrow> bool) \<Rightarrow> bool. is_filter net}"
 proof
-  show "UNIV \<in> ?net" by auto
+  show "(\<lambda>x. True) \<in> ?net" by (auto intro: is_filter.intro)
 qed
 
-lemma Rep_net_nonempty: "\<exists>A. A \<in> Rep_net net"
-using Rep_net [of net] by simp
-
-lemma Rep_net_directed:
-  "A \<in> Rep_net net \<Longrightarrow> B \<in> Rep_net net \<Longrightarrow> \<exists>C\<in>Rep_net net. C \<subseteq> A \<and> C \<subseteq> B"
+lemma is_filter_Rep_net: "is_filter (Rep_net net)"
 using Rep_net [of net] by simp
 
 lemma Abs_net_inverse':
-  assumes "\<exists>A. A \<in> net"
-  assumes "\<And>A B. A \<in> net \<Longrightarrow> B \<in> net \<Longrightarrow> \<exists>C\<in>net. C \<subseteq> A \<and> C \<subseteq> B" 
-  shows "Rep_net (Abs_net net) = net"
+  assumes "is_filter net" shows "Rep_net (Abs_net net) = net"
 using assms by (simp add: Abs_net_inverse)
 
-lemma image_nonempty: "\<exists>x. x \<in> A \<Longrightarrow> \<exists>x. x \<in> f ` A"
-by auto
-
 
 subsection {* Eventually *}
 
 definition
   eventually :: "('a \<Rightarrow> bool) \<Rightarrow> 'a net \<Rightarrow> bool" where
-  [code del]: "eventually P net \<longleftrightarrow> (\<exists>A\<in>Rep_net net. \<forall>x\<in>A. P x)"
+  [code del]: "eventually P net \<longleftrightarrow> Rep_net net P"
+
+lemma eventually_Abs_net:
+  assumes "is_filter net" shows "eventually P (Abs_net net) = net P"
+unfolding eventually_def using assms by (simp add: Abs_net_inverse)
+
+lemma expand_net_eq:
+  shows "net = net' \<longleftrightarrow> (\<forall>P. eventually P net = eventually P net')"
+unfolding Rep_net_inject [symmetric] expand_fun_eq eventually_def ..
 
 lemma eventually_True [simp]: "eventually (\<lambda>x. True) net"
-unfolding eventually_def using Rep_net_nonempty [of net] by fast
+unfolding eventually_def
+by (rule is_filter.True [OF is_filter_Rep_net])
 
 lemma eventually_mono:
   "(\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually P net \<Longrightarrow> eventually Q net"
-unfolding eventually_def by blast
+unfolding eventually_def
+by (rule is_filter.mono [OF is_filter_Rep_net])
 
 lemma eventually_conj:
   assumes P: "eventually (\<lambda>x. P x) net"
   assumes Q: "eventually (\<lambda>x. Q x) net"
   shows "eventually (\<lambda>x. P x \<and> Q x) net"
-proof -
-  obtain A where A: "A \<in> Rep_net net" "\<forall>x\<in>A. P x"
-    using P unfolding eventually_def by fast
-  obtain B where B: "B \<in> Rep_net net" "\<forall>x\<in>B. Q x"
-    using Q unfolding eventually_def by fast
-  obtain C where C: "C \<in> Rep_net net" "C \<subseteq> A" "C \<subseteq> B"
-    using Rep_net_directed [OF A(1) B(1)] by fast
-  then have "\<forall>x\<in>C. P x \<and> Q x" "C \<in> Rep_net net"
-    using A(2) B(2) by auto
-  then show ?thesis unfolding eventually_def ..
-qed
+using assms unfolding eventually_def
+by (rule is_filter.conj [OF is_filter_Rep_net])
 
 lemma eventually_mp:
   assumes "eventually (\<lambda>x. P x \<longrightarrow> Q x) net"
@@ -102,60 +99,116 @@
 using assms by (auto elim!: eventually_rev_mp)
 
 
+subsection {* Finer-than relation *}
+
+text {* @{term "net \<le> net'"} means that @{term net'} is finer than
+@{term net}. *}
+
+instantiation net :: (type) "{order,top}"
+begin
+
+definition
+  le_net_def [code del]:
+    "net \<le> net' \<longleftrightarrow> (\<forall>P. eventually P net \<longrightarrow> eventually P net')"
+
+definition
+  less_net_def [code del]:
+    "(net :: 'a net) < net' \<longleftrightarrow> net \<le> net' \<and> \<not> net' \<le> net"
+
+definition
+  top_net_def [code del]:
+    "top = Abs_net (\<lambda>P. True)"
+
+lemma eventually_top [simp]: "eventually P top"
+unfolding top_net_def
+by (subst eventually_Abs_net, rule is_filter.intro, auto)
+
+instance proof
+  fix x y :: "'a net" show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
+    by (rule less_net_def)
+next
+  fix x :: "'a net" show "x \<le> x"
+    unfolding le_net_def by simp
+next
+  fix x y z :: "'a net" assume "x \<le> y" and "y \<le> z" thus "x \<le> z"
+    unfolding le_net_def by simp
+next
+  fix x y :: "'a net" assume "x \<le> y" and "y \<le> x" thus "x = y"
+    unfolding le_net_def expand_net_eq by fast
+next
+  fix x :: "'a net" show "x \<le> top"
+    unfolding le_net_def by simp
+qed
+
+end
+
+lemma net_leD:
+  "net \<le> net' \<Longrightarrow> eventually P net \<Longrightarrow> eventually P net'"
+unfolding le_net_def by simp
+
+lemma net_leI:
+  "(\<And>P. eventually P net \<Longrightarrow> eventually P net') \<Longrightarrow> net \<le> net'"
+unfolding le_net_def by simp
+
+lemma eventually_False:
+  "eventually (\<lambda>x. False) net \<longleftrightarrow> net = top"
+unfolding expand_net_eq by (auto elim: eventually_rev_mp)
+
+
 subsection {* Standard Nets *}
 
 definition
-  sequentially :: "nat net" where
-  [code del]: "sequentially = Abs_net (range (\<lambda>n. {n..}))"
-
-definition
-  within :: "'a net \<Rightarrow> 'a set \<Rightarrow> 'a net" (infixr "within" 70) where
-  [code del]: "net within S = Abs_net ((\<lambda>A. A \<inter> S) ` Rep_net net)"
+  sequentially :: "nat net"
+where [code del]:
+  "sequentially = Abs_net (\<lambda>P. \<exists>k. \<forall>n\<ge>k. P n)"
 
 definition
-  at :: "'a::topological_space \<Rightarrow> 'a net" where
-  [code del]: "at a = Abs_net ((\<lambda>S. S - {a}) ` {S. open S \<and> a \<in> S})"
-
-lemma Rep_net_sequentially:
-  "Rep_net sequentially = range (\<lambda>n. {n..})"
-unfolding sequentially_def
-apply (rule Abs_net_inverse')
-apply (rule image_nonempty, simp)
-apply (clarsimp, rename_tac m n)
-apply (rule_tac x="max m n" in exI, auto)
-done
+  within :: "'a net \<Rightarrow> 'a set \<Rightarrow> 'a net" (infixr "within" 70)
+where [code del]:
+  "net within S = Abs_net (\<lambda>P. eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) net)"
 
-lemma Rep_net_within:
-  "Rep_net (net within S) = (\<lambda>A. A \<inter> S) ` Rep_net net"
-unfolding within_def
-apply (rule Abs_net_inverse')
-apply (rule image_nonempty, rule Rep_net_nonempty)
-apply (clarsimp, rename_tac A B)
-apply (drule (1) Rep_net_directed)
-apply (clarify, rule_tac x=C in bexI, auto)
-done
-
-lemma Rep_net_at:
-  "Rep_net (at a) = ((\<lambda>S. S - {a}) ` {S. open S \<and> a \<in> S})"
-unfolding at_def
-apply (rule Abs_net_inverse')
-apply (rule image_nonempty)
-apply (rule_tac x="UNIV" in exI, simp)
-apply (clarsimp, rename_tac S T)
-apply (rule_tac x="S \<inter> T" in exI, auto simp add: open_Int)
-done
+definition
+  at :: "'a::topological_space \<Rightarrow> 'a net"
+where [code del]:
+  "at a = Abs_net (\<lambda>P. \<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x))"
 
 lemma eventually_sequentially:
   "eventually P sequentially \<longleftrightarrow> (\<exists>N. \<forall>n\<ge>N. P n)"
-unfolding eventually_def Rep_net_sequentially by auto
+unfolding sequentially_def
+proof (rule eventually_Abs_net, rule is_filter.intro)
+  fix P Q :: "nat \<Rightarrow> bool"
+  assume "\<exists>i. \<forall>n\<ge>i. P n" and "\<exists>j. \<forall>n\<ge>j. Q n"
+  then obtain i j where "\<forall>n\<ge>i. P n" and "\<forall>n\<ge>j. Q n" by auto
+  then have "\<forall>n\<ge>max i j. P n \<and> Q n" by simp
+  then show "\<exists>k. \<forall>n\<ge>k. P n \<and> Q n" ..
+qed auto
 
 lemma eventually_within:
   "eventually P (net within S) = eventually (\<lambda>x. x \<in> S \<longrightarrow> P x) net"
-unfolding eventually_def Rep_net_within by auto
+unfolding within_def
+by (rule eventually_Abs_net, rule is_filter.intro)
+   (auto elim!: eventually_rev_mp)
+
+lemma within_UNIV: "net within UNIV = net"
+  unfolding expand_net_eq eventually_within by simp
 
 lemma eventually_at_topological:
   "eventually P (at a) \<longleftrightarrow> (\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x))"
-unfolding eventually_def Rep_net_at by auto
+unfolding at_def
+proof (rule eventually_Abs_net, rule is_filter.intro)
+  have "open UNIV \<and> a \<in> UNIV \<and> (\<forall>x\<in>UNIV. x \<noteq> a \<longrightarrow> True)" by simp
+  thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> True)" by - rule
+next
+  fix P Q
+  assume "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x)"
+     and "\<exists>T. open T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<noteq> a \<longrightarrow> Q x)"
+  then obtain S T where
+    "open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x)"
+    "open T \<and> a \<in> T \<and> (\<forall>x\<in>T. x \<noteq> a \<longrightarrow> Q x)" by auto
+  hence "open (S \<inter> T) \<and> a \<in> S \<inter> T \<and> (\<forall>x\<in>(S \<inter> T). x \<noteq> a \<longrightarrow> P x \<and> Q x)"
+    by (simp add: open_Int)
+  thus "\<exists>S. open S \<and> a \<in> S \<and> (\<forall>x\<in>S. x \<noteq> a \<longrightarrow> P x \<and> Q x)" by - rule
+qed auto
 
 lemma eventually_at:
   fixes a :: "'a::metric_space"
--- a/src/HOL/MicroJava/J/Example.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/MicroJava/J/Example.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -371,7 +371,7 @@
 done
 
 ML {* val t = resolve_tac @{thms ty_expr_ty_exprs_wt_stmt.intros} 1 *}
-lemma wt_test: "(tprg, empty(e\<mapsto>Class Base))\<turnstile>  
+schematic_lemma wt_test: "(tprg, empty(e\<mapsto>Class Base))\<turnstile>  
   Expr(e::=NewC Ext);; Expr({Base}LAcc e..foo({?pTs'}[Lit Null]))\<surd>"
 apply (tactic t) -- ";;"
 apply  (tactic t) -- "Expr"
@@ -400,7 +400,7 @@
 
 declare split_if [split del]
 declare init_vars_def [simp] c_hupd_def [simp] cast_ok_def [simp]
-lemma exec_test: 
+schematic_lemma exec_test: 
 " [|new_Addr (heap (snd s0)) = (a, None)|] ==>  
   tprg\<turnstile>s0 -test-> ?s"
 apply (unfold test_def)
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -22,8 +22,6 @@
   imports Convex_Euclidean_Space
 begin
 
-declare norm_scaleR[simp]
- 
 lemma brouwer_compactness_lemma:
   assumes "compact s" "continuous_on s f" "\<not> (\<exists>x\<in>s. (f x = (0::real^'n)))"
   obtains d where "0 < d" "\<forall>x\<in>s. d \<le> norm(f x)" proof(cases "s={}") case False
@@ -131,7 +129,7 @@
 lemma image_lemma_2: assumes "finite s" "finite t" "card s = card t" "f ` s \<subseteq> t" "f ` s \<noteq> t" "b \<in> t"
   shows "(card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 0) \<or>
          (card {s'. \<exists>a\<in>s. (s' = s - {a}) \<and> f ` s' = t - {b}} = 2)" proof(cases "{a\<in>s. f ` (s - {a}) = t - {b}} = {}")
-  case True thus ?thesis apply-apply(rule disjI1, rule image_lemma_0) using assms(1) by(auto simp add:card_0_eq)
+  case True thus ?thesis apply-apply(rule disjI1, rule image_lemma_0) using assms(1) by auto
 next let ?M = "{a\<in>s. f ` (s - {a}) = t - {b}}"
   case False then obtain a where "a\<in>?M" by auto hence a:"a\<in>s" "f ` (s - {a}) = t - {b}" by auto
   have "f a \<in> t - {b}" using a and assms by auto
@@ -350,7 +348,7 @@
         (\<forall>x\<in>s. \<forall>j. j\<notin>{1..n} \<longrightarrow> (x j = p)) \<and>
         (\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x))"
 
-lemma ksimplexI:"card s = n + 1 \<Longrightarrow>  \<forall>x\<in>s. \<forall>j. x j \<le> p \<Longrightarrow> \<forall>x\<in>s. \<forall>j. j \<notin> {1..?n} \<longrightarrow> x j = ?p \<Longrightarrow> \<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x \<Longrightarrow> ksimplex p n s"
+lemma ksimplexI:"card s = n + 1 \<Longrightarrow>  \<forall>x\<in>s. \<forall>j. x j \<le> p \<Longrightarrow> \<forall>x\<in>s. \<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p \<Longrightarrow> \<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x \<Longrightarrow> ksimplex p n s"
   unfolding ksimplex_def by auto
 
 lemma ksimplex_eq: "ksimplex p n (s::(nat \<Rightarrow> nat) set) \<longleftrightarrow>
@@ -393,7 +391,7 @@
 
 lemma ksimplexD:
   assumes "ksimplex p n s"
-  shows "card s = n + 1" "finite s" "card s = n + 1" "\<forall>x\<in>s. \<forall>j. x j \<le> p" "\<forall>x\<in>s. \<forall>j. j \<notin> {1..?n} \<longrightarrow> x j = p"
+  shows "card s = n + 1" "finite s" "card s = n + 1" "\<forall>x\<in>s. \<forall>j. x j \<le> p" "\<forall>x\<in>s. \<forall>j. j \<notin> {1..n} \<longrightarrow> x j = p"
   "\<forall>x\<in>s. \<forall>y\<in>s. kle n x y \<or> kle n y x" using assms unfolding ksimplex_eq by auto
 
 lemma ksimplex_successor:
@@ -1350,7 +1348,7 @@
   using assms by auto
 
 text {*Still more general form; could derive this directly without using the 
-  rather involved HOMEOMORPHIC_CONVEX_COMPACT theorem, just using
+  rather involved @{text "HOMEOMORPHIC_CONVEX_COMPACT"} theorem, just using
   a scaling and translation to put the set inside the unit cube. *}
 
 lemma brouwer: fixes f::"real^'n \<Rightarrow> real^'n"
@@ -1383,7 +1381,7 @@
     apply(rule_tac x=x in bexI) apply assumption+ apply(rule continuous_on_intros)+
     unfolding frontier_cball subset_eq Ball_def image_iff apply(rule,rule,erule bexE)
     unfolding vector_dist_norm apply(simp add: * norm_minus_commute) . note x = this
-  hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:group_simps)
+  hence "scaleR 2 a = scaleR 1 x + scaleR 1 x" by(auto simp add:algebra_simps)
   hence "a = x" unfolding scaleR_left_distrib[THEN sym] by auto 
   thus False using x using assms by auto qed
 
@@ -1396,7 +1394,7 @@
  "interval_bij (a,b) (u,v) = (\<lambda>x. (\<chi> i. (v$i - u$i) / (b$i - a$i) * x$i) +
             (\<chi> i. u$i - (v$i - u$i) / (b$i - a$i) * a$i))"
   apply rule unfolding Cart_eq interval_bij_def vector_component_simps
-  by(auto simp add:group_simps field_simps add_divide_distrib[THEN sym]) 
+  by(auto simp add: field_simps add_divide_distrib[THEN sym]) 
 
 lemma continuous_interval_bij:
   "continuous (at x) (interval_bij (a,b::real^'n) (u,v))" 
@@ -1444,7 +1442,7 @@
 lemma infnorm_eq_1_imp: assumes "infnorm (x::real^2) = 1" shows "abs(x$1) \<le> 1" "abs(x$2) \<le> 1"
   using assms unfolding infnorm_eq_1_2 by auto
 
-lemma fashoda_unit: fixes f g::"real^1 \<Rightarrow> real^2"
+lemma fashoda_unit: fixes f g::"real \<Rightarrow> real^2"
   assumes "f ` {- 1..1} \<subseteq> {- 1..1}" "g ` {- 1..1} \<subseteq> {- 1..1}"
   "continuous_on {- 1..1} f"  "continuous_on {- 1..1} g"
   "f (- 1)$1 = - 1" "f 1$1 = 1" "g (- 1) $2 = -1" "g 1 $2 = 1"
@@ -1457,13 +1455,13 @@
   have lem2:"\<forall>z. z\<noteq>0 \<longrightarrow> infnorm(sqprojection z) = 1" unfolding sqprojection_def
     unfolding infnorm_mul[unfolded smult_conv_scaleR] unfolding abs_inverse real_abs_infnorm
     unfolding infnorm_eq_0[THEN sym] by auto
-  let ?F = "(\<lambda>w::real^2. (f \<circ> vec1 \<circ> (\<lambda>x. x$1)) w - (g \<circ> vec1 \<circ> (\<lambda>x. x$2)) w)"
-  have *:"\<And>i. vec1 ` (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real^1}"
+  let ?F = "(\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w)"
+  have *:"\<And>i. (\<lambda>x::real^2. x $ i) ` {- 1..1} = {- 1..1::real}"
     apply(rule set_ext) unfolding image_iff Bex_def mem_interval apply rule defer 
-    apply(rule_tac x="dest_vec1 x" in exI) apply rule apply(rule_tac x="vec (dest_vec1 x)" in exI) by auto
-  { fix x assume "x \<in> (\<lambda>w. (f \<circ> vec1 \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> vec1 \<circ> (\<lambda>x. x $ 2)) w) ` {- 1..1::real^2}"
+    apply(rule_tac x="vec x" in exI) by auto
+  { fix x assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` {- 1..1::real^2}"
     then guess w unfolding image_iff .. note w = this
-    hence "x \<noteq> 0" using as[of "vec1 (w$1)" "vec1 (w$2)"] unfolding mem_interval by auto} note x0=this
+    hence "x \<noteq> 0" using as[of "w$1" "w$2"] unfolding mem_interval by auto} note x0=this
   have 21:"\<And>i::2. i\<noteq>1 \<Longrightarrow> i=2" using UNIV_2 by auto
   have 1:"{- 1<..<1::real^2} \<noteq> {}" unfolding interval_eq_empty by auto
   have 2:"continuous_on {- 1..1} (negatex \<circ> sqprojection \<circ> ?F)" apply(rule continuous_on_intros continuous_on_component continuous_on_vec1)+
@@ -1494,50 +1492,50 @@
       unfolding sqprojection_def vector_component_simps Cart_nth.scaleR real_scaleR_def
       unfolding zero_less_mult_iff mult_less_0_iff by(auto simp add:field_simps) qed
   note lem3 = this[rule_format]
-  have x1:"vec1 (x $ 1) \<in> {- 1..1::real^1}" "vec1 (x $ 2) \<in> {- 1..1::real^1}" using x(1) unfolding mem_interval by auto
-  hence nz:"f (vec1 (x $ 1)) - g (vec1 (x $ 2)) \<noteq> 0" unfolding right_minus_eq apply-apply(rule as) by auto
+  have x1:"x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}" using x(1) unfolding mem_interval by auto
+  hence nz:"f (x $ 1) - g (x $ 2) \<noteq> 0" unfolding right_minus_eq apply-apply(rule as) by auto
   have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1" using nx unfolding infnorm_eq_1_2 by auto 
   thus False proof- fix P Q R S 
     presume "P \<or> Q \<or> R \<or> S" "P\<Longrightarrow>False" "Q\<Longrightarrow>False" "R\<Longrightarrow>False" "S\<Longrightarrow>False" thus False by auto
-  next assume as:"x$1 = 1" hence "vec1 (x$1) = 1" unfolding Cart_eq by auto
-    hence *:"f (vec1 (x $ 1)) $ 1 = 1" using assms(6) by auto
-    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 1 < 0"
+  next assume as:"x$1 = 1"
+    hence *:"f (x $ 1) $ 1 = 1" using assms(6) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0"
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]]
       unfolding as negatex_def vector_2 by auto moreover
-    from x1 have "g (vec1 (x $ 2)) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
+    from x1 have "g (x $ 2) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
     ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
       apply(erule_tac x=1 in allE) by auto 
-  next assume as:"x$1 = -1" hence "vec1 (x$1) = - 1" unfolding Cart_eq by auto
-    hence *:"f (vec1 (x $ 1)) $ 1 = - 1" using assms(5) by auto
-    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 1 > 0"
+  next assume as:"x$1 = -1"
+    hence *:"f (x $ 1) $ 1 = - 1" using assms(5) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0"
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=1]]
       unfolding as negatex_def vector_2 by auto moreover
-    from x1 have "g (vec1 (x $ 2)) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
+    from x1 have "g (x $ 2) \<in> {- 1..1}" apply-apply(rule assms(2)[unfolded subset_eq,rule_format]) by auto
     ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
       apply(erule_tac x=1 in allE) by auto
-  next assume as:"x$2 = 1" hence "vec1 (x$2) = 1" unfolding Cart_eq by auto
-    hence *:"g (vec1 (x $ 2)) $ 2 = 1" using assms(8) by auto
-    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 2 > 0"
+  next assume as:"x$2 = 1"
+    hence *:"g (x $ 2) $ 2 = 1" using assms(8) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0"
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]]
       unfolding as negatex_def vector_2 by auto moreover
-    from x1 have "f (vec1 (x $ 1)) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
+    from x1 have "f (x $ 1) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
     ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
      apply(erule_tac x=2 in allE) by auto
- next assume as:"x$2 = -1" hence "vec1 (x$2) = - 1" unfolding Cart_eq by auto
-    hence *:"g (vec1 (x $ 2)) $ 2 = - 1" using assms(7) by auto
-    have "sqprojection (f (vec1 (x$1)) - g (vec1 (x$2))) $ 2 < 0"
+ next assume as:"x$2 = -1"
+    hence *:"g (x $ 2) $ 2 = - 1" using assms(7) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0"
       using x(2)[unfolded o_def Cart_eq,THEN spec[where x=2]]
       unfolding as negatex_def vector_2 by auto moreover
-    from x1 have "f (vec1 (x $ 1)) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
+    from x1 have "f (x $ 1) \<in> {- 1..1}" apply-apply(rule assms(1)[unfolded subset_eq,rule_format]) by auto
     ultimately show False unfolding lem3[OF nz] vector_component_simps * mem_interval 
       apply(erule_tac x=2 in allE) by auto qed(auto) qed
 
-lemma fashoda_unit_path: fixes f ::"real^1 \<Rightarrow> real^2" and g ::"real^1 \<Rightarrow> real^2"
+lemma fashoda_unit_path: fixes f ::"real \<Rightarrow> real^2" and g ::"real \<Rightarrow> real^2"
   assumes "path f" "path g" "path_image f \<subseteq> {- 1..1}" "path_image g \<subseteq> {- 1..1}"
   "(pathstart f)$1 = -1" "(pathfinish f)$1 = 1"  "(pathstart g)$2 = -1" "(pathfinish g)$2 = 1"
   obtains z where "z \<in> path_image f" "z \<in> path_image g" proof-
   note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]
-  def iscale \<equiv> "\<lambda>z::real^1. inverse 2 *\<^sub>R (z + 1)"
+  def iscale \<equiv> "\<lambda>z::real. inverse 2 *\<^sub>R (z + 1)"
   have isc:"iscale ` {- 1..1} \<subseteq> {0..1}" unfolding iscale_def by(auto)
   have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t" proof(rule fashoda_unit) 
     show "(f \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}" "(g \<circ> iscale) ` {- 1..1} \<subseteq> {- 1..1}"
@@ -1691,11 +1689,11 @@
       path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union>
       path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union>
       path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2)
-      by(auto simp add: pathstart_join pathfinish_join path_image_join path_image_linepath path_join path_linepath) 
+      by(auto simp add: path_image_join path_linepath)
   have abab: "{a..b} \<subseteq> {?a..?b}" by(auto simp add:vector_le_def forall_2 vector_2)
   guess z apply(rule fashoda[of ?P1 ?P2 ?a ?b])
     unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2 proof-
-    show "path ?P1" "path ?P2" using assms by(auto simp add: pathstart_join pathfinish_join path_join)
+    show "path ?P1" "path ?P2" using assms by auto
     have "path_image ?P1 \<subseteq> {?a .. ?b}" unfolding P1P2 path_image_linepath apply(rule Un_least)+ defer 3
       apply(rule_tac[1-4] convex_interval(1)[unfolded convex_contains_segment,rule_format])
       unfolding mem_interval forall_2 vector_2 using ab startfin abab assms(3)
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -24,7 +24,7 @@
 lemma dest_vec1_simps[simp]: fixes a::"real^1"
   shows "a$1 = 0 \<longleftrightarrow> a = 0" (*"a \<le> 1 \<longleftrightarrow> dest_vec1 a \<le> 1" "0 \<le> a \<longleftrightarrow> 0 \<le> dest_vec1 a"*)
   "a \<le> b \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 b" "dest_vec1 (1::real^1) = 1"
-  by(auto simp add:vector_component_simps forall_1 Cart_eq)
+  by(auto simp add: vector_le_def Cart_eq)
 
 lemma norm_not_0:"(x::real^'n)\<noteq>0 \<Longrightarrow> norm x \<noteq> 0" by auto
 
@@ -50,7 +50,7 @@
 lemma mem_interval_1: fixes x :: "real^1" shows
  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
  "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp_all add: Cart_eq vector_less_def vector_le_def forall_1)
+by(simp_all add: Cart_eq vector_less_def vector_le_def)
 
 lemma image_smult_interval:"(\<lambda>x. m *\<^sub>R (x::real^'n)) ` {a..b} =
   (if {a..b} = {} then {} else if 0 \<le> m then {m *\<^sub>R a..m *\<^sub>R b} else {m *\<^sub>R b..m *\<^sub>R a})"
@@ -66,8 +66,7 @@
   apply(rule_tac [!] allI)apply(rule_tac [!] impI)
   apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI)
   apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI)
-  by (auto simp add: vector_less_def vector_le_def forall_1
-    vec1_dest_vec1[unfolded One_nat_def])
+  by (auto simp add: vector_less_def vector_le_def)
 
 lemma dest_vec1_setsum: assumes "finite S"
   shows " dest_vec1 (setsum f S) = setsum (\<lambda>x. dest_vec1 (f x)) S"
@@ -90,7 +89,7 @@
   using one_le_card_finite by auto
 
 lemma real_dimindex_ge_1:"real (CARD('n::finite)) \<ge> 1" 
-  by(metis dimindex_ge_1 linorder_not_less real_eq_of_nat real_le_trans real_of_nat_1 real_of_nat_le_iff) 
+  by(metis dimindex_ge_1 real_eq_of_nat real_of_nat_1 real_of_nat_le_iff) 
 
 lemma real_dimindex_gt_0:"real (CARD('n::finite)) > 0" apply(rule less_le_trans[OF _ real_dimindex_ge_1]) by auto
 
@@ -401,9 +400,38 @@
 lemma convex_halfspace_gt: "convex {x. inner a x > b}"
    using convex_halfspace_lt[of "-a" "-b"] by auto
 
+lemma convex_real_interval:
+  fixes a b :: "real"
+  shows "convex {a..}" and "convex {..b}"
+  and "convex {a<..}" and "convex {..<b}"
+  and "convex {a..b}" and "convex {a<..b}"
+  and "convex {a..<b}" and "convex {a<..<b}"
+proof -
+  have "{a..} = {x. a \<le> inner 1 x}" by auto
+  thus 1: "convex {a..}" by (simp only: convex_halfspace_ge)
+  have "{..b} = {x. inner 1 x \<le> b}" by auto
+  thus 2: "convex {..b}" by (simp only: convex_halfspace_le)
+  have "{a<..} = {x. a < inner 1 x}" by auto
+  thus 3: "convex {a<..}" by (simp only: convex_halfspace_gt)
+  have "{..<b} = {x. inner 1 x < b}" by auto
+  thus 4: "convex {..<b}" by (simp only: convex_halfspace_lt)
+  have "{a..b} = {a..} \<inter> {..b}" by auto
+  thus "convex {a..b}" by (simp only: convex_Int 1 2)
+  have "{a<..b} = {a<..} \<inter> {..b}" by auto
+  thus "convex {a<..b}" by (simp only: convex_Int 3 2)
+  have "{a..<b} = {a..} \<inter> {..<b}" by auto
+  thus "convex {a..<b}" by (simp only: convex_Int 1 4)
+  have "{a<..<b} = {a<..} \<inter> {..<b}" by auto
+  thus "convex {a<..<b}" by (simp only: convex_Int 3 4)
+qed
+
+lemma convex_box:
+  assumes "\<And>i. convex {x. P i x}"
+  shows "convex {x. \<forall>i. P i (x$i)}"
+using assms unfolding convex_def by auto
+
 lemma convex_positive_orthant: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
-  unfolding convex_def apply auto apply(erule_tac x=i in allE)+
-  apply(rule add_nonneg_nonneg) by(auto simp add: mult_nonneg_nonneg)
+by (rule convex_box, simp add: atLeast_def [symmetric] convex_real_interval)
 
 subsection {* Explicit expressions for convexity in terms of arbitrary sums. *}
 
@@ -451,8 +479,8 @@
 next 
   fix t u assume asm:"\<forall>x\<in>s. \<forall>y\<in>s. \<forall>u\<ge>0. \<forall>v\<ge>0. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> s" "finite (t::'a set)"
   (*"finite t" "t \<subseteq> s" "\<forall>x\<in>t. (0::real) \<le> u x" "setsum u t = 1"*)
-  from this(2) have "\<forall>u. t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" apply(induct_tac t rule:finite_induct)
-    prefer 3 apply (rule,rule) apply(erule conjE)+ proof-
+  from this(2) have "\<forall>u. t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s" apply(induct t rule:finite_induct)
+    prefer 2 apply (rule,rule) apply(erule conjE)+ proof-
     fix x f u assume ind:"\<forall>u. f \<subseteq> s \<and> (\<forall>x\<in>f. 0 \<le> u x) \<and> setsum u f = 1 \<longrightarrow> (\<Sum>x\<in>f. u x *\<^sub>R x) \<in> s"
     assume as:"finite f" "x \<notin> f" "insert x f \<subseteq> s" "\<forall>x\<in>insert x f. 0 \<le> u x" "setsum u (insert x f) = (1::real)"
     show "(\<Sum>x\<in>insert x f. u x *\<^sub>R x) \<in> s" proof(cases "u x = 1")
@@ -617,7 +645,7 @@
       using assms(1)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
       using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x=u]]
       apply - apply(rule add_mono) by auto
-    hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: ring_simps)  }
+    hence "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \<le> u * (f x + g x) + v * (f y + g y)" by (simp add: field_simps)  }
   thus ?thesis unfolding convex_on_def by auto 
 qed
 
@@ -625,7 +653,7 @@
   assumes "0 \<le> (c::real)" "convex_on s f"
   shows "convex_on s (\<lambda>x. c * f x)"
 proof-
-  have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: ring_simps)
+  have *:"\<And>u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: field_simps)
   show ?thesis using assms(2) and mult_mono1[OF _ assms(1)] unfolding convex_on_def and * by auto
 qed
 
@@ -747,7 +775,7 @@
 lemma convex_cball:
   fixes x :: "'a::real_normed_vector"
   shows "convex(cball x e)"
-proof(auto simp add: convex_def Ball_def mem_cball)
+proof(auto simp add: convex_def Ball_def)
   fix y z assume yz:"dist x y \<le> e" "dist x z \<le> e"
   fix u v ::real assume uv:" 0 \<le> u" "0 \<le> v" "u + v = 1"
   have "dist x (u *\<^sub>R y + v *\<^sub>R z) \<le> u * dist x y + v * dist x z" using uv yz
@@ -1031,7 +1059,7 @@
 proof-
   have fin:"finite {a,b,c}" "finite {b,c}" "finite {c}" by auto
   have *:"\<And>x y z ::real. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z"
-         "\<And>x y z ::real^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: ring_simps)
+         "\<And>x y z ::real^_. x + y + z = 1 \<longleftrightarrow> x = 1 - y - z" by (auto simp add: field_simps)
   show ?thesis unfolding convex_hull_finite[OF fin(1)] and Collect_def and convex_hull_finite_step[OF fin(2)] and *
     unfolding convex_hull_finite_step[OF fin(3)] apply(rule ext) apply simp apply auto
     apply(rule_tac x=va in exI) apply (rule_tac x="u c" in exI) apply simp
@@ -1115,7 +1143,7 @@
     hence "x + y \<in> s" using `?lhs`[unfolded convex_def, THEN conjunct1]
       apply(erule_tac x="2*\<^sub>R x" in ballE) apply(erule_tac x="2*\<^sub>R y" in ballE)
       apply(erule_tac x="1/2" in allE) apply simp apply(erule_tac x="1/2" in allE) by auto  }
-  thus ?thesis unfolding convex_def cone_def by auto
+  thus ?thesis unfolding convex_def cone_def by blast
 qed
 
 lemma affine_dependent_biggerset: fixes s::"(real^'n) set"
@@ -1230,7 +1258,7 @@
   fixes s :: "'a::real_normed_vector set"
   assumes "open s"
   shows "open(convex hull s)"
-  unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) 
+  unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10)
 proof(rule, rule) fix a
   assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
   then obtain t u where obt:"finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a" by auto
@@ -1250,7 +1278,7 @@
       hence "Min i \<le> b x" unfolding i_def apply(rule_tac Min_le) using obt(1) by auto
       hence "x + (y - a) \<in> cball x (b x)" using y unfolding mem_cball dist_norm by auto
       moreover from `x\<in>t` have "x\<in>s" using obt(2) by auto
-      ultimately have "x + (y - a) \<in> s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by auto }
+      ultimately have "x + (y - a) \<in> s" using y and b[THEN bspec[where x=x]] unfolding subset_eq by fast }
     moreover
     have *:"inj_on (\<lambda>v. v + (y - a)) t" unfolding inj_on_def by auto
     have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
@@ -1320,7 +1348,7 @@
   show ?thesis unfolding caratheodory[of s]
   proof(induct ("CARD('n) + 1"))
     have *:"{x.\<exists>sa. finite sa \<and> sa \<subseteq> s \<and> card sa \<le> 0 \<and> x \<in> convex hull sa} = {}" 
-      using compact_empty by (auto simp add: convex_hull_empty)
+      using compact_empty by auto
     case 0 thus ?case unfolding * by simp
   next
     case (Suc n)
@@ -1330,11 +1358,11 @@
         fix x assume "\<exists>t. finite t \<and> t \<subseteq> s \<and> card t \<le> Suc n \<and> x \<in> convex hull t"
         then obtain t where t:"finite t" "t \<subseteq> s" "card t \<le> Suc n" "x \<in> convex hull t" by auto
         show "x\<in>s" proof(cases "card t = 0")
-          case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by(simp add: convex_hull_empty)
+          case True thus ?thesis using t(4) unfolding card_0_eq[OF t(1)] by simp
         next
           case False hence "card t = Suc 0" using t(3) `n=0` by auto
           then obtain a where "t = {a}" unfolding card_Suc_eq by auto
-          thus ?thesis using t(2,4) by (simp add: convex_hull_singleton)
+          thus ?thesis using t(2,4) by simp
         qed
       next
         fix x assume "x\<in>s"
@@ -1369,7 +1397,7 @@
           show ?P proof(cases "u={}")
             case True hence "x=a" using t(4)[unfolded au] by auto
             show ?P unfolding `x=a` apply(rule_tac x=a in exI, rule_tac x=a in exI, rule_tac x=1 in exI)
-              using t and `n\<noteq>0` unfolding au by(auto intro!: exI[where x="{a}"] simp add: convex_hull_singleton)
+              using t and `n\<noteq>0` unfolding au by(auto intro!: exI[where x="{a}"])
           next
             case False obtain ux vx b where obt:"ux\<ge>0" "vx\<ge>0" "ux + vx = 1" "b \<in> convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b"
               using t(4)[unfolded au convex_hull_insert[OF False]] by auto
@@ -1382,7 +1410,7 @@
       qed
       thus ?thesis using compact_convex_combinations[OF assms Suc] by simp 
     qed
-  qed 
+  qed
 qed
 
 lemma finite_imp_compact_convex_hull:
@@ -1519,7 +1547,7 @@
 subsection {* Closest point of a convex set is unique, with a continuous projection. *}
 
 definition
-  closest_point :: "(real ^ 'n) set \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 'n" where
+  closest_point :: "'a::{real_inner,heine_borel} set \<Rightarrow> 'a \<Rightarrow> 'a" where
  "closest_point s a = (SOME x. x \<in> s \<and> (\<forall>y\<in>s. dist a x \<le> dist a y))"
 
 lemma closest_point_exists:
@@ -1553,7 +1581,7 @@
 lemma norm_le: "norm x \<le> norm y \<longleftrightarrow> inner x x \<le> inner y y"
   unfolding norm_eq_sqrt_inner by simp
 
-lemma closer_points_lemma: fixes y::"real^'n"
+lemma closer_points_lemma:
   assumes "inner y z > 0"
   shows "\<exists>u>0. \<forall>v>0. v \<le> u \<longrightarrow> norm(v *\<^sub>R z - y) < norm y"
 proof- have z:"inner z z > 0" unfolding inner_gt_zero_iff using assms by auto
@@ -1564,7 +1592,6 @@
   qed(rule divide_pos_pos, auto) qed
 
 lemma closer_point_lemma:
-  fixes x y z :: "real ^ 'n"
   assumes "inner (y - x) (z - x) > 0"
   shows "\<exists>u>0. u \<le> 1 \<and> dist (x + u *\<^sub>R (z - x)) y < dist x y"
 proof- obtain u where "u>0" and u:"\<forall>v>0. v \<le> u \<longrightarrow> norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)"
@@ -1573,7 +1600,6 @@
     unfolding dist_norm by(auto simp add: norm_minus_commute field_simps) qed
 
 lemma any_closest_point_dot:
-  fixes s :: "(real ^ _) set"
   assumes "convex s" "closed s" "x \<in> s" "y \<in> s" "\<forall>z\<in>s. dist a x \<le> dist a z"
   shows "inner (a - x) (y - x) \<le> 0"
 proof(rule ccontr) assume "\<not> inner (a - x) (y - x) \<le> 0"
@@ -1582,7 +1608,7 @@
   thus False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp add: dist_commute algebra_simps) qed
 
 lemma any_closest_point_unique:
-  fixes s :: "(real ^ _) set"
+  fixes x :: "'a::real_inner"
   assumes "convex s" "closed s" "x \<in> s" "y \<in> s"
   "\<forall>z\<in>s. dist a x \<le> dist a z" "\<forall>z\<in>s. dist a y \<le> dist a z"
   shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)]
@@ -1634,7 +1660,7 @@
 subsection {* Various point-to-set separating/supporting hyperplane theorems. *}
 
 lemma supporting_hyperplane_closed_point:
-  fixes s :: "(real ^ _) set"
+  fixes z :: "'a::{real_inner,heine_borel}"
   assumes "convex s" "closed s" "s \<noteq> {}" "z \<notin> s"
   shows "\<exists>a b. \<exists>y\<in>s. inner a z < b \<and> (inner a y = b) \<and> (\<forall>x\<in>s. inner a x \<ge> b)"
 proof-
@@ -1653,7 +1679,7 @@
 qed
 
 lemma separating_hyperplane_closed_point:
-  fixes s :: "(real ^ _) set"
+  fixes z :: "'a::{real_inner,heine_borel}"
   assumes "convex s" "closed s" "z \<notin> s"
   shows "\<exists>a b. inner a z < b \<and> (\<forall>x\<in>s. inner a x > b)"
 proof(cases "s={}")
@@ -1824,7 +1850,7 @@
 lemma convex_hull_scaling:
   "convex hull ((\<lambda>x. c *\<^sub>R x) ` s) = (\<lambda>x. c *\<^sub>R x) ` (convex hull s)"
   apply(cases "c=0") defer apply(rule convex_hull_bilemma[rule_format, of _ _ inverse]) apply(rule convex_hull_scaling_lemma)
-  unfolding image_image scaleR_scaleR by(auto simp add:image_constant_conv convex_hull_eq_empty)
+  unfolding image_image scaleR_scaleR by(auto simp add:image_constant_conv)
 
 lemma convex_hull_affinity:
   "convex hull ((\<lambda>x. a + c *\<^sub>R x) ` s) = (\<lambda>x. a + c *\<^sub>R x) ` (convex hull s)"
@@ -2202,9 +2228,9 @@
 
 subsection {* Epigraphs of convex functions. *}
 
-definition "epigraph s (f::real^'n \<Rightarrow> real) = {xy. fstcart xy \<in> s \<and> f(fstcart xy) \<le> dest_vec1 (sndcart xy)}"
-
-lemma mem_epigraph: "(pastecart x (vec1 y)) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" unfolding epigraph_def by auto
+definition "epigraph s (f::_ \<Rightarrow> real) = {xy. fst xy \<in> s \<and> f (fst xy) \<le> snd xy}"
+
+lemma mem_epigraph: "(x, y) \<in> epigraph s f \<longleftrightarrow> x \<in> s \<and> f x \<le> y" unfolding epigraph_def by auto
 
 (** move this**)
 lemma forall_dest_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P(dest_vec1 x))" 
@@ -2213,12 +2239,12 @@
 (** This might break sooner or later. In fact it did already once. **)
 lemma convex_epigraph: 
   "convex(epigraph s f) \<longleftrightarrow> convex_on s f \<and> convex s"
-  unfolding convex_def convex_on_def unfolding Ball_def forall_pastecart epigraph_def
-  unfolding mem_Collect_eq fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR] Ball_def[symmetric] unfolding vector_add_component vector_scaleR_component
-  apply(subst forall_dest_vec1[THEN sym])+ 
+  unfolding convex_def convex_on_def
+  unfolding Ball_def split_paired_All epigraph_def
+  unfolding mem_Collect_eq fst_conv snd_conv fst_add snd_add fst_scaleR snd_scaleR Ball_def[symmetric]
   apply safe defer apply(erule_tac x=x in allE,erule_tac x="f x" in allE) apply safe
   apply(erule_tac x=xa in allE,erule_tac x="f xa" in allE) prefer 3
-  apply(rule_tac y="u * f x + v * f xb" in order_trans) defer by(auto intro!:mult_left_mono add_mono)
+  apply(rule_tac y="u * f a + v * f aa" in order_trans) defer by(auto intro!:mult_left_mono add_mono)
 
 lemma convex_epigraphI:
   "convex_on s f \<Longrightarrow> convex s \<Longrightarrow> convex(epigraph s f)"
@@ -2245,21 +2271,29 @@
   apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule 
   apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto
 
+lemma fst_setsum: "fst (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. fst (f x))"
+by (cases "finite A", induct set: finite, simp_all)
+
+lemma snd_setsum: "snd (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. snd (f x))"
+by (cases "finite A", induct set: finite, simp_all)
+
 lemma convex_on:
-  fixes s :: "(real ^ _) set"
   assumes "convex s"
   shows "convex_on s f \<longleftrightarrow> (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow>
    f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k} ) "
   unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
-  unfolding sndcart_setsum[OF finite_atLeastAtMost] fstcart_setsum[OF finite_atLeastAtMost] dest_vec1_setsum[OF finite_atLeastAtMost]
-  unfolding fstcart_pastecart sndcart_pastecart sndcart_add sndcart_cmul [where 'a=real, unfolded smult_conv_scaleR] fstcart_add fstcart_cmul [where 'a=real, unfolded smult_conv_scaleR]  unfolding vector_scaleR_component
-  apply(subst forall_of_pastecart)+ apply(subst forall_of_dest_vec1)+ apply rule
-  using assms[unfolded convex] apply simp apply(rule,rule,rule)
-  apply(erule_tac x=k in allE, erule_tac x=u in allE, erule_tac x=x in allE) apply rule apply rule apply rule defer
-  apply(rule_tac j="\<Sum>i = 1..k. u i * f (x i)" in real_le_trans)
-  defer apply(rule setsum_mono) apply(erule conjE)+ apply(erule_tac x=i in allE) unfolding real_scaleR_def
+  unfolding fst_setsum snd_setsum fst_scaleR snd_scaleR
+  apply safe
+  apply (drule_tac x=k in spec)
+  apply (drule_tac x=u in spec)
+  apply (drule_tac x="\<lambda>i. (x i, f (x i))" in spec)
+  apply simp
+  using assms[unfolded convex] apply simp
+  apply(rule_tac j="\<Sum>i = 1..k. u i * f (fst (x i))" in real_le_trans)
+  defer apply(rule setsum_mono) apply(erule_tac x=i in allE) unfolding real_scaleR_def
   apply(rule mult_left_mono)using assms[unfolded convex] by auto
 
+
 subsection {* Convexity of general and special intervals. *}
 
 lemma is_interval_convex:
@@ -2275,10 +2309,10 @@
   } moreover
   { fix a b assume "\<not> u * a + v * b \<le> a"
     hence "v * b > (1 - u) * a" unfolding not_le using as(4) by(auto simp add: field_simps)
-    hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: ring_simps)
+    hence "a < b" unfolding * using as(4) apply(rule_tac mult_left_less_imp_less) by(auto simp add: field_simps)
     hence "u * a + v * b \<le> b" unfolding ** using **(2) as(3) by(auto simp add: field_simps intro!:mult_right_mono) }
   ultimately show "u *\<^sub>R x + v *\<^sub>R y \<in> s" apply- apply(rule assms[unfolded is_interval_def, rule_format, OF as(1,2)])
-    using as(3-) dimindex_ge_1 apply- by(auto simp add: vector_component) qed
+    using as(3-) dimindex_ge_1 by auto qed
 
 lemma is_interval_connected:
   fixes s :: "(real ^ _) set"
@@ -2301,7 +2335,7 @@
   hence *:"dest_vec1 a < dest_vec1 x" "dest_vec1 x < dest_vec1 b" apply(rule_tac [!] ccontr) unfolding not_less by auto
   let ?halfl = "{z. inner (basis 1) z < dest_vec1 x} " and ?halfr = "{z. inner (basis 1) z > dest_vec1 x} "
   { fix y assume "y \<in> s" have "y \<in> ?halfr \<union> ?halfl" apply(rule ccontr)
-    using as(6) `y\<in>s` by (auto simp add: inner_vector_def dest_vec1_eq) }
+    using as(6) `y\<in>s` by (auto simp add: inner_vector_def) }
   moreover have "a\<in>?halfl" "b\<in>?halfr" using * by (auto simp add: inner_vector_def)
   hence "?halfl \<inter> s \<noteq> {}" "?halfr \<inter> s \<noteq> {}"  using as(2-3) by auto
   ultimately show False apply(rule_tac notE[OF as(1)[unfolded connected_def]])
@@ -2339,7 +2373,7 @@
   shows "\<exists>x\<in>{a..b}. (f x)$k = y"
   apply(subst neg_equal_iff_equal[THEN sym]) unfolding vector_uminus_component[THEN sym]
   apply(rule ivt_increasing_component_on_1) using assms using continuous_on_neg
-  by(auto simp add:vector_uminus_component)
+  by auto
 
 lemma ivt_decreasing_component_1: fixes f::"real^1 \<Rightarrow> real^'n"
   shows "dest_vec1 a \<le> dest_vec1 b \<Longrightarrow> \<forall>x\<in>{a .. b}. continuous (at x) f
@@ -2350,7 +2384,6 @@
 subsection {* A bound within a convex hull, and so an interval. *}
 
 lemma convex_on_convex_hull_bound:
-  fixes s :: "(real ^ _) set"
   assumes "convex_on (convex hull s) f" "\<forall>x\<in>s. f x \<le> b"
   shows "\<forall>x\<in> convex hull s. f x \<le> b" proof
   fix x assume "x\<in>convex hull s"
@@ -2381,7 +2414,7 @@
         unfolding i'(1) xi_def apply(rule_tac Min_le) unfolding image_iff
         defer apply(rule_tac x=j in bexI) using i' by auto
       have i01:"x$i \<le> 1" "x$i > 0" using Suc(2)[unfolded mem_interval,rule_format,of i] using i'(2) `x$i \<noteq> 0`
-        by(auto simp add: Cart_lambda_beta) 
+        by auto
       show ?thesis proof(cases "x$i=1")
         case True have "\<forall>j\<in>{i. x$i \<noteq> 0}. x$j = 1" apply(rule, rule ccontr) unfolding mem_Collect_eq proof-
           fix j assume "x $ j \<noteq> 0" "x $ j \<noteq> 1"
@@ -2390,21 +2423,21 @@
           hence "x$j \<ge> x$i" unfolding i'(1) xi_def apply(rule_tac Min_le) by auto
           thus False using True Suc(2) j by(auto simp add: vector_le_def elim!:ballE[where x=j]) qed
         thus "x\<in>convex hull ?points" apply(rule_tac hull_subset[unfolded subset_eq, rule_format])
-          by(auto simp add: Cart_lambda_beta)
+          by auto
       next let ?y = "\<lambda>j. if x$j = 0 then 0 else (x$j - x$i) / (1 - x$i)"
         case False hence *:"x = x$i *\<^sub>R (\<chi> j. if x$j = 0 then 0 else 1) + (1 - x$i) *\<^sub>R (\<chi> j. ?y j)" unfolding Cart_eq
-          by(auto simp add: Cart_lambda_beta vector_add_component vector_smult_component vector_minus_component field_simps)
+          by(auto simp add: field_simps)
         { fix j have "x$j \<noteq> 0 \<Longrightarrow> 0 \<le> (x $ j - x $ i) / (1 - x $ i)" "(x $ j - x $ i) / (1 - x $ i) \<le> 1"
             apply(rule_tac divide_nonneg_pos) using i(1)[of j] using False i01
-            using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps Cart_lambda_beta) 
+            using Suc(2)[unfolded mem_interval, rule_format, of j] by(auto simp add:field_simps)
           hence "0 \<le> ?y j \<and> ?y j \<le> 1" by auto }
-        moreover have "i\<in>{j. x$j \<noteq> 0} - {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" using i01 by(auto simp add: Cart_lambda_beta)
+        moreover have "i\<in>{j. x$j \<noteq> 0} - {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" using i01 by auto
         hence "{j. x$j \<noteq> 0} \<noteq> {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0}" by auto
-        hence **:"{j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by(auto simp add: Cart_lambda_beta)  
+        hence **:"{j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<subset> {j. x$j \<noteq> 0}" apply - apply rule by auto
         have "card {j. ((\<chi> j. ?y j)::real^'n) $ j \<noteq> 0} \<le> n" using less_le_trans[OF psubset_card_mono[OF _ **] Suc(4)] by auto
         ultimately show ?thesis apply(subst *) apply(rule convex_convex_hull[unfolded convex_def, rule_format])
           apply(rule_tac hull_subset[unfolded subset_eq, rule_format]) defer apply(rule Suc(1))
-          unfolding mem_interval using i01 Suc(3) by (auto simp add: Cart_lambda_beta)
+          unfolding mem_interval using i01 Suc(3) by auto
       qed qed qed } note * = this
   show ?thesis apply rule defer apply(rule hull_minimal) unfolding subset_eq prefer 3 apply rule 
     apply(rule_tac n2="CARD('n)" in *) prefer 3 apply(rule card_mono) using 01 and convex_interval(1) prefer 5 apply - apply rule
@@ -2419,7 +2452,7 @@
   prefer 3 apply(rule unit_interval_convex_hull) apply rule unfolding mem_Collect_eq proof-
   fix x::"real^'n" assume as:"\<forall>i. x $ i = 0 \<or> x $ i = 1"
   show "x \<in> (\<lambda>s. \<chi> i. if i \<in> s then 1 else 0) ` UNIV" apply(rule image_eqI[where x="{i. x$i = 1}"])
-    unfolding Cart_eq using as by(auto simp add:Cart_lambda_beta) qed auto
+    unfolding Cart_eq using as by auto qed auto
 
 subsection {* Hence any cube (could do any nonempty interval). *}
 
@@ -2430,30 +2463,30 @@
     unfolding image_iff defer apply(erule bexE) proof-
     fix y assume as:"y\<in>{x - ?d .. x + ?d}"
     { fix i::'n have "x $ i \<le> d + y $ i" "y $ i \<le> d + x $ i" using as[unfolded mem_interval, THEN spec[where x=i]]
-        by(auto simp add: vector_component)
+        by auto
       hence "1 \<ge> inverse d * (x $ i - y $ i)" "1 \<ge> inverse d * (y $ i - x $ i)"
         apply(rule_tac[!] mult_left_le_imp_le[OF _ assms]) unfolding mult_assoc[THEN sym]
-        using assms by(auto simp add: field_simps right_inverse) 
+        using assms by(auto simp add: field_simps)
       hence "inverse d * (x $ i * 2) \<le> 2 + inverse d * (y $ i * 2)"
             "inverse d * (y $ i * 2) \<le> 2 + inverse d * (x $ i * 2)" by(auto simp add:field_simps) }
     hence "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \<in> {0..1}" unfolding mem_interval using assms
-      by(auto simp add: Cart_eq vector_component_simps field_simps)
+      by(auto simp add: Cart_eq field_simps)
     thus "\<exists>z\<in>{0..1}. y = x - ?d + (2 * d) *\<^sub>R z" apply- apply(rule_tac x="inverse (2 * d) *\<^sub>R (y - (x - ?d))" in bexI) 
-      using assms by(auto simp add: Cart_eq vector_le_def Cart_lambda_beta)
+      using assms by(auto simp add: Cart_eq vector_le_def)
   next
     fix y z assume as:"z\<in>{0..1}" "y = x - ?d + (2*d) *\<^sub>R z" 
     have "\<And>i. 0 \<le> d * z $ i \<and> d * z $ i \<le> d" using assms as(1)[unfolded mem_interval] apply(erule_tac x=i in allE)
       apply rule apply(rule mult_nonneg_nonneg) prefer 3 apply(rule mult_right_le_one_le)
-      using assms by(auto simp add: vector_component_simps Cart_eq)
+      using assms by(auto simp add: Cart_eq)
     thus "y \<in> {x - ?d..x + ?d}" unfolding as(2) mem_interval apply- apply rule using as(1)[unfolded mem_interval]
-      apply(erule_tac x=i in allE) using assms by(auto simp add:  vector_component_simps Cart_eq) qed
+      apply(erule_tac x=i in allE) using assms by(auto simp add: Cart_eq) qed
   obtain s where "finite s" "{0..1::real^'n} = convex hull s" using unit_cube_convex_hull by auto
   thus ?thesis apply(rule_tac that[of "(\<lambda>y. x - ?d + (2 * d) *\<^sub>R y)` s"]) unfolding * and convex_hull_affinity by auto qed
 
 subsection {* Bounded convex function on open set is continuous. *}
 
 lemma convex_on_bounded_continuous:
-  fixes s :: "(real ^ _) set"
+  fixes s :: "('a::real_normed_vector) set"
   assumes "open s" "convex_on s f" "\<forall>x\<in>s. abs(f x) \<le> b"
   shows "continuous_on s f"
   apply(rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof(rule,rule,rule)
@@ -2503,13 +2536,18 @@
 
 subsection {* Upper bound on a ball implies upper and lower bounds. *}
 
+lemma scaleR_2:
+  fixes x :: "'a::real_vector"
+  shows "scaleR 2 x = x + x"
+unfolding one_add_one_is_two [symmetric] scaleR_left_distrib by simp
+
 lemma convex_bounds_lemma:
-  fixes x :: "real ^ _"
+  fixes x :: "'a::real_normed_vector"
   assumes "convex_on (cball x e) f"  "\<forall>y \<in> cball x e. f y \<le> b"
   shows "\<forall>y \<in> cball x e. abs(f y) \<le> b + 2 * abs(f x)"
   apply(rule) proof(cases "0 \<le> e") case True
   fix y assume y:"y\<in>cball x e" def z \<equiv> "2 *\<^sub>R x - y"
-  have *:"x - (2 *\<^sub>R x - y) = y - x" by vector
+  have *:"x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2)
   have z:"z\<in>cball x e" using y unfolding z_def mem_cball dist_norm * by(auto simp add: norm_minus_commute)
   have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp add: algebra_simps)
   thus "\<bar>f y\<bar> \<le> b + 2 * \<bar>f x\<bar>" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"]
@@ -2531,8 +2569,8 @@
   have "0 < d" unfolding d_def using `e>0` dimge1 by(rule_tac divide_pos_pos, auto) 
   let ?d = "(\<chi> i. d)::real^'n"
   obtain c where c:"finite c" "{x - ?d..x + ?d} = convex hull c" using cube_convex_hull[OF `d>0`, of x] by auto
-  have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by(auto simp add:vector_component_simps)
-  hence "c\<noteq>{}" using c by(auto simp add:convex_hull_empty)
+  have "x\<in>{x - ?d..x + ?d}" using `d>0` unfolding mem_interval by auto
+  hence "c\<noteq>{}" using c by auto
   def k \<equiv> "Max (f ` c)"
   have "convex_on {x - ?d..x + ?d} f" apply(rule convex_on_subset[OF assms(2)])
     apply(rule subset_trans[OF _ e(1)]) unfolding subset_eq mem_cball proof 
@@ -2540,7 +2578,7 @@
     have e:"e = setsum (\<lambda>i. d) (UNIV::'n set)" unfolding setsum_constant d_def using dimge1
       by (metis eq_divide_imp mult_frac_num real_dimindex_gt_0 real_eq_of_nat real_less_def real_mult_commute)
     show "dist x z \<le> e" unfolding dist_norm e apply(rule_tac order_trans[OF norm_le_l1], rule setsum_mono)
-      using z[unfolded mem_interval] apply(erule_tac x=i in allE) by(auto simp add:field_simps vector_component_simps) qed
+      using z[unfolded mem_interval] apply(erule_tac x=i in allE) by auto qed
   hence k:"\<forall>y\<in>{x - ?d..x + ?d}. f y \<le> k" unfolding c(2) apply(rule_tac convex_on_convex_hull_bound) apply assumption
     unfolding k_def apply(rule, rule Max_ge) using c(1) by auto
   have "d \<le> e" unfolding d_def apply(rule mult_imp_div_pos_le) using `e>0` dimge1 unfolding mult_le_cancel_left1 using real_dimindex_ge_1 by auto
@@ -2549,9 +2587,9 @@
   hence "\<forall>y\<in>cball x d. abs (f y) \<le> k + 2 * abs (f x)" apply(rule_tac convex_bounds_lemma) apply assumption proof
     fix y assume y:"y\<in>cball x d"
     { fix i::'n have "x $ i - d \<le> y $ i"  "y $ i \<le> x $ i + d" 
-        using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by(auto simp add: vector_component)  }
+        using order_trans[OF component_le_norm y[unfolded mem_cball dist_norm], of i] by auto  }
     thus "f y \<le> k" apply(rule_tac k[rule_format]) unfolding mem_cball mem_interval dist_norm 
-      by(auto simp add: vector_component_simps) qed
+      by auto qed
   hence "continuous_on (ball x d) f" apply(rule_tac convex_on_bounded_continuous)
     apply(rule open_ball, rule convex_on_subset[OF conv], rule ball_subset_cball)
     apply force
@@ -2566,15 +2604,15 @@
    segment[a,b] is closed and segment(a,b) is open relative to affine hull. *)
 
 definition
-  midpoint :: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 'n" where
+  midpoint :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a" where
   "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)"
 
 definition
-  open_segment :: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> (real ^ 'n) set" where
+  open_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
   "open_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real.  0 < u \<and> u < 1}"
 
 definition
-  closed_segment :: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> (real ^ 'n) set" where
+  closed_segment :: "'a::real_vector \<Rightarrow> 'a \<Rightarrow> 'a set" where
   "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \<le> u \<and> u \<le> 1}"
 
 definition "between = (\<lambda> (a,b). closed_segment a b)"
@@ -2588,24 +2626,38 @@
 
 lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by (auto simp add: scaleR_right_distrib)
 
+lemma midpoint_eq_iff: "midpoint a b = c \<longleftrightarrow> a + b = c + c"
+proof -
+  have "midpoint a b = c \<longleftrightarrow> scaleR 2 (midpoint a b) = scaleR 2 c"
+    by simp
+  thus ?thesis
+    unfolding midpoint_def scaleR_2 [symmetric] by simp
+qed
+
 lemma dist_midpoint:
+  fixes a b :: "'a::real_normed_vector" shows
   "dist a (midpoint a b) = (dist a b) / 2" (is ?t1)
   "dist b (midpoint a b) = (dist a b) / 2" (is ?t2)
   "dist (midpoint a b) a = (dist a b) / 2" (is ?t3)
   "dist (midpoint a b) b = (dist a b) / 2" (is ?t4)
 proof-
-  have *: "\<And>x y::real^'n. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2" unfolding equation_minus_iff by auto
-  have **:"\<And>x y::real^'n. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2" by auto
+  have *: "\<And>x y::'a. 2 *\<^sub>R x = - y \<Longrightarrow> norm x = (norm y) / 2" unfolding equation_minus_iff by auto
+  have **:"\<And>x y::'a. 2 *\<^sub>R x =   y \<Longrightarrow> norm x = (norm y) / 2" by auto
   note scaleR_right_distrib [simp]
-  show ?t1 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector)
-  show ?t2 unfolding midpoint_def dist_norm apply (rule *)  by(auto,vector)
-  show ?t3 unfolding midpoint_def dist_norm apply (rule *)  by(auto,vector)
-  show ?t4 unfolding midpoint_def dist_norm apply (rule **) by(auto,vector) qed
+  show ?t1 unfolding midpoint_def dist_norm apply (rule **)
+    by (simp add: scaleR_right_diff_distrib, simp add: scaleR_2)
+  show ?t2 unfolding midpoint_def dist_norm apply (rule *)
+    by (simp add: scaleR_right_diff_distrib, simp add: scaleR_2)
+  show ?t3 unfolding midpoint_def dist_norm apply (rule *)
+    by (simp add: scaleR_right_diff_distrib, simp add: scaleR_2)
+  show ?t4 unfolding midpoint_def dist_norm apply (rule **)
+    by (simp add: scaleR_right_diff_distrib, simp add: scaleR_2)
+qed
 
 lemma midpoint_eq_endpoint:
-  "midpoint a b = a \<longleftrightarrow> a = (b::real^'n)"
+  "midpoint a b = a \<longleftrightarrow> a = b"
   "midpoint a b = b \<longleftrightarrow> a = b"
-  unfolding dist_eq_0_iff[where 'a="real^'n", THEN sym] dist_midpoint by auto
+  unfolding midpoint_eq_iff by auto
 
 lemma convex_contains_segment:
   "convex s \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. closed_segment a b \<subseteq> s)"
@@ -2631,12 +2683,14 @@
   unfolding segment_convex_hull apply(rule_tac[!] hull_subset[unfolded subset_eq, rule_format]) by auto
 
 lemma segment_furthest_le:
+  fixes a b x y :: "real ^ 'n"
   assumes "x \<in> closed_segment a b" shows "norm(y - x) \<le> norm(y - a) \<or>  norm(y - x) \<le> norm(y - b)" proof-
   obtain z where "z\<in>{a, b}" "norm (x - y) \<le> norm (z - y)" using simplex_furthest_le[of "{a, b}" y]
     using assms[unfolded segment_convex_hull] by auto
   thus ?thesis by(auto simp add:norm_minus_commute) qed
 
 lemma segment_bound:
+  fixes x a b :: "real ^ 'n"
   assumes "x \<in> closed_segment a b"
   shows "norm(x - a) \<le> norm(b - a)" "norm(x - b) \<le> norm(b - a)"
   using segment_furthest_le[OF assms, of a]
@@ -2661,17 +2715,17 @@
         unfolding as(1) by(auto simp add:algebra_simps)
       show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)"
         unfolding norm_minus_commute[of x a] * Cart_eq using as(2,3)
-        by(auto simp add: vector_component_simps field_simps)
+        by(auto simp add: field_simps)
     next assume as:"dist a b = dist a x + dist x b"
       have "norm (a - x) / norm (a - b) \<le> 1" unfolding divide_le_eq_1_pos[OF Fal2] unfolding as[unfolded dist_norm] norm_ge_zero by auto 
       thus "\<exists>u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \<and> 0 \<le> u \<and> u \<le> 1" apply(rule_tac x="dist a x / dist a b" in exI)
         unfolding dist_norm Cart_eq apply- apply rule defer apply(rule, rule divide_nonneg_pos) prefer 4 proof rule
           fix i::'n have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i =
             ((norm (a - b) - norm (a - x)) * (a $ i) + norm (a - x) * (b $ i)) / norm (a - b)"
-            using Fal by(auto simp add:vector_component_simps field_simps)
+            using Fal by(auto simp add: field_simps)
           also have "\<dots> = x$i" apply(rule divide_eq_imp[OF Fal])
             unfolding as[unfolded dist_norm] using as[unfolded dist_triangle_eq Cart_eq,rule_format, of i]
-            by(auto simp add:field_simps vector_component_simps)
+            by(auto simp add:field_simps)
           finally show "x $ i = ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) $ i" by auto
         qed(insert Fal2, auto) qed qed
 
@@ -2680,7 +2734,7 @@
   "between (b,a) (midpoint a b)" (is ?t2)
 proof- have *:"\<And>x y z. x = (1/2::real) *\<^sub>R z \<Longrightarrow> y = (1/2) *\<^sub>R z \<Longrightarrow> norm z = norm x + norm y" by auto
   show ?t1 ?t2 unfolding between midpoint_def dist_norm apply(rule_tac[!] *)
-    by(auto simp add:field_simps Cart_eq vector_component_simps) qed
+    by(auto simp add:field_simps Cart_eq) qed
 
 lemma between_mem_convex_hull:
   "between (a,b) x \<longleftrightarrow> x \<in> convex hull {a,b}"
@@ -2699,7 +2753,7 @@
     have *:"y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using `e>0` by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib)
     have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = abs(1/e) * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)"
       unfolding dist_norm unfolding norm_scaleR[THEN sym] apply(rule norm_eqI) using `e>0`
-      by(auto simp add:vector_component_simps Cart_eq field_simps) 
+      by(auto simp add: Cart_eq field_simps) 
     also have "\<dots> = abs(1/e) * norm (x - e *\<^sub>R (x - c) - y)" by(auto intro!:norm_eqI simp add: algebra_simps)
     also have "\<dots> < d" using as[unfolded dist_norm] and `e>0`
       by(auto simp add:pos_divide_less_eq[OF `e>0`] real_mult_commute)
@@ -2771,11 +2825,11 @@
   fix x::"real^'n" and e assume "0<e" and as:"\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x. 0 \<le> xa $ x) \<and> setsum (op $ xa) UNIV \<le> 1"
   show "(\<forall>xa. 0 < x $ xa) \<and> setsum (op $ x) UNIV < 1" apply(rule,rule) proof-
     fix i::'n show "0 < x $ i" using as[THEN spec[where x="x - (e / 2) *\<^sub>R basis i"]] and `e>0`
-      unfolding dist_norm by(auto simp add: norm_basis vector_component_simps basis_component elim:allE[where x=i])
+      unfolding dist_norm by(auto simp add: norm_basis elim:allE[where x=i])
   next guess a using UNIV_witness[where 'a='n] ..
     have **:"dist x (x + (e / 2) *\<^sub>R basis a) < e" using  `e>0` and norm_basis[of a]
-      unfolding dist_norm by(auto simp add: vector_component_simps basis_component intro!: mult_strict_left_mono_comm)
-    have "\<And>i. (x + (e / 2) *\<^sub>R basis a) $ i = x$i + (if i = a then e/2 else 0)" by(auto simp add:vector_component_simps)
+      unfolding dist_norm by(auto intro!: mult_strict_left_mono_comm)
+    have "\<And>i. (x + (e / 2) *\<^sub>R basis a) $ i = x$i + (if i = a then e/2 else 0)" by auto
     hence *:"setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV = setsum (\<lambda>i. x$i + (if a = i then e/2 else 0)) UNIV" by(rule_tac setsum_cong, auto) 
     have "setsum (op $ x) UNIV < setsum (op $ (x + (e / 2) *\<^sub>R basis a)) UNIV" unfolding * setsum_addf
       using `0<e` dimindex_ge_1 by(auto simp add: setsum_delta')
@@ -2792,13 +2846,13 @@
     fix y assume y:"dist x y < min (Min (op $ x ` UNIV)) ?d"
     have "setsum (op $ y) UNIV \<le> setsum (\<lambda>i. x$i + ?d) UNIV" proof(rule setsum_mono)
       fix i::'n have "abs (y$i - x$i) < ?d" apply(rule le_less_trans) using component_le_norm[of "y - x" i]
-        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add:vector_component_simps norm_minus_commute)
+        using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] by(auto simp add: norm_minus_commute)
       thus "y $ i \<le> x $ i + ?d" by auto qed
     also have "\<dots> \<le> 1" unfolding setsum_addf setsum_constant card_enum real_eq_of_nat using dimindex_ge_1 by(auto simp add: Suc_le_eq)
     finally show "(\<forall>i. 0 \<le> y $ i) \<and> setsum (op $ y) UNIV \<le> 1" apply- proof(rule,rule)
       fix i::'n have "norm (x - y) < x$i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]
-        using Min_gr_iff[of "op $ x ` dimset x"] dimindex_ge_1 by auto
-      thus "0 \<le> y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by(auto simp add: vector_component_simps)
+        by auto
+      thus "0 \<le> y$i" using component_le_norm[of "x - y" i] and as(1)[rule_format, of i] by auto
     qed auto qed auto qed
 
 lemma interior_std_simplex_nonempty: obtains a::"real^'n" where
@@ -2818,23 +2872,42 @@
 
 subsection {* Paths. *}
 
-definition "path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow> continuous_on {0 .. 1} g"
-
-definition "pathstart (g::real^1 \<Rightarrow> real^'n) = g 0"
-
-definition "pathfinish (g::real^1 \<Rightarrow> real^'n) = g 1"
-
-definition "path_image (g::real^1 \<Rightarrow> real^'n) = g ` {0 .. 1}"
-
-definition "reversepath (g::real^1 \<Rightarrow> real^'n) = (\<lambda>x. g(1 - x))"
-
-definition joinpaths:: "(real^1 \<Rightarrow> real^'n) \<Rightarrow> (real^1 \<Rightarrow> real^'n) \<Rightarrow> (real^1 \<Rightarrow> real^'n)" (infixr "+++" 75)
-  where "joinpaths g1 g2 = (\<lambda>x. if dest_vec1 x \<le> ((1 / 2)::real) then g1 (2 *\<^sub>R x) else g2(2 *\<^sub>R x - 1))"
-definition "simple_path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow>
+text {* TODO: Once @{const continuous_on} is generalized to arbitrary
+topological spaces, all of these concepts should be similarly generalized. *}
+
+definition
+  path :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> bool"
+  where "path g \<longleftrightarrow> continuous_on {0 .. 1} g"
+
+definition
+  pathstart :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> 'a"
+  where "pathstart g = g 0"
+
+definition
+  pathfinish :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> 'a"
+  where "pathfinish g = g 1"
+
+definition
+  path_image :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> 'a set"
+  where "path_image g = g ` {0 .. 1}"
+
+definition
+  reversepath :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> (real \<Rightarrow> 'a)"
+  where "reversepath g = (\<lambda>x. g(1 - x))"
+
+definition
+  joinpaths :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a)"
+    (infixr "+++" 75)
+  where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
+
+definition
+  simple_path :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> bool"
+  where "simple_path g \<longleftrightarrow>
   (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
 
-definition "injective_path (g::real^1 \<Rightarrow> real^'n) \<longleftrightarrow>
-  (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y)"
+definition
+  injective_path :: "(real \<Rightarrow> 'a::metric_space) \<Rightarrow> bool"
+  where "injective_path g \<longleftrightarrow> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y)"
 
 subsection {* Some lemmas about these concepts. *}
 
@@ -2846,20 +2919,19 @@
   unfolding path_image_def image_is_empty interval_eq_empty by auto 
 
 lemma pathstart_in_path_image[intro]: "(pathstart g) \<in> path_image g"
-  unfolding pathstart_def path_image_def apply(rule imageI)
-  unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto
+  unfolding pathstart_def path_image_def by auto
 
 lemma pathfinish_in_path_image[intro]: "(pathfinish g) \<in> path_image g"
-  unfolding pathfinish_def path_image_def apply(rule imageI)
-  unfolding mem_interval_1 vec_1[THEN sym] dest_vec1_vec by auto
+  unfolding pathfinish_def path_image_def by auto
 
 lemma connected_path_image[intro]: "path g \<Longrightarrow> connected(path_image g)"
-  unfolding path_def path_image_def apply(rule connected_continuous_image, assumption)
-  by(rule convex_connected, rule convex_interval)
+  unfolding path_def path_image_def
+  apply (erule connected_continuous_image)
+  by(rule convex_connected, rule convex_real_interval)
 
 lemma compact_path_image[intro]: "path g \<Longrightarrow> compact(path_image g)"
-  unfolding path_def path_image_def apply(rule compact_continuous_image, assumption)
-  by(rule compact_interval)
+  unfolding path_def path_image_def
+  by (erule compact_continuous_image, rule compact_real_interval)
 
 lemma reversepath_reversepath[simp]: "reversepath(reversepath g) = g"
   unfolding reversepath_def by auto
@@ -2873,15 +2945,13 @@
 lemma pathstart_join[simp]: "pathstart(g1 +++ g2) = pathstart g1"
   unfolding pathstart_def joinpaths_def pathfinish_def by auto
 
-lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2" proof-
-  have "2 *\<^sub>R 1 - 1 = (1::real^1)" unfolding Cart_eq by(auto simp add:vector_component_simps)
-  thus ?thesis unfolding pathstart_def joinpaths_def pathfinish_def
-    unfolding vec_1[THEN sym] dest_vec1_vec by auto qed
+lemma pathfinish_join[simp]:"pathfinish(g1 +++ g2) = pathfinish g2"
+  unfolding pathstart_def joinpaths_def pathfinish_def by auto
 
 lemma path_image_reversepath[simp]: "path_image(reversepath g) = path_image g" proof-
   have *:"\<And>g. path_image(reversepath g) \<subseteq> path_image g"
     unfolding path_image_def subset_eq reversepath_def Ball_def image_iff apply(rule,rule,erule bexE)  
-    apply(rule_tac x="1 - xa" in bexI) by(auto simp add:vector_le_def vector_component_simps elim!:ballE)
+    apply(rule_tac x="1 - xa" in bexI) by auto
   show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
 
 lemma path_reversepath[simp]: "path(reversepath g) \<longleftrightarrow> path g" proof-
@@ -2889,7 +2959,7 @@
     apply(rule continuous_on_compose[unfolded o_def, of _ "\<lambda>x. 1 - x"])
     apply(rule continuous_on_sub, rule continuous_on_const, rule continuous_on_id)
     apply(rule continuous_on_subset[of "{0..1}"], assumption) by auto
-  show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed
+  show ?thesis using *[of "reversepath g"] *[of g] unfolding reversepath_reversepath by (rule iffI) qed
 
 lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
 
@@ -2897,48 +2967,50 @@
   unfolding path_def pathfinish_def pathstart_def apply rule defer apply(erule conjE) proof-
   assume as:"continuous_on {0..1} (g1 +++ g2)"
   have *:"g1 = (\<lambda>x. g1 (2 *\<^sub>R x)) \<circ> (\<lambda>x. (1/2) *\<^sub>R x)" 
-         "g2 = (\<lambda>x. g2 (2 *\<^sub>R x - 1)) \<circ> (\<lambda>x. (1/2) *\<^sub>R (x + 1))" unfolding o_def by auto
-  have "op *\<^sub>R (1 / 2) ` {0::real^1..1} \<subseteq> {0..1}"  "(\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {(0::real^1)..1} \<subseteq> {0..1}"
-    unfolding image_smult_interval by auto 
+         "g2 = (\<lambda>x. g2 (2 *\<^sub>R x - 1)) \<circ> (\<lambda>x. (1/2) *\<^sub>R (x + 1))"
+    unfolding o_def by (auto simp add: add_divide_distrib)
+  have "op *\<^sub>R (1 / 2) ` {0::real..1} \<subseteq> {0..1}"  "(\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {(0::real)..1} \<subseteq> {0..1}"
+    by auto
   thus "continuous_on {0..1} g1 \<and> continuous_on {0..1} g2" apply -apply rule
     apply(subst *) defer apply(subst *) apply (rule_tac[!] continuous_on_compose)
     apply (rule continuous_on_cmul, rule continuous_on_add, rule continuous_on_id, rule continuous_on_const) defer
     apply (rule continuous_on_cmul, rule continuous_on_id) apply(rule_tac[!] continuous_on_eq[of _ "g1 +++ g2"]) defer prefer 3
     apply(rule_tac[1-2] continuous_on_subset[of "{0 .. 1}"]) apply(rule as, assumption, rule as, assumption)
     apply(rule) defer apply rule proof-
-    fix x assume "x \<in> op *\<^sub>R (1 / 2) ` {0::real^1..1}"
-    hence "dest_vec1 x \<le> 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps)
+    fix x assume "x \<in> op *\<^sub>R (1 / 2) ` {0::real..1}"
+    hence "x \<le> 1 / 2" unfolding image_iff by auto
     thus "(g1 +++ g2) x = g1 (2 *\<^sub>R x)" unfolding joinpaths_def by auto next
-    fix x assume "x \<in> (\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {0::real^1..1}"
-    hence "dest_vec1 x \<ge> 1 / 2" unfolding image_iff by(auto simp add: vector_component_simps)
-    thus "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" proof(cases "dest_vec1 x = 1 / 2")
-      case True hence "x = (1/2) *\<^sub>R 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps)
-      thus ?thesis unfolding joinpaths_def using assms[unfolded pathstart_def pathfinish_def] by auto
+    fix x assume "x \<in> (\<lambda>x. (1 / 2) *\<^sub>R (x + 1)) ` {0::real..1}"
+    hence "x \<ge> 1 / 2" unfolding image_iff by auto
+    thus "(g1 +++ g2) x = g2 (2 *\<^sub>R x - 1)" proof(cases "x = 1 / 2")
+      case True hence "x = (1/2) *\<^sub>R 1" unfolding Cart_eq by auto
+      thus ?thesis unfolding joinpaths_def using assms[unfolded pathstart_def pathfinish_def] by (auto simp add: mult_ac)
     qed (auto simp add:le_less joinpaths_def) qed
 next assume as:"continuous_on {0..1} g1" "continuous_on {0..1} g2"
-  have *:"{0 .. 1::real^1} = {0.. (1/2)*\<^sub>R 1} \<union> {(1/2) *\<^sub>R 1 .. 1}" by(auto simp add: vector_component_simps) 
-  have **:"op *\<^sub>R 2 ` {0..(1 / 2) *\<^sub>R 1} = {0..1::real^1}" apply(rule set_ext, rule) unfolding image_iff 
-    defer apply(rule_tac x="(1/2)*\<^sub>R x" in bexI) by(auto simp add: vector_component_simps)
-  have ***:"(\<lambda>x. 2 *\<^sub>R x - 1) ` {(1 / 2) *\<^sub>R 1..1} = {0..1::real^1}"
-    unfolding image_affinity_interval[of _ "- 1", unfolded diff_def[symmetric]] and interval_eq_empty_1
-    by(auto simp add: vector_component_simps)
-  have ****:"\<And>x::real^1. x $ 1 * 2 = 1 \<longleftrightarrow> x = (1/2) *\<^sub>R 1" unfolding Cart_eq by(auto simp add: forall_1 vector_component_simps)
-  show "continuous_on {0..1} (g1 +++ g2)" unfolding * apply(rule continuous_on_union) apply(rule closed_interval)+ proof-
+  have *:"{0 .. 1::real} = {0.. (1/2)*\<^sub>R 1} \<union> {(1/2) *\<^sub>R 1 .. 1}" by auto
+  have **:"op *\<^sub>R 2 ` {0..(1 / 2) *\<^sub>R 1} = {0..1::real}" apply(rule set_ext, rule) unfolding image_iff 
+    defer apply(rule_tac x="(1/2)*\<^sub>R x" in bexI) by auto
+  have ***:"(\<lambda>x. 2 *\<^sub>R x - 1) ` {(1 / 2) *\<^sub>R 1..1} = {0..1::real}"
+    apply (auto simp add: image_def)
+    apply (rule_tac x="(x + 1) / 2" in bexI)
+    apply (auto simp add: add_divide_distrib)
+    done
+  show "continuous_on {0..1} (g1 +++ g2)" unfolding * apply(rule continuous_on_union) apply (rule closed_real_atLeastAtMost)+ proof-
     show "continuous_on {0..(1 / 2) *\<^sub>R 1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "\<lambda>x. g1 (2 *\<^sub>R x)"]) defer
       unfolding o_def[THEN sym] apply(rule continuous_on_compose) apply(rule continuous_on_cmul, rule continuous_on_id)
-      unfolding ** apply(rule as(1)) unfolding joinpaths_def by(auto simp add: vector_component_simps) next
+      unfolding ** apply(rule as(1)) unfolding joinpaths_def by auto next
     show "continuous_on {(1/2)*\<^sub>R1..1} (g1 +++ g2)" apply(rule continuous_on_eq[of _ "g2 \<circ> (\<lambda>x. 2 *\<^sub>R x - 1)"]) defer
       apply(rule continuous_on_compose) apply(rule continuous_on_sub, rule continuous_on_cmul, rule continuous_on_id, rule continuous_on_const)
       unfolding *** o_def joinpaths_def apply(rule as(2)) using assms[unfolded pathstart_def pathfinish_def]
-      by(auto simp add: vector_component_simps ****) qed qed
+      by (auto simp add: mult_ac) qed qed
 
 lemma path_image_join_subset: "path_image(g1 +++ g2) \<subseteq> (path_image g1 \<union> path_image g2)" proof
   fix x assume "x \<in> path_image (g1 +++ g2)"
-  then obtain y where y:"y\<in>{0..1}" "x = (if dest_vec1 y \<le> 1 / 2 then g1 (2 *\<^sub>R y) else g2 (2 *\<^sub>R y - 1))"
+  then obtain y where y:"y\<in>{0..1}" "x = (if y \<le> 1 / 2 then g1 (2 *\<^sub>R y) else g2 (2 *\<^sub>R y - 1))"
     unfolding path_image_def image_iff joinpaths_def by auto
-  thus "x \<in> path_image g1 \<union> path_image g2" apply(cases "dest_vec1 y \<le> 1/2")
+  thus "x \<in> path_image g1 \<union> path_image g2" apply(cases "y \<le> 1/2")
     apply(rule_tac UnI1) defer apply(rule_tac UnI2) unfolding y(2) path_image_def using y(1)
-    by(auto intro!: imageI simp add: vector_component_simps) qed
+    by(auto intro!: imageI) qed
 
 lemma subset_path_image_join:
   assumes "path_image g1 \<subseteq> s" "path_image g2 \<subseteq> s" shows "path_image(g1 +++ g2) \<subseteq> s"
@@ -2951,13 +3023,12 @@
   fix x assume "x \<in> path_image g1"
   then obtain y where y:"y\<in>{0..1}" "x = g1 y" unfolding path_image_def image_iff by auto
   thus "x \<in> path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff
-    apply(rule_tac x="(1/2) *\<^sub>R y" in bexI) by(auto simp add: vector_component_simps) next
+    apply(rule_tac x="(1/2) *\<^sub>R y" in bexI) by auto next
   fix x assume "x \<in> path_image g2"
   then obtain y where y:"y\<in>{0..1}" "x = g2 y" unfolding path_image_def image_iff by auto
-  moreover have *:"y $ 1 = 0 \<Longrightarrow> y = 0" unfolding Cart_eq by auto
-  ultimately show "x \<in> path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff
+  then show "x \<in> path_image (g1 +++ g2)" unfolding joinpaths_def path_image_def image_iff
     apply(rule_tac x="(1/2) *\<^sub>R (y + 1)" in bexI) using assms(3)[unfolded pathfinish_def pathstart_def]
-    by(auto simp add: vector_component_simps) qed 
+    by (auto simp add: add_divide_distrib) qed
 
 lemma not_in_path_image_join:
   assumes "x \<notin> path_image g1" "x \<notin> path_image g2" shows "x \<notin> path_image(g1 +++ g2)"
@@ -2966,10 +3037,7 @@
 lemma simple_path_reversepath: assumes "simple_path g" shows "simple_path (reversepath g)"
   using assms unfolding simple_path_def reversepath_def apply- apply(rule ballI)+
   apply(erule_tac x="1-x" in ballE, erule_tac x="1-y" in ballE)
-  unfolding mem_interval_1 by(auto simp add:vector_component_simps)
-
-(** move this **)
-declare vector_scaleR_component[simp]
+  unfolding mem_interval_1 by auto
 
 lemma simple_path_join_loop:
   assumes "injective_path g1" "injective_path g2" "pathfinish g2 = pathstart g1"
@@ -2977,42 +3045,42 @@
   shows "simple_path(g1 +++ g2)"
 unfolding simple_path_def proof((rule ballI)+, rule impI) let ?g = "g1 +++ g2"
   note inj = assms(1,2)[unfolded injective_path_def, rule_format]
-  fix x y::"real^1" assume xy:"x \<in> {0..1}" "y \<in> {0..1}" "?g x = ?g y"
-  show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0" proof(case_tac "x$1 \<le> 1/2",case_tac[!] "y$1 \<le> 1/2", unfold not_le)
-    assume as:"x $ 1 \<le> 1 / 2" "y $ 1 \<le> 1 / 2"
+  fix x y::"real" assume xy:"x \<in> {0..1}" "y \<in> {0..1}" "?g x = ?g y"
+  show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0" proof(case_tac "x \<le> 1/2",case_tac[!] "y \<le> 1/2", unfold not_le)
+    assume as:"x \<le> 1 / 2" "y \<le> 1 / 2"
     hence "g1 (2 *\<^sub>R x) = g1 (2 *\<^sub>R y)" using xy(3) unfolding joinpaths_def by auto
     moreover have "2 *\<^sub>R x \<in> {0..1}" "2 *\<^sub>R y \<in> {0..1}" using xy(1,2) as
-      unfolding mem_interval_1 by(auto simp add:vector_component_simps)
+      unfolding mem_interval_1 by auto
     ultimately show ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] by auto
-  next assume as:"x $ 1 > 1 / 2" "y $ 1 > 1 / 2"
+  next assume as:"x > 1 / 2" "y > 1 / 2"
     hence "g2 (2 *\<^sub>R x - 1) = g2 (2 *\<^sub>R y - 1)" using xy(3) unfolding joinpaths_def by auto
-    moreover have "2 *\<^sub>R x - 1 \<in> {0..1}" "2 *\<^sub>R y - 1 \<in> {0..1}" using xy(1,2) as unfolding mem_interval_1 by(auto simp add:vector_component_simps)
+    moreover have "2 *\<^sub>R x - 1 \<in> {0..1}" "2 *\<^sub>R y - 1 \<in> {0..1}" using xy(1,2) as unfolding mem_interval_1 by auto
     ultimately show ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by auto
-  next assume as:"x $ 1 \<le> 1 / 2" "y $ 1 > 1 / 2"
+  next assume as:"x \<le> 1 / 2" "y > 1 / 2"
     hence "?g x \<in> path_image g1" "?g y \<in> path_image g2" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+      using xy(1,2)[unfolded mem_interval_1] by auto
     moreover have "?g y \<noteq> pathstart g2" using as(2) unfolding pathstart_def joinpaths_def
       using inj(2)[of "2 *\<^sub>R y - 1" 0] and xy(2)[unfolded mem_interval_1]
-      by(auto simp add:vector_component_simps field_simps Cart_eq)
+      by (auto simp add: field_simps)
     ultimately have *:"?g x = pathstart g1" using assms(4) unfolding xy(3) by auto
     hence "x = 0" unfolding pathstart_def joinpaths_def using as(1) and xy(1)[unfolded mem_interval_1]
-      using inj(1)[of "2 *\<^sub>R x" 0] by(auto simp add:vector_component_simps)
+      using inj(1)[of "2 *\<^sub>R x" 0] by auto
     moreover have "y = 1" using * unfolding xy(3) assms(3)[THEN sym]
       unfolding joinpaths_def pathfinish_def using as(2) and xy(2)[unfolded mem_interval_1]
-      using inj(2)[of "2 *\<^sub>R y - 1" 1] by (auto simp add:vector_component_simps Cart_eq)
-    ultimately show ?thesis by auto 
-  next assume as:"x $ 1 > 1 / 2" "y $ 1 \<le> 1 / 2"
+      using inj(2)[of "2 *\<^sub>R y - 1" 1] by auto
+    ultimately show ?thesis by auto
+  next assume as:"x > 1 / 2" "y \<le> 1 / 2"
     hence "?g x \<in> path_image g2" "?g y \<in> path_image g1" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+      using xy(1,2)[unfolded mem_interval_1] by auto
     moreover have "?g x \<noteq> pathstart g2" using as(1) unfolding pathstart_def joinpaths_def
       using inj(2)[of "2 *\<^sub>R x - 1" 0] and xy(1)[unfolded mem_interval_1]
-      by(auto simp add:vector_component_simps field_simps Cart_eq)
+      by (auto simp add: field_simps)
     ultimately have *:"?g y = pathstart g1" using assms(4) unfolding xy(3) by auto
     hence "y = 0" unfolding pathstart_def joinpaths_def using as(2) and xy(2)[unfolded mem_interval_1]
-      using inj(1)[of "2 *\<^sub>R y" 0] by(auto simp add:vector_component_simps)
+      using inj(1)[of "2 *\<^sub>R y" 0] by auto
     moreover have "x = 1" using * unfolding xy(3)[THEN sym] assms(3)[THEN sym]
       unfolding joinpaths_def pathfinish_def using as(1) and xy(1)[unfolded mem_interval_1]
-      using inj(2)[of "2 *\<^sub>R x - 1" 1] by(auto simp add:vector_component_simps Cart_eq)
+      using inj(2)[of "2 *\<^sub>R x - 1" 1] by auto
     ultimately show ?thesis by auto qed qed
 
 lemma injective_path_join:
@@ -3021,45 +3089,41 @@
   shows "injective_path(g1 +++ g2)"
   unfolding injective_path_def proof(rule,rule,rule) let ?g = "g1 +++ g2"
   note inj = assms(1,2)[unfolded injective_path_def, rule_format]
-  have *:"\<And>x y::real^1. 2 *\<^sub>R x = 1 \<Longrightarrow> 2 *\<^sub>R y = 1 \<Longrightarrow> x = y" unfolding Cart_eq forall_1 by(auto simp del:dest_vec1_eq)
   fix x y assume xy:"x \<in> {0..1}" "y \<in> {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y"
-  show "x = y" proof(cases "x$1 \<le> 1/2", case_tac[!] "y$1 \<le> 1/2", unfold not_le)
-    assume "x $ 1 \<le> 1 / 2" "y $ 1 \<le> 1 / 2" thus ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
-      unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps)
-  next assume "x $ 1 > 1 / 2" "y $ 1 > 1 / 2" thus ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
-      unfolding mem_interval_1 joinpaths_def by(auto simp add:vector_component_simps)
-  next assume as:"x $ 1 \<le> 1 / 2" "y $ 1 > 1 / 2" 
+  show "x = y" proof(cases "x \<le> 1/2", case_tac[!] "y \<le> 1/2", unfold not_le)
+    assume "x \<le> 1 / 2" "y \<le> 1 / 2" thus ?thesis using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
+      unfolding mem_interval_1 joinpaths_def by auto
+  next assume "x > 1 / 2" "y > 1 / 2" thus ?thesis using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
+      unfolding mem_interval_1 joinpaths_def by auto
+  next assume as:"x \<le> 1 / 2" "y > 1 / 2" 
     hence "?g x \<in> path_image g1" "?g y \<in> path_image g2" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+      using xy(1,2)[unfolded mem_interval_1] by auto
     hence "?g x = pathfinish g1" "?g y = pathstart g2" using assms(4) unfolding assms(3) xy(3) by auto
     thus ?thesis using as and inj(1)[of "2 *\<^sub>R x" 1] inj(2)[of "2 *\<^sub>R y - 1" 0] and xy(1,2)
       unfolding pathstart_def pathfinish_def joinpaths_def mem_interval_1
-      by(auto simp add:vector_component_simps intro:*)
-  next assume as:"x $ 1 > 1 / 2" "y $ 1 \<le> 1 / 2" 
+      by auto
+  next assume as:"x > 1 / 2" "y \<le> 1 / 2" 
     hence "?g x \<in> path_image g2" "?g y \<in> path_image g1" unfolding path_image_def joinpaths_def
-      using xy(1,2)[unfolded mem_interval_1] by(auto simp add:vector_component_simps intro!: imageI)
+      using xy(1,2)[unfolded mem_interval_1] by auto
     hence "?g x = pathstart g2" "?g y = pathfinish g1" using assms(4) unfolding assms(3) xy(3) by auto
     thus ?thesis using as and inj(2)[of "2 *\<^sub>R x - 1" 0] inj(1)[of "2 *\<^sub>R y" 1] and xy(1,2)
       unfolding pathstart_def pathfinish_def joinpaths_def mem_interval_1
-      by(auto simp add:vector_component_simps intro:*) qed qed
+      by auto qed qed
 
 lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join
  
 subsection {* Reparametrizing a closed curve to start at some chosen point. *}
 
-definition "shiftpath a (f::real^1 \<Rightarrow> real^'n) =
-  (\<lambda>x. if dest_vec1 (a + x) \<le> 1 then f(a + x) else f(a + x - 1))"
+definition "shiftpath a (f::real \<Rightarrow> 'a::metric_space) =
+  (\<lambda>x. if (a + x) \<le> 1 then f(a + x) else f(a + x - 1))"
 
 lemma pathstart_shiftpath: "a \<le> 1 \<Longrightarrow> pathstart(shiftpath a g) = g a"
   unfolding pathstart_def shiftpath_def by auto
 
-(** move this **)
-declare forall_1[simp] ex_1[simp]
-
 lemma pathfinish_shiftpath: assumes "0 \<le> a" "pathfinish g = pathstart g"
   shows "pathfinish(shiftpath a g) = g a"
   using assms unfolding pathstart_def pathfinish_def shiftpath_def
-  by(auto simp add: vector_component_simps)
+  by auto
 
 lemma endpoints_shiftpath:
   assumes "pathfinish g = pathstart g" "a \<in> {0 .. 1}" 
@@ -3074,39 +3138,38 @@
 lemma path_shiftpath:
   assumes "path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
   shows "path(shiftpath a g)" proof-
-  have *:"{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}" using assms(3) by(auto simp add: vector_component_simps)
+  have *:"{0 .. 1} = {0 .. 1-a} \<union> {1-a .. 1}" using assms(3) by auto
   have **:"\<And>x. x + a = 1 \<Longrightarrow> g (x + a - 1) = g (x + a)"
     using assms(2)[unfolded pathfinish_def pathstart_def] by auto
   show ?thesis unfolding path_def shiftpath_def * apply(rule continuous_on_union)
-    apply(rule closed_interval)+ apply(rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a + x)"]) prefer 3
+    apply(rule closed_real_atLeastAtMost)+ apply(rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a + x)"]) prefer 3
     apply(rule continuous_on_eq[of _ "g \<circ> (\<lambda>x. a - 1 + x)"]) defer prefer 3
     apply(rule continuous_on_intros)+ prefer 2 apply(rule continuous_on_intros)+
     apply(rule_tac[1-2] continuous_on_subset[OF assms(1)[unfolded path_def]])
-    using assms(3) and ** by(auto simp add:vector_component_simps field_simps Cart_eq) qed
+    using assms(3) and ** by(auto, auto simp add: field_simps) qed
 
 lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" "a \<in> {0..1}" "x \<in> {0..1}" 
   shows "shiftpath (1 - a) (shiftpath a g) x = g x"
-  using assms unfolding pathfinish_def pathstart_def shiftpath_def 
-  by(auto simp add: vector_component_simps)
+  using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto
 
 lemma path_image_shiftpath:
   assumes "a \<in> {0..1}" "pathfinish g = pathstart g"
   shows "path_image(shiftpath a g) = path_image g" proof-
-  { fix x assume as:"g 1 = g 0" "x \<in> {0..1::real^1}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a $ 1 + x $ 1 \<le> 1}. g x \<noteq> g (a + y - 1)" 
-    hence "\<exists>y\<in>{0..1} \<inter> {x. a $ 1 + x $ 1 \<le> 1}. g x = g (a + y)" proof(cases "a \<le> x")
+  { fix x assume as:"g 1 = g 0" "x \<in> {0..1::real}" " \<forall>y\<in>{0..1} \<inter> {x. \<not> a + x \<le> 1}. g x \<noteq> g (a + y - 1)" 
+    hence "\<exists>y\<in>{0..1} \<inter> {x. a + x \<le> 1}. g x = g (a + y)" proof(cases "a \<le> x")
       case False thus ?thesis apply(rule_tac x="1 + x - a" in bexI)
         using as(1,2) and as(3)[THEN bspec[where x="1 + x - a"]] and assms(1)
-        by(auto simp add:vector_component_simps field_simps atomize_not) next
+        by(auto simp add: field_simps atomize_not) next
       case True thus ?thesis using as(1-2) and assms(1) apply(rule_tac x="x - a" in bexI)
-        by(auto simp add:vector_component_simps field_simps) qed }
-  thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def 
-    by(auto simp add:vector_component_simps image_iff) qed
+        by(auto simp add: field_simps) qed }
+  thus ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def
+    by(auto simp add: image_iff) qed
 
 subsection {* Special case of straight-line paths. *}
 
 definition
-  linepath :: "real ^ 'n \<Rightarrow> real ^ 'n \<Rightarrow> real ^ 1 \<Rightarrow> real ^ 'n" where
-  "linepath a b = (\<lambda>x. (1 - dest_vec1 x) *\<^sub>R a + dest_vec1 x *\<^sub>R b)"
+  linepath :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a" where
+  "linepath a b = (\<lambda>x. (1 - x) *\<^sub>R a + x *\<^sub>R b)"
 
 lemma pathstart_linepath[simp]: "pathstart(linepath a b) = a"
   unfolding pathstart_def linepath_def by auto
@@ -3127,30 +3190,36 @@
 lemma path_image_linepath[simp]: "path_image(linepath a b) = (closed_segment a b)"
   unfolding path_image_def segment linepath_def apply (rule set_ext, rule) defer
   unfolding mem_Collect_eq image_iff apply(erule exE) apply(rule_tac x="u *\<^sub>R 1" in bexI)
-  by(auto simp add:vector_component_simps)
+  by auto
 
 lemma reversepath_linepath[simp]:  "reversepath(linepath a b) = linepath b a"
-  unfolding reversepath_def linepath_def by(rule ext, auto simp add:vector_component_simps)
-
-lemma injective_path_linepath: assumes "a \<noteq> b" shows "injective_path(linepath a b)" proof- 
-  { obtain i where i:"a$i \<noteq> b$i" using assms[unfolded Cart_eq] by auto
-    fix x y::"real^1" assume "x $ 1 *\<^sub>R b + y $ 1 *\<^sub>R a = x $ 1 *\<^sub>R a + y $ 1 *\<^sub>R b"
-    hence "x$1 * (b$i - a$i) = y$1 * (b$i - a$i)" unfolding Cart_eq by(auto simp add:field_simps vector_component_simps)
-    hence "x = y" unfolding mult_cancel_right Cart_eq using i(1) by(auto simp add:field_simps) }
-  thus ?thesis unfolding injective_path_def linepath_def by(auto simp add:vector_component_simps algebra_simps) qed
+  unfolding reversepath_def linepath_def by(rule ext, auto)
+
+lemma injective_path_linepath:
+  assumes "a \<noteq> b" shows "injective_path(linepath a b)"
+proof -
+  { fix x y :: "real"
+    assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b"
+    hence "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps)
+    with assms have "x = y" by simp }
+  thus ?thesis unfolding injective_path_def linepath_def by(auto simp add: algebra_simps) qed
 
 lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path(linepath a b)" by(auto intro!: injective_imp_simple_path injective_path_linepath)
 
 subsection {* Bounding a point away from a path. *}
 
-lemma not_on_path_ball: assumes "path g" "z \<notin> path_image g"
+lemma not_on_path_ball:
+  fixes g :: "real \<Rightarrow> 'a::heine_borel"
+  assumes "path g" "z \<notin> path_image g"
   shows "\<exists>e>0. ball z e \<inter> (path_image g) = {}" proof-
   obtain a where "a\<in>path_image g" "\<forall>y\<in>path_image g. dist z a \<le> dist z y"
     using distance_attains_inf[OF _ path_image_nonempty, of g z]
     using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto
   thus ?thesis apply(rule_tac x="dist z a" in exI) using assms(2) by(auto intro!: dist_pos_lt) qed
 
-lemma not_on_path_cball: assumes "path g" "z \<notin> path_image g"
+lemma not_on_path_cball:
+  fixes g :: "real \<Rightarrow> 'a::heine_borel"
+  assumes "path g" "z \<notin> path_image g"
   shows "\<exists>e>0. cball z e \<inter> (path_image g) = {}" proof-
   obtain e where "ball z e \<inter> path_image g = {}" "e>0" using not_on_path_ball[OF assms] by auto
   moreover have "cball z (e/2) \<subseteq> ball z e" using `e>0` by auto
@@ -3167,14 +3236,14 @@
 
 lemma path_component_refl: assumes "x \<in> s" shows "path_component s x x"
   unfolding path_defs apply(rule_tac x="\<lambda>u. x" in exI) using assms 
-  by(auto intro!:continuous_on_intros)    
+  by(auto intro!:continuous_on_intros)
 
 lemma path_component_refl_eq: "path_component s x x \<longleftrightarrow> x \<in> s"
-  by(auto intro!: path_component_mem path_component_refl) 
+  by(auto intro!: path_component_mem path_component_refl)
 
 lemma path_component_sym: "path_component s x y \<Longrightarrow> path_component s y x"
-  using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI) 
-  by(auto simp add: reversepath_simps)
+  using assms unfolding path_component_def apply(erule exE) apply(rule_tac x="reversepath g" in exI)
+  by auto
 
 lemma path_component_trans: assumes "path_component s x y" "path_component s y z" shows "path_component s x z"
   using assms unfolding path_component_def apply- apply(erule exE)+ apply(rule_tac x="g +++ ga" in exI) by(auto simp add: path_image_join)
@@ -3209,7 +3278,9 @@
 
 subsection {* Some useful lemmas about path-connectedness. *}
 
-lemma convex_imp_path_connected: assumes "convex s" shows "path_connected s"
+lemma convex_imp_path_connected:
+  fixes s :: "'a::real_normed_vector set"
+  assumes "convex s" shows "path_connected s"
   unfolding path_connected_def apply(rule,rule,rule_tac x="linepath x y" in exI)
   unfolding path_image_linepath using assms[unfolded convex_contains_segment] by auto
 
@@ -3219,15 +3290,18 @@
   then obtain x1 x2 where obt:"x1\<in>e1\<inter>s" "x2\<in>e2\<inter>s" by auto
   then obtain g where g:"path g" "path_image g \<subseteq> s" "pathstart g = x1" "pathfinish g = x2"
     using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
-  have *:"connected {0..1::real^1}" by(auto intro!: convex_connected convex_interval)
+  have *:"connected {0..1::real}" by(auto intro!: convex_connected convex_real_interval)
   have "{0..1} \<subseteq> {x \<in> {0..1}. g x \<in> e1} \<union> {x \<in> {0..1}. g x \<in> e2}" using as(3) g(2)[unfolded path_defs] by blast
   moreover have "{x \<in> {0..1}. g x \<in> e1} \<inter> {x \<in> {0..1}. g x \<in> e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto 
-  moreover have "{x \<in> {0..1}. g x \<in> e1} \<noteq> {} \<and> {x \<in> {0..1}. g x \<in> e2} \<noteq> {}" using g(3,4)[unfolded path_defs] using obt by(auto intro!: exI)
+  moreover have "{x \<in> {0..1}. g x \<in> e1} \<noteq> {} \<and> {x \<in> {0..1}. g x \<in> e2} \<noteq> {}" using g(3,4)[unfolded path_defs] using obt
+    by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl)
   ultimately show False using *[unfolded connected_local not_ex,rule_format, of "{x\<in>{0..1}. g x \<in> e1}" "{x\<in>{0..1}. g x \<in> e2}"]
     using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(1)]
     using continuous_open_in_preimage[OF g(1)[unfolded path_def] as(2)] by auto qed
 
-lemma open_path_component: assumes "open s" shows "open(path_component s x)"
+lemma open_path_component:
+  fixes s :: "'a::real_normed_vector set" (*TODO: generalize to metric_space*)
+  assumes "open s" shows "open(path_component s x)"
   unfolding open_contains_ball proof
   fix y assume as:"y \<in> path_component s x"
   hence "y\<in>s" apply- apply(rule path_component_mem(2)) unfolding mem_def by auto
@@ -3237,7 +3311,10 @@
       apply(rule path_component_of_subset[OF e(2)]) apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) using `e>0`
       using as[unfolded mem_def] by auto qed qed
 
-lemma open_non_path_component: assumes "open s" shows "open(s - path_component s x)" unfolding open_contains_ball proof
+lemma open_non_path_component:
+  fixes s :: "'a::real_normed_vector set" (*TODO: generalize to metric_space*)
+  assumes "open s" shows "open(s - path_component s x)"
+  unfolding open_contains_ball proof
   fix y assume as:"y\<in>s - path_component s x" 
   then obtain e where e:"e>0" "ball y e \<subseteq> s" using assms[unfolded open_contains_ball] by auto
   show "\<exists>e>0. ball y e \<subseteq> s - path_component s x" apply(rule_tac x=e in exI) apply(rule,rule `e>0`,rule,rule) defer proof(rule ccontr)
@@ -3247,7 +3324,9 @@
       apply(rule convex_imp_path_connected[OF convex_ball, unfolded path_connected_component, rule_format]) by auto
     thus False using as by auto qed(insert e(2), auto) qed
 
-lemma connected_open_path_connected: assumes "open s" "connected s" shows "path_connected s"
+lemma connected_open_path_connected:
+  fixes s :: "'a::real_normed_vector set" (*TODO: generalize to metric_space*)
+  assumes "open s" "connected s" shows "path_connected s"
   unfolding path_connected_component_set proof(rule,rule,rule path_component_subset, rule)
   fix x y assume "x \<in> s" "y \<in> s" show "y \<in> path_component s x" proof(rule ccontr)
     assume "y \<notin> path_component s x" moreover
@@ -3276,8 +3355,10 @@
   unfolding path_connected_def by auto
 
 lemma path_connected_singleton: "path_connected {a}"
-  unfolding path_connected_def apply(rule,rule)
-  apply(rule_tac x="linepath a a" in exI) by(auto simp add:segment scaleR_left_diff_distrib)
+  unfolding path_connected_def pathstart_def pathfinish_def path_image_def
+  apply (clarify, rule_tac x="\<lambda>x. a" in exI, simp add: image_constant_conv)
+  apply (simp add: path_def continuous_on_const)
+  done
 
 lemma path_connected_Un: assumes "path_connected s" "path_connected t" "s \<inter> t \<noteq> {}"
   shows "path_connected (s \<union> t)" unfolding path_connected_component proof(rule,rule)
@@ -3305,7 +3386,7 @@
         hence "\<psi> k \<noteq> \<psi> (Suc k)" using \<psi>[unfolded bij_betw_def inj_on_def, THEN conjunct1, THEN bspec[where x=k]] by auto
         hence **:"?basis k + ?basis (Suc k) \<in> {x. 0 < inner (?basis (Suc k)) x} \<inter> (?A k)" 
           "?basis k - ?basis (Suc k) \<in> {x. 0 > inner (?basis (Suc k)) x} \<inter> ({x. 0 < inner (?basis (Suc k)) x} \<union> (?A k))" using d
-          by(auto simp add: inner_basis vector_component_simps intro!:bexI[where x=k])
+          by(auto simp add: inner_basis intro!:bexI[where x=k])
         show ?thesis unfolding * Un_assoc apply(rule path_connected_Un) defer apply(rule path_connected_Un) 
           prefer 5 apply(rule_tac[1-2] convex_imp_path_connected, rule convex_halfspace_lt, rule convex_halfspace_gt)
           apply(rule Suc(1)) using d ** False by auto
@@ -3319,7 +3400,7 @@
           apply(rule_tac[5] x=" ?basis 1 + ?basis 2" in nequals0I)
           apply(rule_tac[6] x="-?basis 1 + ?basis 2" in nequals0I)
           apply(rule_tac[7] x="-?basis 1 - ?basis 2" in nequals0I)
-          using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add:vector_component_simps inner_basis)
+          using d unfolding *** by(auto intro!: convex_halfspace_gt convex_halfspace_lt, auto simp add: inner_basis)
   qed qed auto qed note lem = this
 
   have ***:"\<And>x::real^'n. (\<exists>i\<in>{1..CARD('n)}. inner (basis (\<psi> i)) x \<noteq> 0) \<longleftrightarrow> (\<exists>i. inner (basis i) x \<noteq> 0)"
@@ -3347,7 +3428,7 @@
     apply(rule, rule continuous_at_within_inv[unfolded o_def inverse_eq_divide]) apply(rule continuous_at_within)
     apply(rule continuous_at_norm[unfolded o_def]) by auto
   thus ?thesis unfolding * ** using path_connected_punctured_universe[OF assms]
-    by(auto intro!: path_connected_continuous_image continuous_on_intros continuous_on_mul) qed
+    by(auto intro!: path_connected_continuous_image continuous_on_intros) qed
 
 lemma connected_sphere: "2 \<le> CARD('n) \<Longrightarrow> connected {x::real^'n. norm(x - a) = r}"
   using path_connected_sphere path_connected_imp_connected by auto
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -1,11 +1,12 @@
-(*  Title:      HOL/Library/Convex_Euclidean_Space.thy
-    Author:                     John Harrison
-    Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+(*  Title:                       HOL/Multivariate_Analysis/Derivative.thy
+    Author:                      John Harrison
+    Translation from HOL Light:  Robert Himmelmann, TU Muenchen
+*)
 
 header {* Multivariate calculus in Euclidean space. *}
 
 theory Derivative
-  imports Brouwer_Fixpoint RealVector
+imports Brouwer_Fixpoint RealVector
 begin
 
 
@@ -40,7 +41,7 @@
   show ?l unfolding deriv_def LIM_def apply safe apply(drule as,safe)
     apply(rule_tac x=d in exI,safe) apply(erule_tac x="xa + x" in allE)
     unfolding vector_dist_norm diff_0_right norm_scaleR
-    unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps *) qed
+    unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:algebra_simps *) qed
 
 lemma FDERIV_conv_has_derivative:"FDERIV f (x::'a::{real_normed_vector,perfect_space}) :> f' = (f has_derivative f') (at x)" (is "?l = ?r") proof 
   assume ?l note as = this[unfolded fderiv_def]
@@ -50,14 +51,14 @@
     thus "\<exists>d>0. \<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow>
       dist ((1 / norm (xa - netlimit (at x))) *\<^sub>R (f xa - (f (netlimit (at x)) + f' (xa - netlimit (at x))))) (0) < e"
       apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa - x" in allE)
-      unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed next
+      unfolding vector_dist_norm netlimit_at[of x] by (auto simp add: diff_diff_eq) qed next
   assume ?r note as = this[unfolded has_derivative_def]
   show ?l unfolding fderiv_def LIM_def apply-apply(rule,rule as[THEN conjunct1]) proof(rule,rule)
     fix e::real assume "e>0"
     guess d using as[THEN conjunct2,unfolded Lim_at,rule_format,OF`e>0`] ..
     thus "\<exists>s>0. \<forall>xa. xa \<noteq> 0 \<and> dist xa 0 < s \<longrightarrow> dist (norm (f (x + xa) - f x - f' xa) / norm xa) 0 < e" apply-
       apply(rule_tac x=d in exI) apply(erule conjE,rule,assumption) apply rule apply(erule_tac x="xa + x" in allE)
-      unfolding vector_dist_norm netlimit_at[of x] by(auto simp add:group_simps) qed qed
+      unfolding vector_dist_norm netlimit_at[of x] by (auto simp add: diff_diff_eq add.commute) qed qed
 
 subsection {* These are the only cases we'll care about, probably. *}
 
@@ -76,7 +77,7 @@
         (\<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. 0 < norm(x' - x) \<and> norm(x' - x) < d
         \<longrightarrow> norm(f x' - f x - f'(x' - x)) / norm(x' - x) < e)"
   unfolding has_derivative_within Lim_within vector_dist_norm
-  unfolding diff_0_right norm_mul by(simp add: group_simps)
+  unfolding diff_0_right norm_mul by (simp add: diff_diff_eq)
 
 lemma has_derivative_at':
  "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -91,9 +92,9 @@
   "a \<in> s \<Longrightarrow> open s \<Longrightarrow> ((f has_derivative f') (at a within s) \<longleftrightarrow> (f has_derivative f') (at a))"
   unfolding has_derivative_within has_derivative_at using Lim_within_open by auto
 
-subsection {* Derivatives on real = Derivatives on real^1 *}
+subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *}
 
-lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding vector_dist_norm by(auto simp add:vec1_dest_vec1_simps)
+lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding vector_dist_norm by auto
 
 lemma bounded_linear_vec1_dest_vec1: fixes f::"real \<Rightarrow> real"
   shows "linear (vec1 \<circ> f \<circ> dest_vec1) = bounded_linear f" (is "?l = ?r") proof-
@@ -186,14 +187,14 @@
   note as = assms[unfolded has_derivative_def]
   show ?thesis unfolding has_derivative_def apply(rule,rule bounded_linear_add)
     using Lim_add[OF as(1)[THEN conjunct2] as(2)[THEN conjunct2]] and as
-    by(auto simp add:group_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
+    by (auto simp add:algebra_simps scaleR_right_diff_distrib scaleR_right_distrib) qed
 
 lemma has_derivative_add_const:"(f has_derivative f') net \<Longrightarrow> ((\<lambda>x. f x + c) has_derivative f') net"
   apply(drule has_derivative_add) apply(rule has_derivative_const) by auto
 
 lemma has_derivative_sub:
  "(f has_derivative f') net \<Longrightarrow> (g has_derivative g') net \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_derivative (\<lambda>h. f'(h) - g'(h))) net"
-  apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:group_simps)
+  apply(drule has_derivative_add) apply(drule has_derivative_neg,assumption) by(simp add:algebra_simps)
 
 lemma has_derivative_setsum: assumes "finite s" "\<forall>a\<in>s. ((f a) has_derivative (f' a)) net"
   shows "((\<lambda>x. setsum (\<lambda>a. f a x) s) has_derivative (\<lambda>h. setsum (\<lambda>a. f' a h) s)) net"
@@ -281,6 +282,8 @@
 
 subsection {* differentiability. *}
 
+no_notation Deriv.differentiable (infixl "differentiable" 60)
+
 definition differentiable :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a net \<Rightarrow> bool" (infixr "differentiable" 30) where
   "f differentiable net \<equiv> (\<exists>f'. (f has_derivative f') net)"
 
@@ -391,8 +394,8 @@
       case False hence "norm (f y - (f x + f' (y - x))) < e * norm (y - x)" using as(4)[rule_format, OF `y\<in>s`]
 	unfolding vector_dist_norm diff_0_right norm_mul using as(3)
 	using pos_divide_less_eq[OF False[unfolded dist_nz], unfolded vector_dist_norm]
-	by(auto simp add:linear_0 linear_sub group_simps)
-      thus ?thesis by(auto simp add:group_simps) qed qed next
+	by (auto simp add: linear_0 linear_sub)
+      thus ?thesis by(auto simp add:algebra_simps) qed qed next
   assume ?rhs thus ?lhs unfolding has_derivative_within Lim_within apply-apply(erule conjE,rule,assumption)
     apply(rule,erule_tac x="e/2" in allE,rule,erule impE) defer apply(erule exE,rule_tac x=d in exI)
     apply(erule conjE,rule,assumption,rule,rule) unfolding vector_dist_norm diff_0_right norm_scaleR
@@ -400,7 +403,7 @@
     fix e d y assume "bounded_linear f'" "0 < e" "0 < d" "y \<in> s" "0 < norm (y - x) \<and> norm (y - x) < d"
         "norm (f y - f x - f' (y - x)) \<le> e / 2 * norm (y - x)"
     thus "\<bar>1 / norm (y - x)\<bar> * norm (f y - (f x + f' (y - x))) < e"
-      apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:group_simps) qed auto qed
+      apply(rule_tac le_less_trans[of _ "e/2"]) by(auto intro!:mult_imp_div_pos_le simp add:algebra_simps) qed auto qed
 
 lemma has_derivative_at_alt:
   "(f has_derivative f') (at x) \<longleftrightarrow> bounded_linear f' \<and>
@@ -435,8 +438,8 @@
     hence 1:"norm (f y - f x - f' (y - x)) \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x))" using d1 d2 d by auto
 
     have "norm (f y - f x) \<le> norm (f y - f x - f' (y - x)) + norm (f' (y - x))"
-      using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:group_simps)
-    also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:group_simps)
+      using norm_triangle_sub[of "f y - f x" "f' (y - x)"] by(auto simp add:algebra_simps)
+    also have "\<dots> \<le> norm (f y - f x - f' (y - x)) + B1 * norm (y - x)" apply(rule add_left_mono) using B1 by(auto simp add:algebra_simps)
     also have "\<dots> \<le> min (norm (y - x)) (e / 2 / B2 * norm (y - x)) + B1 * norm (y - x)" apply(rule add_right_mono) using d1 d2 d as by auto
     also have "\<dots> \<le> norm (y - x) + B1 * norm (y - x)" by auto
     also have "\<dots> = norm (y - x) * (1 + B1)" by(auto simp add:field_simps)
@@ -453,8 +456,8 @@
     interpret g': bounded_linear g' using assms(2) by auto
     interpret f': bounded_linear f' using assms(1) by auto
     have "norm (- g' (f' (y - x)) + g' (f y - f x)) = norm (g' (f y - f x - f' (y - x)))"
-      by(auto simp add:group_simps f'.diff g'.diff g'.add)
-    also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:group_simps)
+      by(auto simp add:algebra_simps f'.diff g'.diff g'.add)
+    also have "\<dots> \<le> B2 * norm (f y - f x - f' (y - x))" using B2 by(auto simp add:algebra_simps)
     also have "\<dots> \<le> B2 * (e / 2 / B2 * norm (y - x))" apply(rule mult_left_mono) using as d1 d2 d B2 by auto 
     also have "\<dots> \<le> e / 2 * norm (y - x)" using B2 by auto
     finally have 5:"norm (- g' (f' (y - x)) + g' (f y - f x)) \<le> e / 2 * norm (y - x)" by auto
@@ -535,7 +538,7 @@
       unfolding scaleR_right_distrib by auto
     also have "\<dots> = norm ((1 / abs c) *\<^sub>R (c *\<^sub>R (- (f' (basis i)) + f'' (basis i))))"  
       unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto
-    also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by(auto simp add:group_simps)
+    also have "\<dots> = e" unfolding e_def norm_mul using c[THEN conjunct1] using norm_minus_cancel[of "f' (basis i) - f'' (basis i)"] by (auto simp add: add.commute ab_diff_minus)
     finally show False using c using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R basis i"] using norm_basis[of i] unfolding vector_dist_norm 
       unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib by auto qed qed
 
@@ -623,10 +626,10 @@
   have ***:"\<And>y y1 y2 d dx::real. (y1\<le>y\<and>y2\<le>y) \<or> (y\<le>y1\<and>y\<le>y2) \<Longrightarrow> d < abs dx \<Longrightarrow> abs(y1 - y - - dx) \<le> d \<Longrightarrow> (abs (y2 - y - dx) \<le> d) \<Longrightarrow> False" by arith
   show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\<bar>D $ k $ j\<bar> / 2 * \<bar>d\<bar>"]) 
     using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left
-    unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding group_simps by (auto intro: mult_pos_pos)
+    unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding algebra_simps by (auto intro: mult_pos_pos)
 qed
 
-subsection {* In particular if we have a mapping into R^1. *}
+subsection {* In particular if we have a mapping into @{typ "real^1"}. *}
 
 lemma differential_zero_maxmin: fixes f::"real^'a \<Rightarrow> real"
   assumes "x \<in> s" "open s" "(f has_derivative f') (at x)"
@@ -727,7 +730,7 @@
   shows "norm(f x - f y) \<le> B * norm(x - y)" proof-
   let ?p = "\<lambda>u. x + u *\<^sub>R (y - x)"
   have *:"\<And>u. u\<in>{0..1} \<Longrightarrow> x + u *\<^sub>R (y - x) \<in> s"
-    using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:group_simps)
+    using assms(1)[unfolded convex_alt,rule_format,OF x y] unfolding scaleR_left_diff_distrib scaleR_right_diff_distrib by(auto simp add:algebra_simps)
   hence 1:"continuous_on {0..1} (f \<circ> ?p)" apply- apply(rule continuous_on_intros continuous_on_vmul)+
     unfolding continuous_on_eq_continuous_within apply(rule,rule differentiable_imp_continuous_within)
     unfolding differentiable_def apply(rule_tac x="f' xa" in exI)
@@ -754,11 +757,11 @@
 lemma onorm_vec1: fixes f::"real \<Rightarrow> real"
   shows "onorm (\<lambda>x. vec1 (f (dest_vec1 x))) = onorm f" proof-
   have "\<forall>x::real^1. norm x = 1 \<longleftrightarrow> x\<in>{vec1 -1, vec1 (1::real)}" unfolding forall_vec1 by(auto simp add:Cart_eq)
-  hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by(auto simp add:norm_vec1)
+  hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto
   have 2:"{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = (\<lambda>x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" by auto
   have "\<forall>x::real. norm x = 1 \<longleftrightarrow> x\<in>{-1, 1}" by auto hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto
   have 4:"{norm (f x) |x. norm x = 1} = (\<lambda>x. norm (f x)) ` {x. norm x=1}" by auto
-  show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max norm_vec1) qed
+  show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max) qed
 
 lemma differentiable_bound_real: fixes f::"real \<Rightarrow> real"
   assumes "convex s" "\<forall>x\<in>s. (f has_derivative f' x) (at x within s)" "\<forall>x\<in>s. onorm(f' x) \<le> B" and x:"x\<in>s" and y:"y\<in>s"
@@ -862,7 +865,7 @@
   assumes "compact t" "convex t"  "t \<noteq> {}" "continuous_on t f"
   "\<forall>x\<in>s. \<forall>y\<in>t. x + (y - f y) \<in> t" "x\<in>s"
   shows "\<exists>y\<in>t. f y = x" proof-
-  have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:group_simps)
+  have *:"\<And>x y. f y = x \<longleftrightarrow> x + (y - f y) = y" by(auto simp add:algebra_simps)
   show ?thesis  unfolding * apply(rule brouwer[OF assms(1-3), of "\<lambda>y. x + (y - f y)"])
     apply(rule continuous_on_intros assms)+ using assms(4-6) by auto qed
 
@@ -907,8 +910,8 @@
     finally have *:"norm (x + g' (z - f x) - x) < e0" by auto
     have **:"f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto
     have "norm (f x - (y + (z - f (x + g' (z - f x))))) \<le> norm (f (x + g' (z - f x)) - z) + norm (f x - y)"
-      using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:group_simps)
-    also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding group_simps ** by auto 
+      using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by(auto simp add:algebra_simps)
+    also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0[THEN conjunct2,rule_format,OF *] unfolding algebra_simps ** by auto 
     also have "\<dots> \<le> 1 / (B * 2) * norm (g' (z - f x)) + e/2" using as(1)[unfolded mem_cball vector_dist_norm] by auto
     also have "\<dots> \<le> 1 / (B * 2) * B * norm (z - f x) + e/2" using * and B by(auto simp add:field_simps)
     also have "\<dots> \<le> 1 / 2 * norm (z - f x) + e/2" by auto
@@ -983,7 +986,7 @@
 (* we know for some other reason that the inverse function exists, it's OK. *}
 
 lemma bounded_linear_sub: "bounded_linear f \<Longrightarrow> bounded_linear g ==> bounded_linear (\<lambda>x. f x - g x)"
-  using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:group_simps)
+  using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g] by(auto simp add:algebra_simps)
 
 lemma has_derivative_locally_injective: fixes f::"real^'n \<Rightarrow> real^'m"
   assumes "a \<in> s" "open s" "bounded_linear g'" "g' o f'(a) = id"
@@ -1004,7 +1007,7 @@
     show "\<forall>x\<in>ball a d. \<forall>x'\<in>ball a d. f x' = f x \<longrightarrow> x' = x" proof(intro strip)
       fix x y assume as:"x\<in>ball a d" "y\<in>ball a d" "f x = f y"
       def ph \<equiv> "\<lambda>w. w - g'(f w - f x)" have ph':"ph = g' \<circ> (\<lambda>w. f' a w - (f w - f x))"
-	unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:group_simps)
+	unfolding ph_def o_def unfolding diff using f'g' by(auto simp add:algebra_simps)
       have "norm (ph x - ph y) \<le> (1/2) * norm (x - y)"
 	apply(rule differentiable_bound[OF convex_ball _ _ as(1-2), where f'="\<lambda>x v. v - g'(f' x v)"])
 	apply(rule_tac[!] ballI) proof- fix u assume u:"u \<in> ball a d" hence "u\<in>s" using d d2 by auto
@@ -1013,14 +1016,14 @@
 	  unfolding ph' * apply(rule diff_chain_within) defer apply(rule bounded_linear.has_derivative[OF assms(3)])
 	  apply(rule has_derivative_intros) defer apply(rule has_derivative_sub[where g'="\<lambda>x.0",unfolded diff_0_right])
 	  apply(rule has_derivative_at_within) using assms(5) and `u\<in>s` `a\<in>s`
-	  by(auto intro!: has_derivative_intros derivative_linear)
+          by(auto intro!: has_derivative_intros derivative_linear)
 	have **:"bounded_linear (\<lambda>x. f' u x - f' a x)" "bounded_linear (\<lambda>x. f' a x - f' u x)" apply(rule_tac[!] bounded_linear_sub)
 	  apply(rule_tac[!] derivative_linear) using assms(5) `u\<in>s` `a\<in>s` by auto
 	have "onorm (\<lambda>v. v - g' (f' u v)) \<le> onorm g' * onorm (\<lambda>w. f' a w - f' u w)" unfolding * apply(rule onorm_compose)
 	  unfolding linear_conv_bounded_linear by(rule assms(3) **)+ 
 	also have "\<dots> \<le> onorm g' * k" apply(rule mult_left_mono) 
 	  using d1[THEN conjunct2,rule_format,of u] using onorm_neg[OF **(1)[unfolded linear_linear]]
-	  using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:group_simps) 
+	  using d and u and onorm_pos_le[OF assms(3)[unfolded linear_linear]] by(auto simp add:algebra_simps) 
 	also have "\<dots> \<le> 1/2" unfolding k_def by auto
 	finally show "onorm (\<lambda>v. v - g' (f' u v)) \<le> 1 / 2" by assumption qed
       moreover have "norm (ph y - ph x) = norm (y - x)" apply(rule arg_cong[where f=norm])
@@ -1039,7 +1042,7 @@
     fix x assume "x\<in>s" show "((\<lambda>a. f m a - f n a) has_derivative (\<lambda>h. f' m x h - f' n x h)) (at x within s)"
       by(rule has_derivative_intros assms(2)[rule_format] `x\<in>s`)+
     { fix h have "norm (f' m x h - f' n x h) \<le> norm (f' m x h - g' x h) + norm (f' n x h - g' x h)"
-	using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:group_simps) 
+	using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] unfolding norm_minus_commute by(auto simp add:algebra_simps) 
       also have "\<dots> \<le> e * norm h+ e * norm h"  using assms(3)[rule_format,OF `N\<le>m` `x\<in>s`, of h] assms(3)[rule_format,OF `N\<le>n` `x\<in>s`, of h]
 	by(auto simp add:field_simps)
       finally have "norm (f' m x h - f' n x h) \<le> 2 * e * norm h" by auto }
@@ -1083,7 +1086,7 @@
       have "eventually (\<lambda>xa. norm (f n x - f n y - (f xa x - f xa y)) \<le> e * norm (x - y)) sequentially" 
 	unfolding eventually_sequentially apply(rule_tac x=N in exI) proof(rule,rule)
 	fix m assume "N\<le>m" thus "norm (f n x - f n y - (f m x - f m y)) \<le> e * norm (x - y)"
-	  using N[rule_format, of n m x y] and as by(auto simp add:group_simps) qed
+	  using N[rule_format, of n m x y] and as by(auto simp add:algebra_simps) qed
       thus "norm (f n x - f n y - (g x - g y)) \<le> e * norm (x - y)" apply-
 	apply(rule Lim_norm_ubound[OF trivial_limit_sequentially, where f="\<lambda>m. (f n x - f n y) - (f m x - f m y)"])
 	apply(rule Lim_sub Lim_const g[rule_format] as)+ by assumption qed qed
@@ -1120,10 +1123,10 @@
 	have "norm (f ?N y - f ?N x - f' ?N x (y - x)) \<le> e / 3 * norm (y - x)" using d1 and as by auto ultimately
 	have "norm (g y - g x - f' ?N x (y - x)) \<le> 2 * e / 3 * norm (y - x)" 
 	  using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] 
-	  by (auto simp add:group_simps) moreover
+	  by (auto simp add:algebra_simps) moreover
 	have " norm (f' ?N x (y - x) - g' x (y - x)) \<le> e / 3 * norm (y - x)" using N1 `x\<in>s` by auto
 	ultimately show "norm (g y - g x - g' x (y - x)) \<le> e * norm (y - x)"
-	  using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:group_simps)
+	  using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by(auto simp add:algebra_simps)
 	qed qed qed qed
 
 subsection {* Can choose to line up antiderivatives if we want. *}
@@ -1212,7 +1215,7 @@
   shows "((\<lambda>x. h (f x) (g x)) has_derivative (\<lambda>d. h (f x) (g' d) + h (f' d) (g x))) (at x)"
   using has_derivative_bilinear_within[of f f' x UNIV g g' h] unfolding within_UNIV using assms by auto
 
-subsection {* Considering derivative R(^1)->R^n as a vector. *}
+subsection {* Considering derivative @{typ "real^1 \<Rightarrow> real^'n"} as a vector. *}
 
 definition has_vector_derivative :: "(real \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> ('b) \<Rightarrow> (real net \<Rightarrow> bool)"
 (infixl "has'_vector'_derivative" 12) where
@@ -1274,7 +1277,7 @@
   unfolding has_vector_derivative_def using has_derivative_id by auto
 
 lemma has_vector_derivative_cmul:  "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net"
-  unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:group_simps)
+  unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:algebra_simps)
 
 lemma has_vector_derivative_cmul_eq: assumes "c \<noteq> 0"
   shows "(((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net \<longleftrightarrow> (f has_vector_derivative f') net)"
--- a/src/HOL/Multivariate_Analysis/Determinants.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -55,7 +55,7 @@
 done
 
   (* FIXME: In Finite_Set there is a useless further assumption *)
-lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})"
+lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_ring_inverse_zero, field})"
   apply (erule finite_induct)
   apply (simp)
   apply simp
@@ -352,13 +352,13 @@
     apply (rule setprod_insert)
     apply simp
     by blast
-  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: ring_simps)
+  also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" by (simp add: field_simps)
   also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2)
   also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
     unfolding  setprod_insert[OF th3] by simp
   finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" unfolding kU[symmetric] .
   then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U"
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
 qed
 
 lemma det_row_mul:
@@ -389,14 +389,14 @@
     apply (rule setprod_insert)
     apply simp
     by blast
-  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: ring_simps)
+  also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" by (simp add: field_simps)
   also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
     unfolding th1 by (simp add: mult_ac)
   also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
     unfolding  setprod_insert[OF th3] by simp
   finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] .
   then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)"
-    by (simp add: ring_simps)
+    by (simp add: field_simps)
 qed
 
 lemma det_row_0:
@@ -604,7 +604,7 @@
   have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
     unfolding setprod_timesf ..
   then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
-        setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
+        setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: field_simps)
 qed
 
 lemma det_mul:
@@ -681,7 +681,7 @@
         using permutes_in_image[OF q] by vector
       show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)"
         using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
-        by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose)
+        by (simp add: sign_nz th00 field_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"
@@ -772,7 +772,7 @@
   have fUk: "finite ?Uk" by simp
   have kUk: "k \<notin> ?Uk" by simp
   have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s"
-    by (vector ring_simps)
+    by (vector field_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
@@ -793,7 +793,7 @@
     unfolding thd0
     unfolding det_row_mul
     unfolding th001[of k "\<lambda>i. row i A"]
-    unfolding thd1  by (simp add: ring_simps)
+    unfolding thd1  by (simp add: field_simps)
 qed
 
 lemma cramer_lemma:
@@ -813,7 +813,7 @@
 lemma cramer:
   fixes A ::"real^'n^'n"
   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)"
+  shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)"
 proof-
   from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
     unfolding invertible_det_nz[symmetric] invertible_def by blast
@@ -821,7 +821,7 @@
   hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc)
   then have xe: "\<exists>x. A*v x = b" by blast
   {fix x assume x: "A *v x = b"
-  have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
+  have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)"
     unfolding x[symmetric]
     using d0 by (simp add: Cart_eq cramer_lemma field_simps)}
   with xe show ?thesis by auto
@@ -901,7 +901,7 @@
   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 th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: field_simps)
     have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0"
       apply (subst eq_iff_diff_eq_0) by simp
     have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp
@@ -929,7 +929,7 @@
       unfolding dot_norm_neg dist_norm[symmetric]
       unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)}
   note fc = this
-  show ?thesis unfolding linear_def vector_eq smult_conv_scaleR by (simp add: inner_simps fc ring_simps)
+  show ?thesis unfolding linear_def vector_eq smult_conv_scaleR by (simp add: inner_simps fc field_simps)
 qed
 
 lemma isometry_linear:
@@ -980,7 +980,7 @@
       using H(5-9)
       apply (simp add: norm_eq norm_eq_1)
       apply (simp add: inner_simps smult_conv_scaleR) unfolding *
-      by (simp add: ring_simps) }
+      by (simp add: field_simps) }
   note th0 = this
   let ?g = "\<lambda>x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)"
   {fix x:: "real ^'n" assume nx: "norm x = 1"
@@ -993,15 +993,15 @@
     moreover
     {assume "x = 0" "y \<noteq> 0"
       then have "dist (?g x) (?g y) = dist x y"
-        apply (simp add: dist_norm norm_mul)
+        apply (simp add: dist_norm)
         apply (rule f1[rule_format])
-        by(simp add: norm_mul field_simps)}
+        by(simp add: field_simps)}
     moreover
     {assume "x \<noteq> 0" "y = 0"
       then have "dist (?g x) (?g y) = dist x y"
-        apply (simp add: dist_norm norm_mul)
+        apply (simp add: dist_norm)
         apply (rule f1[rule_format])
-        by(simp add: norm_mul field_simps)}
+        by(simp add: field_simps)}
     moreover
     {assume z: "x \<noteq> 0" "y \<noteq> 0"
       have th00: "x = norm x *s (inverse (norm x) *s x)" "y = norm y *s (inverse (norm y) *s y)" "norm x *s f ((inverse (norm x) *s x)) = norm x *s f (inverse (norm x) *s x)"
@@ -1013,7 +1013,7 @@
         "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) =
         norm (inverse (norm x) *s x - inverse (norm y) *s y)"
         using z
-        by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_norm])
+        by (auto simp add: vector_smult_assoc field_simps intro: f1[rule_format] fd1[rule_format, unfolded dist_norm])
       from z th0[OF th00] have "dist (?g x) (?g y) = dist x y"
         by (simp add: dist_norm)}
     ultimately have "dist (?g x) (?g y) = dist x y" by blast}
@@ -1047,7 +1047,7 @@
   by (simp add: nat_number setprod_numseg mult_commute)
 
 lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1"
-  by (simp add: det_def permutes_sing sign_id UNIV_1)
+  by (simp add: det_def sign_id UNIV_1)
 
 lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1"
 proof-
@@ -1057,7 +1057,7 @@
   unfolding setsum_over_permutations_insert[OF f12]
   unfolding permutes_sing
   apply (simp add: sign_swap_id sign_id swap_id_eq)
-  by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
+  by (simp add: arith_simps(31)[symmetric] del: arith_simps(31))
 qed
 
 lemma det_3: "det (A::'a::comm_ring_1^3^3) =
@@ -1078,8 +1078,8 @@
 
   unfolding permutes_sing
   apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
-  apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31))
-  by (simp add: ring_simps)
+  apply (simp add: arith_simps(31)[symmetric] del: arith_simps(31))
+  by (simp add: field_simps)
 qed
 
 end
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -7,8 +7,8 @@
 theory Euclidean_Space
 imports
   Complex_Main "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
-  Finite_Cartesian_Product Glbs Infinite_Set Numeral_Type
-  Inner_Product
+  Finite_Cartesian_Product Infinite_Set Numeral_Type
+  Inner_Product L2_Norm
 uses "positivstellensatz.ML" ("normarith.ML")
 begin
 
@@ -114,7 +114,7 @@
   instance by (intro_classes)
 end
 
-text{* The ordering on real^1 is linear. *}
+text{* The ordering on @{typ "real^1"} is linear. *}
 
 class cart_one = assumes UNIV_one: "card (UNIV \<Colon> 'a set) = Suc 0"
 begin
@@ -257,14 +257,14 @@
   | "vector_power x (Suc n) = x * vector_power x n"
 
 instance cart :: (semiring,finite) semiring
-  apply (intro_classes) by (vector ring_simps)+
+  apply (intro_classes) by (vector field_simps)+
 
 instance cart :: (semiring_0,finite) semiring_0
-  apply (intro_classes) by (vector ring_simps)+
+  apply (intro_classes) by (vector field_simps)+
 instance cart :: (semiring_1,finite) semiring_1
   apply (intro_classes) by vector
 instance cart :: (comm_semiring,finite) comm_semiring
-  apply (intro_classes) by (vector ring_simps)+
+  apply (intro_classes) by (vector field_simps)+
 
 instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes)
 instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
@@ -278,7 +278,7 @@
 
 instance cart :: (real_algebra,finite) real_algebra
   apply intro_classes
-  apply (simp_all add: vector_scaleR_def ring_simps)
+  apply (simp_all add: vector_scaleR_def field_simps)
   apply vector
   apply vector
   done
@@ -318,19 +318,19 @@
 lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
   by (vector mult_assoc)
 lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
-  by (vector ring_simps)
+  by (vector field_simps)
 lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
-  by (vector ring_simps)
+  by (vector field_simps)
 lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
 lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
 lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
-  by (vector ring_simps)
+  by (vector field_simps)
 lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
 lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
 lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector
 lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
 lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
-  by (vector ring_simps)
+  by (vector field_simps)
 
 lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
   by (simp add: Cart_eq)
@@ -406,184 +406,6 @@
     by simp
 qed
 
-subsection {* Square root of sum of squares *}
-
-definition
-  "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"
-
-lemma setL2_cong:
-  "\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
-  unfolding setL2_def by simp
-
-lemma strong_setL2_cong:
-  "\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
-  unfolding setL2_def simp_implies_def by simp
-
-lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"
-  unfolding setL2_def by simp
-
-lemma setL2_empty [simp]: "setL2 f {} = 0"
-  unfolding setL2_def by simp
-
-lemma setL2_insert [simp]:
-  "\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
-    setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"
-  unfolding setL2_def by (simp add: setsum_nonneg)
-
-lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
-  unfolding setL2_def by (simp add: setsum_nonneg)
-
-lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"
-  unfolding setL2_def by simp
-
-lemma setL2_constant: "setL2 (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"
-  unfolding setL2_def by (simp add: real_sqrt_mult)
-
-lemma setL2_mono:
-  assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"
-  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
-  shows "setL2 f K \<le> setL2 g K"
-  unfolding setL2_def
-  by (simp add: setsum_nonneg setsum_mono power_mono prems)
-
-lemma setL2_strict_mono:
-  assumes "finite K" and "K \<noteq> {}"
-  assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i"
-  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
-  shows "setL2 f K < setL2 g K"
-  unfolding setL2_def
-  by (simp add: setsum_strict_mono power_strict_mono assms)
-
-lemma setL2_right_distrib:
-  "0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
-  unfolding setL2_def
-  apply (simp add: power_mult_distrib)
-  apply (simp add: setsum_right_distrib [symmetric])
-  apply (simp add: real_sqrt_mult setsum_nonneg)
-  done
-
-lemma setL2_left_distrib:
-  "0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
-  unfolding setL2_def
-  apply (simp add: power_mult_distrib)
-  apply (simp add: setsum_left_distrib [symmetric])
-  apply (simp add: real_sqrt_mult setsum_nonneg)
-  done
-
-lemma setsum_nonneg_eq_0_iff:
-  fixes f :: "'a \<Rightarrow> 'b::ordered_ab_group_add"
-  shows "\<lbrakk>finite A; \<forall>x\<in>A. 0 \<le> f x\<rbrakk> \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
-  apply (induct set: finite, simp)
-  apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)
-  done
-
-lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
-  unfolding setL2_def
-  by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)
-
-lemma setL2_triangle_ineq:
-  shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"
-proof (cases "finite A")
-  case False
-  thus ?thesis by simp
-next
-  case True
-  thus ?thesis
-  proof (induct set: finite)
-    case empty
-    show ?case by simp
-  next
-    case (insert x F)
-    hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>
-           sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"
-      by (intro real_sqrt_le_mono add_left_mono power_mono insert
-                setL2_nonneg add_increasing zero_le_power2)
-    also have
-      "\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"
-      by (rule real_sqrt_sum_squares_triangle_ineq)
-    finally show ?case
-      using insert by simp
-  qed
-qed
-
-lemma sqrt_sum_squares_le_sum:
-  "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"
-  apply (rule power2_le_imp_le)
-  apply (simp add: power2_sum)
-  apply (simp add: mult_nonneg_nonneg)
-  apply (simp add: add_nonneg_nonneg)
-  done
-
-lemma setL2_le_setsum [rule_format]:
-  "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply simp
-  apply clarsimp
-  apply (erule order_trans [OF sqrt_sum_squares_le_sum])
-  apply simp
-  apply simp
-  apply simp
-  done
-
-lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
-  apply (rule power2_le_imp_le)
-  apply (simp add: power2_sum)
-  apply (simp add: mult_nonneg_nonneg)
-  apply (simp add: add_nonneg_nonneg)
-  done
-
-lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply simp
-  apply simp
-  apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])
-  apply simp
-  apply simp
-  done
-
-lemma setL2_mult_ineq_lemma:
-  fixes a b c d :: real
-  shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
-proof -
-  have "0 \<le> (a * d - b * c)\<twosuperior>" by simp
-  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"
-    by (simp only: power2_diff power_mult_distrib)
-  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"
-    by simp
-  finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
-    by simp
-qed
-
-lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"
-  apply (cases "finite A")
-  apply (induct set: finite)
-  apply simp
-  apply (rule power2_le_imp_le, simp)
-  apply (rule order_trans)
-  apply (rule power_mono)
-  apply (erule add_left_mono)
-  apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg)
-  apply (simp add: power2_sum)
-  apply (simp add: power_mult_distrib)
-  apply (simp add: right_distrib left_distrib)
-  apply (rule ord_le_eq_trans)
-  apply (rule setL2_mult_ineq_lemma)
-  apply simp
-  apply (intro mult_nonneg_nonneg setL2_nonneg)
-  apply simp
-  done
-
-lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"
-  apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
-  apply fast
-  apply (subst setL2_insert)
-  apply simp
-  apply simp
-  apply simp
-  done
-
 subsection {* Metric *}
 
 (* TODO: move somewhere else *)
@@ -848,7 +670,7 @@
 subsection{* The collapse of the general concepts to dimension one. *}
 
 lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
-  by (simp add: Cart_eq forall_1)
+  by (simp add: Cart_eq)
 
 lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"
   apply auto
@@ -930,7 +752,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
-  thus ?thesis by (simp add: ring_simps power2_eq_square)
+  thus ?thesis by (simp add: field_simps power2_eq_square)
 qed
 
 lemma square_continuous: "0 < (e::real) ==> \<exists>d. 0 < d \<and> (\<forall>y. abs(y - x) < d \<longrightarrow> abs(y * y - x * x) < e)"
@@ -953,8 +775,7 @@
 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_mult_two_ex ..
   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
@@ -963,7 +784,7 @@
 lemma real_div_sqrt: "0 <= x ==> x / sqrt(x) = sqrt(x)"
   apply (cases "x = 0", simp_all)
   using sqrt_divide_self_eq[of x]
-  apply (simp add: inverse_eq_divide real_sqrt_ge_0_iff field_simps)
+  apply (simp add: inverse_eq_divide field_simps)
   done
 
 text{* Hence derive more interesting properties of the norm. *}
@@ -976,8 +797,8 @@
   by (rule norm_zero)
 
 lemma norm_mul[simp]: "norm(a *s x) = abs(a) * norm x"
-  by (simp add: norm_vector_def vector_component setL2_right_distrib
-           abs_mult cong: strong_setL2_cong)
+  by (simp add: norm_vector_def setL2_right_distrib abs_mult)
+
 lemma norm_eq_0_dot: "(norm x = 0) \<longleftrightarrow> (inner x x = (0::real))"
   by (simp add: norm_vector_def setL2_def power2_eq_square)
 lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero)
@@ -1006,7 +827,7 @@
 lemma norm_triangle_sub:
   fixes x y :: "'a::real_normed_vector"
   shows "norm x \<le> norm y  + norm (x - y)"
-  using norm_triangle_ineq[of "y" "x - y"] by (simp add: ring_simps)
+  using norm_triangle_ineq[of "y" "x - y"] by (simp add: field_simps)
 
 lemma component_le_norm: "\<bar>x$i\<bar> <= norm x"
   apply (simp add: norm_vector_def)
@@ -1044,10 +865,14 @@
   by (auto simp add: norm_eq_sqrt_inner)
 
 lemma real_abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> (x::real)^2 \<le> y^2"
-proof-
-  have "x^2 \<le> y^2 \<longleftrightarrow> (x -y) * (y + x) \<le> 0" by (simp add: ring_simps power2_eq_square)
-  also have "\<dots> \<longleftrightarrow> \<bar>x\<bar> \<le> \<bar>y\<bar>" apply (simp add: zero_compare_simps real_abs_def not_less) by arith
-finally show ?thesis ..
+proof
+  assume "\<bar>x\<bar> \<le> \<bar>y\<bar>"
+  then have "\<bar>x\<bar>\<twosuperior> \<le> \<bar>y\<bar>\<twosuperior>" by (rule power_mono, simp)
+  then show "x\<twosuperior> \<le> y\<twosuperior>" by simp
+next
+  assume "x\<twosuperior> \<le> y\<twosuperior>"
+  then have "sqrt (x\<twosuperior>) \<le> sqrt (y\<twosuperior>)" by (rule real_sqrt_le_mono)
+  then show "\<bar>x\<bar> \<le> \<bar>y\<bar>" by simp
 qed
 
 lemma norm_le_square: "norm(x) <= a \<longleftrightarrow> 0 <= a \<and> x \<bullet> x <= a^2"
@@ -1076,7 +901,7 @@
   unfolding power2_norm_eq_inner inner_simps inner_commute by auto 
 
 lemma dot_norm_neg: "x \<bullet> y = ((norm x ^ 2 + norm y ^ 2) - norm(x - y) ^ 2) / 2"
-  unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:group_simps)
+  unfolding power2_norm_eq_inner inner_simps inner_commute by(auto simp add:algebra_simps)
 
 text{* Equality of vectors in terms of @{term "op \<bullet>"} products.    *}
 
@@ -1087,7 +912,7 @@
   assume ?rhs
   then have "x \<bullet> x - x \<bullet> y = 0 \<and> x \<bullet> y - y \<bullet> y = 0" by simp
   hence "x \<bullet> (x - y) = 0 \<and> y \<bullet> (x - y) = 0" by (simp add: inner_simps inner_commute)
-  then have "(x - y) \<bullet> (x - y) = 0" by (simp add: ring_simps inner_simps inner_commute)
+  then have "(x - y) \<bullet> (x - y) = 0" by (simp add: field_simps inner_simps inner_commute)
   then show "x = y" by (simp)
 qed
 
@@ -1108,7 +933,7 @@
   by (rule order_trans [OF norm_triangle_ineq add_mono])
 
 lemma ge_iff_diff_ge_0: "(a::'a::linordered_ring) \<ge> b == a - b \<ge> 0"
-  by (simp add: ring_simps)
+  by (simp add: field_simps)
 
 lemma pth_1:
   fixes x :: "'a::real_normed_vector"
@@ -1357,7 +1182,7 @@
   assumes fS: "finite S"
   shows "setsum f S *s v = setsum (\<lambda>x. f x *s v) S"
 proof(induct rule: finite_induct[OF fS])
-  case 1 then show ?case by (simp add: vector_smult_lzero)
+  case 1 then show ?case by simp
 next
   case (2 x F)
   from "2.hyps" have "setsum f (insert x F) *s v = (f x + setsum f F) *s v"
@@ -1576,7 +1401,7 @@
   apply (subgoal_tac "vector [v$1] = v")
   apply simp
   apply (vector vector_def)
-  apply (simp add: forall_1)
+  apply simp
   done
 
 lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"
@@ -1608,15 +1433,15 @@
   shows "linear f" using assms unfolding linear_def by auto
 
 lemma linear_compose_cmul: "linear f ==> linear (\<lambda>x. (c::'a::comm_semiring) *s f x)"
-  by (vector linear_def Cart_eq ring_simps)
+  by (vector linear_def Cart_eq field_simps)
 
 lemma linear_compose_neg: "linear (f :: 'a ^'n \<Rightarrow> 'a::comm_ring ^'m) ==> linear (\<lambda>x. -(f(x)))" by (vector linear_def Cart_eq)
 
 lemma linear_compose_add: "linear (f :: 'a ^'n \<Rightarrow> 'a::semiring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f(x) + g(x))"
-  by (vector linear_def Cart_eq ring_simps)
+  by (vector linear_def Cart_eq field_simps)
 
 lemma linear_compose_sub: "linear (f :: 'a ^'n \<Rightarrow> 'a::ring_1 ^'m) \<Longrightarrow> linear g ==> linear (\<lambda>x. f x - g x)"
-  by (vector linear_def Cart_eq ring_simps)
+  by (vector linear_def Cart_eq field_simps)
 
 lemma linear_compose: "linear f \<Longrightarrow> linear g ==> linear (g o f)"
   by (simp add: linear_def)
@@ -1638,7 +1463,7 @@
   shows "linear (\<lambda>x. f x $ k *s v)"
   using lf
   apply (auto simp add: linear_def )
-  by (vector ring_simps)+
+  by (vector field_simps)+
 
 lemma linear_0: "linear f ==> f 0 = (0::'a::semiring_1 ^'n)"
   unfolding linear_def
@@ -1714,7 +1539,7 @@
       unfolding norm_mul
       apply (simp only: mult_commute)
       apply (rule mult_mono)
-      by (auto simp add: ring_simps norm_ge_zero) }
+      by (auto simp add: field_simps) }
     then have th: "\<forall>i\<in> ?S. norm ((x$i) *s f (basis i :: real ^'m)) \<le> norm (f (basis i)) * norm x" by metis
     from real_setsum_norm_le[OF fS, of "\<lambda>i. (x$i) *s (f (basis i))", OF th]
     have "norm (f x) \<le> ?B * norm x" unfolding th0 setsum_left_distrib by metis}
@@ -1731,16 +1556,16 @@
   let ?K = "\<bar>B\<bar> + 1"
   have Kp: "?K > 0" by arith
     {assume C: "B < 0"
-      have "norm (1::real ^ 'n) > 0" by (simp add: zero_less_norm_iff)
+      have "norm (1::real ^ 'n) > 0" by simp
       with C have "B * norm (1:: real ^ 'n) < 0"
-        by (simp add: zero_compare_simps)
+        by (simp add: mult_less_0_iff)
       with B[rule_format, of 1] norm_ge_zero[of "f 1"] have False by simp
     }
     then have Bp: "B \<ge> 0" by ferrack
     {fix x::"real ^ 'n"
       have "norm (f x) \<le> ?K *  norm x"
       using B[rule_format, of x] norm_ge_zero[of x] norm_ge_zero[of "f x"] Bp
-      apply (auto simp add: ring_simps split add: abs_split)
+      apply (auto simp add: field_simps split add: abs_split)
       apply (erule order_trans, simp)
       done
   }
@@ -1809,12 +1634,12 @@
 lemma bilinear_lzero:
   fixes h :: "'a::ring^'n \<Rightarrow> _" assumes bh: "bilinear h" shows "h 0 x = 0"
   using bilinear_ladd[OF bh, of 0 0 x]
-    by (simp add: eq_add_iff ring_simps)
+    by (simp add: eq_add_iff field_simps)
 
 lemma bilinear_rzero:
   fixes h :: "'a::ring^_ \<Rightarrow> _" assumes bh: "bilinear h" shows "h x 0 = 0"
   using bilinear_radd[OF bh, of x 0 0 ]
-    by (simp add: eq_add_iff ring_simps)
+    by (simp add: eq_add_iff field_simps)
 
 lemma bilinear_lsub: "bilinear h ==> h (x - (y:: 'a::ring_1 ^ _)) z = h x z - h y z"
   by (simp  add: diff_def bilinear_ladd bilinear_lneg)
@@ -1855,11 +1680,11 @@
       apply (rule real_setsum_norm_le)
       using fN fM
       apply simp
-      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] norm_mul ring_simps)
+      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh] field_simps)
       apply (rule mult_mono)
-      apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
+      apply (auto simp add: zero_le_mult_iff component_le_norm)
       apply (rule mult_mono)
-      apply (auto simp add: norm_ge_zero zero_le_mult_iff component_le_norm)
+      apply (auto simp add: zero_le_mult_iff component_le_norm)
       done}
   then show ?thesis by metis
 qed
@@ -1879,7 +1704,7 @@
     have "B * norm x * norm y \<le> ?K * norm x * norm y"
       apply -
       apply (rule mult_right_mono, rule mult_right_mono)
-      by (auto simp add: norm_ge_zero)
+      by auto
     then have "norm (h x y) \<le> ?K * norm x * norm y"
       using B[rule_format, of x y] by simp}
   with Kp show ?thesis by blast
@@ -1945,7 +1770,7 @@
         by (simp add: linear_cmul[OF lf])
       finally have "f x \<bullet> y = x \<bullet> ?w"
         apply (simp only: )
-        apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] ring_simps)
+        apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps)
         done}
   }
   then show ?thesis unfolding adjoint_def
@@ -2010,7 +1835,7 @@
 
 lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
-  by (vector matrix_matrix_mult_def setsum_addf[symmetric] ring_simps)
+  by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps)
 
 lemma matrix_mul_lid:
   fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
@@ -2129,7 +1954,7 @@
 where "matrix f = (\<chi> i j. (f(basis j))$i)"
 
 lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::'a::comm_semiring_1 ^ _))"
-  by (simp add: linear_def matrix_vector_mult_def Cart_eq ring_simps setsum_right_distrib setsum_addf)
+  by (simp add: linear_def matrix_vector_mult_def Cart_eq field_simps setsum_right_distrib setsum_addf)
 
 lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::'a::comm_ring_1 ^ 'n)"
 apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute)
@@ -2183,9 +2008,9 @@
 proof-
   have uv': "u = 0 \<longrightarrow> v \<noteq> 0" using u v uv by arith
   have "a = a * (u + v)" unfolding uv  by simp
-  hence th: "u * a + v * a = a" by (simp add: ring_simps)
-  from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_compare_simps)
-  from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_compare_simps)
+  hence th: "u * a + v * a = a" by (simp add: field_simps)
+  from xa u have "u \<noteq> 0 \<Longrightarrow> u*x < u*a" by (simp add: mult_strict_left_mono)
+  from ya v have "v \<noteq> 0 \<Longrightarrow> v * y < v * a" by (simp add: mult_strict_left_mono)
   from xa ya u v have "u * x + v * y < u * a + v * a"
     apply (cases "u = 0", simp_all add: uv')
     apply(rule mult_strict_left_mono)
@@ -2206,7 +2031,7 @@
   shows "u * x + v * y \<le> a"
 proof-
   from xa ya u v have "u * x + v * y \<le> u * a + v * a" by (simp add: add_mono mult_left_mono)
-  also have "\<dots> \<le> (u + v) * a" by (simp add: ring_simps)
+  also have "\<dots> \<le> (u + v) * a" by (simp add: field_simps)
   finally show ?thesis unfolding uv by simp
 qed
 
@@ -2226,8 +2051,8 @@
   assumes x: "0 <= (x::real)" and y:"0 <= y" and z: "0 <= z" and xy: "x^2 <= y^2 + z^2"
   shows "x <= y + z"
 proof-
-  have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y  by (simp add: zero_compare_simps)
-  with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square ring_simps)
+  have "y^2 + z^2 \<le> y^2 + 2*y*z + z^2" using z y by (simp add: mult_nonneg_nonneg)
+  with xy have th: "x ^2 \<le> (y+z)^2" by (simp add: power2_eq_square field_simps)
   from y z have yz: "y + z \<ge> 0" by arith
   from power2_le_imp_le[OF th yz] show ?thesis .
 qed
@@ -2278,10 +2103,10 @@
       {assume x0: "x \<noteq> 0"
         hence n0: "norm x \<noteq> 0" by (metis norm_eq_zero)
         let ?c = "1/ norm x"
-        have "norm (?c*s x) = 1" using x0 by (simp add: n0 norm_mul)
+        have "norm (?c*s x) = 1" using x0 by (simp add: n0)
         with H have "norm (f(?c*s x)) \<le> b" by blast
         hence "?c * norm (f x) \<le> b"
-          by (simp add: linear_cmul[OF lf] norm_mul)
+          by (simp add: linear_cmul[OF lf])
         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}
@@ -2399,18 +2224,24 @@
   where "dest_vec1 x \<equiv> (x$1)"
 
 lemma vec1_component[simp]: "(vec1 x)$1 = x"
-  by (simp add: )
-
-lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
-  by (simp_all add:  Cart_eq forall_1)
-
-lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))" by (metis vec1_dest_vec1)
-
-lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))" by (metis vec1_dest_vec1)
-
-lemma vec1_eq[simp]:  "vec1 x = vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)
-
-lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y" by (metis vec1_dest_vec1)
+  by simp
+
+lemma vec1_dest_vec1: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y"
+  by (simp_all add:  Cart_eq)
+
+declare vec1_dest_vec1(1) [simp]
+
+lemma forall_vec1: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (vec1 x))"
+  by (metis vec1_dest_vec1(1))
+
+lemma exists_vec1: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P(vec1 x))"
+  by (metis vec1_dest_vec1(1))
+
+lemma vec1_eq[simp]:  "vec1 x = vec1 y \<longleftrightarrow> x = y"
+  by (metis vec1_dest_vec1(2))
+
+lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \<longleftrightarrow> x = y"
+  by (metis vec1_dest_vec1(1))
 
 lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
 
@@ -2438,8 +2269,8 @@
 lemma dest_vec1_sum: assumes fS: "finite S"
   shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S"
   apply (induct rule: finite_induct[OF fS])
-  apply (simp add: dest_vec1_vec)
-  apply (auto simp add:vector_minus_component)
+  apply simp
+  apply auto
   done
 
 lemma norm_vec1: "norm(vec1 x) = abs(x)"
@@ -2448,7 +2279,7 @@
 lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)"
   by (simp only: dist_real vec1_component)
 lemma abs_dest_vec1: "norm x = \<bar>dest_vec1 x\<bar>"
-  by (metis vec1_dest_vec1 norm_vec1)
+  by (metis vec1_dest_vec1(1) norm_vec1)
 
 lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component
    vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def
@@ -2476,7 +2307,7 @@
   shows "f = (\<lambda>x. vec1(row 1 (matrix f) \<bullet> x))"
   apply (rule ext)
   apply (subst matrix_works[OF lf, symmetric])
-  apply (simp add: Cart_eq matrix_vector_mult_def row_def inner_vector_def mult_commute forall_1)
+  apply (simp add: Cart_eq matrix_vector_mult_def row_def inner_vector_def mult_commute)
   done
 
 lemma dest_vec1_eq_0: "dest_vec1 x = 0 \<longleftrightarrow> x = 0"
@@ -2544,13 +2375,13 @@
   by (induct rule: finite_induct[OF fS], simp_all add: vec_0[symmetric] del: vec_0)
 
 lemma pastecart_vec[simp]: "pastecart (vec x) (vec x) = vec x"
-  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
+  by (simp add: pastecart_eq)
 
 lemma pastecart_add[simp]:"pastecart (x1::'a::{plus,times}^_) y1 + pastecart x2 y2 = pastecart (x1 + x2) (y1 + y2)"
-  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
+  by (simp add: pastecart_eq)
 
 lemma pastecart_cmul[simp]: "pastecart (c *s (x1::'a::{plus,times}^_)) (c *s y1) = c *s pastecart x1 y1"
-  by (simp add: pastecart_eq fstcart_pastecart sndcart_pastecart)
+  by (simp add: pastecart_eq)
 
 lemma pastecart_neg[simp]: "pastecart (- (x::'a::ring_1^_)) (- y) = - pastecart x y"
   unfolding vector_sneg_minus1 pastecart_cmul ..
@@ -2562,7 +2393,7 @@
   fixes f:: "'d \<Rightarrow> 'a::semiring_1^_"
   assumes fS: "finite S"
   shows "pastecart (setsum f S) (setsum g S) = setsum (\<lambda>i. pastecart (f i) (g i)) S"
-  by (simp  add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS] fstcart_pastecart sndcart_pastecart)
+  by (simp  add: pastecart_eq fstcart_setsum[OF fS] sndcart_setsum[OF fS])
 
 lemma setsum_Plus:
   "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow>
@@ -2580,7 +2411,7 @@
 lemma norm_fstcart: "norm(fstcart x) <= norm (x::real ^('n::finite + 'm::finite))"
 proof-
   have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
-    by (simp add: pastecart_fst_snd)
+    by simp
   have th1: "fstcart x \<bullet> fstcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
     by (simp add: inner_vector_def setsum_UNIV_sum pastecart_def setsum_nonneg)
   then show ?thesis
@@ -2595,7 +2426,7 @@
 lemma norm_sndcart: "norm(sndcart x) <= norm (x::real ^('n::finite + 'm::finite))"
 proof-
   have th0: "norm x = norm (pastecart (fstcart x) (sndcart x))"
-    by (simp add: pastecart_fst_snd)
+    by simp
   have th1: "sndcart x \<bullet> sndcart x \<le> pastecart (fstcart x) (sndcart x) \<bullet> pastecart (fstcart x) (sndcart x)"
     by (simp add: inner_vector_def setsum_UNIV_sum pastecart_def setsum_nonneg)
   then show ?thesis
@@ -2697,7 +2528,7 @@
 
 lemma real_arch_inv: "0 < e \<longleftrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
   using reals_Archimedean
-  apply (auto simp add: field_simps inverse_positive_iff_positive)
+  apply (auto simp add: field_simps)
   apply (subgoal_tac "inverse (real n) > 0")
   apply arith
   apply simp
@@ -2712,9 +2543,9 @@
   from h have p: "1 \<le> (1 + x) ^ n" using Suc.prems by simp
   from h have "1 + real n * x + x \<le> (1 + x) ^ n + x" by simp
   also have "\<dots> \<le> (1 + x) ^ Suc n" apply (subst diff_le_0_iff_le[symmetric])
-    apply (simp add: ring_simps)
+    apply (simp add: field_simps)
     using mult_left_mono[OF p Suc.prems] by simp
-  finally show ?case  by (simp add: real_of_nat_Suc ring_simps)
+  finally show ?case  by (simp add: real_of_nat_Suc field_simps)
 qed
 
 lemma real_arch_pow: assumes x: "1 < (x::real)" shows "\<exists>n. y < x^n"
@@ -2781,9 +2612,9 @@
     have "(- (1 - x)) * setsum (\<lambda>i. x^i) {0 .. n} = - (1 - x^(Suc n))"
       unfolding atLeastLessThanSuc_atLeastAtMost
       using x1' apply (auto simp only: field_simps)
-      apply (simp add: ring_simps)
+      apply (simp add: field_simps)
       done
-    then have ?thesis by (simp add: ring_simps) }
+    then have ?thesis by (simp add: field_simps) }
   ultimately show ?thesis by metis
 qed
 
@@ -2802,7 +2633,7 @@
   from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
   have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})" by simp
   then show ?thesis unfolding sum_gp_basic using mn
-    by (simp add: ring_simps power_add[symmetric])
+    by (simp add: field_simps power_add[symmetric])
 qed
 
 lemma sum_gp: "setsum (op ^ (x::'a::{field})) {m .. n} =
@@ -2824,7 +2655,7 @@
 lemma sum_gp_offset: "setsum (op ^ (x::'a::{field})) {m .. m+n} =
   (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
   unfolding sum_gp[of x m "m + n"] power_Suc
-  by (simp add: ring_simps power_add)
+  by (simp add: field_simps power_add)
 
 
 subsection{* A bit of linear algebra. *}
@@ -2910,7 +2741,8 @@
   "0 \<in> span S"
   "x\<in> span S \<Longrightarrow> y \<in> span S ==> x + y \<in> span S"
   "x \<in> span S \<Longrightarrow> c *s x \<in> span S"
-  by (metis span_def hull_subset subset_eq subspace_span subspace_def)+
+  by (metis span_def hull_subset subset_eq)
+     (metis subspace_span subspace_def)+
 
 lemma span_induct: assumes SP: "\<And>x. x \<in> S ==> P x"
   and P: "subspace P" and x: "x \<in> span S" shows "P x"
@@ -3008,7 +2840,7 @@
 
 (* Individual closure properties. *)
 
-lemma span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses)
+lemma span_superset: "x \<in> S ==> x \<in> span S" by (metis span_clauses(1))
 
 lemma span_0: "0 \<in> span S" by (metis subspace_span subspace_0)
 
@@ -3025,8 +2857,7 @@
   by (metis subspace_span subspace_sub)
 
 lemma span_setsum: "finite A \<Longrightarrow> \<forall>x \<in> A. f x \<in> span S ==> setsum f A \<in> span S"
-  apply (rule subspace_setsum)
-  by (metis subspace_span subspace_setsum)+
+  by (rule subspace_setsum, rule subspace_span)
 
 lemma span_add_eq: "(x::'a::ring_1^_) \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"
   apply (auto simp only: span_add span_sub)
@@ -3098,14 +2929,14 @@
     apply (simp only: )
     apply (rule span_add[unfolded mem_def])
     apply assumption+
-    apply (vector ring_simps)
+    apply (vector field_simps)
     apply (clarsimp simp add: mem_def)
     apply (rule_tac x= "c*k" in exI)
     apply (subgoal_tac "c *s x - (c * k) *s b = c*s (x - k*s b)")
     apply (simp only: )
     apply (rule span_mul[unfolded mem_def])
     apply assumption
-    by (vector ring_simps)
+    by (vector field_simps)
   ultimately show "?P a" using aS span_induct[where S=S and P= "?P"] by metis
 qed
 
@@ -3251,7 +3082,7 @@
           setsum_clauses(2)[OF fS] cong del: if_weak_cong)
       also have "\<dots> = (\<Sum>v\<in>S. u v *s v) + c *s x"
         apply (simp add: setsum_Un_disjoint[OF Sss, unfolded S1[symmetric]])
-        by (vector ring_simps)
+        by (vector field_simps)
       also have "\<dots> = c*s x + y"
         by (simp add: add_commute u)
       finally have "setsum (\<lambda>v. ?u v *s v) ?S = c*s x + y" .
@@ -3288,7 +3119,7 @@
     from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0" by auto
     have s0: "setsum (\<lambda>v. ?u v *s v) ?S = 0"
       using fS aS
-      apply (simp add: vector_smult_lneg vector_smult_lid setsum_clauses ring_simps )
+      apply (simp add: vector_smult_lneg setsum_clauses field_simps)
       apply (subst (2) ua[symmetric])
       apply (rule setsum_cong2)
       by auto
@@ -3657,7 +3488,7 @@
 
 lemma basis_card_eq_dim:
   "B \<subseteq> (V:: (real ^'n) set) \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> finite B \<and> card B = dim V"
-  by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_mono independent_bound)
+  by (metis order_eq_iff independent_card_le_dim span_card_ge_dim independent_bound)
 
 lemma dim_unique: "(B::(real ^'n) set) \<subseteq> V \<Longrightarrow> V \<subseteq> span B \<Longrightarrow> independent B \<Longrightarrow> card B = n \<Longrightarrow> dim V = n"
   by (metis basis_card_eq_dim)
@@ -3821,7 +3652,7 @@
   from C(1) have fC: "finite ?C" by simp
   from fB aB C(1,2) have cC: "card ?C \<le> card (insert a B)" by (simp add: card_insert_if)
   {fix x k
-    have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: ring_simps)
+    have th0: "\<And>(a::'b::comm_ring) b c. a - (b - c) = c + (a - b)" by (simp add: field_simps)
     have "x - k *s (a - (\<Sum>x\<in>C. (x \<bullet> a / (x \<bullet> x)) *s x)) \<in> span C \<longleftrightarrow> x - k *s a \<in> span C"
       apply (simp only: vector_ssub_ldistrib th0)
       apply (rule span_add_eq)
@@ -3847,7 +3678,7 @@
         apply (subst Cy)
         using C(1) fth
         apply (simp only: setsum_clauses) unfolding smult_conv_scaleR
-        apply (auto simp add: inner_simps inner_eq_zero_iff inner_commute[of y a] dot_lsum[OF fth])
+        apply (auto simp add: inner_simps inner_commute[of y a] dot_lsum[OF fth])
         apply (rule setsum_0')
         apply clarsimp
         apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
@@ -3864,7 +3695,7 @@
         using C(1) fth
         apply (simp only: setsum_clauses) unfolding smult_conv_scaleR
         apply (subst inner_commute[of x])
-        apply (auto simp add: inner_simps inner_eq_zero_iff inner_commute[of x a] dot_rsum[OF fth])
+        apply (auto simp add: inner_simps inner_commute[of x a] dot_rsum[OF fth])
         apply (rule setsum_0')
         apply clarsimp
         apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
@@ -3937,10 +3768,10 @@
         apply (subst B') using fB fth
         unfolding setsum_clauses(2)[OF fth]
         apply simp unfolding inner_simps smult_conv_scaleR
-        apply (clarsimp simp add: inner_simps inner_eq_zero_iff smult_conv_scaleR dot_lsum)
+        apply (clarsimp simp add: inner_simps smult_conv_scaleR dot_lsum)
         apply (rule setsum_0', rule ballI)
         unfolding inner_commute
-        by (auto simp add: x field_simps inner_eq_zero_iff intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
+        by (auto simp add: x field_simps intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])}
     then show "\<forall>x \<in> B. ?a \<bullet> x = 0" by blast
   qed
   with a0 show ?thesis unfolding sSB by (auto intro: exI[where x="?a"])
@@ -4041,7 +3872,7 @@
       using z .
     {fix k assume k: "z - k *s a \<in> span b"
       have eq: "z - ?h z *s a - (z - k*s a) = (k - ?h z) *s a"
-        by (simp add: ring_simps vector_sadd_rdistrib[symmetric])
+        by (simp add: field_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
@@ -4055,7 +3886,7 @@
   let ?g = "\<lambda>z. ?h z *s f a + g (z - ?h z *s a)"
   {fix x y assume x: "x \<in> span (insert a b)" and y: "y \<in> span (insert a b)"
     have tha: "\<And>(x::'a^'n) y a k l. (x + y) - (k + l) *s a = (x - k *s a) + (y - l *s a)"
-      by (vector ring_simps)
+      by (vector field_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])
@@ -4068,14 +3899,14 @@
   moreover
   {fix x:: "'a^'n" and c:: 'a  assume x: "x \<in> span (insert a b)"
     have tha: "\<And>(x::'a^'n) c k a. c *s x - (c * k) *s a = c *s (x - k *s a)"
-      by (vector ring_simps)
+      by (vector field_simps)
     have hc: "?h (c *s x) = c * ?h x"
       apply (rule conjunct2[OF h, rule_format, symmetric])
       apply (metis span_mul x)
       by (metis tha span_mul x conjunct1[OF h])
     have "?g (c *s x) = c*s ?g x"
       unfolding hc tha g(2)[rule_format, OF conjunct1[OF h, OF x]]
-      by (vector ring_simps)}
+      by (vector field_simps)}
   moreover
   {fix x assume x: "x \<in> (insert a b)"
     {assume xa: "x = a"
@@ -4093,7 +3924,7 @@
     {assume xb: "x \<in> b"
       have h0: "0 = ?h x"
         apply (rule conjunct2[OF h, rule_format])
-        apply (metis  span_superset insertI1 xb x)
+        apply (metis  span_superset x)
         apply simp
         apply (metis span_superset xb)
         done
@@ -4440,8 +4271,7 @@
     {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)
+          by (rule exI[where x=0], simp)
       next
         fix c y1 y2 assume y1: "y1 \<in> columns A" and y2: "?P y2"
         from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
@@ -4454,7 +4284,7 @@
             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)
+              by (simp add: field_simps)
             have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
            else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
               apply (rule setsum_cong[OF refl])
@@ -4797,7 +4627,7 @@
   from infnorm_triangle[of "x - y" " y"] infnorm_triangle[of "x - y" "-x"]
   have ths: "infnorm x \<le> infnorm (x - y) + infnorm y"
     "infnorm y \<le> infnorm (x - y) + infnorm x"
-    by (simp_all add: ring_simps infnorm_neg diff_def[symmetric])
+    by (simp_all add: field_simps infnorm_neg diff_def[symmetric])
   from th[OF ths]  show ?thesis .
 qed
 
@@ -4865,7 +4695,7 @@
   hence d2: "(sqrt (real ?d))^2 = real ?d"
     by (auto intro: real_sqrt_pow2)
   have th: "sqrt (real ?d) * infnorm x \<ge> 0"
-    by (simp add: zero_le_mult_iff real_sqrt_ge_0_iff infnorm_pos_le)
+    by (simp add: zero_le_mult_iff infnorm_pos_le)
   have th1: "x \<bullet> x \<le> (sqrt (real ?d) * infnorm x)^2"
     unfolding power_mult_distrib d2
     unfolding real_of_nat_def inner_vector_def
@@ -4896,9 +4726,9 @@
       using x y
       unfolding inner_simps smult_conv_scaleR
       unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq apply (simp add: inner_commute)
-      apply (simp add: ring_simps) by metis
+      apply (simp add: field_simps) by metis
     also have "\<dots> \<longleftrightarrow> (2 * norm x * norm y * (norm x * norm y - x \<bullet> y) = 0)" using x y
-      by (simp add: ring_simps inner_commute)
+      by (simp add: field_simps inner_commute)
     also have "\<dots> \<longleftrightarrow> ?lhs" using x y
       apply simp
       by metis
@@ -4944,7 +4774,7 @@
     also have "\<dots> \<longleftrightarrow> norm x *s y = norm y *s x"
       unfolding norm_cauchy_schwarz_eq[symmetric]
       unfolding power2_norm_eq_inner inner_simps
-      by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute ring_simps)
+      by (simp add: power2_norm_eq_inner[symmetric] power2_eq_square inner_commute field_simps)
     finally have ?thesis .}
   ultimately show ?thesis by blast
 qed
@@ -5030,13 +4860,12 @@
 unfolding vector_smult_assoc
 unfolding norm_mul
 apply (subgoal_tac "norm x * c = \<bar>c\<bar> * norm x \<or> norm x * c = - \<bar>c\<bar> * norm x")
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
-apply (case_tac "c <= 0", simp add: ring_simps)
-apply (simp add: ring_simps)
+apply (case_tac "c <= 0", simp add: field_simps)
+apply (simp add: field_simps)
+apply (case_tac "c <= 0", simp add: field_simps)
+apply (simp add: field_simps)
 apply simp
 apply simp
 done
 
 end
- 
\ No newline at end of file
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -933,7 +933,7 @@
 lemma has_integral_vec1: assumes "(f has_integral k) {a..b}"
   shows "((\<lambda>x. vec1 (f x)) has_integral (vec1 k)) {a..b}"
 proof- have *:"\<And>p. (\<Sum>(x, k)\<in>p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k)"
-    unfolding vec_sub Cart_eq by(auto simp add:vec1_dest_vec1_simps split_beta)
+    unfolding vec_sub Cart_eq by(auto simp add: split_beta)
   show ?thesis using assms unfolding has_integral apply safe
     apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe)
     apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed
@@ -1131,7 +1131,7 @@
     guess d2 by(rule has_integralD[OF goal1(2) e]) note d2=this
     guess p by(rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)],of a b]) note p=this
     let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
-      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:group_simps norm_minus_commute)
+      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"] by(auto simp add:algebra_simps norm_minus_commute)
     also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
       apply(rule add_strict_mono) apply(rule_tac[!] d2(2) d1(2)) using p unfolding fine_def by auto
     finally show False by auto
@@ -1244,7 +1244,7 @@
           unfolding scaleR_right_distrib setsum_addf[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,THEN sym]
           by(rule setsum_cong2,auto)
         have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) = norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
-          unfolding * by(auto simp add:group_simps) also let ?res = "\<dots>"
+          unfolding * by(auto simp add:algebra_simps) also let ?res = "\<dots>"
         from as have *:"d1 fine p" "d2 fine p" unfolding fine_inter by auto
         have "?res < e/2 + e/2" apply(rule le_less_trans[OF norm_triangle_ineq])
           apply(rule add_strict_mono) using d1(2)[OF as(1) *(1)] and d2(2)[OF as(1) *(2)] by auto
@@ -1268,7 +1268,7 @@
 
 lemma has_integral_sub:
   shows "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow> ((\<lambda>x. f(x) - g(x)) has_integral (k - l)) s"
-  using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding group_simps by auto
+  using has_integral_add[OF _ has_integral_neg,of f k s g l] unfolding algebra_simps by auto
 
 lemma integral_0: "integral s (\<lambda>x::real^'n. 0::real^'m) = 0"
   by(rule integral_unique has_integral_0)+
@@ -1356,7 +1356,7 @@
 
 lemma has_integral_eq_eq:
   shows "\<forall>x\<in>s. f x = g x \<Longrightarrow> ((f has_integral k) s \<longleftrightarrow> (g has_integral k) s)"
-  using has_integral_eq[of s f g] has_integral_eq[of s g f] by auto 
+  using has_integral_eq[of s f g] has_integral_eq[of s g f] by rule auto
 
 lemma has_integral_null[dest]:
   assumes "content({a..b}) = 0" shows  "(f has_integral 0) ({a..b})"
@@ -1653,7 +1653,7 @@
 proof- have *:"{a..b} = ({a..b} \<inter> {x. x$k \<le> c}) \<union> ({a..b} \<inter> {x. x$k \<ge> c})" by auto
   show ?thesis apply(subst *) apply(rule tagged_division_union[OF assms])
     unfolding interval_split interior_closed_interval
-    by(auto simp add: vector_less_def Cart_lambda_beta elim!:allE[where x=k]) qed
+    by(auto simp add: vector_less_def elim!:allE[where x=k]) qed
 
 lemma has_integral_separate_sides: fixes f::"real^'m \<Rightarrow> 'a::real_normed_vector"
   assumes "(f has_integral i) ({a..b})" "e>0"
@@ -1703,7 +1703,7 @@
       proof- guess p using fine_division_exists[OF d(1), of a' b] . note p=this
         show ?thesis using norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
           using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
-          using p using assms by(auto simp add:group_simps)
+          using p using assms by(auto simp add:algebra_simps)
       qed qed  
     show "?P {x. x $ k \<ge> c}" apply(rule_tac x=d in exI) apply(rule,rule d) apply(rule,rule,rule)
     proof- fix p1 p2 assume as:"p1 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p1 \<and> p2 tagged_division_of {a..b} \<inter> {x. x $ k \<ge> c} \<and> d fine p2"
@@ -1711,7 +1711,7 @@
       proof- guess p using fine_division_exists[OF d(1), of a b'] . note p=this
         show ?thesis using norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
           using as unfolding interval_split b'_def[symmetric] a'_def[symmetric]
-          using p using assms by(auto simp add:group_simps) qed qed qed qed
+          using p using assms by(auto simp add:algebra_simps) qed qed qed qed
 
 subsection {* Generalized notion of additivity. *}
 
@@ -1743,11 +1743,11 @@
 subsection {* Using additivity of lifted function to encode definedness. *}
 
 lemma forall_option: "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>x. P(Some x))"
-  by (metis map_of.simps option.nchotomy)
+  by (metis option.nchotomy)
 
 lemma exists_option:
  "(\<exists>x. P x) \<longleftrightarrow> P None \<or> (\<exists>x. P(Some x))" 
-  by (metis map_of.simps option.nchotomy)
+  by (metis option.nchotomy)
 
 fun lifted where 
   "lifted (opp::'a\<Rightarrow>'a\<Rightarrow>'b) (Some x) (Some y) = Some(opp x y)" |
@@ -1842,13 +1842,12 @@
 lemma operative_content[intro]: "operative (op +) content"
   unfolding operative_def content_split[THEN sym] neutral_add by auto
 
-lemma neutral_monoid[simp]: "neutral ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a) = 0"
-  unfolding neutral_def apply(rule some_equality) defer
-  apply(erule_tac x=0 in allE) by auto
+lemma neutral_monoid: "neutral ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a) = 0"
+  by (rule neutral_add) (* FIXME: duplicate *)
 
 lemma monoidal_monoid[intro]:
   shows "monoidal ((op +)::('a::comm_monoid_add) \<Rightarrow> 'a \<Rightarrow> 'a)"
-  unfolding monoidal_def neutral_monoid by(auto simp add: group_simps) 
+  unfolding monoidal_def neutral_monoid by(auto simp add: algebra_simps) 
 
 lemma operative_integral: fixes f::"real^'n \<Rightarrow> 'a::banach"
   shows "operative (lifted(op +)) (\<lambda>i. if f integrable_on i then Some(integral i f) else None)"
@@ -1941,7 +1940,7 @@
     apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
     apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
     unfolding division_points_def unfolding interval_bounds[OF ab]
-    apply (auto simp add:interval_bounds) unfolding * by auto
+    apply auto unfolding * by auto
   thus "?D1 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto
 
   have *:"interval_lowerbound ({a..b} \<inter> {x. x $ k \<ge> interval_lowerbound l $ k}) $ k = interval_lowerbound l $ k"
@@ -1952,7 +1951,7 @@
     apply(rule_tac x="(k,(interval_lowerbound l)$k)" in exI) defer
     apply(rule_tac x="(k,(interval_upperbound l)$k)" in exI)
     unfolding division_points_def unfolding interval_bounds[OF ab]
-    apply (auto simp add:interval_bounds) unfolding * by auto
+    apply auto unfolding * by auto
   thus "?D2 \<subset> ?D" apply-apply(rule,rule division_points_subset[OF assms(1-4)]) by auto qed
 
 subsection {* Preservation by divisions and tagged divisions. *}
@@ -2181,7 +2180,7 @@
   unfolding operative_tagged_division[OF monoidal_monoid operative_content assms,THEN sym]
   apply(subst setsum_iterate) using assms by auto
 
-subsection {* Finally, the integral of a constant\<forall> *}
+subsection {* Finally, the integral of a constant *}
 
 lemma has_integral_const[intro]:
   "((\<lambda>x. c) has_integral (content({a..b::real^'n}) *\<^sub>R c)) ({a..b})"
@@ -2254,7 +2253,7 @@
   assumes "p tagged_division_of {a..b}"  "\<forall>x\<in>{a..b}. (f x)$i \<le> (g x)$i"
   shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)$i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)$i"
   unfolding setsum_component apply(rule setsum_mono)
-  apply(rule mp) defer apply assumption apply(induct_tac x,rule) unfolding split_conv
+  apply(rule mp) defer apply assumption unfolding split_paired_all apply rule unfolding split_conv
 proof- fix a b assume ab:"(a,b) \<in> p" note assm = tagged_division_ofD(2-4)[OF assms(1) ab]
   from this(3) guess u v apply-by(erule exE)+ note b=this
   show "(content b *\<^sub>R f a) $ i \<le> (content b *\<^sub>R g a) $ i" unfolding b
@@ -2381,8 +2380,8 @@
       have lem2:"\<And>s1 s2 i1 i2. norm(s2 - s1) \<le> e/2 \<Longrightarrow> norm(s1 - i1) < e / 4 \<Longrightarrow> norm(s2 - i2) < e / 4 \<Longrightarrow>norm(i1 - i2) < e"
       proof- case goal1 have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
           using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
-          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:group_simps)
-        also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by(auto simp add:algebra_simps)
+        also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps)
         finally show ?case .
       qed
       show ?case unfolding vector_dist_norm apply(rule lem2) defer
@@ -2399,7 +2398,7 @@
         also have "\<dots> = 2 / real M" unfolding real_divide_def by auto
         finally show "norm (g n x - g m x) \<le> 2 / real M"
           using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
-          by(auto simp add:group_simps simp add:norm_minus_commute)
+          by(auto simp add:algebra_simps simp add:norm_minus_commute)
       qed qed qed
   from this[unfolded convergent_eq_cauchy[THEN sym]] guess s .. note s=this
 
@@ -2413,8 +2412,8 @@
     have lem:"\<And>sf sg i. norm(sf - sg) \<le> e / 3 \<Longrightarrow> norm(i - s) < e / 3 \<Longrightarrow> norm(sg - i) < e / 3 \<Longrightarrow> norm(sf - s) < e"
     proof- case goal1 have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
         using norm_triangle_ineq[of "sf - sg" "sg - s"]
-        using norm_triangle_ineq[of "sg -  i" " i - s"] by(auto simp add:group_simps)
-      also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:group_simps)
+        using norm_triangle_ineq[of "sg -  i" " i - s"] by(auto simp add:algebra_simps)
+      also have "\<dots> < e" using goal1 unfolding norm_minus_commute by(auto simp add:algebra_simps)
       finally show ?case .
     qed
     show ?case apply(rule_tac x=g' in exI) apply(rule,rule g')
@@ -2903,7 +2902,7 @@
   shows "setsum (\<lambda>(x,k). f(interval_upperbound k) - f(interval_lowerbound k)) p = f b - f a"
 proof- let ?f = "(\<lambda>k::(real^1) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
   have *:"operative op + ?f" unfolding operative_1_lt[OF monoidal_monoid] interval_eq_empty_1
-    by(auto simp add:not_less interval_bound_1 vector_less_def)
+    by(auto simp add:not_less vector_less_def)
   have **:"{a..b} \<noteq> {}" using assms(1) by auto note operative_tagged_division[OF monoidal_monoid * assms(2)]
   note * = this[unfolded if_not_P[OF **] interval_bound_1[OF assms(1)],THEN sym ]
   show ?thesis unfolding * apply(subst setsum_iterate[THEN sym]) defer
@@ -2956,7 +2955,7 @@
       have ball:"\<forall>xa\<in>k. xa \<in> ball x (d (dest_vec1 x))" using as(2)[unfolded fine_def,rule_format,OF `(x,k)\<in>p`,unfolded split_conv subset_eq] .
       have "norm ((v$1 - u$1) *\<^sub>R f' x - (f v - f u)) \<le> norm (f u - f x - (u$1 - x$1) *\<^sub>R f' x) + norm (f v - f x - (v$1 - x$1) *\<^sub>R f' x)"
         apply(rule order_trans[OF _ norm_triangle_ineq4]) apply(rule eq_refl) apply(rule arg_cong[where f=norm])
-        unfolding scaleR.diff_left by(auto simp add:group_simps)
+        unfolding scaleR.diff_left by(auto simp add:algebra_simps)
       also have "... \<le> e * norm (dest_vec1 u - dest_vec1 x) + e * norm (dest_vec1 v - dest_vec1 x)"
         apply(rule add_mono) apply(rule d(2)[of "x$1" "u$1",unfolded o_def vec1_dest_vec1]) prefer 4
         apply(rule d(2)[of "x$1" "v$1",unfolded o_def vec1_dest_vec1])
@@ -3098,7 +3097,7 @@
   proof(rule,rule,rule d,safe) case goal1 show ?case proof(cases "y < x")
       case False have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 y}" apply(rule integrable_subinterval,rule integrable_continuous)
         apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
-      hence *:"?I a y - ?I a x = ?I x y" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+      hence *:"?I a y - ?I a x = ?I x y" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
         using False unfolding not_less using assms(2) goal1 by auto
       have **:"norm (y - x) = content {vec1 x..vec1 y}" apply(subst content_1) using False unfolding not_less by auto
       show ?thesis unfolding ** apply(rule has_integral_bound[where f="(\<lambda>u. f u - f x) o dest_vec1"]) unfolding * unfolding o_def
@@ -3112,7 +3111,7 @@
       qed(insert e,auto)
     next case True have "f \<circ> dest_vec1 integrable_on {vec1 a..vec1 x}" apply(rule integrable_subinterval,rule integrable_continuous)
         apply(rule continuous_on_o_dest_vec1 assms)+  unfolding not_less using assms(2) goal1 by auto
-      hence *:"?I a x - ?I a y = ?I y x" unfolding group_simps apply(subst eq_commute) apply(rule integral_combine)
+      hence *:"?I a x - ?I a y = ?I y x" unfolding algebra_simps apply(subst eq_commute) apply(rule integral_combine)
         using True using assms(2) goal1 by auto
       have **:"norm (y - x) = content {vec1 y..vec1 x}" apply(subst content_1) using True unfolding not_less by auto
       have ***:"\<And>fy fx c::'a. fx - fy - (y - x) *\<^sub>R c = -(fy - fx - (x - y) *\<^sub>R c)" unfolding scaleR_left.diff by auto 
@@ -3194,7 +3193,7 @@
           apply(rule_tac X="g ` X" in UnionI) defer apply(rule_tac x="h x" in image_eqI)
           using X(2) assms(3)[rule_format,of x] by auto
       qed note ** = d(2)[OF this] have *:"inj_on (\<lambda>(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastsimp
-       have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding group_simps add_left_cancel
+       have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") unfolding algebra_simps add_left_cancel
         unfolding setsum_reindex[OF *] apply(subst scaleR_right.setsum) defer apply(rule setsum_cong2) unfolding o_def split_paired_all split_conv
         apply(drule p(4)) apply safe unfolding assms(7)[rule_format] using p by auto
       also have "... = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") unfolding scaleR.diff_right scaleR.scaleR_left[THEN sym]
@@ -3327,12 +3326,12 @@
   using additive_tagged_division_1[OF _ assms(2), of "f o dest_vec1"]
   unfolding o_def vec1_dest_vec1 using assms(1) by auto
 
-lemma split_minus[simp]:"(\<lambda>(x, k). ?f x k) x - (\<lambda>(x, k). ?g x k) x = (\<lambda>(x, k). ?f x k - ?g x k) x"
+lemma split_minus[simp]:"(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x"
   unfolding split_def by(rule refl)
 
 lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
   apply(subst(asm)(2) norm_minus_cancel[THEN sym])
-  apply(drule norm_triangle_le) by(auto simp add:group_simps)
+  apply(drule norm_triangle_le) by(auto simp add:algebra_simps)
 
 lemma fundamental_theorem_of_calculus_interior:
   assumes"a \<le> b" "continuous_on {a..b} f" "\<forall>x\<in>{a<..<b}. (f has_vector_derivative f'(x)) (at x)"
@@ -3340,7 +3339,7 @@
 proof- { presume *:"a < b \<Longrightarrow> ?thesis" 
     show ?thesis proof(cases,rule *,assumption)
       assume "\<not> a < b" hence "a = b" using assms(1) by auto
-      hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" apply(auto simp add: Cart_simps) by smt
+      hence *:"{vec a .. vec b} = {vec b}" "f b - f a = 0" by(auto simp add: Cart_eq vector_le_def order_antisym)
       show ?thesis unfolding *(2) apply(rule has_integral_null) unfolding content_eq_0_1 using * `a=b` by auto
     qed } assume ab:"a < b"
   let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {vec1 a..vec1 b} \<and> d fine p \<longrightarrow>
@@ -3422,7 +3421,7 @@
         hence "\<forall>i. u$i \<le> v$i" and uv:"{u,v}\<subseteq>{u..v}" using p(2)[OF as(1)] by auto note this(1) this(1)[unfolded forall_1]
         note result = as(2)[unfolded k interval_bounds[OF this(1)] content_1[OF this(2)]]
 
-        assume as':"x \<noteq> vec1 a" "x \<noteq> vec1 b" hence "x$1 \<in> {a<..<b}" using p(2-3)[OF as(1)] by(auto simp add:Cart_simps) note  * = d(2)[OF this] 
+        assume as':"x \<noteq> vec1 a" "x \<noteq> vec1 b" hence "x$1 \<in> {a<..<b}" using p(2-3)[OF as(1)] by(auto simp add: Cart_eq) note  * = d(2)[OF this]
         have "norm ((v$1 - u$1) *\<^sub>R f' (x$1) - (f (v$1) - f (u$1))) =
           norm ((f (u$1) - f (x$1) - (u$1 - x$1) *\<^sub>R f' (x$1)) - (f (v$1) - f (x$1) - (v$1 - x$1) *\<^sub>R f' (x$1)))" 
           apply(rule arg_cong[of _ _ norm]) unfolding scaleR_left.diff by auto 
@@ -3641,17 +3640,17 @@
   proof safe show "0 < ?d" using d(1) assms(3) unfolding Cart_simps by auto
     fix t::"_^1" assume as:"c \<le> t" "t$1 < c$1 + ?d"
     have *:"integral{a..c} f = integral{a..b} f - integral{c..b} f"
-      "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding group_simps
+      "integral{a..t} f = integral{a..b} f - integral{t..b} f" unfolding algebra_simps
       apply(rule_tac[!] integral_combine) using assms as unfolding Cart_simps by auto
     have "(- c)$1 - d < (- t)$1 \<and> - t \<le> - c" using as by auto note d(2)[rule_format,OF this]
     thus "norm (integral {a..c} f - integral {a..t} f) < e" unfolding * 
-      unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:group_simps) qed qed
+      unfolding integral_reflect apply-apply(subst norm_minus_commute) by(auto simp add:algebra_simps) qed qed
 
 declare dest_vec1_eq[simp del] not_less[simp] not_le[simp]
 
 lemma indefinite_integral_continuous: fixes f::"real^1 \<Rightarrow> 'a::banach"
   assumes "f integrable_on {a..b}" shows  "continuous_on {a..b} (\<lambda>x. integral {a..x} f)"
-proof(unfold continuous_on_def, safe)  fix x e assume as:"x\<in>{a..b}" "0<(e::real)"
+proof(unfold continuous_on_iff, safe)  fix x e assume as:"x\<in>{a..b}" "0<(e::real)"
   let ?thesis = "\<exists>d>0. \<forall>x'\<in>{a..b}. dist x' x < d \<longrightarrow> dist (integral {a..x'} f) (integral {a..x} f) < e"
   { presume *:"a<b \<Longrightarrow> ?thesis"
     show ?thesis apply(cases,rule *,assumption)
@@ -3669,7 +3668,7 @@
     from indefinite_integral_continuous_left[OF assms(1) `a<b` this `e>0`] guess d . note d=this
     show ?thesis apply(rule,rule,rule d,safe) apply(subst dist_commute)
       unfolding `x=b` vector_dist_norm apply(rule d(2)[rule_format]) unfolding norm_real by auto
-  next assume "a<x \<and> x<b" hence xl:"a<x" "x\<le>b" and xr:"a\<le>x" "x<b" by(auto simp add:Cart_simps)
+  next assume "a<x \<and> x<b" hence xl:"a<x" "x\<le>b" and xr:"a\<le>x" "x<b" by(auto simp add: vector_less_def)
     from indefinite_integral_continuous_left [OF assms(1) xl `e>0`] guess d1 . note d1=this
     from indefinite_integral_continuous_right[OF assms(1) xr `e>0`] guess d2 . note d2=this
     show ?thesis apply(rule_tac x="min d1 d2" in exI)
@@ -3715,7 +3714,7 @@
     apply safe apply(rule conv) using assms(4,7) by auto
   have *:"\<And>t xa. (1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x \<Longrightarrow> t = xa"
   proof- case goal1 hence "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c" 
-      unfolding scaleR_simps by(auto simp add:group_simps)
+      unfolding scaleR_simps by(auto simp add:algebra_simps)
     thus ?case using `x\<noteq>c` by auto qed
   have as2:"finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \<in> k}" using assms(2) 
     apply(rule finite_surj[where f="\<lambda>z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"])
@@ -3726,7 +3725,7 @@
     unfolding o_def using assms(5) defer apply-apply(rule)
   proof- fix t assume as:"t\<in>{0..1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \<in> k}"
     have *:"c - t *\<^sub>R c + t *\<^sub>R x \<in> s - k" apply safe apply(rule conv[unfolded scaleR_simps]) 
-      using `x\<in>s` `c\<in>s` as by(auto simp add:scaleR_simps)
+      using `x\<in>s` `c\<in>s` as by(auto simp add: algebra_simps)
     have "(f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\<lambda>x. 0) \<circ> (\<lambda>z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})"
       apply(rule diff_chain_within) apply(rule has_derivative_add)
       unfolding scaleR_simps apply(rule has_derivative_sub) apply(rule has_derivative_const)
@@ -3949,7 +3948,7 @@
 
 lemma has_integral_spike_set_eq: fixes f::"real^'n \<Rightarrow> 'a::banach" 
   assumes "negligible((s - t) \<union> (t - s))" shows "((f has_integral y) s \<longleftrightarrow> (f has_integral y) t)"
-  unfolding has_integral_restrict_univ[THEN sym,of f] apply(rule has_integral_spike_eq[OF assms]) by auto
+  unfolding has_integral_restrict_univ[THEN sym,of f] apply(rule has_integral_spike_eq[OF assms]) by (safe, auto split: split_if_asm)
 
 lemma has_integral_spike_set[dest]: fixes f::"real^'n \<Rightarrow> 'a::banach"
   assumes "negligible((s - t) \<union> (t - s))" "(f has_integral y) s"
@@ -4390,7 +4389,7 @@
   have *:"\<And>ir ip i cr cp. norm((cp + cr) - i) < e \<Longrightarrow> norm(cr - ir) < k \<Longrightarrow> 
     ip + ir = i \<Longrightarrow> norm(cp - ip) \<le> e + k" 
   proof- case goal1 thus ?case  using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]  
-      unfolding goal1(3)[THEN sym] norm_minus_cancel by(auto simp add:group_simps) qed
+      unfolding goal1(3)[THEN sym] norm_minus_cancel by(auto simp add:algebra_simps) qed
   
   have "?x =  norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p. integral k f))"
     unfolding split_def setsum_subtractf ..
@@ -4501,7 +4500,7 @@
             norm(c - d) < e / 4 \<longrightarrow> norm(a - d) < e" 
       proof safe case goal1 thus ?case using norm_triangle_lt[of "a - b" "b - c" "3* e/4"]
           norm_triangle_lt[of "a - b + (b - c)" "c - d" e] unfolding norm_minus_cancel
-          by(auto simp add:group_simps) qed
+          by(auto simp add:algebra_simps) qed
       show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - i) < e" apply(rule *[rule_format,where
           b="\<Sum>(x, k)\<in>p. content k *\<^sub>R f (m x) x" and c="\<Sum>(x, k)\<in>p. integral k (f (m x))"])
       proof safe case goal1
@@ -5152,7 +5151,7 @@
   assumes "f absolutely_integrable_on s" "g absolutely_integrable_on s"
   shows "(\<lambda>x. f(x) - g(x)) absolutely_integrable_on s"
   using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
-  unfolding group_simps .
+  unfolding algebra_simps .
 
 lemma absolutely_integrable_linear: fixes f::"real^'m \<Rightarrow> real^'n" and h::"real^'n \<Rightarrow> real^'p"
   assumes "f absolutely_integrable_on s" "bounded_linear h"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/L2_Norm.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -0,0 +1,187 @@
+(*  Title:      Multivariate_Analysis/L2_Norm.thy
+    Author:     Brian Huffman, Portland State University
+*)
+
+header {* Square root of sum of squares *}
+
+theory L2_Norm
+imports NthRoot
+begin
+
+definition
+  "setL2 f A = sqrt (\<Sum>i\<in>A. (f i)\<twosuperior>)"
+
+lemma setL2_cong:
+  "\<lbrakk>A = B; \<And>x. x \<in> B \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
+  unfolding setL2_def by simp
+
+lemma strong_setL2_cong:
+  "\<lbrakk>A = B; \<And>x. x \<in> B =simp=> f x = g x\<rbrakk> \<Longrightarrow> setL2 f A = setL2 g B"
+  unfolding setL2_def simp_implies_def by simp
+
+lemma setL2_infinite [simp]: "\<not> finite A \<Longrightarrow> setL2 f A = 0"
+  unfolding setL2_def by simp
+
+lemma setL2_empty [simp]: "setL2 f {} = 0"
+  unfolding setL2_def by simp
+
+lemma setL2_insert [simp]:
+  "\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
+    setL2 f (insert a F) = sqrt ((f a)\<twosuperior> + (setL2 f F)\<twosuperior>)"
+  unfolding setL2_def by (simp add: setsum_nonneg)
+
+lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
+  unfolding setL2_def by (simp add: setsum_nonneg)
+
+lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"
+  unfolding setL2_def by simp
+
+lemma setL2_constant: "setL2 (\<lambda>x. y) A = sqrt (of_nat (card A)) * \<bar>y\<bar>"
+  unfolding setL2_def by (simp add: real_sqrt_mult)
+
+lemma setL2_mono:
+  assumes "\<And>i. i \<in> K \<Longrightarrow> f i \<le> g i"
+  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
+  shows "setL2 f K \<le> setL2 g K"
+  unfolding setL2_def
+  by (simp add: setsum_nonneg setsum_mono power_mono prems)
+
+lemma setL2_strict_mono:
+  assumes "finite K" and "K \<noteq> {}"
+  assumes "\<And>i. i \<in> K \<Longrightarrow> f i < g i"
+  assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
+  shows "setL2 f K < setL2 g K"
+  unfolding setL2_def
+  by (simp add: setsum_strict_mono power_strict_mono assms)
+
+lemma setL2_right_distrib:
+  "0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
+  unfolding setL2_def
+  apply (simp add: power_mult_distrib)
+  apply (simp add: setsum_right_distrib [symmetric])
+  apply (simp add: real_sqrt_mult setsum_nonneg)
+  done
+
+lemma setL2_left_distrib:
+  "0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
+  unfolding setL2_def
+  apply (simp add: power_mult_distrib)
+  apply (simp add: setsum_left_distrib [symmetric])
+  apply (simp add: real_sqrt_mult setsum_nonneg)
+  done
+
+lemma setsum_nonneg_eq_0_iff:
+  fixes f :: "'a \<Rightarrow> 'b::ordered_ab_group_add"
+  shows "\<lbrakk>finite A; \<forall>x\<in>A. 0 \<le> f x\<rbrakk> \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
+  apply (induct set: finite, simp)
+  apply (simp add: add_nonneg_eq_0_iff setsum_nonneg)
+  done
+
+lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
+  unfolding setL2_def
+  by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)
+
+lemma setL2_triangle_ineq:
+  shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"
+proof (cases "finite A")
+  case False
+  thus ?thesis by simp
+next
+  case True
+  thus ?thesis
+  proof (induct set: finite)
+    case empty
+    show ?case by simp
+  next
+    case (insert x F)
+    hence "sqrt ((f x + g x)\<twosuperior> + (setL2 (\<lambda>i. f i + g i) F)\<twosuperior>) \<le>
+           sqrt ((f x + g x)\<twosuperior> + (setL2 f F + setL2 g F)\<twosuperior>)"
+      by (intro real_sqrt_le_mono add_left_mono power_mono insert
+                setL2_nonneg add_increasing zero_le_power2)
+    also have
+      "\<dots> \<le> sqrt ((f x)\<twosuperior> + (setL2 f F)\<twosuperior>) + sqrt ((g x)\<twosuperior> + (setL2 g F)\<twosuperior>)"
+      by (rule real_sqrt_sum_squares_triangle_ineq)
+    finally show ?case
+      using insert by simp
+  qed
+qed
+
+lemma sqrt_sum_squares_le_sum:
+  "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) \<le> x + y"
+  apply (rule power2_le_imp_le)
+  apply (simp add: power2_sum)
+  apply (simp add: mult_nonneg_nonneg)
+  apply (simp add: add_nonneg_nonneg)
+  done
+
+lemma setL2_le_setsum [rule_format]:
+  "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply simp
+  apply clarsimp
+  apply (erule order_trans [OF sqrt_sum_squares_le_sum])
+  apply simp
+  apply simp
+  apply simp
+  done
+
+lemma sqrt_sum_squares_le_sum_abs: "sqrt (x\<twosuperior> + y\<twosuperior>) \<le> \<bar>x\<bar> + \<bar>y\<bar>"
+  apply (rule power2_le_imp_le)
+  apply (simp add: power2_sum)
+  apply (simp add: mult_nonneg_nonneg)
+  apply (simp add: add_nonneg_nonneg)
+  done
+
+lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply simp
+  apply simp
+  apply (rule order_trans [OF sqrt_sum_squares_le_sum_abs])
+  apply simp
+  apply simp
+  done
+
+lemma setL2_mult_ineq_lemma:
+  fixes a b c d :: real
+  shows "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+proof -
+  have "0 \<le> (a * d - b * c)\<twosuperior>" by simp
+  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * d) * (b * c)"
+    by (simp only: power2_diff power_mult_distrib)
+  also have "\<dots> = a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior> - 2 * (a * c) * (b * d)"
+    by simp
+  finally show "2 * (a * c) * (b * d) \<le> a\<twosuperior> * d\<twosuperior> + b\<twosuperior> * c\<twosuperior>"
+    by simp
+qed
+
+lemma setL2_mult_ineq: "(\<Sum>i\<in>A. \<bar>f i\<bar> * \<bar>g i\<bar>) \<le> setL2 f A * setL2 g A"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply simp
+  apply (rule power2_le_imp_le, simp)
+  apply (rule order_trans)
+  apply (rule power_mono)
+  apply (erule add_left_mono)
+  apply (simp add: add_nonneg_nonneg mult_nonneg_nonneg setsum_nonneg)
+  apply (simp add: power2_sum)
+  apply (simp add: power_mult_distrib)
+  apply (simp add: right_distrib left_distrib)
+  apply (rule ord_le_eq_trans)
+  apply (rule setL2_mult_ineq_lemma)
+  apply simp
+  apply (intro mult_nonneg_nonneg setL2_nonneg)
+  apply simp
+  done
+
+lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"
+  apply (rule_tac s="insert i (A - {i})" and t="A" in subst)
+  apply fast
+  apply (subst setL2_insert)
+  apply simp
+  apply simp
+  apply simp
+  done
+
+end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -6,7 +6,7 @@
 header {* Elementary topology in Euclidean space. *}
 
 theory Topology_Euclidean_Space
-imports SEQ Euclidean_Space Product_Vector
+imports SEQ Euclidean_Space Product_Vector Glbs
 begin
 
 subsection{* General notion of a topology *}
@@ -48,16 +48,17 @@
   "\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"
   "\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"
   using openin[of U] unfolding istopology_def Collect_def mem_def
-  by (metis mem_def subset_eq)+
+  unfolding subset_eq Ball_def mem_def by auto
 
 lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"
   unfolding topspace_def by blast
 lemma openin_empty[simp]: "openin U {}" by (simp add: openin_clauses)
 
 lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)"
-  by (simp add: openin_clauses)
-
-lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)" by (simp add: openin_clauses)
+  using openin_clauses by simp
+
+lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)"
+  using openin_clauses by simp
 
 lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"
   using openin_Union[of "{S,T}" U] by auto
@@ -946,7 +947,7 @@
   by (metis frontier_def closure_closed Diff_subset)
 
 lemma frontier_empty[simp]: "frontier {} = {}"
-  by (simp add: frontier_def closure_empty)
+  by (simp add: frontier_def)
 
 lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"
 proof-
@@ -954,7 +955,7 @@
     hence "closure S \<subseteq> S" using interior_subset unfolding frontier_def by auto
     hence "closed S" using closure_subset_eq by auto
   }
-  thus ?thesis using frontier_subset_closed[of S] by auto
+  thus ?thesis using frontier_subset_closed[of S] ..
 qed
 
 lemma frontier_complement: "frontier(- S) = frontier S"
@@ -968,7 +969,7 @@
 
 definition
   at_infinity :: "'a::real_normed_vector net" where
-  "at_infinity = Abs_net (range (\<lambda>r. {x. r \<le> norm x}))"
+  "at_infinity = Abs_net (\<lambda>P. \<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x)"
 
 definition
   indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a net" (infixr "indirection" 70) where
@@ -976,23 +977,23 @@
 
 text{* Prove That They are all nets. *}
 
-lemma Rep_net_at_infinity:
-  "Rep_net at_infinity = range (\<lambda>r. {x. r \<le> norm x})"
+lemma eventually_at_infinity:
+  "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)"
 unfolding at_infinity_def
-apply (rule Abs_net_inverse')
-apply (rule image_nonempty, simp)
-apply (clarsimp, rename_tac r s)
-apply (rule_tac x="max r s" in exI, auto)
-done
-
-lemma within_UNIV: "net within UNIV = net"
-  by (simp add: Rep_net_inject [symmetric] Rep_net_within)
+proof (rule eventually_Abs_net, rule is_filter.intro)
+  fix P Q :: "'a \<Rightarrow> bool"
+  assume "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x" and "\<exists>s. \<forall>x. s \<le> norm x \<longrightarrow> Q x"
+  then obtain r s where
+    "\<forall>x. r \<le> norm x \<longrightarrow> P x" and "\<forall>x. s \<le> norm x \<longrightarrow> Q x" by auto
+  then have "\<forall>x. max r s \<le> norm x \<longrightarrow> P x \<and> Q x" by simp
+  then show "\<exists>r. \<forall>x. r \<le> norm x \<longrightarrow> P x \<and> Q x" ..
+qed auto
 
 subsection{* Identify Trivial limits, where we can't approach arbitrarily closely. *}
 
 definition
   trivial_limit :: "'a net \<Rightarrow> bool" where
-  "trivial_limit net \<longleftrightarrow> {} \<in> Rep_net net"
+  "trivial_limit net \<longleftrightarrow> eventually (\<lambda>x. False) net"
 
 lemma trivial_limit_within:
   shows "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"
@@ -1000,21 +1001,21 @@
   assume "trivial_limit (at a within S)"
   thus "\<not> a islimpt S"
     unfolding trivial_limit_def
-    unfolding Rep_net_within Rep_net_at
+    unfolding eventually_within eventually_at_topological
     unfolding islimpt_def
     apply (clarsimp simp add: expand_set_eq)
     apply (rename_tac T, rule_tac x=T in exI)
-    apply (clarsimp, drule_tac x=y in spec, simp)
+    apply (clarsimp, drule_tac x=y in bspec, simp_all)
     done
 next
   assume "\<not> a islimpt S"
   thus "trivial_limit (at a within S)"
     unfolding trivial_limit_def
-    unfolding Rep_net_within Rep_net_at
+    unfolding eventually_within eventually_at_topological
     unfolding islimpt_def
-    apply (clarsimp simp add: image_image)
-    apply (rule_tac x=T in image_eqI)
-    apply (auto simp add: expand_set_eq)
+    apply clarsimp
+    apply (rule_tac x=T in exI)
+    apply auto
     done
 qed
 
@@ -1030,14 +1031,14 @@
 lemma trivial_limit_at_infinity:
   "\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,zero_neq_one}) net)"
   (* FIXME: find a more appropriate type class *)
-  unfolding trivial_limit_def Rep_net_at_infinity
-  apply (clarsimp simp add: expand_set_eq)
-  apply (drule_tac x="scaleR r (sgn 1)" in spec)
+  unfolding trivial_limit_def eventually_at_infinity
+  apply clarsimp
+  apply (rule_tac x="scaleR b (sgn 1)" in exI)
   apply (simp add: norm_sgn)
   done
 
 lemma trivial_limit_sequentially[intro]: "\<not> trivial_limit sequentially"
-  by (auto simp add: trivial_limit_def Rep_net_sequentially)
+  by (auto simp add: trivial_limit_def eventually_sequentially)
 
 subsection{* Some property holds "sufficiently close" to the limit point. *}
 
@@ -1045,10 +1046,6 @@
   "eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
 unfolding eventually_at dist_nz by auto
 
-lemma eventually_at_infinity:
-  "eventually P at_infinity \<longleftrightarrow> (\<exists>b. \<forall>x. norm x >= b \<longrightarrow> P x)"
-unfolding eventually_def Rep_net_at_infinity by auto
-
 lemma eventually_within: "eventually P (at a within S) \<longleftrightarrow>
         (\<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"
 unfolding eventually_within eventually_at dist_nz by auto
@@ -1059,18 +1056,20 @@
 by auto (metis Rats_dense_in_nn_real order_le_less_trans order_refl) 
 
 lemma eventually_happens: "eventually P net ==> trivial_limit net \<or> (\<exists>x. P x)"
-  unfolding eventually_def trivial_limit_def
-  using Rep_net_nonempty [of net] by auto
+  unfolding trivial_limit_def
+  by (auto elim: eventually_rev_mp)
 
 lemma always_eventually: "(\<forall>x. P x) ==> eventually P net"
-  unfolding eventually_def trivial_limit_def
-  using Rep_net_nonempty [of net] by auto
+proof -
+  assume "\<forall>x. P x" hence "P = (\<lambda>x. True)" by (simp add: ext)
+  thus "eventually P net" by simp
+qed
 
 lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"
-  unfolding trivial_limit_def eventually_def by auto
+  unfolding trivial_limit_def by (auto elim: eventually_rev_mp)
 
 lemma eventually_False: "eventually (\<lambda>x. False) net \<longleftrightarrow> trivial_limit net"
-  unfolding trivial_limit_def eventually_def by auto
+  unfolding trivial_limit_def ..
 
 lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)"
   apply (safe elim!: trivial_limit_eventually)
@@ -1590,7 +1589,7 @@
 
 (* FIXME: Only one congruence rule for tendsto can be used at a time! *)
 
-lemma Lim_cong_within[cong add]:
+lemma Lim_cong_within(*[cong add]*):
   fixes a :: "'a::metric_space"
   fixes l :: "'b::metric_space" (* TODO: generalize *)
   shows "(\<And>x. x \<noteq> a \<Longrightarrow> f x = g x) ==> ((\<lambda>x. f x) ---> l) (at a within S) \<longleftrightarrow> ((g ---> l) (at a within S))"
@@ -1669,7 +1668,7 @@
   { fix e::real assume "e>0"
     hence "\<exists>N::nat. \<forall>n::nat\<ge>N. inverse (real n) < e"
       using real_arch_inv[of e] apply auto apply(rule_tac x=n in exI)
-      by (metis not_le le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7))
+      by (metis le_imp_inverse_le not_less real_of_nat_gt_zero_cancel_iff real_of_nat_less_iff xt1(7))
   }
   thus ?thesis unfolding Lim_sequentially dist_norm by simp
 qed
@@ -1704,7 +1703,7 @@
   apply (simp add: interior_def, safe)
   apply (force simp add: open_contains_cball)
   apply (rule_tac x="ball x e" in exI)
-  apply (simp add: open_ball centre_in_ball subset_trans [OF ball_subset_cball])
+  apply (simp add: subset_trans [OF ball_subset_cball])
   done
 
 lemma islimpt_ball:
@@ -1879,14 +1878,14 @@
 lemma frontier_ball:
   fixes a :: "'a::real_normed_vector"
   shows "0 < e ==> frontier(ball a e) = {x. dist a x = e}"
-  apply (simp add: frontier_def closure_ball interior_open open_ball order_less_imp_le)
+  apply (simp add: frontier_def closure_ball interior_open order_less_imp_le)
   apply (simp add: expand_set_eq)
   by arith
 
 lemma frontier_cball:
   fixes a :: "'a::{real_normed_vector, perfect_space}"
   shows "frontier(cball a e) = {x. dist a x = e}"
-  apply (simp add: frontier_def interior_cball closed_cball closure_closed order_less_imp_le)
+  apply (simp add: frontier_def interior_cball closed_cball order_less_imp_le)
   apply (simp add: expand_set_eq)
   by arith
 
@@ -2006,9 +2005,10 @@
 lemma bounded_ball[simp,intro]: "bounded(ball x e)"
   by (metis ball_subset_cball bounded_cball bounded_subset)
 
-lemma finite_imp_bounded[intro]: assumes "finite S" shows "bounded S"
+lemma finite_imp_bounded[intro]:
+  fixes S :: "'a::metric_space set" assumes "finite S" shows "bounded S"
 proof-
-  { fix a F assume as:"bounded F"
+  { fix a and F :: "'a set" assume as:"bounded F"
     then obtain x e where "\<forall>y\<in>F. dist x y \<le> e" unfolding bounded_def by auto
     hence "\<forall>y\<in>(insert a F). dist x y \<le> max e (dist x a)" by auto
     hence "bounded (insert a F)" unfolding bounded_def by (intro exI)
@@ -2218,7 +2218,7 @@
 apply (rule allI, rule impI, rule ext)
 apply (erule conjE)
 apply (induct_tac x)
-apply (simp add: nat_rec_0)
+apply simp
 apply (erule_tac x="n" in allE)
 apply (simp)
 done
@@ -2650,7 +2650,8 @@
   { fix x y assume "x\<in>t" "y\<in>t" "f x = f y"
     hence "x \<in> f x"  "y \<in> f x \<longrightarrow> y = x" using f[THEN bspec[where x=x]] and `t\<subseteq>s` by auto
     hence "x = y" using `f x = f y` and f[THEN bspec[where x=y]] and `y\<in>t` and `t\<subseteq>s` by auto  }
-  hence "infinite (f ` t)" using assms(2) using finite_imageD[unfolded inj_on_def, of f t] by auto
+  hence "inj_on f t" unfolding inj_on_def by simp
+  hence "infinite (f ` t)" using assms(2) using finite_imageD by auto
   moreover
   { fix x assume "x\<in>t" "f x \<notin> g"
     from g(3) assms(3) `x\<in>t` obtain h where "h\<in>g" and "x\<in>h" by auto
@@ -3145,7 +3146,7 @@
       using `?lhs`[unfolded continuous_within Lim_within] by auto
     { fix y assume "y\<in>f ` (ball x d \<inter> s)"
       hence "y \<in> ball (f x) e" using d(2) unfolding dist_nz[THEN sym]
-        apply (auto simp add: dist_commute mem_ball) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
+        apply (auto simp add: dist_commute) apply(erule_tac x=xa in ballE) apply auto using `e>0` by auto
     }
     hence "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" using `d>0` unfolding subset_eq ball_def by (auto simp add: dist_commute)  }
   thus ?rhs by auto
@@ -3168,9 +3169,32 @@
 text{* For setwise continuity, just start from the epsilon-delta definitions. *}
 
 definition
-  continuous_on :: "'a::metric_space set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> bool" where
-  "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d::real>0. \<forall>x' \<in> s. dist x' x < d --> dist (f x') (f x) < e)"
-
+  continuous_on ::
+    "'a set \<Rightarrow> ('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool"
+where
+  "continuous_on s f \<longleftrightarrow>
+    (\<forall>x\<in>s. \<forall>B. open B \<longrightarrow> f x \<in> B \<longrightarrow>
+      (\<exists>A. open A \<and> x \<in> A \<and> (\<forall>x'\<in>s. x' \<in> A \<longrightarrow> f x' \<in> B)))"
+
+lemma continuous_on_iff:
+  "continuous_on s f \<longleftrightarrow>
+    (\<forall>x\<in>s. \<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. dist x' x < d --> dist (f x') (f x) < e)"
+unfolding continuous_on_def
+apply safe
+apply (drule (1) bspec)
+apply (drule_tac x="ball (f x) e" in spec, simp, clarify)
+apply (drule (1) open_dist [THEN iffD1, THEN bspec])
+apply (clarify, rename_tac d)
+apply (rule_tac x=d in exI, clarify)
+apply (drule_tac x=x' in bspec, assumption)
+apply (drule_tac x=x' in spec, simp add: dist_commute)
+apply (drule_tac x=x in bspec, assumption)
+apply (drule (1) open_dist [THEN iffD1, THEN bspec], clarify)
+apply (drule_tac x=e in spec, simp, clarify)
+apply (rule_tac x="ball x d" in exI, simp, clarify)
+apply (drule_tac x=x' in bspec, assumption)
+apply (simp add: dist_commute)
+done
 
 definition
   uniformly_continuous_on ::
@@ -3184,14 +3208,14 @@
 
 lemma continuous_on_o_dest_vec1: fixes f::"real \<Rightarrow> 'a::real_normed_vector"
   assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)"
-  using assms unfolding continuous_on_def apply safe
+  using assms unfolding continuous_on_iff apply safe
   apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe
   apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real 
   apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:vector_le_def)
 
 lemma continuous_on_o_vec1: fixes f::"real^1 \<Rightarrow> 'a::real_normed_vector"
   assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)"
-  using assms unfolding continuous_on_def apply safe
+  using assms unfolding continuous_on_iff apply safe
   apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe
   apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real 
   apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:vector_le_def)
@@ -3200,15 +3224,17 @@
 
 lemma uniformly_continuous_imp_continuous:
  " uniformly_continuous_on s f ==> continuous_on s f"
-  unfolding uniformly_continuous_on_def continuous_on_def by blast
+  unfolding uniformly_continuous_on_def continuous_on_iff by blast
 
 lemma continuous_at_imp_continuous_within:
  "continuous (at x) f ==> continuous (at x within s) f"
   unfolding continuous_within continuous_at using Lim_at_within by auto
 
-lemma continuous_at_imp_continuous_on: assumes "(\<forall>x \<in> s. continuous (at x) f)"
+lemma continuous_at_imp_continuous_on:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  assumes "(\<forall>x \<in> s. continuous (at x) f)"
   shows "continuous_on s f"
-proof(simp add: continuous_at continuous_on_def, rule, rule, rule)
+proof(simp add: continuous_at continuous_on_iff, rule, rule, rule)
   fix x and e::real assume "x\<in>s" "e>0"
   hence "eventually (\<lambda>xa. dist (f xa) (f x) < e) (at x)" using assms unfolding continuous_at tendsto_iff by auto
   then obtain d where d:"d>0" "\<forall>xa. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" unfolding eventually_at by auto
@@ -3221,7 +3247,9 @@
 qed
 
 lemma continuous_on_eq_continuous_within:
- "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)" (is "?lhs = ?rhs")
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x within s) f)"
+    (is "?lhs = ?rhs")
 proof
   assume ?rhs
   { fix x assume "x\<in>s"
@@ -3232,18 +3260,21 @@
       hence "dist (f x') (f x) < e" using `e>0` d `x'\<in>s` dist_eq_0_iff[of x' x] zero_le_dist[of x' x] as(2) by (metis dist_eq_0_iff dist_nz) }
     hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" using `d>0` by auto
   }
-  thus ?lhs using `?rhs` unfolding continuous_on_def continuous_within Lim_within by auto
+  thus ?lhs using `?rhs` unfolding continuous_on_iff continuous_within Lim_within by auto
 next
   assume ?lhs
-  thus ?rhs unfolding continuous_on_def continuous_within Lim_within by blast
+  thus ?rhs unfolding continuous_on_iff continuous_within Lim_within by blast
 qed
 
 lemma continuous_on:
- "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. (f ---> f(x)) (at x within s))"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "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)
+  (* BH: maybe this should be the definition? *)
 
 lemma continuous_on_eq_continuous_at:
- "open s ==> (continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. continuous (at x) f))"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "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:
@@ -3252,19 +3283,22 @@
   unfolding continuous_within by(metis Lim_within_subset)
 
 lemma continuous_on_subset:
- "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "continuous_on s f \<Longrightarrow> t \<subseteq> s ==> continuous_on t f"
   unfolding continuous_on by (metis subset_eq Lim_within_subset)
 
 lemma continuous_on_interior:
- "continuous_on s f \<Longrightarrow> x \<in> interior s ==> continuous (at x) f"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "continuous_on s f \<Longrightarrow> x \<in> interior s ==> continuous (at x) f"
 unfolding interior_def
 apply simp
 by (meson continuous_on_eq_continuous_at continuous_on_subset)
 
 lemma continuous_on_eq:
- "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f
            ==> continuous_on s g"
-  by (simp add: continuous_on_def)
+  by (simp add: continuous_on_iff)
 
 text{* Characterization of various kinds of continuity in terms of sequences.  *}
 
@@ -3317,7 +3351,9 @@
   using continuous_within_sequentially[of a UNIV f] unfolding within_UNIV by auto
 
 lemma continuous_on_sequentially:
- "continuous_on s f \<longleftrightarrow>  (\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "continuous_on s f \<longleftrightarrow>
+    (\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially
                     --> ((f o x) ---> f(a)) sequentially)" (is "?lhs = ?rhs")
 proof
   assume ?rhs thus ?lhs using continuous_within_sequentially[of _ s f] unfolding continuous_on_eq_continuous_within by auto
@@ -3436,7 +3472,7 @@
 
 lemma continuous_on_const:
  "continuous_on s (\<lambda>x. c)"
-  unfolding continuous_on_eq_continuous_within using continuous_const by blast
+  unfolding continuous_on_def by auto
 
 lemma continuous_on_cmul:
   fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"
@@ -3524,7 +3560,7 @@
 
 lemma continuous_on_id:
  "continuous_on s (\<lambda>x. x)"
-  unfolding continuous_on Lim_within by auto
+  unfolding continuous_on_def by auto
 
 lemma uniformly_continuous_on_id:
  "uniformly_continuous_on s (\<lambda>x. x)"
@@ -3559,7 +3595,9 @@
 qed
 
 lemma continuous_on_compose:
- "continuous_on s f \<Longrightarrow> continuous_on (f ` s) g \<Longrightarrow> continuous_on s (g o f)"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  fixes g :: "'b::metric_space \<Rightarrow> 'c::metric_space" (* TODO: generalize *)
+  shows "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:
@@ -3610,7 +3648,8 @@
 qed
 
 lemma continuous_on_open:
- "continuous_on s f \<longleftrightarrow>
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+   shows "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
@@ -3643,7 +3682,8 @@
 (* ------------------------------------------------------------------------- *)
 
 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")
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "continuous_on s f \<longleftrightarrow>  (\<forall>t. closedin (subtopology euclidean (f ` s)) t  --> closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})" (is "?lhs = ?rhs")
 proof
   assume ?lhs
   { fix t
@@ -3667,6 +3707,7 @@
 text{* Half-global and completely global cases.                                  *}
 
 lemma continuous_open_in_preimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on s f"  "open t"
   shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
 proof-
@@ -3677,6 +3718,7 @@
 qed
 
 lemma continuous_closed_in_preimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on s f"  "closed t"
   shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
 proof-
@@ -3688,6 +3730,7 @@
 qed
 
 lemma continuous_open_preimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on s f" "open s" "open t"
   shows "open {x \<in> s. f x \<in> t}"
 proof-
@@ -3697,6 +3740,7 @@
 qed
 
 lemma continuous_closed_preimage:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on s f" "closed s" "closed t"
   shows "closed {x \<in> s. f x \<in> t}"
 proof-
@@ -3739,14 +3783,17 @@
 text{* Equality of continuous functions on closure and related results.          *}
 
 lemma continuous_closed_in_preimage_constant:
- "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "continuous_on s f ==> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
   using continuous_closed_in_preimage[of s f "{a}"] closed_sing by auto
 
 lemma continuous_closed_preimage_constant:
- "continuous_on s f \<Longrightarrow> closed s ==> closed {x \<in> s. f x = a}"
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "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:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on (closure s) f"
           "\<forall>x \<in> s. f x = a"
   shows "\<forall>x \<in> (closure s). f x = a"
@@ -3754,6 +3801,7 @@
     assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset unfolding subset_eq by auto
 
 lemma image_closure_subset:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on (closure s) f"  "closed t"  "(f ` s) \<subseteq> t"
   shows "f ` (closure s) \<subseteq> t"
 proof-
@@ -3798,11 +3846,13 @@
 using assms using continuous_within_avoid[of x UNIV f a, unfolded within_UNIV] by auto
 
 lemma continuous_on_avoid:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   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:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   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
@@ -3810,22 +3860,25 @@
 text{* Proving a function is constant by proving open-ness of level set.         *}
 
 lemma continuous_levelset_open_in_cases:
- "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
         openin (subtopology euclidean s) {x \<in> s. f x = a}
         ==> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"
 unfolding connected_clopen using continuous_closed_in_preimage_constant by auto
 
 lemma continuous_levelset_open_in:
- "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
+  shows "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:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   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
+using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by fast
 
 text{* Some arithmetical combinations (more to prove).                           *}
 
@@ -3906,7 +3959,7 @@
     then obtain y where y:"\<forall>n. y n \<in> s \<and> x n = f (y n)" unfolding image_iff Bex_def using choice[of "\<lambda>n xa. xa \<in> s \<and> x n = f xa"] by auto
     then obtain l r where "l\<in>s" and r:"subseq r" and lr:"((y \<circ> r) ---> l) sequentially" using assms(2)[unfolded compact_def, THEN spec[where x=y]] by auto
     { fix e::real assume "e>0"
-      then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' l < d \<longrightarrow> dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=l], OF `l\<in>s`] by auto
+      then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' l < d \<longrightarrow> dist (f x') (f l) < e" using assms(1)[unfolded continuous_on_iff, THEN bspec[where x=l], OF `l\<in>s`] by auto
       then obtain N::nat where N:"\<forall>n\<ge>N. dist ((y \<circ> r) n) l < d" using lr[unfolded Lim_sequentially, THEN spec[where x=d]] by auto
       { fix n::nat assume "n\<ge>N" hence "dist ((x \<circ> r) n) (f l) < e" using N[THEN spec[where x=n]] d[THEN bspec[where x="y (r n)"]] y[THEN spec[where x="r n"]] by auto  }
       hence "\<exists>N. \<forall>n\<ge>N. dist ((x \<circ> r) n) (f l) < e" by auto  }
@@ -3915,6 +3968,7 @@
 qed
 
 lemma connected_continuous_image:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "continuous_on s f"  "connected s"
   shows "connected(f ` s)"
 proof-
@@ -3935,7 +3989,7 @@
   shows "uniformly_continuous_on s f"
 proof-
     { fix x assume x:"x\<in>s"
-      hence "\<forall>xa. \<exists>y. 0 < xa \<longrightarrow> (y > 0 \<and> (\<forall>x'\<in>s. dist x' x < y \<longrightarrow> dist (f x') (f x) < xa))" using assms(1)[unfolded continuous_on_def, THEN bspec[where x=x]] by auto
+      hence "\<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_iff, 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)"
@@ -4007,7 +4061,7 @@
       using eventually_and[of "(\<lambda>n. \<forall>x\<in>s. norm (f n x - g x) < e / 3)" "(\<lambda>n. continuous_on s (f n))" net] assms(1,2) eventually_happens by blast
     have "e / 3 > 0" using `e>0` by auto
     then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3"
-      using n(2)[unfolded continuous_on_def, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast
+      using n(2)[unfolded continuous_on_iff, 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"]
@@ -4015,7 +4069,7 @@
       hence "dist (g y) (g x) < e" unfolding dist_norm using n(1)[THEN bspec[where x=y], OF `y\<in>s`]
         unfolding norm_minus_cancel[of "f n y - g y", THEN sym] using norm_triangle_lt[of "f n y - g x" "g y - f n y" e] by (auto simp add: uminus_add_conv_diff)  }
     hence "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" using `d>0` by auto  }
-  thus ?thesis unfolding continuous_on_def by auto
+  thus ?thesis unfolding continuous_on_iff by auto
 qed
 
 subsection{* Topological properties of linear functions.                               *}
@@ -4060,6 +4114,7 @@
   unfolding continuous_within using Lim_bilinear[of f "f x"] by auto
 
 lemma bilinear_continuous_on_compose:
+  fixes s :: "'a::metric_space set" (* TODO: generalize *)
   shows "continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h
              ==> continuous_on s (\<lambda>x. h (f x) (g x))"
   unfolding continuous_on_eq_continuous_within apply auto apply(erule_tac x=x in ballE) apply auto apply(erule_tac x=x in ballE) apply auto
@@ -4098,7 +4153,7 @@
 lemma continuous_on_real_range:
   fixes f :: "'a::real_normed_vector \<Rightarrow> real"
   shows "continuous_on s f \<longleftrightarrow> (\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d --> abs(f x' - f x) < e))"
-  unfolding continuous_on_def dist_norm by simp
+  unfolding continuous_on_iff dist_norm by simp
 
 lemma continuous_at_norm: "continuous (at x) norm"
   unfolding continuous_at by (intro tendsto_intros)
@@ -4109,7 +4164,9 @@
 lemma continuous_at_component: "continuous (at a) (\<lambda>x. x $ i)"
 unfolding continuous_at by (intro tendsto_intros)
 
-lemma continuous_on_component: "continuous_on s (\<lambda>x. x $ i)"
+lemma continuous_on_component:
+  fixes s :: "('a::metric_space ^ 'n) set" (* TODO: generalize *)
+  shows "continuous_on s (\<lambda>x. x $ i)"
 unfolding continuous_on by (intro ballI tendsto_intros)
 
 lemma continuous_at_infnorm: "continuous (at x) infnorm"
@@ -4333,7 +4390,7 @@
       using assms(2)[unfolded closed_sequential_limits, THEN spec[where x="\<lambda>n. sndcart (x n)"], THEN spec[where x="sndcart l"]]
       unfolding Lim_sequentially by auto
     hence "l \<in> {pastecart x y |x y. x \<in> s \<and> y \<in> t}" apply- unfolding mem_Collect_eq apply(rule_tac x="fstcart l" in exI,rule_tac x="sndcart l" in exI) by auto }
-  thus ?thesis unfolding closed_sequential_limits by auto
+  thus ?thesis unfolding closed_sequential_limits by blast
 qed
 
 lemma compact_pastecart:
@@ -4424,7 +4481,7 @@
   have "{x - y | x y . x\<in>s \<and> y\<in>s} \<noteq> {}" using `s \<noteq> {}` by auto
   then obtain x where x:"x\<in>{x - y |x y. x \<in> s \<and> y \<in> s}"  "\<forall>y\<in>{x - y |x y. x \<in> s \<and> y \<in> s}. norm y \<le> norm x"
     using compact_differences[OF assms(1) assms(1)]
-    using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by(auto simp add: norm_minus_cancel)
+    using distance_attains_sup[where 'a="'a", unfolded dist_norm, of "{x - y | x y . x\<in>s \<and> y\<in>s}" 0] by auto
   from x(1) obtain a b where "a\<in>s" "b\<in>s" "x = a - b" by auto
   thus ?thesis using x(2)[unfolded `x = a - b`] by blast
 qed
@@ -4442,10 +4499,10 @@
   let ?D = "{norm (x - y) |x y. x \<in> s \<and> y \<in> s}"
   obtain a where a:"\<forall>x\<in>s. norm x \<le> a" using assms[unfolded bounded_iff] by auto
   { fix x y assume "x \<in> s" "y \<in> s"
-    hence "norm (x - y) \<le> 2 * a" using norm_triangle_ineq[of x "-y", unfolded norm_minus_cancel] a[THEN bspec[where x=x]] a[THEN bspec[where x=y]] by (auto simp add: ring_simps)  }
+    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: field_simps)  }
   note * = this
   { fix x y assume "x\<in>s" "y\<in>s"  hence "s \<noteq> {}" by auto
-    have "norm(x - y) \<le> diameter s" unfolding diameter_def using `s\<noteq>{}` *[OF `x\<in>s` `y\<in>s`] `x\<in>s` `y\<in>s`  
+    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`
       by simp (blast intro!: Sup_upper *) }
   moreover
   { fix d::real assume "d>0" "d < diameter s"
@@ -4476,10 +4533,10 @@
 proof-
   have b:"bounded s" using assms(1) by (rule compact_imp_bounded)
   then obtain x y where xys:"x\<in>s" "y\<in>s" and xy:"\<forall>u\<in>s. \<forall>v\<in>s. norm (u - v) \<le> norm (x - y)" using compact_sup_maxdistance[OF assms] by auto
-  hence "diameter s \<le> norm (x - y)" 
-    by (force simp add: diameter_def intro!: Sup_least) 
+  hence "diameter s \<le> norm (x - y)"
+    unfolding diameter_def by clarsimp (rule Sup_least, fast+)
   thus ?thesis
-    by (metis b diameter_bounded_bound order_antisym xys) 
+    by (metis b diameter_bounded_bound order_antisym xys)
 qed
 
 text{* Related results with closure as the conclusion.                           *}
@@ -4668,7 +4725,7 @@
 lemma mem_interval_1: fixes x :: "real^1" shows
  "(x \<in> {a .. b} \<longleftrightarrow> dest_vec1 a \<le> dest_vec1 x \<and> dest_vec1 x \<le> dest_vec1 b)"
  "(x \<in> {a<..<b} \<longleftrightarrow> dest_vec1 a < dest_vec1 x \<and> dest_vec1 x < dest_vec1 b)"
-by(simp_all add: Cart_eq vector_less_def vector_le_def forall_1)
+by(simp_all add: Cart_eq vector_less_def vector_le_def)
 
 lemma vec1_interval:fixes a::"real" shows
   "vec1 ` {a .. b} = {vec1 a .. vec1 b}"
@@ -4694,7 +4751,7 @@
       have "a$i < b$i" using as[THEN spec[where x=i]] by auto
       hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i"
         unfolding vector_smult_component and vector_add_component
-        by (auto simp add: less_divide_eq_number_of1)  }
+        by auto  }
     hence "{a <..< b} \<noteq> {}" using mem_interval(1)[of "?x" a b] by auto  }
   ultimately show ?th1 by blast
 
@@ -4709,7 +4766,7 @@
       have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
       hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i"
         unfolding vector_smult_component and vector_add_component
-        by (auto simp add: less_divide_eq_number_of1)  }
+        by auto  }
     hence "{a .. b} \<noteq> {}" using mem_interval(2)[of "?x" a b] by auto  }
   ultimately show ?th2 by blast
 qed
@@ -4772,13 +4829,13 @@
       { fix j
         have "c $ j < ?x $ j \<and> ?x $ j < d $ j" unfolding Cart_lambda_beta
           apply(cases "j=i") using as(2)[THEN spec[where x=j]]
-          by (auto simp add: less_divide_eq_number_of1 as2)  }
+          by (auto simp add: as2)  }
       hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
       moreover
       have "?x\<notin>{a .. b}"
         unfolding mem_interval apply auto apply(rule_tac x=i in exI)
         using as(2)[THEN spec[where x=i]] and as2
-        by (auto simp add: less_divide_eq_number_of1)
+        by auto
       ultimately have False using as by auto  }
     hence "a$i \<le> c$i" by(rule ccontr)auto
     moreover
@@ -4787,13 +4844,13 @@
       { fix j
         have "d $ j > ?x $ j \<and> ?x $ j > c $ j" unfolding Cart_lambda_beta
           apply(cases "j=i") using as(2)[THEN spec[where x=j]]
-          by (auto simp add: less_divide_eq_number_of1 as2)  }
+          by (auto simp add: as2)  }
       hence "?x\<in>{c<..<d}" unfolding mem_interval by auto
       moreover
       have "?x\<notin>{a .. b}"
         unfolding mem_interval apply auto apply(rule_tac x=i in exI)
         using as(2)[THEN spec[where x=i]] and as2
-        by (auto simp add: less_divide_eq_number_of1)
+        by auto
       ultimately have False using as by auto  }
     hence "b$i \<ge> d$i" by(rule ccontr)auto
     ultimately
@@ -4824,7 +4881,7 @@
 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: less_divide_eq_number_of1 intro!: bexI)
+  by auto
 
 (* Moved interval_open_subset_closed a bit upwards *)
 
@@ -4864,7 +4921,7 @@
   using open_interval[of "vec1 a" "vec1 b"] unfolding open_contains_ball
   apply-apply(rule,erule_tac x="vec1 x" in ballE) apply(erule exE,rule_tac x=e in exI)
   unfolding subset_eq mem_ball apply(rule) defer apply(rule,erule conjE,erule_tac x="vec1 xa" in ballE)
-  by(auto simp add: vec1_dest_vec1_simps vector_less_def forall_1) 
+  by(auto simp add: dist_vec1 dist_real_def vector_less_def)
 
 lemma closed_interval[intro]: fixes a :: "real^'n" shows "closed {a .. b}"
 proof-
@@ -4945,7 +5002,7 @@
     have "a $ i < ((1 / 2) *\<^sub>R (a + b)) $ i \<and> ((1 / 2) *\<^sub>R (a + b)) $ i < b $ i"
       using assms[unfolded interval_ne_empty, THEN spec[where x=i]]
       unfolding vector_smult_component and vector_add_component
-      by(auto simp add: less_divide_eq_number_of1)  }
+      by auto  }
   thus ?thesis unfolding mem_interval by auto
 qed
 
@@ -4987,7 +5044,7 @@
         x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)"
         by (auto simp add: algebra_simps)
       hence "f n < b" and "a < f n" using open_closed_interval_convex[OF open_interval_midpoint[OF assms] as *] unfolding f_def by auto
-      hence False using fn unfolding f_def using xc by(auto simp add: vector_mul_lcancel vector_ssub_ldistrib)  }
+      hence False using fn unfolding f_def using xc by(auto simp add: vector_ssub_ldistrib)  }
     moreover
     { assume "\<not> (f ---> x) sequentially"
       { fix e::real assume "e>0"
@@ -5261,12 +5318,14 @@
   by (auto simp add: eventually_within_Un)
 
 lemma continuous_on_union:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "closed s" "closed t" "continuous_on s f" "continuous_on t f"
   shows "continuous_on (s \<union> t) f"
   using assms unfolding continuous_on unfolding Lim_within_union
   unfolding Lim unfolding trivial_limit_within unfolding closed_limpt by auto
 
 lemma continuous_on_cases:
+  fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" (* TODO: generalize *)
   assumes "closed s" "closed t" "continuous_on s f" "continuous_on t g"
           "\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x"
   shows "continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)"
@@ -5401,7 +5460,7 @@
       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  }
+    ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..  }
   hence "g ` t = s" by auto
   ultimately
   show ?thesis unfolding homeomorphism_def homeomorphic_def
@@ -5543,7 +5602,7 @@
   let ?S' = "{x::real^'m. x\<in>s \<and> norm x = norm a}"
   let ?S'' = "{x::real^'m. norm x = norm a}"
 
-  have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by (auto simp add: norm_minus_cancel)
+  have "?S'' = frontier(cball 0 (norm a))" unfolding frontier_cball and dist_norm by auto
   hence "compact ?S''" using compact_frontier[OF compact_cball, of 0 "norm a"] by auto
   moreover have "?S' = s \<inter> ?S''" by auto
   ultimately have "compact ?S'" using closed_inter_compact[of s ?S''] using s(1) by auto
@@ -5590,7 +5649,7 @@
 
 lemma subspace_substandard:
  "subspace {x::real^_. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
-  unfolding subspace_def by(auto simp add: vector_add_component vector_smult_component elim!: ballE)
+  unfolding subspace_def by auto
 
 lemma closed_substandard:
  "closed {x::real^'n. \<forall>i. P i --> x$i = 0}" (is "closed ?A")
@@ -5607,7 +5666,7 @@
         then obtain B where BB:"B \<in> ?Bs" and B:"B = {x::real^'n. inner (basis i) x = 0}" by auto
         hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto  }
       hence "x\<in>?A" by auto }
-    ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" by auto }
+    ultimately have "x\<in>?A \<longleftrightarrow> x\<in> \<Inter>?Bs" .. }
   hence "?A = \<Inter> ?Bs" by auto
   thus ?thesis by(auto simp add: closed_Inter closed_hyperplane)
 qed
@@ -5752,7 +5811,7 @@
   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 "m *s x = y - c" by (simp add: field_simps)
   hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
   then show "x = inverse m *s y + - (inverse m *s c)"
     using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
@@ -5783,23 +5842,23 @@
   case False
   { fix y assume "a \<le> y" "y \<le> b" "m > 0"
     hence "m *\<^sub>R a + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R b + c"
-      unfolding vector_le_def by(auto simp add: vector_smult_component vector_add_component)
+      unfolding vector_le_def by auto
   } moreover
   { fix y assume "a \<le> y" "y \<le> b" "m < 0"
     hence "m *\<^sub>R b + c \<le> m *\<^sub>R y + c"  "m *\<^sub>R y + c \<le> m *\<^sub>R a + c"
-      unfolding vector_le_def by(auto simp add: vector_smult_component vector_add_component mult_left_mono_neg elim!:ballE)
+      unfolding vector_le_def by(auto simp add: mult_left_mono_neg)
   } moreover
   { fix y assume "m > 0"  "m *\<^sub>R a + c \<le> y"  "y \<le> m *\<^sub>R b + c"
     hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval vector_le_def
-      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
+      apply(auto simp add: vector_smult_assoc pth_3[symmetric]
         intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
       by(auto simp add: pos_le_divide_eq pos_divide_le_eq real_mult_commute diff_le_iff)
   } moreover
   { fix y assume "m *\<^sub>R b + c \<le> y" "y \<le> m *\<^sub>R a + c" "m < 0"
     hence "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}"
       unfolding image_iff Bex_def mem_interval vector_le_def
-      apply(auto simp add: vector_smult_component vector_add_component vector_minus_component vector_smult_assoc pth_3[symmetric]
+      apply(auto simp add: vector_smult_assoc pth_3[symmetric]
         intro!: exI[where x="(1 / m) *\<^sub>R (y - c)"])
       by(auto simp add: neg_le_divide_eq neg_divide_le_eq real_mult_commute diff_le_iff)
   }
@@ -5854,11 +5913,11 @@
       also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"
         using cf_z[of "m + k"] and c by auto
       also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"
-        using Suc by (auto simp add: ring_simps)
+        using Suc by (auto simp add: field_simps)
       also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"
-        unfolding power_add by (auto simp add: ring_simps)
+        unfolding power_add by (auto simp add: field_simps)
       also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"
-        using c by (auto simp add: ring_simps)
+        using c by (auto simp add: field_simps)
       finally show ?case by auto
     qed
   } note cf_z2 = this
@@ -6015,7 +6074,7 @@
         apply(erule_tac x="Na+Nb+n" in allE) apply simp
         using dist_triangle_add_half[of a "f (r (Na + Nb + n)) x" "dist a b - dist (f n x) (f n y)"
           "-b"  "- f (r (Na + Nb + n)) y"]
-        unfolding ** unfolding group_simps(12) by (auto simp add: dist_commute)
+        unfolding ** by (auto simp add: algebra_simps dist_commute)
       moreover
       have "dist (f (r (Na + Nb + n)) x - f (r (Na + Nb + n)) y) (a - b) \<ge> dist a b - dist (f n x) (f n y)"
         using distf[of n "r (Na+Nb+n)", OF _ `x\<in>s` `y\<in>s`]
@@ -6045,7 +6104,7 @@
   { fix x y assume "x\<in>s" "y\<in>s" moreover
     fix e::real assume "e>0" ultimately
     have "dist y x < e \<longrightarrow> dist (g y) (g x) < e" using dist by fastsimp }
-  hence "continuous_on s g" unfolding continuous_on_def by auto
+  hence "continuous_on s g" unfolding continuous_on_iff by auto
 
   hence "((snd \<circ> h \<circ> r) ---> g a) sequentially" unfolding continuous_on_sequentially
     apply (rule allE[where x="\<lambda>n. (fst \<circ> h \<circ> r) n"]) apply (erule ballE[where x=a])
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/document/root.tex	Mon Apr 26 21:25:32 2010 +0200
@@ -0,0 +1,31 @@
+
+% HOL/Multivariate_Analysis/document/root.tex
+
+\documentclass[11pt,a4paper]{article}
+\usepackage{graphicx,isabelle,isabellesym,latexsym}
+\usepackage[only,bigsqcap]{stmaryrd}
+\usepackage[latin1]{inputenc}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+\isabellestyle{it}
+\pagestyle{myheadings}
+
+\begin{document}
+
+\title{Multivariate Analysis}
+\maketitle
+
+\tableofcontents
+
+\begin{center}
+  \includegraphics[scale=0.45]{session_graph}
+\end{center}
+
+\renewcommand{\isamarkupheader}[1]%
+{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- a/src/HOL/NSA/HDeriv.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/NSA/HDeriv.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -204,7 +204,7 @@
 apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib)
 apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
 apply (auto simp add: times_divide_eq_right [symmetric]
-            simp del: times_divide_eq)
+            simp del: times_divide_eq_right times_divide_eq_left)
 apply (drule_tac D = Db in lemma_nsderiv2, assumption+)
 apply (drule_tac
      approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]])
--- a/src/HOL/NSA/HyperDef.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/NSA/HyperDef.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -140,12 +140,12 @@
 
 lemma of_hypreal_inverse [simp]:
   "\<And>x. of_hypreal (inverse x) =
-   inverse (of_hypreal x :: 'a::{real_div_algebra,division_by_zero} star)"
+   inverse (of_hypreal x :: 'a::{real_div_algebra,division_ring_inverse_zero} star)"
 by transfer (rule of_real_inverse)
 
 lemma of_hypreal_divide [simp]:
   "\<And>x y. of_hypreal (x / y) =
-   (of_hypreal x / of_hypreal y :: 'a::{real_field,division_by_zero} star)"
+   (of_hypreal x / of_hypreal y :: 'a::{real_field,division_ring_inverse_zero} star)"
 by transfer (rule of_real_divide)
 
 lemma of_hypreal_eq_iff [simp]:
@@ -454,7 +454,7 @@
 by transfer (rule field_power_not_zero)
 
 lemma hyperpow_inverse:
-  "\<And>r n. r \<noteq> (0::'a::{division_by_zero,field} star)
+  "\<And>r n. r \<noteq> (0::'a::{division_ring_inverse_zero,field} star)
    \<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
 by transfer (rule power_inverse)
   
--- a/src/HOL/NSA/NSA.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/NSA/NSA.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -145,12 +145,12 @@
 by transfer (rule nonzero_norm_inverse)
 
 lemma hnorm_inverse:
-  "\<And>a::'a::{real_normed_div_algebra,division_by_zero} star.
+  "\<And>a::'a::{real_normed_div_algebra,division_ring_inverse_zero} star.
    hnorm (inverse a) = inverse (hnorm a)"
 by transfer (rule norm_inverse)
 
 lemma hnorm_divide:
-  "\<And>a b::'a::{real_normed_field,division_by_zero} star.
+  "\<And>a b::'a::{real_normed_field,division_ring_inverse_zero} star.
    hnorm (a / b) = hnorm a / hnorm b"
 by transfer (rule norm_divide)
 
--- a/src/HOL/NSA/StarDef.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/NSA/StarDef.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -902,7 +902,7 @@
 apply (transfer, rule divide_inverse)
 done
 
-instance star :: (division_by_zero) division_by_zero
+instance star :: (division_ring_inverse_zero) division_ring_inverse_zero
 by (intro_classes, transfer, rule inverse_zero)
 
 instance star :: (ordered_semiring) ordered_semiring
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -325,27 +325,27 @@
 
 subsection {* Schematic Variables *}
 
-lemma "x = ?x"
+schematic_lemma "x = ?x"
 nitpick [expect = none]
 by auto
 
-lemma "\<forall>x. x = ?x"
+schematic_lemma "\<forall>x. x = ?x"
 nitpick [expect = genuine]
 oops
 
-lemma "\<exists>x. x = ?x"
+schematic_lemma "\<exists>x. x = ?x"
 nitpick [expect = none]
 by auto
 
-lemma "\<exists>x\<Colon>'a \<Rightarrow> 'b. x = ?x"
+schematic_lemma "\<exists>x\<Colon>'a \<Rightarrow> 'b. x = ?x"
 nitpick [expect = none]
 by auto
 
-lemma "\<forall>x. ?x = ?y"
+schematic_lemma "\<forall>x. ?x = ?y"
 nitpick [expect = none]
 by auto
 
-lemma "\<exists>x. ?x = ?y"
+schematic_lemma "\<exists>x. ?x = ?y"
 nitpick [expect = none]
 by auto
 
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -340,12 +340,12 @@
 
 subsubsection {* Schematic Variables *}
 
-lemma "?P"
+schematic_lemma "?P"
 nitpick [expect = none]
 apply auto
 done
 
-lemma "x = ?y"
+schematic_lemma "x = ?y"
 nitpick [expect = none]
 apply auto
 done
--- a/src/HOL/Nominal/nominal_inductive.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -543,7 +543,7 @@
 
   in
     ctxt'' |>
-    Proof.theorem_i NONE (fn thss => fn ctxt =>
+    Proof.theorem NONE (fn thss => fn ctxt =>
       let
         val rec_name = space_implode "_" (map Long_Name.base_name names);
         val rec_qualified = Binding.qualify false rec_name;
--- a/src/HOL/Nominal/nominal_inductive2.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -445,7 +445,7 @@
 
   in
     ctxt'' |>
-    Proof.theorem_i NONE (fn thss => fn ctxt =>
+    Proof.theorem NONE (fn thss => fn ctxt =>
       let
         val rec_name = space_implode "_" (map Long_Name.base_name names);
         val rec_qualified = Binding.qualify false rec_name;
--- a/src/HOL/Nominal/nominal_primrec.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Nominal/nominal_primrec.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -363,7 +363,7 @@
   in
     lthy' |>
     Variable.add_fixes (map fst lsrs) |> snd |>
-    Proof.theorem_i NONE
+    Proof.theorem NONE
       (fn thss => fn goal_ctxt =>
          let
            val simps = ProofContext.export goal_ctxt lthy' (flat thss);
--- a/src/HOL/Number_Theory/Binomial.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Number_Theory/Binomial.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -208,7 +208,7 @@
   have "fact (k + 1) * fact (n - k) * (n + 1 choose (k + 1)) =
       fact (k + 1) * fact (n - k) * (n choose (k + 1)) + 
       fact (k + 1) * fact (n - k) * (n choose k)" 
-    by (subst choose_reduce_nat, auto simp add: ring_simps)
+    by (subst choose_reduce_nat, auto simp add: field_simps)
   also note one
   also note two
   also with less have "(n - k) * fact n + (k + 1) * fact n= fact (n + 1)" 
@@ -279,7 +279,7 @@
   also have "... = (SUM k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) +
                   (SUM k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))"
     by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le 
-             power_Suc ring_simps One_nat_def del:setsum_cl_ivl_Suc)
+             power_Suc field_simps One_nat_def del:setsum_cl_ivl_Suc)
   also have "... = a^(n+1) + b^(n+1) +
                   (SUM k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) +
                   (SUM k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))"
@@ -287,10 +287,10 @@
   also have
       "... = a^(n+1) + b^(n+1) + 
          (SUM k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
-    by (auto simp add: ring_simps setsum_addf [symmetric]
+    by (auto simp add: field_simps setsum_addf [symmetric]
       choose_reduce_nat)
   also have "... = (SUM k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
-    using decomp by (simp add: ring_simps)
+    using decomp by (simp add: field_simps)
   finally show "?P (n + 1)" by simp
 qed
 
--- a/src/HOL/Number_Theory/Cong.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Number_Theory/Cong.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -350,7 +350,7 @@
   apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
   (* any way around this? *)
   apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
 done
 
 lemma cong_mult_rcancel_int:
@@ -416,7 +416,7 @@
 done
 
 lemma cong_iff_lin_int: "([(a::int) = b] (mod m)) = (\<exists>k. b = a + m * k)"
-  apply (auto simp add: cong_altdef_int dvd_def ring_simps)
+  apply (auto simp add: cong_altdef_int dvd_def field_simps)
   apply (rule_tac [!] x = "-k" in exI, auto)
 done
 
@@ -428,14 +428,14 @@
   apply (unfold dvd_def, auto)
   apply (rule_tac x = k in exI)
   apply (rule_tac x = 0 in exI)
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
   apply (subst (asm) cong_sym_eq_nat)
   apply (subst (asm) cong_altdef_nat)
   apply force
   apply (unfold dvd_def, auto)
   apply (rule_tac x = 0 in exI)
   apply (rule_tac x = k in exI)
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
   apply (unfold cong_nat_def)
   apply (subgoal_tac "a mod m = (a + k2 * m) mod m")
   apply (erule ssubst)back
@@ -533,7 +533,7 @@
   apply (auto simp add: cong_iff_lin_nat dvd_def)
   apply (rule_tac x="k1 * k" in exI)
   apply (rule_tac x="k2 * k" in exI)
-  apply (simp add: ring_simps)
+  apply (simp add: field_simps)
 done
 
 lemma cong_dvd_modulus_int: "[(x::int) = y] (mod m) \<Longrightarrow> n dvd m \<Longrightarrow> 
@@ -559,7 +559,7 @@
 lemma neg_cong_int: "([(a::int) = b] (mod m)) = ([-a = -b] (mod m))"
   apply (simp add: cong_altdef_int)
   apply (subst dvd_minus_iff [symmetric])
-  apply (simp add: ring_simps)
+  apply (simp add: field_simps)
 done
 
 lemma cong_modulus_neg_int: "([(a::int) = b] (mod m)) = ([a = b] (mod -m))"
@@ -603,7 +603,7 @@
   apply (unfold dvd_def)
   apply auto [1]
   apply (rule_tac x = k in exI)
-  apply (auto simp add: ring_simps) [1]
+  apply (auto simp add: field_simps) [1]
   apply (subst cong_altdef_nat)
   apply (auto simp add: dvd_def)
 done
@@ -611,7 +611,7 @@
 lemma cong_le_nat: "(y::nat) <= x \<Longrightarrow> [x = y] (mod n) \<longleftrightarrow> (\<exists>q. x = q * n + y)"
   apply (subst cong_altdef_nat)
   apply assumption
-  apply (unfold dvd_def, auto simp add: ring_simps)
+  apply (unfold dvd_def, auto simp add: field_simps)
   apply (rule_tac x = k in exI)
   apply auto
 done
--- a/src/HOL/Number_Theory/Fib.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Number_Theory/Fib.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -143,9 +143,9 @@
   apply (induct n rule: fib_induct_nat)
   apply auto
   apply (subst fib_reduce_nat)
-  apply (auto simp add: ring_simps)
+  apply (auto simp add: field_simps)
   apply (subst (1 3 5) fib_reduce_nat)
-  apply (auto simp add: ring_simps Suc_eq_plus1)
+  apply (auto simp add: field_simps Suc_eq_plus1)
 (* hmmm. Why doesn't "n + (1 + (1 + k))" simplify to "n + k + 2"? *)
   apply (subgoal_tac "n + (k + 2) = n + (1 + (1 + k))")
   apply (erule ssubst) back back
@@ -184,7 +184,7 @@
 lemma fib_Cassini_aux_int: "fib (int n + 2) * fib (int n) - 
     (fib (int n + 1))^2 = (-1)^(n + 1)"
   apply (induct n)
-  apply (auto simp add: ring_simps power2_eq_square fib_reduce_int
+  apply (auto simp add: field_simps power2_eq_square fib_reduce_int
       power_add)
 done
 
@@ -222,7 +222,7 @@
   apply (subst (2) fib_reduce_nat)
   apply (auto simp add: Suc_eq_plus1) (* again, natdiff_cancel *)
   apply (subst add_commute, auto)
-  apply (subst gcd_commute_nat, auto simp add: ring_simps)
+  apply (subst gcd_commute_nat, auto simp add: field_simps)
 done
 
 lemma coprime_fib_Suc_nat: "coprime (fib n) (fib (Suc n))"
@@ -305,7 +305,7 @@
 theorem fib_mult_eq_setsum_nat:
     "fib ((n::nat) + 1) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
   apply (induct n)
-  apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat ring_simps)
+  apply (auto simp add: atMost_plus_one_nat fib_plus_2_nat field_simps)
 done
 
 theorem fib_mult_eq_setsum'_nat:
--- a/src/HOL/Number_Theory/Residues.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Number_Theory/Residues.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -69,7 +69,7 @@
   apply (subst mod_add_eq [symmetric])
   apply (subst mult_commute)
   apply (subst zmod_zmult1_eq [symmetric])
-  apply (simp add: ring_simps)
+  apply (simp add: field_simps)
 done
 
 end
--- a/src/HOL/Power.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Power.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -400,7 +400,7 @@
 
 text{*Perhaps these should be simprules.*}
 lemma power_inverse:
-  fixes a :: "'a::{division_ring,division_by_zero,power}"
+  fixes a :: "'a::{division_ring,division_ring_inverse_zero,power}"
   shows "inverse (a ^ n) = (inverse a) ^ n"
 apply (cases "a = 0")
 apply (simp add: power_0_left)
@@ -408,11 +408,11 @@
 done (* TODO: reorient or rename to inverse_power *)
 
 lemma power_one_over:
-  "1 / (a::'a::{field,division_by_zero, power}) ^ n =  (1 / a) ^ n"
+  "1 / (a::'a::{field,division_ring_inverse_zero, power}) ^ n =  (1 / a) ^ n"
   by (simp add: divide_inverse) (rule power_inverse)
 
 lemma power_divide:
-  "(a / b) ^ n = (a::'a::{field,division_by_zero}) ^ n / b ^ n"
+  "(a / b) ^ n = (a::'a::{field,division_ring_inverse_zero}) ^ n / b ^ n"
 apply (cases "b = 0")
 apply (simp add: power_0_left)
 apply (rule nonzero_power_divide)
--- a/src/HOL/Prolog/Func.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Prolog/Func.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -58,11 +58,11 @@
 
 lemmas prog_Func = eval
 
-lemma "eval ((S (S Z)) + (S Z)) ?X"
+schematic_lemma "eval ((S (S Z)) + (S Z)) ?X"
   apply (prolog prog_Func)
   done
 
-lemma "eval (app (fix (%fact. abs(%n. cond (n eq Z) (S Z)
+schematic_lemma "eval (app (fix (%fact. abs(%n. cond (n eq Z) (S Z)
                         (n * (app fact (n - (S Z))))))) (S (S (S Z)))) ?X"
   apply (prolog prog_Func)
   done
--- a/src/HOL/Prolog/Test.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Prolog/Test.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -81,7 +81,7 @@
 
 lemmas prog_Test = append reverse mappred mapfun age eq bag_appl
 
-lemma "append ?x ?y [a,b,c,d]"
+schematic_lemma "append ?x ?y [a,b,c,d]"
   apply (prolog prog_Test)
   back
   back
@@ -89,56 +89,56 @@
   back
   done
 
-lemma "append [a,b] y ?L"
+schematic_lemma "append [a,b] y ?L"
   apply (prolog prog_Test)
   done
 
-lemma "!y. append [a,b] y (?L y)"
+schematic_lemma "!y. append [a,b] y (?L y)"
   apply (prolog prog_Test)
   done
 
-lemma "reverse [] ?L"
+schematic_lemma "reverse [] ?L"
   apply (prolog prog_Test)
   done
 
-lemma "reverse [23] ?L"
+schematic_lemma "reverse [23] ?L"
   apply (prolog prog_Test)
   done
 
-lemma "reverse [23,24,?x] ?L"
+schematic_lemma "reverse [23,24,?x] ?L"
   apply (prolog prog_Test)
   done
 
-lemma "reverse ?L [23,24,?x]"
+schematic_lemma "reverse ?L [23,24,?x]"
   apply (prolog prog_Test)
   done
 
-lemma "mappred age ?x [23,24]"
+schematic_lemma "mappred age ?x [23,24]"
   apply (prolog prog_Test)
   back
   done
 
-lemma "mappred (%x y. ? z. age z y) ?x [23,24]"
+schematic_lemma "mappred (%x y. ? z. age z y) ?x [23,24]"
   apply (prolog prog_Test)
   done
 
-lemma "mappred ?P [bob,sue] [24,23]"
+schematic_lemma "mappred ?P [bob,sue] [24,23]"
   apply (prolog prog_Test)
   done
 
-lemma "mapfun f [bob,bob,sue] [?x,?y,?z]"
+schematic_lemma "mapfun f [bob,bob,sue] [?x,?y,?z]"
   apply (prolog prog_Test)
   done
 
-lemma "mapfun (%x. h x 25) [bob,sue] ?L"
+schematic_lemma "mapfun (%x. h x 25) [bob,sue] ?L"
   apply (prolog prog_Test)
   done
 
-lemma "mapfun ?F [24,25] [h bob 24,h bob 25]"
+schematic_lemma "mapfun ?F [24,25] [h bob 24,h bob 25]"
   apply (prolog prog_Test)
   done
 
-lemma "mapfun ?F [24] [h 24 24]"
+schematic_lemma "mapfun ?F [24] [h 24 24]"
   apply (prolog prog_Test)
   back
   back
@@ -149,12 +149,12 @@
   apply (prolog prog_Test)
   done
 
-lemma "age ?x 24 & age ?y 23"
+schematic_lemma "age ?x 24 & age ?y 23"
   apply (prolog prog_Test)
   back
   done
 
-lemma "age ?x 24 | age ?x 23"
+schematic_lemma "age ?x 24 | age ?x 23"
   apply (prolog prog_Test)
   back
   back
@@ -168,7 +168,7 @@
   apply (prolog prog_Test)
   done
 
-lemma "age sue 24 .. age bob 23 => age ?x ?y"
+schematic_lemma "age sue 24 .. age bob 23 => age ?x ?y"
   apply (prolog prog_Test)
   back
   back
@@ -182,7 +182,7 @@
   done
 (*reset trace_DEPTH_FIRST;*)
 
-lemma "(!x. age x 25 :- age x 23) => age ?x 25 & age ?y 25"
+schematic_lemma "(!x. age x 25 :- age x 23) => age ?x 25 & age ?y 25"
   apply (prolog prog_Test)
   back
   back
@@ -193,7 +193,7 @@
   apply (prolog prog_Test)
   done
 
-lemma "? P. P & eq P ?x"
+schematic_lemma "? P. P & eq P ?x"
   apply (prolog prog_Test)
 (*
   back
@@ -248,14 +248,14 @@
   by fast
 *)
 
-lemma "!Emp Stk.(
+schematic_lemma "!Emp Stk.(
                        empty    (Emp::'b) .. 
          (!(X::nat) S. add    X (S::'b)         (Stk X S)) .. 
          (!(X::nat) S. remove X ((Stk X S)::'b) S))
  => bag_appl 23 24 ?X ?Y"
   oops
 
-lemma "!Qu. ( 
+schematic_lemma "!Qu. ( 
           (!L.            empty    (Qu L L)) .. 
           (!(X::nat) L K. add    X (Qu L (X#K)) (Qu L K)) ..
           (!(X::nat) L K. remove X (Qu (X#L) K) (Qu L K)))
@@ -266,7 +266,7 @@
   apply (prolog prog_Test)
   done
 
-lemma "P x .. P y => P ?X"
+schematic_lemma "P x .. P y => P ?X"
   apply (prolog prog_Test)
   back
   done
--- a/src/HOL/Prolog/Type.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Prolog/Type.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -47,41 +47,41 @@
 
 lemmas prog_Type = prog_Func good_typeof common_typeof
 
-lemma "typeof (abs(%n. abs(%m. abs(%p. p and (n eq m))))) ?T"
+schematic_lemma "typeof (abs(%n. abs(%m. abs(%p. p and (n eq m))))) ?T"
   apply (prolog prog_Type)
   done
 
-lemma "typeof (fix (%x. x)) ?T"
+schematic_lemma "typeof (fix (%x. x)) ?T"
   apply (prolog prog_Type)
   done
 
-lemma "typeof (fix (%fact. abs(%n. (app fact (n - Z))))) ?T"
+schematic_lemma "typeof (fix (%fact. abs(%n. (app fact (n - Z))))) ?T"
   apply (prolog prog_Type)
   done
 
-lemma "typeof (fix (%fact. abs(%n. cond (n eq Z) (S Z)
+schematic_lemma "typeof (fix (%fact. abs(%n. cond (n eq Z) (S Z)
   (n * (app fact (n - (S Z))))))) ?T"
   apply (prolog prog_Type)
   done
 
-lemma "typeof (abs(%v. Z)) ?T" (*correct only solution (?A1 -> nat) *)
+schematic_lemma "typeof (abs(%v. Z)) ?T" (*correct only solution (?A1 -> nat) *)
   apply (prolog prog_Type)
   done
 
-lemma "typeof (abs(%v. Z)) ?T"
+schematic_lemma "typeof (abs(%v. Z)) ?T"
   apply (prolog bad1_typeof common_typeof) (* 1st result ok*)
   done
 
-lemma "typeof (abs(%v. Z)) ?T"
+schematic_lemma "typeof (abs(%v. Z)) ?T"
   apply (prolog bad1_typeof common_typeof)
   back (* 2nd result (?A1 -> ?A1) wrong *)
   done
 
-lemma "typeof (abs(%v. abs(%v. app v v))) ?T"
+schematic_lemma "typeof (abs(%v. abs(%v. app v v))) ?T"
   apply (prolog prog_Type)?  (*correctly fails*)
   oops
 
-lemma "typeof (abs(%v. abs(%v. app v v))) ?T"
+schematic_lemma "typeof (abs(%v. abs(%v. app v v))) ?T"
   apply (prolog bad2_typeof common_typeof) (* wrong result ((?A3 -> ?B3) -> ?A3 -> ?B3)*)
   done
 
--- a/src/HOL/Quotient_Examples/FSet.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Quotient_Examples/FSet.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -387,7 +387,7 @@
   apply (simp add: memb_def[symmetric] memb_finter_raw)
   by (auto simp add: memb_def)
 
-instantiation fset :: (type) "{bot,distrib_lattice}"
+instantiation fset :: (type) "{bounded_lattice_bot,distrib_lattice}"
 begin
 
 quotient_definition
@@ -922,14 +922,13 @@
 
 text {* funion *}
 
-lemma funion_simps[simp]:
-  shows "{||} |\<union>| S = S"
-  and   "finsert x S |\<union>| T = finsert x (S |\<union>| T)"
-  by (lifting append.simps)
+lemmas [simp] =
+  sup_bot_left[where 'a="'a fset"]
+  sup_bot_right[where 'a="'a fset"]
 
-lemma funion_empty[simp]:
-  shows "S |\<union>| {||} = S"
-  by (lifting append_Nil2)
+lemma funion_finsert[simp]:
+  shows "finsert x S |\<union>| T = finsert x (S |\<union>| T)"
+  by (lifting append.simps(2))
 
 lemma singleton_union_left:
   "{|a|} |\<union>| S = finsert a S"
--- a/src/HOL/Rat.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Rat.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -411,7 +411,7 @@
 
 subsubsection {* The field of rational numbers *}
 
-instantiation rat :: "{field, division_by_zero}"
+instantiation rat :: field
 begin
 
 definition
@@ -433,9 +433,6 @@
   by (simp add: divide_rat_def)
 
 instance proof
-  show "inverse 0 = (0::rat)" by (simp add: rat_number_expand)
-    (simp add: rat_number_collapse)
-next
   fix q :: rat
   assume "q \<noteq> 0"
   then show "inverse q * q = 1" by (cases q rule: Rat_cases_nonzero)
@@ -447,6 +444,9 @@
 
 end
 
+instance rat :: division_ring_inverse_zero proof
+qed (simp add: rat_number_expand, simp add: rat_number_collapse)
+
 
 subsubsection {* Various *}
 
@@ -818,7 +818,7 @@
 done
 
 lemma of_rat_inverse:
-  "(of_rat (inverse a)::'a::{field_char_0,division_by_zero}) =
+  "(of_rat (inverse a)::'a::{field_char_0,division_ring_inverse_zero}) =
    inverse (of_rat a)"
 by (cases "a = 0", simp_all add: nonzero_of_rat_inverse)
 
@@ -827,7 +827,7 @@
 by (simp add: divide_inverse of_rat_mult nonzero_of_rat_inverse)
 
 lemma of_rat_divide:
-  "(of_rat (a / b)::'a::{field_char_0,division_by_zero})
+  "(of_rat (a / b)::'a::{field_char_0,division_ring_inverse_zero})
    = of_rat a / of_rat b"
 by (cases "b = 0") (simp_all add: nonzero_of_rat_divide)
 
@@ -968,7 +968,7 @@
 done
 
 lemma Rats_inverse [simp]:
-  fixes a :: "'a::{field_char_0,division_by_zero}"
+  fixes a :: "'a::{field_char_0,division_ring_inverse_zero}"
   shows "a \<in> Rats \<Longrightarrow> inverse a \<in> Rats"
 apply (auto simp add: Rats_def)
 apply (rule range_eqI)
@@ -984,7 +984,7 @@
 done
 
 lemma Rats_divide [simp]:
-  fixes a b :: "'a::{field_char_0,division_by_zero}"
+  fixes a b :: "'a::{field_char_0,division_ring_inverse_zero}"
   shows "\<lbrakk>a \<in> Rats; b \<in> Rats\<rbrakk> \<Longrightarrow> a / b \<in> Rats"
 apply (auto simp add: Rats_def)
 apply (rule range_eqI)
--- a/src/HOL/RealDef.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/RealDef.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -279,7 +279,7 @@
 lemma INVERSE_ZERO: "inverse 0 = (0::real)"
 by (simp add: real_inverse_def)
 
-instance real :: division_by_zero
+instance real :: division_ring_inverse_zero
 proof
   show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
 qed
--- a/src/HOL/RealVector.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/RealVector.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -207,7 +207,7 @@
 by (rule inverse_unique, simp)
 
 lemma inverse_scaleR_distrib:
-  fixes x :: "'a::{real_div_algebra,division_by_zero}"
+  fixes x :: "'a::{real_div_algebra,division_ring_inverse_zero}"
   shows "inverse (scaleR a x) = scaleR (inverse a) (inverse x)"
 apply (case_tac "a = 0", simp)
 apply (case_tac "x = 0", simp)
@@ -250,7 +250,7 @@
 
 lemma of_real_inverse [simp]:
   "of_real (inverse x) =
-   inverse (of_real x :: 'a::{real_div_algebra,division_by_zero})"
+   inverse (of_real x :: 'a::{real_div_algebra,division_ring_inverse_zero})"
 by (simp add: of_real_def inverse_scaleR_distrib)
 
 lemma nonzero_of_real_divide:
@@ -260,7 +260,7 @@
 
 lemma of_real_divide [simp]:
   "of_real (x / y) =
-   (of_real x / of_real y :: 'a::{real_field,division_by_zero})"
+   (of_real x / of_real y :: 'a::{real_field,division_ring_inverse_zero})"
 by (simp add: divide_inverse)
 
 lemma of_real_power [simp]:
@@ -370,7 +370,7 @@
 done
 
 lemma Reals_inverse [simp]:
-  fixes a :: "'a::{real_div_algebra,division_by_zero}"
+  fixes a :: "'a::{real_div_algebra,division_ring_inverse_zero}"
   shows "a \<in> Reals \<Longrightarrow> inverse a \<in> Reals"
 apply (auto simp add: Reals_def)
 apply (rule range_eqI)
@@ -386,7 +386,7 @@
 done
 
 lemma Reals_divide [simp]:
-  fixes a b :: "'a::{real_field,division_by_zero}"
+  fixes a b :: "'a::{real_field,division_ring_inverse_zero}"
   shows "\<lbrakk>a \<in> Reals; b \<in> Reals\<rbrakk> \<Longrightarrow> a / b \<in> Reals"
 apply (auto simp add: Reals_def)
 apply (rule range_eqI)
@@ -726,7 +726,7 @@
 done
 
 lemma norm_inverse:
-  fixes a :: "'a::{real_normed_div_algebra,division_by_zero}"
+  fixes a :: "'a::{real_normed_div_algebra,division_ring_inverse_zero}"
   shows "norm (inverse a) = inverse (norm a)"
 apply (case_tac "a = 0", simp)
 apply (erule nonzero_norm_inverse)
@@ -738,7 +738,7 @@
 by (simp add: divide_inverse norm_mult nonzero_norm_inverse)
 
 lemma norm_divide:
-  fixes a b :: "'a::{real_normed_field,division_by_zero}"
+  fixes a b :: "'a::{real_normed_field,division_ring_inverse_zero}"
   shows "norm (a / b) = norm a / norm b"
 by (simp add: divide_inverse norm_mult norm_inverse)
 
--- a/src/HOL/Rings.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Rings.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -14,8 +14,8 @@
 begin
 
 class semiring = ab_semigroup_add + semigroup_mult +
-  assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c"
-  assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c"
+  assumes left_distrib[algebra_simps, field_simps]: "(a + b) * c = a * c + b * c"
+  assumes right_distrib[algebra_simps, field_simps]: "a * (b + c) = a * b + a * c"
 begin
 
 text{*For the @{text combine_numerals} simproc*}
@@ -230,18 +230,15 @@
 lemma minus_mult_commute: "- a * b = a * - b"
 by simp
 
-lemma right_diff_distrib[algebra_simps]: "a * (b - c) = a * b - a * c"
+lemma right_diff_distrib[algebra_simps, field_simps]: "a * (b - c) = a * b - a * c"
 by (simp add: right_distrib diff_minus)
 
-lemma left_diff_distrib[algebra_simps]: "(a - b) * c = a * c - b * c"
+lemma left_diff_distrib[algebra_simps, field_simps]: "(a - b) * c = a * c - b * c"
 by (simp add: left_distrib diff_minus)
 
 lemmas ring_distribs[no_atp] =
   right_distrib left_distrib left_diff_distrib right_diff_distrib
 
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[no_atp] = algebra_simps
-
 lemma eq_add_iff1:
   "a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
 by (simp add: algebra_simps)
@@ -487,6 +484,125 @@
   "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse a - inverse b = inverse a * (b - a) * inverse b"
 by (simp add: algebra_simps)
 
+lemma right_inverse_eq: "b \<noteq> 0 \<Longrightarrow> a / b = 1 \<longleftrightarrow> a = b"
+proof
+  assume neq: "b \<noteq> 0"
+  {
+    hence "a = (a / b) * b" by (simp add: divide_inverse mult_assoc)
+    also assume "a / b = 1"
+    finally show "a = b" by simp
+  next
+    assume "a = b"
+    with neq show "a / b = 1" by (simp add: divide_inverse)
+  }
+qed
+
+lemma nonzero_inverse_eq_divide: "a \<noteq> 0 \<Longrightarrow> inverse a = 1 / a"
+by (simp add: divide_inverse)
+
+lemma divide_self [simp]: "a \<noteq> 0 \<Longrightarrow> a / a = 1"
+by (simp add: divide_inverse)
+
+lemma divide_zero_left [simp]: "0 / a = 0"
+by (simp add: divide_inverse)
+
+lemma inverse_eq_divide: "inverse a = 1 / a"
+by (simp add: divide_inverse)
+
+lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
+by (simp add: divide_inverse algebra_simps)
+
+lemma divide_1 [simp]: "a / 1 = a"
+  by (simp add: divide_inverse)
+
+lemma times_divide_eq_right [simp]: "a * (b / c) = (a * b) / c"
+  by (simp add: divide_inverse mult_assoc)
+
+lemma minus_divide_left: "- (a / b) = (-a) / b"
+  by (simp add: divide_inverse)
+
+lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
+  by (simp add: divide_inverse nonzero_inverse_minus_eq)
+
+lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
+  by (simp add: divide_inverse nonzero_inverse_minus_eq)
+
+lemma divide_minus_left [simp, no_atp]: "(-a) / b = - (a / b)"
+  by (simp add: divide_inverse)
+
+lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
+  by (simp add: diff_minus add_divide_distrib)
+
+lemma nonzero_eq_divide_eq [field_simps]: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
+proof -
+  assume [simp]: "c \<noteq> 0"
+  have "a = b / c \<longleftrightarrow> a * c = (b / c) * c" by simp
+  also have "... \<longleftrightarrow> a * c = b" by (simp add: divide_inverse mult_assoc)
+  finally show ?thesis .
+qed
+
+lemma nonzero_divide_eq_eq [field_simps]: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
+proof -
+  assume [simp]: "c \<noteq> 0"
+  have "b / c = a \<longleftrightarrow> (b / c) * c = a * c" by simp
+  also have "... \<longleftrightarrow> b = a * c" by (simp add: divide_inverse mult_assoc) 
+  finally show ?thesis .
+qed
+
+lemma divide_eq_imp: "c \<noteq> 0 \<Longrightarrow> b = a * c \<Longrightarrow> b / c = a"
+  by (simp add: divide_inverse mult_assoc)
+
+lemma eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
+  by (drule sym) (simp add: divide_inverse mult_assoc)
+
+end
+
+class division_ring_inverse_zero = division_ring +
+  assumes inverse_zero [simp]: "inverse 0 = 0"
+begin
+
+lemma divide_zero [simp]:
+  "a / 0 = 0"
+  by (simp add: divide_inverse)
+
+lemma divide_self_if [simp]:
+  "a / a = (if a = 0 then 0 else 1)"
+  by simp
+
+lemma inverse_nonzero_iff_nonzero [simp]:
+  "inverse a = 0 \<longleftrightarrow> a = 0"
+  by rule (fact inverse_zero_imp_zero, simp)
+
+lemma inverse_minus_eq [simp]:
+  "inverse (- a) = - inverse a"
+proof cases
+  assume "a=0" thus ?thesis by simp
+next
+  assume "a\<noteq>0" 
+  thus ?thesis by (simp add: nonzero_inverse_minus_eq)
+qed
+
+lemma inverse_eq_imp_eq:
+  "inverse a = inverse b \<Longrightarrow> a = b"
+apply (cases "a=0 | b=0") 
+ apply (force dest!: inverse_zero_imp_zero
+              simp add: eq_commute [of "0::'a"])
+apply (force dest!: nonzero_inverse_eq_imp_eq) 
+done
+
+lemma inverse_eq_iff_eq [simp]:
+  "inverse a = inverse b \<longleftrightarrow> a = b"
+  by (force dest!: inverse_eq_imp_eq)
+
+lemma inverse_inverse_eq [simp]:
+  "inverse (inverse a) = a"
+proof cases
+  assume "a=0" thus ?thesis by simp
+next
+  assume "a\<noteq>0" 
+  thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
+qed
+
 end
 
 class mult_mono = times + zero + ord +
@@ -533,17 +649,17 @@
 subclass ordered_semiring ..
 
 lemma mult_nonneg_nonneg: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
-using mult_left_mono [of zero b a] by simp
+using mult_left_mono [of 0 b a] by simp
 
 lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
-using mult_left_mono [of b zero a] by simp
+using mult_left_mono [of b 0 a] by simp
 
 lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
-using mult_right_mono [of a zero b] by simp
+using mult_right_mono [of a 0 b] by simp
 
 text {* Legacy - use @{text mult_nonpos_nonneg} *}
 lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0" 
-by (drule mult_right_mono [of b zero], auto)
+by (drule mult_right_mono [of b 0], auto)
 
 lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> 0" 
 by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
@@ -597,17 +713,17 @@
 by (force simp add: mult_strict_right_mono not_less [symmetric])
 
 lemma mult_pos_pos: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
-using mult_strict_left_mono [of zero b a] by simp
+using mult_strict_left_mono [of 0 b a] by simp
 
 lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
-using mult_strict_left_mono [of b zero a] by simp
+using mult_strict_left_mono [of b 0 a] by simp
 
 lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
-using mult_strict_right_mono [of a zero b] by simp
+using mult_strict_right_mono [of a 0 b] by simp
 
 text {* Legacy - use @{text mult_neg_pos} *}
 lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0" 
-by (drule mult_strict_right_mono [of b zero], auto)
+by (drule mult_strict_right_mono [of b 0], auto)
 
 lemma zero_less_mult_pos:
   "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
@@ -742,9 +858,6 @@
 
 subclass ordered_ab_group_add ..
 
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[no_atp] = algebra_simps
-
 lemma less_add_iff1:
   "a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
 by (simp add: algebra_simps)
@@ -763,18 +876,18 @@
 
 lemma mult_left_mono_neg:
   "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
-  apply (drule mult_left_mono [of _ _ "uminus c"])
+  apply (drule mult_left_mono [of _ _ "- c"])
   apply simp_all
   done
 
 lemma mult_right_mono_neg:
   "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a * c \<le> b * c"
-  apply (drule mult_right_mono [of _ _ "uminus c"])
+  apply (drule mult_right_mono [of _ _ "- c"])
   apply simp_all
   done
 
 lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
-using mult_right_mono_neg [of a zero b] by simp
+using mult_right_mono_neg [of a 0 b] by simp
 
 lemma split_mult_pos_le:
   "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
@@ -821,7 +934,7 @@
 using mult_strict_right_mono [of b a "- c"] by simp
 
 lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
-using mult_strict_right_mono_neg [of a zero b] by simp
+using mult_strict_right_mono_neg [of a 0 b] by simp
 
 subclass ring_no_zero_divisors
 proof
@@ -949,9 +1062,6 @@
 
 end
 
-text{*Legacy - use @{text algebra_simps} *}
-lemmas ring_simps[no_atp] = algebra_simps
-
 lemmas mult_sign_intros =
   mult_nonneg_nonneg mult_nonneg_nonpos
   mult_nonpos_nonneg mult_nonpos_nonpos
@@ -973,7 +1083,7 @@
 
 lemma pos_add_strict:
   shows "0 < a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
-  using add_strict_mono [of zero a b c] by simp
+  using add_strict_mono [of 0 a b c] by simp
 
 lemma zero_le_one [simp]: "0 \<le> 1"
 by (rule zero_less_one [THEN less_imp_le]) 
@@ -1074,7 +1184,7 @@
   "sgn (a * b) = sgn a * sgn b"
 by (auto simp add: sgn_if zero_less_mult_iff)
 
-lemma abs_sgn: "abs k = k * sgn k"
+lemma abs_sgn: "\<bar>k\<bar> = k * sgn k"
 unfolding sgn_if abs_if by auto
 
 lemma sgn_greater [simp]:
@@ -1085,14 +1195,14 @@
   "sgn a < 0 \<longleftrightarrow> a < 0"
   unfolding sgn_if by auto
 
-lemma abs_dvd_iff [simp]: "(abs m) dvd k \<longleftrightarrow> m dvd k"
+lemma abs_dvd_iff [simp]: "\<bar>m\<bar> dvd k \<longleftrightarrow> m dvd k"
   by (simp add: abs_if)
 
-lemma dvd_abs_iff [simp]: "m dvd (abs k) \<longleftrightarrow> m dvd k"
+lemma dvd_abs_iff [simp]: "m dvd \<bar>k\<bar> \<longleftrightarrow> m dvd k"
   by (simp add: abs_if)
 
 lemma dvd_if_abs_eq:
-  "abs l = abs (k) \<Longrightarrow> l dvd k"
+  "\<bar>l\<bar> = \<bar>k\<bar> \<Longrightarrow> l dvd k"
 by(subst abs_dvd_iff[symmetric]) simp
 
 end
@@ -1110,17 +1220,7 @@
     mult_cancel_right1 mult_cancel_right2
     mult_cancel_left1 mult_cancel_left2
 
--- {* FIXME continue localization here *}
-
-subsection {* Reasoning about inequalities with division *}
-
-lemma mult_right_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1
-    ==> x * y <= x"
-by (auto simp add: mult_le_cancel_left2)
-
-lemma mult_left_le_one_le: "0 <= (x::'a::linordered_idom) ==> 0 <= y ==> y <= 1
-    ==> y * x <= x"
-by (auto simp add: mult_le_cancel_right2)
+text {* Reasoning about inequalities with division *}
 
 context linordered_semidom
 begin
@@ -1137,20 +1237,34 @@
 
 end
 
+context linordered_idom
+begin
 
-subsection {* Absolute Value *}
+lemma mult_right_le_one_le:
+  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> x * y \<le> x"
+  by (auto simp add: mult_le_cancel_left2)
+
+lemma mult_left_le_one_le:
+  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> y * x \<le> x"
+  by (auto simp add: mult_le_cancel_right2)
+
+end
+
+text {* Absolute Value *}
 
 context linordered_idom
 begin
 
-lemma mult_sgn_abs: "sgn x * abs x = x"
+lemma mult_sgn_abs:
+  "sgn x * \<bar>x\<bar> = x"
   unfolding abs_if sgn_if by auto
 
+lemma abs_one [simp]:
+  "\<bar>1\<bar> = 1"
+  by (simp add: abs_if zero_less_one [THEN less_not_sym])
+
 end
 
-lemma abs_one [simp]: "abs 1 = (1::'a::linordered_idom)"
-by (simp add: abs_if zero_less_one [THEN order_less_not_sym])
-
 class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs +
   assumes abs_eq_mult:
     "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0) \<Longrightarrow> \<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
@@ -1162,39 +1276,35 @@
 qed (auto simp add: abs_if not_less mult_less_0_iff)
 
 lemma abs_mult:
-  "abs (a * b) = abs a * abs b" 
+  "\<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>" 
   by (rule abs_eq_mult) auto
 
 lemma abs_mult_self:
-  "abs a * abs a = a * a"
+  "\<bar>a\<bar> * \<bar>a\<bar> = a * a"
   by (simp add: abs_if) 
 
-end
-
 lemma abs_mult_less:
-     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::linordered_idom)"
+  "\<bar>a\<bar> < c \<Longrightarrow> \<bar>b\<bar> < d \<Longrightarrow> \<bar>a\<bar> * \<bar>b\<bar> < c * d"
 proof -
-  assume ac: "abs a < c"
-  hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
-  assume "abs b < d"
+  assume ac: "\<bar>a\<bar> < c"
+  hence cpos: "0<c" by (blast intro: le_less_trans abs_ge_zero)
+  assume "\<bar>b\<bar> < d"
   thus ?thesis by (simp add: ac cpos mult_strict_mono) 
 qed
 
-lemmas eq_minus_self_iff[no_atp] = equal_neg_zero
-
-lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::linordered_idom))"
-  unfolding order_less_le less_eq_neg_nonpos equal_neg_zero ..
+lemma less_minus_self_iff:
+  "a < - a \<longleftrightarrow> a < 0"
+  by (simp only: less_le less_eq_neg_nonpos equal_neg_zero)
 
-lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::linordered_idom))" 
-apply (simp add: order_less_le abs_le_iff)  
-apply (auto simp add: abs_if)
-done
+lemma abs_less_iff:
+  "\<bar>a\<bar> < b \<longleftrightarrow> a < b \<and> - a < b" 
+  by (simp add: less_le abs_le_iff) (auto simp add: abs_if)
 
-lemma abs_mult_pos: "(0::'a::linordered_idom) <= x ==> 
-    (abs y) * x = abs (y * x)"
-  apply (subst abs_mult)
-  apply simp
-done
+lemma abs_mult_pos:
+  "0 \<le> x \<Longrightarrow> \<bar>y\<bar> * x = \<bar>y * x\<bar>"
+  by (simp add: abs_mult)
+
+end
 
 code_modulename SML
   Rings Arith
--- a/src/HOL/SMT/Tools/z3_proof_rules.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/SMT/Tools/z3_proof_rules.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -1137,7 +1137,8 @@
   handle THM _ => NONE
 in
 val z3_simpset = HOL_ss addsimps @{thms array_rules}
-  addsimps @{thms ring_distribs} addsimps @{thms field_eq_simps}
+  addsimps @{thms ring_distribs} addsimps @{thms field_simps}
+  addsimps [@{thm times_divide_eq_right}, @{thm times_divide_eq_left}]
   addsimps @{thms arith_special} addsimps @{thms less_bin_simps}
   addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps}
   addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps}
--- a/src/HOL/SMT/Z3.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/SMT/Z3.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -19,7 +19,7 @@
 
 lemmas [z3_rewrite] =
   refl eq_commute conj_commute disj_commute simp_thms nnf_simps
-  ring_distribs field_eq_simps if_True if_False
+  ring_distribs field_simps times_divide_eq_right times_divide_eq_left if_True if_False
 
 lemma [z3_rewrite]: "(P \<noteq> Q) = (Q = (\<not>P))" by fast
 
--- a/src/HOL/Series.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Series.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -381,7 +381,7 @@
   shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"
 by (rule geometric_sums [THEN sums_summable])
 
-lemma half: "0 < 1 / (2::'a::{number_ring,division_by_zero,linordered_field})"
+lemma half: "0 < 1 / (2::'a::{number_ring,division_ring_inverse_zero,linordered_field})"
   by arith 
 
 lemma power_half_series: "(\<lambda>n. (1/2::real)^Suc n) sums 1"
--- a/src/HOL/SetInterval.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/SetInterval.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -54,22 +54,22 @@
 @{term"{m..<n}"} may not exist in @{term"{..<n}"}-form as well. *}
 
 syntax
-  "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3UN _<=_./ _)" 10)
-  "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3UN _<_./ _)" 10)
-  "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3INT _<=_./ _)" 10)
-  "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3INT _<_./ _)" 10)
+  "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3UN _<=_./ _)" [0, 0, 10] 10)
+  "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3UN _<_./ _)" [0, 0, 10] 10)
+  "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3INT _<=_./ _)" [0, 0, 10] 10)
+  "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3INT _<_./ _)" [0, 0, 10] 10)
 
 syntax (xsymbols)
-  "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _\<le>_./ _)" 10)
-  "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _<_./ _)" 10)
-  "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _\<le>_./ _)" 10)
-  "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _<_./ _)" 10)
+  "_UNION_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _\<le>_./ _)" [0, 0, 10] 10)
+  "_UNION_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Union> _<_./ _)" [0, 0, 10] 10)
+  "_INTER_le"   :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _\<le>_./ _)" [0, 0, 10] 10)
+  "_INTER_less" :: "'a => 'a => 'b set => 'b set"       ("(3\<Inter> _<_./ _)" [0, 0, 10] 10)
 
 syntax (latex output)
-  "_UNION_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ \<le> _)/ _)" 10)
-  "_UNION_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ < _)/ _)" 10)
-  "_INTER_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ \<le> _)/ _)" 10)
-  "_INTER_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ < _)/ _)" 10)
+  "_UNION_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ \<le> _)/ _)" [0, 0, 10] 10)
+  "_UNION_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00_ < _)/ _)" [0, 0, 10] 10)
+  "_INTER_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ \<le> _)/ _)" [0, 0, 10] 10)
+  "_INTER_less" :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Inter>(00_ < _)/ _)" [0, 0, 10] 10)
 
 translations
   "UN i<=n. A"  == "UN i:{..n}. A"
@@ -1012,7 +1012,7 @@
 qed
 
 lemma setsum_le_included:
-  fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add,ordered_ab_semigroup_add_imp_le}"
+  fixes f :: "'a \<Rightarrow> 'b::ordered_comm_monoid_add"
   assumes "finite s" "finite t"
   and "\<forall>y\<in>t. 0 \<le> g y" "(\<forall>x\<in>s. \<exists>y\<in>t. i y = x \<and> f x \<le> g y)"
   shows "setsum f s \<le> setsum g t"
@@ -1085,9 +1085,21 @@
 subsection {* The formula for geometric sums *}
 
 lemma geometric_sum:
-  "x ~= 1 ==> (\<Sum>i=0..<n. x ^ i) =
-  (x ^ n - 1) / (x - 1::'a::{field})"
-by (induct "n") (simp_all add: field_simps)
+  assumes "x \<noteq> 1"
+  shows "(\<Sum>i=0..<n. x ^ i) = (x ^ n - 1) / (x - 1::'a::field)"
+proof -
+  from assms obtain y where "y = x - 1" and "y \<noteq> 0" by simp_all
+  moreover have "(\<Sum>i=0..<n. (y + 1) ^ i) = ((y + 1) ^ n - 1) / y"
+  proof (induct n)
+    case 0 then show ?case by simp
+  next
+    case (Suc n)
+    moreover with `y \<noteq> 0` have "(1 + y) ^ n = (y * inverse y) * (1 + y) ^ n" by simp 
+    ultimately show ?case by (simp add: field_simps divide_inverse)
+  qed
+  ultimately show ?thesis by simp
+qed
+
 
 subsection {* The formula for arithmetic sums *}
 
--- a/src/HOL/Tools/Datatype/datatype_codegen.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_codegen.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -12,6 +12,137 @@
 structure Datatype_Codegen : DATATYPE_CODEGEN =
 struct
 
+(** generic code generator **)
+
+(* liberal addition of code data for datatypes *)
+
+fun mk_constr_consts thy vs dtco cos =
+  let
+    val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
+    val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
+  in if is_some (try (Code.constrset_of_consts thy) cs')
+    then SOME cs
+    else NONE
+  end;
+
+
+(* case certificates *)
+
+fun mk_case_cert thy tyco =
+  let
+    val raw_thms =
+      (#case_rewrites o Datatype_Data.the_info thy) tyco;
+    val thms as hd_thm :: _ = raw_thms
+      |> Conjunction.intr_balanced
+      |> Thm.unvarify_global
+      |> Conjunction.elim_balanced (length raw_thms)
+      |> map Simpdata.mk_meta_eq
+      |> map Drule.zero_var_indexes
+    val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
+      | _ => I) (Thm.prop_of hd_thm) [];
+    val rhs = hd_thm
+      |> Thm.prop_of
+      |> Logic.dest_equals
+      |> fst
+      |> Term.strip_comb
+      |> apsnd (fst o split_last)
+      |> list_comb;
+    val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
+    val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
+  in
+    thms
+    |> Conjunction.intr_balanced
+    |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
+    |> Thm.implies_intr asm
+    |> Thm.generalize ([], params) 0
+    |> AxClass.unoverload thy
+    |> Thm.varifyT_global
+  end;
+
+
+(* equality *)
+
+fun mk_eq_eqns thy dtco =
+  let
+    val (vs, cos) = Datatype_Data.the_spec thy dtco;
+    val { descr, index, inject = inject_thms, distinct = distinct_thms, ... } =
+      Datatype_Data.the_info thy dtco;
+    val ty = Type (dtco, map TFree vs);
+    fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
+      $ t1 $ t2;
+    fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
+    fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
+    val triv_injects = map_filter
+     (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
+       | _ => NONE) cos;
+    fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
+      trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
+    val injects = map prep_inject (nth (Datatype_Prop.make_injs [descr] vs) index);
+    fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
+      [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
+    val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index));
+    val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
+    val simpset = Simplifier.global_context thy (HOL_basic_ss addsimps 
+      (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
+    fun prove prop = Skip_Proof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
+      |> Simpdata.mk_eq;
+  in (map prove (triv_injects @ injects @ distincts), prove refl) end;
+
+fun add_equality vs dtcos thy =
+  let
+    fun add_def dtco lthy =
+      let
+        val ty = Type (dtco, map TFree vs);
+        fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
+          $ Free ("x", ty) $ Free ("y", ty);
+        val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
+          (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
+        val def' = Syntax.check_term lthy def;
+        val ((_, (_, thm)), lthy') = Specification.definition
+          (NONE, (Attrib.empty_binding, def')) lthy;
+        val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
+        val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
+      in (thm', lthy') end;
+    fun tac thms = Class.intro_classes_tac []
+      THEN ALLGOALS (ProofContext.fact_tac thms);
+    fun add_eq_thms dtco =
+      Theory.checkpoint
+      #> `(fn thy => mk_eq_eqns thy dtco)
+      #-> (fn (thms, thm) =>
+        Code.add_nbe_eqn thm
+        #> fold_rev Code.add_eqn thms);
+  in
+    thy
+    |> Theory_Target.instantiation (dtcos, vs, [HOLogic.class_eq])
+    |> fold_map add_def dtcos
+    |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
+         (fn _ => fn def_thms => tac def_thms) def_thms)
+    |-> (fn def_thms => fold Code.del_eqn def_thms)
+    |> fold add_eq_thms dtcos
+  end;
+
+
+(* register a datatype etc. *)
+
+fun add_all_code config dtcos thy =
+  let
+    val (vs :: _, coss) = (split_list o map (Datatype_Data.the_spec thy)) dtcos;
+    val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
+    val css = if exists is_none any_css then []
+      else map_filter I any_css;
+    val case_rewrites = maps (#case_rewrites o Datatype_Data.the_info thy) dtcos;
+    val certs = map (mk_case_cert thy) dtcos;
+  in
+    if null css then thy
+    else thy
+      |> tap (fn _ => Datatype_Aux.message config "Registering datatype for code generator ...")
+      |> fold Code.add_datatype css
+      |> fold_rev Code.add_default_eqn case_rewrites
+      |> fold Code.add_case certs
+      |> add_equality vs dtcos
+   end;
+
+
 (** SML code generator **)
 
 open Codegen;
@@ -288,142 +419,11 @@
   | datatype_tycodegen _ _ _ _ _ _ _ = NONE;
 
 
-(** generic code generator **)
-
-(* liberal addition of code data for datatypes *)
-
-fun mk_constr_consts thy vs dtco cos =
-  let
-    val cs = map (fn (c, tys) => (c, tys ---> Type (dtco, map TFree vs))) cos;
-    val cs' = map (fn c_ty as (_, ty) => (AxClass.unoverload_const thy c_ty, ty)) cs;
-  in if is_some (try (Code.constrset_of_consts thy) cs')
-    then SOME cs
-    else NONE
-  end;
-
-
-(* case certificates *)
-
-fun mk_case_cert thy tyco =
-  let
-    val raw_thms =
-      (#case_rewrites o Datatype_Data.the_info thy) tyco;
-    val thms as hd_thm :: _ = raw_thms
-      |> Conjunction.intr_balanced
-      |> Thm.unvarify_global
-      |> Conjunction.elim_balanced (length raw_thms)
-      |> map Simpdata.mk_meta_eq
-      |> map Drule.zero_var_indexes
-    val params = fold_aterms (fn (Free (v, _)) => insert (op =) v
-      | _ => I) (Thm.prop_of hd_thm) [];
-    val rhs = hd_thm
-      |> Thm.prop_of
-      |> Logic.dest_equals
-      |> fst
-      |> Term.strip_comb
-      |> apsnd (fst o split_last)
-      |> list_comb;
-    val lhs = Free (Name.variant params "case", Term.fastype_of rhs);
-    val asm = (Thm.cterm_of thy o Logic.mk_equals) (lhs, rhs);
-  in
-    thms
-    |> Conjunction.intr_balanced
-    |> MetaSimplifier.rewrite_rule [(Thm.symmetric o Thm.assume) asm]
-    |> Thm.implies_intr asm
-    |> Thm.generalize ([], params) 0
-    |> AxClass.unoverload thy
-    |> Thm.varifyT_global
-  end;
-
-
-(* equality *)
-
-fun mk_eq_eqns thy dtco =
-  let
-    val (vs, cos) = Datatype_Data.the_spec thy dtco;
-    val { descr, index, inject = inject_thms, distinct = distinct_thms, ... } =
-      Datatype_Data.the_info thy dtco;
-    val ty = Type (dtco, map TFree vs);
-    fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT)
-      $ t1 $ t2;
-    fun true_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.true_const);
-    fun false_eq t12 = HOLogic.mk_eq (mk_eq t12, HOLogic.false_const);
-    val triv_injects = map_filter
-     (fn (c, []) => SOME (HOLogic.mk_Trueprop (true_eq (Const (c, ty), Const (c, ty))))
-       | _ => NONE) cos;
-    fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) =
-      trueprop $ (equiv $ mk_eq (t1, t2) $ rhs);
-    val injects = map prep_inject (nth (Datatype_Prop.make_injs [descr] vs) index);
-    fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) =
-      [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)];
-    val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index));
-    val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty)));
-    val simpset = Simplifier.global_context thy (HOL_basic_ss addsimps 
-      (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms)));
-    fun prove prop = Skip_Proof.prove_global thy [] [] prop (K (ALLGOALS (simp_tac simpset)))
-      |> Simpdata.mk_eq;
-  in (map prove (triv_injects @ injects @ distincts), prove refl) end;
-
-fun add_equality vs dtcos thy =
-  let
-    fun add_def dtco lthy =
-      let
-        val ty = Type (dtco, map TFree vs);
-        fun mk_side const_name = Const (const_name, ty --> ty --> HOLogic.boolT)
-          $ Free ("x", ty) $ Free ("y", ty);
-        val def = HOLogic.mk_Trueprop (HOLogic.mk_eq
-          (mk_side @{const_name eq_class.eq}, mk_side @{const_name "op ="}));
-        val def' = Syntax.check_term lthy def;
-        val ((_, (_, thm)), lthy') = Specification.definition
-          (NONE, (Attrib.empty_binding, def')) lthy;
-        val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
-        val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm;
-      in (thm', lthy') end;
-    fun tac thms = Class.intro_classes_tac []
-      THEN ALLGOALS (ProofContext.fact_tac thms);
-    fun add_eq_thms dtco =
-      Theory.checkpoint
-      #> `(fn thy => mk_eq_eqns thy dtco)
-      #-> (fn (thms, thm) =>
-        Code.add_nbe_eqn thm
-        #> fold_rev Code.add_eqn thms);
-  in
-    thy
-    |> Theory_Target.instantiation (dtcos, vs, [HOLogic.class_eq])
-    |> fold_map add_def dtcos
-    |-> (fn def_thms => Class.prove_instantiation_exit_result (map o Morphism.thm)
-         (fn _ => fn def_thms => tac def_thms) def_thms)
-    |-> (fn def_thms => fold Code.del_eqn def_thms)
-    |> fold add_eq_thms dtcos
-  end;
-
-
-(* register a datatype etc. *)
-
-fun add_all_code config dtcos thy =
-  let
-    val (vs :: _, coss) = (split_list o map (Datatype_Data.the_spec thy)) dtcos;
-    val any_css = map2 (mk_constr_consts thy vs) dtcos coss;
-    val css = if exists is_none any_css then []
-      else map_filter I any_css;
-    val case_rewrites = maps (#case_rewrites o Datatype_Data.the_info thy) dtcos;
-    val certs = map (mk_case_cert thy) dtcos;
-  in
-    if null css then thy
-    else thy
-      |> tap (fn _ => Datatype_Aux.message config "Registering datatype for code generator ...")
-      |> fold Code.add_datatype css
-      |> fold_rev Code.add_default_eqn case_rewrites
-      |> fold Code.add_case certs
-      |> add_equality vs dtcos
-   end;
-
-
 (** theory setup **)
 
 val setup = 
-  add_codegen "datatype" datatype_codegen
-  #> add_tycodegen "datatype" datatype_tycodegen
-  #> Datatype_Data.interpretation add_all_code
+  Datatype_Data.interpretation add_all_code
+  #> add_codegen "datatype" datatype_codegen
+  #> add_tycodegen "datatype" datatype_tycodegen;
 
 end;
--- a/src/HOL/Tools/Datatype/datatype_data.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/Datatype/datatype_data.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -436,7 +436,7 @@
   in
     thy
     |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules))
+    |> Proof.theorem NONE after_qed' ((map o map) (rpair []) (flat rules))
   end;
 
 val rep_datatype = gen_rep_datatype Sign.cert_term;
--- a/src/HOL/Tools/Function/function.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/Function/function.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -120,7 +120,7 @@
       end
   in
     lthy
-    |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
+    |> Proof.theorem NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]]
     |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd
   end
 
@@ -177,7 +177,7 @@
     |> ProofContext.note_thmss ""
        [((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]),
          [([Goal.norm_result termination], [])])] |> snd
-    |> Proof.theorem_i NONE afterqed [[(goal, [])]]
+    |> Proof.theorem NONE afterqed [[(goal, [])]]
   end
 
 val termination_proof = gen_termination_proof Syntax.check_term
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -3059,7 +3059,7 @@
            ) options [const]))
       end
   in
-    Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
+    Proof.theorem NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
   end;
 
 val code_pred = generic_code_pred (K I);
--- a/src/HOL/Tools/Quotient/quotient_typ.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/Quotient/quotient_typ.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -45,7 +45,7 @@
   val goals' = map (rpair []) goals
   fun after_qed' thms = after_qed (the_single thms)
 in
-  Proof.theorem_i NONE after_qed' [goals'] ctxt
+  Proof.theorem NONE after_qed' [goals'] ctxt
 end
 
 
--- a/src/HOL/Tools/Sledgehammer/named_thm_set.ML	Mon Apr 26 21:20:43 2010 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-(*  Title:      HOL/Tools/Sledgehammer/named_thm_set.ML
-    Author:     Makarius
-    Author:     Jasmin Blanchette, TU Muenchen
-
-Named sets of theorems.
-*)
-
-signature NAMED_THM_SET =
-sig
-  val member: Proof.context -> thm -> bool
-  val add_thm: thm -> Context.generic -> Context.generic
-  val del_thm: thm -> Context.generic -> Context.generic
-  val add: attribute
-  val del: attribute
-  val setup: theory -> theory
-end;
-
-functor Named_Thm_Set(val name: string val description: string)
-        : NAMED_THM_SET =
-struct
-
-structure Data = Generic_Data
-(
-  type T = thm Termtab.table;
-  val empty = Termtab.empty;
-  val extend = I;
-  fun merge tabs = Termtab.merge (K true) tabs;
-);
-
-fun member ctxt th =
-  Termtab.defined (Data.get (Context.Proof ctxt)) (Thm.full_prop_of th);
-
-fun add_thm th = Data.map (Termtab.update (Thm.full_prop_of th, th));
-fun del_thm th = Data.map (Termtab.delete_safe (Thm.full_prop_of th));
-
-val add = Thm.declaration_attribute add_thm;
-val del = Thm.declaration_attribute del_thm;
-
-val setup =
-  Attrib.setup (Binding.name name) (Attrib.add_del add del)
-               ("declaration of " ^ description) #>
-  PureThy.add_thms_dynamic (Binding.name name,
-                            map #2 o Termtab.dest o Data.get);
-
-end;
--- a/src/HOL/Tools/choice_specification.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/choice_specification.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -226,7 +226,7 @@
     in
       thy
       |> ProofContext.init
-      |> Proof.theorem_i NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
+      |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]]
     end;
 
 
--- a/src/HOL/Tools/numeral_simprocs.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/numeral_simprocs.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -332,8 +332,8 @@
 val field_combine_numerals =
   Arith_Data.prep_simproc @{theory}
     ("field_combine_numerals", 
-     ["(i::'a::{number_ring,field,division_by_zero}) + j",
-      "(i::'a::{number_ring,field,division_by_zero}) - j"], 
+     ["(i::'a::{number_ring,field,division_ring_inverse_zero}) + j",
+      "(i::'a::{number_ring,field,division_ring_inverse_zero}) - j"], 
      K FieldCombineNumerals.proc);
 
 (** Constant folding for multiplication in semirings **)
@@ -442,9 +442,9 @@
       "(l::'a::{semiring_div,number_ring}) div (m * n)"],
      K DivCancelNumeralFactor.proc),
     ("divide_cancel_numeral_factor",
-     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
-      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
-      "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
+     ["((l::'a::{division_ring_inverse_zero,field,number_ring}) * m) / n",
+      "(l::'a::{division_ring_inverse_zero,field,number_ring}) / (m * n)",
+      "((number_of v)::'a::{division_ring_inverse_zero,field,number_ring}) / (number_of w)"],
      K DivideCancelNumeralFactor.proc)];
 
 val field_cancel_numeral_factors =
@@ -454,9 +454,9 @@
       "(l::'a::{field,number_ring}) = m * n"],
      K EqCancelNumeralFactor.proc),
     ("field_cancel_numeral_factor",
-     ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
-      "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
-      "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
+     ["((l::'a::{division_ring_inverse_zero,field,number_ring}) * m) / n",
+      "(l::'a::{division_ring_inverse_zero,field,number_ring}) / (m * n)",
+      "((number_of v)::'a::{division_ring_inverse_zero,field,number_ring}) / (number_of w)"],
      K DivideCancelNumeralFactor.proc)]
 
 
@@ -598,8 +598,8 @@
      ["((l::'a::idom) * m) dvd n", "(l::'a::idom) dvd (m * n)"],
      K DvdCancelFactor.proc),
     ("divide_cancel_factor",
-     ["((l::'a::{division_by_zero,field}) * m) / n",
-      "(l::'a::{division_by_zero,field}) / (m * n)"],
+     ["((l::'a::{division_ring_inverse_zero,field}) * m) / n",
+      "(l::'a::{division_ring_inverse_zero,field}) / (m * n)"],
      K DivideCancelFactor.proc)];
 
 end;
--- a/src/HOL/Tools/typedef.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/Tools/typedef.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -282,7 +282,7 @@
     val ((goal, goal_pat, typedef_result), lthy') =
       prepare_typedef prep_term def name (b, args, mx) set opt_morphs lthy;
     fun after_qed [[th]] = snd o typedef_result th;
-  in Proof.theorem_i NONE after_qed [[(goal, [goal_pat])]] lthy' end;
+  in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end;
 
 in
 
--- a/src/HOL/ex/Classical.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/ex/Classical.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -348,7 +348,7 @@
 
 text{*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
   fast DISCOVERS who killed Agatha. *}
-lemma "lives(agatha) & lives(butler) & lives(charles) &
+schematic_lemma "lives(agatha) & lives(butler) & lives(charles) &
    (killed agatha agatha | killed butler agatha | killed charles agatha) &
    (\<forall>x y. killed x y --> hates x y & ~richer x y) &
    (\<forall>x. hates agatha x --> ~hates charles x) &
--- a/src/HOL/ex/Groebner_Examples.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/ex/Groebner_Examples.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -10,13 +10,13 @@
 
 subsection {* Basic examples *}
 
-lemma "3 ^ 3 == (?X::'a::{number_ring})"
+schematic_lemma "3 ^ 3 == (?X::'a::{number_ring})"
   by sring_norm
 
-lemma "(x - (-2))^5 == ?X::int"
+schematic_lemma "(x - (-2))^5 == ?X::int"
   by sring_norm
 
-lemma "(x - (-2))^5  * (y - 78) ^ 8 == ?X::int"
+schematic_lemma "(x - (-2))^5  * (y - 78) ^ 8 == ?X::int"
   by sring_norm
 
 lemma "((-3) ^ (Suc (Suc (Suc 0)))) == (X::'a::{number_ring})"
--- a/src/HOL/ex/Lagrange.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/ex/Lagrange.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -34,7 +34,7 @@
    sq (x1*y2 + x2*y1 + x3*y4 - x4*y3)  +
    sq (x1*y3 - x2*y4 + x3*y1 + x4*y2)  +
    sq (x1*y4 + x2*y3 - x3*y2 + x4*y1)"
-by (simp only: sq_def ring_simps)
+by (simp only: sq_def field_simps)
 
 
 text {* A challenge by John Harrison. Takes about 12s on a 1.6GHz machine. *}
@@ -50,6 +50,6 @@
       sq (p1*u2 + q1*t2 - r1*w2 + s1*v2 - t1*q2 + u1*p2 - v1*s2 + w1*r2) +
       sq (p1*v2 + q1*w2 + r1*t2 - s1*u2 - t1*r2 + u1*s2 + v1*p2 - w1*q2) +
       sq (p1*w2 - q1*v2 + r1*u2 + s1*t2 - t1*s2 - u1*r2 + v1*q2 + w1*p2)"
-by (simp only: sq_def ring_simps)
+by (simp only: sq_def field_simps)
 
 end
--- a/src/HOL/ex/Refute_Examples.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/ex/Refute_Examples.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -347,12 +347,12 @@
 
 subsubsection {* Schematic variables *}
 
-lemma "?P"
+schematic_lemma "?P"
   refute
   apply auto
 done
 
-lemma "x = ?y"
+schematic_lemma "x = ?y"
   refute
   apply auto
 done
--- a/src/HOL/ex/set.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOL/ex/set.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -26,7 +26,7 @@
   Trivial example of term synthesis: apparently hard for some provers!
 *}
 
-lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
+schematic_lemma "a \<noteq> b \<Longrightarrow> a \<in> ?X \<and> b \<notin> ?X"
   by blast
 
 
@@ -60,15 +60,15 @@
   -- {* Requires best-first search because it is undirectional. *}
   by best
 
-lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
+schematic_lemma "\<forall>f:: 'a \<Rightarrow> 'a set. \<forall>x. f x \<noteq> ?S f"
   -- {*This form displays the diagonal term. *}
   by best
 
-lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
+schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
   -- {* This form exploits the set constructs. *}
   by (rule notI, erule rangeE, best)
 
-lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
+schematic_lemma "?S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
   -- {* Or just this! *}
   by best
 
--- a/src/HOLCF/Tools/pcpodef.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/HOLCF/Tools/pcpodef.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -328,7 +328,7 @@
       prepare_cpodef prep_term def name (b, args, mx) set opt_morphs thy;
     fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
       | after_qed _ = raise Fail "cpodef_proof";
-  in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
+  in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
 
 fun gen_pcpodef_proof prep_term prep_constraint
     ((def, name), (b, raw_args, mx), set, opt_morphs) thy =
@@ -339,7 +339,7 @@
       prepare_pcpodef prep_term def name (b, args, mx) set opt_morphs thy;
     fun after_qed [[th1, th2]] = ProofContext.theory (snd o make_result th1 th2)
       | after_qed _ = raise Fail "pcpodef_proof";
-  in Proof.theorem_i NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
+  in Proof.theorem NONE after_qed [[(goal1, []), (goal2, [])]] ctxt end;
 
 in
 
--- a/src/Pure/Isar/calculation.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/calculation.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -13,11 +13,11 @@
   val sym_add: attribute
   val sym_del: attribute
   val symmetric: attribute
-  val also: (Facts.ref * Attrib.src list) list option -> bool -> Proof.state -> Proof.state Seq.seq
-  val also_i: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
-  val finally: (Facts.ref * Attrib.src list) list option -> bool ->
+  val also: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
+  val also_cmd: (Facts.ref * Attrib.src list) list option -> bool -> Proof.state -> Proof.state Seq.seq
+  val finally: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
+  val finally_cmd: (Facts.ref * Attrib.src list) list option -> bool ->
     Proof.state -> Proof.state Seq.seq
-  val finally_i: thm list option -> bool -> Proof.state -> Proof.state Seq.seq
   val moreover: bool -> Proof.state -> Proof.state
   val ultimately: bool -> Proof.state -> Proof.state
 end;
@@ -148,10 +148,10 @@
         state |> maintain_calculation final calc))
   end;
 
-val also = calculate Proof.get_thmss false;
-val also_i = calculate (K I) false;
-val finally = calculate Proof.get_thmss true;
-val finally_i = calculate (K I) true;
+val also = calculate (K I) false;
+val also_cmd = calculate Proof.get_thmss_cmd false;
+val finally = calculate (K I) true;
+val finally_cmd = calculate Proof.get_thmss_cmd true;
 
 
 (* moreover and ultimately *)
--- a/src/Pure/Isar/class_target.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/class_target.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -163,7 +163,7 @@
       Symtab.empty
       |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
            Symtab.map_default (class, []) (insert (op =) tyco)) arities)
-             ((#arities o Sorts.rep_algebra) algebra);
+             (Sorts.arities_of algebra);
     val the_arities = these o Symtab.lookup arities;
     fun mk_arity class tyco =
       let
@@ -370,7 +370,7 @@
   in
     thy
     |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed [[(mk_prop thy classrel, [])]]
+    |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]]
   end;
 
 in
@@ -539,7 +539,7 @@
   end;
 
 val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
-  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
+  Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
 
 fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
   fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
@@ -595,7 +595,7 @@
   in
     thy
     |> ProofContext.init
-    |> Proof.theorem_i NONE after_qed (map (fn t => [(t, [])]) arities)
+    |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities)
   end;
 
 
--- a/src/Pure/Isar/element.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/element.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -283,10 +283,10 @@
 in
 
 fun witness_proof after_qed wit_propss =
-  gen_witness_proof (Proof.theorem_i NONE) (fn wits => fn _ => after_qed wits)
+  gen_witness_proof (Proof.theorem NONE) (fn wits => fn _ => after_qed wits)
     wit_propss [];
 
-val witness_proof_eqs = gen_witness_proof (Proof.theorem_i NONE);
+val witness_proof_eqs = gen_witness_proof (Proof.theorem NONE);
 
 fun witness_local_proof after_qed cmd wit_propss goal_ctxt int =
   gen_witness_proof (fn after_qed' => fn propss =>
--- a/src/Pure/Isar/isar_cmd.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/isar_cmd.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -219,10 +219,10 @@
 fun goal opt_chain goal stmt int =
   opt_chain #> goal NONE (K I) stmt int;
 
-val have = goal I Proof.have;
-val hence = goal Proof.chain Proof.have;
-val show = goal I Proof.show;
-val thus = goal Proof.chain Proof.show;
+val have = goal I Proof.have_cmd;
+val hence = goal Proof.chain Proof.have_cmd;
+val show = goal I Proof.show_cmd;
+val thus = goal Proof.chain Proof.show_cmd;
 
 
 (* local endings *)
@@ -393,7 +393,7 @@
   let
     val thy = Toplevel.theory_of state;
     val {classes = (space, algebra), ...} = Type.rep_tsig (Sign.tsig_of thy);
-    val {classes, ...} = Sorts.rep_algebra algebra;
+    val classes = Sorts.classes_of algebra;
     fun entry (c, (i, (_, cs))) =
       (i, {name = Name_Space.extern space c, ID = c, parents = cs,
             dir = "", unfold = true, path = ""});
@@ -403,7 +403,7 @@
   in Present.display_graph gr end);
 
 fun thm_deps args = Toplevel.unknown_theory o Toplevel.keep (fn state =>
-  Thm_Deps.thm_deps (Proof.get_thmss (Toplevel.enter_proof_body state) args));
+  Thm_Deps.thm_deps (Proof.get_thmss_cmd (Toplevel.enter_proof_body state) args));
 
 
 (* find unused theorems *)
@@ -437,12 +437,12 @@
 local
 
 fun string_of_stmts state args =
-  Proof.get_thmss state args
+  Proof.get_thmss_cmd state args
   |> map (Element.pretty_statement (Proof.context_of state) Thm.theoremK)
   |> Pretty.chunks2 |> Pretty.string_of;
 
 fun string_of_thms state args =
-  Pretty.string_of (Display.pretty_thms (Proof.context_of state) (Proof.get_thmss state args));
+  Pretty.string_of (Display.pretty_thms (Proof.context_of state) (Proof.get_thmss_cmd state args));
 
 fun string_of_prfs full state arg =
   Pretty.string_of
@@ -460,7 +460,7 @@
         end
     | SOME args => Pretty.chunks
         (map (Proof_Syntax.pretty_proof_of (Proof.context_of state) full)
-          (Proof.get_thmss state args)));
+          (Proof.get_thmss_cmd state args)));
 
 fun string_of_prop state s =
   let
--- a/src/Pure/Isar/isar_syn.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/isar_syn.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -491,16 +491,23 @@
 
 (* statements *)
 
-fun gen_theorem kind =
-  OuterSyntax.local_theory_to_proof' kind ("state " ^ kind) K.thy_goal
+fun gen_theorem schematic kind =
+  OuterSyntax.local_theory_to_proof'
+    (if schematic then "schematic_" ^ kind else kind)
+    ("state " ^ (if schematic then "schematic " ^ kind else kind))
+    (if schematic then K.thy_schematic_goal else K.thy_goal)
     (Scan.optional (SpecParse.opt_thm_name ":" --|
       Scan.ahead (SpecParse.locale_keyword || SpecParse.statement_keyword)) Attrib.empty_binding --
       SpecParse.general_statement >> (fn (a, (elems, concl)) =>
-        (Specification.theorem_cmd kind NONE (K I) a elems concl)));
+        ((if schematic then Specification.schematic_theorem_cmd else Specification.theorem_cmd)
+          kind NONE (K I) a elems concl)));
 
-val _ = gen_theorem Thm.theoremK;
-val _ = gen_theorem Thm.lemmaK;
-val _ = gen_theorem Thm.corollaryK;
+val _ = gen_theorem false Thm.theoremK;
+val _ = gen_theorem false Thm.lemmaK;
+val _ = gen_theorem false Thm.corollaryK;
+val _ = gen_theorem true Thm.theoremK;
+val _ = gen_theorem true Thm.lemmaK;
+val _ = gen_theorem true Thm.corollaryK;
 
 val _ =
   OuterSyntax.command "have" "state local goal"
@@ -535,27 +542,27 @@
 val _ =
   OuterSyntax.command "from" "forward chaining from given facts"
     (K.tag_proof K.prf_chain)
-    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss)));
+    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.from_thmss_cmd)));
 
 val _ =
   OuterSyntax.command "with" "forward chaining from given and current facts"
     (K.tag_proof K.prf_chain)
-    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss)));
+    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.with_thmss_cmd)));
 
 val _ =
   OuterSyntax.command "note" "define facts"
     (K.tag_proof K.prf_decl)
-    (SpecParse.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss)));
+    (SpecParse.name_facts >> (Toplevel.print oo (Toplevel.proof o Proof.note_thmss_cmd)));
 
 val _ =
   OuterSyntax.command "using" "augment goal facts"
     (K.tag_proof K.prf_decl)
-    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.using)));
+    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.using_cmd)));
 
 val _ =
   OuterSyntax.command "unfolding" "unfold definitions in goal and facts"
     (K.tag_proof K.prf_decl)
-    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding)));
+    (facts >> (Toplevel.print oo (Toplevel.proof o Proof.unfolding_cmd)));
 
 
 (* proof context *)
@@ -563,17 +570,17 @@
 val _ =
   OuterSyntax.command "fix" "fix local variables (Skolem constants)"
     (K.tag_proof K.prf_asm)
-    (P.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix)));
+    (P.fixes >> (Toplevel.print oo (Toplevel.proof o Proof.fix_cmd)));
 
 val _ =
   OuterSyntax.command "assume" "assume propositions"
     (K.tag_proof K.prf_asm)
-    (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume)));
+    (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.assume_cmd)));
 
 val _ =
   OuterSyntax.command "presume" "assume propositions, to be established later"
     (K.tag_proof K.prf_asm)
-    (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume)));
+    (SpecParse.statement >> (Toplevel.print oo (Toplevel.proof o Proof.presume_cmd)));
 
 val _ =
   OuterSyntax.command "def" "local definition"
@@ -581,24 +588,24 @@
     (P.and_list1
       (SpecParse.opt_thm_name ":" --
         ((P.binding -- P.opt_mixfix) -- ((P.$$$ "\\<equiv>" || P.$$$ "==") |-- P.!!! P.termp)))
-    >> (Toplevel.print oo (Toplevel.proof o Proof.def)));
+    >> (Toplevel.print oo (Toplevel.proof o Proof.def_cmd)));
 
 val _ =
   OuterSyntax.command "obtain" "generalized existence"
     (K.tag_proof K.prf_asm_goal)
     (P.parname -- Scan.optional (P.fixes --| P.where_) [] -- SpecParse.statement
-      >> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain x y z)));
+      >> (fn ((x, y), z) => Toplevel.print o Toplevel.proof' (Obtain.obtain_cmd x y z)));
 
 val _ =
   OuterSyntax.command "guess" "wild guessing (unstructured)"
     (K.tag_proof K.prf_asm_goal)
-    (Scan.optional P.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess)));
+    (Scan.optional P.fixes [] >> (Toplevel.print oo (Toplevel.proof' o Obtain.guess_cmd)));
 
 val _ =
   OuterSyntax.command "let" "bind text variables"
     (K.tag_proof K.prf_decl)
     (P.and_list1 (P.and_list1 P.term -- (P.$$$ "=" |-- P.term))
-      >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind)));
+      >> (Toplevel.print oo (Toplevel.proof o Proof.let_bind_cmd)));
 
 val case_spec =
   (P.$$$ "(" |-- P.!!! (P.xname -- Scan.repeat1 (P.maybe P.name) --| P.$$$ ")") ||
@@ -607,7 +614,7 @@
 val _ =
   OuterSyntax.command "case" "invoke local context"
     (K.tag_proof K.prf_asm)
-    (case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case)));
+    (case_spec >> (Toplevel.print oo (Toplevel.proof o Proof.invoke_case_cmd)));
 
 
 (* proof structure *)
@@ -704,12 +711,12 @@
 val _ =
   OuterSyntax.command "also" "combine calculation and current facts"
     (K.tag_proof K.prf_decl)
-    (calc_args >> (Toplevel.proofs' o Calculation.also));
+    (calc_args >> (Toplevel.proofs' o Calculation.also_cmd));
 
 val _ =
   OuterSyntax.command "finally" "combine calculation and current facts, exhibit result"
     (K.tag_proof K.prf_chain)
-    (calc_args >> (Toplevel.proofs' o Calculation.finally));
+    (calc_args >> (Toplevel.proofs' o Calculation.finally_cmd));
 
 val _ =
   OuterSyntax.command "moreover" "augment calculation by current facts"
--- a/src/Pure/Isar/obtain.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/obtain.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -39,14 +39,14 @@
 signature OBTAIN =
 sig
   val thatN: string
-  val obtain: string -> (binding * string option * mixfix) list ->
+  val obtain: string -> (binding * typ option * mixfix) list ->
+    (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
+  val obtain_cmd: string -> (binding * string option * mixfix) list ->
     (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
-  val obtain_i: string -> (binding * typ option * mixfix) list ->
-    (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
   val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
     ((string * cterm) list * thm list) * Proof.context
-  val guess: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
-  val guess_i: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+  val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
+  val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
 end;
 
 structure Obtain: OBTAIN =
@@ -148,26 +148,26 @@
     fun after_qed _ =
       Proof.local_qed (NONE, false)
       #> `Proof.the_fact #-> (fn rule =>
-        Proof.fix_i vars
-        #> Proof.assm_i (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms);
+        Proof.fix vars
+        #> Proof.assm (obtain_export fix_ctxt rule (map (cert o Free) parms)) asms);
   in
     state
     |> Proof.enter_forward
-    |> Proof.have_i NONE (K I) [(Thm.empty_binding, [(obtain_prop, [])])] int
+    |> Proof.have NONE (K I) [(Thm.empty_binding, [(obtain_prop, [])])] int
     |> Proof.proof (SOME Method.succeed_text) |> Seq.hd
-    |> Proof.fix_i [(Binding.name thesisN, NONE, NoSyn)]
-    |> Proof.assume_i
+    |> Proof.fix [(Binding.name thesisN, NONE, NoSyn)]
+    |> Proof.assume
       [((Binding.name that_name, [Context_Rules.intro_query NONE]), [(that_prop, [])])]
     |> `Proof.the_facts
     ||> Proof.chain_facts chain_facts
-    ||> Proof.show_i NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false
+    ||> Proof.show NONE after_qed [(Thm.empty_binding, [(thesis, [])])] false
     |-> Proof.refine_insert
   end;
 
 in
 
-val obtain = gen_obtain Attrib.attribute ProofContext.read_vars ProofContext.read_propp;
-val obtain_i = gen_obtain (K I) ProofContext.cert_vars ProofContext.cert_propp;
+val obtain = gen_obtain (K I) ProofContext.cert_vars ProofContext.cert_propp;
+val obtain_cmd = gen_obtain Attrib.attribute ProofContext.read_vars ProofContext.read_propp;
 
 end;
 
@@ -290,8 +290,8 @@
       in
         state'
         |> Proof.map_context (K ctxt')
-        |> Proof.fix_i (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
-        |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm_i
+        |> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
+        |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm
           (obtain_export fix_ctxt rule (map cert ts)) [(Thm.empty_binding, asms)])
         |> Proof.bind_terms Auto_Bind.no_facts
       end;
@@ -307,7 +307,7 @@
     state
     |> Proof.enter_forward
     |> Proof.begin_block
-    |> Proof.fix_i [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
+    |> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
     |> Proof.chain_facts chain_facts
     |> Proof.local_goal print_result (K I) (apsnd (rpair I))
       "guess" before_qed after_qed [(Thm.empty_binding, [Logic.mk_term goal, goal])]
@@ -316,8 +316,8 @@
 
 in
 
-val guess = gen_guess ProofContext.read_vars;
-val guess_i = gen_guess ProofContext.cert_vars;
+val guess = gen_guess ProofContext.cert_vars;
+val guess_cmd = gen_guess ProofContext.read_vars;
 
 end;
 
--- a/src/Pure/Isar/outer_keyword.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/outer_keyword.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -17,6 +17,7 @@
   val thy_decl: T
   val thy_script: T
   val thy_goal: T
+  val thy_schematic_goal: T
   val qed: T
   val qed_block: T
   val qed_global: T
@@ -55,6 +56,7 @@
   val is_proof: string -> bool
   val is_theory_goal: string -> bool
   val is_proof_goal: string -> bool
+  val is_schematic_goal: string -> bool
   val is_qed: string -> bool
   val is_qed_global: string -> bool
 end;
@@ -81,6 +83,7 @@
 val thy_decl = kind "theory-decl";
 val thy_script = kind "theory-script";
 val thy_goal = kind "theory-goal";
+val thy_schematic_goal = kind "theory-schematic-goal";
 val qed = kind "qed";
 val qed_block = kind "qed-block";
 val qed_global = kind "qed-global";
@@ -95,9 +98,11 @@
 val prf_asm_goal = kind "proof-asm-goal";
 val prf_script = kind "proof-script";
 
-val kinds = [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script,
-  thy_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
-  prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script] |> map kind_of;
+val kinds =
+  [control, diag, thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal,
+    thy_schematic_goal, qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open,
+    prf_close, prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script]
+ |> map kind_of;
 
 
 (* tags *)
@@ -191,14 +196,15 @@
 val is_theory_begin = command_category [thy_begin];
 
 val is_theory = command_category
-  [thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal];
+  [thy_begin, thy_switch, thy_end, thy_heading, thy_decl, thy_script, thy_goal, thy_schematic_goal];
 
 val is_proof = command_category
   [qed, qed_block, qed_global, prf_heading, prf_goal, prf_block, prf_open, prf_close,
     prf_chain, prf_decl, prf_asm, prf_asm_goal, prf_script];
 
-val is_theory_goal = command_category [thy_goal];
+val is_theory_goal = command_category [thy_goal, thy_schematic_goal];
 val is_proof_goal = command_category [prf_goal, prf_asm_goal];
+val is_schematic_goal = command_category [thy_schematic_goal];
 val is_qed = command_category [qed, qed_block];
 val is_qed_global = command_category [qed_global];
 
--- a/src/Pure/Isar/proof.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/proof.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -41,37 +41,35 @@
   val raw_goal: state -> {context: context, facts: thm list, goal: thm}
   val goal: state -> {context: context, facts: thm list, goal: thm}
   val simple_goal: state -> {context: context, goal: thm}
-  val match_bind: (string list * string) list -> state -> state
-  val match_bind_i: (term list * term) list -> state -> state
-  val let_bind: (string list * string) list -> state -> state
-  val let_bind_i: (term list * term) list -> state -> state
-  val fix: (binding * string option * mixfix) list -> state -> state
-  val fix_i: (binding * typ option * mixfix) list -> state -> state
+  val let_bind: (term list * term) list -> state -> state
+  val let_bind_cmd: (string list * string) list -> state -> state
+  val fix: (binding * typ option * mixfix) list -> state -> state
+  val fix_cmd: (binding * string option * mixfix) list -> state -> state
   val assm: Assumption.export ->
+    (Thm.binding * (term * term list) list) list -> state -> state
+  val assm_cmd: Assumption.export ->
     (Attrib.binding * (string * string list) list) list -> state -> state
-  val assm_i: Assumption.export ->
-    (Thm.binding * (term * term list) list) list -> state -> state
-  val assume: (Attrib.binding * (string * string list) list) list -> state -> state
-  val assume_i: (Thm.binding * (term * term list) list) list -> state -> state
-  val presume: (Attrib.binding * (string * string list) list) list -> state -> state
-  val presume_i: (Thm.binding * (term * term list) list) list -> state -> state
-  val def: (Attrib.binding * ((binding * mixfix) * (string * string list))) list -> state -> state
-  val def_i: (Thm.binding * ((binding * mixfix) * (term * term list))) list -> state -> state
+  val assume: (Thm.binding * (term * term list) list) list -> state -> state
+  val assume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state
+  val presume: (Thm.binding * (term * term list) list) list -> state -> state
+  val presume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state
+  val def: (Thm.binding * ((binding * mixfix) * (term * term list))) list -> state -> state
+  val def_cmd: (Attrib.binding * ((binding * mixfix) * (string * string list))) list -> state -> state
   val chain: state -> state
   val chain_facts: thm list -> state -> state
-  val get_thmss: state -> (Facts.ref * Attrib.src list) list -> thm list
-  val note_thmss: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
-  val note_thmss_i: (Thm.binding * (thm list * attribute list) list) list -> state -> state
-  val from_thmss: ((Facts.ref * Attrib.src list) list) list -> state -> state
-  val from_thmss_i: ((thm list * attribute list) list) list -> state -> state
-  val with_thmss: ((Facts.ref * Attrib.src list) list) list -> state -> state
-  val with_thmss_i: ((thm list * attribute list) list) list -> state -> state
-  val using: ((Facts.ref * Attrib.src list) list) list -> state -> state
-  val using_i: ((thm list * attribute list) list) list -> state -> state
-  val unfolding: ((Facts.ref * Attrib.src list) list) list -> state -> state
-  val unfolding_i: ((thm list * attribute list) list) list -> state -> state
-  val invoke_case: string * string option list * Attrib.src list -> state -> state
-  val invoke_case_i: string * string option list * attribute list -> state -> state
+  val get_thmss_cmd: state -> (Facts.ref * Attrib.src list) list -> thm list
+  val note_thmss: (Thm.binding * (thm list * attribute list) list) list -> state -> state
+  val note_thmss_cmd: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
+  val from_thmss: ((thm list * attribute list) list) list -> state -> state
+  val from_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val with_thmss: ((thm list * attribute list) list) list -> state -> state
+  val with_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val using: ((thm list * attribute list) list) list -> state -> state
+  val using_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val unfolding: ((thm list * attribute list) list) list -> state -> state
+  val unfolding_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val invoke_case: string * string option list * attribute list -> state -> state
+  val invoke_case_cmd: string * string option list * Attrib.src list -> state -> state
   val begin_block: state -> state
   val next_block: state -> state
   val end_block: state -> state
@@ -87,9 +85,9 @@
     ((binding * 'a list) * 'b) list -> state -> state
   val local_qed: Method.text option * bool -> state -> state
   val theorem: Method.text option -> (thm list list -> context -> context) ->
+    (term * term list) list list -> context -> state
+  val theorem_cmd: Method.text option -> (thm list list -> context -> context) ->
     (string * string list) list list -> context -> state
-  val theorem_i: Method.text option -> (thm list list -> context -> context) ->
-    (term * term list) list list -> context -> state
   val global_qed: Method.text option * bool -> state -> context
   val local_terminal_proof: Method.text * Method.text option -> state -> state
   val local_default_proof: state -> state
@@ -102,13 +100,13 @@
   val global_skip_proof: bool -> state -> context
   val global_done_proof: state -> context
   val have: Method.text option -> (thm list list -> state -> state) ->
+    (Thm.binding * (term * term list) list) list -> bool -> state -> state
+  val have_cmd: Method.text option -> (thm list list -> state -> state) ->
     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
-  val have_i: Method.text option -> (thm list list -> state -> state) ->
-    (Thm.binding * (term * term list) list) list -> bool -> state -> state
   val show: Method.text option -> (thm list list -> state -> state) ->
+    (Thm.binding * (term * term list) list) list -> bool -> state -> state
+  val show_cmd: Method.text option -> (thm list list -> state -> state) ->
     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
-  val show_i: Method.text option -> (thm list list -> state -> state) ->
-    (Thm.binding * (term * term list) list) list -> bool -> state -> state
   val schematic_goal: state -> bool
   val is_relevant: state -> bool
   val local_future_proof: (state -> ('a * state) Future.future) ->
@@ -523,22 +521,20 @@
 
 (** context elements **)
 
-(* bindings *)
+(* let bindings *)
 
 local
 
 fun gen_bind bind args state =
   state
   |> assert_forward
-  |> map_context (bind args #> snd)
+  |> map_context (bind true args #> snd)
   |> put_facts NONE;
 
 in
 
-val match_bind = gen_bind (ProofContext.match_bind false);
-val match_bind_i = gen_bind (ProofContext.match_bind_i false);
-val let_bind = gen_bind (ProofContext.match_bind true);
-val let_bind_i = gen_bind (ProofContext.match_bind_i true);
+val let_bind = gen_bind ProofContext.match_bind_i;
+val let_bind_cmd = gen_bind ProofContext.match_bind;
 
 end;
 
@@ -554,8 +550,8 @@
 
 in
 
-val fix = gen_fix (fn ctxt => fn args => fst (ProofContext.read_vars args ctxt));
-val fix_i = gen_fix (K I);
+val fix = gen_fix (K I);
+val fix_cmd = gen_fix (fn ctxt => fn args => fst (ProofContext.read_vars args ctxt));
 
 end;
 
@@ -572,12 +568,12 @@
 
 in
 
-val assm      = gen_assume ProofContext.add_assms Attrib.attribute;
-val assm_i    = gen_assume ProofContext.add_assms_i (K I);
-val assume    = assm Assumption.assume_export;
-val assume_i  = assm_i Assumption.assume_export;
-val presume   = assm Assumption.presume_export;
-val presume_i = assm_i Assumption.presume_export;
+val assm = gen_assume ProofContext.add_assms_i (K I);
+val assm_cmd = gen_assume ProofContext.add_assms Attrib.attribute;
+val assume = assm Assumption.assume_export;
+val assume_cmd = assm_cmd Assumption.assume_export;
+val presume = assm Assumption.presume_export;
+val presume_cmd = assm_cmd Assumption.presume_export;
 
 end;
 
@@ -605,8 +601,8 @@
 
 in
 
-val def = gen_def Attrib.attribute ProofContext.read_vars ProofContext.match_bind;
-val def_i = gen_def (K I) ProofContext.cert_vars ProofContext.match_bind_i;
+val def = gen_def (K I) ProofContext.cert_vars ProofContext.match_bind_i;
+val def_cmd = gen_def Attrib.attribute ProofContext.read_vars ProofContext.match_bind;
 
 end;
 
@@ -646,18 +642,18 @@
 
 in
 
-val note_thmss = gen_thmss (K []) I #2 Attrib.attribute ProofContext.get_fact;
-val note_thmss_i = gen_thmss (K []) I #2 (K I) (K I);
+val note_thmss = gen_thmss (K []) I #2 (K I) (K I);
+val note_thmss_cmd = gen_thmss (K []) I #2 Attrib.attribute ProofContext.get_fact;
 
-val from_thmss = gen_thmss (K []) chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
-val from_thmss_i = gen_thmss (K []) chain #2 (K I) (K I) o no_binding;
+val from_thmss = gen_thmss (K []) chain #2 (K I) (K I) o no_binding;
+val from_thmss_cmd = gen_thmss (K []) chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
 
-val with_thmss = gen_thmss the_facts chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
-val with_thmss_i = gen_thmss the_facts chain #2 (K I) (K I) o no_binding;
+val with_thmss = gen_thmss the_facts chain #2 (K I) (K I) o no_binding;
+val with_thmss_cmd = gen_thmss the_facts chain #2 Attrib.attribute ProofContext.get_fact o no_binding;
 
 val local_results = gen_thmss (K []) I I (K I) (K I) o map (apsnd Thm.simple_fact);
 
-fun get_thmss state srcs = the_facts (note_thmss [((Binding.empty, []), srcs)] state);
+fun get_thmss_cmd state srcs = the_facts (note_thmss_cmd [((Binding.empty, []), srcs)] state);
 
 end;
 
@@ -686,10 +682,10 @@
 
 in
 
-val using = gen_using append_using (K (K I)) Attrib.attribute ProofContext.get_fact;
-val using_i = gen_using append_using (K (K I)) (K I) (K I);
-val unfolding = gen_using unfold_using unfold_goals Attrib.attribute ProofContext.get_fact;
-val unfolding_i = gen_using unfold_using unfold_goals (K I) (K I);
+val using = gen_using append_using (K (K I)) (K I) (K I);
+val using_cmd = gen_using append_using (K (K I)) Attrib.attribute ProofContext.get_fact;
+val unfolding = gen_using unfold_using unfold_goals (K I) (K I);
+val unfolding_cmd = gen_using unfold_using unfold_goals Attrib.attribute ProofContext.get_fact;
 
 end;
 
@@ -709,15 +705,15 @@
     val assumptions = asms |> map (fn (a, ts) => ((qualified_binding a, atts), map (rpair []) ts));
   in
     state'
-    |> assume_i assumptions
+    |> assume assumptions
     |> bind_terms Auto_Bind.no_facts
-    |> `the_facts |-> (fn thms => note_thmss_i [((Binding.name name, []), [(thms, [])])])
+    |> `the_facts |-> (fn thms => note_thmss [((Binding.name name, []), [(thms, [])])])
   end;
 
 in
 
-val invoke_case = gen_invoke_case Attrib.attribute;
-val invoke_case_i = gen_invoke_case (K I);
+val invoke_case = gen_invoke_case (K I);
+val invoke_case_cmd = gen_invoke_case Attrib.attribute;
 
 end;
 
@@ -790,15 +786,16 @@
 
 local
 
-fun implicit_vars dest add props =
+val is_var =
+  can (dest_TVar o Logic.dest_type o Logic.dest_term) orf
+  can (dest_Var o Logic.dest_term);
+
+fun implicit_vars props =
   let
-    val (explicit_vars, props') = take_prefix (can dest) props |>> map dest;
-    val vars = rev (subtract (op =) explicit_vars (fold add props []));
-    val _ =
-      if null vars then ()
-      else warning ("Goal statement contains unbound schematic variable(s): " ^
-        commas_quote (map (Term.string_of_vname o fst) vars));
-  in (rev vars, props') end;
+    val (var_props, props') = take_prefix is_var props;
+    val explicit_vars = fold Term.add_vars var_props [];
+    val vars = filter_out (member (op =) explicit_vars) (fold Term.add_vars props []);
+  in map (Logic.mk_term o Var) vars end;
 
 fun refine_terms n =
   refine (Method.Basic (K (RAW_METHOD
@@ -823,11 +820,8 @@
       |> map_context_result (fn ctxt => swap (prepp (ctxt, raw_propp)));
     val props = flat propss;
 
-    val (_, props') =
-      implicit_vars (dest_TVar o Logic.dest_type o Logic.dest_term) Term.add_tvars props;
-    val (vars, _) = implicit_vars (dest_Var o Logic.dest_term) Term.add_vars props';
-
-    val propss' = map (Logic.mk_term o Var) vars :: propss;
+    val vars = implicit_vars props;
+    val propss' = vars :: propss;
     val goal_propss = filter_out null propss';
     val goal =
       cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
@@ -903,8 +897,8 @@
   init ctxt
   |> generic_goal (prepp #> ProofContext.auto_fixes) "" before_qed (K I, after_qed) propp;
 
-val theorem = global_goal ProofContext.bind_propp_schematic;
-val theorem_i = global_goal ProofContext.bind_propp_schematic_i;
+val theorem = global_goal ProofContext.bind_propp_schematic_i;
+val theorem_cmd = global_goal ProofContext.bind_propp_schematic;
 
 fun global_qeds txt =
   end_proof true txt
@@ -976,10 +970,10 @@
 
 in
 
-val have = gen_have Attrib.attribute ProofContext.bind_propp;
-val have_i = gen_have (K I) ProofContext.bind_propp_i;
-val show = gen_show Attrib.attribute ProofContext.bind_propp;
-val show_i = gen_show (K I) ProofContext.bind_propp_i;
+val have = gen_have (K I) ProofContext.bind_propp_i;
+val have_cmd = gen_have Attrib.attribute ProofContext.bind_propp;
+val show = gen_show (K I) ProofContext.bind_propp_i;
+val show_cmd = gen_show Attrib.attribute ProofContext.bind_propp;
 
 end;
 
--- a/src/Pure/Isar/specification.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/specification.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -62,6 +62,14 @@
     (thm list list -> local_theory -> local_theory) -> Attrib.binding ->
     Element.context list -> Element.statement ->
     bool -> local_theory -> Proof.state
+  val schematic_theorem: string -> Method.text option ->
+    (thm list list -> local_theory -> local_theory) -> Attrib.binding ->
+    Element.context_i list -> Element.statement_i ->
+    bool -> local_theory -> Proof.state
+  val schematic_theorem_cmd: string -> Method.text option ->
+    (thm list list -> local_theory -> local_theory) -> Attrib.binding ->
+    Element.context list -> Element.statement ->
+    bool -> local_theory -> Proof.state
   val add_theorem_hook: (bool -> Proof.state -> Proof.state) -> Context.generic -> Context.generic
 end;
 
@@ -361,7 +369,7 @@
   fun merge hooks : T = Library.merge (eq_snd op =) hooks;
 );
 
-fun gen_theorem prep_att prep_stmt
+fun gen_theorem schematic prep_att prep_stmt
     kind before_qed after_qed (name, raw_atts) raw_elems raw_concl int lthy =
   let
     val _ = Local_Theory.affirm lthy;
@@ -395,15 +403,20 @@
     goal_ctxt
     |> ProofContext.note_thmss "" [((Binding.name Auto_Bind.assmsN, []), [(prems, [])])]
     |> snd
-    |> Proof.theorem_i before_qed after_qed' (map snd stmt)
+    |> Proof.theorem before_qed after_qed' (map snd stmt)
     |> (case facts of NONE => I | SOME ths => Proof.refine_insert ths)
+    |> tap (fn state => not schematic andalso Proof.schematic_goal state andalso
+        error "Illegal schematic goal statement")
     |> Library.apply (map (fn (f, _) => f int) (rev (TheoremHook.get (Context.Proof goal_ctxt))))
   end;
 
 in
 
-val theorem = gen_theorem (K I) Expression.cert_statement;
-val theorem_cmd = gen_theorem Attrib.intern_src Expression.read_statement;
+val theorem = gen_theorem false (K I) Expression.cert_statement;
+val theorem_cmd = gen_theorem false Attrib.intern_src Expression.read_statement;
+
+val schematic_theorem = gen_theorem true (K I) Expression.cert_statement;
+val schematic_theorem_cmd = gen_theorem true Attrib.intern_src Expression.read_statement;
 
 fun add_theorem_hook f = TheoremHook.map (cons (f, stamp ()));
 
--- a/src/Pure/Isar/toplevel.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Isar/toplevel.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -658,8 +658,11 @@
 
 fun proof_result immediate (tr, proof_trs) st =
   let val st' = command tr st in
-    if immediate orelse null proof_trs orelse
-      not (can proof_of st') orelse Proof.is_relevant (proof_of st')
+    if immediate orelse
+      null proof_trs orelse
+      OuterKeyword.is_schematic_goal (name_of tr) orelse
+      not (can proof_of st') orelse
+      Proof.is_relevant (proof_of st')
     then
       let val (states, st'') = fold_map command_result proof_trs st'
       in (Lazy.value ((tr, st') :: (proof_trs ~~ states)), st'') end
--- a/src/Pure/ML/ml_thms.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/ML/ml_thms.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -64,7 +64,7 @@
         fun after_qed res goal_ctxt =
           put_thms (i, map prep_result (ProofContext.export goal_ctxt ctxt (flat res))) goal_ctxt;
         val ctxt' = ctxt
-          |> Proof.theorem_i NONE after_qed propss
+          |> Proof.theorem NONE after_qed propss
           |> Proof.global_terminal_proof methods;
         val (a, background') = background
           |> ML_Antiquote.variant "lemma" ||> put_thms (i, the_thms ctxt' i);
--- a/src/Pure/ProofGeneral/pgip_parser.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/ProofGeneral/pgip_parser.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -63,6 +63,7 @@
   |> command K.thy_decl         thy_decl
   |> command K.thy_script       thy_decl
   |> command K.thy_goal         goal
+  |> command K.thy_schematic_goal goal
   |> command K.qed              closegoal
   |> command K.qed_block        closegoal
   |> command K.qed_global       (fn text => [D.Giveupgoal {text = text}])
--- a/src/Pure/Thy/thy_output.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Thy/thy_output.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -574,7 +574,7 @@
       val prop_src =
         (case Args.dest_src source of ((a, arg :: _), pos) => Args.src ((a, [arg]), pos));
       val _ = context
-        |> Proof.theorem_i NONE (K I) [[(prop, [])]]
+        |> Proof.theorem NONE (K I) [[(prop, [])]]
         |> Proof.global_terminal_proof methods;
     in output (maybe_pretty_source (pretty_term context) prop_src [prop]) end);
 
--- a/src/Pure/Tools/named_thms.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/Tools/named_thms.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -6,6 +6,7 @@
 
 signature NAMED_THMS =
 sig
+  val member: Proof.context -> thm -> bool
   val get: Proof.context -> thm list
   val add_thm: thm -> Context.generic -> Context.generic
   val del_thm: thm -> Context.generic -> Context.generic
@@ -25,6 +26,8 @@
   val merge = Item_Net.merge;
 );
 
+val member = Item_Net.member o Data.get o Context.Proof;
+
 val content = Item_Net.content o Data.get;
 val get = content o Context.Proof;
 
--- a/src/Pure/axclass.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/axclass.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -13,9 +13,8 @@
   val add_arity: thm -> theory -> theory
   val prove_classrel: class * class -> tactic -> theory -> theory
   val prove_arity: string * sort list * sort -> tactic -> theory -> theory
-  val get_info: theory -> class ->
-    {def: thm, intro: thm, axioms: thm list, params: (string * typ) list}
-  val class_intros: theory -> thm list
+  type info = {def: thm, intro: thm, axioms: thm list, params: (string * typ) list}
+  val get_info: theory -> class -> info
   val class_of_param: theory -> string -> class option
   val cert_classrel: theory -> class * class -> class * class
   val read_classrel: theory -> xstring * xstring -> class * class
@@ -37,8 +36,6 @@
   val param_of_inst: theory -> string * string -> string
   val inst_of_param: theory -> string -> (string * string) option
   val thynames_of_arity: theory -> class * string -> string list
-  val introN: string
-  val axiomsN: string
 end;
 
 structure AxClass: AX_CLASS =
@@ -46,6 +43,18 @@
 
 (** theory data **)
 
+(* axclass info *)
+
+type info =
+ {def: thm,
+  intro: thm,
+  axioms: thm list,
+  params: (string * typ) list};
+
+fun make_axclass (def, intro, axioms, params): info =
+  {def = def, intro = intro, axioms = axioms, params = params};
+
+
 (* class parameters (canonical order) *)
 
 type param = string * class;
@@ -57,188 +66,256 @@
       " for " ^ Pretty.string_of_sort pp [c] ^
       (if c = c' then "" else " and " ^ Pretty.string_of_sort pp [c'])));
 
-fun merge_params _ ([], qs) = qs
-  | merge_params pp (ps, qs) =
-      fold_rev (fn q => if member (op =) ps q then I else add_param pp q) qs ps;
+
+(* setup data *)
+
+datatype data = Data of
+ {axclasses: info Symtab.table,
+  params: param list,
+  proven_classrels: (thm * proof) Symreltab.table,
+  proven_arities: ((class * sort list) * ((thm * string) * proof)) list Symtab.table,
+    (*arity theorems with theory name*)
+  inst_params:
+    (string * thm) Symtab.table Symtab.table *
+      (*constant name ~> type constructor ~> (constant name, equation)*)
+    (string * string) Symtab.table (*constant name ~> (constant name, type constructor)*),
+  diff_merge_classrels: (class * class) list};
+
+fun make_data
+    (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =
+  Data {axclasses = axclasses, params = params, proven_classrels = proven_classrels,
+    proven_arities = proven_arities, inst_params = inst_params,
+    diff_merge_classrels = diff_merge_classrels};
+
+structure Data = Theory_Data_PP
+(
+  type T = data;
+  val empty =
+    make_data (Symtab.empty, [], Symreltab.empty, Symtab.empty, (Symtab.empty, Symtab.empty), []);
+  val extend = I;
+  fun merge pp
+      (Data {axclasses = axclasses1, params = params1, proven_classrels = proven_classrels1,
+        proven_arities = proven_arities1, inst_params = inst_params1,
+        diff_merge_classrels = diff_merge_classrels1},
+       Data {axclasses = axclasses2, params = params2, proven_classrels = proven_classrels2,
+        proven_arities = proven_arities2, inst_params = inst_params2,
+        diff_merge_classrels = diff_merge_classrels2}) =
+    let
+      val axclasses' = Symtab.merge (K true) (axclasses1, axclasses2);
+      val params' =
+        if null params1 then params2
+        else fold_rev (fn q => if member (op =) params1 q then I else add_param pp q) params2 params1;
+
+      (*transitive closure of classrels and arity completion is done in Theory.at_begin hook*)
+      val proven_classrels' = Symreltab.join (K #1) (proven_classrels1, proven_classrels2);
+      val proven_arities' =
+        Symtab.join (K (Library.merge (eq_fst op =))) (proven_arities1, proven_arities2);
+
+      val classrels1 = Symreltab.keys proven_classrels1;
+      val classrels2 = Symreltab.keys proven_classrels2;
+      val diff_merge_classrels' =
+        subtract (op =) classrels1 classrels2 @
+        subtract (op =) classrels2 classrels1 @
+        diff_merge_classrels1 @ diff_merge_classrels2;
+
+      val inst_params' =
+        (Symtab.join (K (Symtab.merge (K true))) (#1 inst_params1, #1 inst_params2),
+          Symtab.merge (K true) (#2 inst_params1, #2 inst_params2));
+    in
+      make_data (axclasses', params', proven_classrels', proven_arities', inst_params',
+        diff_merge_classrels')
+    end;
+);
+
+fun map_data f =
+  Data.map (fn Data {axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels} =>
+    make_data (f (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels)));
+
+fun map_axclasses f =
+  map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+    (f axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels));
+
+fun map_params f =
+  map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+    (axclasses, f params, proven_classrels, proven_arities, inst_params, diff_merge_classrels));
+
+fun map_proven_classrels f =
+  map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+    (axclasses, params, f proven_classrels, proven_arities, inst_params, diff_merge_classrels));
+
+fun map_proven_arities f =
+  map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+    (axclasses, params, proven_classrels, f proven_arities, inst_params, diff_merge_classrels));
+
+fun map_inst_params f =
+  map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, diff_merge_classrels) =>
+    (axclasses, params, proven_classrels, proven_arities, f inst_params, diff_merge_classrels));
+
+val clear_diff_merge_classrels =
+  map_data (fn (axclasses, params, proven_classrels, proven_arities, inst_params, _) =>
+    (axclasses, params, proven_classrels, proven_arities, inst_params, []));
+
+val rep_data = Data.get #> (fn Data args => args);
+
+val axclasses_of = #axclasses o rep_data;
+val params_of = #params o rep_data;
+val proven_classrels_of = #proven_classrels o rep_data;
+val proven_arities_of = #proven_arities o rep_data;
+val inst_params_of = #inst_params o rep_data;
+val diff_merge_classrels_of = #diff_merge_classrels o rep_data;
 
 
-(* axclasses *)
-
-val introN = "intro";
-val superN = "super";
-val axiomsN = "axioms";
+(* axclasses with parameters *)
 
-datatype axclass = AxClass of
- {def: thm,
-  intro: thm,
-  axioms: thm list,
-  params: (string * typ) list};
+fun get_info thy c =
+  (case Symtab.lookup (axclasses_of thy) c of
+    SOME info => info
+  | NONE => error ("No such axclass: " ^ quote c));
 
-type axclasses = axclass Symtab.table * param list;
+fun all_params_of thy S =
+  let val params = params_of thy;
+  in fold (fn (x, c) => if Sign.subsort thy (S, [c]) then cons x else I) params [] end;
 
-fun make_axclass ((def, intro, axioms), params) = AxClass
-  {def = def, intro = intro, axioms = axioms, params = params};
-
-fun merge_axclasses pp ((tab1, params1), (tab2, params2)) : axclasses =
-  (Symtab.merge (K true) (tab1, tab2), merge_params pp (params1, params2));
+fun class_of_param thy = AList.lookup (op =) (params_of thy);
 
 
-(* instances *)
+(* maintain instances *)
 
 val classrel_prefix = "classrel_";
 val arity_prefix = "arity_";
 
-type instances =
-  ((class * class) * thm) list *  (*classrel theorems*)
-  ((class * sort list) * (thm * string)) list Symtab.table;  (*arity theorems with theory name*)
-
-fun merge_instances ((classrel1, arities1): instances, (classrel2, arities2)) =
- (merge (eq_fst op =) (classrel1, classrel2),
-  Symtab.join (K (merge (eq_fst op =))) (arities1, arities2));
-
-
-(* instance parameters *)
-
-type inst_params =
-  (string * thm) Symtab.table Symtab.table
-    (*constant name ~> type constructor ~> (constant name, equation)*)
-  * (string * string) Symtab.table; (*constant name ~> (constant name, type constructor)*)
-
-fun merge_inst_params ((const_param1, param_const1), (const_param2, param_const2)) =
-  (Symtab.join  (K (Symtab.merge (K true))) (const_param1, const_param2),
-    Symtab.merge (K true) (param_const1, param_const2));
-
-
-(* setup data *)
-
-structure AxClassData = Theory_Data_PP
-(
-  type T = axclasses * (instances * inst_params);
-  val empty = ((Symtab.empty, []), (([], Symtab.empty), (Symtab.empty, Symtab.empty)));
-  val extend = I;
-  fun merge pp ((axclasses1, (instances1, inst_params1)), (axclasses2, (instances2, inst_params2))) =
-    (merge_axclasses pp (axclasses1, axclasses2),
-      (merge_instances (instances1, instances2), merge_inst_params (inst_params1, inst_params2)));
-);
-
-
-(* maintain axclasses *)
-
-val get_axclasses = #1 o AxClassData.get;
-val map_axclasses = AxClassData.map o apfst;
-
-val lookup_def = Symtab.lookup o #1 o get_axclasses;
-
-fun get_info thy c =
-  (case lookup_def thy c of
-    SOME (AxClass info) => info
-  | NONE => error ("No such axclass: " ^ quote c));
-
-fun class_intros thy =
-  let
-    fun add_intro c =
-      (case lookup_def thy c of SOME (AxClass {intro, ...}) => cons intro | _ => I);
-    val classes = Sign.all_classes thy;
-  in map (Thm.class_triv thy) classes @ fold add_intro classes [] end;
-
-
-fun get_params thy pred =
-  let val params = #2 (get_axclasses thy);
-  in fold (fn (x, c) => if pred c then cons x else I) params [] end;
-
-fun all_params_of thy S = get_params thy (fn c => Sign.subsort thy (S, [c]));
-
-fun class_of_param thy = AList.lookup (op =) (#2 (get_axclasses thy));
-
-
-(* maintain instances *)
-
 fun instance_name (a, c) = Long_Name.base_name c ^ "_" ^ Long_Name.base_name a;
 
-val get_instances = #1 o #2 o AxClassData.get;
-val map_instances = AxClassData.map o apsnd o apfst;
-
 
 fun the_classrel thy (c1, c2) =
-  (case AList.lookup (op =) (#1 (get_instances thy)) (c1, c2) of
-    SOME th => Thm.transfer thy th
+  (case Symreltab.lookup (proven_classrels_of thy) (c1, c2) of
+    SOME classrel => classrel
   | NONE => error ("Unproven class relation " ^
       Syntax.string_of_classrel (ProofContext.init thy) [c1, c2]));
 
-fun put_classrel arg = map_instances (fn (classrel, arities) =>
-  (insert (eq_fst op =) arg classrel, arities));
+fun the_classrel_thm thy = Thm.transfer thy o #1 o the_classrel thy;
+fun the_classrel_prf thy = #2 o the_classrel thy;
+
+fun put_trancl_classrel ((c1, c2), th) thy =
+  let
+    val cert = Thm.cterm_of thy;
+    val certT = Thm.ctyp_of thy;
+
+    val classes = Sorts.classes_of (Sign.classes_of thy);
+    val classrels = proven_classrels_of thy;
+
+    fun reflcl_classrel (c1', c2') =
+      if c1' = c2'
+      then Thm.trivial (cert (Logic.mk_of_class (TVar ((Name.aT, 0), []), c1')))
+      else the_classrel_thm thy (c1', c2');
+    fun gen_classrel (c1_pred, c2_succ) =
+      let
+        val th' = ((reflcl_classrel (c1_pred, c1) RS th) RS reflcl_classrel (c2, c2_succ))
+          |> Drule.instantiate' [SOME (certT (TVar ((Name.aT, 0), [])))] []
+          |> Thm.close_derivation;
+        val prf' = Thm.proof_of th';
+      in ((c1_pred, c2_succ), (th', prf')) end;
+
+    val new_classrels =
+      Library.map_product pair (c1 :: Graph.imm_preds classes c1) (c2 :: Graph.imm_succs classes c2)
+      |> filter_out (Symreltab.defined classrels)
+      |> map gen_classrel;
+    val needed = not (null new_classrels);
+  in
+    (needed,
+      if needed then map_proven_classrels (fold Symreltab.update new_classrels) thy
+      else thy)
+  end;
+
+fun complete_classrels thy =
+  let
+    val classrels = proven_classrels_of thy;
+    val diff_merge_classrels = diff_merge_classrels_of thy;
+    val (needed, thy') = (false, thy) |>
+      fold (fn c12 => fn (needed, thy) =>
+          put_trancl_classrel (c12, Symreltab.lookup classrels c12 |> the |> #1) thy
+          |>> (fn b => needed orelse b))
+        diff_merge_classrels;
+  in
+    if null diff_merge_classrels then NONE
+    else SOME (clear_diff_merge_classrels thy')
+  end;
 
 
 fun the_arity thy a (c, Ss) =
-  (case AList.lookup (op =) (Symtab.lookup_list (#2 (get_instances thy)) a) (c, Ss) of
-    SOME (th, _) => Thm.transfer thy th
+  (case AList.lookup (op =) (Symtab.lookup_list (proven_arities_of thy) a) (c, Ss) of
+    SOME arity => arity
   | NONE => error ("Unproven type arity " ^
       Syntax.string_of_arity (ProofContext.init thy) (a, Ss, [c])));
 
+fun the_arity_thm thy a c_Ss = the_arity thy a c_Ss |> #1 |> #1 |> Thm.transfer thy;
+fun the_arity_prf thy a c_Ss = the_arity thy a c_Ss |> #2;
+
 fun thynames_of_arity thy (c, a) =
-  Symtab.lookup_list (#2 (get_instances thy)) a
-  |> map_filter (fn ((c', _), (_, name)) => if c = c' then SOME name else NONE)
+  Symtab.lookup_list (proven_arities_of thy) a
+  |> map_filter (fn ((c', _), ((_, name),_)) => if c = c' then SOME name else NONE)
   |> rev;
 
-fun insert_arity_completions thy (t, ((c, Ss), (th, thy_name))) arities =
+fun insert_arity_completions thy (t, ((c, Ss), ((th, thy_name), _))) arities =
   let
     val algebra = Sign.classes_of thy;
     val super_class_completions =
       Sign.super_classes thy c
       |> filter_out (fn c1 => exists (fn ((c2, Ss2), _) => c1 = c2
           andalso Sorts.sorts_le algebra (Ss2, Ss)) (Symtab.lookup_list arities t));
-    val completions = map (fn c1 => (Sorts.classrel_derivation algebra
-      (fn (th, c2) => fn c3 => th RS the_classrel thy (c2, c3)) (th, c) c1
-        |> Thm.close_derivation, c1)) super_class_completions;
-    val arities' = fold (fn (th1, c1) => Symtab.cons_list (t, ((c1, Ss), (th1, thy_name))))
+    val names_and_Ss = Name.names Name.context Name.aT (map (K []) Ss);
+    val completions = super_class_completions |> map (fn c1 =>
+      let
+        val th1 = (th RS the_classrel_thm thy (c, c1))
+          |> Drule.instantiate' (map (SOME o ctyp_of thy o TVar o apfst (rpair 0)) names_and_Ss) []
+          |> Thm.close_derivation;
+        val prf1 = Thm.proof_of th1;
+      in (((th1, thy_name), prf1), c1) end);
+    val arities' = fold (fn (th_thy_prf1, c1) => Symtab.cons_list (t, ((c1, Ss), th_thy_prf1)))
       completions arities;
   in (null completions, arities') end;
 
 fun put_arity ((t, Ss, c), th) thy =
   let
-    val arity' = (t, ((c, Ss), (th, Context.theory_name thy)));
+    val arity' = (t, ((c, Ss), ((th, Context.theory_name thy), Thm.proof_of th)));
   in
     thy
-    |> map_instances (fn (classrel, arities) => (classrel,
-      arities
-      |> Symtab.insert_list (eq_fst op =) arity'
-      |> insert_arity_completions thy arity'
-      |> snd))
+    |> map_proven_arities
+      (Symtab.insert_list (eq_fst op =) arity' #>
+        insert_arity_completions thy arity' #> snd)
   end;
 
 fun complete_arities thy =
   let
-    val arities = snd (get_instances thy);
+    val arities = proven_arities_of thy;
     val (finished, arities') = arities
       |> fold_map (insert_arity_completions thy) (Symtab.dest_list arities);
   in
-    if forall I finished then NONE
-    else SOME (thy |> map_instances (fn (classrel, _) => (classrel, arities')))
+    if forall I finished
+    then NONE
+    else SOME (map_proven_arities (K arities') thy)
   end;
 
-val _ = Context.>> (Context.map_theory (Theory.at_begin complete_arities));
+val _ = Context.>> (Context.map_theory
+  (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities));
 
 
 (* maintain instance parameters *)
 
-val get_inst_params = #2 o #2 o AxClassData.get;
-val map_inst_params = AxClassData.map o apsnd o apsnd;
+fun get_inst_param thy (c, tyco) =
+  (case Symtab.lookup (the_default Symtab.empty (Symtab.lookup (#1 (inst_params_of thy)) c)) tyco of
+    SOME c' => c'
+  | NONE => error ("No instance parameter for constant " ^ quote c ^ " on type " ^ quote tyco));
 
-fun get_inst_param thy (c, tyco) =
-  case Symtab.lookup ((the_default Symtab.empty o Symtab.lookup (fst (get_inst_params thy))) c) tyco
-   of SOME c' => c'
-    | NONE => error ("No instance parameter for constant " ^ quote c
-        ^ " on type constructor " ^ quote tyco);
-
-fun add_inst_param (c, tyco) inst = (map_inst_params o apfst
-      o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
+fun add_inst_param (c, tyco) inst =
+  (map_inst_params o apfst o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
   #> (map_inst_params o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
 
-val inst_of_param = Symtab.lookup o snd o get_inst_params;
+val inst_of_param = Symtab.lookup o #2 o inst_params_of;
 val param_of_inst = fst oo get_inst_param;
 
-fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)
-  (get_inst_params thy) [];
+fun inst_thms thy =
+  Symtab.fold (Symtab.fold (cons o #2 o #2) o #2) (#1 (inst_params_of thy)) [];
 
 fun get_inst_tyco consts = try (fst o dest_Type o the_single o Consts.typargs consts);
 
@@ -248,18 +325,20 @@
 fun unoverload_conv thy = MetaSimplifier.rewrite true (inst_thms thy);
 fun overload_conv thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));
 
-fun lookup_inst_param consts params (c, T) = case get_inst_tyco consts (c, T)
- of SOME tyco => AList.lookup (op =) params (c, tyco)
-  | NONE => NONE;
+fun lookup_inst_param consts params (c, T) =
+  (case get_inst_tyco consts (c, T) of
+    SOME tyco => AList.lookup (op =) params (c, tyco)
+  | NONE => NONE);
 
 fun unoverload_const thy (c_ty as (c, _)) =
-  if is_some (class_of_param thy c)
-  then case get_inst_tyco (Sign.consts_of thy) c_ty
-   of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c
-    | NONE => c
+  if is_some (class_of_param thy c) then
+    (case get_inst_tyco (Sign.consts_of thy) c_ty of
+      SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c
+    | NONE => c)
   else c;
 
 
+
 (** instances **)
 
 (* class relations *)
@@ -331,6 +410,8 @@
 
 (* primitive rules *)
 
+val shyps_topped = forall null o #shyps o Thm.rep_thm;
+
 fun add_classrel raw_th thy =
   let
     val th = Thm.strip_shyps (Thm.transfer thy raw_th);
@@ -338,10 +419,14 @@
     fun err () = raise THM ("add_classrel: malformed class relation", 0, [th]);
     val rel = Logic.dest_classrel prop handle TERM _ => err ();
     val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();
+    val th' = th
+      |> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [c1])))] []
+      |> Thm.unconstrain_allTs;
+    val _ = shyps_topped th' orelse raise Fail "add_classrel: nontop shyps after unconstrain";
   in
     thy
     |> Sign.primitive_classrel (c1, c2)
-    |> put_classrel ((c1, c2), Thm.close_derivation (Drule.unconstrainTs th))
+    |> (snd oo put_trancl_classrel) ((c1, c2), th')
     |> perhaps complete_arities
   end;
 
@@ -351,17 +436,22 @@
     val prop = Thm.plain_prop_of th;
     fun err () = raise THM ("add_arity: malformed type arity", 0, [th]);
     val (t, Ss, c) = Logic.dest_arity prop handle TERM _ => err ();
-    val T = Type (t, map TFree (Name.names Name.context Name.aT Ss));
+    val names = Name.names Name.context Name.aT Ss;
+    val T = Type (t, map TFree names);
     val missing_params = Sign.complete_sort thy [c]
       |> maps (these o Option.map #params o try (get_info thy))
       |> filter_out (fn (const, _) => can (get_inst_param thy) (const, t))
       |> (map o apsnd o map_atyps) (K T);
     val _ = map (Sign.certify_sort thy) Ss = Ss orelse err ();
+    val th' = th
+      |> Drule.instantiate' (map (SOME o ctyp_of thy o TVar o apfst (rpair 0)) names) []
+      |> Thm.unconstrain_allTs;
+    val _ = shyps_topped th' orelse raise Fail "add_arity: nontop shyps after unconstrain";
   in
     thy
     |> fold (snd oo declare_overloaded) missing_params
     |> Sign.primitive_arity (t, Ss, [c])
-    |> put_arity ((t, Ss, c), Thm.close_derivation (Drule.unconstrainTs th))
+    |> put_arity ((t, Ss, c), th')
   end;
 
 
@@ -483,25 +573,24 @@
       def_thy
       |> Sign.qualified_path true bconst
       |> PureThy.note_thmss ""
-        [((Binding.name introN, []), [([Drule.export_without_context raw_intro], [])]),
-         ((Binding.name superN, []), [(map Drule.export_without_context raw_classrel, [])]),
-         ((Binding.name axiomsN, []),
+        [((Binding.name "intro", []), [([Drule.export_without_context raw_intro], [])]),
+         ((Binding.name "super", []), [(map Drule.export_without_context raw_classrel, [])]),
+         ((Binding.name "axioms", []),
            [(map (fn th => Drule.export_without_context (class_triv RS th)) raw_axioms, [])])]
       ||> Sign.restore_naming def_thy;
 
 
     (* result *)
 
-    val axclass = make_axclass ((def, intro, axioms), params);
+    val axclass = make_axclass (def, intro, axioms, params);
     val result_thy =
       facts_thy
-      |> fold put_classrel (map (pair class) super ~~ classrel)
+      |> fold (snd oo put_trancl_classrel) (map (pair class) super ~~ classrel)
       |> Sign.qualified_path false bconst
       |> PureThy.note_thmss "" (name_atts ~~ map Thm.simple_fact (unflat axiomss axioms)) |> snd
       |> Sign.restore_naming facts_thy
-      |> map_axclasses (fn (axclasses, parameters) =>
-        (Symtab.update (class, axclass) axclasses,
-          fold (fn (x, _) => add_param pp (x, class)) params parameters));
+      |> map_axclasses (Symtab.update (class, axclass))
+      |> map_params (fold (fn (x, _) => add_param pp (x, class)) params);
 
   in (class, result_thy) end;
 
--- a/src/Pure/display.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/display.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -182,7 +182,8 @@
     val extern_const = Name_Space.extern (#1 constants);
     val {classes, default, types, ...} = Type.rep_tsig tsig;
     val (class_space, class_algebra) = classes;
-    val {classes, arities} = Sorts.rep_algebra class_algebra;
+    val classes = Sorts.classes_of class_algebra;
+    val arities = Sorts.arities_of class_algebra;
 
     val clsses = Name_Space.dest_table (class_space, Symtab.make (Graph.dest classes));
     val tdecls = Name_Space.dest_table types;
--- a/src/Pure/drule.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/drule.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -106,7 +106,6 @@
   val dummy_thm: thm
   val sort_constraintI: thm
   val sort_constraint_eq: thm
-  val unconstrainTs: thm -> thm
   val with_subgoal: int -> (thm -> thm) -> thm -> thm
   val comp_no_flatten: thm * int -> int -> thm -> thm
   val rename_bvars: (string * string) list -> thm -> thm
@@ -204,12 +203,6 @@
 
 (** Standardization of rules **)
 
-(* type classes and sorts *)
-
-fun unconstrainTs th =
-  fold (Thm.unconstrainT o Thm.ctyp_of (Thm.theory_of_thm th) o TVar)
-    (Thm.fold_terms Term.add_tvars th []) th;
-
 (*Generalization over a list of variables*)
 val forall_intr_list = fold_rev forall_intr;
 
--- a/src/Pure/item_net.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/item_net.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -11,6 +11,7 @@
   val init: ('a * 'a -> bool) -> ('a -> term list) -> 'a T
   val content: 'a T -> 'a list
   val retrieve: 'a T -> term -> 'a list
+  val member: 'a T -> 'a -> bool
   val merge: 'a T * 'a T -> 'a T
   val remove: 'a -> 'a T -> 'a T
   val update: 'a -> 'a T -> 'a T
@@ -49,7 +50,6 @@
   mk_items eq index (x :: content) (next - 1)
     (fold (fn t => Net.insert_term (K false) (t, (next, x))) (index x) net);
 
-
 fun merge (items1, Items {content = content2, ...}) =
   fold_rev (fn y => if member items1 y then I else cons y) content2 items1;
 
--- a/src/Pure/sorts.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/sorts.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -25,9 +25,8 @@
   val insert_term: term -> sort OrdList.T -> sort OrdList.T
   val insert_terms: term list -> sort OrdList.T -> sort OrdList.T
   type algebra
-  val rep_algebra: algebra ->
-   {classes: serial Graph.T,
-    arities: (class * (class * sort list)) list Symtab.table}
+  val classes_of: algebra -> serial Graph.T
+  val arities_of: algebra -> (class * (class * sort list)) list Symtab.table
   val all_classes: algebra -> class list
   val super_classes: algebra -> class -> class list
   val class_less: algebra -> class * class -> bool
@@ -116,10 +115,8 @@
  {classes: serial Graph.T,
   arities: (class * (class * sort list)) list Symtab.table};
 
-fun rep_algebra (Algebra args) = args;
-
-val classes_of = #classes o rep_algebra;
-val arities_of = #arities o rep_algebra;
+fun classes_of (Algebra {classes, ...}) = classes;
+fun arities_of (Algebra {arities, ...}) = arities;
 
 fun make_algebra (classes, arities) =
   Algebra {classes = classes, arities = arities};
--- a/src/Pure/thm.ML	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Pure/thm.ML	Mon Apr 26 21:25:32 2010 +0200
@@ -92,8 +92,6 @@
   val instantiate: (ctyp * ctyp) list * (cterm * cterm) list -> thm -> thm
   val instantiate_cterm: (ctyp * ctyp) list * (cterm * cterm) list -> cterm -> cterm
   val trivial: cterm -> thm
-  val of_class: ctyp * class -> thm
-  val unconstrainT: ctyp -> thm -> thm
   val dest_state: thm * int -> (term * term) list * term list * term * term
   val lift_rule: cterm -> thm -> thm
   val incr_indexes: int -> thm -> thm
@@ -139,6 +137,9 @@
   val adjust_maxidx_thm: int -> thm -> thm
   val varifyT_global: thm -> thm
   val varifyT_global': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm
+  val of_class: ctyp * class -> thm
+  val unconstrainT: ctyp -> thm -> thm
+  val unconstrain_allTs: thm -> thm
   val freezeT: thm -> thm
   val assumption: int -> thm -> thm Seq.seq
   val eq_assumption: int -> thm -> thm
@@ -1240,6 +1241,11 @@
       prop = Logic.list_implies (constraints, unconstrain prop)})
   end;
 
+fun unconstrain_allTs th =
+  fold (unconstrainT o ctyp_of (theory_of_thm th) o TVar)
+    (fold_terms Term.add_tvars th []) th;
+
+
 (* Replace all TFrees not fixed or in the hyps by new TVars *)
 fun varifyT_global' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
   let
--- a/src/Sequents/LK/Quantifiers.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/Sequents/LK/Quantifiers.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -90,13 +90,13 @@
   oops
 
 (*INVALID*)
-lemma "|- P(?a) --> (ALL x. P(x))"
+schematic_lemma "|- P(?a) --> (ALL x. P(x))"
   apply fast?
   apply (rule _)
   oops
 
 (*INVALID*)
-lemma "|- (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
+schematic_lemma "|- (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
   apply fast?
   apply (rule _)
   oops
@@ -114,7 +114,7 @@
 
 text "Solving for a Var"
 
-lemma "|- (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
+schematic_lemma "|- (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
   by fast
 
 
--- a/src/ZF/AC/Cardinal_aux.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/ZF/AC/Cardinal_aux.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -71,7 +71,7 @@
 apply (rule Un_upper1 [THEN subset_imp_lepoll]) 
 done
 
-lemma paired_bij: "?f \<in> bij({{y,z}. y \<in> x}, x)"
+schematic_lemma paired_bij: "?f \<in> bij({{y,z}. y \<in> x}, x)"
 apply (rule RepFun_bijective)
 apply (simp add: doubleton_eq_iff, blast)
 done
--- a/src/ZF/Constructible/Reflection.thy	Mon Apr 26 21:20:43 2010 +0200
+++ b/src/ZF/Constructible/Reflection.thy	Mon Apr 26 21:25:32 2010 +0200
@@ -268,7 +268,7 @@
 text{*Example 1: reflecting a simple formula.  The reflecting class is first
 given as the variable @{text ?Cl} and later retrieved from the final 
 proof state.*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & x \<in> y, 
                \<lambda>a x. \<exists>y\<in>Mset(a). x \<in> y)"
@@ -285,7 +285,7 @@
 
 
 text{*Example 2*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
                \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
@@ -302,14 +302,14 @@
 by fast
 
 text{*Example 2''.  We expand the subset relation.*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
   "Reflects(?Cl,
         \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> (\<forall>w. M(w) --> w\<in>z --> w\<in>x) --> z\<in>y),
         \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). (\<forall>w\<in>Mset(a). w\<in>z --> w\<in>x) --> z\<in>y)"
 by fast
 
 text{*Example 2'''.  Single-step version, to reveal the reflecting class.*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<subseteq> x --> z \<in> y), 
                \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<subseteq> x --> z \<in> y)" 
@@ -329,21 +329,21 @@
 
 text{*Example 3.  Warning: the following examples make sense only
 if @{term P} is quantifier-free, since it is not being relativized.*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & (\<forall>z. M(z) --> z \<in> y <-> z \<in> x & P(z)), 
                \<lambda>a x. \<exists>y\<in>Mset(a). \<forall>z\<in>Mset(a). z \<in> y <-> z \<in> x & P(z))"
 by fast
 
 text{*Example 3'*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & y = Collect(x,P),
                \<lambda>a x. \<exists>y\<in>Mset(a). y = Collect(x,P))";
 by fast
 
 text{*Example 3''*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>x. \<exists>y. M(y) & y = Replace(x,P),
                \<lambda>a x. \<exists>y\<in>Mset(a). y = Replace(x,P))";
@@ -351,7 +351,7 @@
 
 text{*Example 4: Axiom of Choice.  Possibly wrong, since @{text \<Pi>} needs
 to be relativized.*}
-lemma (in reflection) 
+schematic_lemma (in reflection) 
      "Reflects(?Cl,
                \<lambda>A. 0\<notin>A --> (\<exists>f. M(f) & f \<in> (\<Pi> X \<in> A. X)),
                \<lambda>a A. 0\<notin>A --> (\<exists>f\<in>Mset(a). f \<in> (\<Pi> X \<in> A. X)))"