Automated merge with ssh://macbroy25.informatik.tu-muenchen.de//home/isabelle-repository/repos/isabelle
authorAndreas Lochbihler
Fri, 13 Apr 2012 13:30:27 +0200
changeset 47453 598604c91036
parent 47449 5e1482296b12 (diff)
parent 47452 60da1ee5363f (current diff)
child 47454 479b4d6b9562
Automated merge with ssh://macbroy25.informatik.tu-muenchen.de//home/isabelle-repository/repos/isabelle
NEWS
--- a/Admin/java/README	Fri Apr 13 13:29:55 2012 +0200
+++ b/Admin/java/README	Fri Apr 13 13:30:27 2012 +0200
@@ -1,3 +1,3 @@
-This is JDK 1.7.0_03 for Linux and Linux x86 from
+This is JDK 1.7.0_03 for Linux from
 http://www.oracle.com/technetwork/java/javase/downloads/index.html
 
--- a/Admin/java/etc/settings	Fri Apr 13 13:29:55 2012 +0200
+++ b/Admin/java/etc/settings	Fri Apr 13 13:30:27 2012 +0200
@@ -1,4 +1,4 @@
 # -*- shell-script -*- :mode=shellscript:
 
-ISABELLE_JDK_HOME="$COMPONENT/${ISABELLE_PLATFORM64:-$ISABELLE_PLATFORM}/jdk1.7.0_03"
+ISABELLE_JDK_HOME="$COMPONENT/jdk1.7.0_03"
 
--- a/CONTRIBUTORS	Fri Apr 13 13:29:55 2012 +0200
+++ b/CONTRIBUTORS	Fri Apr 13 13:30:27 2012 +0200
@@ -6,6 +6,9 @@
 Contributions to this Isabelle version
 --------------------------------------
 
+* March 2012: Christian Sternagel, Japan Advanced Institute of Science and Technology
+  Consolidated theory of relation composition.
+
 * March 2012: Nik Sultana, University of Cambridge
   HOL/TPTP parser and import facilities.
 
--- a/NEWS	Fri Apr 13 13:29:55 2012 +0200
+++ b/NEWS	Fri Apr 13 13:30:27 2012 +0200
@@ -294,6 +294,42 @@
  
 INCOMPATIBILITY.
 
+* Theory Relation: Consolidated constant name for relation composition
+  and corresponding theorem names:
+  - Renamed constant rel_comp to relcomp
+  - Dropped abbreviation pred_comp. Use relcompp instead.
+  - Renamed theorems:
+  Relation:
+    rel_compI ~> relcompI
+    rel_compEpair ~> relcompEpair
+    rel_compE ~> relcompE
+    pred_comp_rel_comp_eq ~> relcompp_relcomp_eq
+    rel_comp_empty1 ~> relcomp_empty1
+    rel_comp_mono ~> relcomp_mono
+    rel_comp_subset_Sigma ~> relcomp_subset_Sigma
+    rel_comp_distrib ~> relcomp_distrib
+    rel_comp_distrib2 ~> relcomp_distrib2
+    rel_comp_UNION_distrib ~> relcomp_UNION_distrib
+    rel_comp_UNION_distrib2 ~> relcomp_UNION_distrib2
+    single_valued_rel_comp ~> single_valued_relcomp
+    rel_comp_unfold ~> relcomp_unfold
+    converse_rel_comp ~> converse_relcomp
+    pred_compI ~> relcomppI
+    pred_compE ~> relcomppE
+    pred_comp_bot1 ~> relcompp_bot1
+    pred_comp_bot2 ~> relcompp_bot2
+    transp_pred_comp_less_eq ~> transp_relcompp_less_eq
+    pred_comp_mono ~> relcompp_mono
+    pred_comp_distrib ~> relcompp_distrib
+    pred_comp_distrib2 ~> relcompp_distrib2
+    converse_pred_comp ~> converse_relcompp
+  Transitive_Closure:
+    finite_rel_comp ~> finite_relcomp
+  List:
+    set_rel_comp ~> set_relcomp
+
+INCOMPATIBILITY.
+
 * New theory HOL/Library/DAList provides an abstract type for association
   lists with distinct keys.
 
--- a/README_REPOSITORY	Fri Apr 13 13:29:55 2012 +0200
+++ b/README_REPOSITORY	Fri Apr 13 13:30:27 2012 +0200
@@ -32,10 +32,9 @@
 
 Mercurial provides nice web presentation of incoming changes with a
 digest of log entries; this also includes RSS/Atom news feeds.  There
-are add-on browsers, notably hgtk that is part of the TortoiseHg
-distribution and works for generic Python/GTk platforms.  The
-alternative "view" utility helps to inspect the semantic content of
-merge nodes.
+are add-on history browsers such as "hg view" and TortoiseHg.  Unlike
+the default web view, some of these tools help to inspect the semantic
+content of non-trivial merge nodes.
 
 
 Initial configuration
@@ -107,27 +106,27 @@
 time of writing and many years later.
 
 Push access to the Isabelle repository requires an account at TUM,
-with properly configured ssh to the local machines (e.g. macbroy20,
-atbroy100).  You also need to be a member of the "isabelle" Unix
+with properly configured ssh to the local machines (e.g. macbroy20 ..
+macbroy29).  You also need to be a member of the "isabelle" Unix
 group.
 
 Sharing a locally modified clone then works as follows, using your
 user name instead of "wenzelm":
 
-  hg out ssh://wenzelm@atbroy100//home/isabelle-repository/repos/isabelle
+  hg out ssh://wenzelm@macbroy20//home/isabelle-repository/repos/isabelle
 
 In fact, the "out" or "outgoing" command performs only a dry run: it
 displays the changesets that would get published.  An actual "push",
 with a lasting effect on the Isabelle repository, works like this:
 
-  hg push ssh://wenzelm@atbroy100//home/isabelle-repository/repos/isabelle
+  hg push ssh://wenzelm@macbroy20//home/isabelle-repository/repos/isabelle
 
 
 Default paths for push and pull can be configured in
 isabelle/.hg/hgrc, for example:
 
   [paths]
-  default = ssh://wenzelm@atbroy100//home/isabelle-repository/repos/isabelle
+  default = ssh://wenzelm@macbroy20//home/isabelle-repository/repos/isabelle
 
 Now "hg pull" or "hg push" will use that shared file space, unless a
 different URL is specified explicitly.
@@ -136,7 +135,7 @@
 source URL.  So we could have cloned via that ssh URL in the first
 place, to get exactly to the same point:
 
-  hg clone ssh://wenzelm@atbroy100//home/isabelle-repository/repos/isabelle
+  hg clone ssh://wenzelm@macbroy20//home/isabelle-repository/repos/isabelle
 
 
 Simple merges
--- a/lib/Tools/scala	Fri Apr 13 13:29:55 2012 +0200
+++ b/lib/Tools/scala	Fri Apr 13 13:30:27 2012 +0200
@@ -6,6 +6,8 @@
 
 [ -e "$ISABELLE_HOME/Admin/build" ] && { "$ISABELLE_HOME/Admin/build" jars || exit $?; }
 
+export JAVA_HOME="$ISABELLE_JDK_HOME"
+
 CLASSPATH="$(jvmpath "$CLASSPATH")"
 isabelle_scala scala -Dfile.encoding=UTF-8 \
   "-Djava.ext.dirs=$("$ISABELLE_HOME/src/Tools/JVM/java_ext_dirs")" "$@"
--- a/lib/Tools/scalac	Fri Apr 13 13:29:55 2012 +0200
+++ b/lib/Tools/scalac	Fri Apr 13 13:30:27 2012 +0200
@@ -6,6 +6,8 @@
 
 [ -e "$ISABELLE_HOME/Admin/build" ] && { "$ISABELLE_HOME/Admin/build" jars || exit $?; }
 
+export JAVA_HOME="$ISABELLE_JDK_HOME"
+
 CLASSPATH="$(jvmpath "$CLASSPATH")"
 isabelle_scala scalac -Dfile.encoding=UTF-8 \
   "-Djava.ext.dirs=$("$ISABELLE_HOME/src/Tools/JVM/java_ext_dirs")" "$@"
--- a/src/CCL/Wfd.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/CCL/Wfd.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -493,26 +493,19 @@
 subsection {* Evaluation *}
 
 ML {*
-
-local
-  structure Data = Named_Thms(val name = @{binding eval} val description = "evaluation rules");
-in
+structure Eval_Rules =
+  Named_Thms(val name = @{binding eval} val description = "evaluation rules");
 
 fun eval_tac ths =
   Subgoal.FOCUS_PREMS (fn {context, prems, ...} =>
-    DEPTH_SOLVE_1 (resolve_tac (ths @ prems @ Data.get context) 1));
+    DEPTH_SOLVE_1 (resolve_tac (ths @ prems @ Eval_Rules.get context) 1));
+*}
+setup Eval_Rules.setup
 
-val eval_setup =
-  Data.setup #>
-  Method.setup @{binding eval}
-    (Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ths ctxt)))
-    "evaluation";
-
-end;
-
+method_setup eval = {*
+  Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ths ctxt))
 *}
 
-setup eval_setup
 
 lemmas eval_rls [eval] = trueV falseV pairV lamV caseVtrue caseVfalse caseVpair caseVlam
 
--- a/src/HOL/Auth/Smartcard/ShoupRubin.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -762,7 +762,7 @@
 
 method_setup prepare = {*
     Scan.succeed (K (SIMPLE_METHOD ShoupRubin.prepare_tac)) *}
-  "to launch a few simple facts that'll help the simplifier"
+  "to launch a few simple facts that will help the simplifier"
 
 method_setup parts_prepare = {*
     Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubin.parts_prepare_tac ctxt)) *}
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -771,7 +771,7 @@
 
 method_setup prepare = {*
     Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.prepare_tac ctxt)) *}
-  "to launch a few simple facts that'll help the simplifier"
+  "to launch a few simple facts that will help the simplifier"
 
 method_setup parts_prepare = {*
     Scan.succeed (fn ctxt => SIMPLE_METHOD (ShoupRubinBella.parts_prepare_tac ctxt)) *}
