Added remote-SOS changes by Philipp Meyer
authornipkow
Wed, 29 Jul 2009 09:06:49 +0200
changeset 32269 b642e4813e53
parent 32267 99711ef9d582 (diff)
parent 32268 d50f0cb67578 (current diff)
child 32270 615c524bd9e4
child 32272 cc1bf9077167
Added remote-SOS changes by Philipp Meyer
--- a/Admin/isatest/settings/mac-poly-M4	Fri Jul 24 13:56:02 2009 +0200
+++ b/Admin/isatest/settings/mac-poly-M4	Wed Jul 29 09:06:49 2009 +0200
@@ -23,6 +23,6 @@
 ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
 ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
 
-ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -M 4 -t true"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 0"
--- a/Admin/isatest/settings/mac-poly-M8	Fri Jul 24 13:56:02 2009 +0200
+++ b/Admin/isatest/settings/mac-poly-M8	Wed Jul 29 09:06:49 2009 +0200
@@ -23,6 +23,6 @@
 ISABELLE_OUTPUT="$ISABELLE_HOME_USER/heaps"
 ISABELLE_BROWSER_INFO="$ISABELLE_HOME_USER/browser_info"
 
-ISABELLE_USEDIR_OPTIONS="-i false -d false -M 8"
+ISABELLE_USEDIR_OPTIONS="-i false -d false -M 8 -t true"
 
 HOL_USEDIR_OPTIONS="-p 2 -q 0"
--- a/NEWS	Fri Jul 24 13:56:02 2009 +0200
+++ b/NEWS	Wed Jul 29 09:06:49 2009 +0200
@@ -18,6 +18,8 @@
 
 *** HOL ***
 
+* Set.UNIV and Set.empty are mere abbreviations for top and bot.  INCOMPATIBILITY.
+
 * More convenient names for set intersection and union.  INCOMPATIBILITY:
 
     Set.Int ~>  Set.inter
@@ -31,7 +33,7 @@
 
 * New quickcheck implementation using new code generator.
 
-* New type class boolean_algebra.
+* New class "boolean_algebra".
 
 * Refinements to lattices classes:
   - added boolean_algebra type class
@@ -40,6 +42,19 @@
   - lemma ge_sup_conv renamed to le_sup_iff, in accordance with le_inf_iff
   - dropped lemma alias inf_ACI for inf_aci (same for sup_ACI and sup_aci)
   - renamed ACI to inf_sup_aci
+  - class "complete_lattice" moved to separate theory "complete_lattice";
+    corresponding constants renamed:
+    
+    Set.Inf ~>      Complete_Lattice.Inf
+    Set.Sup ~>      Complete_Lattice.Sup
+    Set.INFI ~>     Complete_Lattice.INFI
+    Set.SUPR ~>     Complete_Lattice.SUPR
+    Set.Inter ~>    Complete_Lattice.Inter
+    Set.Union ~>    Complete_Lattice.Union
+    Set.INTER ~>    Complete_Lattice.INTER
+    Set.UNION ~>    Complete_Lattice.UNION
+
+  INCOMPATIBILITY.
 
 * Class semiring_div requires superclass no_zero_divisors and proof of
 div_mult_mult1; theorems div_mult_mult1, div_mult_mult2,
@@ -54,6 +69,13 @@
 multiplicative monoids retains syntax "^" and is now defined generic
 in class power.  INCOMPATIBILITY.
 
+* Relation composition "R O S" now has a "more standard" argument order,
+i.e., "R O S = {(x,z). EX y. (x,y) : R & (y,z) : S }".
+INCOMPATIBILITY: Rewrite propositions with "S O R" --> "R O S". Proofs
+may occationally break, since the O_assoc rule was not rewritten like this.
+Fix using O_assoc[symmetric].
+The same applies to the curried version "R OO S".
+
 * GCD now has functions Gcd/GCD and Lcm/LCM for the gcd and lcm of finite and
 infinite sets. It is shown that they form a complete lattice.
 
@@ -118,6 +140,11 @@
 
 *** ML ***
 
+* Tactical FOCUS is similar to SUBPROOF, but allows the body tactic to
+introduce new subgoals and schematic variables.  FOCUS_PARAMS is
+similar, but focuses on the parameter prefix only, leaving subgoal
+premises unchanged.
+
 * Renamed functor TableFun to Table, and GraphFun to Graph.  (Since
 functors have their own ML name space there is no point to mark them
 separately.)  Minor INCOMPATIBILITY.
--- a/doc-src/IsarImplementation/Thy/Integration.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Integration.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -239,20 +239,20 @@
 
 text %mlref {*
   \begin{mldecls}
-  @{index_ML the_context: "unit -> theory"} \\
+  @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
   @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\
   \end{mldecls}
 
   \begin{description}
 
-  \item @{ML "the_context ()"} refers to the theory context of the
-  {\ML} toplevel --- at compile time!  {\ML} code needs to take care
-  to refer to @{ML "the_context ()"} correctly.  Recall that
-  evaluation of a function body is delayed until actual runtime.
-  Moreover, persistent {\ML} toplevel bindings to an unfinished theory
-  should be avoided: code should either project out the desired
-  information immediately, or produce an explicit @{ML_type
-  theory_ref} (cf.\ \secref{sec:context-theory}).
+  \item @{ML "ML_Context.the_generic_context ()"} refers to the theory
+  context of the {\ML} toplevel --- at compile time!  {\ML} code needs
+  to take care to refer to @{ML "ML_Context.the_generic_context ()"}
+  correctly.  Recall that evaluation of a function body is delayed
+  until actual runtime.  Moreover, persistent {\ML} toplevel bindings
+  to an unfinished theory should be avoided: code should either
+  project out the desired information immediately, or produce an
+  explicit @{ML_type theory_ref} (cf.\ \secref{sec:context-theory}).
 
   \item @{ML "Context.>>"}~@{text f} applies context transformation
   @{text f} to the implicit context of the {\ML} toplevel.
--- a/doc-src/IsarImplementation/Thy/Proof.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Proof.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -95,7 +95,7 @@
   @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
   @{index_ML Variable.import: "bool -> thm list -> Proof.context ->
   ((ctyp list * cterm list) * thm list) * Proof.context"} \\
-  @{index_ML Variable.focus: "cterm -> Proof.context -> (cterm list * cterm) * Proof.context"} \\
+  @{index_ML Variable.focus: "cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context"} \\
   \end{mldecls}
 
   \begin{description}
@@ -284,10 +284,7 @@
 
 text %mlref {*
   \begin{mldecls}
-  @{index_ML SUBPROOF:
-  "({context: Proof.context, schematics: ctyp list * cterm list,
-    params: cterm list, asms: cterm list, concl: cterm,
-    prems: thm list} -> tactic) -> Proof.context -> int -> tactic"} \\
+  @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\
   \end{mldecls}
   \begin{mldecls}
   @{index_ML Goal.prove: "Proof.context -> string list -> term list -> term ->
@@ -297,7 +294,7 @@
   \end{mldecls}
   \begin{mldecls}
   @{index_ML Obtain.result: "(Proof.context -> tactic) ->
-  thm list -> Proof.context -> (cterm list * thm list) * Proof.context"} \\
+  thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarImplementation/Thy/Tactic.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/Tactic.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -63,7 +63,7 @@
 text %mlref {*
   \begin{mldecls}
   @{index_ML Goal.init: "cterm -> thm"} \\
-  @{index_ML Goal.finish: "thm -> thm"} \\
+  @{index_ML Goal.finish: "Proof.context -> thm -> thm"} \\
   @{index_ML Goal.protect: "thm -> thm"} \\
   @{index_ML Goal.conclude: "thm -> thm"} \\
   \end{mldecls}
@@ -73,9 +73,10 @@
   \item @{ML "Goal.init"}~@{text C} initializes a tactical goal from
   the well-formed proposition @{text C}.
 
-  \item @{ML "Goal.finish"}~@{text "thm"} checks whether theorem
+  \item @{ML "Goal.finish"}~@{text "ctxt thm"} checks whether theorem
   @{text "thm"} is a solved goal (no subgoals), and concludes the
-  result by removing the goal protection.
+  result by removing the goal protection.  The context is only
+  required for printing error messages.
 
   \item @{ML "Goal.protect"}~@{text "thm"} protects the full statement
   of theorem @{text "thm"}.
--- a/doc-src/IsarImplementation/Thy/document/Integration.tex	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Integration.tex	Wed Jul 29 09:06:49 2009 +0200
@@ -300,19 +300,20 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
-  \indexdef{}{ML}{the\_context}\verb|the_context: unit -> theory| \\
+  \indexdef{}{ML}{ML\_Context.the\_generic\_context}\verb|ML_Context.the_generic_context: unit -> Context.generic| \\
   \indexdef{}{ML}{Context.$>$$>$ }\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\
   \end{mldecls}
 
   \begin{description}
 
-  \item \verb|the_context ()| refers to the theory context of the
-  {\ML} toplevel --- at compile time!  {\ML} code needs to take care
-  to refer to \verb|the_context ()| correctly.  Recall that
-  evaluation of a function body is delayed until actual runtime.
-  Moreover, persistent {\ML} toplevel bindings to an unfinished theory
-  should be avoided: code should either project out the desired
-  information immediately, or produce an explicit \verb|theory_ref| (cf.\ \secref{sec:context-theory}).
+  \item \verb|ML_Context.the_generic_context ()| refers to the theory
+  context of the {\ML} toplevel --- at compile time!  {\ML} code needs
+  to take care to refer to \verb|ML_Context.the_generic_context ()|
+  correctly.  Recall that evaluation of a function body is delayed
+  until actual runtime.  Moreover, persistent {\ML} toplevel bindings
+  to an unfinished theory should be avoided: code should either
+  project out the desired information immediately, or produce an
+  explicit \verb|theory_ref| (cf.\ \secref{sec:context-theory}).
 
   \item \verb|Context.>>|~\isa{f} applies context transformation
   \isa{f} to the implicit context of the {\ML} toplevel.
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex	Wed Jul 29 09:06:49 2009 +0200
@@ -113,7 +113,7 @@
   \indexdef{}{ML}{Variable.polymorphic}\verb|Variable.polymorphic: Proof.context -> term list -> term list| \\
   \indexdef{}{ML}{Variable.import}\verb|Variable.import: bool -> thm list -> Proof.context ->|\isasep\isanewline%
 \verb|  ((ctyp list * cterm list) * thm list) * Proof.context| \\
-  \indexdef{}{ML}{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context| \\
+  \indexdef{}{ML}{Variable.focus}\verb|Variable.focus: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context| \\
   \end{mldecls}
 
   \begin{description}
@@ -324,9 +324,7 @@
 %
 \begin{isamarkuptext}%
 \begin{mldecls}
-  \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: ({context: Proof.context, schematics: ctyp list * cterm list,|\isasep\isanewline%
-\verb|    params: cterm list, asms: cterm list, concl: cterm,|\isasep\isanewline%
-\verb|    prems: thm list} -> tactic) -> Proof.context -> int -> tactic| \\
+  \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) -> Proof.context -> int -> tactic| \\
   \end{mldecls}
   \begin{mldecls}
   \indexdef{}{ML}{Goal.prove}\verb|Goal.prove: Proof.context -> string list -> term list -> term ->|\isasep\isanewline%
@@ -336,7 +334,7 @@
   \end{mldecls}
   \begin{mldecls}
   \indexdef{}{ML}{Obtain.result}\verb|Obtain.result: (Proof.context -> tactic) ->|\isasep\isanewline%
-\verb|  thm list -> Proof.context -> (cterm list * thm list) * Proof.context| \\
+\verb|  thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context| \\
   \end{mldecls}
 
   \begin{description}
--- a/doc-src/IsarImplementation/Thy/document/Tactic.tex	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/IsarImplementation/Thy/document/Tactic.tex	Wed Jul 29 09:06:49 2009 +0200
@@ -84,7 +84,7 @@
 \begin{isamarkuptext}%
 \begin{mldecls}
   \indexdef{}{ML}{Goal.init}\verb|Goal.init: cterm -> thm| \\
-  \indexdef{}{ML}{Goal.finish}\verb|Goal.finish: thm -> thm| \\
+  \indexdef{}{ML}{Goal.finish}\verb|Goal.finish: Proof.context -> thm -> thm| \\
   \indexdef{}{ML}{Goal.protect}\verb|Goal.protect: thm -> thm| \\
   \indexdef{}{ML}{Goal.conclude}\verb|Goal.conclude: thm -> thm| \\
   \end{mldecls}
@@ -94,9 +94,10 @@
   \item \verb|Goal.init|~\isa{C} initializes a tactical goal from
   the well-formed proposition \isa{C}.
 
-  \item \verb|Goal.finish|~\isa{thm} checks whether theorem
+  \item \verb|Goal.finish|~\isa{ctxt\ thm} checks whether theorem
   \isa{thm} is a solved goal (no subgoals), and concludes the
-  result by removing the goal protection.
+  result by removing the goal protection.  The context is only
+  required for printing error messages.
 
   \item \verb|Goal.protect|~\isa{thm} protects the full statement
   of theorem \isa{thm}.
--- a/doc-src/Main/Docs/Main_Doc.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/Main/Docs/Main_Doc.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -79,8 +79,8 @@
 \begin{tabular}{@ {} l @ {~::~} l @ {}}
 @{const Lattices.inf} & @{typeof Lattices.inf}\\
 @{const Lattices.sup} & @{typeof Lattices.sup}\\
-@{const Set.Inf} & @{term_type_only Set.Inf "'a set \<Rightarrow> 'a::complete_lattice"}\\
-@{const Set.Sup} & @{term_type_only Set.Sup "'a set \<Rightarrow> 'a::complete_lattice"}\\
+@{const Complete_Lattice.Inf} & @{term_type_only Complete_Lattice.Inf "'a set \<Rightarrow> 'a::complete_lattice"}\\
+@{const Complete_Lattice.Sup} & @{term_type_only Complete_Lattice.Sup "'a set \<Rightarrow> 'a::complete_lattice"}\\
 \end{tabular}
 
 \subsubsection*{Syntax}
@@ -110,8 +110,8 @@
 @{const Set.insert} & @{term_type_only insert "'a\<Rightarrow>'a set\<Rightarrow>'a set"}\\
 @{const Collect} & @{term_type_only Collect "('a\<Rightarrow>bool)\<Rightarrow>'a set"}\\
 @{const "op :"} & @{term_type_only "op :" "'a\<Rightarrow>'a set\<Rightarrow>bool"} & (\texttt{:})\\
-@{const Set.union} & @{term_type_only Set.Un "'a set\<Rightarrow>'a set \<Rightarrow> 'a set"} & (\texttt{Un})\\
-@{const Set.inter} & @{term_type_only Set.Int "'a set\<Rightarrow>'a set \<Rightarrow> 'a set"} & (\texttt{Int})\\
+@{const Set.union} & @{term_type_only Set.union "'a set\<Rightarrow>'a set \<Rightarrow> 'a set"} & (\texttt{Un})\\
+@{const Set.inter} & @{term_type_only Set.inter "'a set\<Rightarrow>'a set \<Rightarrow> 'a set"} & (\texttt{Int})\\
 @{const UNION} & @{term_type_only UNION "'a set\<Rightarrow>('a \<Rightarrow> 'b set) \<Rightarrow> 'b set"}\\
 @{const INTER} & @{term_type_only INTER "'a set\<Rightarrow>('a \<Rightarrow> 'b set) \<Rightarrow> 'b set"}\\
 @{const Union} & @{term_type_only Union "'a set set\<Rightarrow>'a set"}\\
@@ -220,7 +220,7 @@
 
 \begin{supertabular}{@ {} l @ {~::~} l @ {}}
 @{const Relation.converse} & @{term_type_only Relation.converse "('a * 'b)set \<Rightarrow> ('b*'a)set"}\\
-@{const Relation.rel_comp} & @{term_type_only Relation.rel_comp "('a*'b)set\<Rightarrow>('c*'a)set\<Rightarrow>('c*'b)set"}\\
+@{const Relation.rel_comp} & @{term_type_only Relation.rel_comp "('a*'b)set\<Rightarrow>('b*'c)set\<Rightarrow>('a*'c)set"}\\
 @{const Relation.Image} & @{term_type_only Relation.Image "('a*'b)set\<Rightarrow>'a set\<Rightarrow>'b set"}\\
 @{const Relation.inv_image} & @{term_type_only Relation.inv_image "('a*'a)set\<Rightarrow>('b\<Rightarrow>'a)\<Rightarrow>('b*'b)set"}\\
 @{const Relation.Id_on} & @{term_type_only Relation.Id_on "'a set\<Rightarrow>('a*'a)set"}\\
--- a/doc-src/Main/Docs/document/Main_Doc.tex	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/Main/Docs/document/Main_Doc.tex	Wed Jul 29 09:06:49 2009 +0200
@@ -231,7 +231,7 @@
 
 \begin{supertabular}{@ {} l @ {~::~} l @ {}}
 \isa{converse} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}b\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}\\
-\isa{op\ O} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}c\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}c\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ set}\\
+\isa{op\ O} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}b\ {\isasymtimes}\ {\isacharprime}c{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}c{\isacharparenright}\ set}\\
 \isa{op\ {\isacharbackquote}{\isacharbackquote}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharprime}b\ set}\\
 \isa{inv{\isacharunderscore}image} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}b\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}b\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ set}\\
 \isa{Id{\isacharunderscore}on} & \isa{{\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set}\\
--- a/doc-src/TutorialI/Protocol/Message.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Message.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -841,6 +841,8 @@
 (*Apply rules to break down assumptions of the form
   Y \<in> parts(insert X H)  and  Y \<in> analz(insert X H)
 *)
+fun impOfSubs th = th RSN (2, @{thm rev_subsetD})
+
 val Fake_insert_tac = 
     dresolve_tac [impOfSubs Fake_analz_insert,
                   impOfSubs Fake_parts_insert] THEN'
--- a/doc-src/rail.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/rail.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -1,3 +1,12 @@
+(*  Title:      doc-src/rail.ML
+    Author:     Michael Kerscher, TUM
+
+Railroad diagrams for LaTeX.
+*)
+
+structure Rail =
+struct
+
 datatype token =
   Identifier of string |
   Special_Identifier of string |
@@ -9,7 +18,7 @@
 fun is_identifier (Identifier _) = true
   | is_identifier (Special_Identifier _ ) = true
   | is_identifier _ = false;
-  
+
 fun is_text (Text _) = true
   | is_text _ = false;
 
@@ -65,7 +74,7 @@
   ("executable", ("isatt", no_check, true)),
   ("tool", ("isatt", K check_tool, true)),
   ("file", ("isatt", K (File.exists o Path.explode), true)),
-  ("theory", ("", K ThyInfo.known_thy, true)) 
+  ("theory", ("", K ThyInfo.known_thy, true))
   ];
 
 in
@@ -123,7 +132,7 @@
 end;
 
 (* escaped version *)
-fun scan_identifier ctxt = 
+fun scan_identifier ctxt =
   let fun is_identifier_start s =
     Symbol.is_letter s orelse
     s = "_"
@@ -197,13 +206,13 @@
   CR | EMPTY | ANNOTE | IDENT | STRING |
   Null_Kind;
 
-datatype body_type =	
+datatype body_type =
   Body of body_kind * string * string * id_type * body_type list |
   Body_Pos of body_kind * string * string * id_type * body_type list * int * int |
   Empty_Body |
   Null_Body;
 
-datatype rule = 
+datatype rule =
   Rule of id_type * body_type;
 
 fun new_id id kind = Id (id, kind);
@@ -219,7 +228,7 @@
 fun add_list (Body(kind, text, annot, id, bodies), body) =
   Body(kind, text, annot, id, bodies @ [body]);
 
-fun cat_body_lists (Body(kind, text, annot, id, bodies1), Body(_,_,_,_, bodies2)) = 
+fun cat_body_lists (Body(kind, text, annot, id, bodies1), Body(_,_,_,_, bodies2)) =
       Body(kind, text, annot, id, bodies1 @ bodies2);
 
 fun add_body (kind, body1 as Body(kind1,_,_,_,_), body2 as Body(kind2,_,_,_,_)) =
@@ -299,7 +308,7 @@
 	add_body(PLUS, new_empty_body, rev_body body1)
       else
 	add_body(BAR, new_empty_body, add_body (PLUS, body1, rev_body body2)) ) ||
-  parse_body2 -- ($$$ "+" |-- !!! "body4e expected" (parse_body4e)) >> 
+  parse_body2 -- ($$$ "+" |-- !!! "body4e expected" (parse_body4e)) >>
     (fn (body1, body2) => new_body (PLUS, body1, rev_body body2) ) ||
   parse_body2e
   ) x
@@ -326,7 +335,7 @@
   (
   $$$ "(" |-- !!! "body0 or ')' expected" (parse_body --| $$$ ")") ||
   Scan.one is_text -- (Scan.optional (Scan.one is_anot) (Anot(""))) >>
-    (fn (text, anot) => new_text_annote_body (text,anot)) ||    
+    (fn (text, anot) => new_text_annote_body (text,anot)) ||
   Scan.one is_identifier -- (Scan.optional (Scan.one is_anot) (Anot(""))) >>
     (fn (id, anot) => new_ident_body (id,anot)) ||
   $$$ "\\" >> (fn _ => new_body (CR, Null_Body, Null_Body))
@@ -401,7 +410,7 @@
     in
       format_bodies(bodies)
     end
-  | format_body (Body_Pos(BAR,_,_,_,bodies,_,_),cent) = 
+  | format_body (Body_Pos(BAR,_,_,_,bodies,_,_),cent) =
     let fun format_bodies([]) = "\\rail@endbar\n"
 	  | format_bodies(x::xs) =
 	      "\\rail@nextbar{" ^ string_of_int(getystart(x)) ^"}\n" ^
@@ -466,3 +475,6 @@
 
 val _ = ThyOutput.antiquotation "rail" (Scan.lift ( OuterParse.position Args.name ))
   (fn {context = ctxt,...} => fn txt => process txt ctxt);
+
+end;
+
--- a/doc-src/railsetup.sty	Fri Jul 24 13:56:02 2009 +0200
+++ b/doc-src/railsetup.sty	Wed Jul 29 09:06:49 2009 +0200
@@ -1,10 +1,4 @@
-
-\railalias{percent}{\%}
-\railalias{ppercent}{\%\%}
-\railalias{underscore}{\_}
-\railalias{lbrace}{\texttt{\ttlbrace}}
-\railalias{rbrace}{\texttt{\ttrbrace}}
-\railalias{atsign}{{\at}}
+%% dimensions
 
 \setlength\railextra{3.6ex}
 \setlength\railboxleft{0.9ex}
@@ -20,6 +14,22 @@
 \setlength\railtextup{5pt}
 \setlength\railjoinsize{16pt}
 
+
+%% rail antiquotation environment
+
+\newenvironment{railoutput}%
+{\begin{list}{}{\rail@param}\def\rail@expand{\relax}\makeatletter}{\end{list}}
+
+
+%% old-style content markup
+
+\railalias{percent}{\%}
+\railalias{ppercent}{\%\%}
+\railalias{underscore}{\_}
+\railalias{lbrace}{\texttt{\ttlbrace}}
+\railalias{rbrace}{\texttt{\ttrbrace}}
+\railalias{atsign}{{\at}}
+
 \def\rail@termfont{\small\ttfamily\upshape}
 \def\rail@tokenfont{\small\ttfamily\upshape}
 \def\rail@nontfont{\small\rmfamily\upshape}
--- a/src/CCL/CCL.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/CCL/CCL.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -255,14 +255,20 @@
 val caseB_lemmas = mk_lemmas @{thms caseBs}
 
 val ccl_dstncts =
-        let fun mk_raw_dstnct_thm rls s =
-                  prove_goal @{theory} s (fn _=> [rtac notI 1,eresolve_tac rls 1])
-        in map (mk_raw_dstnct_thm caseB_lemmas)
-                (mk_dstnct_rls @{theory} ["bot","true","false","pair","lambda"]) end
+  let
+    fun mk_raw_dstnct_thm rls s =
+      Goal.prove_global @{theory} [] [] (Syntax.read_prop_global @{theory} s)
+        (fn _=> rtac notI 1 THEN eresolve_tac rls 1)
+  in map (mk_raw_dstnct_thm caseB_lemmas)
+    (mk_dstnct_rls @{theory} ["bot","true","false","pair","lambda"]) end
 
 fun mk_dstnct_thms thy defs inj_rls xs =
-  let fun mk_dstnct_thm rls s = prove_goalw thy defs s
-    (fn _ => [simp_tac (global_simpset_of thy addsimps (rls@inj_rls)) 1])
+  let
+    fun mk_dstnct_thm rls s =
+      Goal.prove_global thy [] [] (Syntax.read_prop_global thy s)
+        (fn _ =>
+          rewrite_goals_tac defs THEN
+          simp_tac (global_simpset_of thy addsimps (rls @ inj_rls)) 1)
   in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end
 
 fun mkall_dstnct_thms thy defs i_rls xss = maps (mk_dstnct_thms thy defs i_rls) xss
--- a/src/CCL/Wfd.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/CCL/Wfd.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -498,14 +498,13 @@
   structure Data = Named_Thms(val name = "eval" val description = "evaluation rules");
 in
 
-fun eval_tac ctxt ths =
-  METAHYPS (fn prems =>
-    DEPTH_SOLVE_1 (resolve_tac (ths @ prems @ Data.get ctxt) 1)) 1;
+fun eval_tac ctxt ths i =
+  DEPTH_SOLVE_1 (resolve_tac (ths @ Data.get ctxt) i ORELSE assume_tac i);
 
 val eval_setup =
   Data.setup #>
   Method.setup @{binding eval}
-    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (CHANGED (eval_tac ctxt ths))))
+    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ctxt ths)))
     "evaluation";
 
 end;
--- a/src/FOL/FOL.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/FOL/FOL.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -12,7 +12,6 @@
   "~~/src/Provers/clasimp.ML"
   "~~/src/Tools/induct.ML"
   ("cladata.ML")
-  ("blastdata.ML")
   ("simpdata.ML")
 begin
 
@@ -171,7 +170,25 @@
 setup Cla.setup
 setup cla_setup
 
-use "blastdata.ML"
+ML {*
+  structure Blast = Blast
+  (
+    val thy = @{theory}
+    type claset	= Cla.claset
+    val equality_name = @{const_name "op ="}
+    val not_name = @{const_name Not}
+    val notE	= @{thm notE}
+    val ccontr	= @{thm ccontr}
+    val contr_tac = Cla.contr_tac
+    val dup_intr	= Cla.dup_intr
+    val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
+    val rep_cs = Cla.rep_cs
+    val cla_modifiers = Cla.cla_modifiers
+    val cla_meth' = Cla.cla_meth'
+  );
+  val blast_tac = Blast.blast_tac;
+*}
+
 setup Blast.setup
 
 
@@ -360,7 +377,7 @@
 text {* Method setup. *}
 
 ML {*
-  structure Induct = InductFun
+  structure Induct = Induct
   (
     val cases_default = @{thm case_split}
     val atomize = @{thms induct_atomize}
--- a/src/FOL/IFOL.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/FOL/IFOL.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -591,12 +591,12 @@
   done
 
 ML {*
-structure ProjectRule = ProjectRuleFun
-(struct
+structure Project_Rule = Project_Rule
+(
   val conjunct1 = @{thm conjunct1}
   val conjunct2 = @{thm conjunct2}
   val mp = @{thm mp}
-end)
+)
 *}
 
 use "fologic.ML"
--- a/src/FOL/IsaMakefile	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/FOL/IsaMakefile	Wed Jul 29 09:06:49 2009 +0200
@@ -36,8 +36,8 @@
   $(SRC)/Provers/hypsubst.ML $(SRC)/Tools/induct.ML			\
   $(SRC)/Tools/intuitionistic.ML $(SRC)/Tools/atomize_elim.ML		\
   $(SRC)/Tools/project_rule.ML $(SRC)/Provers/quantifier1.ML		\
-  $(SRC)/Provers/splitter.ML FOL.thy IFOL.thy ROOT.ML blastdata.ML	\
-  cladata.ML document/root.tex fologic.ML hypsubstdata.ML intprover.ML	\
+  $(SRC)/Provers/splitter.ML FOL.thy IFOL.thy ROOT.ML cladata.ML	\
+  document/root.tex fologic.ML hypsubstdata.ML intprover.ML		\
   simpdata.ML
 	@$(ISABELLE_TOOL) usedir -p 2 -b $(OUT)/Pure FOL
 
--- a/src/FOL/blastdata.ML	Fri Jul 24 13:56:02 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-
-(*** Applying BlastFun to create blast_tac ***)
-structure Blast_Data = 
-  struct
-  type claset	= Cla.claset
-  val equality_name = @{const_name "op ="}
-  val not_name = @{const_name Not}
-  val notE	= @{thm notE}
-  val ccontr	= @{thm ccontr}
-  val contr_tac = Cla.contr_tac
-  val dup_intr	= Cla.dup_intr
-  val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
-  val rep_cs = Cla.rep_cs
-  val cla_modifiers = Cla.cla_modifiers;
-  val cla_meth' = Cla.cla_meth'
-  end;
-
-structure Blast = BlastFun(Blast_Data);
-val blast_tac = Blast.blast_tac;
--- a/src/FOL/simpdata.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/FOL/simpdata.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -95,27 +95,25 @@
 
 (*** Case splitting ***)
 
-structure SplitterData =
-  struct
-  structure Simplifier = Simplifier
-  val mk_eq          = mk_eq
+structure Splitter = Splitter
+(
+  val thy = @{theory}
+  val mk_eq = mk_eq
   val meta_eq_to_iff = @{thm meta_eq_to_iff}
-  val iffD           = @{thm iffD2}
-  val disjE          = @{thm disjE}
-  val conjE          = @{thm conjE}
-  val exE            = @{thm exE}
-  val contrapos      = @{thm contrapos}
-  val contrapos2     = @{thm contrapos2}
-  val notnotD        = @{thm notnotD}
-  end;
+  val iffD = @{thm iffD2}
+  val disjE = @{thm disjE}
+  val conjE = @{thm conjE}
+  val exE = @{thm exE}
+  val contrapos = @{thm contrapos}
+  val contrapos2 = @{thm contrapos2}
+  val notnotD = @{thm notnotD}
+);
 
-structure Splitter = SplitterFun(SplitterData);
-
-val split_tac        = Splitter.split_tac;
+val split_tac = Splitter.split_tac;
 val split_inside_tac = Splitter.split_inside_tac;
-val split_asm_tac    = Splitter.split_asm_tac;
-val op addsplits     = Splitter.addsplits;
-val op delsplits     = Splitter.delsplits;
+val split_asm_tac = Splitter.split_asm_tac;
+val op addsplits = Splitter.addsplits;
+val op delsplits = Splitter.delsplits;
 
 
 (*** Standard simpsets ***)
--- a/src/HOL/Decision_Procs/Approximation.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -3232,7 +3232,7 @@
 code_const approx' (Eval "Float'_Arith.approx'")
 
 ML {*
-  fun reorder_bounds_tac i prems =
+  fun reorder_bounds_tac prems i =
     let
       fun variable_of_bound (Const ("Trueprop", _) $
                              (Const (@{const_name "op :"}, _) $
@@ -3337,7 +3337,7 @@
       REPEAT (FIRST' [etac @{thm intervalE},
                       etac @{thm meta_eqE},
                       rtac @{thm impI}] i)
-      THEN METAHYPS (reorder_bounds_tac i) i
+      THEN FOCUS (fn {prems, ...} => reorder_bounds_tac prems i) @{context} i
       THEN TRY (filter_prems_tac (K false) i)
       THEN DETERM (Reflection.genreify_tac ctxt form_equations NONE i)
       THEN rewrite_interpret_form_tac ctxt prec splitting taylor i
--- a/src/HOL/Finite_Set.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Finite_Set.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -2960,13 +2960,9 @@
   "ab_semigroup_idem_mult max"
   proof qed (auto simp add: max_def)
 
-lemma min_lattice:
-  "lower_semilattice (op \<le>) (op <) min"
-  proof qed (auto simp add: min_def)
-
 lemma max_lattice:
   "lower_semilattice (op \<ge>) (op >) max"
-  proof qed (auto simp add: max_def)
+  by (fact min_max.dual_semilattice)
 
 lemma dual_max:
   "ord.max (op \<ge>) = min"
@@ -3126,11 +3122,7 @@
 lemma Min_le [simp]:
   assumes "finite A" and "x \<in> A"
   shows "Min A \<le> x"
-proof -
-  interpret lower_semilattice "op \<le>" "op <" min
-    by (rule min_lattice)
-  from assms show ?thesis by (simp add: Min_def fold1_belowI)
-qed
+  using assms by (simp add: Min_def min_max.fold1_belowI)
 
 lemma Max_ge [simp]:
   assumes "finite A" and "x \<in> A"
@@ -3144,11 +3136,7 @@
 lemma Min_ge_iff [simp, noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "x \<le> Min A \<longleftrightarrow> (\<forall>a\<in>A. x \<le> a)"
-proof -
-  interpret lower_semilattice "op \<le>" "op <" min
-    by (rule min_lattice)
-  from assms show ?thesis by (simp add: Min_def below_fold1_iff)
-qed
+  using assms by (simp add: Min_def min_max.below_fold1_iff)
 
 lemma Max_le_iff [simp, noatp]:
   assumes "finite A" and "A \<noteq> {}"
@@ -3162,63 +3150,46 @@
 lemma Min_gr_iff [simp, noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "x < Min A \<longleftrightarrow> (\<forall>a\<in>A. x < a)"
-proof -
-  interpret lower_semilattice "op \<le>" "op <" min
-    by (rule min_lattice)
-  from assms show ?thesis by (simp add: Min_def strict_below_fold1_iff)
-qed
+  using assms by (simp add: Min_def strict_below_fold1_iff)
 
 lemma Max_less_iff [simp, noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "Max A < x \<longleftrightarrow> (\<forall>a\<in>A. a < x)"
 proof -
-  note Max = Max_def
-  interpret linorder "op \<ge>" "op >"
+  interpret dual: linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
-    by (simp add: Max strict_below_fold1_iff [folded dual_max])
+    by (simp add: Max_def dual.strict_below_fold1_iff [folded dual.dual_max])
 qed
 
 lemma Min_le_iff [noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A \<le> x \<longleftrightarrow> (\<exists>a\<in>A. a \<le> x)"
-proof -
-  interpret lower_semilattice "op \<le>" "op <" min
-    by (rule min_lattice)
-  from assms show ?thesis
-    by (simp add: Min_def fold1_below_iff)
-qed
+  using assms by (simp add: Min_def fold1_below_iff)
 
 lemma Max_ge_iff [noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "x \<le> Max A \<longleftrightarrow> (\<exists>a\<in>A. x \<le> a)"
 proof -
-  note Max = Max_def
-  interpret linorder "op \<ge>" "op >"
+  interpret dual: linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
-    by (simp add: Max fold1_below_iff [folded dual_max])
+    by (simp add: Max_def dual.fold1_below_iff [folded dual.dual_max])
 qed
 
 lemma Min_less_iff [noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "Min A < x \<longleftrightarrow> (\<exists>a\<in>A. a < x)"
-proof -
-  interpret lower_semilattice "op \<le>" "op <" min
-    by (rule min_lattice)
-  from assms show ?thesis
-    by (simp add: Min_def fold1_strict_below_iff)
-qed
+  using assms by (simp add: Min_def fold1_strict_below_iff)
 
 lemma Max_gr_iff [noatp]:
   assumes "finite A" and "A \<noteq> {}"
   shows "x < Max A \<longleftrightarrow> (\<exists>a\<in>A. x < a)"
 proof -
-  note Max = Max_def
-  interpret linorder "op \<ge>" "op >"
+  interpret dual: linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
-    by (simp add: Max fold1_strict_below_iff [folded dual_max])
+    by (simp add: Max_def dual.fold1_strict_below_iff [folded dual.dual_max])
 qed
 
 lemma Min_eqI:
@@ -3248,21 +3219,16 @@
 lemma Min_antimono:
   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
   shows "Min N \<le> Min M"
-proof -
-  interpret distrib_lattice "op \<le>" "op <" min max
-    by (rule distrib_lattice_min_max)
-  from assms show ?thesis by (simp add: Min_def fold1_antimono)
-qed
+  using assms by (simp add: Min_def fold1_antimono)
 
 lemma Max_mono:
   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
   shows "Max M \<le> Max N"
 proof -
-  note Max = Max_def
-  interpret linorder "op \<ge>" "op >"
+  interpret dual: linorder "op \<ge>" "op >"
     by (rule dual_linorder)
   from assms show ?thesis
-    by (simp add: Max fold1_antimono [folded dual_max])
+    by (simp add: Max_def dual.fold1_antimono [folded dual.dual_max])
 qed
 
 lemma finite_linorder_max_induct[consumes 1, case_names empty insert]:
--- a/src/HOL/FunDef.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/FunDef.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -193,9 +193,9 @@
 subsection {* Reduction Pairs *}
 
 definition
-  "reduction_pair P = (wf (fst P) \<and> snd P O fst P \<subseteq> fst P)"
+  "reduction_pair P = (wf (fst P) \<and> fst P O snd P \<subseteq> fst P)"
 
-lemma reduction_pairI[intro]: "wf R \<Longrightarrow> S O R \<subseteq> R \<Longrightarrow> reduction_pair (R, S)"
+lemma reduction_pairI[intro]: "wf R \<Longrightarrow> R O S \<subseteq> R \<Longrightarrow> reduction_pair (R, S)"
 unfolding reduction_pair_def by auto
 
 lemma reduction_pair_lemma:
@@ -205,7 +205,7 @@
   assumes "wf S"
   shows "wf (R \<union> S)"
 proof -
-  from rp `S \<subseteq> snd P` have "wf (fst P)" "S O fst P \<subseteq> fst P"
+  from rp `S \<subseteq> snd P` have "wf (fst P)" "fst P O S \<subseteq> fst P"
     unfolding reduction_pair_def by auto
   with `wf S` have "wf (fst P \<union> S)" 
     by (auto intro: wf_union_compatible)
@@ -266,8 +266,8 @@
 text {* Reduction Pairs *}
 
 lemma max_ext_compat: 
-  assumes "S O R \<subseteq> R"
-  shows "(max_ext S \<union> {({},{})}) O max_ext R \<subseteq> max_ext R"
+  assumes "R O S \<subseteq> R"
+  shows "max_ext R O (max_ext S \<union> {({},{})}) \<subseteq> max_ext R"
 using assms 
 apply auto
 apply (elim max_ext.cases)
@@ -287,8 +287,8 @@
 by (auto simp: pair_less_def pair_leq_def)
 
 lemma min_ext_compat: 
-  assumes "S O R \<subseteq> R"
-  shows "(min_ext S \<union> {({},{})}) O min_ext R \<subseteq> min_ext R"
+  assumes "R O S \<subseteq> R"
+  shows "min_ext R O  (min_ext S \<union> {({},{})}) \<subseteq> min_ext R"
 using assms 
 apply (auto simp: min_ext_def)
 apply (drule_tac x=ya in bspec, assumption)
--- a/src/HOL/HOL.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/HOL.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -910,8 +910,9 @@
 done
 
 ML {*
-structure Blast = BlastFun
+structure Blast = Blast
 (
+  val thy = @{theory}
   type claset = Classical.claset
   val equality_name = @{const_name "op ="}
   val not_name = @{const_name Not}
@@ -1390,7 +1391,7 @@
 text {* Rule projections: *}
 
 ML {*
-structure ProjectRule = ProjectRuleFun
+structure Project_Rule = Project_Rule
 (
   val conjunct1 = @{thm conjunct1}
   val conjunct2 = @{thm conjunct2}
@@ -1446,7 +1447,7 @@
 text {* Method setup. *}
 
 ML {*
-structure Induct = InductFun
+structure Induct = Induct
 (
   val cases_default = @{thm case_split}
   val atomize = @{thms induct_atomize}
--- a/src/HOL/IMP/Denotation.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/IMP/Denotation.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -12,14 +12,14 @@
 
 definition
   Gamma :: "[bexp,com_den] => (com_den => com_den)" where
-  "Gamma b cd = (\<lambda>phi. {(s,t). (s,t) \<in> (phi O cd) \<and> b s} \<union>
+  "Gamma b cd = (\<lambda>phi. {(s,t). (s,t) \<in> (cd O phi) \<and> b s} \<union>
                        {(s,t). s=t \<and> \<not>b s})"
 
 primrec C :: "com => com_den"
 where
   C_skip:   "C \<SKIP>   = Id"
 | C_assign: "C (x :== a) = {(s,t). t = s[x\<mapsto>a(s)]}"
-| C_comp:   "C (c0;c1)   = C(c1) O C(c0)"
+| C_comp:   "C (c0;c1)   = C(c0) O C(c1)"
 | C_if:     "C (\<IF> b \<THEN> c1 \<ELSE> c2) = {(s,t). (s,t) \<in> C c1 \<and> b s} \<union>
                                                 {(s,t). (s,t) \<in> C c2 \<and> \<not>b s}"
 | C_while:  "C(\<WHILE> b \<DO> c) = lfp (Gamma b (C c))"
--- a/src/HOL/IMP/Machines.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/IMP/Machines.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -6,7 +6,7 @@
   "((x,z) \<in> R ^^ n) = (n=0 \<and> z=x \<or> (\<exists>m y. n = Suc m \<and> (x,y) \<in> R \<and> (y,z) \<in> R ^^ m))"
 apply(rule iffI)
  apply(blast elim:rel_pow_E2)
-apply (fastsimp simp add:gr0_conv_Suc rel_pow_commute)
+apply (auto simp: rel_pow_commute[symmetric])
 done
 
 subsection "Instructions"
--- a/src/HOL/Import/proof_kernel.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Import/proof_kernel.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -245,7 +245,8 @@
 
 fun prth th = writeln (PrintMode.setmp [] Display.string_of_thm_without_context th)
 fun prc ct = writeln (PrintMode.setmp [] Display.string_of_cterm ct)
-fun prin t = writeln (PrintMode.setmp [] (fn () => Syntax.string_of_term_global (the_context ()) t) ());
+fun prin t = writeln (PrintMode.setmp []
+  (fn () => Syntax.string_of_term (ML_Context.the_local_context ()) t) ());
 fun pth (HOLThm(ren,thm)) =
     let
         (*val _ = writeln "Renaming:"
--- a/src/HOL/Lambda/Eta.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Lambda/Eta.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -371,7 +371,7 @@
 
 theorem eta_postponement:
   assumes "(sup beta eta)\<^sup>*\<^sup>* s t"
-  shows "(eta\<^sup>*\<^sup>* OO beta\<^sup>*\<^sup>*) s t" using assms
+  shows "(beta\<^sup>*\<^sup>* OO eta\<^sup>*\<^sup>*) s t" using assms
 proof induct
   case base
   show ?case by blast
--- a/src/HOL/Lattices.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Lattices.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -413,20 +413,14 @@
 subsection {* @{const min}/@{const max} on linear orders as
   special case of @{const inf}/@{const sup} *}
 
-lemma (in linorder) distrib_lattice_min_max:
-  "distrib_lattice (op \<le>) (op <) min max"
+sublocale linorder < min_max!: distrib_lattice less_eq less "Orderings.ord.min less_eq" "Orderings.ord.max less_eq"
 proof
-  have aux: "\<And>x y \<Colon> 'a. x < y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
-    by (auto simp add: less_le antisym)
   fix x y z
-  show "max x (min y z) = min (max x y) (max x z)"
-  unfolding min_def max_def
-  by auto
+  show "Orderings.ord.max less_eq x (Orderings.ord.min less_eq y z) =
+    Orderings.ord.min less_eq (Orderings.ord.max less_eq x y) (Orderings.ord.max less_eq x z)"
+  unfolding min_def max_def by auto
 qed (auto simp add: min_def max_def not_le less_imp_le)
 
-interpretation min_max: distrib_lattice "op \<le> :: 'a::linorder \<Rightarrow> 'a \<Rightarrow> bool" "op <" min max
-  by (rule distrib_lattice_min_max)
-
 lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{lower_semilattice, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
   by (rule ext)+ (auto intro: antisym)
 
--- a/src/HOL/Library/Kleene_Algebra.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Library/Kleene_Algebra.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -274,7 +274,19 @@
 oops
 
 lemma star_decomp: "star (a + b) = star a * star (b * star a)"
-oops
+proof (rule antisym)
+  have "1 + (a + b) * star a * star (b * star a) \<le>
+    1 + a * star a * star (b * star a) + b * star a * star (b * star a)"
+    by (metis add_commute add_left_commute eq_iff left_distrib mult_assoc)
+  also have "\<dots> \<le> star a * star (b * star a)"
+    by (metis add_commute add_est1 add_left_commute ka18 plus_leI star_unfold_left x_less_star)
+  finally show "star (a + b) \<le> star a * star (b * star a)"
+    by (metis mult_1_right mult_assoc star3')
+next
+  show "star a * star (b * star a) \<le> star (a + b)"
+    by (metis add_assoc add_est1 add_est2 add_left_commute less_star mult_mono'
+      star_absorb_one star_absorb_one' star_idemp star_mono star_mult_idem zero_minimum)
+qed
 
 lemma ka22: "y * star x \<le> star x * star y \<Longrightarrow>  star y * star x \<le> star x * star y"
 by (metis mult_assoc mult_right_mono plus_leI star3' star_mult_idem x_less_star zero_minimum)
--- a/src/HOL/Library/RBT.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Library/RBT.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -13,7 +13,7 @@
 datatype color = R | B
 datatype ('a,'b)"rbt" = Empty | Tr color "('a,'b)rbt" 'a 'b "('a,'b)rbt"
 
-(* Suchbaum-Eigenschaften *)
+text {* Search tree properties *}
 
 primrec
   pin_tree :: "'a \<Rightarrow> 'b \<Rightarrow> ('a,'b) rbt \<Rightarrow> bool"
@@ -340,9 +340,6 @@
 
 subsection {* Deletion *}
 
-(*definition
-  [simp]: "ibn t = (bh t > 0 \<and> treec t = B)"
-*)
 lemma bh_paintR'[simp]: "treec t = B \<Longrightarrow> bh (paint R t) = bh t - 1"
 by (cases t rule: rbt_cases) auto
 
@@ -916,9 +913,72 @@
   "alist_of Empty = []"
 | "alist_of (Tr _ l k v r) = alist_of l @ (k,v) # alist_of r"
 
+lemma map_of_alist_of_aux: "st (Tr c t1 k v t2) \<Longrightarrow> RBT.map_of (Tr c t1 k v t2) = RBT.map_of t2 ++ [k\<mapsto>v] ++ RBT.map_of t1"
+proof (rule ext)
+  fix x
+  assume ST: "st (Tr c t1 k v t2)"
+  let ?thesis = "RBT.map_of (Tr c t1 k v t2) x = (RBT.map_of t2 ++ [k \<mapsto> v] ++ RBT.map_of t1) x"
+
+  have DOM_T1: "!!k'. k'\<in>dom (RBT.map_of t1) \<Longrightarrow> k>k'"
+  proof -
+    fix k'
+    from ST have "t1 |\<guillemotleft> k" by simp
+    with tlt_prop have "\<forall>k'\<in>keys t1. k>k'" by auto
+    moreover assume "k'\<in>dom (RBT.map_of t1)"
+    ultimately show "k>k'" using RBT.mapof_keys ST by auto
+  qed
+
+  have DOM_T2: "!!k'. k'\<in>dom (RBT.map_of t2) \<Longrightarrow> k<k'"
+  proof -
+    fix k'
+    from ST have "k \<guillemotleft>| t2" by simp
+    with tgt_prop have "\<forall>k'\<in>keys t2. k<k'" by auto
+    moreover assume "k'\<in>dom (RBT.map_of t2)"
+    ultimately show "k<k'" using RBT.mapof_keys ST by auto
+  qed
+
+  {
+    assume C: "x<k"
+    hence "RBT.map_of (Tr c t1 k v t2) x = RBT.map_of t1 x" by simp
+    moreover from C have "x\<notin>dom [k\<mapsto>v]" by simp
+    moreover have "x\<notin>dom (RBT.map_of t2)" proof
+      assume "x\<in>dom (RBT.map_of t2)"
+      with DOM_T2 have "k<x" by blast
+      with C show False by simp
+    qed
+    ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
+  } moreover {
+    assume [simp]: "x=k"
+    hence "RBT.map_of (Tr c t1 k v t2) x = [k \<mapsto> v] x" by simp
+    moreover have "x\<notin>dom (RBT.map_of t1)" proof
+      assume "x\<in>dom (RBT.map_of t1)"
+      with DOM_T1 have "k>x" by blast
+      thus False by simp
+    qed
+    ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
+  } moreover {
+    assume C: "x>k"
+    hence "RBT.map_of (Tr c t1 k v t2) x = RBT.map_of t2 x" by (simp add: less_not_sym[of k x])
+    moreover from C have "x\<notin>dom [k\<mapsto>v]" by simp
+    moreover have "x\<notin>dom (RBT.map_of t1)" proof
+      assume "x\<in>dom (RBT.map_of t1)"
+      with DOM_T1 have "k>x" by simp
+      with C show False by simp
+    qed
+    ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps)
+  } ultimately show ?thesis using less_linear by blast
+qed
+
 lemma map_of_alist_of:
   shows "st t \<Longrightarrow> Map.map_of (alist_of t) = map_of t"
-  oops
+proof (induct t)
+  case Empty thus ?case by (simp add: RBT.map_of_Empty)
+next
+  case (Tr c t1 k v t2)
+  hence "Map.map_of (alist_of (Tr c t1 k v t2)) = RBT.map_of t2 ++ [k \<mapsto> v] ++ RBT.map_of t1" by simp
+  also note map_of_alist_of_aux[OF Tr.prems,symmetric]
+  finally show ?case .
+qed
 
 lemma fold_alist_fold:
   "foldwithkey f t x = foldl (\<lambda>x (k,v). f k v x) x (alist_of t)"
--- a/src/HOL/Map.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Map.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -307,6 +307,9 @@
 lemma map_add_upds[simp]: "m1 ++ (m2(xs[\<mapsto>]ys)) = (m1++m2)(xs[\<mapsto>]ys)"
 by (simp add: map_upds_def)
 
+lemma map_add_upd_left: "m\<notin>dom e2 \<Longrightarrow> e1(m \<mapsto> u1) ++ e2 = (e1 ++ e2)(m \<mapsto> u1)"
+by (rule ext) (auto simp: map_add_def dom_def split: option.split)
+
 lemma map_of_append[simp]: "map_of (xs @ ys) = map_of ys ++ map_of xs"
 unfolding map_add_def
 apply (induct xs)
@@ -508,6 +511,12 @@
 lemma map_add_comm: "dom m1 \<inter> dom m2 = {} \<Longrightarrow> m1++m2 = m2++m1"
 by (rule ext) (force simp: map_add_def dom_def split: option.split)
 
+lemma map_add_dom_app_simps:
+  "\<lbrakk> m\<in>dom l2 \<rbrakk> \<Longrightarrow> (l1++l2) m = l2 m"
+  "\<lbrakk> m\<notin>dom l1 \<rbrakk> \<Longrightarrow> (l1++l2) m = l2 m"
+  "\<lbrakk> m\<notin>dom l2 \<rbrakk> \<Longrightarrow> (l1++l2) m = l1 m"
+by (auto simp add: map_add_def split: option.split_asm)
+
 lemma dom_const [simp]:
   "dom (\<lambda>x. Some y) = UNIV"
   by auto
--- a/src/HOL/MicroJava/BV/Listn.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/MicroJava/BV/Listn.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/Listn.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   2000 TUM
 
@@ -8,7 +7,9 @@
 
 header {* \isaheader{Fixed Length Lists} *}
 
-theory Listn imports Err begin
+theory Listn
+imports Err
+begin
 
 constdefs
 
@@ -317,6 +318,10 @@
 apply (simp add: nth_Cons split: nat.split)
 done
 
+lemma equals0I_aux:
+  "(\<And>y. A y \<Longrightarrow> False) \<Longrightarrow> A = bot_class.bot"
+  by (rule equals0I) (auto simp add: mem_def)
+
 lemma acc_le_listI [intro!]:
   "\<lbrakk> order r; acc r \<rbrakk> \<Longrightarrow> acc(Listn.le r)"
 apply (unfold acc_def)
@@ -330,7 +335,7 @@
  apply (rename_tac m n)
  apply (case_tac "m=n")
   apply simp
- apply (fast intro!: equals0I [to_pred bot_empty_eq pred_equals_eq] dest: not_sym)
+ apply (fast intro!: equals0I_aux dest: not_sym)
 apply clarify
 apply (rename_tac n)
 apply (induct_tac n)
--- a/src/HOL/Modelcheck/MuckeSyn.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Modelcheck/MuckeSyn.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -118,7 +118,7 @@
 
 local
 
-  val move_thm = prove_goal @{theory} "[| a = b ==> P a; a = b |] ==> P b"
+  val move_thm = OldGoals.prove_goal @{theory} "[| a = b ==> P a; a = b |] ==> P b"
 	(fn prems => [cut_facts_tac prems 1, dtac sym 1, hyp_subst_tac 1,
 		     REPEAT (resolve_tac prems 1)]);
 
@@ -159,9 +159,9 @@
 
 (* transforming fun-defs into lambda-defs *)
 
-val [eq] = goal Pure.thy "(!! x. f x == g x) ==> f == g";
- by (rtac (extensional eq) 1);
-qed "ext_rl";
+val [eq] = OldGoals.goal Pure.thy "(!! x. f x == g x) ==> f == g";
+ OldGoals.by (rtac (extensional eq) 1);
+OldGoals.qed "ext_rl";
 
 infix cc;
 
@@ -196,7 +196,7 @@
 	val rhs = Syntax.string_of_term_global thy (freeze_thaw RHS)
 	val gl = delete_bold_string (fnam ^" == % " ^ (getargs LHS) ^" . " ^ rhs);
     in
-	SOME (prove_goal thy gl (fn prems =>
+	SOME (OldGoals.prove_goal thy gl (fn prems =>
   		[(REPEAT (rtac ext_rl 1)), (rtac t 1) ]))
     end
 | mk_lam_def [] _ t= NONE; 
--- a/src/HOL/Modelcheck/mucke_oracle.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Modelcheck/mucke_oracle.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -109,9 +109,9 @@
 (
 OldGoals.push_proof();
 OldGoals.goalw_cterm [] (cterm_of sign trm);
-by (simp_tac (global_simpset_of sign) 1);
+OldGoals.by (simp_tac (global_simpset_of sign) 1);
         let
-        val if_tmp_result = result()
+        val if_tmp_result = OldGoals.result()
         in
         (
         OldGoals.pop_proof();
--- a/src/HOL/NatTransfer.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/NatTransfer.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -382,7 +382,7 @@
 
 lemma UNIV_apply:
   "UNIV x = True"
-  by (simp add: top_set_eq [symmetric] top_fun_eq top_bool_eq)
+  by (simp add: top_fun_eq top_bool_eq)
 
 declare TransferMorphism_int_nat[transfer add return:
   transfer_int_nat_numerals
--- a/src/HOL/Nominal/nominal_datatype.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Nominal/nominal_datatype.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -152,7 +152,7 @@
 val meta_spec = thm "meta_spec";
 
 fun projections rule =
-  ProjectRule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
+  Project_Rule.projections (ProofContext.init (Thm.theory_of_thm rule)) rule
   |> map (standard #> RuleCases.save rule);
 
 val supp_prod = thm "supp_prod";
@@ -1257,7 +1257,7 @@
             [resolve_tac exists_fresh' 1,
              simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms @
                fin_set_supp @ ths)) 1]);
-        val (([cx], ths), ctxt') = Obtain.result
+        val (([(_, cx)], ths), ctxt') = Obtain.result
           (fn _ => EVERY
             [etac exE 1,
              full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
@@ -1324,7 +1324,7 @@
                        val _ $ (_ $ _ $ u) = concl';
                        val U = fastype_of u;
                        val (xs, params') =
-                         chop (length cargs) (map term_of params);
+                         chop (length cargs) (map (term_of o #2) params);
                        val Ts = map fastype_of xs;
                        val cnstr = Const (cname, Ts ---> U);
                        val (pis, z) = split_last params';
@@ -1640,7 +1640,7 @@
                     REPEAT_DETERM (resolve_tac [allI, impI] 1),
                     REPEAT_DETERM (etac conjE 1),
                     rtac unique 1,
-                    SUBPROOF (fn {prems = prems', params = [a, b], ...} => EVERY
+                    SUBPROOF (fn {prems = prems', params = [(_, a), (_, b)], ...} => EVERY
                       [cut_facts_tac [rec_prem] 1,
                        rtac (Thm.instantiate ([],
                          [(cterm_of thy11 (Var (("pi", 0), mk_permT aT)),
@@ -1693,7 +1693,7 @@
              REPEAT_DETERM (dresolve_tac (the (AList.lookup op = rec_fin_supp T)) 1),
              resolve_tac exists_fresh' 1,
              asm_simp_tac (HOL_ss addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]);
-        val (([cx], ths), ctxt') = Obtain.result
+        val (([(_, cx)], ths), ctxt') = Obtain.result
           (fn _ => EVERY
             [etac exE 1,
              full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
@@ -1763,7 +1763,7 @@
                     val cargsr' = partition_cargs idxs cargsr;
                     val boundsr = List.concat (map fst cargsr');
                     val (params1, _ :: params2) =
-                      chop (length params div 2) (map term_of params);
+                      chop (length params div 2) (map (term_of o #2) params);
                     val params' = params1 @ params2;
                     val rec_prems = filter (fn th => case prop_of th of
                         _ $ p => (case head_of p of
--- a/src/HOL/Nominal/nominal_inductive.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -301,7 +301,7 @@
           (fn _ => EVERY
             [resolve_tac exists_fresh' 1,
              resolve_tac fs_atoms 1]);
-        val (([cx], ths), ctxt') = Obtain.result
+        val (([(_, cx)], ths), ctxt') = Obtain.result
           (fn _ => EVERY
             [etac exE 1,
              full_simp_tac (HOL_ss addsimps (fresh_prod :: fresh_atm)) 1,
@@ -319,7 +319,7 @@
              SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
                let
                  val (params', (pis, z)) =
-                   chop (length params - length atomTs - 1) (map term_of params) ||>
+                   chop (length params - length atomTs - 1) (map (term_of o #2) params) ||>
                    split_last;
                  val bvars' = map
                    (fn (Bound i, T) => (nth params' (length params' - i), T)
@@ -474,7 +474,7 @@
                 rtac (first_order_mrs case_hyps case_hyp) 1
               else
                 let
-                  val params' = map (term_of o nth (rev params)) is;
+                  val params' = map (term_of o #2 o nth (rev params)) is;
                   val tab = params' ~~ map fst qs;
                   val (hyps1, hyps2) = chop (length args) case_hyps;
                   (* turns a = t and [x1 # t, ..., xn # t] *)
@@ -566,7 +566,7 @@
             map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct])
           ctxt;
         val strong_inducts =
-          ProjectRule.projects ctxt (1 upto length names) strong_induct'
+          Project_Rule.projects ctxt (1 upto length names) strong_induct'
       in
         ctxt' |>
         LocalTheory.note Thm.generatedK
@@ -635,7 +635,7 @@
             val prems'' = map (fn th => Simplifier.simplify eqvt_ss
               (mk_perm_bool (cterm_of thy pi) th)) prems';
             val intr' = Drule.cterm_instantiate (map (cterm_of thy) vs ~~
-               map (cterm_of thy o NominalDatatype.mk_perm [] pi o term_of) params)
+               map (cterm_of thy o NominalDatatype.mk_perm [] pi o term_of o #2) params)
                intr
           in (rtac intr' THEN_ALL_NEW (TRY o resolve_tac prems'')) 1
           end) ctxt' 1 st
--- a/src/HOL/Nominal/nominal_inductive2.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Nominal/nominal_inductive2.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -323,7 +323,7 @@
         val avoid_th = Drule.instantiate'
           [SOME (ctyp_of thy (fastype_of p))] [SOME (cterm_of thy p)]
           ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding});
-        val (([cx], th1 :: th2 :: ths), ctxt') = Obtain.result
+        val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result
           (fn _ => EVERY
             [rtac avoid_th 1,
              full_simp_tac (HOL_ss addsimps [@{thm fresh_star_prod_set}]) 1,
@@ -359,7 +359,7 @@
              SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} =>
                let
                  val (cparams', (pis, z)) =
-                   chop (length params - length atomTs - 1) params ||>
+                   chop (length params - length atomTs - 1) (map #2 params) ||>
                    (map term_of #> split_last);
                  val params' = map term_of cparams'
                  val sets' = map (apfst (curry subst_bounds (rev params'))) sets;
@@ -466,7 +466,7 @@
             map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct])
           ctxt;
         val strong_inducts =
-          ProjectRule.projects ctxt' (1 upto length names) strong_induct'
+          Project_Rule.projects ctxt' (1 upto length names) strong_induct'
       in
         ctxt' |>
         LocalTheory.note Thm.generatedK
--- a/src/HOL/Nominal/nominal_primrec.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Nominal/nominal_primrec.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -380,7 +380,7 @@
       [goals] |>
     Proof.apply (Method.Basic (fn _ => RAW_METHOD (fn _ =>
       rewrite_goals_tac defs_thms THEN
-      compose_tac (false, rule, length rule_prems) 1), Position.none)) |>
+      compose_tac (false, rule, length rule_prems) 1))) |>
     Seq.hd
   end;
 
--- a/src/HOL/Orderings.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Orderings.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -6,7 +6,9 @@
 
 theory Orderings
 imports HOL
-uses "~~/src/Provers/order.ML"
+uses
+  "~~/src/Provers/order.ML"
+  "~~/src/Provers/quasi.ML"  (* FIXME unused? *)
 begin
 
 subsection {* Quasi orders *}
@@ -289,7 +291,7 @@
 sig
   val print_structures: Proof.context -> unit
   val setup: theory -> theory
-  val order_tac: thm list -> Proof.context -> int -> tactic
+  val order_tac: Proof.context -> thm list -> int -> tactic
 end;
 
 structure Orders: ORDERS =
@@ -329,7 +331,7 @@
 
 (** Method **)
 
-fun struct_tac ((s, [eq, le, less]), thms) prems =
+fun struct_tac ((s, [eq, le, less]), thms) ctxt prems =
   let
     fun decomp thy (@{const Trueprop} $ t) =
       let
@@ -354,13 +356,13 @@
       | decomp thy _ = NONE;
   in
     case s of
-      "order" => Order_Tac.partial_tac decomp thms prems
-    | "linorder" => Order_Tac.linear_tac decomp thms prems
+      "order" => Order_Tac.partial_tac decomp thms ctxt prems
+    | "linorder" => Order_Tac.linear_tac decomp thms ctxt prems
     | _ => error ("Unknown kind of order `" ^ s ^ "' encountered in transitivity reasoner.")
   end
 
-fun order_tac prems ctxt =
-  FIRST' (map (fn s => CHANGED o struct_tac s prems) (Data.get (Context.Proof ctxt)));
+fun order_tac ctxt prems =
+  FIRST' (map (fn s => CHANGED o struct_tac s ctxt prems) (Data.get (Context.Proof ctxt)));
 
 
 (** Attribute **)
@@ -394,7 +396,7 @@
 (** Setup **)
 
 val setup =
-  Method.setup @{binding order} (Scan.succeed (SIMPLE_METHOD' o order_tac []))
+  Method.setup @{binding order} (Scan.succeed (fn ctxt => SIMPLE_METHOD' (order_tac ctxt [])))
     "transitivity reasoner" #>
   attrib_setup;
 
@@ -532,7 +534,7 @@
       Simplifier.simproc thy name raw_ts proc) procs)) thy;
 fun add_solver name tac =
   Simplifier.map_simpset (fn ss => ss addSolver
-    mk_solver' name (fn ss => tac (Simplifier.prems_of_ss ss) (Simplifier.the_context ss)));
+    mk_solver' name (fn ss => tac (Simplifier.the_context ss) (Simplifier.prems_of_ss ss)));
 
 in
   add_simprocs [
--- a/src/HOL/Predicate.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Predicate.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -221,11 +221,11 @@
 subsubsection {* Composition  *}
 
 inductive
-  pred_comp  :: "['b => 'c => bool, 'a => 'b => bool] => 'a => 'c => bool"
+  pred_comp  :: "['a => 'b => bool, 'b => 'c => bool] => 'a => 'c => bool"
     (infixr "OO" 75)
-  for r :: "'b => 'c => bool" and s :: "'a => 'b => bool"
+  for r :: "'a => 'b => bool" and s :: "'b => 'c => bool"
 where
-  pred_compI [intro]: "s a b ==> r b c ==> (r OO s) a c"
+  pred_compI [intro]: "r a b ==> s b c ==> (r OO s) a c"
 
 inductive_cases pred_compE [elim!]: "(r OO s) a c"
 
--- a/src/HOL/Recdef.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Recdef.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -79,6 +79,32 @@
 use "Tools/recdef.ML"
 setup Recdef.setup
 
+text {*Wellfoundedness of @{text same_fst}*}
+
+definition
+ same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
+where
+    "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
+   --{*For @{text rec_def} declarations where the first n parameters
+       stay unchanged in the recursive call. *}
+
+lemma same_fstI [intro!]:
+     "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R"
+by (simp add: same_fst_def)
+
+lemma wf_same_fst:
+  assumes prem: "(!!x. P x ==> wf(R x))"
+  shows "wf(same_fst P R)"
+apply (simp cong del: imp_cong add: wf_def same_fst_def)
+apply (intro strip)
+apply (rename_tac a b)
+apply (case_tac "wf (R a)")
+ apply (erule_tac a = b in wf_induct, blast)
+apply (blast intro: prem)
+done
+
+text {*Rule setup*}
+
 lemmas [recdef_simp] =
   inv_image_def
   measure_def
--- a/src/HOL/Relation.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Relation.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -21,9 +21,9 @@
   converse  ("(_\<inverse>)" [1000] 999)
 
 definition
-  rel_comp  :: "[('b * 'c) set, ('a * 'b) set] => ('a * 'c) set"
+  rel_comp  :: "[('a * 'b) set, ('b * 'c) set] => ('a * 'c) set"
     (infixr "O" 75) where
-  "r O s == {(x,z). EX y. (x, y) : s & (y, z) : r}"
+  "r O s == {(x,z). EX y. (x, y) : r & (y, z) : s}"
 
 definition
   Image :: "[('a * 'b) set, 'a set] => 'b set"
@@ -140,15 +140,15 @@
 subsection {* Composition of two relations *}
 
 lemma rel_compI [intro]:
-  "(a, b) : s ==> (b, c) : r ==> (a, c) : r O s"
+  "(a, b) : r ==> (b, c) : s ==> (a, c) : r O s"
 by (unfold rel_comp_def) blast
 
 lemma rel_compE [elim!]: "xz : r O s ==>
-  (!!x y z. xz = (x, z) ==> (x, y) : s ==> (y, z) : r  ==> P) ==> P"
+  (!!x y z. xz = (x, z) ==> (x, y) : r ==> (y, z) : s  ==> P) ==> P"
 by (unfold rel_comp_def) (iprover elim!: CollectE splitE exE conjE)
 
 lemma rel_compEpair:
-  "(a, c) : r O s ==> (!!y. (a, y) : s ==> (y, c) : r ==> P) ==> P"
+  "(a, c) : r O s ==> (!!y. (a, y) : r ==> (y, c) : s ==> P) ==> P"
 by (iprover elim: rel_compE Pair_inject ssubst)
 
 lemma R_O_Id [simp]: "R O Id = R"
@@ -173,7 +173,7 @@
 by blast
 
 lemma rel_comp_subset_Sigma:
-    "s \<subseteq> A \<times> B ==> r \<subseteq> B \<times> C ==> (r O s) \<subseteq> A \<times> C"
+    "r \<subseteq> A \<times> B ==> s \<subseteq> B \<times> C ==> (r O s) \<subseteq> A \<times> C"
 by blast
 
 lemma rel_comp_distrib[simp]: "R O (S \<union> T) = (R O S) \<union> (R O T)" 
--- a/src/HOL/SAT.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/SAT.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -25,12 +25,12 @@
   maxtime=60,
   satsolver="auto"]
 
-ML {* structure sat = SATFunc(structure cnf = cnf); *}
+ML {* structure sat = SATFunc(cnf) *}
 
-method_setup sat = {* Scan.succeed (K (SIMPLE_METHOD' sat.sat_tac)) *}
+method_setup sat = {* Scan.succeed (SIMPLE_METHOD' o sat.sat_tac) *}
   "SAT solver"
 
-method_setup satx = {* Scan.succeed (K (SIMPLE_METHOD' sat.satx_tac)) *}
+method_setup satx = {* Scan.succeed (SIMPLE_METHOD' o sat.satx_tac) *}
   "SAT solver (with definitional CNF)"
 
 end
--- a/src/HOL/Set.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Set.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -100,8 +100,8 @@
 
 text {* Set enumerations *}
 
-definition empty :: "'a set" ("{}") where
-  bot_set_eq [symmetric]: "{} = bot"
+abbreviation empty :: "'a set" ("{}") where
+  "{} \<equiv> bot"
 
 definition insert :: "'a \<Rightarrow> 'a set \<Rightarrow> 'a set" where
   insert_compr: "insert a B = {x. x = a \<or> x \<in> B}"
@@ -541,12 +541,12 @@
 
 subsubsection {* The universal set -- UNIV *}
 
-definition UNIV :: "'a set" where
-  top_set_eq [symmetric]: "UNIV = top"
+abbreviation UNIV :: "'a set" where
+  "UNIV \<equiv> top"
 
 lemma UNIV_def:
   "UNIV = {x. True}"
-  by (simp add: top_set_eq [symmetric] top_fun_eq top_bool_eq Collect_def)
+  by (simp add: top_fun_eq top_bool_eq Collect_def)
 
 lemma UNIV_I [simp]: "x : UNIV"
   by (simp add: UNIV_def)
@@ -579,7 +579,7 @@
 
 lemma empty_def:
   "{} = {x. False}"
-  by (simp add: bot_set_eq [symmetric] bot_fun_eq bot_bool_eq Collect_def)
+  by (simp add: bot_fun_eq bot_bool_eq Collect_def)
 
 lemma empty_iff [simp]: "(c : {}) = False"
   by (simp add: empty_def)
--- a/src/HOL/SizeChange/sct.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/SizeChange/sct.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -210,7 +210,7 @@
 
     in
       if Thm.no_prems no_step
-      then NoStep (Goal.finish no_step RS no_stepI)
+      then NoStep (Goal.finish ctxt no_step RS no_stepI)
       else
         let
           fun set_m1 i =
@@ -240,10 +240,10 @@
                       val thm1 = decr with_m2
                     in
                       if Thm.no_prems thm1
-                      then ((rtac (inst_nums thy i j approx_less) 1) THEN (simp_nth_tac 1) THEN (rtac (Goal.finish thm1) 1))
+                      then ((rtac (inst_nums thy i j approx_less) 1) THEN (simp_nth_tac 1) THEN (rtac (Goal.finish ctxt thm1) 1))
                       else let val thm2 = decreq with_m2 in
                              if Thm.no_prems thm2
-                             then ((rtac (inst_nums thy i j approx_leq) 1) THEN (simp_nth_tac 1) THEN (rtac (Goal.finish thm2) 1))
+                             then ((rtac (inst_nums thy i j approx_leq) 1) THEN (simp_nth_tac 1) THEN (rtac (Goal.finish ctxt thm2) 1))
                              else all_tac end
                     end
               in set_m2 end
@@ -257,7 +257,7 @@
                     |> cterm_of thy
                     |> Goal.init
                     |> tac |> Seq.hd
-                    |> Goal.finish
+                    |> Goal.finish ctxt
 
           val _ $ (_ $ G $ _ $ _ $ _ $ _) = prop_of approx_thm
         in
--- a/src/HOL/Statespace/state_space.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Statespace/state_space.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -149,7 +149,7 @@
    thy
    |> Expression.sublocale_cmd name expr
    |> Proof.global_terminal_proof
-         (Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt),Position.none), NONE)
+         (Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt)), NONE)
    |> ProofContext.theory_of
 
 fun add_locale name expr elems thy =
--- a/src/HOL/Tools/Datatype/datatype.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Datatype/datatype.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -191,7 +191,7 @@
 
 fun add_cases_induct infos induction thy =
   let
-    val inducts = ProjectRule.projections (ProofContext.init thy) induction;
+    val inducts = Project_Rule.projections (ProofContext.init thy) induction;
 
     fun named_rules (name, {index, exhaustion, ...}: info) =
       [((Binding.empty, nth inducts index), [Induct.induct_type name]),
--- a/src/HOL/Tools/Function/fundef_datatype.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Function/fundef_datatype.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -208,13 +208,12 @@
 
 val by_pat_completeness_auto =
     Proof.global_future_terminal_proof
-      (Method.Basic (pat_completeness, Position.none),
+      (Method.Basic pat_completeness,
        SOME (Method.Source_i (Args.src (("HOL.auto", []), Position.none))))
 
 fun termination_by method int =
     Fundef.termination_proof NONE
-    #> Proof.global_future_terminal_proof
-      (Method.Basic (method, Position.none), NONE) int
+    #> Proof.global_future_terminal_proof (Method.Basic method, NONE) int
 
 fun mk_catchall fixes arity_of =
     let
--- a/src/HOL/Tools/Function/fundef_lib.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Function/fundef_lib.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -126,7 +126,10 @@
 fun try_proof cgoal tac = 
     case SINGLE tac (Goal.init cgoal) of
       NONE => Fail
-    | SOME st => if Thm.no_prems st then Solved (Goal.finish st) else Stuck st
+    | SOME st =>
+        if Thm.no_prems st
+        then Solved (Goal.finish (Syntax.init_pretty_global (Thm.theory_of_cterm cgoal)) st)
+        else Stuck st
 
 
 fun dest_binop_list cn (t as (Const (n, _) $ a $ b)) = 
--- a/src/HOL/Tools/Function/induction_scheme.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Function/induction_scheme.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -324,7 +324,7 @@
                          THEN CONVERSION ind_rulify 1)
                      |> Seq.hd
                      |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep)
-                     |> Goal.finish
+                     |> Goal.finish ctxt
                      |> implies_intr (cprop_of branch_hyp)
                      |> fold_rev (forall_intr o cert) fxs
           in
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -140,7 +140,7 @@
 fun pr_table table = writeln (cat_lines (map (fn r => concat (map pr_cell r)) table))
 
 fun pr_goals ctxt st =
-    Display_Goal.pretty_goals ctxt Markup.none (true, false) (Thm.nprems_of st) st
+    Goal_Display.pretty_goals ctxt {total = true, main = false, maxgoals = Thm.nprems_of st} st
      |> Pretty.chunks
      |> Pretty.string_of
 
--- a/src/HOL/Tools/Qelim/cooper.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -428,7 +428,7 @@
    in
     fun provein x S = 
      case term_of S of
-        Const(@{const_name Set.empty}, _) => error "Unexpected error in Cooper, please email Amine Chaieb"
+        Const(@{const_name Orderings.bot}, _) => error "Unexpected error in Cooper, please email Amine Chaieb"
       | Const(@{const_name insert}, _) $ y $ _ => 
          let val (cy,S') = Thm.dest_binop S
          in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
--- a/src/HOL/Tools/Qelim/ferrante_rackoff.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Qelim/ferrante_rackoff.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -99,7 +99,7 @@
    in
     fun provein x S =
      case term_of S of
-        Const(@{const_name Set.empty}, _) => raise CTERM ("provein : not a member!", [S])
+        Const(@{const_name Orderings.bot}, _) => raise CTERM ("provein : not a member!", [S])
       | Const(@{const_name insert}, _) $ y $_ =>
          let val (cy,S') = Thm.dest_binop S
          in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1
--- a/src/HOL/Tools/Qelim/langford.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/Qelim/langford.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -15,7 +15,7 @@
  let 
   fun h acc ct = 
    case term_of ct of
-     Const (@{const_name Set.empty}, _) => acc
+     Const (@{const_name Orderings.bot}, _) => acc
    | Const (@{const_name insert}, _) $ _ $ t => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct)
 in h [] end;
 
@@ -48,11 +48,11 @@
        in (ne, f) end
 
      val qe = case (term_of L, term_of U) of 
-      (Const (@{const_name Set.empty}, _),_) =>  
+      (Const (@{const_name Orderings.bot}, _),_) =>  
         let
           val (neU,fU) = proveneF U 
         in simp_rule (transitive ths (dlo_qeth_nolb OF [neU, fU])) end
-    | (_,Const (@{const_name Set.empty}, _)) =>  
+    | (_,Const (@{const_name Orderings.bot}, _)) =>  
         let
           val (neL,fL) = proveneF L
         in simp_rule (transitive ths (dlo_qeth_noub OF [neL, fL])) end
--- a/src/HOL/Tools/atp_wrapper.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/atp_wrapper.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -59,9 +59,7 @@
     val (ctxt, (chain_ths, th)) = goal
     val thy = ProofContext.theory_of ctxt
     val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths
-    val goal_cls = #1 (ResAxioms.neg_conjecture_clauses th subgoalno)
-      handle THM ("assume: variables", _, _) =>
-        error "Sledgehammer: Goal contains type variables (TVars)"
+    val goal_cls = #1 (ResAxioms.neg_conjecture_clauses ctxt th subgoalno)
     val _ = app (fn th => Output.debug (fn _ => Display.string_of_thm ctxt th)) goal_cls
     val the_filtered_clauses =
       case filtered_clauses of
@@ -71,7 +69,8 @@
       case axiom_clauses of
           NONE => the_filtered_clauses
         | SOME axcls => axcls
-    val (thm_names, clauses) = preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy
+    val (thm_names, clauses) =
+      preparer goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy
 
     (* write out problem file and call prover *)
     val probfile = prob_pathname subgoalno
@@ -85,7 +84,7 @@
     val _ =
       if destdir' = "" then File.rm probfile
       else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof
-    
+
     (* check for success and print out some information on failure *)
     val failure = find_failure proof
     val success = rc = 0 andalso is_none failure
@@ -133,7 +132,7 @@
 fun vampire_opts max_new theory_const timeout = tptp_prover_opts
   max_new theory_const
   (Path.explode "$VAMPIRE_HOME/vampire",
-               ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
+    ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
   timeout;
 
 val vampire = vampire_opts 60 false;
@@ -141,7 +140,7 @@
 fun vampire_opts_full max_new theory_const timeout = full_prover_opts
   max_new theory_const
   (Path.explode "$VAMPIRE_HOME/vampire",
-               ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
+    ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
   timeout;
 
 val vampire_full = vampire_opts_full 60 false;
@@ -173,7 +172,8 @@
   (ResAtp.prepare_clauses true)
   (ResHolClause.dfg_write_file (AtpManager.get_full_types()))
   (Path.explode "$SPASS_HOME/SPASS",
-    "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^ string_of_int timeout)
+    "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^
+      string_of_int timeout)
   ResReconstruct.find_failure
   (ResReconstruct.lemma_list true)
   timeout ax_clauses fcls name n goal;
@@ -196,22 +196,23 @@
   end;
 
 fun refresh_systems () = Synchronized.change systems (fn _ =>
-  get_systems());
+  get_systems ());
 
 fun get_system prefix = Synchronized.change_result systems (fn systems =>
   let val systems = if null systems then get_systems() else systems
   in (find_first (String.isPrefix prefix) systems, systems) end);
 
 fun remote_prover_opts max_new theory_const args prover_prefix timeout =
-  let val sys = case get_system prover_prefix of
+  let val sys =
+    case get_system prover_prefix of
       NONE => error ("No system like " ^ quote prover_prefix ^ " at SystemOnTPTP")
     | SOME sys => sys
   in tptp_prover_opts max_new theory_const
     (Path.explode "$ISABELLE_HOME/lib/scripts/SystemOnTPTP",
-      args ^ " -t " ^ string_of_int timeout ^ " -s " ^ sys) timeout end;
+      args ^ " -t " ^ string_of_int timeout ^ " -s " ^ sys) timeout
+  end;
 
 val remote_prover = remote_prover_opts 60 false;
 
 end;
 
-
--- a/src/HOL/Tools/cnf_funcs.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/cnf_funcs.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -33,19 +33,20 @@
 
 signature CNF =
 sig
-	val is_atom           : Term.term -> bool
-	val is_literal        : Term.term -> bool
-	val is_clause         : Term.term -> bool
-	val clause_is_trivial : Term.term -> bool
+	val is_atom: term -> bool
+	val is_literal: term -> bool
+	val is_clause: term -> bool
+	val clause_is_trivial: term -> bool
 
-	val clause2raw_thm : Thm.thm -> Thm.thm
+	val clause2raw_thm: thm -> thm
 
-	val weakening_tac : int -> Tactical.tactic  (* removes the first hypothesis of a subgoal *)
+	val weakening_tac: int -> tactic  (* removes the first hypothesis of a subgoal *)
 
-	val make_cnf_thm  : theory -> Term.term -> Thm.thm
-	val make_cnfx_thm : theory -> Term.term ->  Thm.thm
-	val cnf_rewrite_tac  : int -> Tactical.tactic  (* converts all prems of a subgoal to CNF *)
-	val cnfx_rewrite_tac : int -> Tactical.tactic  (* converts all prems of a subgoal to (almost) definitional CNF *)
+	val make_cnf_thm: theory -> term -> thm
+	val make_cnfx_thm: theory -> term -> thm
+	val cnf_rewrite_tac: Proof.context -> int -> tactic  (* converts all prems of a subgoal to CNF *)
+	val cnfx_rewrite_tac: Proof.context -> int -> tactic
+	  (* converts all prems of a subgoal to (almost) definitional CNF *)
 end;
 
 structure cnf : CNF =
@@ -505,8 +506,6 @@
 (* weakening_tac: removes the first hypothesis of the 'i'-th subgoal         *)
 (* ------------------------------------------------------------------------- *)
 
-(* int -> Tactical.tactic *)
-
 fun weakening_tac i =
 	dtac weakening_thm i THEN atac (i+1);
 
@@ -516,17 +515,16 @@
 (*      premise)                                                             *)
 (* ------------------------------------------------------------------------- *)
 
-(* int -> Tactical.tactic *)
-
-fun cnf_rewrite_tac i =
+fun cnf_rewrite_tac ctxt i =
 	(* cut the CNF formulas as new premises *)
-	METAHYPS (fn prems =>
+	FOCUS (fn {prems, ...} =>
 		let
-			val cnf_thms = map (fn pr => make_cnf_thm (theory_of_thm pr) ((HOLogic.dest_Trueprop o prop_of) pr)) prems
+		  val thy = ProofContext.theory_of ctxt
+			val cnf_thms = map (make_cnf_thm thy o HOLogic.dest_Trueprop o Thm.prop_of) prems
 			val cut_thms = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnf_thms ~~ prems)
 		in
 			cut_facts_tac cut_thms 1
-		end) i
+		end) ctxt i
 	(* remove the original premises *)
 	THEN SELECT_GOAL (fn thm =>
 		let
@@ -540,17 +538,16 @@
 (*      (possibly introducing new literals)                                  *)
 (* ------------------------------------------------------------------------- *)
 
-(* int -> Tactical.tactic *)
-
-fun cnfx_rewrite_tac i =
+fun cnfx_rewrite_tac ctxt i =
 	(* cut the CNF formulas as new premises *)
-	METAHYPS (fn prems =>
+	FOCUS (fn {prems, ...} =>
 		let
-			val cnfx_thms = map (fn pr => make_cnfx_thm (theory_of_thm pr) ((HOLogic.dest_Trueprop o prop_of) pr)) prems
+		  val thy = ProofContext.theory_of ctxt;
+			val cnfx_thms = map (make_cnfx_thm thy o HOLogic.dest_Trueprop o prop_of) prems
 			val cut_thms  = map (fn (th, pr) => cnftac_eq_imp OF [th, pr]) (cnfx_thms ~~ prems)
 		in
 			cut_facts_tac cut_thms 1
-		end) i
+		end) ctxt i
 	(* remove the original premises *)
 	THEN SELECT_GOAL (fn thm =>
 		let
--- a/src/HOL/Tools/hologic.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/hologic.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -153,13 +153,13 @@
 fun mk_set T ts =
   let
     val sT = mk_setT T;
-    val empty = Const ("Set.empty", sT);
+    val empty = Const ("Orderings.bot_class.bot", sT);
     fun insert t u = Const ("Set.insert", T --> sT --> sT) $ t $ u;
   in fold_rev insert ts empty end;
 
-fun mk_UNIV T = Const ("Set.UNIV", mk_setT T);
+fun mk_UNIV T = Const ("Orderings.top_class.top", mk_setT T);
 
-fun dest_set (Const ("Set.empty", _)) = []
+fun dest_set (Const ("Orderings.bot_class.bot", _)) = []
   | dest_set (Const ("Set.insert", _) $ t $ u) = t :: dest_set u
   | dest_set t = raise TERM ("dest_set", [t]);
 
--- a/src/HOL/Tools/inductive.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/inductive.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -96,11 +96,12 @@
 
 val notTrueE = TrueI RSN (2, notE);
 val notFalseI = Seq.hd (atac 1 notI);
-val simp_thms' = map (fn s => mk_meta_eq (the (find_first
-  (equal (OldGoals.read_prop @{theory HOL} s) o prop_of) simp_thms)))
-  ["(~True) = False", "(~False) = True",
-   "(True --> ?P) = ?P", "(False --> ?P) = True",
-   "(?P & True) = ?P", "(True & ?P) = ?P"];
+
+val simp_thms' = map mk_meta_eq
+  @{lemma "(~True) = False" "(~False) = True"
+      "(True --> P) = P" "(False --> P) = True"
+      "(P & True) = P" "(True & P) = P"
+    by (fact simp_thms)+};
 
 
 
@@ -712,7 +713,7 @@
           map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]);
 
     val ctxt3 = if no_ind orelse coind then ctxt2 else
-      let val inducts = cnames ~~ ProjectRule.projects ctxt2 (1 upto length cnames) induct'
+      let val inducts = cnames ~~ Project_Rule.projects ctxt2 (1 upto length cnames) induct'
       in
         ctxt2 |>
         LocalTheory.notes kind [((rec_qualified (Binding.name "inducts"), []),
--- a/src/HOL/Tools/meson.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/meson.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -132,7 +132,7 @@
                 Display.string_of_thm_without_context st ::
                 "Premises:" :: map Display.string_of_thm_without_context prems))
   in
-    case Seq.pull (ALLGOALS (METAHYPS forward_tacf) st)
+    case Seq.pull (ALLGOALS (OldGoals.METAHYPS forward_tacf) st)
     of SOME(th,_) => th
      | NONE => raise THM("forward_res", 0, [st])
   end;
@@ -226,7 +226,7 @@
 fun forward_res2 nf hyps st =
   case Seq.pull
         (REPEAT
-         (METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
+         (OldGoals.METAHYPS (fn major::minors => rtac (nf (minors@hyps) major) 1) 1)
          st)
   of SOME(th,_) => th
    | NONE => raise THM("forward_res2", 0, [st]);
@@ -335,8 +335,8 @@
               (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using
                 all combinations of converting P, Q to CNF.*)
               let val tac =
-                  (METAHYPS (resop cnf_nil) 1) THEN
-                   (fn st' => st' |> METAHYPS (resop cnf_nil) 1)
+                  (OldGoals.METAHYPS (resop cnf_nil) 1) THEN
+                   (fn st' => st' |> OldGoals.METAHYPS (resop cnf_nil) 1)
               in  Seq.list_of (tac (th RS disj_forward)) @ ths  end
           | _ => nodups th :: ths  (*no work to do*)
       and cnf_nil th = cnf_aux (th,[])
@@ -584,9 +584,9 @@
   SELECT_GOAL
     (EVERY [ObjectLogic.atomize_prems_tac 1,
             rtac ccontr 1,
-            METAHYPS (fn negs =>
+            OldGoals.METAHYPS (fn negs =>
                       EVERY1 [skolemize_prems_tac negs,
-                              METAHYPS (cltac o mkcl)]) 1]) i st
+                              OldGoals.METAHYPS (cltac o mkcl)]) 1]) i st
   handle THM _ => no_tac st;    (*probably from make_meta_clause, not first-order*)
 
 (** Best-first search versions **)
@@ -698,7 +698,7 @@
 fun skolemize_tac i st =
   let val ts = Logic.strip_assums_hyp (List.nth (prems_of st, i-1))
   in
-     EVERY' [METAHYPS
+     EVERY' [OldGoals.METAHYPS
             (fn hyps => (cut_facts_tac (skolemize_nnf_list hyps) 1
                          THEN REPEAT (etac exE 1))),
             REPEAT_DETERM_N (length ts) o (etac thin_rl)] i st
--- a/src/HOL/Tools/record.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/record.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -2161,7 +2161,7 @@
       fn st => let val [s, s'] = map #1 (rev (Tactic.innermost_params 1 st)) in
         st |> (res_inst_tac context [((rN, 0), s)] cases_scheme 1
         THEN res_inst_tac context [((rN, 0), s')] cases_scheme 1
-        THEN (METAHYPS equality_tac 1))
+        THEN (FOCUS (fn {prems, ...} => equality_tac prems) context 1))
              (* simp_all_tac ss (sel_convs) would also work but is less efficient *)
       end);
      val equality = timeit_msg "record equality proof:" equality_prf;
--- a/src/HOL/Tools/res_axioms.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/res_axioms.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -14,7 +14,7 @@
   val neg_clausify: thm list -> thm list
   val expand_defs_tac: thm -> tactic
   val combinators: thm -> thm
-  val neg_conjecture_clauses: thm -> int -> thm list * (string * typ) list
+  val neg_conjecture_clauses: Proof.context -> thm -> int -> thm list * (string * typ) list
   val atpset_rules_of: Proof.context -> (string * thm) list
   val suppress_endtheory: bool ref     (*for emergency use where endtheory causes problems*)
   val setup: theory -> theory
@@ -498,32 +498,30 @@
 
 val neg_skolemize_tac = EVERY' [rtac ccontr, ObjectLogic.atomize_prems_tac, Meson.skolemize_tac];
 
-fun neg_clausify sts =
-  sts |> Meson.make_clauses |> map combinators |> Meson.finish_cnf;
+val neg_clausify = Meson.make_clauses #> map combinators #> Meson.finish_cnf;
 
-fun neg_conjecture_clauses st0 n =
-  let val st = Seq.hd (neg_skolemize_tac n st0)
-      val (params,_,_) = strip_context (Logic.nth_prem (n, Thm.prop_of st))
-  in (neg_clausify (the (metahyps_thms n st)), params) end
-  handle Option => error "unable to Skolemize subgoal";
+fun neg_conjecture_clauses ctxt st0 n =
+  let
+    val st = Seq.hd (neg_skolemize_tac n st0)
+    val ({params, prems, ...}, _) = Subgoal.focus (Variable.set_body false ctxt) n st
+  in (neg_clausify prems, map (Term.dest_Free o Thm.term_of o #2) params) end;
 
 (*Conversion of a subgoal to conjecture clauses. Each clause has
   leading !!-bound universal variables, to express generality. *)
-val neg_clausify_tac =
+fun neg_clausify_tac ctxt =
   neg_skolemize_tac THEN'
-  SUBGOAL
-    (fn (prop,_) =>
-     let val ts = Logic.strip_assums_hyp prop
-     in EVERY1
-         [METAHYPS
-            (fn hyps =>
-              (Method.insert_tac
-                (map forall_intr_vars (neg_clausify hyps)) 1)),
-          REPEAT_DETERM_N (length ts) o (etac thin_rl)]
+  SUBGOAL (fn (prop, i) =>
+    let val ts = Logic.strip_assums_hyp prop in
+      EVERY'
+       [FOCUS
+         (fn {prems, ...} =>
+           (Method.insert_tac
+             (map forall_intr_vars (neg_clausify prems)) i)) ctxt,
+        REPEAT_DETERM_N (length ts) o etac thin_rl] i
      end);
 
 val neg_clausify_setup =
-  Method.setup @{binding neg_clausify} (Scan.succeed (K (SIMPLE_METHOD' neg_clausify_tac)))
+  Method.setup @{binding neg_clausify} (Scan.succeed (SIMPLE_METHOD' o neg_clausify_tac))
   "conversion of goal to conjecture clauses";
 
 
--- a/src/HOL/Tools/res_clause.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/res_clause.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -99,20 +99,10 @@
 val const_trans_table =
       Symtab.make [(@{const_name "op ="}, "equal"),
                    (@{const_name HOL.less_eq}, "lessequals"),
-                   (@{const_name HOL.less}, "less"),
                    (@{const_name "op &"}, "and"),
                    (@{const_name "op |"}, "or"),
-                   (@{const_name HOL.plus}, "plus"),
-                   (@{const_name HOL.minus}, "minus"),
-                   (@{const_name HOL.times}, "times"),
-                   (@{const_name Divides.div}, "div"),
-                   (@{const_name HOL.divide}, "divide"),
                    (@{const_name "op -->"}, "implies"),
-                   (@{const_name Set.empty}, "emptyset"),
                    (@{const_name "op :"}, "in"),
-                   (@{const_name union}, "union"),
-                   (@{const_name inter}, "inter"),
-                   ("List.append", "append"),
                    ("ATP_Linkup.fequal", "fequal"),
                    ("ATP_Linkup.COMBI", "COMBI"),
                    ("ATP_Linkup.COMBK", "COMBK"),
--- a/src/HOL/Tools/res_reconstruct.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/res_reconstruct.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -442,7 +442,7 @@
       val _ = trace (Int.toString (length nonnull_lines) ^ " nonnull_lines extracted\n")
       val (_,lines) = List.foldr (add_wanted_prfline ctxt) (0,[]) nonnull_lines
       val _ = trace (Int.toString (length lines) ^ " lines extracted\n")
-      val (ccls,fixes) = ResAxioms.neg_conjecture_clauses th sgno
+      val (ccls,fixes) = ResAxioms.neg_conjecture_clauses ctxt th sgno
       val _ = trace (Int.toString (length ccls) ^ " conjecture clauses\n")
       val ccls = map forall_intr_vars ccls
       val _ =
@@ -452,10 +452,7 @@
       val _ = trace "\ndecode_tstp_file: finishing\n"
   in
     isar_header (map #1 fixes) ^ String.concat ilines ^ "qed\n"
-  end
-  handle e => (*FIXME: exn handler is too general!*)
-    let val msg = "Translation of TSTP raised an exception: " ^ ML_Compiler.exn_message e
-    in  trace msg; msg  end;
+  end;
 
 
   (*=== EXTRACTING PROOF-TEXT === *)
--- a/src/HOL/Tools/sat_funcs.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/sat_funcs.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -48,16 +48,16 @@
 
 signature SAT =
 sig
-	val trace_sat  : bool ref    (* input: print trace messages *)
-	val solver     : string ref  (* input: name of SAT solver to be used *)
-	val counter    : int ref     (* output: number of resolution steps during last proof replay *)
-	val rawsat_thm : cterm list -> thm
-	val rawsat_tac : int -> Tactical.tactic
-	val sat_tac    : int -> Tactical.tactic
-	val satx_tac   : int -> Tactical.tactic
+	val trace_sat: bool ref    (* input: print trace messages *)
+	val solver: string ref  (* input: name of SAT solver to be used *)
+	val counter: int ref     (* output: number of resolution steps during last proof replay *)
+	val rawsat_thm: cterm list -> thm
+	val rawsat_tac: Proof.context -> int -> tactic
+	val sat_tac: Proof.context -> int -> tactic
+	val satx_tac: Proof.context -> int -> tactic
 end
 
-functor SATFunc (structure cnf : CNF) : SAT =
+functor SATFunc(cnf : CNF) : SAT =
 struct
 
 val trace_sat = ref false;
@@ -410,7 +410,8 @@
 (*      or "True"                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-fun rawsat_tac i = METAHYPS (fn prems => rtac (rawsat_thm (map cprop_of prems)) 1) i;
+fun rawsat_tac ctxt i =
+  FOCUS (fn {prems, ...} => rtac (rawsat_thm (map cprop_of prems)) 1) ctxt i;
 
 (* ------------------------------------------------------------------------- *)
 (* pre_cnf_tac: converts the i-th subgoal                                    *)
@@ -436,8 +437,8 @@
 (*      subgoal                                                              *)
 (* ------------------------------------------------------------------------- *)
 
-fun cnfsat_tac i =
-	(etac FalseE i) ORELSE (REPEAT_DETERM (etac conjE i) THEN rawsat_tac i);
+fun cnfsat_tac ctxt i =
+	(etac FalseE i) ORELSE (REPEAT_DETERM (etac conjE i) THEN rawsat_tac ctxt i);
 
 (* ------------------------------------------------------------------------- *)
 (* cnfxsat_tac: checks if the empty clause "False" occurs among the          *)
@@ -446,9 +447,9 @@
 (*      then applies 'rawsat_tac' to solve the subgoal                       *)
 (* ------------------------------------------------------------------------- *)
 
-fun cnfxsat_tac i =
+fun cnfxsat_tac ctxt i =
 	(etac FalseE i) ORELSE
-		(REPEAT_DETERM (etac conjE i ORELSE etac exE i) THEN rawsat_tac i);
+		(REPEAT_DETERM (etac conjE i ORELSE etac exE i) THEN rawsat_tac ctxt i);
 
 (* ------------------------------------------------------------------------- *)
 (* sat_tac: tactic for calling an external SAT solver, taking as input an    *)
@@ -456,8 +457,8 @@
 (*      an exponential blowup.                                               *)
 (* ------------------------------------------------------------------------- *)
 
-fun sat_tac i =
-	pre_cnf_tac i THEN cnf.cnf_rewrite_tac i THEN cnfsat_tac i;
+fun sat_tac ctxt i =
+	pre_cnf_tac i THEN cnf.cnf_rewrite_tac ctxt i THEN cnfsat_tac ctxt i;
 
 (* ------------------------------------------------------------------------- *)
 (* satx_tac: tactic for calling an external SAT solver, taking as input an   *)
@@ -465,7 +466,7 @@
 (*      introducing new literals.                                            *)
 (* ------------------------------------------------------------------------- *)
 
-fun satx_tac i =
-	pre_cnf_tac i THEN cnf.cnfx_rewrite_tac i THEN cnfxsat_tac i;
+fun satx_tac ctxt i =
+	pre_cnf_tac i THEN cnf.cnfx_rewrite_tac ctxt i THEN cnfxsat_tac ctxt i;
 
 end;
--- a/src/HOL/Tools/simpdata.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Tools/simpdata.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -126,27 +126,25 @@
 
 val safe_solver = mk_solver "HOL safe" safe_solver_tac;
 
-structure SplitterData =
-struct
-  structure Simplifier = Simplifier
-  val mk_eq           = mk_eq
-  val meta_eq_to_iff  = @{thm meta_eq_to_obj_eq}
-  val iffD            = @{thm iffD2}
-  val disjE           = @{thm disjE}
-  val conjE           = @{thm conjE}
-  val exE             = @{thm exE}
-  val contrapos       = @{thm contrapos_nn}
-  val contrapos2      = @{thm contrapos_pp}
-  val notnotD         = @{thm notnotD}
-end;
+structure Splitter = Splitter
+(
+  val thy = @{theory}
+  val mk_eq = mk_eq
+  val meta_eq_to_iff = @{thm meta_eq_to_obj_eq}
+  val iffD = @{thm iffD2}
+  val disjE = @{thm disjE}
+  val conjE = @{thm conjE}
+  val exE = @{thm exE}
+  val contrapos = @{thm contrapos_nn}
+  val contrapos2 = @{thm contrapos_pp}
+  val notnotD = @{thm notnotD}
+);
 
-structure Splitter = SplitterFun(SplitterData);
-
-val split_tac        = Splitter.split_tac;
+val split_tac = Splitter.split_tac;
 val split_inside_tac = Splitter.split_inside_tac;
 
-val op addsplits     = Splitter.addsplits;
-val op delsplits     = Splitter.delsplits;
+val op addsplits = Splitter.addsplits;
+val op delsplits = Splitter.delsplits;
 
 
 (* integration of simplifier with classical reasoner *)
--- a/src/HOL/Transitive_Closure.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Transitive_Closure.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -159,7 +159,7 @@
   apply (erule asm_rl exE disjE conjE base step)+
   done
 
-lemma rtrancl_Int_subset: "[| Id \<subseteq> s; r O (r^* \<inter> s) \<subseteq> s|] ==> r^* \<subseteq> s"
+lemma rtrancl_Int_subset: "[| Id \<subseteq> s; (r^* \<inter> s) O r \<subseteq> s|] ==> r^* \<subseteq> s"
   apply (rule subsetI)
   apply (rule_tac p="x" in PairE, clarify)
   apply (erule rtrancl_induct, auto) 
@@ -291,7 +291,7 @@
   by (blast elim: rtranclE converse_rtranclE
     intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl)
 
-lemma rtrancl_unfold: "r^* = Id Un r O r^*"
+lemma rtrancl_unfold: "r^* = Id Un r^* O r"
   by (auto intro: rtrancl_into_rtrancl elim: rtranclE)
 
 lemma rtrancl_Un_separatorE:
@@ -384,7 +384,7 @@
   | (step) c where "(a, c) : r^+" and "(c, b) : r"
   using assms by cases simp_all
 
-lemma trancl_Int_subset: "[| r \<subseteq> s; r O (r^+ \<inter> s) \<subseteq> s|] ==> r^+ \<subseteq> s"
+lemma trancl_Int_subset: "[| r \<subseteq> s; (r^+ \<inter> s) O r \<subseteq> s|] ==> r^+ \<subseteq> s"
   apply (rule subsetI)
   apply (rule_tac p = x in PairE)
   apply clarify
@@ -392,7 +392,7 @@
    apply auto
   done
 
-lemma trancl_unfold: "r^+ = r Un r O r^+"
+lemma trancl_unfold: "r^+ = r Un r^+ O r"
   by (auto intro: trancl_into_trancl elim: tranclE)
 
 text {* Transitivity of @{term "r^+"} *}
@@ -676,7 +676,7 @@
 
 primrec relpow :: "nat \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set" where
     "relpow 0 R = Id"
-  | "relpow (Suc n) R = R O (R ^^ n)"
+  | "relpow (Suc n) R = (R ^^ n) O R"
 
 end
 
@@ -734,11 +734,11 @@
   apply (cut_tac n=nat and R=R in rel_pow_Suc_D2', simp, blast)
   done
 
-lemma rel_pow_add: "R ^^ (m+n) = R^^n O R^^m"
+lemma rel_pow_add: "R ^^ (m+n) = R^^m O R^^n"
 by(induct n) auto
 
 lemma rel_pow_commute: "R O R ^^ n = R ^^ n O R"
-  by (induct n) (simp, simp add: O_assoc [symmetric])
+by (induct n) (simp, simp add: O_assoc [symmetric])
 
 lemma rtrancl_imp_UN_rel_pow:
   assumes "p \<in> R^*"
@@ -811,16 +811,16 @@
 
 ML {*
 
-structure Trancl_Tac = Trancl_Tac_Fun (
-  struct
-    val r_into_trancl = @{thm trancl.r_into_trancl};
-    val trancl_trans  = @{thm trancl_trans};
-    val rtrancl_refl = @{thm rtrancl.rtrancl_refl};
-    val r_into_rtrancl = @{thm r_into_rtrancl};
-    val trancl_into_rtrancl = @{thm trancl_into_rtrancl};
-    val rtrancl_trancl_trancl = @{thm rtrancl_trancl_trancl};
-    val trancl_rtrancl_trancl = @{thm trancl_rtrancl_trancl};
-    val rtrancl_trans = @{thm rtrancl_trans};
+structure Trancl_Tac = Trancl_Tac
+(
+  val r_into_trancl = @{thm trancl.r_into_trancl};
+  val trancl_trans  = @{thm trancl_trans};
+  val rtrancl_refl = @{thm rtrancl.rtrancl_refl};
+  val r_into_rtrancl = @{thm r_into_rtrancl};
+  val trancl_into_rtrancl = @{thm trancl_into_rtrancl};
+  val rtrancl_trancl_trancl = @{thm rtrancl_trancl_trancl};
+  val trancl_rtrancl_trancl = @{thm trancl_rtrancl_trancl};
+  val rtrancl_trans = @{thm rtrancl_trans};
 
   fun decomp (@{const Trueprop} $ t) =
     let fun dec (Const ("op :", _) $ (Const ("Pair", _) $ a $ b) $ rel ) =
@@ -832,19 +832,18 @@
       | dec _ =  NONE
     in dec t end
     | decomp _ = NONE;
-
-  end);
+);
 
-structure Tranclp_Tac = Trancl_Tac_Fun (
-  struct
-    val r_into_trancl = @{thm tranclp.r_into_trancl};
-    val trancl_trans  = @{thm tranclp_trans};
-    val rtrancl_refl = @{thm rtranclp.rtrancl_refl};
-    val r_into_rtrancl = @{thm r_into_rtranclp};
-    val trancl_into_rtrancl = @{thm tranclp_into_rtranclp};
-    val rtrancl_trancl_trancl = @{thm rtranclp_tranclp_tranclp};
-    val trancl_rtrancl_trancl = @{thm tranclp_rtranclp_tranclp};
-    val rtrancl_trans = @{thm rtranclp_trans};
+structure Tranclp_Tac = Trancl_Tac
+(
+  val r_into_trancl = @{thm tranclp.r_into_trancl};
+  val trancl_trans  = @{thm tranclp_trans};
+  val rtrancl_refl = @{thm rtranclp.rtrancl_refl};
+  val r_into_rtrancl = @{thm r_into_rtranclp};
+  val trancl_into_rtrancl = @{thm tranclp_into_rtranclp};
+  val rtrancl_trancl_trancl = @{thm rtranclp_tranclp_tranclp};
+  val trancl_rtrancl_trancl = @{thm tranclp_rtranclp_tranclp};
+  val rtrancl_trans = @{thm rtranclp_trans};
 
   fun decomp (@{const Trueprop} $ t) =
     let fun dec (rel $ a $ b) =
@@ -856,31 +855,31 @@
       | dec _ =  NONE
     in dec t end
     | decomp _ = NONE;
-
-  end);
+);
 *}
 
 declaration {* fn _ =>
   Simplifier.map_ss (fn ss => ss
-    addSolver (mk_solver "Trancl" (fn _ => Trancl_Tac.trancl_tac))
-    addSolver (mk_solver "Rtrancl" (fn _ => Trancl_Tac.rtrancl_tac))
-    addSolver (mk_solver "Tranclp" (fn _ => Tranclp_Tac.trancl_tac))
-    addSolver (mk_solver "Rtranclp" (fn _ => Tranclp_Tac.rtrancl_tac)))
+    addSolver (mk_solver' "Trancl" (Trancl_Tac.trancl_tac o Simplifier.the_context))
+    addSolver (mk_solver' "Rtrancl" (Trancl_Tac.rtrancl_tac o Simplifier.the_context))
+    addSolver (mk_solver' "Tranclp" (Tranclp_Tac.trancl_tac o Simplifier.the_context))
+    addSolver (mk_solver' "Rtranclp" (Tranclp_Tac.rtrancl_tac o Simplifier.the_context)))
 *}
 
-(* Optional methods *)
+
+text {* Optional methods. *}
 
 method_setup trancl =
-  {* Scan.succeed (K (SIMPLE_METHOD' Trancl_Tac.trancl_tac)) *}
+  {* Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.trancl_tac) *}
   {* simple transitivity reasoner *}
 method_setup rtrancl =
-  {* Scan.succeed (K (SIMPLE_METHOD' Trancl_Tac.rtrancl_tac)) *}
+  {* Scan.succeed (SIMPLE_METHOD' o Trancl_Tac.rtrancl_tac) *}
   {* simple transitivity reasoner *}
 method_setup tranclp =
-  {* Scan.succeed (K (SIMPLE_METHOD' Tranclp_Tac.trancl_tac)) *}
+  {* Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.trancl_tac) *}
   {* simple transitivity reasoner (predicate version) *}
 method_setup rtranclp =
-  {* Scan.succeed (K (SIMPLE_METHOD' Tranclp_Tac.rtrancl_tac)) *}
+  {* Scan.succeed (SIMPLE_METHOD' o Tranclp_Tac.rtrancl_tac) *}
   {* simple transitivity reasoner (predicate version) *}
 
 end
--- a/src/HOL/UNITY/ListOrder.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/UNITY/ListOrder.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -117,7 +117,7 @@
 (*Lemma proving transitivity and more*)
 lemma genPrefix_trans_O [rule_format]: 
      "(x, y) : genPrefix r  
-      ==> ALL z. (y,z) : genPrefix s --> (x, z) : genPrefix (s O r)"
+      ==> ALL z. (y,z) : genPrefix s --> (x, z) : genPrefix (r O s)"
 apply (erule genPrefix.induct)
   prefer 3 apply (blast dest: append_genPrefix)
  prefer 2 apply (blast intro: genPrefix.prepend, blast)
@@ -134,13 +134,15 @@
 lemma prefix_genPrefix_trans [rule_format]: 
      "[| x<=y;  (y,z) : genPrefix r |] ==> (x, z) : genPrefix r"
 apply (unfold prefix_def)
-apply (subst R_O_Id [symmetric], erule genPrefix_trans_O, assumption)
+apply (drule genPrefix_trans_O, assumption)
+apply simp
 done
 
 lemma genPrefix_prefix_trans [rule_format]: 
      "[| (x,y) : genPrefix r;  y<=z |] ==> (x,z) : genPrefix r"
 apply (unfold prefix_def)
-apply (subst Id_O_R [symmetric], erule genPrefix_trans_O, assumption)
+apply (drule genPrefix_trans_O, assumption)
+apply simp
 done
 
 lemma trans_genPrefix: "trans r ==> trans (genPrefix r)"
--- a/src/HOL/Wellfounded.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/Wellfounded.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -187,8 +187,12 @@
 lemma wf_empty [iff]: "wf({})"
   by (simp add: wf_def)
 
-lemmas wfP_empty [iff] =
-  wf_empty [to_pred bot_empty_eq2, simplified bot_fun_eq bot_bool_eq]
+lemma wfP_empty [iff]:
+  "wfP (\<lambda>x y. False)"
+proof -
+  have "wfP bot" by (fact wf_empty [to_pred bot_empty_eq2])
+  then show ?thesis by (simp add: bot_fun_eq bot_bool_eq)
+qed
 
 lemma wf_Int1: "wf r ==> wf (r Int r')"
   apply (erule wf_subset)
@@ -235,7 +239,7 @@
 
 lemma wf_union_compatible:
   assumes "wf R" "wf S"
-  assumes "S O R \<subseteq> R"
+  assumes "R O S \<subseteq> R"
   shows "wf (R \<union> S)"
 proof (rule wfI_min)
   fix x :: 'a and Q 
@@ -254,8 +258,8 @@
       assume "y \<in> Q"
       with `y \<notin> ?Q'` 
       obtain w where "(w, y) \<in> R" and "w \<in> Q" by auto
-      from `(w, y) \<in> R` `(y, z) \<in> S` have "(w, z) \<in> S O R" by (rule rel_compI)
-      with `S O R \<subseteq> R` have "(w, z) \<in> R" ..
+      from `(w, y) \<in> R` `(y, z) \<in> S` have "(w, z) \<in> R O S" by (rule rel_compI)
+      with `R O S \<subseteq> R` have "(w, z) \<in> R" ..
       with `z \<in> ?Q'` have "w \<notin> Q" by blast 
       with `w \<in> Q` show False by contradiction
     qed
@@ -279,8 +283,10 @@
 apply (blast elim!: allE)  
 done
 
-lemmas wfP_SUP = wf_UN [where I=UNIV and r="\<lambda>i. {(x, y). r i x y}",
-  to_pred SUP_UN_eq2 bot_empty_eq pred_equals_eq, simplified, standard]
+lemma wfP_SUP:
+  "\<forall>i. wfP (r i) \<Longrightarrow> \<forall>i j. r i \<noteq> r j \<longrightarrow> inf (DomainP (r i)) (RangeP (r j)) = bot \<Longrightarrow> wfP (SUPR UNIV r)"
+  by (rule wf_UN [where I=UNIV and r="\<lambda>i. {(x, y). r i x y}", to_pred SUP_UN_eq2 pred_equals_eq])
+    (simp_all add: bot_fun_eq bot_bool_eq)
 
 lemma wf_Union: 
  "[| ALL r:R. wf r;  
@@ -308,7 +314,7 @@
   by (auto simp: Un_ac)
 
 lemma wf_union_merge: 
-  "wf (R \<union> S) = wf (R O R \<union> R O S \<union> S)" (is "wf ?A = wf ?B")
+  "wf (R \<union> S) = wf (R O R \<union> S O R \<union> S)" (is "wf ?A = wf ?B")
 proof
   assume "wf ?A"
   with wf_trancl have wfT: "wf (?A^+)" .
@@ -327,7 +333,7 @@
     obtain z where "z \<in> Q" and "\<And>y. (y, z) \<in> ?B \<Longrightarrow> y \<notin> Q" 
       by (erule wfE_min)
     then have A1: "\<And>y. (y, z) \<in> R O R \<Longrightarrow> y \<notin> Q"
-      and A2: "\<And>y. (y, z) \<in> R O S \<Longrightarrow> y \<notin> Q"
+      and A2: "\<And>y. (y, z) \<in> S O R \<Longrightarrow> y \<notin> Q"
       and A3: "\<And>y. (y, z) \<in> S \<Longrightarrow> y \<notin> Q"
       by auto
     
@@ -349,7 +355,7 @@
           with A1 show "y \<notin> Q" .
         next
           assume "(y, z') \<in> S" 
-          then have "(y, z) \<in> R O S" using  `(z', z) \<in> R` ..
+          then have "(y, z) \<in> S O R" using  `(z', z) \<in> R` ..
           with A2 show "y \<notin> Q" .
         qed
       qed
@@ -882,30 +888,6 @@
   qed
 qed
 
-text {*Wellfoundedness of @{text same_fst}*}
-
-definition
- same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
-where
-    "same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
-   --{*For @{text rec_def} declarations where the first n parameters
-       stay unchanged in the recursive call. *}
-
-lemma same_fstI [intro!]:
-     "[| P x; (y',y) : R x |] ==> ((x,y'),(x,y)) : same_fst P R"
-by (simp add: same_fst_def)
-
-lemma wf_same_fst:
-  assumes prem: "(!!x. P x ==> wf(R x))"
-  shows "wf(same_fst P R)"
-apply (simp cong del: imp_cong add: wf_def same_fst_def)
-apply (intro strip)
-apply (rename_tac a b)
-apply (case_tac "wf (R a)")
- apply (erule_tac a = b in wf_induct, blast)
-apply (blast intro: prem)
-done
-
 
 subsection{*Weakly decreasing sequences (w.r.t. some well-founded order) 
    stabilize.*}
--- a/src/HOL/ex/Meson_Test.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/ex/Meson_Test.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -5,6 +5,12 @@
 imports Main
 begin
 
+ML {*
+  val Goal = OldGoals.Goal;
+  val by = OldGoals.by;
+  val gethyps = OldGoals.gethyps;
+*}
+
 text {*
   WARNING: there are many potential conflicts between variables used
   below and constants declared in HOL!
--- a/src/HOL/ex/ROOT.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/ex/ROOT.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -68,7 +68,9 @@
   "Landau"
 ];
 
-setmp Proofterm.proofs 2 use_thy "Hilbert_Classical";
+Future.shutdown ();
+(setmp_noncritical proofs 2 (setmp_noncritical Multithreading.max_threads 1 use_thy))
+  "Hilbert_Classical";
 
 
 use_thy "SVC_Oracle";
--- a/src/HOL/ex/SAT_Examples.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOL/ex/SAT_Examples.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -83,7 +83,7 @@
 ML {* reset quick_and_dirty; *}
 
 method_setup rawsat =
-  {* Scan.succeed (K (SIMPLE_METHOD' sat.rawsat_tac)) *}
+  {* Scan.succeed (SIMPLE_METHOD' o sat.rawsat_tac) *}
   "SAT solver (no preprocessing)"
 
 (* ML {* Toplevel.profiling := 1; *} *)
--- a/src/HOLCF/IOA/Modelcheck/MuIOA.thy	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOLCF/IOA/Modelcheck/MuIOA.thy	Wed Jul 29 09:06:49 2009 +0200
@@ -246,7 +246,7 @@
 in
 (
 OldGoals.push_proof();
-Goal 
+OldGoals.Goal 
 ( "(Internal_of_A = (% a::(" ^ action_type_str ^ "). a:(" ^ ia_str^
   "))) --> (Internal_of_C = (% a::(" ^ action_type_str ^ "). a:(" ^ ic_str ^ 
   "))) --> (Start_of_A = (% (" ^ abs_pre_tupel_typed ^
@@ -275,21 +275,21 @@
  ") --> (! " ^ con_pre_str_typed ^ ". ! " ^ abs_pre_str_typed ^
 	". (Start_of_C (" ^ con_pre_tupel ^ ") & Start_of_A (" ^ abs_pre_tupel ^
 	")) --> isSimCA(" ^ con_pre_tupel ^ "," ^ abs_pre_tupel ^ "))");
-by (REPEAT (rtac impI 1));
-by (REPEAT (dtac eq_reflection 1));
+OldGoals.by (REPEAT (rtac impI 1));
+OldGoals.by (REPEAT (dtac eq_reflection 1));
 (* Bis hierher wird im Kapitel 4 erl"autert, ab hier im Kapitel 5 *)
-by (full_simp_tac ((global_simpset_of sign delcongs (if_weak_cong :: weak_case_congs)
+OldGoals.by (full_simp_tac ((global_simpset_of sign delcongs (if_weak_cong :: weak_case_congs)
                               delsimps [not_iff,split_part])
 			      addsimps (thl @ comp_simps @ restrict_simps @ hide_simps @
                                         rename_simps @ ioa_simps @ asig_simps)) 1);
-by (full_simp_tac
+OldGoals.by (full_simp_tac
   (Mucke_ss delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1);
-by (REPEAT (if_full_simp_tac
+OldGoals.by (REPEAT (if_full_simp_tac
   (global_simpset_of sign delcongs (if_weak_cong :: weak_case_congs) delsimps [not_iff,split_part]) 1));
-by (call_mucke_tac 1);
+OldGoals.by (call_mucke_tac 1);
 (* Bis hierher wird im Kapitel 5 erl"autert, ab hier wieder im Kapitel 4 *)
-by (atac 1);
-result();
+OldGoals.by (atac 1);
+OldGoals.result();
 OldGoals.pop_proof();
 Thm.cterm_of sign (Logic.strip_imp_concl subgoal)
 )
--- a/src/HOLCF/Tools/Domain/domain_theorems.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/HOLCF/Tools/Domain/domain_theorems.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -1033,7 +1033,7 @@
     in pg [] goal (K tacs) end;
 end; (* local *)
 
-val inducts = ProjectRule.projections (ProofContext.init thy) ind;
+val inducts = Project_Rule.projections (ProofContext.init thy) ind;
 fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]);
 val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI);
 
--- a/src/Provers/Arith/fast_lin_arith.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Provers/Arith/fast_lin_arith.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -788,7 +788,7 @@
           all_tac) THEN
           PRIMITIVE (trace_thm ctxt "State after neqE:") THEN
           (* use theorems generated from the actual justifications *)
-          METAHYPS (fn asms => rtac (mkthm ss asms j) 1) i
+          FOCUS (fn {prems, ...} => rtac (mkthm ss prems j) 1) ctxt i
     in
       (* rewrite "[| A1; ...; An |] ==> B" to "[| A1; ...; An; ~B |] ==> False" *)
       DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i) THEN
--- a/src/Provers/blast.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Provers/blast.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -39,7 +39,8 @@
 type netpair = (int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net;
 
 signature BLAST_DATA =
-  sig
+sig
+  val thy: theory
   type claset
   val equality_name: string
   val not_name: string
@@ -57,11 +58,10 @@
                  haz_netpair: netpair, dup_netpair: netpair, xtra_netpair: ContextRules.netpair}
   val cla_modifiers: Method.modifier parser list
   val cla_meth': (claset -> int -> tactic) -> thm list -> Proof.context -> Proof.method
-  end;
-
+end;
 
 signature BLAST =
-  sig
+sig
   type claset
   exception TRANS of string    (*reports translation errors*)
   datatype term =
@@ -90,10 +90,10 @@
   val tryInThy          : theory -> claset -> int -> string ->
                   (int->tactic) list * branch list list * (int*int*exn) list
   val normBr            : branch -> branch
-  end;
+end;
 
 
-functor BlastFun(Data: BLAST_DATA) : BLAST =
+functor Blast(Data: BLAST_DATA) : BLAST =
 struct
 
 type claset = Data.claset;
@@ -181,8 +181,8 @@
 fun isGoal (Const ("*Goal*", _) $ _) = true
   | isGoal _ = false;
 
-val TruepropC = ObjectLogic.judgment_name (the_context ());
-val TruepropT = Sign.the_const_type (the_context ()) TruepropC;
+val TruepropC = ObjectLogic.judgment_name Data.thy;
+val TruepropT = Sign.the_const_type Data.thy TruepropC;
 
 fun mk_Trueprop t = Term.$ (Term.Const (TruepropC, TruepropT), t);
 
--- a/src/Provers/order.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Provers/order.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -35,9 +35,11 @@
 
   (* Tactics *)
   val partial_tac:
-    (theory -> term -> (term * string * term) option) -> less_arith -> thm list -> int -> tactic
+    (theory -> term -> (term * string * term) option) -> less_arith ->
+    Proof.context -> thm list -> int -> tactic
   val linear_tac:
-    (theory -> term -> (term * string * term) option) -> less_arith -> thm list -> int -> tactic
+    (theory -> term -> (term * string * term) option) -> less_arith ->
+    Proof.context -> thm list -> int -> tactic
 end;
 
 structure Order_Tac: ORDER_TAC =
@@ -259,34 +261,28 @@
       less_le_trans, le_less_trans, le_trans, le_neq_trans, neq_le_trans,
       not_sym, not_lessD, not_leD, not_lessI, not_leI, less_imp_neq, thm));
 
-
-(* Extract subgoal with signature *)
-fun SUBGOAL goalfun i st =
-  goalfun (List.nth(prems_of st, i-1),  i, Thm.theory_of_thm st) st
-                             handle Subscript => Seq.empty;
-
 (* Internal datatype for the proof *)
 datatype proof
-  = Asm of int 
-  | Thm of proof list * thm; 
-  
+  = Asm of int
+  | Thm of proof list * thm;
+
 exception Cannot;
  (* Internal exception, raised if conclusion cannot be derived from
      assumptions. *)
 exception Contr of proof;
   (* Internal exception, raised if contradiction ( x < x ) was derived *)
 
-fun prove asms = 
+fun prove asms =
   let fun pr (Asm i) = List.nth (asms, i)
   |       pr (Thm (prfs, thm)) = (map pr prfs) MRS thm
   in pr end;
 
 (* Internal datatype for inequalities *)
-datatype less 
-   = Less  of term * term * proof 
+datatype less
+   = Less  of term * term * proof
    | Le    of term * term * proof
-   | NotEq of term * term * proof; 
-   
+   | NotEq of term * term * proof;
+
 (* Misc functions for datatype less *)
 fun lower (Less (x, _, _)) = x
   | lower (Le (x, _, _)) = x
@@ -314,17 +310,17 @@
 fun mkasm_partial decomp (less_thms : less_arith) sign (t, n) =
   case decomp sign t of
     SOME (x, rel, y) => (case rel of
-      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], #less_reflE less_thms)) 
+      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], #less_reflE less_thms))
                else [Less (x, y, Asm n)]
     | "~<"  => []
     | "<="  => [Le (x, y, Asm n)]
-    | "~<=" => [] 
+    | "~<=" => []
     | "="   => [Le (x, y, Thm ([Asm n], #eqD1 less_thms)),
                 Le (y, x, Thm ([Asm n], #eqD2 less_thms))]
-    | "~="  => if (x aconv y) then 
+    | "~="  => if (x aconv y) then
                   raise Contr (Thm ([(Thm ([(Thm ([], #le_refl less_thms)) ,(Asm n)], #le_neq_trans less_thms))], #less_reflE less_thms))
                else [ NotEq (x, y, Asm n),
-                      NotEq (y, x,Thm ( [Asm n], #not_sym less_thms))] 
+                      NotEq (y, x,Thm ( [Asm n], #not_sym less_thms))]
     | _     => error ("partial_tac: unknown relation symbol ``" ^ rel ^
                  "''returned by decomp."))
   | NONE => [];
@@ -338,23 +334,23 @@
 (* Linear orders only.                                                      *)
 (*                                                                          *)
 (* ************************************************************************ *)
- 
+
 fun mkasm_linear decomp (less_thms : less_arith) sign (t, n) =
   case decomp sign t of
     SOME (x, rel, y) => (case rel of
-      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], #less_reflE less_thms)) 
+      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], #less_reflE less_thms))
                else [Less (x, y, Asm n)]
     | "~<"  => [Le (y, x, Thm ([Asm n], #not_lessD less_thms))]
     | "<="  => [Le (x, y, Asm n)]
-    | "~<=" => if (x aconv y) then 
-                  raise (Contr (Thm ([Thm ([Asm n], #not_leD less_thms)], #less_reflE less_thms))) 
-               else [Less (y, x, Thm ([Asm n], #not_leD less_thms))] 
+    | "~<=" => if (x aconv y) then
+                  raise (Contr (Thm ([Thm ([Asm n], #not_leD less_thms)], #less_reflE less_thms)))
+               else [Less (y, x, Thm ([Asm n], #not_leD less_thms))]
     | "="   => [Le (x, y, Thm ([Asm n], #eqD1 less_thms)),
                 Le (y, x, Thm ([Asm n], #eqD2 less_thms))]
-    | "~="  => if (x aconv y) then 
+    | "~="  => if (x aconv y) then
                   raise Contr (Thm ([(Thm ([(Thm ([], #le_refl less_thms)) ,(Asm n)], #le_neq_trans less_thms))], #less_reflE less_thms))
                else [ NotEq (x, y, Asm n),
-                      NotEq (y, x,Thm ( [Asm n], #not_sym less_thms))] 
+                      NotEq (y, x,Thm ( [Asm n], #not_sym less_thms))]
     | _     => error ("linear_tac: unknown relation symbol ``" ^ rel ^
                  "''returned by decomp."))
   | NONE => [];
@@ -409,7 +405,7 @@
      where Rel is one of "<", "<=", "~<", "~<=", "=" and "~=",
      other relation symbols cause an error message *)
 
-fun gen_order_tac mkasm mkconcl decomp' (less_thms : less_arith) prems =
+fun gen_order_tac mkasm mkconcl decomp' (less_thms : less_arith) ctxt prems =
 
 let
 
@@ -440,19 +436,19 @@
 |   mergeLess (Le (x, _, p) , Less (_ , z, q)) =
       Less (x, z, Thm ([p,q] , #le_less_trans less_thms))
 |   mergeLess (Le (x, z, p) , NotEq (x', z', q)) =
-      if (x aconv x' andalso z aconv z' ) 
+      if (x aconv x' andalso z aconv z' )
       then Less (x, z, Thm ([p,q] , #le_neq_trans less_thms))
       else error "linear/partial_tac: internal error le_neq_trans"
 |   mergeLess (NotEq (x, z, p) , Le (x' , z', q)) =
-      if (x aconv x' andalso z aconv z') 
+      if (x aconv x' andalso z aconv z')
       then Less (x, z, Thm ([p,q] , #neq_le_trans less_thms))
       else error "linear/partial_tac: internal error neq_le_trans"
 |   mergeLess (NotEq (x, z, p) , Less (x' , z', q)) =
-      if (x aconv x' andalso z aconv z') 
+      if (x aconv x' andalso z aconv z')
       then Less ((x' , z', q))
       else error "linear/partial_tac: internal error neq_less_trans"
 |   mergeLess (Less (x, z, p) , NotEq (x', z', q)) =
-      if (x aconv x' andalso z aconv z') 
+      if (x aconv x' andalso z aconv z')
       then Less (x, z, p)
       else error "linear/partial_tac: internal error less_neq_trans"
 |   mergeLess (Le (x, _, p) , Le (_ , z, q)) =
@@ -471,8 +467,8 @@
   | (Less (_, y, _)) tr (Less (x', _, _)) = ( y aconv x' )
   | (Le (_, y, _))   tr (Le (x', _, _))   = ( y aconv x' )
   | _ tr _ = false;
-  
-  
+
+
 (* ******************************************************************* *)
 (*                                                                     *)
 (* transPath (Lesslist, Less): (less list * less) -> less              *)
@@ -488,7 +484,7 @@
 |   transPath (x::xs,lesss) =
       if lesss tr x then transPath (xs, mergeLess(lesss,x))
       else error "linear/partial_tac: internal error transpath";
-  
+
 (* ******************************************************************* *)
 (*                                                                     *)
 (* less1 subsumes less2 : less -> less -> bool                         *)
@@ -496,7 +492,7 @@
 (* subsumes checks whether less1 implies less2                         *)
 (*                                                                     *)
 (* ******************************************************************* *)
-  
+
 infix subsumes;
 fun (Less (x, y, _)) subsumes (Le (x', y', _)) =
       (x aconv x' andalso y aconv y')
@@ -554,7 +550,7 @@
 (*                                                                     *)
 (* ******************************************************************* *)
 
-   
+
 (* *********************************************************** *)
 (* Functions for constructing graphs                           *)
 (* *********************************************************** *)
@@ -562,27 +558,27 @@
 fun addEdge (v,d,[]) = [(v,d)]
 |   addEdge (v,d,((u,dl)::el)) = if v aconv u then ((v,d@dl)::el)
     else (u,dl):: (addEdge(v,d,el));
-    
+
 (* ********************************************************************* *)
 (*                                                                       *)
-(* mkGraphs constructs from a list of objects of type less a graph g,    *) 
+(* mkGraphs constructs from a list of objects of type less a graph g,    *)
 (* by taking all edges that are candidate for a <=, and a list neqE, by  *)
 (* taking all edges that are candiate for a ~=                           *)
 (*                                                                       *)
 (* ********************************************************************* *)
 
 fun mkGraphs [] = ([],[],[])
-|   mkGraphs lessList = 
+|   mkGraphs lessList =
  let
 
 fun buildGraphs ([],leqG,neqG,neqE) = (leqG, neqG, neqE)
-|   buildGraphs (l::ls, leqG,neqG, neqE) = case l of 
-  (Less (x,y,p)) =>    
-       buildGraphs (ls, addEdge (x,[(y,(Less (x, y, p)))],leqG) , 
-                        addEdge (x,[(y,(Less (x, y, p)))],neqG), l::neqE) 
+|   buildGraphs (l::ls, leqG,neqG, neqE) = case l of
+  (Less (x,y,p)) =>
+       buildGraphs (ls, addEdge (x,[(y,(Less (x, y, p)))],leqG) ,
+                        addEdge (x,[(y,(Less (x, y, p)))],neqG), l::neqE)
 | (Le (x,y,p)) =>
-      buildGraphs (ls, addEdge (x,[(y,(Le (x, y,p)))],leqG) , neqG, neqE) 
-| (NotEq  (x,y,p)) => 
+      buildGraphs (ls, addEdge (x,[(y,(Le (x, y,p)))],leqG) , neqG, neqE)
+| (NotEq  (x,y,p)) =>
       buildGraphs (ls, leqG , addEdge (x,[(y,(NotEq (x, y, p)))],neqG), l::neqE) ;
 
   in buildGraphs (lessList, [], [], []) end;
@@ -595,12 +591,12 @@
 (* List of successors of u in graph g                                      *)
 (*                                                                         *)
 (* *********************************************************************** *)
- 
-fun adjacent eq_comp ((v,adj)::el) u = 
+
+fun adjacent eq_comp ((v,adj)::el) u =
     if eq_comp (u, v) then adj else adjacent eq_comp el u
-|   adjacent _  []  _ = []  
-  
-     
+|   adjacent _  []  _ = []
+
+
 (* *********************************************************************** *)
 (*                                                                         *)
 (* transpose g:                                                            *)
@@ -616,7 +612,7 @@
    (* Compute list of reversed edges for each adjacency list *)
    fun flip (u,(v,l)::el) = (v,(u,l)) :: flip (u,el)
      | flip (_,nil) = nil
-   
+
    (* Compute adjacency list for node u from the list of edges
       and return a likewise reduced list of edges.  The list of edges
       is searches for edges starting from u, and these edges are removed. *)
@@ -640,19 +636,19 @@
    (* Compute, for each adjacency list, the list with reversed edges,
       and concatenate these lists. *)
    val flipped = List.foldr (op @) nil (map flip g)
- 
- in assemble g flipped end    
-      
+
+ in assemble g flipped end
+
 (* *********************************************************************** *)
-(*                                                                         *)      
+(*                                                                         *)
 (* scc_term : (term * term list) list -> term list list                    *)
 (*                                                                         *)
 (* The following is based on the algorithm for finding strongly connected  *)
 (* components described in Introduction to Algorithms, by Cormon, Leiserson*)
 (* and Rivest, section 23.5. The input G is an adjacency list description  *)
 (* of a directed graph. The output is a list of the strongly connected     *)
-(* components (each a list of vertices).                                   *)          
-(*                                                                         *)   
+(* components (each a list of vertices).                                   *)
+(*                                                                         *)
 (*                                                                         *)
 (* *********************************************************************** *)
 
@@ -664,11 +660,11 @@
 
   (* List of vertices which have been visited. *)
   val visited : term list ref = ref nil
-  
+
   fun been_visited v = exists (fn w => w aconv v) (!visited)
-  
+
   (* Given the adjacency list rep of a graph (a list of pairs),
-     return just the first element of each pair, yielding the 
+     return just the first element of each pair, yielding the
      vertex list. *)
   val members = map (fn (v,_) => v)
 
@@ -701,7 +697,7 @@
      dfs_visit along with u form a connected component. We
      collect all the connected components together in a
      list, which is what is returned. *)
-  Library.foldl (fn (comps,u) =>  
+  Library.foldl (fn (comps,u) =>
       if been_visited u then comps
       else ((u :: dfs_visit (transpose (op aconv) G) u) :: comps))  (nil,(!finish))
 end;
@@ -710,17 +706,17 @@
 (* *********************************************************************** *)
 (*                                                                         *)
 (* dfs_int_reachable g u:                                                  *)
-(* (int * int list) list -> int -> int list                                *) 
+(* (int * int list) list -> int -> int list                                *)
 (*                                                                         *)
 (* Computes list of all nodes reachable from u in g.                       *)
 (*                                                                         *)
 (* *********************************************************************** *)
 
-fun dfs_int_reachable g u = 
+fun dfs_int_reachable g u =
  let
   (* List of vertices which have been visited. *)
   val visited : int list ref = ref nil
-  
+
   fun been_visited v = exists (fn w => w = v) (!visited)
 
   fun dfs_visit g u : int list =
@@ -731,59 +727,59 @@
             else v :: dfs_visit g v @ ds)
         nil (adjacent (op =) g u)
    in  descendents end
- 
+
  in u :: dfs_visit g u end;
 
-    
-fun indexComps components = 
+
+fun indexComps components =
     ListPair.map (fn (a,b) => (a,b)) (0 upto (length components -1), components);
 
-fun indexNodes IndexComp = 
+fun indexNodes IndexComp =
     List.concat (map (fn (index, comp) => (map (fn v => (v,index)) comp)) IndexComp);
-    
+
 fun getIndex v [] = ~1
-|   getIndex v ((v',k)::vs) = if v aconv v' then k else getIndex v vs; 
-    
+|   getIndex v ((v',k)::vs) = if v aconv v' then k else getIndex v vs;
+
 
 
 (* *********************************************************************** *)
 (*                                                                         *)
 (* dfs eq_comp g u v:                                                       *)
 (* ('a * 'a -> bool) -> ('a  *( 'a * less) list) list ->                   *)
-(* 'a -> 'a -> (bool * ('a * less) list)                                   *) 
+(* 'a -> 'a -> (bool * ('a * less) list)                                   *)
 (*                                                                         *)
 (* Depth first search of v from u.                                         *)
 (* Returns (true, path(u, v)) if successful, otherwise (false, []).        *)
 (*                                                                         *)
 (* *********************************************************************** *)
 
-fun dfs eq_comp g u v = 
- let 
+fun dfs eq_comp g u v =
+ let
     val pred = ref nil;
     val visited = ref nil;
-    
+
     fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
-    
-    fun dfs_visit u' = 
+
+    fun dfs_visit u' =
     let val _ = visited := u' :: (!visited)
-    
+
     fun update (x,l) = let val _ = pred := (x,l) ::(!pred) in () end;
-    
-    in if been_visited v then () 
+
+    in if been_visited v then ()
     else (app (fn (v',l) => if been_visited v' then () else (
-       update (v',l); 
+       update (v',l);
        dfs_visit v'; ()) )) (adjacent eq_comp g u')
      end
-  in 
-    dfs_visit u; 
-    if (been_visited v) then (true, (!pred)) else (false , [])   
+  in
+    dfs_visit u;
+    if (been_visited v) then (true, (!pred)) else (false , [])
   end;
 
-  
+
 (* *********************************************************************** *)
 (*                                                                         *)
 (* completeTermPath u v g:                                                 *)
-(* Term.term -> Term.term -> (Term.term * (Term.term * less) list) list    *) 
+(* Term.term -> Term.term -> (Term.term * (Term.term * less) list) list    *)
 (* -> less list                                                            *)
 (*                                                                         *)
 (* Complete the path from u to v in graph g.  Path search is performed     *)
@@ -793,36 +789,36 @@
 (*                                                                         *)
 (* *********************************************************************** *)
 
-  
-fun completeTermPath u v g  = 
-  let 
+
+fun completeTermPath u v g  =
+  let
    val (found, tmp) =  dfs (op aconv) g u v ;
    val pred = map snd tmp;
-   
+
    fun path x y  =
       let
- 
+
       (* find predecessor u of node v and the edge u -> v *)
 
       fun lookup v [] = raise Cannot
       |   lookup v (e::es) = if (upper e) aconv v then e else lookup v es;
 
-      (* traverse path backwards and return list of visited edges *)   
-      fun rev_path v = 
+      (* traverse path backwards and return list of visited edges *)
+      fun rev_path v =
        let val l = lookup v pred
            val u = lower l;
        in
         if u aconv x then [l]
-        else (rev_path u) @ [l] 
+        else (rev_path u) @ [l]
        end
      in rev_path y end;
-       
-  in 
+
+  in
   if found then (if u aconv v then [(Le (u, v, (Thm ([], #le_refl less_thms))))]
   else path u v ) else raise Cannot
-end;  
+end;
 
-      
+
 (* *********************************************************************** *)
 (*                                                                         *)
 (* findProof (sccGraph, neqE, ntc, sccSubgraphs) subgoal:                  *)
@@ -834,133 +830,133 @@
 (* proof for subgoal.  Raises exception Cannot if this is not possible.    *)
 (*                                                                         *)
 (* *********************************************************************** *)
-     
+
 fun findProof (sccGraph, neqE, ntc, sccSubgraphs) subgoal =
 let
-   
+
  (* complete path x y from component graph *)
- fun completeComponentPath x y predlist = 
-   let         
-	  val xi = getIndex x ntc
-	  val yi = getIndex y ntc 
-	  
-	  fun lookup k [] =  raise Cannot
-	  |   lookup k ((h: int,l)::us) = if k = h then l else lookup k us;	  
-	  
-	  fun rev_completeComponentPath y' = 
-	   let val edge = lookup (getIndex y' ntc) predlist
-	       val u = lower edge
-	       val v = upper edge
-	   in
-             if (getIndex u ntc) = xi then 
-	       (completeTermPath x u (List.nth(sccSubgraphs, xi)) )@[edge]
-	       @(completeTermPath v y' (List.nth(sccSubgraphs, getIndex y' ntc)) )
-	     else (rev_completeComponentPath u)@[edge]
-	          @(completeTermPath v y' (List.nth(sccSubgraphs, getIndex y' ntc)) )
+ fun completeComponentPath x y predlist =
+   let
+          val xi = getIndex x ntc
+          val yi = getIndex y ntc
+
+          fun lookup k [] =  raise Cannot
+          |   lookup k ((h: int,l)::us) = if k = h then l else lookup k us;
+
+          fun rev_completeComponentPath y' =
+           let val edge = lookup (getIndex y' ntc) predlist
+               val u = lower edge
+               val v = upper edge
+           in
+             if (getIndex u ntc) = xi then
+               (completeTermPath x u (List.nth(sccSubgraphs, xi)) )@[edge]
+               @(completeTermPath v y' (List.nth(sccSubgraphs, getIndex y' ntc)) )
+             else (rev_completeComponentPath u)@[edge]
+                  @(completeTermPath v y' (List.nth(sccSubgraphs, getIndex y' ntc)) )
            end
-   in  
-      if x aconv y then 
+   in
+      if x aconv y then
         [(Le (x, y, (Thm ([], #le_refl less_thms))))]
       else ( if xi = yi then completeTermPath x y (List.nth(sccSubgraphs, xi))
-             else rev_completeComponentPath y )  
+             else rev_completeComponentPath y )
    end;
 
-(* ******************************************************************* *) 
+(* ******************************************************************* *)
 (* findLess e x y xi yi xreachable yreachable                          *)
 (*                                                                     *)
 (* Find a path from x through e to y, of weight <                      *)
-(* ******************************************************************* *) 
- 
- fun findLess e x y xi yi xreachable yreachable = 
-  let val u = lower e 
+(* ******************************************************************* *)
+
+ fun findLess e x y xi yi xreachable yreachable =
+  let val u = lower e
       val v = upper e
       val ui = getIndex u ntc
       val vi = getIndex v ntc
-            
-  in 
-      if ui mem xreachable andalso vi mem xreachable andalso 
+
+  in
+      if ui mem xreachable andalso vi mem xreachable andalso
          ui mem yreachable andalso vi mem yreachable then (
-       
-  (case e of (Less (_, _, _)) =>  
+
+  (case e of (Less (_, _, _)) =>
        let
         val (xufound, xupred) =  dfs (op =) sccGraph xi (getIndex u ntc)
-	    in 
-	     if xufound then (
-	      let 
-	       val (vyfound, vypred) =  dfs (op =) sccGraph (getIndex v ntc) yi  
-	      in 
-	       if vyfound then (
-	        let 
-	         val xypath = (completeComponentPath x u xupred)@[e]@(completeComponentPath v y vypred)
-	         val xyLesss = transPath (tl xypath, hd xypath)
-	        in 
-		 if xyLesss subsumes subgoal then SOME (getprf xyLesss) 
+            in
+             if xufound then (
+              let
+               val (vyfound, vypred) =  dfs (op =) sccGraph (getIndex v ntc) yi
+              in
+               if vyfound then (
+                let
+                 val xypath = (completeComponentPath x u xupred)@[e]@(completeComponentPath v y vypred)
+                 val xyLesss = transPath (tl xypath, hd xypath)
+                in
+                 if xyLesss subsumes subgoal then SOME (getprf xyLesss)
                  else NONE
-	       end)
-	       else NONE
-	      end)
-	     else NONE
-	    end
-       |  _   => 
-         let val (uvfound, uvpred) =  dfs (op =) sccGraph (getIndex u ntc) (getIndex v ntc) 
-             in 
-	      if uvfound then (
-	       let 
-	        val (xufound, xupred) = dfs (op =) sccGraph xi (getIndex u ntc)
-	       in
-		if xufound then (
-		 let 
-		  val (vyfound, vypred) =  dfs (op =) sccGraph (getIndex v ntc) yi
-		 in 
-		  if vyfound then (
-		   let
-		    val uvpath = completeComponentPath u v uvpred
-		    val uvLesss = mergeLess ( transPath (tl uvpath, hd uvpath), e)
-		    val xypath = (completeComponentPath  x u xupred)@[uvLesss]@(completeComponentPath v y vypred)
-		    val xyLesss = transPath (tl xypath, hd xypath)
-		   in 
-		    if xyLesss subsumes subgoal then SOME (getprf xyLesss)
+               end)
+               else NONE
+              end)
+             else NONE
+            end
+       |  _   =>
+         let val (uvfound, uvpred) =  dfs (op =) sccGraph (getIndex u ntc) (getIndex v ntc)
+             in
+              if uvfound then (
+               let
+                val (xufound, xupred) = dfs (op =) sccGraph xi (getIndex u ntc)
+               in
+                if xufound then (
+                 let
+                  val (vyfound, vypred) =  dfs (op =) sccGraph (getIndex v ntc) yi
+                 in
+                  if vyfound then (
+                   let
+                    val uvpath = completeComponentPath u v uvpred
+                    val uvLesss = mergeLess ( transPath (tl uvpath, hd uvpath), e)
+                    val xypath = (completeComponentPath  x u xupred)@[uvLesss]@(completeComponentPath v y vypred)
+                    val xyLesss = transPath (tl xypath, hd xypath)
+                   in
+                    if xyLesss subsumes subgoal then SOME (getprf xyLesss)
                     else NONE
-		   end )
-		  else NONE   
-	         end)
-		else NONE
-	       end ) 
-	      else NONE
-	     end )
+                   end )
+                  else NONE
+                 end)
+                else NONE
+               end )
+              else NONE
+             end )
     ) else NONE
-end;  
-   
-         
+end;
+
+
 in
   (* looking for x <= y: any path from x to y is sufficient *)
   case subgoal of (Le (x, y, _)) => (
-  if null sccGraph then raise Cannot else ( 
-   let 
+  if null sccGraph then raise Cannot else (
+   let
     val xi = getIndex x ntc
     val yi = getIndex y ntc
     (* searches in sccGraph a path from xi to yi *)
     val (found, pred) = dfs (op =) sccGraph xi yi
-   in 
+   in
     if found then (
-       let val xypath = completeComponentPath x y pred 
-           val xyLesss = transPath (tl xypath, hd xypath) 
-       in  
-	  (case xyLesss of
-	    (Less (_, _, q)) => if xyLesss subsumes subgoal then (Thm ([q], #less_imp_le less_thms))  
-				else raise Cannot
-	     | _   => if xyLesss subsumes subgoal then (getprf xyLesss) 
-	              else raise Cannot)
+       let val xypath = completeComponentPath x y pred
+           val xyLesss = transPath (tl xypath, hd xypath)
+       in
+          (case xyLesss of
+            (Less (_, _, q)) => if xyLesss subsumes subgoal then (Thm ([q], #less_imp_le less_thms))
+                                else raise Cannot
+             | _   => if xyLesss subsumes subgoal then (getprf xyLesss)
+                      else raise Cannot)
        end )
-     else raise Cannot 
-   end 
+     else raise Cannot
+   end
     )
    )
  (* looking for x < y: particular path required, which is not necessarily
     found by normal dfs *)
  |   (Less (x, y, _)) => (
   if null sccGraph then raise Cannot else (
-   let 
+   let
     val xi = getIndex x ntc
     val yi = getIndex y ntc
     val sccGraph_transpose = transpose (op =) sccGraph
@@ -969,13 +965,13 @@
     (* all comonents reachable from y in the transposed graph sccGraph' *)
     val yreachable = dfs_int_reachable sccGraph_transpose yi
     (* for all edges u ~= v or u < v check if they are part of path x < y *)
-    fun processNeqEdges [] = raise Cannot 
-    |   processNeqEdges (e::es) = 
-      case  (findLess e x y xi yi xreachable yreachable) of (SOME prf) => prf  
+    fun processNeqEdges [] = raise Cannot
+    |   processNeqEdges (e::es) =
+      case  (findLess e x y xi yi xreachable yreachable) of (SOME prf) => prf
       | _ => processNeqEdges es
-        
-    in 
-       processNeqEdges neqE 
+
+    in
+       processNeqEdges neqE
     end
   )
  )
@@ -986,99 +982,99 @@
   else if null sccSubgraphs then (
      (case (Library.find_first (fn fact => fact subsumes subgoal) neqE, subgoal) of
        ( SOME (NotEq (x, y, p)), NotEq (x', y', _)) =>
-          if  (x aconv x' andalso y aconv y') then p 
-	  else Thm ([p], #not_sym less_thms)
-     | ( SOME (Less (x, y, p)), NotEq (x', y', _)) => 
+          if  (x aconv x' andalso y aconv y') then p
+          else Thm ([p], #not_sym less_thms)
+     | ( SOME (Less (x, y, p)), NotEq (x', y', _)) =>
           if x aconv x' andalso y aconv y' then (Thm ([p], #less_imp_neq less_thms))
           else (Thm ([(Thm ([p], #less_imp_neq less_thms))], #not_sym less_thms))
-     | _ => raise Cannot) 
+     | _ => raise Cannot)
    ) else (
-   
-   let  val xi = getIndex x ntc 
+
+   let  val xi = getIndex x ntc
         val yi = getIndex y ntc
-	val sccGraph_transpose = transpose (op =) sccGraph
+        val sccGraph_transpose = transpose (op =) sccGraph
         val xreachable = dfs_int_reachable sccGraph xi
-	val yreachable = dfs_int_reachable sccGraph_transpose yi
-	
-	fun processNeqEdges [] = raise Cannot  
-  	|   processNeqEdges (e::es) = (
-	    let val u = lower e 
-	        val v = upper e
-		val ui = getIndex u ntc
-		val vi = getIndex v ntc
-		
-	    in  
-	        (* if x ~= y follows from edge e *)
-	    	if e subsumes subgoal then (
-		     case e of (Less (u, v, q)) => (
-		       if u aconv x andalso v aconv y then (Thm ([q], #less_imp_neq less_thms))
-		       else (Thm ([(Thm ([q], #less_imp_neq less_thms))], #not_sym less_thms))
-		     )
-		     |    (NotEq (u,v, q)) => (
-		       if u aconv x andalso v aconv y then q
-		       else (Thm ([q],  #not_sym less_thms))
-		     )
-		 )
+        val yreachable = dfs_int_reachable sccGraph_transpose yi
+
+        fun processNeqEdges [] = raise Cannot
+        |   processNeqEdges (e::es) = (
+            let val u = lower e
+                val v = upper e
+                val ui = getIndex u ntc
+                val vi = getIndex v ntc
+
+            in
+                (* if x ~= y follows from edge e *)
+                if e subsumes subgoal then (
+                     case e of (Less (u, v, q)) => (
+                       if u aconv x andalso v aconv y then (Thm ([q], #less_imp_neq less_thms))
+                       else (Thm ([(Thm ([q], #less_imp_neq less_thms))], #not_sym less_thms))
+                     )
+                     |    (NotEq (u,v, q)) => (
+                       if u aconv x andalso v aconv y then q
+                       else (Thm ([q],  #not_sym less_thms))
+                     )
+                 )
                 (* if SCC_x is linked to SCC_y via edge e *)
-		 else if ui = xi andalso vi = yi then (
+                 else if ui = xi andalso vi = yi then (
                    case e of (Less (_, _,_)) => (
-		        let val xypath = (completeTermPath x u (List.nth(sccSubgraphs, ui)) ) @ [e] @ (completeTermPath v y (List.nth(sccSubgraphs, vi)) )
-			    val xyLesss = transPath (tl xypath, hd xypath)
-			in  (Thm ([getprf xyLesss], #less_imp_neq less_thms)) end)
-		   | _ => (   
-		        let 
-			    val xupath = completeTermPath x u  (List.nth(sccSubgraphs, ui))
-			    val uxpath = completeTermPath u x  (List.nth(sccSubgraphs, ui))
-			    val vypath = completeTermPath v y  (List.nth(sccSubgraphs, vi))
-			    val yvpath = completeTermPath y v  (List.nth(sccSubgraphs, vi))
-			    val xuLesss = transPath (tl xupath, hd xupath)     
-			    val uxLesss = transPath (tl uxpath, hd uxpath)			    
-			    val vyLesss = transPath (tl vypath, hd vypath)			
-			    val yvLesss = transPath (tl yvpath, hd yvpath)
-			    val x_eq_u =  (Thm ([(getprf xuLesss),(getprf uxLesss)], #eqI less_thms))
-			    val v_eq_y =  (Thm ([(getprf vyLesss),(getprf yvLesss)], #eqI less_thms))
-			in 
-                           (Thm ([x_eq_u , (getprf e), v_eq_y ], #eq_neq_eq_imp_neq less_thms)) 
-			end
-			)       
-		  ) else if ui = yi andalso vi = xi then (
-		     case e of (Less (_, _,_)) => (
-		        let val xypath = (completeTermPath y u (List.nth(sccSubgraphs, ui)) ) @ [e] @ (completeTermPath v x (List.nth(sccSubgraphs, vi)) )
-			    val xyLesss = transPath (tl xypath, hd xypath)
-			in (Thm ([(Thm ([getprf xyLesss], #less_imp_neq less_thms))] , #not_sym less_thms)) end ) 
-		     | _ => (
-		        
-			let val yupath = completeTermPath y u (List.nth(sccSubgraphs, ui))
-			    val uypath = completeTermPath u y (List.nth(sccSubgraphs, ui))
-			    val vxpath = completeTermPath v x (List.nth(sccSubgraphs, vi))
-			    val xvpath = completeTermPath x v (List.nth(sccSubgraphs, vi))
-			    val yuLesss = transPath (tl yupath, hd yupath)     
-			    val uyLesss = transPath (tl uypath, hd uypath)			    
-			    val vxLesss = transPath (tl vxpath, hd vxpath)			
-			    val xvLesss = transPath (tl xvpath, hd xvpath)
-			    val y_eq_u =  (Thm ([(getprf yuLesss),(getprf uyLesss)], #eqI less_thms))
-			    val v_eq_x =  (Thm ([(getprf vxLesss),(getprf xvLesss)], #eqI less_thms))
-			in
-			    (Thm ([(Thm ([y_eq_u , (getprf e), v_eq_x ], #eq_neq_eq_imp_neq less_thms))], #not_sym less_thms))
-		        end
-		       )
-		  ) else (
+                        let val xypath = (completeTermPath x u (List.nth(sccSubgraphs, ui)) ) @ [e] @ (completeTermPath v y (List.nth(sccSubgraphs, vi)) )
+                            val xyLesss = transPath (tl xypath, hd xypath)
+                        in  (Thm ([getprf xyLesss], #less_imp_neq less_thms)) end)
+                   | _ => (
+                        let
+                            val xupath = completeTermPath x u  (List.nth(sccSubgraphs, ui))
+                            val uxpath = completeTermPath u x  (List.nth(sccSubgraphs, ui))
+                            val vypath = completeTermPath v y  (List.nth(sccSubgraphs, vi))
+                            val yvpath = completeTermPath y v  (List.nth(sccSubgraphs, vi))
+                            val xuLesss = transPath (tl xupath, hd xupath)
+                            val uxLesss = transPath (tl uxpath, hd uxpath)
+                            val vyLesss = transPath (tl vypath, hd vypath)
+                            val yvLesss = transPath (tl yvpath, hd yvpath)
+                            val x_eq_u =  (Thm ([(getprf xuLesss),(getprf uxLesss)], #eqI less_thms))
+                            val v_eq_y =  (Thm ([(getprf vyLesss),(getprf yvLesss)], #eqI less_thms))
+                        in
+                           (Thm ([x_eq_u , (getprf e), v_eq_y ], #eq_neq_eq_imp_neq less_thms))
+                        end
+                        )
+                  ) else if ui = yi andalso vi = xi then (
+                     case e of (Less (_, _,_)) => (
+                        let val xypath = (completeTermPath y u (List.nth(sccSubgraphs, ui)) ) @ [e] @ (completeTermPath v x (List.nth(sccSubgraphs, vi)) )
+                            val xyLesss = transPath (tl xypath, hd xypath)
+                        in (Thm ([(Thm ([getprf xyLesss], #less_imp_neq less_thms))] , #not_sym less_thms)) end )
+                     | _ => (
+
+                        let val yupath = completeTermPath y u (List.nth(sccSubgraphs, ui))
+                            val uypath = completeTermPath u y (List.nth(sccSubgraphs, ui))
+                            val vxpath = completeTermPath v x (List.nth(sccSubgraphs, vi))
+                            val xvpath = completeTermPath x v (List.nth(sccSubgraphs, vi))
+                            val yuLesss = transPath (tl yupath, hd yupath)
+                            val uyLesss = transPath (tl uypath, hd uypath)
+                            val vxLesss = transPath (tl vxpath, hd vxpath)
+                            val xvLesss = transPath (tl xvpath, hd xvpath)
+                            val y_eq_u =  (Thm ([(getprf yuLesss),(getprf uyLesss)], #eqI less_thms))
+                            val v_eq_x =  (Thm ([(getprf vxLesss),(getprf xvLesss)], #eqI less_thms))
+                        in
+                            (Thm ([(Thm ([y_eq_u , (getprf e), v_eq_x ], #eq_neq_eq_imp_neq less_thms))], #not_sym less_thms))
+                        end
+                       )
+                  ) else (
                        (* there exists a path x < y or y < x such that
                           x ~= y may be concluded *)
-	        	case  (findLess e x y xi yi xreachable yreachable) of 
-		              (SOME prf) =>  (Thm ([prf], #less_imp_neq less_thms))  
+                        case  (findLess e x y xi yi xreachable yreachable) of
+                              (SOME prf) =>  (Thm ([prf], #less_imp_neq less_thms))
                              | NONE =>  (
-		               let 
-		                val yr = dfs_int_reachable sccGraph yi
-	                        val xr = dfs_int_reachable sccGraph_transpose xi
-		               in 
-		                case  (findLess e y x yi xi yr xr) of 
-		                      (SOME prf) => (Thm ([(Thm ([prf], #less_imp_neq less_thms))], #not_sym less_thms)) 
+                               let
+                                val yr = dfs_int_reachable sccGraph yi
+                                val xr = dfs_int_reachable sccGraph_transpose xi
+                               in
+                                case  (findLess e y x yi xi yr xr) of
+                                      (SOME prf) => (Thm ([(Thm ([prf], #less_imp_neq less_thms))], #not_sym less_thms))
                                       | _ => processNeqEdges es
-		               end)
-		 ) end) 
+                               end)
+                 ) end)
      in processNeqEdges neqE end)
-  )    
+  )
 end;
 
 
@@ -1110,115 +1106,115 @@
 (* ******************************************************************* *)
 
 fun mk_sccGraphs _ [] _ _ = ([],[])
-|   mk_sccGraphs components leqG neqG ntc = 
+|   mk_sccGraphs components leqG neqG ntc =
     let
     (* Liste (Index der Komponente, Komponente *)
     val IndexComp = indexComps components;
 
-       
-    fun handleContr edge g = 
-       (case edge of 
+
+    fun handleContr edge g =
+       (case edge of
           (Less  (x, y, _)) => (
-	    let 
-	     val xxpath = edge :: (completeTermPath y x g )
-	     val xxLesss = transPath (tl xxpath, hd xxpath)
-	     val q = getprf xxLesss
-	    in 
-	     raise (Contr (Thm ([q], #less_reflE less_thms ))) 
-	    end 
-	  )
+            let
+             val xxpath = edge :: (completeTermPath y x g )
+             val xxLesss = transPath (tl xxpath, hd xxpath)
+             val q = getprf xxLesss
+            in
+             raise (Contr (Thm ([q], #less_reflE less_thms )))
+            end
+          )
         | (NotEq (x, y, _)) => (
-	    let 
-	     val xypath = (completeTermPath x y g )
-	     val yxpath = (completeTermPath y x g )
-	     val xyLesss = transPath (tl xypath, hd xypath)
-	     val yxLesss = transPath (tl yxpath, hd yxpath)
-             val q = getprf (mergeLess ((mergeLess (edge, xyLesss)),yxLesss )) 
-	    in 
-	     raise (Contr (Thm ([q], #less_reflE less_thms )))
-	    end  
-	 )
-	| _ =>  error "trans_tac/handleContr: invalid Contradiction");
- 
-   	
-    (* k is index of the actual component *)   
-       
-    fun processComponent (k, comp) = 
-     let    
-        (* all edges with weight <= of the actual component *)  
+            let
+             val xypath = (completeTermPath x y g )
+             val yxpath = (completeTermPath y x g )
+             val xyLesss = transPath (tl xypath, hd xypath)
+             val yxLesss = transPath (tl yxpath, hd yxpath)
+             val q = getprf (mergeLess ((mergeLess (edge, xyLesss)),yxLesss ))
+            in
+             raise (Contr (Thm ([q], #less_reflE less_thms )))
+            end
+         )
+        | _ =>  error "trans_tac/handleContr: invalid Contradiction");
+
+
+    (* k is index of the actual component *)
+
+    fun processComponent (k, comp) =
+     let
+        (* all edges with weight <= of the actual component *)
         val leqEdges = List.concat (map (adjacent (op aconv) leqG) comp);
-        (* all edges with weight ~= of the actual component *)  
+        (* all edges with weight ~= of the actual component *)
         val neqEdges = map snd (List.concat (map (adjacent (op aconv) neqG) comp));
 
-       (* find an edge leading to a contradiction *)   
+       (* find an edge leading to a contradiction *)
        fun findContr [] = NONE
-       |   findContr (e::es) = 
-		    let val ui = (getIndex (lower e) ntc) 
-			val vi = (getIndex (upper e) ntc) 
-		    in 
-		        if ui = vi then  SOME e
-		        else findContr es
-		    end; 
-		   
-		(* sort edges into component internal edges and 
-		   edges pointing away from the component *)
-		fun sortEdges  [] (intern,extern)  = (intern,extern)
-		|   sortEdges  ((v,l)::es) (intern, extern) = 
-		    let val k' = getIndex v ntc in 
-		        if k' = k then 
-			    sortEdges es (l::intern, extern)
-			else sortEdges  es (intern, (k',l)::extern) end;
-		
-		(* Insert edge into sorted list of edges, where edge is
+       |   findContr (e::es) =
+                    let val ui = (getIndex (lower e) ntc)
+                        val vi = (getIndex (upper e) ntc)
+                    in
+                        if ui = vi then  SOME e
+                        else findContr es
+                    end;
+
+                (* sort edges into component internal edges and
+                   edges pointing away from the component *)
+                fun sortEdges  [] (intern,extern)  = (intern,extern)
+                |   sortEdges  ((v,l)::es) (intern, extern) =
+                    let val k' = getIndex v ntc in
+                        if k' = k then
+                            sortEdges es (l::intern, extern)
+                        else sortEdges  es (intern, (k',l)::extern) end;
+
+                (* Insert edge into sorted list of edges, where edge is
                     only added if
                     - it is found for the first time
                     - it is a <= edge and no parallel < edge was found earlier
                     - it is a < edge
                  *)
-          	 fun insert (h: int,l) [] = [(h,l)]
-		 |   insert (h,l) ((k',l')::es) = if h = k' then (
-		     case l of (Less (_, _, _)) => (h,l)::es
-		     | _  => (case l' of (Less (_, _, _)) => (h,l')::es
-	                      | _ => (k',l)::es) )
-		     else (k',l'):: insert (h,l) es;
-		
-		(* Reorganise list of edges such that
+                 fun insert (h: int,l) [] = [(h,l)]
+                 |   insert (h,l) ((k',l')::es) = if h = k' then (
+                     case l of (Less (_, _, _)) => (h,l)::es
+                     | _  => (case l' of (Less (_, _, _)) => (h,l')::es
+                              | _ => (k',l)::es) )
+                     else (k',l'):: insert (h,l) es;
+
+                (* Reorganise list of edges such that
                     - duplicate edges are removed
                     - if a < edge and a <= edge exist at the same time,
                       remove <= edge *)
-    		 fun reOrganizeEdges [] sorted = sorted: (int * less) list
-		 |   reOrganizeEdges (e::es) sorted = reOrganizeEdges es (insert e sorted); 
-	
+                 fun reOrganizeEdges [] sorted = sorted: (int * less) list
+                 |   reOrganizeEdges (e::es) sorted = reOrganizeEdges es (insert e sorted);
+
                  (* construct the subgraph forming the strongly connected component
-		    from the edge list *)    
-		 fun sccSubGraph [] g  = g
-		 |   sccSubGraph (l::ls) g = 
-		          sccSubGraph ls (addEdge ((lower l),[((upper l),l)],g))
-		 
-		 val (intern, extern) = sortEdges leqEdges ([], []);
+                    from the edge list *)
+                 fun sccSubGraph [] g  = g
+                 |   sccSubGraph (l::ls) g =
+                          sccSubGraph ls (addEdge ((lower l),[((upper l),l)],g))
+
+                 val (intern, extern) = sortEdges leqEdges ([], []);
                  val subGraph = sccSubGraph intern [];
-		  
-     in  
+
+     in
          case findContr neqEdges of SOME e => handleContr e subGraph
-	 | _ => ((k, (reOrganizeEdges (extern) [])), subGraph)
-     end; 
-  
-    val tmp =  map processComponent IndexComp    
-in 
-     ( (map fst tmp), (map snd tmp))  
-end; 
+         | _ => ((k, (reOrganizeEdges (extern) [])), subGraph)
+     end;
+
+    val tmp =  map processComponent IndexComp
+in
+     ( (map fst tmp), (map snd tmp))
+end;
 
 
 (** Find proof if possible. **)
 
 fun gen_solve mkconcl sign (asms, concl) =
- let 
+ let
   val (leqG, neqG, neqE) = mkGraphs asms
-  val components = scc_term leqG   
+  val components = scc_term leqG
   val ntc = indexNodes (indexComps components)
   val (sccGraph, sccSubgraphs) = mk_sccGraphs components leqG neqG ntc
  in
-   let 
+   let
    val (subgoals, prf) = mkconcl decomp less_thms sign concl
    fun solve facts less =
       (case triv_solv less of NONE => findProof (sccGraph, neqE, ntc, sccSubgraphs) less
@@ -1229,24 +1225,25 @@
  end;
 
 in
-
-    SUBGOAL (fn (A, n, sign) =>
-     let
-      val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
-      val Hs = map prop_of prems @ map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
-      val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
-      val lesss = List.concat (ListPair.map (mkasm decomp less_thms sign) (Hs, 0 upto (length Hs - 1)))
-      val prfs = gen_solve mkconcl sign (lesss, C);
-      val (subgoals, prf) = mkconcl decomp less_thms sign C;
-     in
-      METAHYPS (fn asms =>
-	let val thms = map (prove (prems @ asms)) prfs
-	in rtac (prove thms prf) 1 end) n
-     end
-     handle Contr p => METAHYPS (fn asms => rtac (prove asms p) 1) n
-	  | Cannot  => no_tac
-	  )
-
+ SUBGOAL (fn (A, n) => fn st =>
+  let
+   val thy = ProofContext.theory_of ctxt;
+   val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A));
+   val Hs = map prop_of prems @ map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
+   val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
+   val lesss = List.concat (ListPair.map (mkasm decomp less_thms thy) (Hs, 0 upto (length Hs - 1)))
+   val prfs = gen_solve mkconcl thy (lesss, C);
+   val (subgoals, prf) = mkconcl decomp less_thms thy C;
+  in
+   FOCUS (fn {prems = asms, ...} =>
+     let val thms = map (prove (prems @ asms)) prfs
+     in rtac (prove thms prf) 1 end) ctxt n st
+  end
+  handle Contr p =>
+      (FOCUS (fn {prems = asms, ...} => rtac (prove asms p) 1) ctxt n st
+        handle Subscript => Seq.empty)
+   | Cannot => Seq.empty
+   | Subscript => Seq.empty)
 end;
 
 (* partial_tac - solves partial orders *)
@@ -1255,5 +1252,4 @@
 (* linear_tac - solves linear/total orders *)
 val linear_tac = gen_order_tac mkasm_linear mkconcl_linear;
 
-       
 end;
--- a/src/Provers/quasi.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Provers/quasi.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -3,10 +3,10 @@
 Reasoner for simple transitivity and quasi orders.
 *)
 
-(* 
- 
+(*
+
 The package provides tactics trans_tac and quasi_tac that use
-premises of the form 
+premises of the form
 
   t = u, t ~= u, t < u and t <= u
 
@@ -29,15 +29,15 @@
    the assumptions.
    Note that the presence of a strict relation is not necessary for
    quasi_tac.  Configure decomp_quasi to ignore < and ~=.  A list of
-   required theorems for both situations is given below. 
+   required theorems for both situations is given below.
 *)
 
 signature LESS_ARITH =
 sig
   (* Transitivity of <=
-     Note that transitivities for < hold for partial orders only. *) 
+     Note that transitivities for < hold for partial orders only. *)
   val le_trans: thm  (* [| x <= y; y <= z |] ==> x <= z *)
- 
+
   (* Additional theorem for quasi orders *)
   val le_refl: thm  (* x <= x *)
   val eqD1: thm (* x = y ==> x <= y *)
@@ -64,41 +64,36 @@
   val decomp_quasi: theory -> term -> (term * string * term) option
 end;
 
-signature QUASI_TAC = 
+signature QUASI_TAC =
 sig
-  val trans_tac: int -> tactic
-  val quasi_tac: int -> tactic
+  val trans_tac: Proof.context -> int -> tactic
+  val quasi_tac: Proof.context -> int -> tactic
 end;
 
-functor Quasi_Tac_Fun (Less: LESS_ARITH): QUASI_TAC =
+functor Quasi_Tac(Less: LESS_ARITH): QUASI_TAC =
 struct
 
-(* Extract subgoal with signature *)
-fun SUBGOAL goalfun i st =
-  goalfun (List.nth(prems_of st, i-1),  i, Thm.theory_of_thm st) st
-                             handle Subscript => Seq.empty;
-
 (* Internal datatype for the proof *)
 datatype proof
-  = Asm of int 
-  | Thm of proof list * thm; 
-  
+  = Asm of int
+  | Thm of proof list * thm;
+
 exception Cannot;
  (* Internal exception, raised if conclusion cannot be derived from
      assumptions. *)
 exception Contr of proof;
   (* Internal exception, raised if contradiction ( x < x ) was derived *)
 
-fun prove asms = 
+fun prove asms =
   let fun pr (Asm i) = List.nth (asms, i)
   |       pr (Thm (prfs, thm)) = (map pr prfs) MRS thm
   in pr end;
 
 (* Internal datatype for inequalities *)
-datatype less 
-   = Less  of term * term * proof 
+datatype less
+   = Less  of term * term * proof
    | Le    of term * term * proof
-   | NotEq of term * term * proof; 
+   | NotEq of term * term * proof;
 
  (* Misc functions for datatype less *)
 fun lower (Less (x, _, _)) = x
@@ -125,13 +120,13 @@
 
 fun mkasm_trans sign (t, n) =
   case Less.decomp_trans sign t of
-    SOME (x, rel, y) => 
+    SOME (x, rel, y) =>
     (case rel of
      "<="  =>  [Le (x, y, Asm n)]
     | _     => error ("trans_tac: unknown relation symbol ``" ^ rel ^
                  "''returned by decomp_trans."))
   | NONE => [];
-  
+
 (* ************************************************************************ *)
 (*                                                                          *)
 (* mkasm_quasi sign (t, n) : theory -> (Term.term * int) -> less            *)
@@ -145,15 +140,15 @@
 fun mkasm_quasi sign (t, n) =
   case Less.decomp_quasi sign t of
     SOME (x, rel, y) => (case rel of
-      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], Less.less_reflE)) 
+      "<"   => if (x aconv y) then raise Contr (Thm ([Asm n], Less.less_reflE))
                else [Less (x, y, Asm n)]
     | "<="  => [Le (x, y, Asm n)]
     | "="   => [Le (x, y, Thm ([Asm n], Less.eqD1)),
                 Le (y, x, Thm ([Asm n], Less.eqD2))]
-    | "~="  => if (x aconv y) then 
+    | "~="  => if (x aconv y) then
                   raise Contr (Thm ([(Thm ([(Thm ([], Less.le_refl)) ,(Asm n)], Less.le_neq_trans))], Less.less_reflE))
                else [ NotEq (x, y, Asm n),
-                      NotEq (y, x,Thm ( [Asm n], thm "not_sym"))] 
+                      NotEq (y, x,Thm ( [Asm n], thm "not_sym"))]
     | _     => error ("quasi_tac: unknown relation symbol ``" ^ rel ^
                  "''returned by decomp_quasi."))
   | NONE => [];
@@ -168,15 +163,15 @@
 (*                                                                          *)
 (* ************************************************************************ *)
 
-  
+
 fun mkconcl_trans sign t =
   case Less.decomp_trans sign t of
     SOME (x, rel, y) => (case rel of
-     "<="  => (Le (x, y, Asm ~1), Asm 0) 
+     "<="  => (Le (x, y, Asm ~1), Asm 0)
     | _  => raise Cannot)
   | NONE => raise Cannot;
-  
-  
+
+
 (* ************************************************************************ *)
 (*                                                                          *)
 (* mkconcl_quasi sign t : theory -> Term.term -> less                       *)
@@ -194,8 +189,8 @@
     | "~="  => ([NotEq (x,y, Asm ~1)], Asm 0)
     | _  => raise Cannot)
 | NONE => raise Cannot;
-  
-  
+
+
 (* ******************************************************************* *)
 (*                                                                     *)
 (* mergeLess (less1,less2):  less * less -> less                       *)
@@ -211,11 +206,11 @@
 fun mergeLess (Le (x, _, p) , Le (_ , z, q)) =
       Le (x, z, Thm ([p,q] , Less.le_trans))
 |   mergeLess (Le (x, z, p) , NotEq (x', z', q)) =
-      if (x aconv x' andalso z aconv z' ) 
+      if (x aconv x' andalso z aconv z' )
        then Less (x, z, Thm ([p,q] , Less.le_neq_trans))
         else error "quasi_tac: internal error le_neq_trans"
 |   mergeLess (NotEq (x, z, p) , Le (x' , z', q)) =
-      if (x aconv x' andalso z aconv z') 
+      if (x aconv x' andalso z aconv z')
       then Less (x, z, Thm ([p,q] , Less.neq_le_trans))
       else error "quasi_tac: internal error neq_le_trans"
 |   mergeLess (_, _) =
@@ -229,7 +224,7 @@
 infix tr;
 fun (Le (_, y, _))   tr (Le (x', _, _))   = ( y aconv x' )
   | _ tr _ = false;
-  
+
 (* ******************************************************************* *)
 (*                                                                     *)
 (* transPath (Lesslist, Less): (less list * less) -> less              *)
@@ -245,7 +240,7 @@
 |   transPath (x::xs,lesss) =
       if lesss tr x then transPath (xs, mergeLess(lesss,x))
       else error "trans/quasi_tac: internal error transpath";
-  
+
 (* ******************************************************************* *)
 (*                                                                     *)
 (* less1 subsumes less2 : less -> less -> bool                         *)
@@ -253,7 +248,7 @@
 (* subsumes checks whether less1 implies less2                         *)
 (*                                                                     *)
 (* ******************************************************************* *)
-  
+
 infix subsumes;
 fun (Le (x, y, _)) subsumes (Le (x', y', _)) =
       (x aconv x' andalso y aconv y')
@@ -271,7 +266,7 @@
 (* ******************************************************************* *)
 
 fun triv_solv (Le (x, x', _)) =
-    if x aconv x' then  SOME (Thm ([], Less.le_refl)) 
+    if x aconv x' then  SOME (Thm ([], Less.le_refl))
     else NONE
 |   triv_solv _ = NONE;
 
@@ -286,33 +281,33 @@
 fun addEdge (v,d,[]) = [(v,d)]
 |   addEdge (v,d,((u,dl)::el)) = if v aconv u then ((v,d@dl)::el)
     else (u,dl):: (addEdge(v,d,el));
-    
+
 (* ********************************************************************** *)
 (*                                                                        *)
-(* mkQuasiGraph constructs from a list of objects of type less a graph g, *) 
+(* mkQuasiGraph constructs from a list of objects of type less a graph g, *)
 (* by taking all edges that are candidate for a <=, and a list neqE, by   *)
 (* taking all edges that are candiate for a ~=                            *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
 fun mkQuasiGraph [] = ([],[])
-|   mkQuasiGraph lessList = 
+|   mkQuasiGraph lessList =
  let
  fun buildGraphs ([],leG, neqE) = (leG,  neqE)
-  |   buildGraphs (l::ls, leG,  neqE) = case l of 
+  |   buildGraphs (l::ls, leG,  neqE) = case l of
        (Less (x,y,p)) =>
-         let 
-	  val leEdge  = Le (x,y, Thm ([p], Less.less_imp_le)) 
-	  val neqEdges = [ NotEq (x,y, Thm ([p], Less.less_imp_neq)),
-	                   NotEq (y,x, Thm ( [Thm ([p], Less.less_imp_neq)], thm "not_sym"))]
-	 in
-           buildGraphs (ls, addEdge(y,[],(addEdge (x,[(y,leEdge)],leG))), neqEdges@neqE) 
-	 end
-     |  (Le (x,y,p))   => buildGraphs (ls, addEdge(y,[],(addEdge (x,[(y,l)],leG))), neqE) 
+         let
+          val leEdge  = Le (x,y, Thm ([p], Less.less_imp_le))
+          val neqEdges = [ NotEq (x,y, Thm ([p], Less.less_imp_neq)),
+                           NotEq (y,x, Thm ( [Thm ([p], Less.less_imp_neq)], thm "not_sym"))]
+         in
+           buildGraphs (ls, addEdge(y,[],(addEdge (x,[(y,leEdge)],leG))), neqEdges@neqE)
+         end
+     |  (Le (x,y,p))   => buildGraphs (ls, addEdge(y,[],(addEdge (x,[(y,l)],leG))), neqE)
      | _ =>  buildGraphs (ls, leG,  l::neqE) ;
 
 in buildGraphs (lessList, [],  []) end;
-  
+
 (* ********************************************************************** *)
 (*                                                                        *)
 (* mkGraph constructs from a list of objects of type less a graph g       *)
@@ -321,11 +316,11 @@
 (* ********************************************************************** *)
 
 fun mkGraph [] = []
-|   mkGraph lessList = 
+|   mkGraph lessList =
  let
   fun buildGraph ([],g) = g
-  |   buildGraph (l::ls, g) =  buildGraph (ls, (addEdge ((lower l),[((upper l),l)],g))) 
-     
+  |   buildGraph (l::ls, g) =  buildGraph (ls, (addEdge ((lower l),[((upper l),l)],g)))
+
 in buildGraph (lessList, []) end;
 
 (* *********************************************************************** *)
@@ -335,42 +330,42 @@
 (* List of successors of u in graph g                                      *)
 (*                                                                         *)
 (* *********************************************************************** *)
- 
-fun adjacent eq_comp ((v,adj)::el) u = 
+
+fun adjacent eq_comp ((v,adj)::el) u =
     if eq_comp (u, v) then adj else adjacent eq_comp el u
-|   adjacent _  []  _ = []  
+|   adjacent _  []  _ = []
 
 (* *********************************************************************** *)
 (*                                                                         *)
 (* dfs eq_comp g u v:                                                      *)
 (* ('a * 'a -> bool) -> ('a  *( 'a * less) list) list ->                   *)
-(* 'a -> 'a -> (bool * ('a * less) list)                                   *) 
+(* 'a -> 'a -> (bool * ('a * less) list)                                   *)
 (*                                                                         *)
 (* Depth first search of v from u.                                         *)
 (* Returns (true, path(u, v)) if successful, otherwise (false, []).        *)
 (*                                                                         *)
 (* *********************************************************************** *)
 
-fun dfs eq_comp g u v = 
- let 
+fun dfs eq_comp g u v =
+ let
     val pred = ref nil;
     val visited = ref nil;
-    
+
     fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
-    
-    fun dfs_visit u' = 
+
+    fun dfs_visit u' =
     let val _ = visited := u' :: (!visited)
-    
+
     fun update (x,l) = let val _ = pred := (x,l) ::(!pred) in () end;
-    
-    in if been_visited v then () 
+
+    in if been_visited v then ()
     else (app (fn (v',l) => if been_visited v' then () else (
-       update (v',l); 
+       update (v',l);
        dfs_visit v'; ()) )) (adjacent eq_comp g u')
      end
-  in 
-    dfs_visit u; 
-    if (been_visited v) then (true, (!pred)) else (false , [])   
+  in
+    dfs_visit u;
+    if (been_visited v) then (true, (!pred)) else (false , [])
   end;
 
 (* ************************************************************************ *)
@@ -393,8 +388,8 @@
 (* ************************************************************************ *)
 
 
- fun findPath x y g = 
-  let 
+ fun findPath x y g =
+  let
     val (found, tmp) =  dfs (op aconv) g x y ;
     val pred = map snd tmp;
 
@@ -403,38 +398,38 @@
        (* find predecessor u of node v and the edge u -> v *)
        fun lookup v [] = raise Cannot
        |   lookup v (e::es) = if (upper e) aconv v then e else lookup v es;
-		
-       (* traverse path backwards and return list of visited edges *)   
-       fun rev_path v = 
- 	let val l = lookup v pred
+
+       (* traverse path backwards and return list of visited edges *)
+       fun rev_path v =
+        let val l = lookup v pred
             val u = lower l;
- 	in
-           if u aconv x then [l] else (rev_path u) @ [l] 
-	end
+        in
+           if u aconv x then [l] else (rev_path u) @ [l]
+        end
       in rev_path y end;
-		
-  in 
+
+  in
    if found then (
     if x aconv y then (found,[(Le (x, y, (Thm ([], Less.le_refl))))])
-    else (found, (path x y) )) 
+    else (found, (path x y) ))
    else (found,[])
-  end; 
-	
-      
-(* ************************************************************************ *) 
+  end;
+
+
+(* ************************************************************************ *)
 (*                                                                          *)
 (* findQuasiProof (leqG, neqE) subgoal:                                     *)
 (* (Term.term * (Term.term * less list) list) * less list  -> less -> proof *)
 (*                                                                          *)
 (* Constructs a proof for subgoal by searching a special path in leqG and   *)
-(* neqE. Raises Cannot if construction of the proof fails.                  *)   
+(* neqE. Raises Cannot if construction of the proof fails.                  *)
 (*                                                                          *)
-(* ************************************************************************ *) 
+(* ************************************************************************ *)
 
 
 (* As the conlusion can be either of form x <= y, x < y or x ~= y we have        *)
 (* three cases to deal with. Finding a transitivity path from x to y with label  *)
-(* 1. <=                                                                         *) 
+(* 1. <=                                                                         *)
 (*    This is simply done by searching any path from x to y in the graph leG.    *)
 (*    The graph leG contains only edges with label <=.                           *)
 (*                                                                               *)
@@ -450,54 +445,54 @@
 
 fun findQuasiProof (leG, neqE) subgoal =
   case subgoal of (Le (x,y, _)) => (
-   let 
-    val (xyLefound,xyLePath) = findPath x y leG 
+   let
+    val (xyLefound,xyLePath) = findPath x y leG
    in
     if xyLefound then (
-     let 
+     let
       val Le_x_y = (transPath (tl xyLePath, hd xyLePath))
      in getprf Le_x_y end )
     else raise Cannot
    end )
   | (Less (x,y,_))  => (
-   let 
+   let
     fun findParallelNeq []  = NONE
     |   findParallelNeq (e::es)  =
      if      (x aconv (lower e) andalso y aconv (upper e)) then SOME e
      else if (y aconv (lower e) andalso x aconv (upper e)) then SOME (NotEq (x,y, (Thm ([getprf e], thm "not_sym"))))
-     else findParallelNeq es ;  
+     else findParallelNeq es ;
    in
    (* test if there is a edge x ~= y respectivly  y ~= x and     *)
    (* if it possible to find a path x <= y in leG, thus we can conclude x < y *)
-    (case findParallelNeq neqE of (SOME e) => 
-      let 
-       val (xyLeFound,xyLePath) = findPath x y leG 
+    (case findParallelNeq neqE of (SOME e) =>
+      let
+       val (xyLeFound,xyLePath) = findPath x y leG
       in
        if xyLeFound then (
-        let 
+        let
          val Le_x_y = (transPath (tl xyLePath, hd xyLePath))
          val Less_x_y = mergeLess (e, Le_x_y)
         in getprf Less_x_y end
        ) else raise Cannot
-      end 
-    | _ => raise Cannot)    
+      end
+    | _ => raise Cannot)
    end )
- | (NotEq (x,y,_)) => 
+ | (NotEq (x,y,_)) =>
   (* First check if a single premiss is sufficient *)
   (case (Library.find_first (fn fact => fact subsumes subgoal) neqE, subgoal) of
     (SOME (NotEq (x, y, p)), NotEq (x', y', _)) =>
-      if  (x aconv x' andalso y aconv y') then p 
+      if  (x aconv x' andalso y aconv y') then p
       else Thm ([p], thm "not_sym")
-    | _  => raise Cannot 
+    | _  => raise Cannot
   )
 
-      
-(* ************************************************************************ *) 
-(*                                                                          *) 
-(* End: Quasi Order relevant functions                                      *) 
-(*                                                                          *) 
-(*                                                                          *) 
-(* ************************************************************************ *) 
+
+(* ************************************************************************ *)
+(*                                                                          *)
+(* End: Quasi Order relevant functions                                      *)
+(*                                                                          *)
+(*                                                                          *)
+(* ************************************************************************ *)
 
 (* *********************************************************************** *)
 (*                                                                         *)
@@ -509,14 +504,14 @@
 (* *********************************************************************** *)
 
 fun solveLeTrans sign (asms, concl) =
- let 
+ let
   val g = mkGraph asms
  in
-   let 
+   let
     val (subgoal, prf) = mkconcl_trans sign concl
-    val (found, path) = findPath (lower subgoal) (upper subgoal) g 
+    val (found, path) = findPath (lower subgoal) (upper subgoal) g
    in
-    if found then [getprf (transPath (tl path, hd path))] 
+    if found then [getprf (transPath (tl path, hd path))]
     else raise Cannot
   end
  end;
@@ -532,10 +527,10 @@
 (* *********************************************************************** *)
 
 fun solveQuasiOrder sign (asms, concl) =
- let 
+ let
   val (leG, neqE) = mkQuasiGraph asms
  in
-   let 
+   let
    val (subgoals, prf) = mkconcl_quasi sign concl
    fun solve facts less =
        (case triv_solv less of NONE => findQuasiProof (leG, neqE) less
@@ -543,54 +538,56 @@
   in   map (solve asms) subgoals end
  end;
 
-(* ************************************************************************ *) 
-(*                                                                          *) 
+(* ************************************************************************ *)
+(*                                                                          *)
 (* Tactics                                                                  *)
 (*                                                                          *)
-(*  - trans_tac                                                          *)                     
-(*  - quasi_tac, solves quasi orders                                        *)                     
-(* ************************************************************************ *) 
+(*  - trans_tac                                                          *)
+(*  - quasi_tac, solves quasi orders                                        *)
+(* ************************************************************************ *)
 
 
 (* trans_tac - solves transitivity chains over <= *)
-val trans_tac  =  SUBGOAL (fn (A, n, sign) =>
+
+fun trans_tac ctxt = SUBGOAL (fn (A, n) =>
  let
-  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
-  val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
-  val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
-  val lesss = List.concat (ListPair.map (mkasm_trans  sign) (Hs, 0 upto (length Hs - 1)))
-  val prfs = solveLeTrans  sign (lesss, C);
-  
-  val (subgoal, prf) = mkconcl_trans  sign C;
+  val thy = ProofContext.theory_of ctxt;
+  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A));
+  val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A);
+  val C = subst_bounds (rfrees, Logic.strip_assums_concl A);
+  val lesss = List.concat (ListPair.map (mkasm_trans thy) (Hs, 0 upto (length Hs - 1)));
+  val prfs = solveLeTrans thy (lesss, C);
+
+  val (subgoal, prf) = mkconcl_trans thy C;
  in
-  
-  METAHYPS (fn asms =>
-    let val thms = map (prove asms) prfs
-    in rtac (prove thms prf) 1 end) n
-  
+  FOCUS (fn {prems, ...} =>
+    let val thms = map (prove prems) prfs
+    in rtac (prove thms prf) 1 end) ctxt n
  end
- handle Contr p => METAHYPS (fn asms => rtac (prove asms p) 1) n
-      | Cannot  => no_tac
-);
+ handle Contr p => FOCUS (fn {prems, ...} => rtac (prove prems p) 1) ctxt n
+  | Cannot  => no_tac);
+
 
 (* quasi_tac - solves quasi orders *)
-val quasi_tac = SUBGOAL (fn (A, n, sign) =>
+
+fun quasi_tac ctxt = SUBGOAL (fn (A, n) => fn st =>
  let
-  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A))
-  val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A)
-  val C = subst_bounds (rfrees, Logic.strip_assums_concl A)
-  val lesss = List.concat (ListPair.map (mkasm_quasi sign) (Hs, 0 upto (length Hs - 1)))
-  val prfs = solveQuasiOrder sign (lesss, C);
-  val (subgoals, prf) = mkconcl_quasi sign C;
+  val thy = ProofContext.theory_of ctxt;
+  val rfrees = map Free (Term.rename_wrt_term A (Logic.strip_params A));
+  val Hs = map (fn H => subst_bounds (rfrees, H)) (Logic.strip_assums_hyp A);
+  val C = subst_bounds (rfrees, Logic.strip_assums_concl A);
+  val lesss = flat (ListPair.map (mkasm_quasi thy) (Hs, 0 upto (length Hs - 1)));
+  val prfs = solveQuasiOrder thy (lesss, C);
+  val (subgoals, prf) = mkconcl_quasi thy C;
  in
- 
-  METAHYPS (fn asms =>
-    let val thms = map (prove asms) prfs
-    in rtac (prove thms prf) 1 end) n
- 
+  FOCUS (fn {prems, ...} =>
+    let val thms = map (prove prems) prfs
+    in rtac (prove thms prf) 1 end) ctxt n st
  end
- handle Contr p => METAHYPS (fn asms => rtac (prove asms p) 1) n
-      | Cannot  => no_tac
-);
-   
+ handle Contr p =>
+    (FOCUS (fn {prems, ...} => rtac (prove prems p) 1) ctxt n st
+      handle Subscript => Seq.empty)
+  | Cannot => Seq.empty
+  | Subscript => Seq.empty);
+
 end;
--- a/src/Provers/splitter.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Provers/splitter.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -1,5 +1,4 @@
-(*  Title:      Provers/splitter
-    ID:         $Id$
+(*  Title:      Provers/splitter.ML
     Author:     Tobias Nipkow
     Copyright   1995  TU Munich
 
@@ -12,6 +11,7 @@
 
 signature SPLITTER_DATA =
 sig
+  val thy           : theory
   val mk_eq         : thm -> thm
   val meta_eq_to_iff: thm (* "x == y ==> x = y"                      *)
   val iffD          : thm (* "[| P = Q; Q |] ==> P"                  *)
@@ -41,18 +41,18 @@
   val setup: theory -> theory
 end;
 
-functor SplitterFun(Data: SPLITTER_DATA): SPLITTER =
+functor Splitter(Data: SPLITTER_DATA): SPLITTER =
 struct
 
 val Const (const_not, _) $ _ =
-  ObjectLogic.drop_judgment (the_context ())
+  ObjectLogic.drop_judgment Data.thy
     (#1 (Logic.dest_implies (Thm.prop_of Data.notnotD)));
 
 val Const (const_or , _) $ _ $ _ =
-  ObjectLogic.drop_judgment (the_context ())
+  ObjectLogic.drop_judgment Data.thy
     (#1 (Logic.dest_implies (Thm.prop_of Data.disjE)));
 
-val const_Trueprop = ObjectLogic.judgment_name (the_context ());
+val const_Trueprop = ObjectLogic.judgment_name Data.thy;
 
 
 fun split_format_err () = error "Wrong format for split rule";
--- a/src/Provers/trancl.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Provers/trancl.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -26,7 +26,7 @@
 
 *)
 
-signature TRANCL_ARITH = 
+signature TRANCL_ARITH =
 sig
 
   (* theorems for transitive closure *)
@@ -63,25 +63,25 @@
             "r":   the relation itself,
             "r^+": transitive closure of the relation,
             "r^*": reflexive-transitive closure of the relation
-  *)  
+  *)
 
-  val decomp: term ->  (term * term * term * string) option 
+  val decomp: term ->  (term * term * term * string) option
 
 end;
 
-signature TRANCL_TAC = 
+signature TRANCL_TAC =
 sig
-  val trancl_tac: int -> tactic;
-  val rtrancl_tac: int -> tactic;
+  val trancl_tac: Proof.context -> int -> tactic
+  val rtrancl_tac: Proof.context -> int -> tactic
 end;
 
-functor Trancl_Tac_Fun (Cls : TRANCL_ARITH): TRANCL_TAC = 
+functor Trancl_Tac(Cls: TRANCL_ARITH): TRANCL_TAC =
 struct
 
- 
+
 datatype proof
-  = Asm of int 
-  | Thm of proof list * thm; 
+  = Asm of int
+  | Thm of proof list * thm;
 
 exception Cannot; (* internal exception: raised if no proof can be found *)
 
@@ -89,7 +89,7 @@
   (Envir.beta_eta_contract x, Envir.beta_eta_contract y,
    Envir.beta_eta_contract rel, r)) (Cls.decomp t);
 
-fun prove thy r asms = 
+fun prove thy r asms =
   let
     fun inst thm =
       let val SOME (_, _, r', _) = decomp (concl_of thm)
@@ -98,12 +98,12 @@
       | pr (Thm (prfs, thm)) = map pr prfs MRS inst thm
   in pr end;
 
-  
+
 (* Internal datatype for inequalities *)
-datatype rel 
+datatype rel
    = Trans  of term * term * proof  (* R^+ *)
    | RTrans of term * term * proof; (* R^* *)
-   
+
  (* Misc functions for datatype rel *)
 fun lower (Trans (x, _, _)) = x
   | lower (RTrans (x,_,_)) = x;
@@ -112,8 +112,8 @@
   | upper (RTrans (_,y,_)) = y;
 
 fun getprf   (Trans   (_, _, p)) = p
-|   getprf   (RTrans (_,_, p)) = p; 
- 
+|   getprf   (RTrans (_,_, p)) = p;
+
 (* ************************************************************************ *)
 (*                                                                          *)
 (*  mkasm_trancl Rel (t,n): term -> (term , int) -> rel list                *)
@@ -127,16 +127,16 @@
 
 fun mkasm_trancl  Rel  (t, n) =
   case decomp t of
-    SOME (x, y, rel,r) => if rel aconv Rel then  
-    
+    SOME (x, y, rel,r) => if rel aconv Rel then
+
     (case r of
       "r"   => [Trans (x,y, Thm([Asm n], Cls.r_into_trancl))]
     | "r+"  => [Trans (x,y, Asm n)]
     | "r*"  => []
     | _     => error ("trancl_tac: unknown relation symbol"))
-    else [] 
+    else []
   | NONE => [];
-  
+
 (* ************************************************************************ *)
 (*                                                                          *)
 (*  mkasm_rtrancl Rel (t,n): term -> (term , int) -> rel list               *)
@@ -156,7 +156,7 @@
     | "r+"  => [ Trans (x,y, Asm n)]
     | "r*"  => [ RTrans(x,y, Asm n)]
     | _     => error ("rtrancl_tac: unknown relation symbol" ))
-   else [] 
+   else []
   | NONE => [];
 
 (* ************************************************************************ *)
@@ -204,7 +204,7 @@
 fun makeStep (Trans (a,_,p), Trans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.trancl_trans))
 (* refl. + trans. cls. rules *)
 |   makeStep (RTrans (a,_,p), Trans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.rtrancl_trancl_trancl))
-|   makeStep (Trans (a,_,p), RTrans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.trancl_rtrancl_trancl))   
+|   makeStep (Trans (a,_,p), RTrans(_,c,q))  = Trans (a,c, Thm ([p,q], Cls.trancl_rtrancl_trancl))
 |   makeStep (RTrans (a,_,p), RTrans(_,c,q))  = RTrans (a,c, Thm ([p,q], Cls.rtrancl_trans));
 
 (* ******************************************************************* *)
@@ -220,7 +220,7 @@
 
 fun transPath ([],acc) = acc
 |   transPath (x::xs,acc) = transPath (xs, makeStep(acc,x))
-      
+
 (* ********************************************************************* *)
 (* Graph functions                                                       *)
 (* ********************************************************************* *)
@@ -232,7 +232,7 @@
 fun addEdge (v,d,[]) = [(v,d)]
 |   addEdge (v,d,((u,dl)::el)) = if v aconv u then ((v,d@dl)::el)
     else (u,dl):: (addEdge(v,d,el));
-    
+
 (* ********************************************************************** *)
 (*                                                                        *)
 (* mkGraph constructs from a list of objects of type rel  a graph g       *)
@@ -241,13 +241,13 @@
 (* ********************************************************************** *)
 
 fun mkGraph [] = ([],[])
-|   mkGraph ys =  
+|   mkGraph ys =
  let
   fun buildGraph ([],g,zs) = (g,zs)
-  |   buildGraph (x::xs, g, zs) = 
-        case x of (Trans (_,_,_)) => 
-	       buildGraph (xs, addEdge((upper x), [],(addEdge ((lower x),[((upper x),x)],g))), x::zs) 
-	| _ => buildGraph (xs, addEdge((upper x), [],(addEdge ((lower x),[((upper x),x)],g))), zs)
+  |   buildGraph (x::xs, g, zs) =
+        case x of (Trans (_,_,_)) =>
+               buildGraph (xs, addEdge((upper x), [],(addEdge ((lower x),[((upper x),x)],g))), x::zs)
+        | _ => buildGraph (xs, addEdge((upper x), [],(addEdge ((lower x),[((upper x),x)],g))), zs)
 in buildGraph (ys, [], []) end;
 
 (* *********************************************************************** *)
@@ -257,42 +257,42 @@
 (* List of successors of u in graph g                                      *)
 (*                                                                         *)
 (* *********************************************************************** *)
- 
-fun adjacent eq_comp ((v,adj)::el) u = 
+
+fun adjacent eq_comp ((v,adj)::el) u =
     if eq_comp (u, v) then adj else adjacent eq_comp el u
-|   adjacent _  []  _ = []  
+|   adjacent _  []  _ = []
 
 (* *********************************************************************** *)
 (*                                                                         *)
 (* dfs eq_comp g u v:                                                      *)
 (* ('a * 'a -> bool) -> ('a  *( 'a * rel) list) list ->                    *)
-(* 'a -> 'a -> (bool * ('a * rel) list)                                    *) 
+(* 'a -> 'a -> (bool * ('a * rel) list)                                    *)
 (*                                                                         *)
 (* Depth first search of v from u.                                         *)
 (* Returns (true, path(u, v)) if successful, otherwise (false, []).        *)
 (*                                                                         *)
 (* *********************************************************************** *)
 
-fun dfs eq_comp g u v = 
- let 
+fun dfs eq_comp g u v =
+ let
     val pred = ref nil;
     val visited = ref nil;
-    
+
     fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
-    
-    fun dfs_visit u' = 
+
+    fun dfs_visit u' =
     let val _ = visited := u' :: (!visited)
-    
+
     fun update (x,l) = let val _ = pred := (x,l) ::(!pred) in () end;
-    
-    in if been_visited v then () 
+
+    in if been_visited v then ()
     else (app (fn (v',l) => if been_visited v' then () else (
-       update (v',l); 
+       update (v',l);
        dfs_visit v'; ()) )) (adjacent eq_comp g u')
      end
-  in 
-    dfs_visit u; 
-    if (been_visited v) then (true, (!pred)) else (false , [])   
+  in
+    dfs_visit u;
+    if (been_visited v) then (true, (!pred)) else (false , [])
   end;
 
 (* *********************************************************************** *)
@@ -310,7 +310,7 @@
    (* Compute list of reversed edges for each adjacency list *)
    fun flip (u,(v,l)::el) = (v,(u,l)) :: flip (u,el)
      | flip (_,nil) = nil
-   
+
    (* Compute adjacency list for node u from the list of edges
       and return a likewise reduced list of edges.  The list of edges
       is searches for edges starting from u, and these edges are removed. *)
@@ -334,23 +334,23 @@
    (* Compute, for each adjacency list, the list with reversed edges,
       and concatenate these lists. *)
    val flipped = List.foldr (op @) nil (map flip g)
- 
- in assemble g flipped end    
- 
+
+ in assemble g flipped end
+
 (* *********************************************************************** *)
 (*                                                                         *)
 (* dfs_reachable eq_comp g u:                                              *)
-(* (int * int list) list -> int -> int list                                *) 
+(* (int * int list) list -> int -> int list                                *)
 (*                                                                         *)
 (* Computes list of all nodes reachable from u in g.                       *)
 (*                                                                         *)
 (* *********************************************************************** *)
 
-fun dfs_reachable eq_comp g u = 
+fun dfs_reachable eq_comp g u =
  let
   (* List of vertices which have been visited. *)
   val visited  = ref nil;
-  
+
   fun been_visited v = exists (fn w => eq_comp (w, v)) (!visited)
 
   fun dfs_visit g u  =
@@ -361,13 +361,13 @@
             else v :: dfs_visit g v @ ds)
         nil (adjacent eq_comp g u)
    in  descendents end
- 
+
  in u :: dfs_visit g u end;
 
 (* *********************************************************************** *)
 (*                                                                         *)
 (* dfs_term_reachable g u:                                                  *)
-(* (term * term list) list -> term -> term list                            *) 
+(* (term * term list) list -> term -> term list                            *)
 (*                                                                         *)
 (* Computes list of all nodes reachable from u in g.                       *)
 (*                                                                         *)
@@ -375,45 +375,45 @@
 
 fun dfs_term_reachable g u = dfs_reachable (op aconv) g u;
 
-(* ************************************************************************ *) 
+(* ************************************************************************ *)
 (*                                                                          *)
 (* findPath x y g: Term.term -> Term.term ->                                *)
-(*                  (Term.term * (Term.term * rel list) list) ->            *) 
+(*                  (Term.term * (Term.term * rel list) list) ->            *)
 (*                  (bool, rel list)                                        *)
 (*                                                                          *)
 (*  Searches a path from vertex x to vertex y in Graph g, returns true and  *)
 (*  the list of edges if path is found, otherwise false and nil.            *)
 (*                                                                          *)
-(* ************************************************************************ *) 
- 
-fun findPath x y g = 
-  let 
+(* ************************************************************************ *)
+
+fun findPath x y g =
+  let
    val (found, tmp) =  dfs (op aconv) g x y ;
    val pred = map snd tmp;
-	 
+
    fun path x y  =
     let
-	 (* find predecessor u of node v and the edge u -> v *)
-		
+         (* find predecessor u of node v and the edge u -> v *)
+
       fun lookup v [] = raise Cannot
       |   lookup v (e::es) = if (upper e) aconv v then e else lookup v es;
-		
-      (* traverse path backwards and return list of visited edges *)   
-      fun rev_path v = 
-	let val l = lookup v pred
-	    val u = lower l;
-	in
-	  if u aconv x then [l] else (rev_path u) @ [l] 
-	end
-       
+
+      (* traverse path backwards and return list of visited edges *)
+      fun rev_path v =
+        let val l = lookup v pred
+            val u = lower l;
+        in
+          if u aconv x then [l] else (rev_path u) @ [l]
+        end
+
     in rev_path y end;
-		
-   in 
+
+   in
+
 
-     
       if found then ( (found, (path x y) )) else (found,[])
-   
-     
+
+
 
    end;
 
@@ -426,143 +426,142 @@
 (*                                                                          *)
 (* ************************************************************************ *)
 
-fun findRtranclProof g tranclEdges subgoal = 
+fun findRtranclProof g tranclEdges subgoal =
    case subgoal of (RTrans (x,y,_)) => if x aconv y then [Thm ([], Cls.rtrancl_refl)] else (
      let val (found, path) = findPath (lower subgoal) (upper subgoal) g
-     in 
+     in
        if found then (
           let val path' = (transPath (tl path, hd path))
-	  in 
-	   
-	    case path' of (Trans (_,_,p)) => [Thm ([p], Cls.trancl_into_rtrancl )] 
-	    | _ => [getprf path']
-	   
-	  end
+          in
+
+            case path' of (Trans (_,_,p)) => [Thm ([p], Cls.trancl_into_rtrancl )]
+            | _ => [getprf path']
+
+          end
        )
        else raise Cannot
      end
    )
-   
+
 | (Trans (x,y,_)) => (
- 
+
   let
    val Vx = dfs_term_reachable g x;
    val g' = transpose (op aconv) g;
    val Vy = dfs_term_reachable g' y;
-   
+
    fun processTranclEdges [] = raise Cannot
-   |   processTranclEdges (e::es) = 
+   |   processTranclEdges (e::es) =
           if (upper e) mem Vx andalso (lower e) mem Vx
-	  andalso (upper e) mem Vy andalso (lower e) mem Vy
-	  then (
-	      
-	   
-	    if (lower e) aconv x then (
-	      if (upper e) aconv y then (
-	          [(getprf e)] 
-	      )
-	      else (
-	          let 
-		    val (found,path) = findPath (upper e) y g
-		  in
+          andalso (upper e) mem Vy andalso (lower e) mem Vy
+          then (
+
+
+            if (lower e) aconv x then (
+              if (upper e) aconv y then (
+                  [(getprf e)]
+              )
+              else (
+                  let
+                    val (found,path) = findPath (upper e) y g
+                  in
+
+                   if found then (
+                       [getprf (transPath (path, e))]
+                      ) else processTranclEdges es
+
+                  end
+              )
+            )
+            else if (upper e) aconv y then (
+               let val (xufound,xupath) = findPath x (lower e) g
+               in
+
+                  if xufound then (
 
-		   if found then ( 
-		       [getprf (transPath (path, e))]
-		      ) else processTranclEdges es
-		  
-		  end 
-	      )   
-	    )
-	    else if (upper e) aconv y then (
-	       let val (xufound,xupath) = findPath x (lower e) g
-	       in 
-	       
-	          if xufound then (
-		  	    
-		    let val xuRTranclEdge = transPath (tl xupath, hd xupath)
-			    val xyTranclEdge = makeStep(xuRTranclEdge,e)
-				
-				in [getprf xyTranclEdge] end
-				
-	         ) else processTranclEdges es
-	       
-	       end
-	    )
-	    else ( 
-	   
-	        let val (xufound,xupath) = findPath x (lower e) g
-		    val (vyfound,vypath) = findPath (upper e) y g
-		 in 
-		    if xufound then (
-		         if vyfound then ( 
-			    let val xuRTranclEdge = transPath (tl xupath, hd xupath)
-			        val vyRTranclEdge = transPath (tl vypath, hd vypath)
-				val xyTranclEdge = makeStep (makeStep(xuRTranclEdge,e),vyRTranclEdge)
-				
-				in [getprf xyTranclEdge] end
-				
-			 ) else processTranclEdges es
-		    ) 
-		    else processTranclEdges es
-		 end
-	    )
-	  )
-	  else processTranclEdges es;
+                    let val xuRTranclEdge = transPath (tl xupath, hd xupath)
+                            val xyTranclEdge = makeStep(xuRTranclEdge,e)
+
+                                in [getprf xyTranclEdge] end
+
+                 ) else processTranclEdges es
+
+               end
+            )
+            else (
+
+                let val (xufound,xupath) = findPath x (lower e) g
+                    val (vyfound,vypath) = findPath (upper e) y g
+                 in
+                    if xufound then (
+                         if vyfound then (
+                            let val xuRTranclEdge = transPath (tl xupath, hd xupath)
+                                val vyRTranclEdge = transPath (tl vypath, hd vypath)
+                                val xyTranclEdge = makeStep (makeStep(xuRTranclEdge,e),vyRTranclEdge)
+
+                                in [getprf xyTranclEdge] end
+
+                         ) else processTranclEdges es
+                    )
+                    else processTranclEdges es
+                 end
+            )
+          )
+          else processTranclEdges es;
    in processTranclEdges tranclEdges end )
 | _ => raise Cannot
 
-   
-fun solveTrancl (asms, concl) = 
+
+fun solveTrancl (asms, concl) =
  let val (g,_) = mkGraph asms
  in
   let val (_, subgoal, _) = mkconcl_trancl concl
       val (found, path) = findPath (lower subgoal) (upper subgoal) g
   in
     if found then  [getprf (transPath (tl path, hd path))]
-    else raise Cannot 
+    else raise Cannot
   end
  end;
-  
-fun solveRtrancl (asms, concl) = 
+
+fun solveRtrancl (asms, concl) =
  let val (g,tranclEdges) = mkGraph asms
      val (_, subgoal, _) = mkconcl_rtrancl concl
 in
   findRtranclProof g tranclEdges subgoal
 end;
- 
-   
-val trancl_tac =   SUBGOAL (fn (A, n) => fn st =>
+
+
+fun trancl_tac ctxt = SUBGOAL (fn (A, n) =>
  let
-  val thy = theory_of_thm st;
+  val thy = ProofContext.theory_of ctxt;
   val Hs = Logic.strip_assums_hyp A;
   val C = Logic.strip_assums_concl A;
-  val (rel,subgoals, prf) = mkconcl_trancl C;
-  val prems = List.concat (ListPair.map (mkasm_trancl rel) (Hs, 0 upto (length Hs - 1)))
+  val (rel, subgoals, prf) = mkconcl_trancl C;
+
+  val prems = flat (ListPair.map (mkasm_trancl rel) (Hs, 0 upto (length Hs - 1)));
   val prfs = solveTrancl (prems, C);
-
  in
-  METAHYPS (fn asms =>
-    let val thms = map (prove thy rel asms) prfs
-    in rtac (prove thy rel thms prf) 1 end) n st
+  FOCUS (fn {prems, ...} =>
+    let val thms = map (prove thy rel prems) prfs
+    in rtac (prove thy rel thms prf) 1 end) ctxt n
  end
-handle  Cannot  => no_tac st);
+ handle Cannot => no_tac);
 
- 
- 
-val rtrancl_tac =   SUBGOAL (fn (A, n) => fn st =>
+
+fun rtrancl_tac ctxt = SUBGOAL (fn (A, n) => fn st =>
  let
-  val thy = theory_of_thm st;
+  val thy = ProofContext.theory_of ctxt;
   val Hs = Logic.strip_assums_hyp A;
   val C = Logic.strip_assums_concl A;
-  val (rel,subgoals, prf) = mkconcl_rtrancl C;
+  val (rel, subgoals, prf) = mkconcl_rtrancl C;
 
-  val prems = List.concat (ListPair.map (mkasm_rtrancl rel) (Hs, 0 upto (length Hs - 1)))
+  val prems = flat (ListPair.map (mkasm_rtrancl rel) (Hs, 0 upto (length Hs - 1)));
   val prfs = solveRtrancl (prems, C);
  in
-  METAHYPS (fn asms =>
-    let val thms = map (prove thy rel asms) prfs
-    in rtac (prove thy rel thms prf) 1 end) n st
+  FOCUS (fn {prems, ...} =>
+    let val thms = map (prove thy rel prems) prfs
+    in rtac (prove thy rel thms prf) 1 end) ctxt n st
  end
-handle  Cannot  => no_tac st);
+ handle Cannot => Seq.empty | Subscript => Seq.empty);
 
 end;
--- a/src/Pure/Concurrent/future.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Concurrent/future.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -1,7 +1,8 @@
 (*  Title:      Pure/Concurrent/future.ML
     Author:     Makarius
 
-Future values.
+Future values, see also
+http://www4.in.tum.de/~wenzelm/papers/parallel-isabelle.pdf
 
 Notes:
 
@@ -11,7 +12,7 @@
     processes.
 
   * Futures are grouped; failure of one group member causes the whole
-    group to be interrupted eventually.
+    group to be interrupted eventually.  Groups are block-structured.
 
   * Forked futures are evaluated spontaneously by a farm of worker
     threads in the background; join resynchronizes the computation and
@@ -26,7 +27,6 @@
 
 signature FUTURE =
 sig
-  val enabled: unit -> bool
   type task = Task_Queue.task
   type group = Task_Queue.group
   val is_worker: unit -> bool
@@ -46,7 +46,6 @@
   val join: 'a future -> 'a
   val map: ('a -> 'b) -> 'a future -> 'b future
   val interruptible_task: ('a -> 'b) -> 'a -> 'b
-  val interrupt_task: string -> unit
   val cancel_group: group -> unit
   val cancel: 'a future -> unit
   val shutdown: unit -> unit
@@ -57,11 +56,6 @@
 
 (** future values **)
 
-fun enabled () =
-  Multithreading.enabled () andalso
-    not (Multithreading.self_critical ());
-
-
 (* identifiers *)
 
 type task = Task_Queue.task;
@@ -84,18 +78,19 @@
 datatype 'a future = Future of
  {task: task,
   group: group,
-  result: 'a Exn.result option ref};
+  result: 'a Exn.result option Synchronized.var};
 
 fun task_of (Future {task, ...}) = task;
 fun group_of (Future {group, ...}) = group;
+fun result_of (Future {result, ...}) = result;
 
-fun peek (Future {result, ...}) = ! result;
+fun peek x = Synchronized.peek (result_of x);
 fun is_finished x = is_some (peek x);
 
 fun value x = Future
  {task = Task_Queue.new_task 0,
   group = Task_Queue.new_group NONE,
-  result = ref (SOME (Exn.Result x))};
+  result = Synchronized.var "future" (SOME (Exn.Result x))};
 
 
 
@@ -114,58 +109,51 @@
 
 (* synchronization *)
 
+val scheduler_event = ConditionVar.conditionVar ();
+val work_available = ConditionVar.conditionVar ();
+val work_finished = ConditionVar.conditionVar ();
+
 local
   val lock = Mutex.mutex ();
-  val cond = ConditionVar.conditionVar ();
 in
 
 fun SYNCHRONIZED name = SimpleThread.synchronized name lock;
 
-fun wait () = (*requires SYNCHRONIZED*)
-  ConditionVar.wait (cond, lock);
+fun wait cond = (*requires SYNCHRONIZED*)
+  ConditionVar.wait (cond, lock) handle Exn.Interrupt => ();
+
+fun wait_interruptible cond timeout = (*requires SYNCHRONIZED*)
+  interruptible (fn () =>
+    ignore (ConditionVar.waitUntil (cond, lock, Time.+ (Time.now (), timeout)))) ();
 
-fun wait_timeout timeout = (*requires SYNCHRONIZED*)
-  ignore (ConditionVar.waitUntil (cond, lock, Time.+ (Time.now (), timeout)));
+fun signal cond = (*requires SYNCHRONIZED*)
+  ConditionVar.signal cond;
 
-fun notify_all () = (*requires SYNCHRONIZED*)
+fun broadcast cond = (*requires SYNCHRONIZED*)
   ConditionVar.broadcast cond;
 
+fun broadcast_work () = (*requires SYNCHRONIZED*)
+ (ConditionVar.broadcast work_available;
+  ConditionVar.broadcast work_finished);
+
 end;
 
 
-(* worker activity *)
-
-fun count_active ws =
-  fold (fn (_, active) => fn i => if active then i + 1 else i) ws 0;
-
-fun trace_active () = Multithreading.tracing 1 (fn () =>
-  let
-    val ws = ! workers;
-    val m = string_of_int (length ws);
-    val n = string_of_int (count_active ws);
-  in "SCHEDULE: " ^ m ^ " workers, " ^ n ^ " active" end);
-
-fun change_active active = (*requires SYNCHRONIZED*)
-  change workers (AList.update Thread.equal (Thread.self (), active));
-
-fun overloaded () =
-  count_active (! workers) > Multithreading.max_threads_value ();
-
-
 (* execute future jobs *)
 
 fun future_job group (e: unit -> 'a) =
   let
-    val result = ref (NONE: 'a Exn.result option);
+    val result = Synchronized.var "future" (NONE: 'a Exn.result option);
     fun job ok =
       let
         val res =
           if ok then
-            Exn.capture
-              (Multithreading.with_attributes Multithreading.restricted_interrupts
-                (fn _ => fn () => e ())) ()
+            Exn.capture (fn () =>
+             (Thread.testInterrupt ();
+              Multithreading.with_attributes Multithreading.restricted_interrupts
+                (fn _ => fn () => e ())) ()) ()
           else Exn.Exn Exn.Interrupt;
-        val _ = result := SOME res;
+        val _ = Synchronized.change result (K (SOME res));
       in
         (case res of
           Exn.Exn exn => (Task_Queue.cancel_group group exn; false)
@@ -174,38 +162,51 @@
   in (result, job) end;
 
 fun do_cancel group = (*requires SYNCHRONIZED*)
-  change canceled (insert Task_Queue.eq_group group);
+ (change canceled (insert Task_Queue.eq_group group); broadcast scheduler_event);
 
 fun execute name (task, group, jobs) =
   let
-    val _ = trace_active ();
     val valid = not (Task_Queue.is_canceled group);
     val ok = setmp_thread_data (name, task, group) (fn () =>
       fold (fn job => fn ok => job valid andalso ok) jobs true) ();
-    val _ = SYNCHRONIZED "execute" (fn () =>
-     (change queue (Task_Queue.finish task);
-      if ok then ()
-      else if Task_Queue.cancel (! queue) group then ()
-      else do_cancel group;
-      notify_all ()));
+    val _ = SYNCHRONIZED "finish" (fn () =>
+      let
+        val maximal = change_result queue (Task_Queue.finish task);
+        val _ =
+          if ok then ()
+          else if Task_Queue.cancel (! queue) group then ()
+          else do_cancel group;
+        val _ = broadcast work_finished;
+        val _ = if maximal then () else broadcast work_available;
+      in () end);
   in () end;
 
 
+(* worker activity *)
+
+fun count_active () = (*requires SYNCHRONIZED*)
+  fold (fn (_, active) => fn i => if active then i + 1 else i) (! workers) 0;
+
+fun change_active active = (*requires SYNCHRONIZED*)
+  change workers (AList.update Thread.equal (Thread.self (), active));
+
+
 (* worker threads *)
 
-fun worker_wait () = (*requires SYNCHRONIZED*)
-  (change_active false; wait (); change_active true);
+fun worker_wait cond = (*requires SYNCHRONIZED*)
+ (change_active false; wait cond; change_active true);
 
 fun worker_next () = (*requires SYNCHRONIZED*)
   if ! excessive > 0 then
     (dec excessive;
      change workers (filter_out (fn (thread, _) => Thread.equal (thread, Thread.self ())));
-     notify_all ();
+     broadcast scheduler_event;
      NONE)
-  else if overloaded () then (worker_wait (); worker_next ())
+  else if count_active () > Multithreading.max_threads_value () then
+    (worker_wait scheduler_event; worker_next ())
   else
-    (case change_result queue Task_Queue.dequeue of
-      NONE => (worker_wait (); worker_next ())
+    (case change_result queue (Task_Queue.dequeue (Thread.self ())) of
+      NONE => (worker_wait work_available; worker_next ())
     | some => some);
 
 fun worker_loop name =
@@ -214,37 +215,51 @@
   | SOME work => (execute name work; worker_loop name));
 
 fun worker_start name = (*requires SYNCHRONIZED*)
-  change workers (cons (SimpleThread.fork false (fn () => worker_loop name), true));
+  change workers (cons (SimpleThread.fork false (fn () =>
+     (broadcast scheduler_event; worker_loop name)), true));
 
 
 (* scheduler *)
 
+val last_status = ref Time.zeroTime;
+val next_status = Time.fromMilliseconds 500;
+val next_round = Time.fromMilliseconds 50;
+
 fun scheduler_next () = (*requires SYNCHRONIZED*)
   let
-    (*queue status*)
-    val _ = Multithreading.tracing 1 (fn () =>
-      let val {ready, pending, running} = Task_Queue.status (! queue) in
-        "SCHEDULE: " ^
-          string_of_int ready ^ " ready, " ^
-          string_of_int pending ^ " pending, " ^
-          string_of_int running ^ " running"
-      end);
+    (*queue and worker status*)
+    val _ =
+      let val now = Time.now () in
+        if Time.> (Time.+ (! last_status, next_status), now) then ()
+        else
+         (last_status := now; Multithreading.tracing 1 (fn () =>
+            let
+              val {ready, pending, running} = Task_Queue.status (! queue);
+              val total = length (! workers);
+              val active = count_active ();
+            in
+              "SCHEDULE: " ^
+                string_of_int ready ^ " ready, " ^
+                string_of_int pending ^ " pending, " ^
+                string_of_int running ^ " running; " ^
+                string_of_int total ^ " workers, " ^
+                string_of_int active ^ " active"
+            end))
+      end;
 
     (*worker threads*)
-    val ws = ! workers;
     val _ =
-      if forall (Thread.isActive o #1) ws then ()
+      if forall (Thread.isActive o #1) (! workers) then ()
       else
-        (case List.partition (Thread.isActive o #1) ws of
+        (case List.partition (Thread.isActive o #1) (! workers) of
           (_, []) => ()
-        | (active, inactive) =>
-            (workers := active; Multithreading.tracing 0 (fn () =>
-              "SCHEDULE: disposed " ^ string_of_int (length inactive) ^ " dead worker threads")));
-    val _ = trace_active ();
+        | (alive, dead) =>
+            (workers := alive; Multithreading.tracing 0 (fn () =>
+              "SCHEDULE: disposed " ^ string_of_int (length dead) ^ " dead worker threads")));
 
     val m = if ! do_shutdown then 0 else Multithreading.max_threads_value ();
     val mm = (m * 3) div 2;
-    val l = length ws;
+    val l = length (! workers);
     val _ = excessive := l - mm;
     val _ =
       if mm > l then
@@ -252,16 +267,21 @@
       else ();
 
     (*canceled groups*)
-    val _ = change canceled (filter_out (Task_Queue.cancel (! queue)));
+    val _ =
+      if null (! canceled) then ()
+      else (change canceled (filter_out (Task_Queue.cancel (! queue))); broadcast_work ());
+
+    (*delay loop*)
+    val _ = wait_interruptible scheduler_event next_round
+      handle Exn.Interrupt =>
+        (Multithreading.tracing 1 (fn () => "Interrupt");
+          List.app do_cancel (Task_Queue.cancel_all (! queue)));
 
     (*shutdown*)
-    val continue = not (! do_shutdown andalso null ws);
+    val _ = if Task_Queue.is_empty (! queue) then do_shutdown := true else ();
+    val continue = not (! do_shutdown andalso null (! workers));
     val _ = if continue then () else scheduler := NONE;
-
-    val _ = notify_all ();
-    val _ = interruptible (fn () =>
-        wait_timeout (Time.fromMilliseconds (if null (! canceled) then 1000 else 50))) ()
-      handle Exn.Interrupt => List.app do_cancel (Task_Queue.cancel_all (! queue));
+    val _ = broadcast scheduler_event;
   in continue end;
 
 fun scheduler_loop () =
@@ -270,11 +290,10 @@
 fun scheduler_active () = (*requires SYNCHRONIZED*)
   (case ! scheduler of NONE => false | SOME thread => Thread.isActive thread);
 
-fun scheduler_check name = SYNCHRONIZED name (fn () =>
-  if not (scheduler_active ()) then
-    (do_shutdown := false; scheduler := SOME (SimpleThread.fork false scheduler_loop))
-  else if ! do_shutdown then error "Scheduler shutdown in progress"
-  else ());
+fun scheduler_check () = (*requires SYNCHRONIZED*)
+ (do_shutdown := false;
+  if scheduler_active () then ()
+  else scheduler := SOME (SimpleThread.fork false scheduler_loop));
 
 
 
@@ -284,15 +303,17 @@
 
 fun fork_future opt_group deps pri e =
   let
-    val _ = scheduler_check "future check";
-
     val group =
       (case opt_group of
         SOME group => group
       | NONE => Task_Queue.new_group (worker_group ()));
     val (result, job) = future_job group e;
-    val task = SYNCHRONIZED "future" (fn () =>
-      change_result queue (Task_Queue.enqueue group deps pri job) before notify_all ());
+    val task = SYNCHRONIZED "enqueue" (fn () =>
+      let
+        val (task, minimal) = change_result queue (Task_Queue.enqueue group deps pri job);
+        val _ = if minimal then signal work_available else ();
+        val _ = scheduler_check ();
+      in task end);
   in Future {task = task, group = group, result = result} end;
 
 fun fork e = fork_future NONE [] 0 e;
@@ -312,35 +333,33 @@
       Exn.Exn (Exn.EXCEPTIONS (Exn.flatten_list (Task_Queue.group_status (group_of x))))
   | SOME res => res);
 
+fun join_wait x =
+  Synchronized.guarded_access (result_of x) (fn NONE => NONE | some => SOME ((), some));
+
 fun join_next deps = (*requires SYNCHRONIZED*)
-  if overloaded () then (worker_wait (); join_next deps)
-  else change_result queue (Task_Queue.dequeue_towards deps);
+  if null deps then NONE
+  else
+    (case change_result queue (Task_Queue.dequeue_towards (Thread.self ()) deps) of
+      (NONE, []) => NONE
+    | (NONE, deps') => (worker_wait work_finished; join_next deps')
+    | (SOME work, deps') => SOME (work, deps'));
 
-fun join_deps deps =
+fun join_work deps =
   (case SYNCHRONIZED "join" (fn () => join_next deps) of
     NONE => ()
-  | SOME (work, deps') => (execute "join" work; join_deps deps'));
+  | SOME (work, deps') => (execute "join" work; join_work deps'));
 
 in
 
 fun join_results xs =
   if forall is_finished xs then map get_result xs
+  else if Multithreading.self_critical () then
+    error "Cannot join future values within critical section"
   else uninterruptible (fn _ => fn () =>
-    let
-      val _ = scheduler_check "join check";
-      val _ = Multithreading.self_critical () andalso
-        error "Cannot join future values within critical section";
-
-      val worker = is_worker ();
-      val _ = if worker then join_deps (map task_of xs) else ();
-
-      fun join_wait x =
-        if SYNCHRONIZED "join_wait" (fn () =>
-          is_finished x orelse (if worker then worker_wait () else wait (); false))
-        then () else join_wait x;
-      val _ = List.app join_wait xs;
-
-    in map get_result xs end) ();
+     (if is_worker ()
+      then join_work (map task_of xs)
+      else List.app join_wait xs;
+      map get_result xs)) ();
 
 end;
 
@@ -352,13 +371,11 @@
 
 fun map_future f x =
   let
-    val _ = scheduler_check "map_future check";
-
     val task = task_of x;
     val group = Task_Queue.new_group (SOME (group_of x));
     val (result, job) = future_job group (fn () => f (join x));
 
-    val extended = SYNCHRONIZED "map_future" (fn () =>
+    val extended = SYNCHRONIZED "extend" (fn () =>
       (case Task_Queue.extend task job (! queue) of
         SOME queue' => (queue := queue'; true)
       | NONE => false));
@@ -372,38 +389,26 @@
 
 fun interruptible_task f x =
   if Multithreading.available then
+   (Thread.testInterrupt ();
     Multithreading.with_attributes
       (if is_worker ()
        then Multithreading.restricted_interrupts
        else Multithreading.regular_interrupts)
-      (fn _ => f) x
+      (fn _ => fn x => f x) x)
   else interruptible f x;
 
-(*interrupt: permissive signal, may get ignored*)
-fun interrupt_task id = SYNCHRONIZED "interrupt"
-  (fn () => Task_Queue.interrupt_external (! queue) id);
-
 (*cancel: present and future group members will be interrupted eventually*)
-fun cancel_group group =
- (scheduler_check "cancel check";
-  SYNCHRONIZED "cancel" (fn () => (do_cancel group; notify_all ())));
-
+fun cancel_group group = SYNCHRONIZED "cancel" (fn () => do_cancel group);
 fun cancel x = cancel_group (group_of x);
 
 
-(** global join and shutdown **)
+(* shutdown *)
 
 fun shutdown () =
   if Multithreading.available then
-   (scheduler_check "shutdown check";
     SYNCHRONIZED "shutdown" (fn () =>
-     (while not (scheduler_active ()) do wait ();
-      while not (Task_Queue.is_empty (! queue)) do wait ();
-      do_shutdown := true;
-      notify_all ();
-      while not (null (! workers)) do wait ();
-      while scheduler_active () do wait ();
-      OS.Process.sleep (Time.fromMilliseconds 300))))
+     while scheduler_active () do
+      (wait scheduler_event; broadcast_work ()))
   else ();
 
 
--- a/src/Pure/Concurrent/par_list.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Concurrent/par_list.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -27,10 +27,8 @@
 struct
 
 fun raw_map f xs =
-  if Future.enabled () then
-    let val group = Task_Queue.new_group (Future.worker_group ())
-    in Future.join_results (map (fn x => Future.fork_group group (fn () => f x)) xs) end
-  else map (Exn.capture f) xs;
+  let val group = Task_Queue.new_group (Future.worker_group ())
+  in Future.join_results (map (fn x => Future.fork_group group (fn () => f x)) xs) end;
 
 fun map f xs = Exn.release_first (raw_map f xs);
 
--- a/src/Pure/Concurrent/simple_thread.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Concurrent/simple_thread.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -28,13 +28,15 @@
     val immediate =
       if Mutex.trylock lock then true
       else
-       (Multithreading.tracing 3 (fn () => name ^ ": locking ...");
-        Mutex.lock lock;
-        Multithreading.tracing 3 (fn () => name ^ ": locked");
-        false);
+        let
+          val _ = Multithreading.tracing 5 (fn () => name ^ ": locking ...");
+          val time = Multithreading.real_time Mutex.lock lock;
+          val _ = Multithreading.tracing_time true time
+            (fn () => name ^ ": locked after " ^ Time.toString time);
+        in false end;
     val result = Exn.capture (restore_attributes e) ();
     val _ =
-      if immediate then () else Multithreading.tracing 3 (fn () => name ^ ": unlocking ...");
+      if immediate then () else Multithreading.tracing 5 (fn () => name ^ ": unlocking ...");
     val _ = Mutex.unlock lock;
   in result end) ());
 
--- a/src/Pure/Concurrent/synchronized.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Concurrent/synchronized.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -8,6 +8,7 @@
 sig
   type 'a var
   val var: string -> 'a -> 'a var
+  val peek: 'a var -> 'a
   val value: 'a var -> 'a
   val timed_access: 'a var -> ('a -> Time.time option) -> ('a -> ('b * 'a) option) -> 'b option
   val guarded_access: 'a var -> ('a -> ('b * 'a) option) -> 'b
@@ -32,6 +33,8 @@
   cond = ConditionVar.conditionVar (),
   var = ref x};
 
+fun peek (Var {var, ...}) = ! var;  (*unsynchronized!*)
+
 fun value (Var {name, lock, cond, var}) = SimpleThread.synchronized name lock (fn () => ! var);
 
 
--- a/src/Pure/Concurrent/task_queue.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Concurrent/task_queue.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -11,27 +11,25 @@
   val pri_of_task: task -> int
   val str_of_task: task -> string
   type group
+  val new_group: group option -> group
   val group_id: group -> int
   val eq_group: group * group -> bool
-  val new_group: group option -> group
+  val cancel_group: group -> exn -> unit
+  val is_canceled: group -> bool
   val group_status: group -> exn list
   val str_of_group: group -> string
   type queue
   val empty: queue
   val is_empty: queue -> bool
   val status: queue -> {ready: int, pending: int, running: int}
-  val enqueue: group -> task list -> int -> (bool -> bool) -> queue -> task * queue
-  val extend: task -> (bool -> bool) -> queue -> queue option
-  val dequeue: queue -> (task * group * (bool -> bool) list) option * queue
-  val dequeue_towards: task list -> queue ->
-    (((task * group * (bool -> bool) list) * task list) option * queue)
-  val interrupt: queue -> task -> unit
-  val interrupt_external: queue -> string -> unit
-  val is_canceled: group -> bool
-  val cancel_group: group -> exn -> unit
   val cancel: queue -> group -> bool
   val cancel_all: queue -> group list
-  val finish: task -> queue -> queue
+  val enqueue: group -> task list -> int -> (bool -> bool) -> queue -> (task * bool) * queue
+  val extend: task -> (bool -> bool) -> queue -> queue option
+  val dequeue: Thread.thread -> queue -> (task * group * (bool -> bool) list) option * queue
+  val dequeue_towards: Thread.thread -> task list -> queue ->
+    (((task * group * (bool -> bool) list) option * task list) * queue)
+  val finish: task -> queue -> bool * queue
 end;
 
 structure Task_Queue:> TASK_QUEUE =
@@ -54,11 +52,11 @@
 datatype group = Group of
  {parent: group option,
   id: serial,
-  status: exn list ref};
+  status: exn list Synchronized.var};
 
 fun make_group (parent, id, status) = Group {parent = parent, id = id, status = status};
 
-fun new_group parent = make_group (parent, serial (), ref []);
+fun new_group parent = make_group (parent, serial (), Synchronized.var "group" []);
 
 fun group_id (Group {id, ...}) = id;
 fun eq_group (group1, group2) = group_id group1 = group_id group2;
@@ -67,16 +65,22 @@
   id :: (case parent of NONE => [] | SOME group => group_ancestry group);
 
 
-fun cancel_group (Group {status, ...}) exn = CRITICAL (fn () =>
-  (case exn of
-    Exn.Interrupt => if null (! status) then status := [exn] else ()
-  | _ => change status (cons exn)));
+(* group status *)
 
-fun group_status (Group {parent, status, ...}) = (*non-critical*)
-  ! status @ (case parent of NONE => [] | SOME group => group_status group);
+fun cancel_group (Group {status, ...}) exn =
+  Synchronized.change status
+    (fn exns =>
+      (case exn of
+        Exn.Interrupt => if null exns then [exn] else exns
+      | _ => exn :: exns));
 
-fun is_canceled (Group {parent, status, ...}) = (*non-critical*)
-  not (null (! status)) orelse (case parent of NONE => false | SOME group => is_canceled group);
+fun is_canceled (Group {parent, status, ...}) =
+  not (null (Synchronized.value status)) orelse
+    (case parent of NONE => false | SOME group => is_canceled group);
+
+fun group_status (Group {parent, status, ...}) =
+  Synchronized.value status @
+    (case parent of NONE => [] | SOME group => group_status group);
 
 fun str_of_group group =
   (is_canceled group ? enclose "(" ")") (string_of_int (group_id group));
@@ -97,6 +101,9 @@
 fun add_job task dep (jobs: jobs) =
   Task_Graph.add_edge (dep, task) jobs handle Task_Graph.UNDEF _ => jobs;
 
+fun get_deps (jobs: jobs) task =
+  Task_Graph.imm_preds jobs task handle Task_Graph.UNDEF _ => [];
+
 
 (* queue of grouped jobs *)
 
@@ -157,14 +164,16 @@
       |> fold (fn gid => Inttab.cons_list (gid, task)) (group_ancestry group);
     val jobs' = jobs
       |> Task_Graph.new_node (task, (group, Job [job]))
-      |> fold (add_job task) deps;
+      |> fold (add_job task) deps
+      |> fold (fold (add_job task) o get_deps jobs) deps;
+    val minimal = null (get_deps jobs' task);
     val cache' =
       (case cache of
         Result last =>
           if task_ord (last, task) = LESS
           then cache else Unknown
       | _ => Unknown);
-  in (task, make_queue groups' jobs' cache') end;
+  in ((task, minimal), make_queue groups' jobs' cache') end;
 
 fun extend task job (Queue {groups, jobs, cache}) =
   (case try (get_job jobs) task of
@@ -174,7 +183,7 @@
 
 (* dequeue *)
 
-fun dequeue (queue as Queue {groups, jobs, cache}) =
+fun dequeue thread (queue as Queue {groups, jobs, cache}) =
   let
     fun ready (task, ((group, Job list), ([], _))) = SOME (task, group, rev list)
       | ready _ = NONE;
@@ -183,7 +192,7 @@
         NONE => (NONE, make_queue groups jobs No_Result)
       | SOME (result as (task, _, _)) =>
           let
-            val jobs' = set_job task (Running (Thread.self ())) jobs;
+            val jobs' = set_job task (Running thread) jobs;
             val cache' = Result task;
           in (SOME result, make_queue groups jobs' cache') end);
   in
@@ -196,49 +205,32 @@
 
 (* dequeue_towards -- adhoc dependencies *)
 
-fun dequeue_towards deps (queue as Queue {groups, jobs, ...}) =
+fun dequeue_towards thread deps (queue as Queue {groups, jobs, ...}) =
   let
     fun ready task =
       (case Task_Graph.get_node jobs task of
         (group, Job list) =>
-          if null (Task_Graph.imm_preds jobs task)
+          if null (get_deps jobs task)
           then SOME (task, group, rev list)
           else NONE
       | _ => NONE);
-
     val tasks = filter (can (Task_Graph.get_node jobs)) deps;
     fun result (res as (task, _, _)) =
       let
-        val jobs' = set_job task (Running (Thread.self ())) jobs;
+        val jobs' = set_job task (Running thread) jobs;
         val cache' = Unknown;
-      in (SOME (res, tasks), make_queue groups jobs' cache') end;
+      in ((SOME res, tasks), make_queue groups jobs' cache') end;
   in
     (case get_first ready tasks of
       SOME res => result res
     | NONE =>
-        (case get_first ready (Task_Graph.all_preds jobs tasks) of
+        (case get_first (get_first ready o get_deps jobs) tasks of
           SOME res => result res
-        | NONE => (NONE, queue)))
+        | NONE => ((NONE, tasks), queue)))
   end;
 
 
-(* sporadic interrupts *)
-
-fun interrupt (Queue {jobs, ...}) task =
-  (case try (get_job jobs) task of
-    SOME (Running thread) => SimpleThread.interrupt thread
-  | _ => ());
-
-fun interrupt_external (queue as Queue {jobs, ...}) str =
-  (case Int.fromString str of
-    SOME i =>
-      (case Task_Graph.get_first NONE
-          (fn (task as Task (_, j), _) => if i = j then SOME task else NONE) jobs
-        of SOME task => interrupt queue task | NONE => ())
-  | NONE => ());
-
-
-(* termination *)
+(* finish *)
 
 fun finish task (Queue {groups, jobs, cache}) =
   let
@@ -246,9 +238,8 @@
     val groups' = groups
       |> fold (fn gid => Inttab.remove_list (op =) (gid, task)) (group_ancestry group);
     val jobs' = Task_Graph.del_node task jobs;
-    val cache' =
-      if null (Task_Graph.imm_succs jobs task) then cache
-      else Unknown;
-  in make_queue groups' jobs' cache' end;
+    val maximal = null (Task_Graph.imm_succs jobs task);
+    val cache' = if maximal then cache else Unknown;
+  in (maximal, make_queue groups' jobs' cache') end;
 
 end;
--- a/src/Pure/General/position.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/General/position.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -36,7 +36,6 @@
   val range: T -> T -> range
   val thread_data: unit -> T
   val setmp_thread_data: T -> ('a -> 'b) -> 'a -> 'b
-  val setmp_thread_data_seq: T -> ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq
 end;
 
 structure Position: POSITION =
@@ -176,9 +175,6 @@
   if ! Output.debugging then f x
   else Library.setmp_thread_data tag (thread_data ()) pos f x;
 
-fun setmp_thread_data_seq pos f x =
-  setmp_thread_data pos f x |> Seq.wrap (fn pull => setmp_thread_data pos pull ());
-
 end;
 
 end;
--- a/src/Pure/General/seq.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/General/seq.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -33,7 +33,6 @@
   val lift: ('a -> 'b -> 'c) -> 'a seq -> 'b -> 'c seq
   val lifts: ('a -> 'b -> 'c seq) -> 'a seq -> 'b -> 'c seq
   val singleton: ('a list -> 'b list seq) -> 'a -> 'b seq
-  val wrap: ((unit -> ('a * 'a seq) option) -> ('a * 'a seq) option) -> 'a seq -> 'a seq
   val print: (int -> 'a -> unit) -> int -> 'a seq -> unit
   val it_right : ('a * 'b seq -> 'b seq) -> 'a seq * 'b seq -> 'b seq
   val succeed: 'a -> 'a seq
@@ -170,13 +169,6 @@
 
 fun singleton f x = f [x] |> map (fn [y] => y | _ => raise Empty);
 
-(*wrapped lazy evaluation*)
-fun wrap f xq =
-  make (fn () =>
-    (case f (fn () => pull xq) of
-      NONE => NONE
-    | SOME (x, xq') => SOME (x, wrap f xq')));
-
 (*print a sequence, up to "count" elements*)
 fun print print_elem count =
   let
@@ -200,7 +192,7 @@
 
 
 
-(** sequence functions **)      (*cf. Pure/tctical.ML*)
+(** sequence functions **)      (*cf. Pure/tactical.ML*)
 
 fun succeed x = single x;
 fun fail _ = empty;
--- a/src/Pure/IsaMakefile	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/IsaMakefile	Wed Jul 29 09:06:49 2009 +0200
@@ -90,11 +90,11 @@
   Tools/find_theorems.ML Tools/named_thms.ML Tools/xml_syntax.ML	\
   assumption.ML axclass.ML codegen.ML config.ML conjunction.ML		\
   consts.ML context.ML context_position.ML conv.ML defs.ML display.ML	\
-  display_goal.ML drule.ML envir.ML facts.ML goal.ML interpretation.ML	\
+  drule.ML envir.ML facts.ML goal.ML goal_display.ML interpretation.ML	\
   item_net.ML library.ML logic.ML meta_simplifier.ML more_thm.ML	\
   morphism.ML name.ML net.ML old_goals.ML old_term.ML pattern.ML	\
   primitive_defs.ML proofterm.ML pure_setup.ML pure_thy.ML search.ML	\
-  sign.ML simplifier.ML sorts.ML subgoal.ML tactic.ML tctical.ML	\
+  sign.ML simplifier.ML sorts.ML subgoal.ML tactic.ML tactical.ML	\
   term.ML term_ord.ML term_subst.ML theory.ML thm.ML type.ML		\
   type_infer.ML unify.ML variable.ML
 	@./mk
--- a/src/Pure/Isar/class.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Isar/class.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -101,7 +101,7 @@
 
 (* reading and processing class specifications *)
 
-fun prep_class_elems prep_decl thy supexpr sups proto_base_sort raw_elems =
+fun prep_class_elems prep_decl thy sups proto_base_sort raw_elems =
   let
 
     (* user space type system: only permits 'a type variable, improves towards 'a *)
@@ -129,16 +129,19 @@
                   ^ Syntax.string_of_sort_global thy a_sort ^ ".")
             | _ => error "Multiple type variables in class specification.";
       in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
-    fun add_typ_check level name f = Context.proof_map (Syntax.add_typ_check level name (fn xs => fn ctxt =>
-      let val xs' = f xs in if eq_list (op =) (xs, xs') then NONE else SOME (xs', ctxt) end));
+    fun add_typ_check level name f = Context.proof_map
+      (Syntax.add_typ_check level name (fn xs => fn ctxt =>
+        let val xs' = f xs in if eq_list (op =) (xs, xs') then NONE else SOME (xs', ctxt) end));
 
-    (* preprocessing elements, retrieving base sort from type-checked elements *)
+    (* preprocessing elements, retrieving base sort from typechecked elements *)
     val init_class_body = fold (ProofContext.add_const_constraint o apsnd SOME) base_constraints
       #> redeclare_operations thy sups
       #> add_typ_check 10 "reject_bcd_etc" reject_bcd_etc
       #> add_typ_check ~10 "singleton_fixate" (singleton_fixate thy (Sign.classes_of thy));
-    val ((_, _, inferred_elems), _) = ProofContext.init thy
-      |> prep_decl supexpr init_class_body raw_elems;
+    val raw_supexpr = (map (fn sup => (sup, (("", false),
+      Expression.Positional []))) sups, []);
+    val ((raw_supparams, _, inferred_elems), _) = ProofContext.init thy
+      |> prep_decl raw_supexpr init_class_body raw_elems;
     fun fold_element_types f (Element.Fixes fxs) = fold (fn (_, SOME T, _) => f T) fxs
       | fold_element_types f (Element.Constrains cnstrs) = fold (f o snd) cnstrs
       | fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) =>
@@ -151,9 +154,16 @@
       case (fold o fold_element_types) Term.add_tfreesT inferred_elems []
        of [] => error "No type variable in class specification"
         | [(_, sort)] => sort
-        | _ => error "Multiple type variables in class specification"
+        | _ => error "Multiple type variables in class specification";
+    val supparams = map (fn ((c, T), _) =>
+      (c, map_atyps (K (TFree (Name.aT, base_sort))) T)) raw_supparams;
+    val supparam_names = map fst supparams;
+    fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c);
+    val supexpr = (map (fn sup => (sup, (("", false),
+      Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup))))) sups,
+        map (fn (c, T) => (Binding.name c, SOME T, NoSyn)) supparams);
 
-  in (base_sort, inferred_elems) end;
+  in (base_sort, supparam_names, supexpr, inferred_elems) end;
 
 val cert_class_elems = prep_class_elems Expression.cert_declaration;
 val read_class_elems = prep_class_elems Expression.cert_read_declaration;
@@ -168,23 +178,21 @@
     val _ = case filter_out (is_class thy) sups
      of [] => ()
       | no_classes => error ("No (proper) classes: " ^ commas (map quote no_classes));
-    val supparams = (map o apsnd) (snd o snd) (these_params thy sups);
-    val supparam_names = map fst supparams;
-    val _ = if has_duplicates (op =) supparam_names
+    val raw_supparams = (map o apsnd) (snd o snd) (these_params thy sups);
+    val raw_supparam_names = map fst raw_supparams;
+    val _ = if has_duplicates (op =) raw_supparam_names
       then error ("Duplicate parameter(s) in superclasses: "
-        ^ (commas o map quote o duplicates (op =)) supparam_names)
+        ^ (commas o map quote o duplicates (op =)) raw_supparam_names)
       else ();
-    val supexpr = (map (fn sup => (sup, (("", false), Expression.Positional [])))
-      sups, []);
     val given_basesort = fold inter_sort (map (base_sort thy) sups) [];
 
     (* infer types and base sort *)
-    val (base_sort, inferred_elems) = prep_class_elems thy supexpr sups
-      given_basesort raw_elems;
-    val sup_sort = inter_sort base_sort sups
+    val (base_sort, supparam_names, supexpr, inferred_elems) =
+      prep_class_elems thy sups given_basesort raw_elems;
+    val sup_sort = inter_sort base_sort sups;
 
     (* process elements as class specification *)
-    val class_ctxt = begin sups base_sort (ProofContext.init thy)
+    val class_ctxt = begin sups base_sort (ProofContext.init thy);
     val ((_, _, syntax_elems), _) = class_ctxt
       |> Expression.cert_declaration supexpr I inferred_elems;
     fun check_vars e vs = if null vs
@@ -204,11 +212,11 @@
       | fork_syn x = pair x;
     val (elems, global_syntax) = fold_map fork_syn syntax_elems [];
     val constrain = Element.Constrains ((map o apsnd o map_atyps)
-      (K (TFree (Name.aT, base_sort))) supparams);
+      (K (TFree (Name.aT, base_sort))) raw_supparams);
       (*FIXME perhaps better: control type variable by explicit
       parameter instantiation of import expression*)
 
-  in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (constrain :: elems, global_syntax)) end;
+  in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), ((*constrain :: *)elems, global_syntax)) end;
 
 val cert_class_spec = prep_class_spec (K I) cert_class_elems;
 val read_class_spec = prep_class_spec Sign.intern_class read_class_elems;
@@ -216,14 +224,14 @@
 
 (* class establishment *)
 
-fun add_consts class base_sort sups supparams global_syntax thy =
+fun add_consts class base_sort sups supparam_names global_syntax thy =
   let
     (*FIXME simplify*)
-    val supconsts = supparams
+    val supconsts = supparam_names
       |> AList.make (snd o the o AList.lookup (op =) (these_params thy sups))
       |> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]);
     val all_params = Locale.params_of thy class;
-    val raw_params = (snd o chop (length supparams)) all_params;
+    val raw_params = (snd o chop (length supparam_names)) all_params;
     fun add_const ((raw_c, raw_ty), _) thy =
       let
         val b = Binding.name raw_c;
@@ -246,7 +254,7 @@
     |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
   end;
 
-fun adjungate_axclass bname class base_sort sups supsort supparams global_syntax thy =
+fun adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax thy =
   let
     (*FIXME simplify*)
     fun globalize param_map = map_aterms
@@ -260,7 +268,7 @@
       | [thm] => SOME thm;
   in
     thy
-    |> add_consts class base_sort sups supparams global_syntax
+    |> add_consts class base_sort sups supparam_names global_syntax
     |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
           (map (fst o snd) params)
           [(Thm.empty_binding, Option.map (globalize param_map) raw_pred |> the_list)]
@@ -270,16 +278,16 @@
     #> pair (param_map, params, assm_axiom)))
   end;
 
-fun gen_class prep_spec bname raw_supclasses raw_elems thy =
+fun gen_class prep_class_spec bname raw_supclasses raw_elems thy =
   let
     val class = Sign.full_name thy bname;
-    val (((sups, supparams), (supsort, base_sort, supexpr)), (elems, global_syntax)) =
-      prep_spec thy raw_supclasses raw_elems;
+    val (((sups, supparam_names), (supsort, base_sort, supexpr)), (elems, global_syntax)) =
+      prep_class_spec thy raw_supclasses raw_elems;
   in
     thy
     |> Expression.add_locale bname Binding.empty supexpr elems
     |> snd |> LocalTheory.exit_global
-    |> adjungate_axclass bname class base_sort sups supsort supparams global_syntax
+    |> adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax
     ||> Theory.checkpoint
     |-> (fn (param_map, params, assm_axiom) =>
        `(fn thy => calculate thy class sups base_sort param_map assm_axiom)
--- a/src/Pure/Isar/element.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Isar/element.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -213,7 +213,7 @@
   let
     val ((xs, prop'), ctxt') = Variable.focus prop ctxt;
     val As = Logic.strip_imp_prems (Thm.term_of prop');
-  in ((Binding.empty, (map (fix o Term.dest_Free o Thm.term_of) xs, As)), ctxt') end;
+  in ((Binding.empty, (map (fix o Term.dest_Free o Thm.term_of o #2) xs, As)), ctxt') end;
 
 in
 
@@ -272,7 +272,7 @@
   Proof.refine (Method.Basic (K (RAW_METHOD
     (K (ALLGOALS
       (CONJUNCTS (ALLGOALS
-        (CONJUNCTS (TRYALL (Tactic.rtac Drule.protectI)))))))), Position.none));
+        (CONJUNCTS (TRYALL (Tactic.rtac Drule.protectI))))))))));
 
 fun gen_witness_proof proof after_qed wit_propss eq_props =
   let
--- a/src/Pure/Isar/method.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Isar/method.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -15,7 +15,7 @@
 sig
   include BASIC_METHOD
   type method
-  val apply: Position.T -> (Proof.context -> method) -> Proof.context -> thm list -> cases_tactic
+  val apply: (Proof.context -> method) -> Proof.context -> thm list -> cases_tactic
   val RAW_METHOD_CASES: (thm list -> cases_tactic) -> method
   val RAW_METHOD: (thm list -> tactic) -> method
   val METHOD_CASES: (thm list -> cases_tactic) -> method
@@ -55,7 +55,7 @@
   val raw_tactic: string * Position.T -> Proof.context -> method
   type src = Args.src
   datatype text =
-    Basic of (Proof.context -> method) * Position.T |
+    Basic of Proof.context -> method |
     Source of src |
     Source_i of src |
     Then of text list |
@@ -69,7 +69,7 @@
   val this_text: text
   val done_text: text
   val sorry_text: bool -> text
-  val finish_text: text option * bool -> Position.T -> text
+  val finish_text: text option * bool -> text
   val print_methods: theory -> unit
   val intern: theory -> xstring -> string
   val defined: theory -> string -> bool
@@ -106,8 +106,7 @@
 
 datatype method = Meth of thm list -> cases_tactic;
 
-fun apply pos meth_fun ctxt facts goal = Position.setmp_thread_data_seq pos
-  (fn () => let val Meth meth = meth_fun ctxt in meth facts goal end) ();
+fun apply meth ctxt = let val Meth m = meth ctxt in m end;
 
 val RAW_METHOD_CASES = Meth;
 
@@ -297,7 +296,7 @@
 type src = Args.src;
 
 datatype text =
-  Basic of (Proof.context -> method) * Position.T |
+  Basic of Proof.context -> method |
   Source of src |
   Source_i of src |
   Then of text list |
@@ -306,15 +305,15 @@
   Repeat1 of text |
   SelectGoals of int * text;
 
-fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r)), Position.none);
-val succeed_text = Basic (K succeed, Position.none);
+fun primitive_text r = Basic (K (SIMPLE_METHOD (PRIMITIVE r)));
+val succeed_text = Basic (K succeed);
 val default_text = Source (Args.src (("default", []), Position.none));
-val this_text = Basic (K this, Position.none);
-val done_text = Basic (K (SIMPLE_METHOD all_tac), Position.none);
-fun sorry_text int = Basic (cheating int, Position.none);
+val this_text = Basic (K this);
+val done_text = Basic (K (SIMPLE_METHOD all_tac));
+fun sorry_text int = Basic (cheating int);
 
-fun finish_text (NONE, immed) pos = Basic (close immed, pos)
-  | finish_text (SOME txt, immed) pos = Then [txt, Basic (close immed, pos)];
+fun finish_text (NONE, immed) = Basic (close immed)
+  | finish_text (SOME txt, immed) = Then [txt, Basic (close immed)];
 
 
 (* method definitions *)
@@ -477,8 +476,8 @@
 
 end;
 
-structure BasicMethod: BASIC_METHOD = Method;
-open BasicMethod;
+structure Basic_Method: BASIC_METHOD = Method;
+open Basic_Method;
 
 val RAW_METHOD_CASES = Method.RAW_METHOD_CASES;
 val RAW_METHOD = Method.RAW_METHOD;
--- a/src/Pure/Isar/obtain.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Isar/obtain.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -44,7 +44,7 @@
   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 ->
-    (cterm list * 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
 end;
@@ -200,7 +200,7 @@
     val obtain_rule = Thm.forall_elim (cert (Logic.varify (Free thesis_var))) rule';
     val ((params, stmt), fix_ctxt) = Variable.focus (Thm.cprem_of obtain_rule 1) ctxt';
     val (prems, ctxt'') =
-      Assumption.add_assms (obtain_export fix_ctxt obtain_rule params)
+      Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params))
         (Drule.strip_imp_prems stmt) fix_ctxt;
   in ((params, prems), ctxt'') end;
 
--- a/src/Pure/Isar/proof.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Isar/proof.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -344,8 +344,8 @@
             (if mode <> Backward orelse null using then NONE else SOME using) @
           [Pretty.str ("goal" ^
             description [levels_up (i div 2), subgoals (Thm.nprems_of goal)] ^ ":")] @
-          Display_Goal.pretty_goals ctxt Markup.subgoal
-            (true, ! show_main_goal) (! Display.goals_limit) goal @
+          Goal_Display.pretty_goals ctxt
+            {total = true, main = ! show_main_goal, maxgoals = ! Display.goals_limit} goal @
           (map (fn msg => Position.setmp_thread_data pos msg ()) (rev messages))
       | prt_goal NONE = [];
 
@@ -365,8 +365,10 @@
   end;
 
 fun pretty_goals main state =
-  let val (ctxt, (_, goal)) = get_goal state
-  in Display_Goal.pretty_goals ctxt Markup.none (false, main) (! Display.goals_limit) goal end;
+  let val (ctxt, (_, goal)) = get_goal state in
+    Goal_Display.pretty_goals ctxt
+      {total = false, main = main, maxgoals = ! Display.goals_limit} goal
+  end;
 
 
 
@@ -384,13 +386,13 @@
 fun goal_cases st =
   RuleCases.make_common true (Thm.theory_of_thm st, Thm.prop_of st) (map (rpair []) (goals st));
 
-fun apply_method current_context pos meth state =
+fun apply_method current_context meth state =
   let
     val (goal_ctxt, (_, {statement, messages = _, using, goal, before_qed, after_qed})) =
       find_goal state;
     val ctxt = if current_context then context_of state else goal_ctxt;
   in
-    Method.apply pos meth ctxt using goal |> Seq.map (fn (meth_cases, goal') =>
+    Method.apply meth ctxt using goal |> Seq.map (fn (meth_cases, goal') =>
       state
       |> map_goal
           (ProofContext.add_cases false (no_goal_cases goal @ goal_cases goal') #>
@@ -408,16 +410,15 @@
     |> Seq.maps meth
     |> Seq.maps (fn state' => state'
       |> Seq.lift set_goal (Goal.retrofit 1 n (#2 (#2 (get_goal state'))) goal))
-    |> Seq.maps (apply_method true Position.none (K Method.succeed)));
+    |> Seq.maps (apply_method true (K Method.succeed)));
 
 fun apply_text cc text state =
   let
     val thy = theory_of state;
-    val pos_of = #2 o Args.dest_src;
 
-    fun eval (Method.Basic (m, pos)) = apply_method cc pos m
-      | eval (Method.Source src) = apply_method cc (pos_of src) (Method.method thy src)
-      | eval (Method.Source_i src) = apply_method cc (pos_of src) (Method.method_i thy src)
+    fun eval (Method.Basic m) = apply_method cc m
+      | eval (Method.Source src) = apply_method cc (Method.method thy src)
+      | eval (Method.Source_i src) = apply_method cc (Method.method_i thy src)
       | eval (Method.Then txts) = Seq.EVERY (map eval txts)
       | eval (Method.Orelse txts) = Seq.FIRST (map eval txts)
       | eval (Method.Try txt) = Seq.TRY (eval txt)
@@ -431,7 +432,7 @@
 val refine_end = apply_text false;
 
 fun refine_insert [] = I
-  | refine_insert ths = Seq.hd o refine (Method.Basic (K (Method.insert ths), Position.none));
+  | refine_insert ths = Seq.hd o refine (Method.Basic (K (Method.insert ths)));
 
 end;
 
@@ -472,8 +473,8 @@
 
     val ngoals = Thm.nprems_of goal;
     val _ = ngoals = 0 orelse error (Pretty.string_of (Pretty.chunks
-      (Display_Goal.pretty_goals ctxt Markup.none
-          (true, ! show_main_goal) (! Display.goals_limit) goal @
+      (Goal_Display.pretty_goals ctxt
+          {total = true, main = ! show_main_goal, maxgoals = ! Display.goals_limit} goal @
         [Pretty.str (string_of_int ngoals ^ " unsolved goal(s)!")])));
 
     val extra_hyps = Assumption.extra_hyps ctxt goal;
@@ -748,7 +749,7 @@
   |> assert_current_goal true
   |> using_facts []
   |> `before_qed |-> (refine o the_default Method.succeed_text)
-  |> Seq.maps (refine (Method.finish_text txt (Position.thread_data ())));
+  |> Seq.maps (refine (Method.finish_text txt));
 
 fun check_result msg sq =
   (case Seq.pull sq of
@@ -760,8 +761,8 @@
 
 (* unstructured refinement *)
 
-fun defer i = assert_no_chain #> refine (Method.Basic (K (Method.defer i), Position.none));
-fun prefer i = assert_no_chain #> refine (Method.Basic (K (Method.prefer i), Position.none));
+fun defer i = assert_no_chain #> refine (Method.Basic (K (Method.defer i)));
+fun prefer i = assert_no_chain #> refine (Method.Basic (K (Method.prefer i)));
 
 fun apply text = assert_backward #> refine text #> Seq.map (using_facts []);
 fun apply_end text = assert_forward #> refine_end text;
@@ -787,7 +788,7 @@
 fun refine_terms n =
   refine (Method.Basic (K (RAW_METHOD
     (K (HEADGOAL (PRECISE_CONJUNCTS n
-      (HEADGOAL (CONJUNCTS (ALLGOALS (rtac Drule.termI)))))))), Position.none))
+      (HEADGOAL (CONJUNCTS (ALLGOALS (rtac Drule.termI))))))))))
   #> Seq.hd;
 
 in
@@ -830,7 +831,7 @@
     |> put_goal NONE
     |> enter_backward
     |> not (null vars) ? refine_terms (length goal_propss)
-    |> null props ? (refine (Method.Basic (Method.assumption, Position.none)) #> Seq.hd)
+    |> null props ? (refine (Method.Basic Method.assumption) #> Seq.hd)
   end;
 
 fun generic_qed after_ctxt state =
--- a/src/Pure/ML-Systems/multithreading.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/ML-Systems/multithreading.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -15,6 +15,8 @@
   include BASIC_MULTITHREADING
   val trace: int ref
   val tracing: int -> (unit -> string) -> unit
+  val tracing_time: bool -> Time.time -> (unit -> string) -> unit
+  val real_time: ('a -> unit) -> 'a -> Time.time
   val available: bool
   val max_threads: int ref
   val max_threads_value: unit -> int
@@ -31,10 +33,15 @@
 structure Multithreading: MULTITHREADING =
 struct
 
-(* options *)
+(* tracing *)
 
 val trace = ref (0: int);
 fun tracing _ _ = ();
+fun tracing_time _ _ _ = ();
+fun real_time f x = (f x; Time.zeroTime);
+
+
+(* options *)
 
 val available = false;
 val max_threads = ref (1: int);
--- a/src/Pure/ML-Systems/multithreading_polyml.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/ML-Systems/multithreading_polyml.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -27,7 +27,7 @@
 structure Multithreading: MULTITHREADING =
 struct
 
-(* options *)
+(* tracing *)
 
 val trace = ref 0;
 
@@ -36,6 +36,24 @@
   else (TextIO.output (TextIO.stdErr, (">>> " ^ msg () ^ "\n")); TextIO.flushOut TextIO.stdErr)
     handle _ (*sic*) => ();
 
+fun tracing_time detailed time =
+  tracing
+   (if not detailed then 5
+    else if Time.>= (time, Time.fromMilliseconds 1000) then 1
+    else if Time.>= (time, Time.fromMilliseconds 100) then 2
+    else if Time.>= (time, Time.fromMilliseconds 10) then 3
+    else if Time.>= (time, Time.fromMilliseconds 1) then 4 else 5);
+
+fun real_time f x =
+  let
+    val timer = Timer.startRealTimer ();
+    val () = f x;
+    val time = Timer.checkRealTimer timer;
+  in time end;
+
+
+(* options *)
+
 val available = true;
 
 val max_threads = ref 0;
@@ -92,8 +110,8 @@
     val _ = Thread.setAttributes orig_atts;
   in Exn.release result end;
 
-fun interruptible f =
-  with_attributes regular_interrupts (fn _ => fn x => f x);
+fun interruptible f x =
+  (Thread.testInterrupt (); with_attributes regular_interrupts (fn _ => fn x => f x) x);
 
 fun uninterruptible f =
   with_attributes no_interrupts (fn atts => fn x =>
@@ -205,22 +223,16 @@
 fun NAMED_CRITICAL name e =
   if self_critical () then e ()
   else
-    uninterruptible (fn restore_attributes => fn () =>
+    Exn.release (uninterruptible (fn restore_attributes => fn () =>
       let
         val name' = ! critical_name;
         val _ =
           if Mutex.trylock critical_lock then ()
           else
             let
-              val timer = Timer.startRealTimer ();
-              val _ = tracing 4 (fn () => "CRITICAL" ^ show name ^ show' name' ^ ": waiting");
-              val _ = Mutex.lock critical_lock;
-              val time = Timer.checkRealTimer timer;
-              val trace_time =
-                if Time.>= (time, Time.fromMilliseconds 1000) then 1
-                else if Time.>= (time, Time.fromMilliseconds 100) then 2
-                else if Time.>= (time, Time.fromMilliseconds 10) then 3 else 4;
-              val _ = tracing trace_time (fn () =>
+              val _ = tracing 5 (fn () => "CRITICAL" ^ show name ^ show' name' ^ ": waiting");
+              val time = real_time Mutex.lock critical_lock;
+              val _ = tracing_time true time (fn () =>
                 "CRITICAL" ^ show name ^ show' name' ^ ": passed after " ^ Time.toString time);
             in () end;
         val _ = critical_thread := SOME (Thread.self ());
@@ -229,7 +241,7 @@
         val _ = critical_name := "";
         val _ = critical_thread := NONE;
         val _ = Mutex.unlock critical_lock;
-      in Exn.release result end) ();
+      in result end) ());
 
 fun CRITICAL e = NAMED_CRITICAL "" e;
 
--- a/src/Pure/Proof/reconstruct.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/Proof/reconstruct.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -255,7 +255,7 @@
       let
         fun search env [] = error ("Unsolvable constraints:\n" ^
               Pretty.string_of (Pretty.chunks (map (fn (_, p, _) =>
-                Display_Goal.pretty_flexpair (Syntax.init_pretty_global thy) (pairself
+                Goal_Display.pretty_flexpair (Syntax.init_pretty_global thy) (pairself
                   (Envir.norm_term bigenv) p)) cs)))
           | search env ((u, p as (t1, t2), vs)::ps) =
               if u then
--- a/src/Pure/ROOT.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/ROOT.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -119,8 +119,8 @@
 use "morphism.ML";
 use "variable.ML";
 use "conv.ML";
-use "display_goal.ML";
-use "tctical.ML";
+use "goal_display.ML";
+use "tactical.ML";
 use "search.ML";
 use "tactic.ML";
 use "meta_simplifier.ML";
--- a/src/Pure/axclass.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/axclass.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -607,8 +607,10 @@
     val cache' = cache |> fold (insert_type typ) (sort' ~~ ths');
     val ths =
       sort |> map (fn c =>
-        Goal.finish (the (lookup_type cache' typ c) RS
-          Goal.init (Thm.cterm_of thy (Logic.mk_of_class (typ, c))))
+        Goal.finish
+          (Syntax.init_pretty_global thy)
+          (the (lookup_type cache' typ c) RS
+            Goal.init (Thm.cterm_of thy (Logic.mk_of_class (typ, c))))
         |> Thm.adjust_maxidx_thm ~1);
   in (ths, cache') end;
 
--- a/src/Pure/codegen.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/codegen.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -831,7 +831,8 @@
   end;
 
 val generate_code_i = gen_generate_code Sign.cert_term;
-val generate_code = gen_generate_code OldGoals.read_term;
+val generate_code =
+  gen_generate_code (Syntax.read_term o ProofContext.allow_dummies o ProofContext.init);
 
 
 (**** Reflection ****)
--- a/src/Pure/display.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/display.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -42,8 +42,8 @@
 
 (** options **)
 
-val goals_limit = Display_Goal.goals_limit;
-val show_consts = Display_Goal.show_consts;
+val goals_limit = Goal_Display.goals_limit;
+val show_consts = Goal_Display.show_consts;
 
 val show_hyps = ref false;      (*false: print meta-hypotheses as dots*)
 val show_tags = ref false;      (*false: suppress tags*)
@@ -87,7 +87,7 @@
       if hlen = 0 andalso status = "" then []
       else if ! show_hyps orelse show_hyps' then
         [Pretty.brk 2, Pretty.list "[" "]"
-          (map (q o Display_Goal.pretty_flexpair ctxt) tpairs @ map prt_term hyps' @
+          (map (q o Goal_Display.pretty_flexpair ctxt) tpairs @ map prt_term hyps' @
            map (Syntax.pretty_sort ctxt) xshyps @
             (if status = "" then [] else [Pretty.str status]))]
       else [Pretty.brk 2, Pretty.str ("[" ^ implode (replicate hlen ".") ^ status ^ "]")];
--- a/src/Pure/display_goal.ML	Fri Jul 24 13:56:02 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,126 +0,0 @@
-(*  Title:      Pure/display_goal.ML
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Author:     Makarius
-
-Display tactical goal state.
-*)
-
-signature DISPLAY_GOAL =
-sig
-  val goals_limit: int ref
-  val show_consts: bool ref
-  val pretty_flexpair: Proof.context -> term * term -> Pretty.T
-  val pretty_goals: Proof.context -> Markup.T -> bool * bool -> int -> thm -> Pretty.T list
-  val pretty_goals_without_context: int -> thm -> Pretty.T list
-  val print_goals_without_context: int -> thm -> unit
-end;
-
-structure Display_Goal: DISPLAY_GOAL =
-struct
-
-val goals_limit = ref 10;      (*max number of goals to print*)
-val show_consts = ref false;   (*true: show consts with types in proof state output*)
-
-fun pretty_flexpair ctxt (t, u) = Pretty.block
-  [Syntax.pretty_term ctxt t, Pretty.str " =?=", Pretty.brk 1, Syntax.pretty_term ctxt u];
-
-
-(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*)
-
-local
-
-fun ins_entry (x, y) =
-  AList.default (op =) (x, []) #>
-  AList.map_entry (op =) x (insert (op =) y);
-
-val add_consts = Term.fold_aterms
-  (fn Const (c, T) => ins_entry (T, (c, T))
-    | _ => I);
-
-val add_vars = Term.fold_aterms
-  (fn Free (x, T) => ins_entry (T, (x, ~1))
-    | Var (xi, T) => ins_entry (T, xi)
-    | _ => I);
-
-val add_varsT = Term.fold_atyps
-  (fn TFree (x, S) => ins_entry (S, (x, ~1))
-    | TVar (xi, S) => ins_entry (S, xi)
-    | _ => I);
-
-fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs;
-fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs;
-
-fun consts_of t = sort_cnsts (add_consts t []);
-fun vars_of t = sort_idxs (add_vars t []);
-fun varsT_of t = rev (sort_idxs (Term.fold_types add_varsT t []));
-
-in
-
-fun pretty_goals ctxt markup (msg, main) maxgoals state =
-  let
-    val prt_sort = Syntax.pretty_sort ctxt;
-    val prt_typ = Syntax.pretty_typ ctxt;
-    val prt_term = Syntax.pretty_term ctxt;
-
-    fun prt_atoms prt prtT (X, xs) = Pretty.block
-      [Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::",
-        Pretty.brk 1, prtT X];
-
-    fun prt_var (x, ~1) = prt_term (Syntax.free x)
-      | prt_var xi = prt_term (Syntax.var xi);
-
-    fun prt_varT (x, ~1) = prt_typ (TFree (x, []))
-      | prt_varT xi = prt_typ (TVar (xi, []));
-
-    val prt_consts = prt_atoms (prt_term o Const) prt_typ;
-    val prt_vars = prt_atoms prt_var prt_typ;
-    val prt_varsT = prt_atoms prt_varT prt_sort;
-
-
-    fun pretty_list _ _ [] = []
-      | pretty_list name prt lst = [Pretty.big_list name (map prt lst)];
-
-    fun pretty_subgoal (n, A) = Pretty.markup markup
-      [Pretty.str (" " ^ string_of_int n ^ ". "), prt_term A];
-    fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As);
-
-    val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair ctxt);
-
-    val pretty_consts = pretty_list "constants:" prt_consts o consts_of;
-    val pretty_vars = pretty_list "variables:" prt_vars o vars_of;
-    val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of;
-
-
-    val {prop, tpairs, ...} = Thm.rep_thm state;
-    val (As, B) = Logic.strip_horn prop;
-    val ngoals = length As;
-
-    fun pretty_gs (types, sorts) =
-      (if main then [prt_term B] else []) @
-       (if ngoals = 0 then [Pretty.str "No subgoals!"]
-        else if ngoals > maxgoals then
-          pretty_subgoals (Library.take (maxgoals, As)) @
-          (if msg then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")]
-           else [])
-        else pretty_subgoals As) @
-      pretty_ffpairs tpairs @
-      (if ! show_consts then pretty_consts prop else []) @
-      (if types then pretty_vars prop else []) @
-      (if sorts then pretty_varsT prop else []);
-  in
-    setmp show_no_free_types true
-      (setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types)
-        (setmp show_sorts false pretty_gs))
-   (! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts)
-  end;
-
-fun pretty_goals_without_context n th =
-  pretty_goals (Syntax.init_pretty_global (Thm.theory_of_thm th)) Markup.none (true, true) n th;
-
-val print_goals_without_context =
-  (Pretty.writeln o Pretty.chunks) oo pretty_goals_without_context;
-
-end;
-
-end;
-
--- a/src/Pure/goal.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/goal.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -18,7 +18,8 @@
   val init: cterm -> thm
   val protect: thm -> thm
   val conclude: thm -> thm
-  val finish: thm -> thm
+  val check_finished: Proof.context -> thm -> unit
+  val finish: Proof.context -> thm -> thm
   val norm_result: thm -> thm
   val future_enabled: unit -> bool
   val local_future_enabled: unit -> bool
@@ -74,12 +75,15 @@
   --- (finish)
    C
 *)
-fun finish th =
+fun check_finished ctxt th =
   (case Thm.nprems_of th of
-    0 => conclude th
+    0 => ()
   | n => raise THM ("Proof failed.\n" ^
-      Pretty.string_of (Pretty.chunks (Display_Goal.pretty_goals_without_context n th)) ^
-      ("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th]));
+      Pretty.string_of (Pretty.chunks
+        (Goal_Display.pretty_goals ctxt {total = true, main = true, maxgoals = n} th)) ^
+      "\n" ^ string_of_int n ^ " unsolved goal(s)!", 0, [th]));
+
+fun finish ctxt th = (check_finished ctxt th; conclude th);
 
 
 
@@ -99,7 +103,7 @@
 val parallel_proofs = ref 1;
 
 fun future_enabled () =
-  Future.enabled () andalso Future.is_worker () andalso ! parallel_proofs >= 1;
+  Multithreading.enabled () andalso Future.is_worker () andalso ! parallel_proofs >= 1;
 
 fun local_future_enabled () = future_enabled () andalso ! parallel_proofs >= 2;
 
@@ -145,7 +149,8 @@
 
 fun prove_internal casms cprop tac =
   (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
-    SOME th => Drule.implies_intr_list casms (finish th)
+    SOME th => Drule.implies_intr_list casms
+      (finish (Syntax.init_pretty_global (Thm.theory_of_thm th)) th)
   | NONE => error "Tactic failed.");
 
 
@@ -180,7 +185,7 @@
       (case SINGLE (tac {prems = prems, context = ctxt'}) (init stmt) of
         NONE => err "Tactic failed."
       | SOME st =>
-          let val res = finish st handle THM (msg, _, _) => err msg in
+          let val res = finish ctxt' st handle THM (msg, _, _) => err msg in
             if Unify.matches_list thy [Thm.term_of stmt] [Thm.prop_of res]
             then Thm.check_shyps sorts res
             else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/goal_display.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -0,0 +1,124 @@
+(*  Title:      Pure/goal_display.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author:     Makarius
+
+Display tactical goal state.
+*)
+
+signature GOAL_DISPLAY =
+sig
+  val goals_limit: int ref
+  val show_consts: bool ref
+  val pretty_flexpair: Proof.context -> term * term -> Pretty.T
+  val pretty_goals: Proof.context -> {total: bool, main: bool, maxgoals: int} ->
+    thm -> Pretty.T list
+  val pretty_goals_without_context: int -> thm -> Pretty.T list
+end;
+
+structure Goal_Display: GOAL_DISPLAY =
+struct
+
+val goals_limit = ref 10;      (*max number of goals to print*)
+val show_consts = ref false;   (*true: show consts with types in proof state output*)
+
+fun pretty_flexpair ctxt (t, u) = Pretty.block
+  [Syntax.pretty_term ctxt t, Pretty.str " =?=", Pretty.brk 1, Syntax.pretty_term ctxt u];
+
+
+(*print thm A1,...,An/B in "goal style" -- premises as numbered subgoals*)
+
+local
+
+fun ins_entry (x, y) =
+  AList.default (op =) (x, []) #>
+  AList.map_entry (op =) x (insert (op =) y);
+
+val add_consts = Term.fold_aterms
+  (fn Const (c, T) => ins_entry (T, (c, T))
+    | _ => I);
+
+val add_vars = Term.fold_aterms
+  (fn Free (x, T) => ins_entry (T, (x, ~1))
+    | Var (xi, T) => ins_entry (T, xi)
+    | _ => I);
+
+val add_varsT = Term.fold_atyps
+  (fn TFree (x, S) => ins_entry (S, (x, ~1))
+    | TVar (xi, S) => ins_entry (S, xi)
+    | _ => I);
+
+fun sort_idxs vs = map (apsnd (sort (prod_ord string_ord int_ord))) vs;
+fun sort_cnsts cs = map (apsnd (sort_wrt fst)) cs;
+
+fun consts_of t = sort_cnsts (add_consts t []);
+fun vars_of t = sort_idxs (add_vars t []);
+fun varsT_of t = rev (sort_idxs (Term.fold_types add_varsT t []));
+
+in
+
+fun pretty_goals ctxt {total, main, maxgoals} state =
+  let
+    val prt_sort = Syntax.pretty_sort ctxt;
+    val prt_typ = Syntax.pretty_typ ctxt;
+    val prt_term = Syntax.pretty_term ctxt;
+
+    fun prt_atoms prt prtT (X, xs) = Pretty.block
+      [Pretty.block (Pretty.commas (map prt xs)), Pretty.str " ::",
+        Pretty.brk 1, prtT X];
+
+    fun prt_var (x, ~1) = prt_term (Syntax.free x)
+      | prt_var xi = prt_term (Syntax.var xi);
+
+    fun prt_varT (x, ~1) = prt_typ (TFree (x, []))
+      | prt_varT xi = prt_typ (TVar (xi, []));
+
+    val prt_consts = prt_atoms (prt_term o Const) prt_typ;
+    val prt_vars = prt_atoms prt_var prt_typ;
+    val prt_varsT = prt_atoms prt_varT prt_sort;
+
+
+    fun pretty_list _ _ [] = []
+      | pretty_list name prt lst = [Pretty.big_list name (map prt lst)];
+
+    fun pretty_subgoal (n, A) = Pretty.markup Markup.subgoal
+      [Pretty.str (" " ^ string_of_int n ^ ". "), prt_term A];
+    fun pretty_subgoals As = map pretty_subgoal (1 upto length As ~~ As);
+
+    val pretty_ffpairs = pretty_list "flex-flex pairs:" (pretty_flexpair ctxt);
+
+    val pretty_consts = pretty_list "constants:" prt_consts o consts_of;
+    val pretty_vars = pretty_list "variables:" prt_vars o vars_of;
+    val pretty_varsT = pretty_list "type variables:" prt_varsT o varsT_of;
+
+
+    val {prop, tpairs, ...} = Thm.rep_thm state;
+    val (As, B) = Logic.strip_horn prop;
+    val ngoals = length As;
+
+    fun pretty_gs (types, sorts) =
+      (if main then [prt_term B] else []) @
+       (if ngoals = 0 then [Pretty.str "No subgoals!"]
+        else if ngoals > maxgoals then
+          pretty_subgoals (Library.take (maxgoals, As)) @
+          (if total then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")]
+           else [])
+        else pretty_subgoals As) @
+      pretty_ffpairs tpairs @
+      (if ! show_consts then pretty_consts prop else []) @
+      (if types then pretty_vars prop else []) @
+      (if sorts then pretty_varsT prop else []);
+  in
+    setmp show_no_free_types true
+      (setmp show_types (! show_types orelse ! show_sorts orelse ! show_all_types)
+        (setmp show_sorts false pretty_gs))
+   (! show_types orelse ! show_sorts orelse ! show_all_types, ! show_sorts)
+  end;
+
+fun pretty_goals_without_context n th =
+  pretty_goals (Syntax.init_pretty_global (Thm.theory_of_thm th))
+    {total = true, main = true, maxgoals = n} th;
+
+end;
+
+end;
+
--- a/src/Pure/library.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/library.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -109,7 +109,6 @@
   val split_list: ('a * 'b) list -> 'a list * 'b list
   val map_product: ('a -> 'b -> 'c) -> 'a list -> 'b list -> 'c list
   val fold_product: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c
-  val multiply: 'a list -> 'a list list -> 'a list list
   val separate: 'a -> 'a list -> 'a list
   val surround: 'a -> 'a list -> 'a list
   val replicate: int -> 'a -> 'a list
@@ -552,9 +551,6 @@
     else rep (n, [])
   end;
 
-(*multiply [a, b, c, ...] * [xs, ys, zs, ...]*)
-fun multiply [] _ = []
-  | multiply (x :: xs) yss = map (cons x) yss @ multiply xs yss;
 
 (* direct product *)
 
@@ -1138,5 +1134,5 @@
 
 end;
 
-structure BasicLibrary: BASIC_LIBRARY = Library;
-open BasicLibrary;
+structure Basic_Library: BASIC_LIBRARY = Library;
+open Basic_Library;
--- a/src/Pure/more_thm.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/more_thm.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -11,6 +11,8 @@
   include THM
   val aconvc: cterm * cterm -> bool
   val add_cterm_frees: cterm -> cterm list -> cterm list
+  val all_name: string * cterm -> cterm -> cterm
+  val all: cterm -> cterm -> cterm
   val mk_binop: cterm -> cterm -> cterm -> cterm
   val dest_binop: cterm -> cterm * cterm
   val dest_implies: cterm -> cterm * cterm
@@ -102,6 +104,14 @@
 
 (* cterm constructors and destructors *)
 
+fun all_name (x, t) A =
+  let
+    val cert = Thm.cterm_of (Thm.theory_of_cterm t);
+    val T = #T (Thm.rep_cterm t);
+  in Thm.capply (cert (Const ("all", (T --> propT) --> propT))) (Thm.cabs_name (x, t) A) end;
+
+fun all t A = all_name ("", t) A;
+
 fun mk_binop c a b = Thm.capply (Thm.capply c a) b;
 fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct);
 
--- a/src/Pure/old_goals.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/old_goals.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -8,28 +8,50 @@
 may be a stack of pending proofs.
 *)
 
-signature GOALS =
+signature OLD_GOALS =
 sig
-  val the_context: unit -> theory
+  val strip_context: term -> (string * typ) list * term list * term
+  val metahyps_thms: int -> thm -> thm list option
+  val METAHYPS: (thm list -> tactic) -> int -> tactic
+  val simple_read_term: theory -> typ -> string -> term
+  val read_term: theory -> string -> term
+  val read_prop: theory -> string -> term
+  type proof
   val premises: unit -> thm list
+  val reset_goals: unit -> unit
+  val result_error_fn: (thm -> string -> thm) ref
+  val print_sign_exn: theory -> exn -> 'a
+  val prove_goalw_cterm: thm list->cterm->(thm list->tactic list)->thm
+  val prove_goalw_cterm_nocheck: thm list->cterm->(thm list->tactic list)->thm
+  val prove_goalw: theory -> thm list -> string -> (thm list -> tactic list) -> thm
   val prove_goal: theory -> string -> (thm list -> tactic list) -> thm
-  val prove_goalw: theory -> thm list -> string -> (thm list -> tactic list) -> thm
   val topthm: unit -> thm
   val result: unit -> thm
   val uresult: unit -> thm
   val getgoal: int -> term
   val gethyps: int -> thm list
+  val print_exn: exn -> 'a
+  val filter_goal: (term*term->bool) -> thm list -> int -> thm list
   val prlev: int -> unit
   val pr: unit -> unit
   val prlim: int -> unit
+  val goalw_cterm: thm list -> cterm -> thm list
+  val goalw: theory -> thm list -> string -> thm list
   val goal: theory -> string -> thm list
-  val goalw: theory -> thm list -> string -> thm list
+  val Goalw: thm list -> string -> thm list
   val Goal: string -> thm list
-  val Goalw: thm list -> string -> thm list
+  val simple_prove_goal_cterm: cterm->(thm list->tactic list)->thm
   val by: tactic -> unit
+  val byev: tactic list -> unit
   val back: unit -> unit
   val choplev: int -> unit
+  val chop: unit -> unit
   val undo: unit -> unit
+  val save_proof: unit -> proof
+  val restore_proof: proof -> thm list
+  val push_proof: unit -> unit
+  val pop_proof: unit -> thm list
+  val rotate_proof: unit -> thm list
   val qed: string -> unit
   val qed_goal: string -> theory -> string -> (thm list -> tactic list) -> unit
   val qed_goalw: string -> theory -> thm list -> string
@@ -40,37 +62,148 @@
     -> (thm list -> tactic list) -> unit
 end;
 
-signature OLD_GOALS =
-sig
-  include GOALS
-  val simple_read_term: theory -> typ -> string -> term
-  val read_term: theory -> string -> term
-  val read_prop: theory -> string -> term
-  type proof
-  val chop: unit -> unit
-  val reset_goals: unit -> unit
-  val result_error_fn: (thm -> string -> thm) ref
-  val print_sign_exn: theory -> exn -> 'a
-  val prove_goalw_cterm: thm list->cterm->(thm list->tactic list)->thm
-  val prove_goalw_cterm_nocheck: thm list->cterm->(thm list->tactic list)->thm
-  val print_exn: exn -> 'a
-  val filter_goal: (term*term->bool) -> thm list -> int -> thm list
-  val goalw_cterm: thm list -> cterm -> thm list
-  val simple_prove_goal_cterm: cterm->(thm list->tactic list)->thm
-  val byev: tactic list -> unit
-  val save_proof: unit -> proof
-  val restore_proof: proof -> thm list
-  val push_proof: unit -> unit
-  val pop_proof: unit -> thm list
-  val rotate_proof: unit -> thm list
-end;
-
 structure OldGoals: OLD_GOALS =
 struct
 
-(* global context *)
+(**** METAHYPS -- tactical for using hypotheses as meta-level assumptions
+       METAHYPS (fn prems => tac prems) i
+
+converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new
+proof state A==>A, supplying A1,...,An as meta-level assumptions (in
+"prems").  The parameters x1,...,xm become free variables.  If the
+resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An)
+then it is lifted back into the original context, yielding k subgoals.
+
+Replaces unknowns in the context by Frees having the prefix METAHYP_
+New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm.
+DOES NOT HANDLE TYPE UNKNOWNS.
+
+
+NOTE: This version does not observe the proof context, and thus cannot
+work reliably.  See also Subgoal.SUBPROOF and Subgoal.FOCUS for
+properly localized variants of the same idea.
+****)
+
+(*Strips assumptions in goal yielding  ( [x1,...,xm], [H1,...,Hn], B )
+    H1,...,Hn are the hypotheses;  x1...xm are variants of the parameters.
+  Main difference from strip_assums concerns parameters:
+    it replaces the bound variables by free variables.  *)
+fun strip_context_aux (params, Hs, Const ("==>", _) $ H $ B) =
+      strip_context_aux (params, H :: Hs, B)
+  | strip_context_aux (params, Hs, Const ("all",_) $ Abs (a, T, t)) =
+      let val (b, u) = Syntax.variant_abs (a, T, t)
+      in strip_context_aux ((b, T) :: params, Hs, u) end
+  | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B);
+
+fun strip_context A = strip_context_aux ([], [], A);
+
+local
+
+  (*Left-to-right replacements: ctpairs = [...,(vi,ti),...].
+    Instantiates distinct free variables by terms of same type.*)
+  fun free_instantiate ctpairs =
+    forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
+
+  fun free_of s ((a, i), T) =
+    Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T)
+
+  fun mk_inst v = (Var v, free_of "METAHYP1_" v)
+in
+
+(*Common code for METAHYPS and metahyps_thms*)
+fun metahyps_split_prem prem =
+  let (*find all vars in the hyps -- should find tvars also!*)
+      val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) []
+      val insts = map mk_inst hyps_vars
+      (*replace the hyps_vars by Frees*)
+      val prem' = subst_atomic insts prem
+      val (params,hyps,concl) = strip_context prem'
+  in (insts,params,hyps,concl)  end;
 
-val the_context = ML_Context.the_global_context;
+fun metahyps_aux_tac tacf (prem,gno) state =
+  let val (insts,params,hyps,concl) = metahyps_split_prem prem
+      val maxidx = Thm.maxidx_of state
+      val cterm = Thm.cterm_of (Thm.theory_of_thm state)
+      val chyps = map cterm hyps
+      val hypths = map assume chyps
+      val subprems = map (Thm.forall_elim_vars 0) hypths
+      val fparams = map Free params
+      val cparams = map cterm fparams
+      fun swap_ctpair (t,u) = (cterm u, cterm t)
+      (*Subgoal variables: make Free; lift type over params*)
+      fun mk_subgoal_inst concl_vars (v, T) =
+          if member (op =) concl_vars (v, T)
+          then ((v, T), true, free_of "METAHYP2_" (v, T))
+          else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T))
+      (*Instantiate subgoal vars by Free applied to params*)
+      fun mk_ctpair (v, in_concl, u) =
+          if in_concl then (cterm (Var v), cterm u)
+          else (cterm (Var v), cterm (list_comb (u, fparams)))
+      (*Restore Vars with higher type and index*)
+      fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) =
+          if in_concl then (cterm u, cterm (Var ((a, i), T)))
+          else (cterm u, cterm (Var ((a, i + maxidx), U)))
+      (*Embed B in the original context of params and hyps*)
+      fun embed B = list_all_free (params, Logic.list_implies (hyps, B))
+      (*Strip the context using elimination rules*)
+      fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths
+      (*A form of lifting that discharges assumptions.*)
+      fun relift st =
+        let val prop = Thm.prop_of st
+            val subgoal_vars = (*Vars introduced in the subgoals*)
+              fold Term.add_vars (Logic.strip_imp_prems prop) []
+            and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
+            val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
+            val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
+            val emBs = map (cterm o embed) (prems_of st')
+            val Cth  = implies_elim_list st' (map (elim o assume) emBs)
+        in  (*restore the unknowns to the hypotheses*)
+            free_instantiate (map swap_ctpair insts @
+                              map mk_subgoal_swap_ctpair subgoal_insts)
+                (*discharge assumptions from state in same order*)
+                (implies_intr_list emBs
+                  (forall_intr_list cparams (implies_intr_list chyps Cth)))
+        end
+      (*function to replace the current subgoal*)
+      fun next st = Thm.bicompose false (false, relift st, nprems_of st) gno state
+  in Seq.maps next (tacf subprems (trivial (cterm concl))) end;
+
+end;
+
+(*Returns the theorem list that METAHYPS would supply to its tactic*)
+fun metahyps_thms i state =
+  let val prem = Logic.nth_prem (i, Thm.prop_of state)
+      and cterm = cterm_of (Thm.theory_of_thm state)
+      val (_,_,hyps,_) = metahyps_split_prem prem
+  in SOME (map (Thm.forall_elim_vars 0 o Thm.assume o cterm) hyps) end
+  handle TERM ("nth_prem", [A]) => NONE;
+
+local
+
+fun print_vars_terms thy (n,thm) =
+  let
+    fun typed ty = " has type: " ^ Syntax.string_of_typ_global thy ty;
+    fun find_vars thy (Const (c, ty)) =
+          if null (Term.add_tvarsT ty []) then I
+          else insert (op =) (c ^ typed ty)
+      | find_vars thy (Var (xi, ty)) = insert (op =) (Term.string_of_vname xi ^ typed ty)
+      | find_vars _ (Free _) = I
+      | find_vars _ (Bound _) = I
+      | find_vars thy (Abs (_, _, t)) = find_vars thy t
+      | find_vars thy (t1 $ t2) =
+          find_vars thy t1 #> find_vars thy t1;
+    val prem = Logic.nth_prem (n, Thm.prop_of thm)
+    val tms = find_vars thy prem []
+  in
+    (warning "Found schematic vars in assumptions:"; warning (cat_lines tms))
+  end;
+
+in
+
+fun METAHYPS tacf n thm = SUBGOAL (metahyps_aux_tac tacf) n thm
+  handle THM("assume: variables",_,_) => (print_vars_terms (theory_of_thm thm) (n,thm); Seq.empty)
+
+end;
 
 
 (* old ways of reading terms *)
@@ -136,7 +269,7 @@
   special applications.*)
 fun result_error_default state msg : thm =
   Pretty.str "Bad final proof state:" ::
-      Display_Goal.pretty_goals_without_context (!goals_limit) state @
+      Goal_Display.pretty_goals_without_context (!goals_limit) state @
     [Pretty.str msg, Pretty.str "Proof failed!"] |> Pretty.chunks |> Pretty.string_of |> error;
 
 val result_error_fn = ref result_error_default;
@@ -278,7 +411,7 @@
       (if ngoals > 0 then " (" ^ string_of_int ngoals ^ " subgoal" ^
         (if ngoals <> 1 then "s" else "") ^ ")"
       else ""))] @
-    Display_Goal.pretty_goals_without_context m th
+    Goal_Display.pretty_goals_without_context m th
   end |> Pretty.chunks |> Pretty.writeln;
 
 (*Printing can raise exceptions, so the assignment occurs last.
@@ -457,12 +590,13 @@
   in proofstack := ps;  pr();  prems end;
 
 (* rotate the stack so that the top element goes to the bottom *)
-fun rotate_proof() = let val (p,ps) = top_proof()
-                    in proofstack := ps@[save_proof()];
-                       restore_proof p;
-                       pr();
-                       !curr_prems
-                    end;
+fun rotate_proof() =
+  let val (p,ps) = top_proof()
+  in proofstack := ps@[save_proof()];
+     restore_proof p;
+     pr();
+     !curr_prems
+  end;
 
 
 (** theorem bindings **)
@@ -480,5 +614,3 @@
 
 end;
 
-structure Goals: GOALS = OldGoals;
-open Goals;
--- a/src/Pure/proofterm.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/proofterm.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -1334,8 +1334,7 @@
 fun get_name hyps prop prf =
   let val prop = Logic.list_implies (hyps, prop) in
     (case strip_combt (fst (strip_combP prf)) of
-      (PAxm (name, prop', _), _) => if prop = prop' then name else ""   (* FIXME !? *)
-    | (PThm (_, ((name, prop', _), _)), _) => if prop = prop' then name else ""
+      (PThm (_, ((name, prop', _), _)), _) => if prop = prop' then name else ""
     | _ => "")
   end;
 
--- a/src/Pure/subgoal.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/subgoal.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -1,58 +1,122 @@
 (*  Title:      Pure/subgoal.ML
     Author:     Makarius
 
-Tactical operations depending on local subgoal structure.
+Tactical operations with explicit subgoal focus, based on
+canonical proof decomposition (similar to fix/assume/show).
 *)
 
-signature BASIC_SUBGOAL =
-sig
-  val SUBPROOF:
-    ({context: Proof.context, schematics: ctyp list * cterm list,
-      params: cterm list, asms: cterm list, concl: cterm,
-      prems: thm list} -> tactic) -> Proof.context -> int -> tactic
-end
-
 signature SUBGOAL =
 sig
-  include BASIC_SUBGOAL
-  val focus: Proof.context -> int -> thm ->
-   {context: Proof.context, schematics: ctyp list * cterm list,
-    params: cterm list, asms: cterm list, concl: cterm, prems: thm list} * thm
-
+  type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list,
+    asms: cterm list, concl: cterm, schematics: ctyp list * cterm list}
+  val focus: Proof.context -> int -> thm -> focus * thm
+  val focus_params: Proof.context -> int -> thm -> focus * thm
+  val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list ->
+    int -> thm -> thm -> thm Seq.seq
+  val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic
+  val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic
+  val SUBPROOF: (focus -> tactic) -> Proof.context -> int -> tactic
 end;
 
 structure Subgoal: SUBGOAL =
 struct
 
-(* canonical proof decomposition -- similar to fix/assume/show *)
+(* focus *)
 
-fun focus ctxt i st =
+type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list,
+  asms: cterm list, concl: cterm, schematics: ctyp list * cterm list};
+
+fun gen_focus params_only ctxt i st =
   let
     val ((schematics, [st']), ctxt') =
       Variable.import true [Simplifier.norm_hhf_protect st] ctxt;
     val ((params, goal), ctxt'') = Variable.focus_subgoal i st' ctxt';
-    val asms = Drule.strip_imp_prems goal;
-    val concl = Drule.strip_imp_concl goal;
+    val (asms, concl) =
+      if params_only then ([], goal)
+      else (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal);
     val (prems, context) = Assumption.add_assumes asms ctxt'';
   in
-    ({context = context, schematics = schematics, params = params,
-      asms = asms, concl = concl, prems = prems}, Goal.init concl)
+    ({context = context, params = params, prems = prems,
+      asms = asms, concl = concl, schematics = schematics}, Goal.init concl)
   end;
 
-fun SUBPROOF tac ctxt i st =
+val focus = gen_focus false;
+val focus_params = gen_focus true;
+
+
+(* lift and retrofit *)
+
+(*
+     B [?'b, ?y]
+  ----------------
+  B ['b, y params]
+*)
+fun lift_import params th ctxt =
+  let
+    val cert = Thm.cterm_of (Thm.theory_of_thm th);
+    val ((_, [th']), ctxt') = Variable.importT [th] ctxt;
+
+    val Ts = map (#T o Thm.rep_cterm) params;
+    val ts = map Thm.term_of params;
+
+    val vars = rev (Term.add_vars (Thm.full_prop_of th') []);
+    val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt';
+    fun var_inst (v as (_, T)) y =
+      (cert (Var v), cert (list_comb (Free (y, Ts ---> T), ts)));
+    val th'' = Thm.instantiate ([], map2 var_inst vars ys) th';
+  in (th'', ctxt'') end;
+
+(*
+      [x, A x]
+          :
+      B x ==> C
+  ------------------
+  [!!x. A x ==> B x]
+         :
+         C
+*)
+fun lift_subgoals params asms th =
+  let
+    val lift = fold_rev Thm.all_name params o curry Drule.list_implies asms;
+    val unlift =
+      fold (Thm.elim_implies o Thm.assume) asms o
+      Drule.forall_elim_list (map #2 params) o Thm.assume;
+    val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th));
+    val th' = fold (Thm.elim_implies o unlift) subgoals th;
+  in (subgoals, th') end;
+
+fun retrofit ctxt1 ctxt0 params asms i st1 st0 =
+  let
+    val ps = map #2 params;
+    val (st2, ctxt2) = lift_import ps st1 ctxt1;
+    val (subgoals, st3) = lift_subgoals params asms st2;
+    val result = st3
+      |> Goal.conclude
+      |> Drule.implies_intr_list asms
+      |> Drule.forall_intr_list ps
+      |> Drule.implies_intr_list subgoals
+      |> singleton (Variable.export ctxt2 ctxt0)
+      |> Drule.zero_var_indexes
+      |> Drule.incr_indexes st0;
+  in Thm.compose_no_flatten false (result, Thm.nprems_of st1) i st0 end;
+
+
+(* tacticals *)
+
+fun GEN_FOCUS params_only tac ctxt i st =
   if Thm.nprems_of st < i then Seq.empty
   else
-    let
-      val (args as {context, params, ...}, st') = focus ctxt i st;
-      fun export_closed th =
-        let
-          val (th' :: params') = ProofContext.export context ctxt (th :: map Drule.mk_term params);
-          val vs = map (Thm.dest_arg o Drule.strip_imp_concl o Thm.cprop_of) params';
-        in Drule.forall_intr_list vs th' end;
-      fun retrofit th = Thm.compose_no_flatten false (th, 0) i;
-    in Seq.lifts retrofit (Seq.map (Goal.finish #> export_closed) (tac args st')) st end
+    let val (args as {context, params, asms, ...}, st') = gen_focus params_only ctxt i st;
+    in Seq.lifts (retrofit context ctxt params asms i) (tac args st') st end;
+
+val FOCUS = GEN_FOCUS false;
+val FOCUS_PARAMS = GEN_FOCUS true;
+
+fun SUBPROOF tac = FOCUS (FILTER Thm.no_prems o tac);
 
 end;
 
-structure BasicSubgoal: BASIC_SUBGOAL = Subgoal;
-open BasicSubgoal;
+val FOCUS = Subgoal.FOCUS;
+val FOCUS_PARAMS = Subgoal.FOCUS_PARAMS;
+val SUBPROOF = Subgoal.SUBPROOF;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/tactical.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -0,0 +1,384 @@
+(*  Title:      Pure/tactical.ML
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+
+Tacticals.
+*)
+
+infix 1 THEN THEN' THEN_ALL_NEW;
+infix 0 ORELSE APPEND INTLEAVE ORELSE' APPEND' INTLEAVE';
+infix 0 THEN_ELSE;
+
+signature TACTICAL =
+sig
+  type tactic = thm -> thm Seq.seq
+  val THEN: tactic * tactic -> tactic
+  val ORELSE: tactic * tactic -> tactic
+  val APPEND: tactic * tactic -> tactic
+  val INTLEAVE: tactic * tactic -> tactic
+  val THEN_ELSE: tactic * (tactic*tactic) -> tactic
+  val THEN': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
+  val ORELSE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
+  val APPEND': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
+  val INTLEAVE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
+  val all_tac: tactic
+  val no_tac: tactic
+  val DETERM: tactic -> tactic
+  val COND: (thm -> bool) -> tactic -> tactic -> tactic
+  val TRY: tactic -> tactic
+  val EVERY: tactic list -> tactic
+  val EVERY': ('a -> tactic) list -> 'a -> tactic
+  val EVERY1: (int -> tactic) list -> tactic
+  val FIRST: tactic list -> tactic
+  val FIRST': ('a -> tactic) list -> 'a -> tactic
+  val FIRST1: (int -> tactic) list -> tactic
+  val RANGE: (int -> tactic) list -> int -> tactic
+  val print_tac: string -> tactic
+  val pause_tac: tactic
+  val trace_REPEAT: bool ref
+  val suppress_tracing: bool ref
+  val tracify: bool ref -> tactic -> tactic
+  val traced_tac: (thm -> (thm * thm Seq.seq) option) -> tactic
+  val DETERM_UNTIL: (thm -> bool) -> tactic -> tactic
+  val REPEAT_DETERM_N: int -> tactic -> tactic
+  val REPEAT_DETERM: tactic -> tactic
+  val REPEAT: tactic -> tactic
+  val REPEAT_DETERM1: tactic -> tactic
+  val REPEAT1: tactic -> tactic
+  val FILTER: (thm -> bool) -> tactic -> tactic
+  val CHANGED: tactic -> tactic
+  val CHANGED_PROP: tactic -> tactic
+  val ALLGOALS: (int -> tactic) -> tactic
+  val SOMEGOAL: (int -> tactic) -> tactic
+  val FIRSTGOAL: (int -> tactic) -> tactic
+  val REPEAT_SOME: (int -> tactic) -> tactic
+  val REPEAT_DETERM_SOME: (int -> tactic) -> tactic
+  val REPEAT_FIRST: (int -> tactic) -> tactic
+  val REPEAT_DETERM_FIRST: (int -> tactic) -> tactic
+  val TRYALL: (int -> tactic) -> tactic
+  val CSUBGOAL: ((cterm * int) -> tactic) -> int -> tactic
+  val SUBGOAL: ((term * int) -> tactic) -> int -> tactic
+  val CHANGED_GOAL: (int -> tactic) -> int -> tactic
+  val THEN_ALL_NEW: (int -> tactic) * (int -> tactic) -> int -> tactic
+  val REPEAT_ALL_NEW: (int -> tactic) -> int -> tactic
+  val PRIMSEQ: (thm -> thm Seq.seq) -> tactic
+  val PRIMITIVE: (thm -> thm) -> tactic
+  val SINGLE: tactic -> thm -> thm option
+  val CONVERSION: conv -> int -> tactic
+end;
+
+structure Tactical : TACTICAL =
+struct
+
+(**** Tactics ****)
+
+(*A tactic maps a proof tree to a sequence of proof trees:
+    if length of sequence = 0 then the tactic does not apply;
+    if length > 1 then backtracking on the alternatives can occur.*)
+
+type tactic = thm -> thm Seq.seq;
+
+
+(*** LCF-style tacticals ***)
+
+(*the tactical THEN performs one tactic followed by another*)
+fun (tac1 THEN tac2) st = Seq.maps tac2 (tac1 st);
+
+
+(*The tactical ORELSE uses the first tactic that returns a nonempty sequence.
+  Like in LCF, ORELSE commits to either tac1 or tac2 immediately.
+  Does not backtrack to tac2 if tac1 was initially chosen. *)
+fun (tac1 ORELSE tac2) st =
+    case Seq.pull(tac1 st) of
+        NONE       => tac2 st
+      | sequencecell => Seq.make(fn()=> sequencecell);
+
+
+(*The tactical APPEND combines the results of two tactics.
+  Like ORELSE, but allows backtracking on both tac1 and tac2.
+  The tactic tac2 is not applied until needed.*)
+fun (tac1 APPEND tac2) st =
+  Seq.append (tac1 st) (Seq.make(fn()=> Seq.pull (tac2 st)));
+
+(*Like APPEND, but interleaves results of tac1 and tac2.*)
+fun (tac1 INTLEAVE tac2) st =
+    Seq.interleave(tac1 st,
+                        Seq.make(fn()=> Seq.pull (tac2 st)));
+
+(*Conditional tactic.
+        tac1 ORELSE tac2 = tac1 THEN_ELSE (all_tac, tac2)
+        tac1 THEN tac2   = tac1 THEN_ELSE (tac2, no_tac)
+*)
+fun (tac THEN_ELSE (tac1, tac2)) st =
+    case Seq.pull(tac st) of
+        NONE    => tac2 st                                   (*failed; try tactic 2*)
+      | seqcell => Seq.maps tac1 (Seq.make(fn()=> seqcell)); (*succeeded; use tactic 1*)
+
+
+(*Versions for combining tactic-valued functions, as in
+     SOMEGOAL (resolve_tac rls THEN' assume_tac) *)
+fun (tac1 THEN' tac2) x = tac1 x THEN tac2 x;
+fun (tac1 ORELSE' tac2) x = tac1 x ORELSE tac2 x;
+fun (tac1 APPEND' tac2) x = tac1 x APPEND tac2 x;
+fun (tac1 INTLEAVE' tac2) x = tac1 x INTLEAVE tac2 x;
+
+(*passes all proofs through unchanged;  identity of THEN*)
+fun all_tac st = Seq.single st;
+
+(*passes no proofs through;  identity of ORELSE and APPEND*)
+fun no_tac st  = Seq.empty;
+
+
+(*Make a tactic deterministic by chopping the tail of the proof sequence*)
+fun DETERM tac = Seq.DETERM tac;
+
+(*Conditional tactical: testfun controls which tactic to use next.
+  Beware: due to eager evaluation, both thentac and elsetac are evaluated.*)
+fun COND testfun thenf elsef = (fn prf =>
+    if testfun prf then  thenf prf   else  elsef prf);
+
+(*Do the tactic or else do nothing*)
+fun TRY tac = tac ORELSE all_tac;
+
+(*** List-oriented tactics ***)
+
+local
+  (*This version of EVERY avoids backtracking over repeated states*)
+
+  fun EVY (trail, []) st =
+        Seq.make (fn()=> SOME(st,
+                        Seq.make (fn()=> Seq.pull (evyBack trail))))
+    | EVY (trail, tac::tacs) st =
+          case Seq.pull(tac st) of
+              NONE    => evyBack trail              (*failed: backtrack*)
+            | SOME(st',q) => EVY ((st',q,tacs)::trail, tacs) st'
+  and evyBack [] = Seq.empty (*no alternatives*)
+    | evyBack ((st',q,tacs)::trail) =
+          case Seq.pull q of
+              NONE        => evyBack trail
+            | SOME(st,q') => if Thm.eq_thm (st',st)
+                             then evyBack ((st',q',tacs)::trail)
+                             else EVY ((st,q',tacs)::trail, tacs) st
+in
+
+(* EVERY [tac1,...,tacn]   equals    tac1 THEN ... THEN tacn   *)
+fun EVERY tacs = EVY ([], tacs);
+end;
+
+
+(* EVERY' [tac1,...,tacn] i  equals    tac1 i THEN ... THEN tacn i   *)
+fun EVERY' tacs i = EVERY (map (fn f => f i) tacs);
+
+(*Apply every tactic to 1*)
+fun EVERY1 tacs = EVERY' tacs 1;
+
+(* FIRST [tac1,...,tacn]   equals    tac1 ORELSE ... ORELSE tacn   *)
+fun FIRST tacs = fold_rev (curry op ORELSE) tacs no_tac;
+
+(* FIRST' [tac1,...,tacn] i  equals    tac1 i ORELSE ... ORELSE tacn i   *)
+fun FIRST' tacs = fold_rev (curry op ORELSE') tacs (K no_tac);
+
+(*Apply first tactic to 1*)
+fun FIRST1 tacs = FIRST' tacs 1;
+
+(*Apply tactics on consecutive subgoals*)
+fun RANGE [] _ = all_tac
+  | RANGE (tac :: tacs) i = RANGE tacs (i + 1) THEN tac i;
+
+
+(*** Tracing tactics ***)
+
+(*Print the current proof state and pass it on.*)
+fun print_tac msg st =
+ (tracing (msg ^ "\n" ^
+    Pretty.string_of (Pretty.chunks
+      (Goal_Display.pretty_goals_without_context (! Goal_Display.goals_limit) st)));
+  Seq.single st);
+
+(*Pause until a line is typed -- if non-empty then fail. *)
+fun pause_tac st =
+  (tracing "** Press RETURN to continue:";
+   if TextIO.inputLine TextIO.stdIn = SOME "\n" then Seq.single st
+   else (tracing "Goodbye";  Seq.empty));
+
+exception TRACE_EXIT of thm
+and TRACE_QUIT;
+
+(*Tracing flags*)
+val trace_REPEAT= ref false
+and suppress_tracing = ref false;
+
+(*Handle all tracing commands for current state and tactic *)
+fun exec_trace_command flag (tac, st) =
+   case TextIO.inputLine TextIO.stdIn of
+       SOME "\n" => tac st
+     | SOME "f\n" => Seq.empty
+     | SOME "o\n" => (flag:=false;  tac st)
+     | SOME "s\n" => (suppress_tracing:=true;  tac st)
+     | SOME "x\n" => (tracing "Exiting now";  raise (TRACE_EXIT st))
+     | SOME "quit\n" => raise TRACE_QUIT
+     | _     => (tracing
+"Type RETURN to continue or...\n\
+\     f    - to fail here\n\
+\     o    - to switch tracing off\n\
+\     s    - to suppress tracing until next entry to a tactical\n\
+\     x    - to exit at this point\n\
+\     quit - to abort this tracing run\n\
+\** Well? "     ;  exec_trace_command flag (tac, st));
+
+
+(*Extract from a tactic, a thm->thm seq function that handles tracing*)
+fun tracify flag tac st =
+  if !flag andalso not (!suppress_tracing) then
+    (tracing (Pretty.string_of (Pretty.chunks
+        (Goal_Display.pretty_goals_without_context (! Goal_Display.goals_limit) st @
+          [Pretty.str "** Press RETURN to continue:"])));
+     exec_trace_command flag (tac, st))
+  else tac st;
+
+(*Create a tactic whose outcome is given by seqf, handling TRACE_EXIT*)
+fun traced_tac seqf st =
+    (suppress_tracing := false;
+     Seq.make (fn()=> seqf st
+                         handle TRACE_EXIT st' => SOME(st', Seq.empty)));
+
+
+(*Deterministic DO..UNTIL: only retains the first outcome; tail recursive.
+  Forces repitition until predicate on state is fulfilled.*)
+fun DETERM_UNTIL p tac =
+let val tac = tracify trace_REPEAT tac
+    fun drep st = if p st then SOME (st, Seq.empty)
+                          else (case Seq.pull(tac st) of
+                                  NONE        => NONE
+                                | SOME(st',_) => drep st')
+in  traced_tac drep  end;
+
+(*Deterministic REPEAT: only retains the first outcome;
+  uses less space than REPEAT; tail recursive.
+  If non-negative, n bounds the number of repetitions.*)
+fun REPEAT_DETERM_N n tac =
+  let val tac = tracify trace_REPEAT tac
+      fun drep 0 st = SOME(st, Seq.empty)
+        | drep n st =
+           (case Seq.pull(tac st) of
+                NONE       => SOME(st, Seq.empty)
+              | SOME(st',_) => drep (n-1) st')
+  in  traced_tac (drep n)  end;
+
+(*Allows any number of repetitions*)
+val REPEAT_DETERM = REPEAT_DETERM_N ~1;
+
+(*General REPEAT: maintains a stack of alternatives; tail recursive*)
+fun REPEAT tac =
+  let val tac = tracify trace_REPEAT tac
+      fun rep qs st =
+        case Seq.pull(tac st) of
+            NONE       => SOME(st, Seq.make(fn()=> repq qs))
+          | SOME(st',q) => rep (q::qs) st'
+      and repq [] = NONE
+        | repq(q::qs) = case Seq.pull q of
+            NONE       => repq qs
+          | SOME(st,q) => rep (q::qs) st
+  in  traced_tac (rep [])  end;
+
+(*Repeat 1 or more times*)
+fun REPEAT_DETERM1 tac = DETERM tac THEN REPEAT_DETERM tac;
+fun REPEAT1 tac = tac THEN REPEAT tac;
+
+
+(** Filtering tacticals **)
+
+fun FILTER pred tac st = Seq.filter pred (tac st);
+
+(*Accept only next states that change the theorem somehow*)
+fun CHANGED tac st =
+  let fun diff st' = not (Thm.eq_thm (st, st'));
+  in Seq.filter diff (tac st) end;
+
+(*Accept only next states that change the theorem's prop field
+  (changes to signature, hyps, etc. don't count)*)
+fun CHANGED_PROP tac st =
+  let fun diff st' = not (Thm.eq_thm_prop (st, st'));
+  in Seq.filter diff (tac st) end;
+
+
+(*** Tacticals based on subgoal numbering ***)
+
+(*For n subgoals, performs tac(n) THEN ... THEN tac(1)
+  Essential to work backwards since tac(i) may add/delete subgoals at i. *)
+fun ALLGOALS tac st =
+  let fun doall 0 = all_tac
+        | doall n = tac(n) THEN doall(n-1)
+  in  doall(nprems_of st)st  end;
+
+(*For n subgoals, performs tac(n) ORELSE ... ORELSE tac(1)  *)
+fun SOMEGOAL tac st =
+  let fun find 0 = no_tac
+        | find n = tac(n) ORELSE find(n-1)
+  in  find(nprems_of st)st  end;
+
+(*For n subgoals, performs tac(1) ORELSE ... ORELSE tac(n).
+  More appropriate than SOMEGOAL in some cases.*)
+fun FIRSTGOAL tac st =
+  let fun find (i,n) = if i>n then no_tac else  tac(i) ORELSE find (i+1,n)
+  in  find(1, nprems_of st)st  end;
+
+(*Repeatedly solve some using tac. *)
+fun REPEAT_SOME tac = REPEAT1 (SOMEGOAL (REPEAT1 o tac));
+fun REPEAT_DETERM_SOME tac = REPEAT_DETERM1 (SOMEGOAL (REPEAT_DETERM1 o tac));
+
+(*Repeatedly solve the first possible subgoal using tac. *)
+fun REPEAT_FIRST tac = REPEAT1 (FIRSTGOAL (REPEAT1 o tac));
+fun REPEAT_DETERM_FIRST tac = REPEAT_DETERM1 (FIRSTGOAL (REPEAT_DETERM1 o tac));
+
+(*For n subgoals, tries to apply tac to n,...1  *)
+fun TRYALL tac = ALLGOALS (TRY o tac);
+
+
+(*Make a tactic for subgoal i, if there is one.  *)
+fun CSUBGOAL goalfun i st =
+  (case SOME (Thm.cprem_of st i) handle THM _ => NONE of
+    SOME goal => goalfun (goal, i) st
+  | NONE => Seq.empty);
+
+fun SUBGOAL goalfun =
+  CSUBGOAL (fn (goal, i) => goalfun (Thm.term_of goal, i));
+
+(*Returns all states that have changed in subgoal i, counted from the LAST
+  subgoal.  For stac, for example.*)
+fun CHANGED_GOAL tac i st =
+    let val np = Thm.nprems_of st
+        val d = np-i                 (*distance from END*)
+        val t = Thm.term_of (Thm.cprem_of st i)
+        fun diff st' =
+            Thm.nprems_of st' - d <= 0   (*the subgoal no longer exists*)
+            orelse
+             not (Pattern.aeconv (t, Thm.term_of (Thm.cprem_of st' (Thm.nprems_of st' - d))))
+    in  Seq.filter diff (tac i st)  end
+    handle Subscript => Seq.empty  (*no subgoal i*);
+
+fun (tac1 THEN_ALL_NEW tac2) i st =
+  st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st) st'));
+
+(*repeatedly dig into any emerging subgoals*)
+fun REPEAT_ALL_NEW tac =
+  tac THEN_ALL_NEW (TRY o (fn i => REPEAT_ALL_NEW tac i));
+
+(*Makes a tactic whose effect on a state is given by thmfun: thm->thm seq.*)
+fun PRIMSEQ thmfun st =  thmfun st handle THM _ => Seq.empty;
+
+(*Makes a tactic whose effect on a state is given by thmfun: thm->thm.*)
+fun PRIMITIVE thmfun = PRIMSEQ (Seq.single o thmfun);
+
+(*Inverse (more or less) of PRIMITIVE*)
+fun SINGLE tacf = Option.map fst o Seq.pull o tacf
+
+(*Conversions as tactics*)
+fun CONVERSION cv i st = Seq.single (Conv.gconv_rule cv i st)
+  handle THM _ => Seq.empty
+    | CTERM _ => Seq.empty
+    | TERM _ => Seq.empty
+    | TYPE _ => Seq.empty;
+
+end;
+
+open Tactical;
--- a/src/Pure/tctical.ML	Fri Jul 24 13:56:02 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,523 +0,0 @@
-(*  Title:      Pure/tctical.ML
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-
-Tacticals.
-*)
-
-infix 1 THEN THEN' THEN_ALL_NEW;
-infix 0 ORELSE APPEND INTLEAVE ORELSE' APPEND' INTLEAVE';
-infix 0 THEN_ELSE;
-
-signature TACTICAL =
-sig
-  type tactic = thm -> thm Seq.seq
-  val THEN: tactic * tactic -> tactic
-  val ORELSE: tactic * tactic -> tactic
-  val APPEND: tactic * tactic -> tactic
-  val INTLEAVE: tactic * tactic -> tactic
-  val THEN_ELSE: tactic * (tactic*tactic) -> tactic
-  val THEN': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
-  val ORELSE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
-  val APPEND': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
-  val INTLEAVE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic
-  val all_tac: tactic
-  val no_tac: tactic
-  val DETERM: tactic -> tactic
-  val COND: (thm -> bool) -> tactic -> tactic -> tactic
-  val TRY: tactic -> tactic
-  val EVERY: tactic list -> tactic
-  val EVERY': ('a -> tactic) list -> 'a -> tactic
-  val EVERY1: (int -> tactic) list -> tactic
-  val FIRST: tactic list -> tactic
-  val FIRST': ('a -> tactic) list -> 'a -> tactic
-  val FIRST1: (int -> tactic) list -> tactic
-  val RANGE: (int -> tactic) list -> int -> tactic
-  val print_tac: string -> tactic
-  val pause_tac: tactic
-  val trace_REPEAT: bool ref
-  val suppress_tracing: bool ref
-  val tracify: bool ref -> tactic -> tactic
-  val traced_tac: (thm -> (thm * thm Seq.seq) option) -> tactic
-  val DETERM_UNTIL: (thm -> bool) -> tactic -> tactic
-  val REPEAT_DETERM_N: int -> tactic -> tactic
-  val REPEAT_DETERM: tactic -> tactic
-  val REPEAT: tactic -> tactic
-  val REPEAT_DETERM1: tactic -> tactic
-  val REPEAT1: tactic -> tactic
-  val FILTER: (thm -> bool) -> tactic -> tactic
-  val CHANGED: tactic -> tactic
-  val CHANGED_PROP: tactic -> tactic
-  val ALLGOALS: (int -> tactic) -> tactic
-  val SOMEGOAL: (int -> tactic) -> tactic
-  val FIRSTGOAL: (int -> tactic) -> tactic
-  val REPEAT_SOME: (int -> tactic) -> tactic
-  val REPEAT_DETERM_SOME: (int -> tactic) -> tactic
-  val REPEAT_FIRST: (int -> tactic) -> tactic
-  val REPEAT_DETERM_FIRST: (int -> tactic) -> tactic
-  val TRYALL: (int -> tactic) -> tactic
-  val CSUBGOAL: ((cterm * int) -> tactic) -> int -> tactic
-  val SUBGOAL: ((term * int) -> tactic) -> int -> tactic
-  val CHANGED_GOAL: (int -> tactic) -> int -> tactic
-  val THEN_ALL_NEW: (int -> tactic) * (int -> tactic) -> int -> tactic
-  val REPEAT_ALL_NEW: (int -> tactic) -> int -> tactic
-  val strip_context: term -> (string * typ) list * term list * term
-  val metahyps_thms: int -> thm -> thm list option
-  val METAHYPS: (thm list -> tactic) -> int -> tactic
-  val PRIMSEQ: (thm -> thm Seq.seq) -> tactic
-  val PRIMITIVE: (thm -> thm) -> tactic
-  val SINGLE: tactic -> thm -> thm option
-  val CONVERSION: conv -> int -> tactic
-end;
-
-structure Tactical : TACTICAL =
-struct
-
-(**** Tactics ****)
-
-(*A tactic maps a proof tree to a sequence of proof trees:
-    if length of sequence = 0 then the tactic does not apply;
-    if length > 1 then backtracking on the alternatives can occur.*)
-
-type tactic = thm -> thm Seq.seq;
-
-
-(*** LCF-style tacticals ***)
-
-(*the tactical THEN performs one tactic followed by another*)
-fun (tac1 THEN tac2) st = Seq.maps tac2 (tac1 st);
-
-
-(*The tactical ORELSE uses the first tactic that returns a nonempty sequence.
-  Like in LCF, ORELSE commits to either tac1 or tac2 immediately.
-  Does not backtrack to tac2 if tac1 was initially chosen. *)
-fun (tac1 ORELSE tac2) st =
-    case Seq.pull(tac1 st) of
-        NONE       => tac2 st
-      | sequencecell => Seq.make(fn()=> sequencecell);
-
-
-(*The tactical APPEND combines the results of two tactics.
-  Like ORELSE, but allows backtracking on both tac1 and tac2.
-  The tactic tac2 is not applied until needed.*)
-fun (tac1 APPEND tac2) st =
-  Seq.append (tac1 st) (Seq.make(fn()=> Seq.pull (tac2 st)));
-
-(*Like APPEND, but interleaves results of tac1 and tac2.*)
-fun (tac1 INTLEAVE tac2) st =
-    Seq.interleave(tac1 st,
-                        Seq.make(fn()=> Seq.pull (tac2 st)));
-
-(*Conditional tactic.
-        tac1 ORELSE tac2 = tac1 THEN_ELSE (all_tac, tac2)
-        tac1 THEN tac2   = tac1 THEN_ELSE (tac2, no_tac)
-*)
-fun (tac THEN_ELSE (tac1, tac2)) st =
-    case Seq.pull(tac st) of
-        NONE    => tac2 st                                   (*failed; try tactic 2*)
-      | seqcell => Seq.maps tac1 (Seq.make(fn()=> seqcell)); (*succeeded; use tactic 1*)
-
-
-(*Versions for combining tactic-valued functions, as in
-     SOMEGOAL (resolve_tac rls THEN' assume_tac) *)
-fun (tac1 THEN' tac2) x = tac1 x THEN tac2 x;
-fun (tac1 ORELSE' tac2) x = tac1 x ORELSE tac2 x;
-fun (tac1 APPEND' tac2) x = tac1 x APPEND tac2 x;
-fun (tac1 INTLEAVE' tac2) x = tac1 x INTLEAVE tac2 x;
-
-(*passes all proofs through unchanged;  identity of THEN*)
-fun all_tac st = Seq.single st;
-
-(*passes no proofs through;  identity of ORELSE and APPEND*)
-fun no_tac st  = Seq.empty;
-
-
-(*Make a tactic deterministic by chopping the tail of the proof sequence*)
-fun DETERM tac = Seq.DETERM tac;
-
-(*Conditional tactical: testfun controls which tactic to use next.
-  Beware: due to eager evaluation, both thentac and elsetac are evaluated.*)
-fun COND testfun thenf elsef = (fn prf =>
-    if testfun prf then  thenf prf   else  elsef prf);
-
-(*Do the tactic or else do nothing*)
-fun TRY tac = tac ORELSE all_tac;
-
-(*** List-oriented tactics ***)
-
-local
-  (*This version of EVERY avoids backtracking over repeated states*)
-
-  fun EVY (trail, []) st =
-        Seq.make (fn()=> SOME(st,
-                        Seq.make (fn()=> Seq.pull (evyBack trail))))
-    | EVY (trail, tac::tacs) st =
-          case Seq.pull(tac st) of
-              NONE    => evyBack trail              (*failed: backtrack*)
-            | SOME(st',q) => EVY ((st',q,tacs)::trail, tacs) st'
-  and evyBack [] = Seq.empty (*no alternatives*)
-    | evyBack ((st',q,tacs)::trail) =
-          case Seq.pull q of
-              NONE        => evyBack trail
-            | SOME(st,q') => if Thm.eq_thm (st',st)
-                             then evyBack ((st',q',tacs)::trail)
-                             else EVY ((st,q',tacs)::trail, tacs) st
-in
-
-(* EVERY [tac1,...,tacn]   equals    tac1 THEN ... THEN tacn   *)
-fun EVERY tacs = EVY ([], tacs);
-end;
-
-
-(* EVERY' [tac1,...,tacn] i  equals    tac1 i THEN ... THEN tacn i   *)
-fun EVERY' tacs i = EVERY (map (fn f => f i) tacs);
-
-(*Apply every tactic to 1*)
-fun EVERY1 tacs = EVERY' tacs 1;
-
-(* FIRST [tac1,...,tacn]   equals    tac1 ORELSE ... ORELSE tacn   *)
-fun FIRST tacs = fold_rev (curry op ORELSE) tacs no_tac;
-
-(* FIRST' [tac1,...,tacn] i  equals    tac1 i ORELSE ... ORELSE tacn i   *)
-fun FIRST' tacs = fold_rev (curry op ORELSE') tacs (K no_tac);
-
-(*Apply first tactic to 1*)
-fun FIRST1 tacs = FIRST' tacs 1;
-
-(*Apply tactics on consecutive subgoals*)
-fun RANGE [] _ = all_tac
-  | RANGE (tac :: tacs) i = RANGE tacs (i + 1) THEN tac i;
-
-
-(*** Tracing tactics ***)
-
-(*Print the current proof state and pass it on.*)
-fun print_tac msg st =
- (tracing (msg ^ "\n" ^
-    Pretty.string_of (Pretty.chunks
-      (Display_Goal.pretty_goals_without_context (! Display_Goal.goals_limit) st)));
-  Seq.single st);
-
-(*Pause until a line is typed -- if non-empty then fail. *)
-fun pause_tac st =
-  (tracing "** Press RETURN to continue:";
-   if TextIO.inputLine TextIO.stdIn = SOME "\n" then Seq.single st
-   else (tracing "Goodbye";  Seq.empty));
-
-exception TRACE_EXIT of thm
-and TRACE_QUIT;
-
-(*Tracing flags*)
-val trace_REPEAT= ref false
-and suppress_tracing = ref false;
-
-(*Handle all tracing commands for current state and tactic *)
-fun exec_trace_command flag (tac, st) =
-   case TextIO.inputLine TextIO.stdIn of
-       SOME "\n" => tac st
-     | SOME "f\n" => Seq.empty
-     | SOME "o\n" => (flag:=false;  tac st)
-     | SOME "s\n" => (suppress_tracing:=true;  tac st)
-     | SOME "x\n" => (tracing "Exiting now";  raise (TRACE_EXIT st))
-     | SOME "quit\n" => raise TRACE_QUIT
-     | _     => (tracing
-"Type RETURN to continue or...\n\
-\     f    - to fail here\n\
-\     o    - to switch tracing off\n\
-\     s    - to suppress tracing until next entry to a tactical\n\
-\     x    - to exit at this point\n\
-\     quit - to abort this tracing run\n\
-\** Well? "     ;  exec_trace_command flag (tac, st));
-
-
-(*Extract from a tactic, a thm->thm seq function that handles tracing*)
-fun tracify flag tac st =
-  if !flag andalso not (!suppress_tracing) then
-   (Display_Goal.print_goals_without_context (! Display_Goal.goals_limit) st;
-    tracing "** Press RETURN to continue:";
-    exec_trace_command flag (tac, st))
-  else tac st;
-
-(*Create a tactic whose outcome is given by seqf, handling TRACE_EXIT*)
-fun traced_tac seqf st =
-    (suppress_tracing := false;
-     Seq.make (fn()=> seqf st
-                         handle TRACE_EXIT st' => SOME(st', Seq.empty)));
-
-
-(*Deterministic DO..UNTIL: only retains the first outcome; tail recursive.
-  Forces repitition until predicate on state is fulfilled.*)
-fun DETERM_UNTIL p tac =
-let val tac = tracify trace_REPEAT tac
-    fun drep st = if p st then SOME (st, Seq.empty)
-                          else (case Seq.pull(tac st) of
-                                  NONE        => NONE
-                                | SOME(st',_) => drep st')
-in  traced_tac drep  end;
-
-(*Deterministic REPEAT: only retains the first outcome;
-  uses less space than REPEAT; tail recursive.
-  If non-negative, n bounds the number of repetitions.*)
-fun REPEAT_DETERM_N n tac =
-  let val tac = tracify trace_REPEAT tac
-      fun drep 0 st = SOME(st, Seq.empty)
-        | drep n st =
-           (case Seq.pull(tac st) of
-                NONE       => SOME(st, Seq.empty)
-              | SOME(st',_) => drep (n-1) st')
-  in  traced_tac (drep n)  end;
-
-(*Allows any number of repetitions*)
-val REPEAT_DETERM = REPEAT_DETERM_N ~1;
-
-(*General REPEAT: maintains a stack of alternatives; tail recursive*)
-fun REPEAT tac =
-  let val tac = tracify trace_REPEAT tac
-      fun rep qs st =
-        case Seq.pull(tac st) of
-            NONE       => SOME(st, Seq.make(fn()=> repq qs))
-          | SOME(st',q) => rep (q::qs) st'
-      and repq [] = NONE
-        | repq(q::qs) = case Seq.pull q of
-            NONE       => repq qs
-          | SOME(st,q) => rep (q::qs) st
-  in  traced_tac (rep [])  end;
-
-(*Repeat 1 or more times*)
-fun REPEAT_DETERM1 tac = DETERM tac THEN REPEAT_DETERM tac;
-fun REPEAT1 tac = tac THEN REPEAT tac;
-
-
-(** Filtering tacticals **)
-
-fun FILTER pred tac st = Seq.filter pred (tac st);
-
-(*Accept only next states that change the theorem somehow*)
-fun CHANGED tac st =
-  let fun diff st' = not (Thm.eq_thm (st, st'));
-  in Seq.filter diff (tac st) end;
-
-(*Accept only next states that change the theorem's prop field
-  (changes to signature, hyps, etc. don't count)*)
-fun CHANGED_PROP tac st =
-  let fun diff st' = not (Thm.eq_thm_prop (st, st'));
-  in Seq.filter diff (tac st) end;
-
-
-(*** Tacticals based on subgoal numbering ***)
-
-(*For n subgoals, performs tac(n) THEN ... THEN tac(1)
-  Essential to work backwards since tac(i) may add/delete subgoals at i. *)
-fun ALLGOALS tac st =
-  let fun doall 0 = all_tac
-        | doall n = tac(n) THEN doall(n-1)
-  in  doall(nprems_of st)st  end;
-
-(*For n subgoals, performs tac(n) ORELSE ... ORELSE tac(1)  *)
-fun SOMEGOAL tac st =
-  let fun find 0 = no_tac
-        | find n = tac(n) ORELSE find(n-1)
-  in  find(nprems_of st)st  end;
-
-(*For n subgoals, performs tac(1) ORELSE ... ORELSE tac(n).
-  More appropriate than SOMEGOAL in some cases.*)
-fun FIRSTGOAL tac st =
-  let fun find (i,n) = if i>n then no_tac else  tac(i) ORELSE find (i+1,n)
-  in  find(1, nprems_of st)st  end;
-
-(*Repeatedly solve some using tac. *)
-fun REPEAT_SOME tac = REPEAT1 (SOMEGOAL (REPEAT1 o tac));
-fun REPEAT_DETERM_SOME tac = REPEAT_DETERM1 (SOMEGOAL (REPEAT_DETERM1 o tac));
-
-(*Repeatedly solve the first possible subgoal using tac. *)
-fun REPEAT_FIRST tac = REPEAT1 (FIRSTGOAL (REPEAT1 o tac));
-fun REPEAT_DETERM_FIRST tac = REPEAT_DETERM1 (FIRSTGOAL (REPEAT_DETERM1 o tac));
-
-(*For n subgoals, tries to apply tac to n,...1  *)
-fun TRYALL tac = ALLGOALS (TRY o tac);
-
-
-(*Make a tactic for subgoal i, if there is one.  *)
-fun CSUBGOAL goalfun i st =
-  (case SOME (Thm.cprem_of st i) handle THM _ => NONE of
-    SOME goal => goalfun (goal, i) st
-  | NONE => Seq.empty);
-
-fun SUBGOAL goalfun =
-  CSUBGOAL (fn (goal, i) => goalfun (Thm.term_of goal, i));
-
-(*Returns all states that have changed in subgoal i, counted from the LAST
-  subgoal.  For stac, for example.*)
-fun CHANGED_GOAL tac i st =
-    let val np = Thm.nprems_of st
-        val d = np-i                 (*distance from END*)
-        val t = Thm.term_of (Thm.cprem_of st i)
-        fun diff st' =
-            Thm.nprems_of st' - d <= 0   (*the subgoal no longer exists*)
-            orelse
-             not (Pattern.aeconv (t, Thm.term_of (Thm.cprem_of st' (Thm.nprems_of st' - d))))
-    in  Seq.filter diff (tac i st)  end
-    handle Subscript => Seq.empty  (*no subgoal i*);
-
-fun (tac1 THEN_ALL_NEW tac2) i st =
-  st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 i (i + nprems_of st' - nprems_of st) st'));
-
-(*repeatedly dig into any emerging subgoals*)
-fun REPEAT_ALL_NEW tac =
-  tac THEN_ALL_NEW (TRY o (fn i => REPEAT_ALL_NEW tac i));
-
-
-(*Strips assumptions in goal yielding  ( [x1,...,xm], [H1,...,Hn], B )
-    H1,...,Hn are the hypotheses;  x1...xm are variants of the parameters.
-  Main difference from strip_assums concerns parameters:
-    it replaces the bound variables by free variables.  *)
-fun strip_context_aux (params, Hs, Const("==>", _) $ H $ B) =
-        strip_context_aux (params, H::Hs, B)
-  | strip_context_aux (params, Hs, Const("all",_)$Abs(a,T,t)) =
-        let val (b,u) = Syntax.variant_abs(a,T,t)
-        in  strip_context_aux ((b,T)::params, Hs, u)  end
-  | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B);
-
-fun strip_context A = strip_context_aux ([],[],A);
-
-
-(**** METAHYPS -- tactical for using hypotheses as meta-level assumptions
-       METAHYPS (fn prems => tac prems) i
-
-converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new
-proof state A==>A, supplying A1,...,An as meta-level assumptions (in
-"prems").  The parameters x1,...,xm become free variables.  If the
-resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An)
-then it is lifted back into the original context, yielding k subgoals.
-
-Replaces unknowns in the context by Frees having the prefix METAHYP_
-New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm.
-DOES NOT HANDLE TYPE UNKNOWNS.
-****)
-
-local
-
-  (*Left-to-right replacements: ctpairs = [...,(vi,ti),...].
-    Instantiates distinct free variables by terms of same type.*)
-  fun free_instantiate ctpairs =
-    forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
-
-  fun free_of s ((a, i), T) =
-    Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T)
-
-  fun mk_inst v = (Var v, free_of "METAHYP1_" v)
-in
-
-(*Common code for METAHYPS and metahyps_thms*)
-fun metahyps_split_prem prem =
-  let (*find all vars in the hyps -- should find tvars also!*)
-      val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) []
-      val insts = map mk_inst hyps_vars
-      (*replace the hyps_vars by Frees*)
-      val prem' = subst_atomic insts prem
-      val (params,hyps,concl) = strip_context prem'
-  in (insts,params,hyps,concl)  end;
-
-fun metahyps_aux_tac tacf (prem,gno) state =
-  let val (insts,params,hyps,concl) = metahyps_split_prem prem
-      val maxidx = Thm.maxidx_of state
-      val cterm = Thm.cterm_of (Thm.theory_of_thm state)
-      val chyps = map cterm hyps
-      val hypths = map assume chyps
-      val subprems = map (Thm.forall_elim_vars 0) hypths
-      val fparams = map Free params
-      val cparams = map cterm fparams
-      fun swap_ctpair (t,u) = (cterm u, cterm t)
-      (*Subgoal variables: make Free; lift type over params*)
-      fun mk_subgoal_inst concl_vars (v, T) =
-          if member (op =) concl_vars (v, T)
-          then ((v, T), true, free_of "METAHYP2_" (v, T))
-          else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T))
-      (*Instantiate subgoal vars by Free applied to params*)
-      fun mk_ctpair (v, in_concl, u) =
-          if in_concl then (cterm (Var v), cterm u)
-          else (cterm (Var v), cterm (list_comb (u, fparams)))
-      (*Restore Vars with higher type and index*)
-      fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) =
-          if in_concl then (cterm u, cterm (Var ((a, i), T)))
-          else (cterm u, cterm (Var ((a, i + maxidx), U)))
-      (*Embed B in the original context of params and hyps*)
-      fun embed B = list_all_free (params, Logic.list_implies (hyps, B))
-      (*Strip the context using elimination rules*)
-      fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths
-      (*A form of lifting that discharges assumptions.*)
-      fun relift st =
-        let val prop = Thm.prop_of st
-            val subgoal_vars = (*Vars introduced in the subgoals*)
-              fold Term.add_vars (Logic.strip_imp_prems prop) []
-            and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
-            val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
-            val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
-            val emBs = map (cterm o embed) (prems_of st')
-            val Cth  = implies_elim_list st' (map (elim o assume) emBs)
-        in  (*restore the unknowns to the hypotheses*)
-            free_instantiate (map swap_ctpair insts @
-                              map mk_subgoal_swap_ctpair subgoal_insts)
-                (*discharge assumptions from state in same order*)
-                (implies_intr_list emBs
-                  (forall_intr_list cparams (implies_intr_list chyps Cth)))
-        end
-      (*function to replace the current subgoal*)
-      fun next st = Thm.bicompose false (false, relift st, nprems_of st) gno state
-  in Seq.maps next (tacf subprems (trivial (cterm concl))) end;
-
-end;
-
-(*Returns the theorem list that METAHYPS would supply to its tactic*)
-fun metahyps_thms i state =
-  let val prem = Logic.nth_prem (i, Thm.prop_of state)
-      and cterm = cterm_of (Thm.theory_of_thm state)
-      val (_,_,hyps,_) = metahyps_split_prem prem
-  in SOME (map (Thm.forall_elim_vars 0 o Thm.assume o cterm) hyps) end
-  handle TERM ("nth_prem", [A]) => NONE;
-
-local
-
-fun print_vars_terms thy (n,thm) =
-  let
-    fun typed ty = " has type: " ^ Syntax.string_of_typ_global thy ty;
-    fun find_vars thy (Const (c, ty)) =
-          if null (Term.add_tvarsT ty []) then I
-          else insert (op =) (c ^ typed ty)
-      | find_vars thy (Var (xi, ty)) = insert (op =) (Term.string_of_vname xi ^ typed ty)
-      | find_vars _ (Free _) = I
-      | find_vars _ (Bound _) = I
-      | find_vars thy (Abs (_, _, t)) = find_vars thy t
-      | find_vars thy (t1 $ t2) =
-          find_vars thy t1 #> find_vars thy t1;
-    val prem = Logic.nth_prem (n, Thm.prop_of thm)
-    val tms = find_vars thy prem []
-  in
-    (warning "Found schematic vars in assumptions:"; warning (cat_lines tms))
-  end;
-
-in
-
-fun METAHYPS tacf n thm = SUBGOAL (metahyps_aux_tac tacf) n thm
-  handle THM("assume: variables",_,_) => (print_vars_terms (theory_of_thm thm) (n,thm); Seq.empty)
-
-end;
-
-(*Makes a tactic whose effect on a state is given by thmfun: thm->thm seq.*)
-fun PRIMSEQ thmfun st =  thmfun st handle THM _ => Seq.empty;
-
-(*Makes a tactic whose effect on a state is given by thmfun: thm->thm.*)
-fun PRIMITIVE thmfun = PRIMSEQ (Seq.single o thmfun);
-
-(*Inverse (more or less) of PRIMITIVE*)
-fun SINGLE tacf = Option.map fst o Seq.pull o tacf
-
-(*Conversions as tactics*)
-fun CONVERSION cv i st = Seq.single (Conv.gconv_rule cv i st)
-  handle THM _ => Seq.empty
-    | CTERM _ => Seq.empty
-    | TERM _ => Seq.empty
-    | TYPE _ => Seq.empty;
-
-end;
-
-open Tactical;
--- a/src/Pure/term.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/term.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -151,6 +151,7 @@
   val match_bvars: (term * term) * (string * string) list -> (string * string) list
   val map_abs_vars: (string -> string) -> term -> term
   val rename_abs: term -> term -> term -> term option
+  val lambda_name: string * term -> term -> term
   val close_schematic_term: term -> term
   val maxidx_typ: typ -> int -> int
   val maxidx_typs: typ list -> int -> int
@@ -719,14 +720,15 @@
         | _ => raise Same.SAME);
   in abs 0 body handle Same.SAME => body end;
 
-fun lambda v t =
-  let val x =
-    (case v of
-      Const (x, _) => Long_Name.base_name x
-    | Free (x, _) => x
-    | Var ((x, _), _) => x
-    | _ => Name.uu)
-  in Abs (x, fastype_of v, abstract_over (v, t)) end;
+fun term_name (Const (x, _)) = Long_Name.base_name x
+  | term_name (Free (x, _)) = x
+  | term_name (Var ((x, _), _)) = x
+  | term_name _ = Name.uu;
+
+fun lambda_name (x, v) t =
+  Abs (if x = "" then term_name v else x, fastype_of v, abstract_over (v, t));
+
+fun lambda v t = lambda_name ("", v) t;
 
 (*Form an abstraction over a free variable.*)
 fun absfree (a,T,body) = Abs (a, T, abstract_over (Free (a, T), body));
--- a/src/Pure/thm.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/thm.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -110,6 +110,7 @@
   val dest_arg1: cterm -> cterm
   val dest_abs: string option -> cterm -> cterm * cterm
   val capply: cterm -> cterm -> cterm
+  val cabs_name: string * cterm -> cterm -> cterm
   val cabs: cterm -> cterm -> cterm
   val adjust_maxidx_cterm: int -> cterm -> cterm
   val incr_indexes_cterm: int -> cterm -> cterm
@@ -274,16 +275,18 @@
       else raise CTERM ("capply: types don't agree", [cf, cx])
   | capply cf cx = raise CTERM ("capply: first arg is not a function", [cf, cx]);
 
-fun cabs
-  (ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...})
+fun cabs_name
+  (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...})
   (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) =
-    let val t = Term.lambda t1 t2 in
+    let val t = Term.lambda_name (x, t1) t2 in
       Cterm {thy_ref = merge_thys0 ct1 ct2,
         t = t, T = T1 --> T2,
         maxidx = Int.max (maxidx1, maxidx2),
         sorts = Sorts.union sorts1 sorts2}
     end;
 
+fun cabs t u = cabs_name ("", t) u;
+
 
 (* indexes *)
 
--- a/src/Pure/variable.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Pure/variable.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -57,8 +57,8 @@
     ((ctyp list * cterm list) * thm list) * Proof.context
   val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
   val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list
-  val focus: cterm -> Proof.context -> (cterm list * cterm) * Proof.context
-  val focus_subgoal: int -> thm -> Proof.context -> (cterm list * cterm) * Proof.context
+  val focus: cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context
+  val focus_subgoal: int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context
   val warn_extra_tfrees: Proof.context -> Proof.context -> unit
   val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list
   val polymorphic: Proof.context -> term list -> term list
@@ -477,7 +477,7 @@
     val ps' = ListPair.map (cert o Free) (xs', Ts);
     val goal' = fold forall_elim_prop ps' goal;
     val ctxt'' = ctxt' |> fold (declare_constraints o Thm.term_of) ps';
-  in ((ps', goal'), ctxt'') end;
+  in ((xs ~~ ps', goal'), ctxt'') end;
 
 fun focus_subgoal i st =
   let
--- a/src/Tools/coherent.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Tools/coherent.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -215,7 +215,7 @@
 fun coherent_tac ctxt rules = SUBPROOF (fn {prems, concl, params, context, ...} =>
   rtac (rulify_elim_conv concl RS equal_elim_rule2) 1 THEN
   SUBPROOF (fn {prems = prems', concl, context, ...} =>
-    let val xs = map term_of params @
+    let val xs = map (term_of o #2) params @
       map (fn (_, s) => Free (s, the (Variable.default_type context s)))
         (Variable.fixes_of context)
     in
--- a/src/Tools/induct.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Tools/induct.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -70,7 +70,7 @@
   val setup: theory -> theory
 end;
 
-functor InductFun(Data: INDUCT_DATA): INDUCT =
+functor Induct(Data: INDUCT_DATA): INDUCT =
 struct
 
 
@@ -568,7 +568,7 @@
 *)
 
 fun get_inductT ctxt insts =
-  fold_rev multiply (insts |> map
+  fold_rev (map_product cons) (insts |> map
       ((fn [] => NONE | ts => List.last ts) #>
         (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #>
         find_inductT ctxt)) [[]]
--- a/src/Tools/project_rule.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/Tools/project_rule.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -24,7 +24,7 @@
   val projections: Proof.context -> thm -> thm list
 end;
 
-functor ProjectRuleFun(Data: PROJECT_RULE_DATA): PROJECT_RULE =
+functor Project_Rule(Data: PROJECT_RULE_DATA): PROJECT_RULE =
 struct
 
 fun conj1 th = th RS Data.conjunct1;
--- a/src/ZF/Tools/induct_tacs.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/ZF/Tools/induct_tacs.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -75,7 +75,7 @@
              (v, #1 (dest_Const (head_of A)))
         | mk_pair _ = raise Match
       val pairs = List.mapPartial (try (mk_pair o FOLogic.dest_Trueprop))
-          (#2 (strip_context Bi))
+          (#2 (OldGoals.strip_context Bi))
   in case AList.lookup (op =) pairs var of
        NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var)
      | SOME t => t
--- a/src/ZF/Tools/typechk.ML	Fri Jul 24 13:56:02 2009 +0200
+++ b/src/ZF/Tools/typechk.ML	Wed Jul 29 09:06:49 2009 +0200
@@ -116,7 +116,7 @@
     (Method.sections
       [Args.add -- Args.colon >> K (I, TC_add),
        Args.del -- Args.colon >> K (I, TC_del)]
-      >> (K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt)))))
+      >> K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt))))
     "ZF type-checking";
 
 val _ =