--- a/src/HOL/Boogie/Tools/boogie_tactics.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -45,12 +45,12 @@
     in tac ctxt (thms @ facts) end))
 
 val setup_boogie = Method.setup @{binding boogie} (boogie_method false)
-  ("Applies an SMT solver to the current goal " ^
-   "using the current set of Boogie background axioms")
+  "apply an SMT solver to the current goal \
+  \using the current set of Boogie background axioms"
 
 val setup_boogie_all = Method.setup @{binding boogie_all} (boogie_method true)
-  ("Applies an SMT solver to all goals " ^
-   "using the current set of Boogie background axioms")
+  "apply an SMT solver to all goals \
+  \using the current set of Boogie background axioms"
 
 
 local
@@ -96,7 +96,7 @@
 in
 val setup_boogie_cases = Method.setup @{binding boogie_cases}
   (Scan.succeed boogie_cases)
-  "Prepares a set of Boogie assertions for case-based proofs"
+  "prepare a set of Boogie assertions for case-based proofs"
 end
 
 
--- a/src/HOL/Decision_Procs/Commutative_Ring.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/Commutative_Ring.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -314,6 +314,9 @@
 
 
 use "commutative_ring_tac.ML"
-setup Commutative_Ring_Tac.setup
+
+method_setup comm_ring = {*
+  Scan.succeed (SIMPLE_METHOD' o Commutative_Ring_Tac.tac)
+*} "reflective decision procedure for equalities over commutative rings"
 
 end
--- a/src/HOL/Decision_Procs/Cooper.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -2005,7 +2005,12 @@
 *}
 
 use "cooper_tac.ML"
-setup "Cooper_Tac.setup"
+
+method_setup cooper = {*
+  Args.mode "no_quantify" >>
+    (fn q => fn ctxt => SIMPLE_METHOD' (Cooper_Tac.linz_tac ctxt (not q)))
+*} "decision procedure for linear integer arithmetic"
+
 
 text {* Tests *}
 
--- a/src/HOL/Decision_Procs/Ferrack.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -2004,7 +2004,12 @@
 *}
 
 use "ferrack_tac.ML"
-setup Ferrack_Tac.setup
+
+method_setup rferrack = {*
+  Args.mode "no_quantify" >>
+    (fn q => fn ctxt => SIMPLE_METHOD' (Ferrack_Tac.linr_tac ctxt (not q)))
+*} "decision procedure for linear real arithmetic"
+
 
 lemma
   fixes x :: real
--- a/src/HOL/Decision_Procs/MIR.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -5635,7 +5635,12 @@
 *}
 
 use "mir_tac.ML"
-setup "Mir_Tac.setup"
+
+method_setup mir = {*
+  Args.mode "no_quantify" >>
+    (fn q => fn ctxt => SIMPLE_METHOD' (Mir_Tac.mir_tac ctxt (not q)))
+*} "decision procedure for MIR arithmetic"
+
 
 lemma "ALL (x::real). (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> = (x = real \<lfloor>x\<rfloor>))"
   by mir
--- a/src/HOL/Decision_Procs/commutative_ring_tac.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/commutative_ring_tac.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -7,7 +7,6 @@
 signature COMMUTATIVE_RING_TAC =
 sig
   val tac: Proof.context -> int -> tactic
-  val setup: theory -> theory
 end
 
 structure Commutative_Ring_Tac: COMMUTATIVE_RING_TAC =
@@ -98,8 +97,4 @@
     THEN (simp_tac cring_ss i)
   end);
 
-val setup =
-  Method.setup @{binding comm_ring} (Scan.succeed (SIMPLE_METHOD' o tac))
-    "reflective decision procedure for equalities over commutative rings";
-
 end;
--- a/src/HOL/Decision_Procs/cooper_tac.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/cooper_tac.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -6,7 +6,6 @@
 sig
   val trace: bool Unsynchronized.ref
   val linz_tac: Proof.context -> bool -> int -> tactic
-  val setup: theory -> theory
 end
 
 structure Cooper_Tac: COOPER_TAC =
@@ -115,15 +114,4 @@
       | _ => (pre_thm, assm_tac i)
   in rtac (((mp_step nh) o (spec_step np)) th) i THEN tac end);
 
-val setup =
-  Method.setup @{binding cooper}
-    let
-      val parse_flag = Args.$$$ "no_quantify" >> K (K false)
-    in
-      Scan.lift (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
-        curry (Library.foldl op |>) true) >>
-      (fn q => fn ctxt => SIMPLE_METHOD' (linz_tac ctxt q))
-    end
-    "decision procedure for linear integer arithmetic";
-
 end
--- a/src/HOL/Decision_Procs/ferrack_tac.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/ferrack_tac.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -6,7 +6,6 @@
 sig
   val trace: bool Unsynchronized.ref
   val linr_tac: Proof.context -> bool -> int -> tactic
-  val setup: theory -> theory
 end
 
 structure Ferrack_Tac =
@@ -96,10 +95,4 @@
       | _ => (pre_thm, assm_tac i)
   in rtac ((mp_step nh o spec_step np) th) i THEN tac end);
 
-val setup =
-  Method.setup @{binding rferrack}
-    (Args.mode "no_quantify" >> (fn q => fn ctxt =>
-      SIMPLE_METHOD' (linr_tac ctxt (not q))))
-    "decision procedure for linear real arithmetic";
-
 end
--- a/src/HOL/Decision_Procs/mir_tac.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Decision_Procs/mir_tac.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -6,7 +6,6 @@
 sig
   val trace: bool Unsynchronized.ref
   val mir_tac: Proof.context -> bool -> int -> tactic
-  val setup: theory -> theory
 end
 
 structure Mir_Tac =
@@ -143,15 +142,4 @@
       | _ => (pre_thm, assm_tac i)
   in rtac (((mp_step nh) o (spec_step np)) th) i THEN tac end);
 
-val setup =
-  Method.setup @{binding mir}
-    let
-      val parse_flag = Args.$$$ "no_quantify" >> K (K false)
-    in
-      Scan.lift (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
-        curry (Library.foldl op |>) true) >>
-      (fn q => fn ctxt => SIMPLE_METHOD' (mir_tac ctxt q))
-    end
-    "decision procedure for MIR arithmetic";
-
 end
--- a/src/HOL/FunDef.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/FunDef.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -110,14 +110,21 @@
 use "Tools/Function/relation.ML"
 use "Tools/Function/function.ML"
 use "Tools/Function/pat_completeness.ML"
+
+method_setup pat_completeness = {*
+  Scan.succeed (SIMPLE_METHOD' o Pat_Completeness.pat_completeness_tac)
+*} "prove completeness of datatype patterns"
+
 use "Tools/Function/fun.ML"
 use "Tools/Function/induction_schema.ML"
 
+method_setup induction_schema = {*
+  Scan.succeed (RAW_METHOD o Induction_Schema.induction_schema_tac)
+*} "prove an induction principle"
+
 setup {* 
   Function.setup
-  #> Pat_Completeness.setup
   #> Function_Fun.setup
-  #> Induction_Schema.setup
 *}
 
 subsection {* Measure Functions *}
@@ -137,6 +144,12 @@
 by (rule is_measure_trivial)
 
 use "Tools/Function/lexicographic_order.ML"
+
+method_setup lexicographic_order = {*
+  Method.sections clasimp_modifiers >>
+  (K (SIMPLE_METHOD o Lexicographic_Order.lexicographic_order_tac false))
+*} "termination prover for lexicographic orderings"
+
 setup Lexicographic_Order.setup 
 
 
--- a/src/HOL/Groebner_Basis.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Groebner_Basis.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -43,8 +43,20 @@
 
 use "Tools/groebner.ML"
 
-method_setup algebra = Groebner.algebra_method
-  "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"
+method_setup algebra = {*
+  let
+    fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+    val addN = "add"
+    val delN = "del"
+    val any_keyword = keyword addN || keyword delN
+    val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
+  in
+    Scan.optional (keyword addN |-- thms) [] --
+     Scan.optional (keyword delN |-- thms) [] >>
+    (fn (add_ths, del_ths) => fn ctxt =>
+      SIMPLE_METHOD' (Groebner.algebra_tac add_ths del_ths ctxt))
+  end
+*} "solve polynomial equations over (semi)rings and ideal membership problems using Groebner bases"
 
 declare dvd_def[algebra]
 declare dvd_eq_mod_eq_0[symmetric, algebra]
--- a/src/HOL/HOLCF/Fixrec.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/HOLCF/Fixrec.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -230,7 +230,9 @@
 use "Tools/holcf_library.ML"
 use "Tools/fixrec.ML"
 
-setup {* Fixrec.setup *}
+method_setup fixrec_simp = {*
+  Scan.succeed (SIMPLE_METHOD' o Fixrec.fixrec_simp_tac)
+*} "pattern prover for fixrec constants"
 
 setup {*
   Fixrec.add_matchers
--- a/src/HOL/HOLCF/Tools/fixrec.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/HOLCF/Tools/fixrec.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -12,7 +12,6 @@
     -> (bool * (Attrib.binding * string)) list -> local_theory -> local_theory
   val add_matchers: (string * string) list -> theory -> theory
   val fixrec_simp_tac: Proof.context -> int -> tactic
-  val setup: theory -> theory
 end
 
 structure Fixrec : FIXREC =
@@ -406,9 +405,4 @@
     (Parse.fixes -- (Parse.where_ |-- Parse.!!! alt_specs')
       >> (fn (fixes, specs) => add_fixrec_cmd fixes specs))
 
-val setup =
-  Method.setup @{binding fixrec_simp}
-    (Scan.succeed (SIMPLE_METHOD' o fixrec_simp_tac))
-    "pattern prover for fixrec constants"
-
 end
--- a/src/HOL/Hahn_Banach/Hahn_Banach.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -151,12 +151,12 @@
         qed
       qed
 
-      def H' \<equiv> "H \<oplus> lin x'"
+      def H' \<equiv> "H + lin x'"
         -- {* Define @{text H'} as the direct sum of @{text H} and the linear closure of @{text x'}. \skp *}
       have HH': "H \<unlhd> H'"
       proof (unfold H'_def)
         from x'E have "vectorspace (lin x')" ..
-        with H show "H \<unlhd> H \<oplus> lin x'" ..
+        with H show "H \<unlhd> H + lin x'" ..
       qed
 
       obtain xi where
--- a/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Hahn_Banach/Hahn_Banach_Ext_Lemmas.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -90,7 +90,7 @@
 lemma h'_lf:
   assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
       SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H \<oplus> lin x0"
+    and H'_def: "H' \<equiv> H + lin x0"
     and HE: "H \<unlhd> E"
   assumes "linearform H h"
   assumes x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
@@ -106,7 +106,7 @@
     proof (unfold H'_def)
       from `x0 \<in> E`
       have "lin x0 \<unlhd> E" ..
-      with HE show "vectorspace (H \<oplus> lin x0)" using E ..
+      with HE show "vectorspace (H + lin x0)" using E ..
     qed
     {
       fix x1 x2 assume x1: "x1 \<in> H'" and x2: "x2 \<in> H'"
@@ -194,7 +194,7 @@
 lemma h'_norm_pres:
   assumes h'_def: "h' \<equiv> \<lambda>x. let (y, a) =
       SOME (y, a). x = y + a \<cdot> x0 \<and> y \<in> H in h y + a * xi"
-    and H'_def: "H' \<equiv> H \<oplus> lin x0"
+    and H'_def: "H' \<equiv> H + lin x0"
     and x0: "x0 \<notin> H"  "x0 \<in> E"  "x0 \<noteq> 0"
   assumes E: "vectorspace E" and HE: "subspace H E"
     and "seminorm E p" and "linearform H h"
--- a/src/HOL/Hahn_Banach/Subspace.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Hahn_Banach/Subspace.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -228,38 +228,38 @@
   set of all sums of elements from @{text U} and @{text V}.
 *}
 
-lemma sum_def: "U \<oplus> V = {u + v | u v. u \<in> U \<and> v \<in> V}"
+lemma sum_def: "U + V = {u + v | u v. u \<in> U \<and> v \<in> V}"
   unfolding set_plus_def by auto
 
 lemma sumE [elim]:
-    "x \<in> U \<oplus> V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
+    "x \<in> U + V \<Longrightarrow> (\<And>u v. x = u + v \<Longrightarrow> u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> C) \<Longrightarrow> C"
   unfolding sum_def by blast
 
 lemma sumI [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U \<oplus> V"
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> x = u + v \<Longrightarrow> x \<in> U + V"
   unfolding sum_def by blast
 
 lemma sumI' [intro]:
-    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U \<oplus> V"
+    "u \<in> U \<Longrightarrow> v \<in> V \<Longrightarrow> u + v \<in> U + V"
   unfolding sum_def by blast
 
-text {* @{text U} is a subspace of @{text "U \<oplus> V"}. *}
+text {* @{text U} is a subspace of @{text "U + V"}. *}
 
 lemma subspace_sum1 [iff]:
   assumes "vectorspace U" "vectorspace V"
-  shows "U \<unlhd> U \<oplus> V"
+  shows "U \<unlhd> U + V"
 proof -
   interpret vectorspace U by fact
   interpret vectorspace V by fact
   show ?thesis
   proof
     show "U \<noteq> {}" ..
-    show "U \<subseteq> U \<oplus> V"
+    show "U \<subseteq> U + V"
     proof
       fix x assume x: "x \<in> U"
       moreover have "0 \<in> V" ..
-      ultimately have "x + 0 \<in> U \<oplus> V" ..
-      with x show "x \<in> U \<oplus> V" by simp
+      ultimately have "x + 0 \<in> U + V" ..
+      with x show "x \<in> U + V" by simp
     qed
     fix x y assume x: "x \<in> U" and "y \<in> U"
     then show "x + y \<in> U" by simp
@@ -271,30 +271,30 @@
 
 lemma sum_subspace [intro?]:
   assumes "subspace U E" "vectorspace E" "subspace V E"
-  shows "U \<oplus> V \<unlhd> E"
+  shows "U + V \<unlhd> E"
 proof -
   interpret subspace U E by fact
   interpret vectorspace E by fact
   interpret subspace V E by fact
   show ?thesis
   proof
-    have "0 \<in> U \<oplus> V"
+    have "0 \<in> U + V"
     proof
       show "0 \<in> U" using `vectorspace E` ..
       show "0 \<in> V" using `vectorspace E` ..
       show "(0::'a) = 0 + 0" by simp
     qed
-    then show "U \<oplus> V \<noteq> {}" by blast
-    show "U \<oplus> V \<subseteq> E"
+    then show "U + V \<noteq> {}" by blast
+    show "U + V \<subseteq> E"
     proof
-      fix x assume "x \<in> U \<oplus> V"
+      fix x assume "x \<in> U + V"
       then obtain u v where "x = u + v" and
         "u \<in> U" and "v \<in> V" ..
       then show "x \<in> E" by simp
     qed
   next
-    fix x y assume x: "x \<in> U \<oplus> V" and y: "y \<in> U \<oplus> V"
-    show "x + y \<in> U \<oplus> V"
+    fix x y assume x: "x \<in> U + V" and y: "y \<in> U + V"
+    show "x + y \<in> U + V"
     proof -
       from x obtain ux vx where "x = ux + vx" and "ux \<in> U" and "vx \<in> V" ..
       moreover
@@ -306,7 +306,7 @@
         using x y by (simp_all add: add_ac)
       then show ?thesis ..
     qed
-    fix a show "a \<cdot> x \<in> U \<oplus> V"
+    fix a show "a \<cdot> x \<in> U + V"
     proof -
       from x obtain u v where "x = u + v" and "u \<in> U" and "v \<in> V" ..
       then have "a \<cdot> u \<in> U" and "a \<cdot> v \<in> V"
@@ -319,7 +319,7 @@
 text{* The sum of two subspaces is a vectorspace. *}
 
 lemma sum_vs [intro?]:
-    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U \<oplus> V)"
+    "U \<unlhd> E \<Longrightarrow> V \<unlhd> E \<Longrightarrow> vectorspace E \<Longrightarrow> vectorspace (U + V)"
   by (rule subspace.vectorspace) (rule sum_subspace)
 
 
@@ -481,7 +481,7 @@
 proof -
   interpret vectorspace E by fact
   interpret subspace H E by fact
-  from x y x' have "x \<in> H \<oplus> lin x'" by auto
+  from x y x' have "x \<in> H + lin x'" by auto
   have "\<exists>!p. (\<lambda>(y, a). x = y + a \<cdot> x' \<and> y \<in> H) p" (is "\<exists>!p. ?P p")
   proof (rule ex_ex1I)
     from x y show "\<exists>p. ?P p" by blast
--- a/src/HOL/Library/BigO.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Library/BigO.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -92,7 +92,7 @@
   by (auto simp add: bigo_def) 
 
 lemma bigo_plus_self_subset [intro]: 
-  "O(f) \<oplus> O(f) <= O(f)"
+  "O(f) + O(f) <= O(f)"
   apply (auto simp add: bigo_alt_def set_plus_def)
   apply (rule_tac x = "c + ca" in exI)
   apply auto
@@ -104,14 +104,14 @@
   apply force
 done
 
-lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)"
+lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)"
   apply (rule equalityI)
   apply (rule bigo_plus_self_subset)
   apply (rule set_zero_plus2) 
   apply (rule bigo_zero)
   done
 
-lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)"
+lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)"
   apply (rule subsetI)
   apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def)
   apply (subst bigo_pos_const [symmetric])+
@@ -153,15 +153,15 @@
   apply simp
   done
 
-lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A \<oplus> B <= O(f)"
-  apply (subgoal_tac "A \<oplus> B <= O(f) \<oplus> O(f)")
+lemma bigo_plus_subset2 [intro]: "A <= O(f) ==> B <= O(f) ==> A + B <= O(f)"
+  apply (subgoal_tac "A + B <= O(f) + O(f)")
   apply (erule order_trans)
   apply simp
   apply (auto del: subsetI simp del: bigo_plus_idemp)
   done
 
 lemma bigo_plus_eq: "ALL x. 0 <= f x ==> ALL x. 0 <= g x ==> 
-    O(f + g) = O(f) \<oplus> O(g)"
+    O(f + g) = O(f) + O(g)"
   apply (rule equalityI)
   apply (rule bigo_plus_subset)
   apply (simp add: bigo_alt_def set_plus_def func_plus)
@@ -273,12 +273,12 @@
 lemma bigo_abs5: "f =o O(g) ==> (%x. abs(f x)) =o O(g)" 
   by (unfold bigo_def, auto)
 
-lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) \<oplus> O(h)"
+lemma bigo_elt_subset2 [intro]: "f : g +o O(h) ==> O(f) <= O(g) + O(h)"
 proof -
   assume "f : g +o O(h)"
-  also have "... <= O(g) \<oplus> O(h)"
+  also have "... <= O(g) + O(h)"
     by (auto del: subsetI)
-  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
+  also have "... = O(%x. abs(g x)) + O(%x. abs(h x))"
     apply (subst bigo_abs3 [symmetric])+
     apply (rule refl)
     done
@@ -287,13 +287,13 @@
   finally have "f : ...".
   then have "O(f) <= ..."
     by (elim bigo_elt_subset)
-  also have "... = O(%x. abs(g x)) \<oplus> O(%x. abs(h x))"
+  also have "... = O(%x. abs(g x)) + O(%x. abs(h x))"
     by (rule bigo_plus_eq, auto)
   finally show ?thesis
     by (simp add: bigo_abs3 [symmetric])
 qed
 
-lemma bigo_mult [intro]: "O(f)\<otimes>O(g) <= O(f * g)"
+lemma bigo_mult [intro]: "O(f)*O(g) <= O(f * g)"
   apply (rule subsetI)
   apply (subst bigo_def)
   apply (auto simp add: bigo_alt_def set_times_def func_times)
@@ -369,7 +369,7 @@
   done
 
 lemma bigo_mult7: "ALL x. f x ~= 0 ==>
-    O(f * g) <= O(f::'a => ('b::linordered_field)) \<otimes> O(g)"
+    O(f * g) <= O(f::'a => ('b::linordered_field)) * O(g)"
   apply (subst bigo_mult6)
   apply assumption
   apply (rule set_times_mono3)
@@ -377,7 +377,7 @@
   done
 
 lemma bigo_mult8: "ALL x. f x ~= 0 ==>
-    O(f * g) = O(f::'a => ('b::linordered_field)) \<otimes> O(g)"
+    O(f * g) = O(f::'a => ('b::linordered_field)) * O(g)"
   apply (rule equalityI)
   apply (erule bigo_mult7)
   apply (rule bigo_mult)
@@ -402,9 +402,9 @@
   show "f +o O(g) <= O(g)"
   proof -
     have "f : O(f)" by auto
-    then have "f +o O(g) <= O(f) \<oplus> O(g)"
+    then have "f +o O(g) <= O(f) + O(g)"
       by (auto del: subsetI)
-    also have "... <= O(g) \<oplus> O(g)"
+    also have "... <= O(g) + O(g)"
     proof -
       from a have "O(f) <= O(g)" by (auto del: subsetI)
       thus ?thesis by (auto del: subsetI)
@@ -656,7 +656,7 @@
 subsection {* Misc useful stuff *}
 
 lemma bigo_useful_intro: "A <= O(f) ==> B <= O(f) ==>
-  A \<oplus> B <= O(f)"
+  A + B <= O(f)"
   apply (subst bigo_plus_idemp [symmetric])
   apply (rule set_plus_mono2)
   apply assumption+
--- a/src/HOL/Library/Countable.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Library/Countable.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -269,8 +269,7 @@
     by - (rule exI)
 qed
 
-method_setup countable_datatype = {*
-let
+ML {*
   fun countable_tac ctxt =
     SUBGOAL (fn (goal, i) =>
       let
@@ -297,9 +296,10 @@
            etac induct_thm i,
            REPEAT (resolve_tac rules i ORELSE atac i)]) 1
       end)
-in
+*}
+
+method_setup countable_datatype = {*
   Scan.succeed (fn ctxt => SIMPLE_METHOD' (countable_tac ctxt))
-end
 *} "prove countable class instances for datatypes"
 
 hide_const (open) finite_item nth_item
--- a/src/HOL/Library/Eval_Witness.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Library/Eval_Witness.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -77,8 +77,8 @@
 
 method_setup eval_witness = {*
   Scan.lift (Scan.repeat Args.name) >>
-  (fn ws => K (SIMPLE_METHOD'
-    (CSUBGOAL (fn (ct, i) => rtac (eval_witness_oracle (ct, ws)) i))))
+    (fn ws => K (SIMPLE_METHOD'
+      (CSUBGOAL (fn (ct, i) => rtac (eval_witness_oracle (ct, ws)) i))))
 *} "evaluation with ML witnesses"
 
 
--- a/src/HOL/Library/Set_Algebras.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Library/Set_Algebras.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -14,11 +14,45 @@
   comments at the top of theory @{text BigO}.
 *}
 
-definition set_plus :: "'a::plus set \<Rightarrow> 'a set \<Rightarrow> 'a set"  (infixl "\<oplus>" 65) where
-  "A \<oplus> B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a + b}"
+instantiation set :: (plus) plus
+begin
+
+definition plus_set :: "'a::plus set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+  set_plus_def: "A + B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a + b}"
+
+instance ..
+
+end
+
+instantiation set :: (times) times
+begin
+
+definition times_set :: "'a::times set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+  set_times_def: "A * B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a * b}"
+
+instance ..
+
+end
 
-definition set_times :: "'a::times set \<Rightarrow> 'a set \<Rightarrow> 'a set"  (infixl "\<otimes>" 70) where
-  "A \<otimes> B = {c. \<exists>a\<in>A. \<exists>b\<in>B. c = a * b}"
+instantiation set :: (zero) zero
+begin
+
+definition
+  set_zero[simp]: "0::('a::zero)set == {0}"
+
+instance ..
+
+end
+ 
+instantiation set :: (one) one
+begin
+
+definition
+  set_one[simp]: "1::('a::one)set == {1}"
+
+instance ..
+
+end
 
 definition elt_set_plus :: "'a::plus \<Rightarrow> 'a set \<Rightarrow> 'a set"  (infixl "+o" 70) where
   "a +o B = {c. \<exists>b\<in>B. c = a + b}"
@@ -29,105 +63,38 @@
 abbreviation (input) elt_set_eq :: "'a \<Rightarrow> 'a set \<Rightarrow> bool"  (infix "=o" 50) where
   "x =o A \<equiv> x \<in> A"
 
-interpretation set_add!: semigroup "set_plus :: 'a::semigroup_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof
-qed (force simp add: set_plus_def add.assoc)
-
-interpretation set_add!: abel_semigroup "set_plus :: 'a::ab_semigroup_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof
-qed (force simp add: set_plus_def add.commute)
+instance set :: (semigroup_add) semigroup_add
+by default (force simp add: set_plus_def add.assoc)
 
-interpretation set_add!: monoid "set_plus :: 'a::monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" proof
-qed (simp_all add: set_plus_def)
-
-interpretation set_add!: comm_monoid "set_plus :: 'a::comm_monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" proof
-qed (simp add: set_plus_def)
-
-definition listsum_set :: "('a::monoid_add set) list \<Rightarrow> 'a set" where
-  "listsum_set = monoid_add.listsum set_plus {0}"
+instance set :: (ab_semigroup_add) ab_semigroup_add
+by default (force simp add: set_plus_def add.commute)
 
-interpretation set_add!: monoid_add "set_plus :: 'a::monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" where
-  "monoid_add.listsum set_plus {0::'a} = listsum_set"
-proof -
-  show "class.monoid_add set_plus {0::'a}" proof
-  qed (simp_all add: set_add.assoc)
-  then interpret set_add!: monoid_add "set_plus :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" .
-  show "monoid_add.listsum set_plus {0::'a} = listsum_set"
-    by (simp only: listsum_set_def)
-qed
-
-definition setsum_set :: "('b \<Rightarrow> ('a::comm_monoid_add) set) \<Rightarrow> 'b set \<Rightarrow> 'a set" where
-  "setsum_set f A = (if finite A then fold_image set_plus f {0} A else {0})"
+instance set :: (monoid_add) monoid_add
+by default (simp_all add: set_plus_def)
 
-interpretation set_add!:
-  comm_monoid_big "set_plus :: 'a::comm_monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" setsum_set 
-proof
-qed (fact setsum_set_def)
-
-interpretation set_add!: comm_monoid_add "set_plus :: 'a::comm_monoid_add set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" where
-  "monoid_add.listsum set_plus {0::'a} = listsum_set"
-  and "comm_monoid_add.setsum set_plus {0::'a} = setsum_set"
-proof -
-  show "class.comm_monoid_add (set_plus :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set) {0}" proof
-  qed (simp_all add: set_add.commute)
-  then interpret set_add!: comm_monoid_add "set_plus :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{0}" .
-  show "monoid_add.listsum set_plus {0::'a} = listsum_set"
-    by (simp only: listsum_set_def)
-  show "comm_monoid_add.setsum set_plus {0::'a} = setsum_set"
-    by (simp add: set_add.setsum_def setsum_set_def fun_eq_iff)
-qed
+instance set :: (comm_monoid_add) comm_monoid_add
+by default (simp_all add: set_plus_def)
 
-interpretation set_mult!: semigroup "set_times :: 'a::semigroup_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof
-qed (force simp add: set_times_def mult.assoc)
-
-interpretation set_mult!: abel_semigroup "set_times :: 'a::ab_semigroup_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" proof
-qed (force simp add: set_times_def mult.commute)
+instance set :: (semigroup_mult) semigroup_mult
+by default (force simp add: set_times_def mult.assoc)
 
-interpretation set_mult!: monoid "set_times :: 'a::monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" proof
-qed (simp_all add: set_times_def)
-
-interpretation set_mult!: comm_monoid "set_times :: 'a::comm_monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" proof
-qed (simp add: set_times_def)
-
-definition power_set :: "nat \<Rightarrow> ('a::monoid_mult set) \<Rightarrow> 'a set" where
-  "power_set n A = power.power {1} set_times A n"
+instance set :: (ab_semigroup_mult) ab_semigroup_mult
+by default (force simp add: set_times_def mult.commute)
 
-interpretation set_mult!: monoid_mult "{1}" "set_times :: 'a::monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
-  "power.power {1} set_times = (\<lambda>A n. power_set n A)"
-proof -
-  show "class.monoid_mult {1} (set_times :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set)" proof
-  qed (simp_all add: set_mult.assoc)
-  show "power.power {1} set_times = (\<lambda>A n. power_set n A)"
-    by (simp add: power_set_def)
-qed
-
-definition setprod_set :: "('b \<Rightarrow> ('a::comm_monoid_mult) set) \<Rightarrow> 'b set \<Rightarrow> 'a set" where
-  "setprod_set f A = (if finite A then fold_image set_times f {1} A else {1})"
+instance set :: (monoid_mult) monoid_mult
+by default (simp_all add: set_times_def)
 
-interpretation set_mult!:
-  comm_monoid_big "set_times :: 'a::comm_monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" setprod_set 
-proof
-qed (fact setprod_set_def)
+instance set :: (comm_monoid_mult) comm_monoid_mult
+by default (simp_all add: set_times_def)
 
-interpretation set_mult!: comm_monoid_mult "set_times :: 'a::comm_monoid_mult set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" where
-  "power.power {1} set_times = (\<lambda>A n. power_set n A)"
-  and "comm_monoid_mult.setprod set_times {1::'a} = setprod_set"
-proof -
-  show "class.comm_monoid_mult (set_times :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set) {1}" proof
-  qed (simp_all add: set_mult.commute)
-  then interpret set_mult!: comm_monoid_mult "set_times :: 'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" "{1}" .
-  show "power.power {1} set_times = (\<lambda>A n. power_set n A)"
-    by (simp add: power_set_def)
-  show "comm_monoid_mult.setprod set_times {1::'a} = setprod_set"
-    by (simp add: set_mult.setprod_def setprod_set_def fun_eq_iff)
-qed
-
-lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C \<oplus> D"
+lemma set_plus_intro [intro]: "a : C ==> b : D ==> a + b : C + D"
   by (auto simp add: set_plus_def)
 
 lemma set_plus_intro2 [intro]: "b : C ==> a + b : a +o C"
   by (auto simp add: elt_set_plus_def)
 
-lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) \<oplus>
-    (b +o D) = (a + b) +o (C \<oplus> D)"
+lemma set_plus_rearrange: "((a::'a::comm_monoid_add) +o C) +
+    (b +o D) = (a + b) +o (C + D)"
   apply (auto simp add: elt_set_plus_def set_plus_def add_ac)
    apply (rule_tac x = "ba + bb" in exI)
   apply (auto simp add: add_ac)
@@ -139,8 +106,8 @@
     (a + b) +o C"
   by (auto simp add: elt_set_plus_def add_assoc)
 
-lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) \<oplus> C =
-    a +o (B \<oplus> C)"
+lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C =
+    a +o (B + C)"
   apply (auto simp add: elt_set_plus_def set_plus_def)
    apply (blast intro: add_ac)
   apply (rule_tac x = "a + aa" in exI)
@@ -151,8 +118,8 @@
    apply (auto simp add: add_ac)
   done
 
-theorem set_plus_rearrange4: "C \<oplus> ((a::'a::comm_monoid_add) +o D) =
-    a +o (C \<oplus> D)"
+theorem set_plus_rearrange4: "C + ((a::'a::comm_monoid_add) +o D) =
+    a +o (C + D)"
   apply (auto simp add: elt_set_plus_def set_plus_def add_ac)
    apply (rule_tac x = "aa + ba" in exI)
    apply (auto simp add: add_ac)
@@ -165,17 +132,17 @@
   by (auto simp add: elt_set_plus_def)
 
 lemma set_plus_mono2 [intro]: "(C::('a::plus) set) <= D ==> E <= F ==>
-    C \<oplus> E <= D \<oplus> F"
+    C + E <= D + F"
   by (auto simp add: set_plus_def)
 
-lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C \<oplus> D"
+lemma set_plus_mono3 [intro]: "a : C ==> a +o D <= C + D"
   by (auto simp add: elt_set_plus_def set_plus_def)
 
 lemma set_plus_mono4 [intro]: "(a::'a::comm_monoid_add) : C ==>
-    a +o D <= D \<oplus> C"
+    a +o D <= D + C"
   by (auto simp add: elt_set_plus_def set_plus_def add_ac)
 
-lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C \<oplus> D"
+lemma set_plus_mono5: "a:C ==> B <= D ==> a +o B <= C + D"
   apply (subgoal_tac "a +o B <= a +o D")
    apply (erule order_trans)
    apply (erule set_plus_mono3)
@@ -188,21 +155,21 @@
   apply auto
   done
 
-lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C \<oplus> E ==>
-    x : D \<oplus> F"
+lemma set_plus_mono2_b: "C <= D ==> E <= F ==> x : C + E ==>
+    x : D + F"
   apply (frule set_plus_mono2)
    prefer 2
    apply force
   apply assumption
   done
 
-lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C \<oplus> D"
+lemma set_plus_mono3_b: "a : C ==> x : a +o D ==> x : C + D"
   apply (frule set_plus_mono3)
   apply auto
   done
 
 lemma set_plus_mono4_b: "(a::'a::comm_monoid_add) : C ==>
-    x : a +o D ==> x : D \<oplus> C"
+    x : a +o D ==> x : D + C"
   apply (frule set_plus_mono4)
   apply auto
   done
@@ -210,7 +177,7 @@
 lemma set_zero_plus [simp]: "(0::'a::comm_monoid_add) +o C = C"
   by (auto simp add: elt_set_plus_def)
 
-lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A \<oplus> B"
+lemma set_zero_plus2: "(0::'a::comm_monoid_add) : A ==> B <= A + B"
   apply (auto simp add: set_plus_def)
   apply (rule_tac x = 0 in bexI)
    apply (rule_tac x = x in bexI)
@@ -231,14 +198,14 @@
   by (rule iffI, rule set_minus_imp_plus, assumption, rule set_plus_imp_minus,
     assumption)
 
-lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C \<otimes> D"
+lemma set_times_intro [intro]: "a : C ==> b : D ==> a * b : C * D"
   by (auto simp add: set_times_def)
 
 lemma set_times_intro2 [intro!]: "b : C ==> a * b : a *o C"
   by (auto simp add: elt_set_times_def)
 
-lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) \<otimes>
-    (b *o D) = (a * b) *o (C \<otimes> D)"
+lemma set_times_rearrange: "((a::'a::comm_monoid_mult) *o C) *
+    (b *o D) = (a * b) *o (C * D)"
   apply (auto simp add: elt_set_times_def set_times_def)
    apply (rule_tac x = "ba * bb" in exI)
    apply (auto simp add: mult_ac)
@@ -250,8 +217,8 @@
     (a * b) *o C"
   by (auto simp add: elt_set_times_def mult_assoc)
 
-lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) \<otimes> C =
-    a *o (B \<otimes> C)"
+lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) * C =
+    a *o (B * C)"
   apply (auto simp add: elt_set_times_def set_times_def)
    apply (blast intro: mult_ac)
   apply (rule_tac x = "a * aa" in exI)
@@ -262,8 +229,8 @@
    apply (auto simp add: mult_ac)
   done
 
-theorem set_times_rearrange4: "C \<otimes> ((a::'a::comm_monoid_mult) *o D) =
-    a *o (C \<otimes> D)"
+theorem set_times_rearrange4: "C * ((a::'a::comm_monoid_mult) *o D) =
+    a *o (C * D)"
   apply (auto simp add: elt_set_times_def set_times_def
     mult_ac)
    apply (rule_tac x = "aa * ba" in exI)
@@ -277,17 +244,17 @@
   by (auto simp add: elt_set_times_def)
 
 lemma set_times_mono2 [intro]: "(C::('a::times) set) <= D ==> E <= F ==>
-    C \<otimes> E <= D \<otimes> F"
+    C * E <= D * F"
   by (auto simp add: set_times_def)
 
-lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C \<otimes> D"
+lemma set_times_mono3 [intro]: "a : C ==> a *o D <= C * D"
   by (auto simp add: elt_set_times_def set_times_def)
 
 lemma set_times_mono4 [intro]: "(a::'a::comm_monoid_mult) : C ==>
-    a *o D <= D \<otimes> C"
+    a *o D <= D * C"
   by (auto simp add: elt_set_times_def set_times_def mult_ac)
 
-lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C \<otimes> D"
+lemma set_times_mono5: "a:C ==> B <= D ==> a *o B <= C * D"
   apply (subgoal_tac "a *o B <= a *o D")
    apply (erule order_trans)
    apply (erule set_times_mono3)
@@ -300,21 +267,21 @@
   apply auto
   done
 
-lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C \<otimes> E ==>
-    x : D \<otimes> F"
+lemma set_times_mono2_b: "C <= D ==> E <= F ==> x : C * E ==>
+    x : D * F"
   apply (frule set_times_mono2)
    prefer 2
    apply force
   apply assumption
   done
 
-lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C \<otimes> D"
+lemma set_times_mono3_b: "a : C ==> x : a *o D ==> x : C * D"
   apply (frule set_times_mono3)
   apply auto
   done
 
 lemma set_times_mono4_b: "(a::'a::comm_monoid_mult) : C ==>
-    x : a *o D ==> x : D \<otimes> C"
+    x : a *o D ==> x : D * C"
   apply (frule set_times_mono4)
   apply auto
   done
@@ -326,16 +293,16 @@
     (a * b) +o (a *o C)"
   by (auto simp add: elt_set_plus_def elt_set_times_def ring_distribs)
 
-lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B \<oplus> C) =
-    (a *o B) \<oplus> (a *o C)"
+lemma set_times_plus_distrib2: "(a::'a::semiring) *o (B + C) =
+    (a *o B) + (a *o C)"
   apply (auto simp add: set_plus_def elt_set_times_def ring_distribs)
    apply blast
   apply (rule_tac x = "b + bb" in exI)
   apply (auto simp add: ring_distribs)
   done
 
-lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) \<otimes> D <=
-    a *o D \<oplus> C \<otimes> D"
+lemma set_times_plus_distrib3: "((a::'a::semiring) +o C) * D <=
+    a *o D + C * D"
   apply (auto simp add:
     elt_set_plus_def elt_set_times_def set_times_def
     set_plus_def ring_distribs)
@@ -355,16 +322,16 @@
   by (auto simp add: elt_set_times_def)
 
 lemma set_plus_image:
-  fixes S T :: "'n::semigroup_add set" shows "S \<oplus> T = (\<lambda>(x, y). x + y) ` (S \<times> T)"
+  fixes S T :: "'n::semigroup_add set" shows "S + T = (\<lambda>(x, y). x + y) ` (S \<times> T)"
   unfolding set_plus_def by (fastforce simp: image_iff)
 
 lemma set_setsum_alt:
   assumes fin: "finite I"
-  shows "setsum_set S I = {setsum s I |s. \<forall>i\<in>I. s i \<in> S i}"
+  shows "setsum S I = {setsum s I |s. \<forall>i\<in>I. s i \<in> S i}"
     (is "_ = ?setsum I")
 using fin proof induct
   case (insert x F)
-  have "setsum_set S (insert x F) = S x \<oplus> ?setsum F"
+  have "setsum S (insert x F) = S x + ?setsum F"
     using insert.hyps by auto
   also have "...= {s x + setsum s F |s. \<forall> i\<in>insert x F. s i \<in> S i}"
     unfolding set_plus_def
@@ -380,15 +347,15 @@
 
 lemma setsum_set_cond_linear:
   fixes f :: "('a::comm_monoid_add) set \<Rightarrow> ('b::comm_monoid_add) set"
-  assumes [intro!]: "\<And>A B. P A  \<Longrightarrow> P B  \<Longrightarrow> P (A \<oplus> B)" "P {0}"
-    and f: "\<And>A B. P A  \<Longrightarrow> P B \<Longrightarrow> f (A \<oplus> B) = f A \<oplus> f B" "f {0} = {0}"
+  assumes [intro!]: "\<And>A B. P A  \<Longrightarrow> P B  \<Longrightarrow> P (A + B)" "P {0}"
+    and f: "\<And>A B. P A  \<Longrightarrow> P B \<Longrightarrow> f (A + B) = f A + f B" "f {0} = {0}"
   assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)"
-  shows "f (setsum_set S I) = setsum_set (f \<circ> S) I"
+  shows "f (setsum S I) = setsum (f \<circ> S) I"
 proof cases
   assume "finite I" from this all show ?thesis
   proof induct
     case (insert x F)
-    from `finite F` `\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)` have "P (setsum_set S F)"
+    from `finite F` `\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)` have "P (setsum S F)"
       by induct auto
     with insert show ?case
       by (simp, subst f) auto
@@ -397,8 +364,18 @@
 
 lemma setsum_set_linear:
   fixes f :: "('a::comm_monoid_add) set => ('b::comm_monoid_add) set"
-  assumes "\<And>A B. f(A) \<oplus> f(B) = f(A \<oplus> B)" "f {0} = {0}"
-  shows "f (setsum_set S I) = setsum_set (f \<circ> S) I"
+  assumes "\<And>A B. f(A) + f(B) = f(A + B)" "f {0} = {0}"
+  shows "f (setsum S I) = setsum (f \<circ> S) I"
   using setsum_set_cond_linear[of "\<lambda>x. True" f I S] assms by auto
 
+lemma set_times_Un_distrib:
+  "A * (B \<union> C) = A * B \<union> A * C"
+  "(A \<union> B) * C = A * C \<union> B * C"
+by (auto simp: set_times_def)
+
+lemma set_times_UNION_distrib:
+  "A * UNION I M = UNION I (%i. A * M i)"
+  "UNION I M * A = UNION I (%i. M i * A)"
+by (auto simp: set_times_def)
+
 end
--- a/src/HOL/Limits.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Limits.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -112,7 +112,8 @@
 qed
 
 ML {*
-  fun ev_elim_tac ctxt thms thm = let
+  fun eventually_elim_tac ctxt thms thm =
+    let
       val thy = Proof_Context.theory_of ctxt
       val mp_thms = thms RL [@{thm eventually_rev_mp}]
       val raw_elim_thm =
@@ -124,13 +125,11 @@
     in
       CASES cases (rtac raw_elim_thm 1) thm
     end
-
-  fun eventually_elim_setup name =
-    Method.setup name (Scan.succeed (fn ctxt => METHOD_CASES (ev_elim_tac ctxt)))
-      "elimination of eventually quantifiers"
 *}
 
-setup {* eventually_elim_setup @{binding "eventually_elim"} *}
+method_setup eventually_elim = {*
+  Scan.succeed (fn ctxt => METHOD_CASES (eventually_elim_tac ctxt))
+*} "elimination of eventually quantifiers"
 
 
 subsection {* Finer-than relation *}
--- a/src/HOL/Metis_Examples/Big_O.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -146,17 +146,17 @@
 by (auto simp add: bigo_def)
 
 lemma bigo_plus_self_subset [intro]:
-  "O(f) \<oplus> O(f) <= O(f)"
+  "O(f) + O(f) <= O(f)"
 apply (auto simp add: bigo_alt_def set_plus_def)
 apply (rule_tac x = "c + ca" in exI)
 apply auto
 apply (simp add: ring_distribs func_plus)
 by (metis order_trans abs_triangle_ineq add_mono)
 
-lemma bigo_plus_idemp [simp]: "O(f) \<oplus> O(f) = O(f)"
+lemma bigo_plus_idemp [simp]: "O(f) + O(f) = O(f)"
 by (metis bigo_plus_self_subset bigo_zero set_eq_subset set_zero_plus2)
 
-lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) \<oplus> O(g)"
+lemma bigo_plus_subset [intro]: "O(f + g) <= O(f) + O(g)"
 apply (rule subsetI)
 apply (auto simp add: bigo_def bigo_pos_const func_plus set_plus_def)
 apply (subst bigo_pos_const [symmetric])+
@@ -187,10 +187,10 @@
  apply (metis abs_triangle_ineq mult_le_cancel_left_pos)
 by (metis abs_ge_zero abs_of_pos zero_le_mult_iff)
 
-lemma bigo_plus_subset2 [intro]: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow> A \<oplus> B <= O(f)"
+lemma bigo_plus_subset2 [intro]: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow> A + B <= O(f)"
 by (metis bigo_plus_idemp set_plus_mono2)
 
-lemma bigo_plus_eq: "\<forall>x. 0 <= f x \<Longrightarrow> \<forall>x. 0 <= g x \<Longrightarrow> O(f + g) = O(f) \<oplus> O(g)"
+lemma bigo_plus_eq: "\<forall>x. 0 <= f x \<Longrightarrow> \<forall>x. 0 <= g x \<Longrightarrow> O(f + g) = O(f) + O(g)"
 apply (rule equalityI)
 apply (rule bigo_plus_subset)
 apply (simp add: bigo_alt_def set_plus_def func_plus)
@@ -284,25 +284,25 @@
 lemma bigo_abs5: "f =o O(g) \<Longrightarrow> (\<lambda>x. abs(f x)) =o O(g)"
 by (unfold bigo_def, auto)
 
-lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \<Longrightarrow> O(f) <= O(g) \<oplus> O(h)"
+lemma bigo_elt_subset2 [intro]: "f : g +o O(h) \<Longrightarrow> O(f) <= O(g) + O(h)"
 proof -
   assume "f : g +o O(h)"
-  also have "... <= O(g) \<oplus> O(h)"
+  also have "... <= O(g) + O(h)"
     by (auto del: subsetI)
-  also have "... = O(\<lambda>x. abs(g x)) \<oplus> O(\<lambda>x. abs(h x))"
+  also have "... = O(\<lambda>x. abs(g x)) + O(\<lambda>x. abs(h x))"
     by (metis bigo_abs3)
   also have "... = O((\<lambda>x. abs(g x)) + (\<lambda>x. abs(h x)))"
     by (rule bigo_plus_eq [symmetric], auto)
   finally have "f : ...".
   then have "O(f) <= ..."
     by (elim bigo_elt_subset)
-  also have "... = O(\<lambda>x. abs(g x)) \<oplus> O(\<lambda>x. abs(h x))"
+  also have "... = O(\<lambda>x. abs(g x)) + O(\<lambda>x. abs(h x))"
     by (rule bigo_plus_eq, auto)
   finally show ?thesis
     by (simp add: bigo_abs3 [symmetric])
 qed
 
-lemma bigo_mult [intro]: "O(f) \<otimes> O(g) <= O(f * g)"
+lemma bigo_mult [intro]: "O(f) * O(g) <= O(f * g)"
 apply (rule subsetI)
 apply (subst bigo_def)
 apply (auto simp del: abs_mult mult_ac
@@ -358,14 +358,14 @@
 declare bigo_mult6 [simp]
 
 lemma bigo_mult7:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) \<otimes> O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) \<le> O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) * O(g)"
 by (metis bigo_refl bigo_mult6 set_times_mono3)
 
 declare bigo_mult6 [simp del]
 declare bigo_mult7 [intro!]
 
 lemma bigo_mult8:
-"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) \<otimes> O(g)"
+"\<forall>x. f x \<noteq> 0 \<Longrightarrow> O(f * g) = O(f\<Colon>'a \<Rightarrow> ('b\<Colon>linordered_field)) * O(g)"
 by (metis bigo_mult bigo_mult7 order_antisym_conv)
 
 lemma bigo_minus [intro]: "f : O(g) \<Longrightarrow> - f : O(g)"
@@ -575,7 +575,7 @@
 subsection {* Misc useful stuff *}
 
 lemma bigo_useful_intro: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow>
-  A \<oplus> B <= O(f)"
+  A + B <= O(f)"
   apply (subst bigo_plus_idemp [symmetric])
   apply (rule set_plus_mono2)
   apply assumption+
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -5428,13 +5428,13 @@
 
 lemma closure_sum:
   fixes S T :: "('n::euclidean_space) set"
-  shows "closure S \<oplus> closure T \<subseteq> closure (S \<oplus> T)"
+  shows "closure S + closure T \<subseteq> closure (S + T)"
 proof-
-  have "(closure S) \<oplus> (closure T) = (\<lambda>(x,y). x + y) ` (closure S \<times> closure T)"
+  have "(closure S) + (closure T) = (\<lambda>(x,y). x + y) ` (closure S \<times> closure T)"
     by (simp add: set_plus_image)
   also have "... = (\<lambda>(x,y). x + y) ` closure (S \<times> T)"
     using closure_direct_sum by auto
-  also have "... \<subseteq> closure (S \<oplus> T)"
+  also have "... \<subseteq> closure (S + T)"
     using fst_snd_linear closure_linear_image[of "(\<lambda>(x,y). x + y)" "S \<times> T"]
     by (auto simp: set_plus_image)
   finally show ?thesis
@@ -5444,7 +5444,7 @@
 lemma convex_oplus:
 fixes S T :: "('n::euclidean_space) set"
 assumes "convex S" "convex T"
-shows "convex (S \<oplus> T)"
+shows "convex (S + T)"
 proof-
 have "{x + y |x y. x : S & y : T} = {c. EX a:S. EX b:T. c = a + b}" by auto
 thus ?thesis unfolding set_plus_def using convex_sums[of S T] assms by auto
@@ -5452,13 +5452,13 @@
 
 lemma convex_hull_sum:
 fixes S T :: "('n::euclidean_space) set"
-shows "convex hull (S \<oplus> T) = (convex hull S) \<oplus> (convex hull T)"
+shows "convex hull (S + T) = (convex hull S) + (convex hull T)"
 proof-
-have "(convex hull S) \<oplus> (convex hull T) =
+have "(convex hull S) + (convex hull T) =
       (%(x,y). x + y) ` ((convex hull S) <*> (convex hull T))"
    by (simp add: set_plus_image)
 also have "... = (%(x,y). x + y) ` (convex hull (S <*> T))" using convex_hull_direct_sum by auto
-also have "...= convex hull (S \<oplus> T)" using fst_snd_linear linear_conv_bounded_linear
+also have "...= convex hull (S + T)" using fst_snd_linear linear_conv_bounded_linear
    convex_hull_linear_image[of "(%(x,y). x + y)" "S <*> T"] by (auto simp add: set_plus_image)
 finally show ?thesis by auto
 qed
@@ -5466,12 +5466,12 @@
 lemma rel_interior_sum:
 fixes S T :: "('n::euclidean_space) set"
 assumes "convex S" "convex T"
-shows "rel_interior (S \<oplus> T) = (rel_interior S) \<oplus> (rel_interior T)"
+shows "rel_interior (S + T) = (rel_interior S) + (rel_interior T)"
 proof-
-have "(rel_interior S) \<oplus> (rel_interior T) = (%(x,y). x + y) ` (rel_interior S <*> rel_interior T)"
+have "(rel_interior S) + (rel_interior T) = (%(x,y). x + y) ` (rel_interior S <*> rel_interior T)"
    by (simp add: set_plus_image)
 also have "... = (%(x,y). x + y) ` rel_interior (S <*> T)" using rel_interior_direct_sum assms by auto
-also have "...= rel_interior (S \<oplus> T)" using fst_snd_linear convex_direct_sum assms
+also have "...= rel_interior (S + T)" using fst_snd_linear convex_direct_sum assms
    rel_interior_convex_linear_image[of "(%(x,y). x + y)" "S <*> T"] by (auto simp add: set_plus_image)
 finally show ?thesis by auto
 qed
@@ -5479,7 +5479,7 @@
 lemma convex_sum_gen:
   fixes S :: "'a \<Rightarrow> 'n::euclidean_space set"
   assumes "\<And>i. i \<in> I \<Longrightarrow> (convex (S i))"
-  shows "convex (setsum_set S I)"
+  shows "convex (setsum S I)"
 proof cases
   assume "finite I" from this assms show ?thesis
     by induct (auto simp: convex_oplus)
@@ -5487,14 +5487,14 @@
 
 lemma convex_hull_sum_gen:
 fixes S :: "'a => ('n::euclidean_space) set"
-shows "convex hull (setsum_set S I) = setsum_set (%i. (convex hull (S i))) I"
+shows "convex hull (setsum S I) = setsum (%i. (convex hull (S i))) I"
 apply (subst setsum_set_linear) using convex_hull_sum convex_hull_singleton by auto
 
 
 lemma rel_interior_sum_gen:
 fixes S :: "'a => ('n::euclidean_space) set"
 assumes "!i:I. (convex (S i))"
-shows "rel_interior (setsum_set S I) = setsum_set (%i. (rel_interior (S i))) I"
+shows "rel_interior (setsum S I) = setsum (%i. (rel_interior (S i))) I"
 apply (subst setsum_set_cond_linear[of convex])
   using rel_interior_sum rel_interior_sing[of "0"] assms by (auto simp add: convex_oplus)
 
@@ -5507,13 +5507,13 @@
 lemma convex_rel_open_sum:
 fixes S T :: "('n::euclidean_space) set"
 assumes "convex S" "rel_open S" "convex T" "rel_open T"
-shows "convex (S \<oplus> T) & rel_open (S \<oplus> T)"
+shows "convex (S + T) & rel_open (S + T)"
 by (metis assms convex_oplus rel_interior_sum rel_open_def)
 
 lemma convex_hull_finite_union_cones:
 assumes "finite I" "I ~= {}"
 assumes "!i:I. (convex (S i) & cone (S i) & (S i) ~= {})"
-shows "convex hull (Union (S ` I)) = setsum_set S I"
+shows "convex hull (Union (S ` I)) = setsum S I"
   (is "?lhs = ?rhs")
 proof-
 { fix x assume "x : ?lhs"
@@ -5547,16 +5547,16 @@
 fixes S T :: "('m::euclidean_space) set"
 assumes "convex S" "cone S" "S ~= {}"
 assumes "convex T" "cone T" "T ~= {}"
-shows "convex hull (S Un T) = S \<oplus> T"
+shows "convex hull (S Un T) = S + T"
 proof-
 def I == "{(1::nat),2}"
 def A == "(%i. (if i=(1::nat) then S else T))"
 have "Union (A ` I) = S Un T" using A_def I_def by auto
 hence "convex hull (Union (A ` I)) = convex hull (S Un T)" by auto
-moreover have "convex hull Union (A ` I) = setsum_set A I"
+moreover have "convex hull Union (A ` I) = setsum A I"
     apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def by auto
 moreover have
-  "setsum_set A I = S \<oplus> T" using A_def I_def
+  "setsum A I = S + T" using A_def I_def
      unfolding set_plus_def apply auto unfolding set_plus_def by auto
 ultimately show ?thesis by auto
 qed
@@ -5600,15 +5600,15 @@
   ultimately have "convex hull (Union (K ` I)) >= K0"
      unfolding K0_def using hull_minimal[of _ "convex hull (Union (K ` I))" "cone"] by blast
   hence "K0 = convex hull (Union (K ` I))" using geq by auto
-  also have "...=setsum_set K I"
+  also have "...=setsum K I"
      apply (subst convex_hull_finite_union_cones[of I K])
      using assms apply blast
      using `I ~= {}` apply blast
      unfolding K_def apply rule
      apply (subst convex_cone_hull) apply (subst convex_direct_sum)
      using assms cone_cone_hull `!i:I. K i ~= {}` K_def by auto
-  finally have "K0 = setsum_set K I" by auto
-  hence *: "rel_interior K0 = setsum_set (%i. (rel_interior (K i))) I"
+  finally have "K0 = setsum K I" by auto
+  hence *: "rel_interior K0 = setsum (%i. (rel_interior (K i))) I"
      using rel_interior_sum_gen[of I K] convK by auto
   { fix x assume "x : ?lhs"
     hence "((1::real),x) : rel_interior K0" using K0_def C0_def
--- a/src/HOL/NSA/StarDef.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/NSA/StarDef.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -91,6 +91,12 @@
 use "transfer.ML"
 setup Transfer_Principle.setup
 
+method_setup transfer = {*
+  Attrib.thms >> (fn ths => fn ctxt =>
+    SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths))
+*} "transfer principle"
+
+
 text {* Transfer introduction rules. *}
 
 lemma transfer_ex [transfer_intro]:
--- a/src/HOL/NSA/transfer.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/NSA/transfer.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -112,8 +112,6 @@
   Attrib.setup @{binding transfer_unfold} (Attrib.add_del unfold_add unfold_del)
     "declaration of transfer unfolding rule" #>
   Attrib.setup @{binding transfer_refold} (Attrib.add_del refold_add refold_del)
-    "declaration of transfer refolding rule" #>
-  Method.setup @{binding transfer} (Attrib.thms >> (fn ths => fn ctxt =>
-    SIMPLE_METHOD' (transfer_tac ctxt ths))) "transfer principle";
+    "declaration of transfer refolding rule"
 
 end;
--- a/src/HOL/Orderings.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Orderings.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -312,7 +312,7 @@
 signature ORDERS =
 sig
   val print_structures: Proof.context -> unit
-  val setup: theory -> theory
+  val attrib_setup: theory -> theory
   val order_tac: Proof.context -> thm list -> int -> tactic
 end;
 
@@ -414,19 +414,15 @@
     (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_context o
         Toplevel.keep (print_structures o Toplevel.context_of)));
 
-
-(** Setup **)
-
-val setup =
-  Method.setup @{binding order} (Scan.succeed (fn ctxt => SIMPLE_METHOD' (order_tac ctxt [])))
-    "transitivity reasoner" #>
-  attrib_setup;
-
 end;
 
 *}
 
-setup Orders.setup
+setup Orders.attrib_setup
+
+method_setup order = {*
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (Orders.order_tac ctxt []))
+*} "transitivity reasoner"
 
 
 text {* Declarations to set up transitivity reasoner of partial and linear orders. *}
--- a/src/HOL/Presburger.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Presburger.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -387,10 +387,25 @@
   by (simp cong: conj_cong)
 
 use "Tools/Qelim/cooper.ML"
-
 setup Cooper.setup
 
-method_setup presburger = "Cooper.method" "Cooper's algorithm for Presburger arithmetic"
+method_setup presburger = {*
+  let
+    fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+    fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
+    val addN = "add"
+    val delN = "del"
+    val elimN = "elim"
+    val any_keyword = keyword addN || keyword delN || simple_keyword elimN
+    val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
+  in
+    Scan.optional (simple_keyword elimN >> K false) true --
+    Scan.optional (keyword addN |-- thms) [] --
+    Scan.optional (keyword delN |-- thms) [] >>
+    (fn ((elim, add_ths), del_ths) => fn ctxt =>
+      SIMPLE_METHOD' (Cooper.tac elim add_ths del_ths ctxt))
+  end
+*} "Cooper's algorithm for Presburger arithmetic"
 
 declare dvd_eq_mod_eq_0[symmetric, presburger]
 declare mod_1[presburger] 
--- a/src/HOL/Tools/Function/induction_schema.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Tools/Function/induction_schema.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -9,7 +9,6 @@
   val mk_ind_tac : (int -> tactic) -> (int -> tactic) -> (int -> tactic)
                    -> Proof.context -> thm list -> tactic
   val induction_schema_tac : Proof.context -> thm list -> tactic
-  val setup : theory -> theory
 end
 
 
@@ -395,8 +394,4 @@
 fun induction_schema_tac ctxt =
   mk_ind_tac (K all_tac) (assume_tac APPEND' Goal.assume_rule_tac ctxt) (K all_tac) ctxt;
 
-val setup =
-  Method.setup @{binding induction_schema} (Scan.succeed (RAW_METHOD o induction_schema_tac))
-    "proves an induction principle"
-
 end
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -9,8 +9,6 @@
 sig
   val lex_order_tac : bool -> Proof.context -> tactic -> tactic
   val lexicographic_order_tac : bool -> Proof.context -> tactic
-  val lexicographic_order : Proof.context -> Proof.method
-
   val setup: theory -> theory
 end
 
@@ -218,12 +216,7 @@
   lex_order_tac quiet ctxt
     (auto_tac (map_simpset (fn ss => ss addsimps Function_Common.Termination_Simps.get ctxt) ctxt))
 
-val lexicographic_order = SIMPLE_METHOD o lexicographic_order_tac false
-
 val setup =
-  Method.setup @{binding lexicographic_order}
-    (Method.sections clasimp_modifiers >> (K lexicographic_order))
-    "termination prover for lexicographic orderings"
-  #> Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false))
+  Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false))
 
 end;
--- a/src/HOL/Tools/Function/pat_completeness.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Tools/Function/pat_completeness.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -7,11 +7,8 @@
 signature PAT_COMPLETENESS =
 sig
     val pat_completeness_tac: Proof.context -> int -> tactic
-    val pat_completeness: Proof.context -> Proof.method
     val prove_completeness : theory -> term list -> term -> term list list ->
       term list list -> thm
-
-    val setup : theory -> theory
 end
 
 structure Pat_Completeness : PAT_COMPLETENESS =
@@ -153,11 +150,4 @@
   end
   handle COMPLETENESS => no_tac)
 
-
-val pat_completeness = SIMPLE_METHOD' o pat_completeness_tac
-
-val setup =
-  Method.setup @{binding pat_completeness} (Scan.succeed pat_completeness)
-    "Completeness prover for datatype patterns"
-
 end
--- a/src/HOL/Tools/Qelim/cooper.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -13,7 +13,6 @@
   exception COOPER of string
   val conv: Proof.context -> conv
   val tac: bool -> thm list -> thm list -> Proof.context -> int -> tactic
-  val method: (Proof.context -> Method.method) context_parser
   val setup: theory -> theory
 end;
 
@@ -857,23 +856,6 @@
   THEN_ALL_NEW finish_tac elim
 end;
 
-val method =
-  let
-    fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
-    fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
-    val addN = "add"
-    val delN = "del"
-    val elimN = "elim"
-    val any_keyword = keyword addN || keyword delN || simple_keyword elimN
-    val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-  in
-    Scan.optional (simple_keyword elimN >> K false) true --
-    Scan.optional (keyword addN |-- thms) [] --
-    Scan.optional (keyword delN |-- thms) [] >>
-    (fn ((elim, add_ths), del_ths) => fn ctxt =>
-      SIMPLE_METHOD' (tac elim add_ths del_ths ctxt))
-  end;
-
 
 (* theory setup *)
 
--- a/src/HOL/Tools/groebner.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/Tools/groebner.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -16,7 +16,6 @@
   val ring_tac: thm list -> thm list -> Proof.context -> int -> tactic
   val ideal_tac: thm list -> thm list -> Proof.context -> int -> tactic
   val algebra_tac: thm list -> thm list -> Proof.context -> int -> tactic
-  val algebra_method: (Proof.context -> Method.method) context_parser
 end
 
 structure Groebner : GROEBNER =
@@ -1027,21 +1026,4 @@
 fun algebra_tac add_ths del_ths ctxt i =
  ring_tac add_ths del_ths ctxt i ORELSE ideal_tac add_ths del_ths ctxt i
 
-local
-
-fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
-val addN = "add"
-val delN = "del"
-val any_keyword = keyword addN || keyword delN
-val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
-
-in
-
-val algebra_method = ((Scan.optional (keyword addN |-- thms) []) --
-   (Scan.optional (keyword delN |-- thms) [])) >>
-  (fn (add_ths, del_ths) => fn ctxt =>
-       SIMPLE_METHOD' (algebra_tac add_ths del_ths ctxt))
-
 end;
-
-end;
--- a/src/HOL/ex/SAT_Examples.thy	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/HOL/ex/SAT_Examples.thy	Fri Apr 13 13:30:27 2012 +0200
@@ -80,11 +80,9 @@
 ML {* sat.trace_sat := false; *}
 ML {* quick_and_dirty := false; *}
 
-method_setup rawsat =
-  {* Scan.succeed (SIMPLE_METHOD' o sat.rawsat_tac) *}
-  "SAT solver (no preprocessing)"
-
-(* ML {* Toplevel.profiling := 1; *} *)
+method_setup rawsat = {*
+  Scan.succeed (SIMPLE_METHOD' o sat.rawsat_tac)
+*} "SAT solver (no preprocessing)"
 
 (* Translated from TPTP problem library: PUZ015-2.006.dimacs *)
 
@@ -512,8 +510,6 @@
    (as of 2006-08-30, on a 2.5 GHz Pentium 4)
 *)
 
-(* ML {* Toplevel.profiling := 0; *} *)
-
 text {*
 Function {\tt benchmark} takes the name of an existing DIMACS CNF
 file, parses this file, passes the problem to a SAT solver, and checks
--- a/src/Pure/Concurrent/lazy.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/Pure/Concurrent/lazy.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -61,15 +61,15 @@
             SOME e =>
               let
                 val res0 = Exn.capture (restore_attributes e) ();
-                val res1 = Exn.capture (fn () => Future.fulfill_result x res0) ();
-                val res2 = Future.join_result x;
-              in
+                val _ = Exn.capture (fn () => Future.fulfill_result x res0) ();
+                val res = Future.join_result x;
                 (*semantic race: some other threads might see the same
                   interrupt, until there is a fresh start*)
-                if Exn.is_interrupt_exn res1 orelse Exn.is_interrupt_exn res2 then
-                  (Synchronized.change var (fn _ => Expr e); Exn.interrupt ())
-                else res2
-              end
+                val _ =
+                  if Exn.is_interrupt_exn res then
+                    Synchronized.change var (fn _ => Expr e)
+                  else ();
+              in res end
           | NONE => Exn.capture (restore_attributes (fn () => Future.join x)) ())
         end) ());
 
--- a/src/Pure/Isar/proof.ML	Fri Apr 13 13:29:55 2012 +0200
+++ b/src/Pure/Isar/proof.ML	Fri Apr 13 13:30:27 2012 +0200
@@ -198,6 +198,9 @@
 val context_of = #context o current;
 val theory_of = Proof_Context.theory_of o context_of;
 
+fun map_node_context f =
+  map_node (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal));
+
 fun map_context f =
   map_current (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal));
 
@@ -292,24 +295,27 @@
 
 (* nested goal *)
 
-fun map_goal f g (State (Node {context, facts, mode, goal = SOME goal}, nodes)) =
+fun map_goal f g h (State (Node {context, facts, mode, goal = SOME goal}, node :: nodes)) =
       let
         val Goal {statement, messages, using, goal, before_qed, after_qed} = goal;
         val goal' = make_goal (g (statement, messages, using, goal, before_qed, after_qed));
-      in State (make_node (f context, facts, mode, SOME goal'), nodes) end
-  | map_goal f g (State (nd, node :: nodes)) =
-      let val State (node', nodes') = map_goal f g (State (node, nodes))
-      in map_context f (State (nd, node' :: nodes')) end
-  | map_goal _ _ state = state;
+        val node' = map_node_context h node;
+      in State (make_node (f context, facts, mode, SOME goal'), node' :: nodes) end
+  | map_goal f g h (State (nd, node :: nodes)) =
+      let
+        val nd' = map_node_context f nd;
+        val State (node', nodes') = map_goal f g h (State (node, nodes));
+      in State (nd', node' :: nodes') end
+  | map_goal _ _ _ state = state;
 
 fun provide_goal goal = map_goal I (fn (statement, _, using, _, before_qed, after_qed) =>
-  (statement, [], using, goal, before_qed, after_qed));
+  (statement, [], using, goal, before_qed, after_qed)) I;
 
 fun goal_message msg = map_goal I (fn (statement, messages, using, goal, before_qed, after_qed) =>
-  (statement, msg :: messages, using, goal, before_qed, after_qed));
+  (statement, msg :: messages, using, goal, before_qed, after_qed)) I;
 
 fun using_facts using = map_goal I (fn (statement, _, _, goal, before_qed, after_qed) =>
-  (statement, [], using, goal, before_qed, after_qed));
+  (statement, [], using, goal, before_qed, after_qed)) I;
 
 local
   fun find i state =
@@ -407,7 +413,7 @@
       |> map_goal
           (Proof_Context.add_cases false (no_goal_cases goal @ goal_cases goal') #>
            Proof_Context.add_cases true meth_cases)
-          (K (statement, [], using, goal', before_qed, after_qed)))
+          (K (statement, [], using, goal', before_qed, after_qed)) I)
   end;
 
 fun select_goals n meth state =
@@ -717,7 +723,7 @@
       let
         val ctxt = context_of state';
         val ths = maps snd named_facts;
-      in (statement, [], f ctxt ths using, g ctxt ths goal, before_qed, after_qed) end));
+      in (statement, [], f ctxt ths using, g ctxt ths goal, before_qed, after_qed) end) I);
 
 fun append_using _ ths using = using @ filter_out Thm.is_dummy ths;
 fun unfold_using ctxt ths = map (Local_Defs.unfold ctxt ths);
@@ -1074,7 +1080,9 @@
     val _ = assert_backward state;
     val (goal_ctxt, (_, goal)) = find_goal state;
     val {statement as (kind, _, prop), messages, using, goal, before_qed, after_qed} = goal;
-    val goal_txt = prop :: map Thm.term_of (Assumption.all_assms_of goal_ctxt);
+    val goal_tfrees =
+      fold Term.add_tfrees
+        (prop :: map Thm.term_of (Assumption.all_assms_of goal_ctxt)) [];
 
     val _ = is_relevant state andalso error "Cannot fork relevant proof";
 
@@ -1090,8 +1098,8 @@
 
     val result_ctxt =
       state
-      |> map_contexts (fold Variable.declare_term goal_txt)
       |> map_goal I (K (statement', messages, using, goal', before_qed, after_qed'))
+        (fold (Variable.declare_typ o TFree) goal_tfrees)
       |> fork_proof;
 
     val future_thm = result_ctxt |> Future.map (fn (_, x) =>
@@ -1099,7 +1107,7 @@
     val finished_goal = Goal.future_result goal_ctxt future_thm prop';
     val state' =
       state
-      |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed))
+      |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed)) I
       |> done;
   in (Future.map #1 result_ctxt, state') end;