merged
authorwenzelm
Wed, 27 Aug 2014 15:52:58 +0200
changeset 58051 be9815d02b10
parent 58050 1b6035697c49 (current diff)
parent 58049 930727de976c (diff)
child 58052 ec66337a7162
merged
src/HOL/Library/Quickcheck_Types.thy
src/HOL/SMT_Examples/SMT_Examples.certs
src/Pure/GUI/jfx_thread.scala
src/Pure/System/isabelle_font.scala
src/Pure/Thy/thm_deps.ML
src/ZF/IntArith.thy
src/ZF/Tools/twos_compl.ML
--- a/Admin/components/components.sha1	Wed Aug 27 11:33:00 2014 +0200
+++ b/Admin/components/components.sha1	Wed Aug 27 15:52:58 2014 +0200
@@ -34,6 +34,7 @@
 ec740ee9ffd43551ddf1e5b91641405116af6291  jdk-7u6.tar.gz
 7d5b152ac70f720bb9e783fa45ecadcf95069584  jdk-7u9.tar.gz
 5442f1015a0657259be0590b04572cd933431df7  jdk-8u11.tar.gz
+cfecb1383faaf027ffbabfcd77a0b6a6521e0969  jdk-8u20.tar.gz
 c95ebf7777beb3e7ef10c0cf3f734cb78f9828e4  jdk-8u5.tar.gz
 44775a22f42a9d665696bfb49e53c79371c394b0  jedit_build-20111217.tar.gz
 a242a688810f2bccf24587b0062ce8027bf77fa2  jedit_build-20120304.tar.gz
--- a/Admin/components/main	Wed Aug 27 11:33:00 2014 +0200
+++ b/Admin/components/main	Wed Aug 27 15:52:58 2014 +0200
@@ -11,6 +11,5 @@
 polyml-5.5.2-1
 scala-2.11.2
 spass-3.8ds
-z3-3.2-1
 z3-4.3.2pre-1
 xz-java-1.2-1
--- a/CONTRIBUTORS	Wed Aug 27 11:33:00 2014 +0200
+++ b/CONTRIBUTORS	Wed Aug 27 15:52:58 2014 +0200
@@ -3,6 +3,13 @@
 who is listed as an author in one of the source files of this Isabelle
 distribution.
 
+Contributions to this Isabelle version
+--------------------------------------
+
+* August 2014: Manuel Eberl, TUM
+  Generic euclidean algorithms for gcd et al.
+
+
 Contributions to Isabelle2014
 -----------------------------
 
--- a/NEWS	Wed Aug 27 11:33:00 2014 +0200
+++ b/NEWS	Wed Aug 27 15:52:58 2014 +0200
@@ -1,6 +1,67 @@
 Isabelle NEWS -- history of user-relevant changes
 =================================================
 
+New in this Isabelle version
+----------------------------
+
+*** General ***
+
+* Commands 'method_setup' and 'attribute_setup' now work within a
+local theory context.
+
+* Command 'named_theorems' declares a dynamic fact within the context,
+together with an attribute to maintain the content incrementally.
+This supersedes functor Named_Thms, but with a subtle change of
+semantics due to external visual order vs. internal reverse order.
+
+
+*** HOL ***
+
+* New (co)datatype package:
+  - Renamed theorems:
+      disc_corec ~> corec_disc
+      disc_corec_iff ~> corec_disc_iff
+      disc_exclude ~> distinct_disc
+      disc_exhaust ~> exhaust_disc
+      disc_map_iff ~> map_disc_iff
+      sel_corec ~> corec_sel
+      sel_exhaust ~> exhaust_sel
+      sel_map ~> map_sel
+      sel_set ~> set_sel
+      sel_split ~> split_sel
+      sel_split_asm ~> split_sel_asm
+      strong_coinduct ~> coinduct_strong
+      weak_case_cong ~> case_cong_weak
+    INCOMPATIBILITY.
+  - The rules "set_empty" have been removed. They are easy
+    consequences of other set rules "by auto".
+    INCOMPATIBILITY.
+  - The rule "set_cases" is now registered with the "[cases set]"
+    attribute. This can influence the behavior of the "cases" proof
+    method when more than one case rule is applicable (e.g., an
+    assumption is of the form "w : set ws" and the method "cases w"
+    is invoked). The solution is to specify the case rule explicitly
+    (e.g. "cases w rule: widget.exhaust").
+    INCOMPATIBILITY.
+
+* Old datatype package:
+  - Renamed theorems:
+      weak_case_cong ~> case_cong_weak
+    INCOMPATIBILITY.
+
+* Sledgehammer:
+  - Minimization is now always enabled by default.
+    Removed subcommand:
+      min
+
+
+*** ML ***
+
+* Tactical PARALLEL_ALLGOALS is the most common way to refer to
+PARALLEL_GOALS.
+
+
+
 New in Isabelle2014 (August 2014)
 ---------------------------------
 
--- a/etc/isar-keywords-ZF.el	Wed Aug 27 11:33:00 2014 +0200
+++ b/etc/isar-keywords-ZF.el	Wed Aug 27 15:52:58 2014 +0200
@@ -97,6 +97,7 @@
     "locale_deps"
     "method_setup"
     "moreover"
+    "named_theorems"
     "next"
     "no_notation"
     "no_syntax"
@@ -378,6 +379,7 @@
     "local_setup"
     "locale"
     "method_setup"
+    "named_theorems"
     "no_notation"
     "no_syntax"
     "no_translations"
--- a/etc/isar-keywords.el	Wed Aug 27 11:33:00 2014 +0200
+++ b/etc/isar-keywords.el	Wed Aug 27 15:52:58 2014 +0200
@@ -139,6 +139,7 @@
     "locale_deps"
     "method_setup"
     "moreover"
+    "named_theorems"
     "next"
     "nitpick"
     "nitpick_params"
@@ -550,6 +551,7 @@
     "local_setup"
     "locale"
     "method_setup"
+    "named_theorems"
     "nitpick_params"
     "no_adhoc_overloading"
     "no_notation"
--- a/src/CCL/Wfd.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/CCL/Wfd.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -483,15 +483,14 @@
 
 subsection {* Evaluation *}
 
+named_theorems eval "evaluation rules"
+
 ML {*
-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 @ Eval_Rules.get context) 1));
+  Subgoal.FOCUS_PREMS (fn {context = ctxt, prems, ...} =>
+    let val eval_rules = Named_Theorems.get ctxt @{named_theorems eval}
+    in DEPTH_SOLVE_1 (resolve_tac (ths @ prems @ rev eval_rules) 1) end)
 *}
-setup Eval_Rules.setup
 
 method_setup eval = {*
   Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ths ctxt))
--- a/src/Cube/Cube.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Cube/Cube.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -10,14 +10,7 @@
 
 setup Pure_Thy.old_appl_syntax_setup
 
-ML {*
-  structure Rules = Named_Thms
-  (
-    val name = @{binding rules}
-    val description = "Cube inference rules"
-  )
-*}
-setup Rules.setup
+named_theorems rules "Cube inference rules"
 
 typedecl "term"
 typedecl "context"
--- a/src/Doc/Datatypes/Datatypes.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Datatypes/Datatypes.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -535,7 +535,7 @@
 variable (@{typ 'a}, @{typ 'b}, \ldots) \cite{isabelle-isar-ref}.
 
 The optional names preceding the type variables allow to override the default
-names of the set functions (@{text set1_t}, \ldots, @{text setM_t}). Type
+names of the set functions (@{text set\<^sub>1_t}, \ldots, @{text set\<^sub>m_t}). Type
 arguments can be marked as dead by entering ``@{text dead}'' in front of the
 type variable (e.g., ``@{text "(dead 'a)"}''); otherwise, they are live or dead
 (and a set function is generated or not) depending on where they occur in the
@@ -647,13 +647,13 @@
 Case combinator: &
   @{text t.case_t} (rendered using the familiar @{text case}--@{text of} syntax) \\
 Discriminators: &
-  @{text "t.is_C\<^sub>1"}$, \ldots, $@{text "t.is_C\<^sub>n"} \\
+  @{text t.is_C\<^sub>1}$, \ldots, $@{text t.is_C\<^sub>n} \\
 Selectors: &
   @{text t.un_C\<^sub>11}$, \ldots, $@{text t.un_C\<^sub>1k\<^sub>1} \\
 & \quad\vdots \\
 & @{text t.un_C\<^sub>n1}$, \ldots, $@{text t.un_C\<^sub>nk\<^sub>n} \\
 Set functions: &
-  @{text t.set1_t}, \ldots, @{text t.setm_t} \\
+  @{text t.set\<^sub>1_t}, \ldots, @{text t.set\<^sub>m_t} \\
 Map function: &
   @{text t.map_t} \\
 Relator: &
@@ -773,8 +773,8 @@
 \item[@{text "t."}\hthm{case_cong} @{text "[fundef_cong]"}\rm:] ~ \\
 @{thm list.case_cong[no_vars]}
 
-\item[@{text "t."}\hthm{weak_case_cong} @{text "[cong]"}\rm:] ~ \\
-@{thm list.weak_case_cong[no_vars]}
+\item[@{text "t."}\hthm{case_cong_weak} @{text "[cong]"}\rm:] ~ \\
+@{thm list.case_cong_weak[no_vars]}
 
 \item[@{text "t."}\hthm{split}\rm:] ~ \\
 @{thm list.split[no_vars]}
@@ -809,27 +809,29 @@
 @{thm list.collapse(1)[no_vars]} \\
 @{thm list.collapse(2)[no_vars]}
 
-\item[@{text "t."}\hthm{disc_exclude} @{text "[dest]"}\rm:] ~ \\
+\item[@{text "t."}\hthm{distinct_disc} @{text "[dest]"}\rm:] ~ \\
 These properties are missing for @{typ "'a list"} because there is only one
 proper discriminator. Had the datatype been introduced with a second
 discriminator called @{const nonnull}, they would have read thusly: \\[\jot]
 @{prop "null list \<Longrightarrow> \<not> nonnull list"} \\
 @{prop "nonnull list \<Longrightarrow> \<not> null list"}
 
-\item[@{text "t."}\hthm{disc_exhaust} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
-@{thm list.disc_exhaust[no_vars]}
-
-\item[@{text "t."}\hthm{sel_exhaust} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
-@{thm list.sel_exhaust[no_vars]}
+\item[@{text "t."}\hthm{exhaust_disc} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
+@{thm list.exhaust_disc[no_vars]}
+
+\item[@{text "t."}\hthm{exhaust_sel} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
+@{thm list.exhaust_sel[no_vars]}
 
 \item[@{text "t."}\hthm{expand}\rm:] ~ \\
 @{thm list.expand[no_vars]}
 
-\item[@{text "t."}\hthm{sel_split}\rm:] ~ \\
-@{thm list.sel_split[no_vars]}
-
-\item[@{text "t."}\hthm{sel_split_asm}\rm:] ~ \\
-@{thm list.sel_split_asm[no_vars]}
+\item[@{text "t."}\hthm{split_sel}\rm:] ~ \\
+@{thm list.split_sel[no_vars]}
+
+\item[@{text "t."}\hthm{split_sel_asm}\rm:] ~ \\
+@{thm list.split_sel_asm[no_vars]}
+
+\item[@{text "t."}\hthm{split_sels} = @{text "split_sel split_sel_asm"}]
 
 \item[@{text "t."}\hthm{case_eq_if}\rm:] ~ \\
 @{thm list.case_eq_if[no_vars]}
@@ -854,25 +856,33 @@
 \begin{indentblock}
 \begin{description}
 
+\item[@{text "t."}\hthm{ctr_transfer}\rm:] ~ \\
+@{thm list.ctr_transfer(1)[no_vars]} \\
+@{thm list.ctr_transfer(2)[no_vars]}
+
 \item[@{text "t."}\hthm{set} @{text "[simp, code]"}\rm:] ~ \\
 @{thm list.set(1)[no_vars]} \\
 @{thm list.set(2)[no_vars]}
 
-\item[@{text "t."}\hthm{set_empty}\rm:] ~ \\
-@{thm list.set_empty[no_vars]}
-
-\item[@{text "t."}\hthm{sel_set}\rm:] ~ \\
-@{thm list.sel_set[no_vars]}
+\item[@{text "t."}\hthm{set_cases} @{text "[consumes 1, cases set: set\<^sub>i_t]"}\rm:] ~ \\
+@{thm list.set_cases[no_vars]}
+
+\item[@{text "t."}\hthm{set_intros}\rm:] ~ \\
+@{thm list.set_intros(1)[no_vars]} \\
+@{thm list.set_intros(2)[no_vars]}
+
+\item[@{text "t."}\hthm{set_sel}\rm:] ~ \\
+@{thm list.set_sel[no_vars]}
 
 \item[@{text "t."}\hthm{map} @{text "[simp, code]"}\rm:] ~ \\
 @{thm list.map(1)[no_vars]} \\
 @{thm list.map(2)[no_vars]}
 
-\item[@{text "t."}\hthm{disc_map_iff} @{text "[simp]"}\rm:] ~ \\
-@{thm list.disc_map_iff[no_vars]}
-
-\item[@{text "t."}\hthm{sel_map}\rm:] ~ \\
-@{thm list.sel_map[no_vars]}
+\item[@{text "t."}\hthm{map_disc_iff} @{text "[simp]"}\rm:] ~ \\
+@{thm list.map_disc_iff[no_vars]}
+
+\item[@{text "t."}\hthm{map_sel}\rm:] ~ \\
+@{thm list.map_sel[no_vars]}
 
 \item[@{text "t."}\hthm{rel_inject} @{text "[simp]"}\rm:] ~ \\
 @{thm list.rel_inject(1)[no_vars]} \\
@@ -906,15 +916,24 @@
 \begin{indentblock}
 \begin{description}
 
+\item[@{text "t."}\hthm{inj_map}\rm:] ~ \\
+@{thm list.inj_map[no_vars]}
+
+\item[@{text "t."}\hthm{inj_map_strong}\rm:] ~ \\
+@{thm list.inj_map_strong[no_vars]}
+
 \item[@{text "t."}\hthm{set_map}\rm:] ~ \\
 @{thm list.set_map[no_vars]}
 
-\item[@{text "t."}\hthm{map_comp}\rm:] ~ \\
+\item[@{text "t."}\hthm{map_comg0}\rm:] ~ \\
 @{thm list.map_cong0[no_vars]}
 
 \item[@{text "t."}\hthm{map_cong} @{text "[fundef_cong]"}\rm:] ~ \\
 @{thm list.map_cong[no_vars]}
 
+\item[@{text "t."}\hthm{map_cong_simp}\rm:] ~ \\
+@{thm list.map_cong_simp[no_vars]}
+
 \item[@{text "t."}\hthm{map_id}\rm:] ~ \\
 @{thm list.map_id[no_vars]}
 
@@ -936,6 +955,10 @@
 \item[@{text "t."}\hthm{rel_flip}\rm:] ~ \\
 @{thm list.rel_flip[no_vars]}
 
+\item[@{text "t."}\hthm{rel_map}\rm:] ~ \\
+@{thm list.rel_map(1)[no_vars]} \\
+@{thm list.rel_map(2)[no_vars]}
+
 \item[@{text "t."}\hthm{rel_mono}\rm:] ~ \\
 @{thm list.rel_mono[no_vars]}
 
@@ -956,14 +979,13 @@
 \item[@{text "t."}\hthm{induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n, induct t]"}\rm:] ~ \\
 @{thm list.induct[no_vars]}
 
-\item[@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
-Given $m > 1$ mutually recursive datatypes, this induction rule can be used to
-prove $m$ properties simultaneously.
-
 \item[@{text "t."}\hthm{rel_induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n, induct pred]"}\rm:] ~ \\
 @{thm list.rel_induct[no_vars]}
 
-\item[@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{rel_induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm:] ~ \\
+\item[\begin{tabular}{@ {}l@ {}}
+  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm: \\
+  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{rel_induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n]"}\rm: \\
+\end{tabular}] ~ \\
 Given $m > 1$ mutually recursive datatypes, this induction rule can be used to
 prove $m$ properties simultaneously.
 
@@ -1753,10 +1775,10 @@
 @{thm llist.coinduct[no_vars]}
 
 \item[\begin{tabular}{@ {}l@ {}}
-  @{text "t."}\hthm{strong_coinduct} @{text "[consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
-  \phantom{@{text "t."}\hthm{strong_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm:
+  @{text "t."}\hthm{coinduct_strong} @{text "[consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
+  \phantom{@{text "t."}\hthm{coinduct_strong} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm:
 \end{tabular}] ~ \\
-@{thm llist.strong_coinduct[no_vars]}
+@{thm llist.coinduct_strong[no_vars]}
 
 \item[\begin{tabular}{@ {}l@ {}}
   @{text "t."}\hthm{rel_coinduct} @{text "[consumes m, case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
@@ -1767,14 +1789,21 @@
 
 \item[\begin{tabular}{@ {}l@ {}}
   @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m, case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"} \\
-  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{strong_coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
-  \phantom{@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{strong_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm: \\
+  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{coinduct_strong} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
+  \phantom{@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{coinduct_strong} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm: \\
   @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{rel_coinduct} @{text "[case_names t\<^sub>1 \<dots> t\<^sub>m,"} \\
   \phantom{@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{rel_coinduct} @{text "["}}@{text "case_conclusion D\<^sub>1 \<dots> D\<^sub>n]"}\rm: \\
 \end{tabular}] ~ \\
 Given $m > 1$ mutually corecursive codatatypes, these coinduction rules can be
 used to prove $m$ properties simultaneously.
 
+\item[\begin{tabular}{@ {}l@ {}}
+  @{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{set_induct} @{text "[case_names C\<^sub>1 \<dots> C\<^sub>n,"} \\
+  \phantom{@{text "t\<^sub>1_\<dots>_t\<^sub>m."}\hthm{set_induct} @{text "["}}@{text "induct set: set\<^sub>j_t\<^sub>1, \<dots>, induct set: set\<^sub>j_t\<^sub>m]"}\rm: \\
+\end{tabular}] ~ \\
+@{thm llist.set_induct[no_vars]} \\
+If $m = 1$, the attribute @{text "[consumes 1]"} is generated as well.
+
 \item[@{text "t."}\hthm{corec}\rm:] ~ \\
 @{thm llist.corec(1)[no_vars]} \\
 @{thm llist.corec(2)[no_vars]}
@@ -1782,17 +1811,17 @@
 \item[@{text "t."}\hthm{corec_code} @{text "[code]"}\rm:] ~ \\
 @{thm llist.corec_code[no_vars]}
 
-\item[@{text "t."}\hthm{disc_corec}\rm:] ~ \\
-@{thm llist.disc_corec(1)[no_vars]} \\
-@{thm llist.disc_corec(2)[no_vars]}
-
-\item[@{text "t."}\hthm{disc_corec_iff} @{text "[simp]"}\rm:] ~ \\
-@{thm llist.disc_corec_iff(1)[no_vars]} \\
-@{thm llist.disc_corec_iff(2)[no_vars]}
-
-\item[@{text "t."}\hthm{sel_corec} @{text "[simp]"}\rm:] ~ \\
-@{thm llist.sel_corec(1)[no_vars]} \\
-@{thm llist.sel_corec(2)[no_vars]}
+\item[@{text "t."}\hthm{corec_disc}\rm:] ~ \\
+@{thm llist.corec_disc(1)[no_vars]} \\
+@{thm llist.corec_disc(2)[no_vars]}
+
+\item[@{text "t."}\hthm{corec_disc_iff} @{text "[simp]"}\rm:] ~ \\
+@{thm llist.corec_disc_iff(1)[no_vars]} \\
+@{thm llist.corec_disc_iff(2)[no_vars]}
+
+\item[@{text "t."}\hthm{corec_sel} @{text "[simp]"}\rm:] ~ \\
+@{thm llist.corec_sel(1)[no_vars]} \\
+@{thm llist.corec_sel(2)[no_vars]}
 
 \end{description}
 \end{indentblock}
@@ -1803,7 +1832,7 @@
 \begin{indentblock}
 \begin{description}
 
-\item[@{text "t."}\hthm{simps} = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.disc_corec_iff}] @{text t.sel_corec} ~ \\
+\item[@{text "t."}\hthm{simps} = @{text t.inject} @{text t.distinct} @{text t.case} @{text t.corec_disc_iff}] @{text t.corec_sel} ~ \\
 @{text t.map} @{text t.rel_inject} @{text t.rel_distinct} @{text t.set}
 
 \end{description}
@@ -2125,7 +2154,7 @@
 @{thm [source] iterate\<^sub>i\<^sub>i.coinduct},
 @{thm [source] iterates\<^sub>i\<^sub>i.coinduct}, and
 @{thm [source] iterate\<^sub>i\<^sub>i_iterates\<^sub>i\<^sub>i.coinduct}
-and analogously for @{text strong_coinduct}. These rules and the
+and analogously for @{text coinduct_strong}. These rules and the
 underlying corecursors are generated on a per-need basis and are kept in a cache
 to speed up subsequent definitions.
 *}
--- a/src/Doc/Implementation/Isar.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Implementation/Isar.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -436,32 +436,25 @@
 end
 
 text {* \medskip Apart from explicit arguments, common proof methods
-  typically work with a default configuration provided by the context.
-  As a shortcut to rule management we use a cheap solution via the
-  functor @{ML_functor Named_Thms} (see also @{file
-  "~~/src/Pure/Tools/named_thms.ML"}).  *}
+  typically work with a default configuration provided by the context. As a
+  shortcut to rule management we use a cheap solution via the @{command
+  named_theorems} command to declare a dynamic fact in the context. *}
 
-ML {*
-  structure My_Simps =
-    Named_Thms(
-      val name = @{binding my_simp}
-      val description = "my_simp rule"
-    )
-*}
-setup My_Simps.setup
+named_theorems my_simp
 
-text {* This provides ML access to a list of theorems in canonical
-  declaration order via @{ML My_Simps.get}.  The user can add or
-  delete rules via the attribute @{attribute my_simp}.  The actual
-  proof method is now defined as before, but we append the explicit
-  arguments and the rules from the context.  *}
+text {* The proof method is now defined as before, but we append the
+  explicit arguments and the rules from the context. *}
 
 method_setup my_simp' =
   \<open>Attrib.thms >> (fn thms => fn ctxt =>
-    SIMPLE_METHOD' (fn i =>
-      CHANGED (asm_full_simp_tac
-        (put_simpset HOL_basic_ss ctxt
-          addsimps (thms @ My_Simps.get ctxt)) i)))\<close>
+    let
+      val my_simps = Named_Theorems.get ctxt @{named_theorems my_simp}
+    in
+      SIMPLE_METHOD' (fn i =>
+        CHANGED (asm_full_simp_tac
+          (put_simpset HOL_basic_ss ctxt
+            addsimps (thms @ my_simps)) i))
+    end)\<close>
   "rewrite subgoal by given rules and my_simp rules from the context"
 
 text {*
@@ -500,7 +493,7 @@
   theory library, for example.
 
   This is an inherent limitation of the simplistic rule management via
-  functor @{ML_functor Named_Thms}, because it lacks tool-specific
+  @{command named_theorems}, because it lacks tool-specific
   storage and retrieval.  More realistic applications require
   efficient index-structures that organize theorems in a customized
   manner, such as a discrimination net that is indexed by the
--- a/src/Doc/Isar_Ref/Proof.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Isar_Ref/Proof.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -915,17 +915,18 @@
 
 text {*
   \begin{matharray}{rcl}
-    @{command_def "method_setup"} & : & @{text "theory \<rightarrow> theory"} \\
+    @{command_def "method_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   \end{matharray}
 
   @{rail \<open>
-    @@{command method_setup} @{syntax name} '=' @{syntax text} @{syntax text}?
+    @@{command method_setup} @{syntax target}?
+      @{syntax name} '=' @{syntax text} @{syntax text}?
   \<close>}
 
   \begin{description}
 
   \item @{command "method_setup"}~@{text "name = text description"}
-  defines a proof method in the current theory.  The given @{text
+  defines a proof method in the current context.  The given @{text
   "text"} has to be an ML expression of type
   @{ML_type "(Proof.context -> Proof.method) context_parser"}, cf.\
   basic parsers defined in structure @{ML_structure Args} and @{ML_structure
--- a/src/Doc/Isar_Ref/Spec.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Isar_Ref/Spec.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1031,7 +1031,7 @@
     @{command_def "ML_command"} & : & @{text "any \<rightarrow>"} \\
     @{command_def "setup"} & : & @{text "theory \<rightarrow> theory"} \\
     @{command_def "local_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
-    @{command_def "attribute_setup"} & : & @{text "theory \<rightarrow> theory"} \\
+    @{command_def "attribute_setup"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   \end{matharray}
   \begin{tabular}{rcll}
     @{attribute_def ML_print_depth} & : & @{text attribute} & default 10 \\
@@ -1045,7 +1045,8 @@
     (@@{command ML} | @@{command ML_prf} | @@{command ML_val} |
       @@{command ML_command} | @@{command setup} | @@{command local_setup}) @{syntax text}
     ;
-    @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}?
+    @@{command attribute_setup} @{syntax target}?
+      @{syntax name} '=' @{syntax text} @{syntax text}?
   \<close>}
 
   \begin{description}
@@ -1093,7 +1094,7 @@
   concrete outer syntax, for example.
 
   \item @{command "attribute_setup"}~@{text "name = text description"}
-  defines an attribute in the current theory.  The given @{text
+  defines an attribute in the current context.  The given @{text
   "text"} has to be an ML expression of type
   @{ML_type "attribute context_parser"}, cf.\ basic parsers defined in
   structure @{ML_structure Args} and @{ML_structure Attrib}.
@@ -1305,12 +1306,16 @@
   \begin{matharray}{rcll}
     @{command_def "lemmas"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
     @{command_def "theorems"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+    @{command_def "named_theorems"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
   \end{matharray}
 
   @{rail \<open>
     (@@{command lemmas} | @@{command theorems}) @{syntax target}? \<newline>
       (@{syntax thmdef}? @{syntax thmrefs} + @'and')
       (@'for' (@{syntax vars} + @'and'))?
+    ;
+    @@{command named_theorems} @{syntax target}?
+      @{syntax name} @{syntax text}?
   \<close>}
 
   \begin{description}
@@ -1324,6 +1329,12 @@
   \item @{command "theorems"} is the same as @{command "lemmas"}, but
   marks the result as a different kind of facts.
 
+  \item @{command "named_theorems"}~@{text "name description"} declares a
+  dynamic fact within the context. The same @{text name} is used to define
+  an attribute with the usual @{text add}/@{text del} syntax (e.g.\ see
+  \secref{sec:simp-rules}) to maintain the content incrementally, in
+  canonical declaration order of the text structure.
+
   \end{description}
 *}
 
--- a/src/Doc/Logics/document/HOL.tex	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Logics/document/HOL.tex	Wed Aug 27 15:52:58 2014 +0200
@@ -1709,7 +1709,7 @@
   the arms of the \texttt{case}-construct exposed and simplified. To ensure
   full simplification of all parts of a \texttt{case}-construct for datatype
   $t$, remove $t$\texttt{.}\ttindexbold{case_weak_cong} from the simpset, for
-  example by \texttt{delcongs [thm "$t$.weak_case_cong"]}.
+  example by \texttt{delcongs [thm "$t$.case_cong_weak"]}.
 \end{warn}
 
 \subsubsection{The function \cdx{size}}\label{sec:HOL:size}
--- a/src/Doc/Main/document/root.tex	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Main/document/root.tex	Wed Aug 27 15:52:58 2014 +0200
@@ -1,4 +1,5 @@
 \documentclass[12pt,a4paper]{article}
+\usepackage[T1]{fontenc}
 
 \oddsidemargin=4.6mm
 \evensidemargin=4.6mm
@@ -15,9 +16,9 @@
 % this should be the last package used
 \usepackage{pdfsetup}
 
-% urls in roman style, theory text in math-similar italics
+% urls in roman style, theory text in math-similar italics, with literal underscore
 \urlstyle{rm}
-\isabellestyle{it}
+\isabellestyle{literal}
 
 % for uniform font size
 \renewcommand{\isastyle}{\isastyleminor}
--- a/src/Doc/Prog_Prove/Basics.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Prog_Prove/Basics.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -73,10 +73,9 @@
 called \concept{type inference}.  Despite type inference, it is sometimes
 necessary to attach an explicit \concept{type constraint} (or \concept{type
 annotation}) to a variable or term.  The syntax is @{text "t :: \<tau>"} as in
-\mbox{\noquotes{@{prop[source] "m < (n::nat)"}}}. Type constraints may be
+\mbox{\noquotes{@{term[source] "m + (n::nat)"}}}. Type constraints may be
 needed to
-disambiguate terms involving overloaded functions such as @{text "+"}, @{text
-"*"} and @{text"\<le>"}.
+disambiguate terms involving overloaded functions such as @{text "+"}.
 
 Finally there are the universal quantifier @{text"\<And>"}\index{$4@\isasymAnd} and the implication
 @{text"\<Longrightarrow>"}\index{$3@\isasymLongrightarrow}. They are part of the Isabelle framework, not the logic
@@ -100,7 +99,7 @@
 
 Roughly speaking, a \concept{theory} is a named collection of types,
 functions, and theorems, much like a module in a programming language.
-All the Isabelle text that you ever type needs to go into a theory.
+All Isabelle text needs to go into a theory.
 The general format of a theory @{text T} is
 \begin{quote}
 \indexed{\isacom{theory}}{theory} @{text T}\\
--- a/src/Doc/Prog_Prove/Bool_nat_list.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Prog_Prove/Bool_nat_list.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -213,6 +213,7 @@
 \input{MyList.thy}\end{alltt}
 \caption{A Theory of Lists}
 \label{fig:MyList}
+\index{comment}
 \end{figure}
 
 \subsubsection{Structural Induction for Lists}
--- a/src/Doc/Prog_Prove/MyList.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Prog_Prove/MyList.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -14,4 +14,6 @@
 
 value "rev(Cons True (Cons False Nil))"
 
+(* a comment *)
+
 end
--- a/src/Doc/Prog_Prove/Types_and_funs.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Prog_Prove/Types_and_funs.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -93,8 +93,9 @@
 text{*
 Note that @{text"\<tau>\<^sub>1 * \<tau>\<^sub>2"} is the type of pairs, also written @{text"\<tau>\<^sub>1 \<times> \<tau>\<^sub>2"}.
 Pairs can be taken apart either by pattern matching (as above) or with the
-projection functions @{const fst} and @{const snd}: @{thm fst_conv} and @{thm snd_conv}. Tuples are simulated by pairs nested to the right: @{term"(a,b,c)"}
-abbreviates @{text"(a, (b, c))"} and @{text "\<tau>\<^sub>1 \<times> \<tau>\<^sub>2 \<times> \<tau>\<^sub>3"} abbreviates
+projection functions @{const fst} and @{const snd}: @{thm fst_conv[of x y]} and @{thm snd_conv[of x y]}.
+Tuples are simulated by pairs nested to the right: @{term"(a,b,c)"}
+is short for @{text"(a, (b, c))"} and @{text "\<tau>\<^sub>1 \<times> \<tau>\<^sub>2 \<times> \<tau>\<^sub>3"} is short for
 @{text "\<tau>\<^sub>1 \<times> (\<tau>\<^sub>2 \<times> \<tau>\<^sub>3)"}.
 
 \subsection{Definitions}
@@ -388,7 +389,7 @@
 \begin{array}{r@ {}c@ {}l@ {\quad}l}
 @{text"(0 + Suc 0"} & \leq & @{text"Suc 0 + x)"}  & \stackrel{(1)}{=} \\
 @{text"(Suc 0"}     & \leq & @{text"Suc 0 + x)"}  & \stackrel{(2)}{=} \\
-@{text"(Suc 0"}     & \leq & @{text"Suc (0 + x)"} & \stackrel{(3)}{=} \\
+@{text"(Suc 0"}     & \leq & @{text"Suc (0 + x))"} & \stackrel{(3)}{=} \\
 @{text"(0"}         & \leq & @{text"0 + x)"}      & \stackrel{(4)}{=} \\[1ex]
  & @{const True}
 \end{array}
--- a/src/Doc/Prog_Prove/document/intro-isabelle.tex	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Prog_Prove/document/intro-isabelle.tex	Wed Aug 27 15:52:58 2014 +0200
@@ -82,13 +82,13 @@
 \emph{Programming and Proving in Isabelle/HOL} constitutes part~I of
 \href{http://www.concrete-semantics.org}{Concrete Semantics}.  The web
 pages for \href{http://www.concrete-semantics.org}{Concrete Semantics}
-also provide a set of \LaTeX-based slides for teaching \emph{Programming and
-Proving in Isabelle/HOL}.
+also provide a set of \LaTeX-based slides and Isabelle demo files
+for teaching \emph{Programming and Proving in Isabelle/HOL}.
 \fi
 
 \ifsem\else
 \paragraph{Acknowledgements}
 I wish to thank the following people for their comments on this document:
-Florian Haftmann, Ren\'{e} Thiemann, Sean Seefried, Christian Sternagel
-and Carl Witty.
+Florian Haftmann, Peter Johnson, Ren\'{e} Thiemann, Sean Seefried,
+Christian Sternagel and Carl Witty.
 \fi
\ No newline at end of file
--- a/src/Doc/Sledgehammer/document/root.tex	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/Sledgehammer/document/root.tex	Wed Aug 27 15:52:58 2014 +0200
@@ -447,36 +447,15 @@
 \point{Why does Metis fail to reconstruct the proof?}
 
 There are many reasons. If Metis runs seemingly forever, that is a sign that the
-proof is too difficult for it. Metis's search is complete, so it should
-eventually find it, but that's little consolation. There are several possible
-solutions:
-
-\begin{enum}
-\item[\labelitemi] Try the \textit{isar\_proofs} option (\S\ref{output-format}) to
-obtain a step-by-step Isar proof. Since the steps are fairly small, \textit{metis}
-and the other Isabelle proof methods are more likely to be able to replay them.
-
-\item[\labelitemi] Try the \textit{smt2} proof method instead of \textit{metis}.
-It is usually stronger, but you need to either have Z3 available to replay the
-proofs, trust the SMT solver, or use certificates. See the documentation in the
-\textit{SMT2} theory (\texttt{\$ISABELLE\_HOME/src/HOL/SMT2.thy}) for details.
-
-\item[\labelitemi] Try the \textit{blast} or \textit{auto} proof methods, passing
-the necessary facts via \textbf{unfolding}, \textbf{using}, \textit{intro}{:},
-\textit{elim}{:}, \textit{dest}{:}, or \textit{simp}{:}, as appropriate.
-\end{enum}
+proof is too difficult for it. Metis's search is complete for first-order logic
+with equality, so if the proof was found by an ATP such as E, SPASS, or Vampire,
+Metis should eventually find it, but that's little consolation.
 
 In some rare cases, \textit{metis} fails fairly quickly, and you get the error
-message
-
-\prew
-\slshape
-One-line proof reconstruction failed.
-\postw
-
-This message indicates that Sledgehammer determined that the goal is provable,
-but the proof is, for technical reasons, beyond \textit{metis}'s power. You can
-then try again with the \textit{strict} option (\S\ref{problem-encoding}).
+message ``One-line proof reconstruction failed.'' This indicates that
+Sledgehammer determined that the goal is provable, but the proof is, for
+technical reasons, beyond \textit{metis}'s power. You can then try again with
+the \textit{strict} option (\S\ref{problem-encoding}).
 
 If the goal is actually unprovable and you did not specify an unsound encoding
 using \textit{type\_enc} (\S\ref{problem-encoding}), this is a bug, and you are
@@ -519,7 +498,7 @@
 generated by Sledgehammer instead of \textit{metis} if the proof obviously
 requires type information or if \textit{metis} failed when Sledgehammer
 preplayed the proof. (By default, Sledgehammer tries to run \textit{metis} with
-various sets of option for up to 2~seconds each time to ensure that the generated
+various sets of option for up to 1~second each time to ensure that the generated
 one-line proofs actually work and to display timing information. This can be
 configured using the \textit{preplay\_timeout} and \textit{dont\_preplay}
 options (\S\ref{timeouts}).)
@@ -554,26 +533,15 @@
 \point{Are generated proofs minimal?}
 
 Automatic provers frequently use many more facts than are necessary.
-Sledgehammer inclues a minimization tool that takes a set of facts returned by a
-given prover and repeatedly calls the same prover, \textit{metis}, or
-\textit{smt2} with subsets of those axioms in order to find a minimal set.
-Reducing the number of axioms typically improves Metis's speed and success rate,
-while also removing superfluous clutter from the proof scripts.
+Sledgehammer includes a minimization tool that takes a set of facts returned by
+a given prover and repeatedly calls a prover or proof method with subsets of
+those facts to find a minimal set. Reducing the number of facts typically helps
+reconstruction, while also removing superfluous clutter from the proof scripts.
 
 In earlier versions of Sledgehammer, generated proofs were systematically
 accompanied by a suggestion to invoke the minimization tool. This step is now
-performed implicitly if it can be done in a reasonable amount of time (something
-that can be guessed from the number of facts in the original proof and the time
-it took to find or preplay it).
-
-In addition, some provers do not provide proofs or sometimes produce incomplete
-proofs. The minimizer is then invoked to find out which facts are actually needed
-from the (large) set of facts that was initially given to the prover. Finally,
-if a prover returns a proof with lots of facts, the minimizer is invoked
-automatically since Metis would be unlikely to re-find the proof.
-%
-Automatic minimization can be forced or disabled using the \textit{minimize}
-option (\S\ref{mode-of-operation}).
+performed by default but can be disabled using the \textit{minimize} option
+(\S\ref{mode-of-operation}).
 
 \point{A strange error occurred---what should I do?}
 
@@ -623,10 +591,6 @@
 \item[\labelitemi] \textbf{\textit{run} (the default):} Runs Sledgehammer on
 subgoal number \qty{num} (1 by default), with the given options and facts.
 
-\item[\labelitemi] \textbf{\textit{min}:} Attempts to minimize the facts
-specified in the \qty{facts\_override} argument to obtain a simpler proof
-involving fewer facts. The options and goal number are as for \textit{run}.
-
 \item[\labelitemi] \textbf{\textit{messages}:} Redisplays recent messages issued
 by Sledgehammer. This allows you to examine results that might have been lost
 due to Sledgehammer's asynchronous nature. The \qty{num} argument specifies a
@@ -973,16 +937,6 @@
 SPASS, and Vampire for 5~seconds yields a similar success rate to running the
 most effective of these for 120~seconds \cite{boehme-nipkow-2010}.
 
-In addition to the local and remote provers, the Isabelle proof methods
-\textit{metis} and \textit{smt2} can be specified as \textbf{\textit{metis}}
-and \textbf{\textit{smt}}, respectively. They are generally not recommended
-for proof search but occasionally arise in Sledgehammer-generated
-minimization commands (e.g.,
-``\textbf{sledgehammer} \textit{min} [\textit{prover} = \textit{metis}]'').
-
-For the \textit{min} subcommand, the default prover is \textit{metis}. If
-several provers are set, the first one is used.
-
 \opnodefault{prover}{string}
 Alias for \textit{provers}.
 
@@ -1008,12 +962,9 @@
 \nopagebreak
 {\small See also \textit{verbose} (\S\ref{output-format}).}
 
-\opsmart{minimize}{dont\_minimize}
+\optrue{minimize}{dont\_minimize}
 Specifies whether the minimization tool should be invoked automatically after
-proof search. By default, automatic minimization takes place only if
-it can be done in a reasonable amount of time (as determined by
-the number of facts in the original proof and the time it took to find or
-preplay it) or the proof involves an unreasonably large number of facts.
+proof search.
 
 \nopagebreak
 {\small See also \textit{preplay\_timeout} (\S\ref{timeouts})
@@ -1321,13 +1272,16 @@
 one-line proofs. If the option is set to \textit{smart} (the default), Isar
 proofs are only generated when no working one-line proof is available.
 
-\opdefault{compress}{int}{\upshape 10}
+\opdefault{compress}{int}{smart}
 Specifies the granularity of the generated Isar proofs if \textit{isar\_proofs}
 is explicitly enabled. A value of $n$ indicates that each Isar proof step should
-correspond to a group of up to $n$ consecutive proof steps in the ATP proof.
+correspond to a group of up to $n$ consecutive proof steps in the ATP proof. If
+the option is set to \textit{smart} (the default), the compression factor is 10
+if the \textit{isar\_proofs} option is explicitly enabled; otherwise, it is
+$\infty$.
 
 \optrueonly{dont\_compress}
-Alias for ``\textit{compress} = 0''.
+Alias for ``\textit{compress} = 1''.
 
 \optrue{try0}{dont\_try0}
 Specifies whether standard proof methods such as \textit{auto} and
@@ -1335,8 +1289,8 @@
 The collection of methods is roughly the same as for the \textbf{try0} command.
 
 \opsmart{smt\_proofs}{no\_smt\_proofs}
-Specifies whether the \textit{smt2} proof method should be tried as an
-alternative to \textit{metis}.  If the option is set to \textit{smart} (the
+Specifies whether the \textit{smt2} proof method should be tried in addition to
+Isabelle's other proof methods. If the option is set to \textit{smart} (the
 default), the \textit{smt2} method is used for one-line proofs but not in Isar
 proofs.
 \end{enum}
@@ -1373,7 +1327,7 @@
 Specifies the maximum number of seconds that the automatic provers should spend
 searching for a proof. This excludes problem preparation and is a soft limit.
 
-\opdefault{preplay\_timeout}{float}{\upshape 2}
+\opdefault{preplay\_timeout}{float}{\upshape 1}
 Specifies the maximum number of seconds that \textit{metis} or other proof
 methods should spend trying to ``preplay'' the found proof. If this option
 is set to 0, no preplaying takes place, and no timing information is displayed
--- a/src/Doc/antiquote_setup.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Doc/antiquote_setup.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -208,7 +208,7 @@
   is_some (Keyword.command_keyword name) andalso
     let
       val markup =
-        Outer_Syntax.scan Position.none name
+        Outer_Syntax.scan (Keyword.get_lexicons ()) Position.none name
         |> maps (Outer_Syntax.command_reports (#2 (Outer_Syntax.get_syntax ())))
         |> map (snd o fst);
       val _ = Context_Position.reports ctxt (map (pair pos) markup);
--- a/src/FOL/IFOL.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/FOL/IFOL.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -710,16 +710,14 @@
 
 subsection {* Atomizing elimination rules *}
 
-setup AtomizeElim.setup
-
 lemma atomize_exL[atomize_elim]: "(!!x. P(x) ==> Q) == ((EX x. P(x)) ==> Q)"
-by rule iprover+
+  by rule iprover+
 
 lemma atomize_conjL[atomize_elim]: "(A ==> B ==> C) == (A & B ==> C)"
-by rule iprover+
+  by rule iprover+
 
 lemma atomize_disjL[atomize_elim]: "((A ==> C) ==> (B ==> C) ==> C) == ((A | B ==> C) ==> C)"
-by rule iprover+
+  by rule iprover+
 
 lemma atomize_elimL[atomize_elim]: "(!!B. (A ==> B) ==> B) == Trueprop(A)" ..
 
--- a/src/HOL/ATP.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/ATP.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -15,6 +15,8 @@
 ML_file "Tools/ATP/atp_problem.ML"
 ML_file "Tools/ATP/atp_proof.ML"
 ML_file "Tools/ATP/atp_proof_redirect.ML"
+ML_file "Tools/ATP/atp_satallax.ML"
+
 
 subsection {* Higher-order reasoning helpers *}
 
--- a/src/HOL/Archimedean_Field.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Archimedean_Field.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -174,6 +174,9 @@
 lemma floor_le_iff: "floor x \<le> z \<longleftrightarrow> x < of_int z + 1"
   by (simp add: not_less [symmetric] less_floor_iff)
 
+lemma floor_split[arith_split]: "P (floor t) \<longleftrightarrow> (\<forall>i. of_int i \<le> t \<and> t < of_int i + 1 \<longrightarrow> P i)"
+  by (metis floor_correct floor_unique less_floor_iff not_le order_refl)
+
 lemma floor_mono: assumes "x \<le> y" shows "floor x \<le> floor y"
 proof -
   have "of_int (floor x) \<le> x" by (rule of_int_floor_le)
@@ -285,7 +288,6 @@
 lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1"
   using floor_diff_of_int [of x 1] by simp
 
-
 subsection {* Ceiling function *}
 
 definition
@@ -426,6 +428,9 @@
 lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1"
   using ceiling_diff_of_int [of x 1] by simp
 
+lemma ceiling_split[arith_split]: "P (ceiling t) \<longleftrightarrow> (\<forall>i. of_int i - 1 < t \<and> t \<le> of_int i \<longrightarrow> P i)"
+  by (auto simp add: ceiling_unique ceiling_correct)
+
 lemma ceiling_diff_floor_le_1: "ceiling x - floor x \<le> 1"
 proof -
   have "of_int \<lceil>x\<rceil> - 1 < x" 
--- a/src/HOL/BNF_Comp.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_Comp.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1,6 +1,7 @@
 (*  Title:      HOL/BNF_Comp.thy
     Author:     Dmitriy Traytel, TU Muenchen
-    Copyright   2012
+    Author:     Jasmin Blanchette, TU Muenchen
+    Copyright   2012, 2013, 2014
 
 Composition of bounded natural functors.
 *)
--- a/src/HOL/BNF_Def.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_Def.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1,7 +1,7 @@
 (*  Title:      HOL/BNF_Def.thy
     Author:     Dmitriy Traytel, TU Muenchen
     Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012
+    Copyright   2012, 2013, 2014
 
 Definition of bounded natural functors.
 *)
@@ -159,6 +159,11 @@
 "case_sum f g \<circ> Inr = g"
 by auto
 
+lemma map_sum_o_inj:
+"map_sum f g o Inl = Inl o f"
+"map_sum f g o Inr = Inr o g"
+by auto
+
 lemma card_order_csum_cone_cexp_def:
   "card_order r \<Longrightarrow> ( |A1| +c cone) ^c r = |Func UNIV (Inl ` A1 \<union> {Inr ()})|"
   unfolding cexp_def cone_def Field_csum Field_card_of by (auto dest: Field_card_order)
--- a/src/HOL/BNF_Examples/Derivation_Trees/DTree.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_Examples/Derivation_Trees/DTree.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -47,13 +47,13 @@
 using Node unfolding Node_def
 by (metis Node Node_root_cont finite_cont)
 
-lemma dtree_sel_ctor[simp]:
+lemma dtree_sel_ctr[simp]:
 "root (Node n as) = n"
 "finite as \<Longrightarrow> cont (Node n as) = as"
 unfolding Node_def cont_def by auto
 
-lemmas root_Node = dtree_sel_ctor(1)
-lemmas cont_Node = dtree_sel_ctor(2)
+lemmas root_Node = dtree_sel_ctr(1)
+lemmas cont_Node = dtree_sel_ctr(2)
 
 lemma dtree_cong:
 assumes "root tr = root tr'" and "cont tr = cont tr'"
@@ -75,7 +75,7 @@
 lemma unfold:
 "root (unfold rt ct b) = rt b"
 "finite (ct b) \<Longrightarrow> cont (unfold rt ct b) = image (id \<oplus> unfold rt ct) (ct b)"
-using dtree.sel_corec[of rt "the_inv fset o image (map_sum id Inr) o ct" b] unfolding unfold_def
+using dtree.corec_sel[of rt "the_inv fset o image (map_sum id Inr) o ct" b] unfolding unfold_def
 apply blast
 unfolding cont_def comp_def
 by (simp add: case_sum_o_inj map_sum.compositionality image_image)
@@ -83,7 +83,7 @@
 lemma corec:
 "root (corec rt ct b) = rt b"
 "finite (ct b) \<Longrightarrow> cont (corec rt ct b) = image (id \<oplus> ([[id, corec rt ct]])) (ct b)"
-using dtree.sel_corec[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
+using dtree.corec_sel[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
 unfolding cont_def comp_def id_def
 by simp_all
 
--- a/src/HOL/BNF_Examples/Misc_Datatype.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_Examples/Misc_Datatype.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -184,4 +184,45 @@
 datatype_new d5'' = is_D: D nat | E int
 datatype_new d5''' = is_D: D nat | is_E: E int
 
+datatype_compat simple
+datatype_compat simple'
+datatype_compat simple''
+datatype_compat mylist
+datatype_compat some_passive
+datatype_compat I1 I2
+datatype_compat tree forest
+datatype_compat tree' branch
+datatype_compat bin_rose_tree
+datatype_compat exp trm factor
+datatype_compat ftree
+datatype_compat nofail1
+datatype_compat kk1 kk2 kk3
+datatype_compat t1 t2 t3
+datatype_compat t1' t2' t3'
+datatype_compat k1 k2 k3 k4
+datatype_compat tt1 tt2 tt3 tt4
+datatype_compat deadbar
+datatype_compat deadbar_option
+datatype_compat bar
+datatype_compat foo
+datatype_compat deadfoo
+datatype_compat dead_foo
+datatype_compat use_dead_foo
+datatype_compat d1
+datatype_compat d1'
+datatype_compat d2
+datatype_compat d2'
+datatype_compat d3
+datatype_compat d3'
+datatype_compat d3''
+datatype_compat d3'''
+datatype_compat d4
+datatype_compat d4'
+datatype_compat d4''
+datatype_compat d4'''
+datatype_compat d5
+datatype_compat d5'
+datatype_compat d5''
+datatype_compat d5'''
+
 end
--- a/src/HOL/BNF_Examples/Process.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_Examples/Process.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -29,7 +29,7 @@
 (* Constructors versus discriminators *)
 theorem isAction_isChoice:
 "isAction p \<or> isChoice p"
-by (rule process.disc_exhaust) auto
+by (rule process.exhaust_disc) auto
 
 theorem not_isAction_isChoice: "\<not> (isAction p \<and> isChoice p)"
 by (cases rule: process.exhaust[of p]) auto
@@ -54,7 +54,7 @@
   Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> (\<phi> p p' \<or> p = p') \<and> (\<phi> q q' \<or> q = q')"
   shows "p = p'"
   using assms
-  by (coinduct rule: process.strong_coinduct) (metis process.collapse(1,2) process.disc(3))
+  by (coinduct rule: process.coinduct_strong) (metis process.collapse(1,2) process.disc(3))
 
 
 subsection {* Coiteration (unfold) *}
--- a/src/HOL/BNF_Examples/Stream.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_Examples/Stream.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -27,39 +27,14 @@
 
 hide_const (open) smember
 
-(* TODO: Provide by the package*)
-theorem sset_induct:
-  assumes Base: "\<And>s. P (shd s) s" and Step: "\<And>s y. \<lbrakk>y \<in> sset (stl s); P y (stl s)\<rbrakk> \<Longrightarrow> P y s"
-  shows "\<forall>y \<in> sset s. P y s"
-proof (rule stream.dtor_set_induct)
-  fix a :: 'a and s :: "'a stream"
-  assume "a \<in> set1_pre_stream (dtor_stream s)"
-  then have "a = shd s"
-    by (cases "dtor_stream s")
-      (auto simp: BNF_Comp.id_bnf_comp_def shd_def fsts_def set1_pre_stream_def stream.dtor_ctor SCons_def
-        split: stream.splits)
-  with Base show "P a s" by simp
-next
-  fix a :: 'a and s' :: "'a stream"  and s :: "'a stream"
-  assume "s' \<in> set2_pre_stream (dtor_stream s)" and prems: "a \<in> sset s'" "P a s'"
-  then have "s' = stl s"
-    by (cases "dtor_stream s")
-      (auto simp: BNF_Comp.id_bnf_comp_def stl_def snds_def set2_pre_stream_def stream.dtor_ctor SCons_def
-        split: stream.splits)
-  with Step prems show "P a s" by simp
-qed
+lemmas smap_simps[simp] = stream.map_sel
+lemmas shd_sset = stream.set_sel(1)
+lemmas stl_sset = stream.set_sel(2)
 
-lemmas smap_simps[simp] = stream.sel_map
-lemmas shd_sset = stream.sel_set(1)
-lemmas stl_sset = stream.sel_set(2)
-
-(* only for the non-mutual case: *)
-theorem sset_induct1[consumes 1, case_names shd stl, induct set: "sset"]:
-  assumes "y \<in> sset s" and "\<And>s. P (shd s) s"
-  and "\<And>s y. \<lbrakk>y \<in> sset (stl s); P y (stl s)\<rbrakk> \<Longrightarrow> P y s"
+theorem sset_induct[consumes 1, case_names shd stl, induct set: sset]:
+  assumes "y \<in> sset s" and "\<And>s. P (shd s) s" and "\<And>s y. \<lbrakk>y \<in> sset (stl s); P y (stl s)\<rbrakk> \<Longrightarrow> P y s"
   shows "P y s"
-  using assms sset_induct by blast
-(* end TODO *)
+using assms by induct (metis stream.sel(1), auto)
 
 
 subsection {* prepend list to stream *}
@@ -456,7 +431,7 @@
   thus "x \<in> ?R" by (auto simp: sset_range dest!: nth_mem)
 next
   fix x xs assume "xs \<in> sset s" ?P "x \<in> set xs" thus "x \<in> ?L"
-    by (induct rule: sset_induct1)
+    by (induct rule: sset_induct)
       (metis UnI1 flat_unfold shift.simps(1) sset_shift,
        metis UnI2 flat_unfold shd_sset stl_sset sset_shift)
 qed
--- a/src/HOL/BNF_FP_Base.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_FP_Base.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -2,7 +2,8 @@
     Author:     Lorenz Panny, TU Muenchen
     Author:     Dmitriy Traytel, TU Muenchen
     Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
+    Author:     Martin Desharnais, TU Muenchen
+    Copyright   2012, 2013, 2014
 
 Shared fixed point operations on bounded natural functors.
 *)
--- a/src/HOL/BNF_GFP.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_GFP.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -2,7 +2,7 @@
     Author:     Dmitriy Traytel, TU Muenchen
     Author:     Lorenz Panny, TU Muenchen
     Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
+    Copyright   2012, 2013, 2014
 
 Greatest fixed point operation on bounded natural functors.
 *)
@@ -22,33 +22,33 @@
 *}
 
 lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
-by simp
+  by simp
 
 lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
-by (cases s) auto
+  by (cases s) auto
 
 lemma not_TrueE: "\<not> True \<Longrightarrow> P"
-by (erule notE, rule TrueI)
+  by (erule notE, rule TrueI)
 
 lemma neq_eq_eq_contradict: "\<lbrakk>t \<noteq> u; s = t; s = u\<rbrakk> \<Longrightarrow> P"
-by fast
+  by fast
 
 lemma case_sum_expand_Inr: "f o Inl = g \<Longrightarrow> f x = case_sum g (f o Inr) x"
-by (auto split: sum.splits)
+  by (auto split: sum.splits)
 
 lemma case_sum_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> case_sum g h = f"
-apply rule
- apply (rule ext, force split: sum.split)
-by (rule ext, metis case_sum_o_inj(2))
+  apply rule
+   apply (rule ext, force split: sum.split)
+  by (rule ext, metis case_sum_o_inj(2))
 
 lemma converse_Times: "(A \<times> B) ^-1 = B \<times> A"
-by fast
+  by fast
 
 lemma equiv_proj:
-  assumes e: "equiv A R" and "z \<in> R"
+  assumes e: "equiv A R" and m: "z \<in> R"
   shows "(proj R o fst) z = (proj R o snd) z"
 proof -
-  from assms(2) have z: "(fst z, snd z) \<in> R" by auto
+  from m have z: "(fst z, snd z) \<in> R" by auto
   with e have "\<And>x. (fst z, x) \<in> R \<Longrightarrow> (snd z, x) \<in> R" "\<And>x. (snd z, x) \<in> R \<Longrightarrow> (fst z, x) \<in> R"
     unfolding equiv_def sym_def trans_def by blast+
   then show ?thesis unfolding proj_def[abs_def] by auto
@@ -58,93 +58,93 @@
 definition image2 where "image2 A f g = {(f a, g a) | a. a \<in> A}"
 
 lemma Id_on_Gr: "Id_on A = Gr A id"
-unfolding Id_on_def Gr_def by auto
+  unfolding Id_on_def Gr_def by auto
 
 lemma image2_eqI: "\<lbrakk>b = f x; c = g x; x \<in> A\<rbrakk> \<Longrightarrow> (b, c) \<in> image2 A f g"
-unfolding image2_def by auto
+  unfolding image2_def by auto
 
 lemma IdD: "(a, b) \<in> Id \<Longrightarrow> a = b"
-by auto
+  by auto
 
 lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)"
-unfolding image2_def Gr_def by auto
+  unfolding image2_def Gr_def by auto
 
 lemma GrD1: "(x, fx) \<in> Gr A f \<Longrightarrow> x \<in> A"
-unfolding Gr_def by simp
+  unfolding Gr_def by simp
 
 lemma GrD2: "(x, fx) \<in> Gr A f \<Longrightarrow> f x = fx"
-unfolding Gr_def by simp
+  unfolding Gr_def by simp
 
 lemma Gr_incl: "Gr A f \<subseteq> A <*> B \<longleftrightarrow> f ` A \<subseteq> B"
-unfolding Gr_def by auto
+  unfolding Gr_def by auto
 
 lemma subset_Collect_iff: "B \<subseteq> A \<Longrightarrow> (B \<subseteq> {x \<in> A. P x}) = (\<forall>x \<in> B. P x)"
-by blast
+  by blast
 
 lemma subset_CollectI: "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> Q x \<Longrightarrow> P x) \<Longrightarrow> ({x \<in> B. Q x} \<subseteq> {x \<in> A. P x})"
-by blast
+  by blast
 
 lemma in_rel_Collect_split_eq: "in_rel (Collect (split X)) = X"
-unfolding fun_eq_iff by auto
+  unfolding fun_eq_iff by auto
 
 lemma Collect_split_in_rel_leI: "X \<subseteq> Y \<Longrightarrow> X \<subseteq> Collect (split (in_rel Y))"
-by auto
+  by auto
 
 lemma Collect_split_in_rel_leE: "X \<subseteq> Collect (split (in_rel Y)) \<Longrightarrow> (X \<subseteq> Y \<Longrightarrow> R) \<Longrightarrow> R"
-by force
+  by force
 
 lemma conversep_in_rel: "(in_rel R)\<inverse>\<inverse> = in_rel (R\<inverse>)"
-unfolding fun_eq_iff by auto
+  unfolding fun_eq_iff by auto
 
 lemma relcompp_in_rel: "in_rel R OO in_rel S = in_rel (R O S)"
-unfolding fun_eq_iff by auto
+  unfolding fun_eq_iff by auto
 
 lemma in_rel_Gr: "in_rel (Gr A f) = Grp A f"
-unfolding Gr_def Grp_def fun_eq_iff by auto
+  unfolding Gr_def Grp_def fun_eq_iff by auto
 
 definition relImage where
-"relImage R f \<equiv> {(f a1, f a2) | a1 a2. (a1,a2) \<in> R}"
+  "relImage R f \<equiv> {(f a1, f a2) | a1 a2. (a1,a2) \<in> R}"
 
 definition relInvImage where
-"relInvImage A R f \<equiv> {(a1, a2) | a1 a2. a1 \<in> A \<and> a2 \<in> A \<and> (f a1, f a2) \<in> R}"
+  "relInvImage A R f \<equiv> {(a1, a2) | a1 a2. a1 \<in> A \<and> a2 \<in> A \<and> (f a1, f a2) \<in> R}"
 
 lemma relImage_Gr:
-"\<lbrakk>R \<subseteq> A \<times> A\<rbrakk> \<Longrightarrow> relImage R f = (Gr A f)^-1 O R O Gr A f"
-unfolding relImage_def Gr_def relcomp_def by auto
+  "\<lbrakk>R \<subseteq> A \<times> A\<rbrakk> \<Longrightarrow> relImage R f = (Gr A f)^-1 O R O Gr A f"
+  unfolding relImage_def Gr_def relcomp_def by auto
 
 lemma relInvImage_Gr: "\<lbrakk>R \<subseteq> B \<times> B\<rbrakk> \<Longrightarrow> relInvImage A R f = Gr A f O R O (Gr A f)^-1"
-unfolding Gr_def relcomp_def image_def relInvImage_def by auto
+  unfolding Gr_def relcomp_def image_def relInvImage_def by auto
 
 lemma relImage_mono:
-"R1 \<subseteq> R2 \<Longrightarrow> relImage R1 f \<subseteq> relImage R2 f"
-unfolding relImage_def by auto
+  "R1 \<subseteq> R2 \<Longrightarrow> relImage R1 f \<subseteq> relImage R2 f"
+  unfolding relImage_def by auto
 
 lemma relInvImage_mono:
-"R1 \<subseteq> R2 \<Longrightarrow> relInvImage A R1 f \<subseteq> relInvImage A R2 f"
-unfolding relInvImage_def by auto
+  "R1 \<subseteq> R2 \<Longrightarrow> relInvImage A R1 f \<subseteq> relInvImage A R2 f"
+  unfolding relInvImage_def by auto
 
 lemma relInvImage_Id_on:
-"(\<And>a1 a2. f a1 = f a2 \<longleftrightarrow> a1 = a2) \<Longrightarrow> relInvImage A (Id_on B) f \<subseteq> Id"
-unfolding relInvImage_def Id_on_def by auto
+  "(\<And>a1 a2. f a1 = f a2 \<longleftrightarrow> a1 = a2) \<Longrightarrow> relInvImage A (Id_on B) f \<subseteq> Id"
+  unfolding relInvImage_def Id_on_def by auto
 
 lemma relInvImage_UNIV_relImage:
-"R \<subseteq> relInvImage UNIV (relImage R f) f"
-unfolding relInvImage_def relImage_def by auto
+  "R \<subseteq> relInvImage UNIV (relImage R f) f"
+  unfolding relInvImage_def relImage_def by auto
 
 lemma relImage_proj:
-assumes "equiv A R"
-shows "relImage R (proj R) \<subseteq> Id_on (A//R)"
-unfolding relImage_def Id_on_def
-using proj_iff[OF assms] equiv_class_eq_iff[OF assms]
-by (auto simp: proj_preserves)
+  assumes "equiv A R"
+  shows "relImage R (proj R) \<subseteq> Id_on (A//R)"
+  unfolding relImage_def Id_on_def
+  using proj_iff[OF assms] equiv_class_eq_iff[OF assms]
+  by (auto simp: proj_preserves)
 
 lemma relImage_relInvImage:
-assumes "R \<subseteq> f ` A <*> f ` A"
-shows "relImage (relInvImage A R f) f = R"
-using assms unfolding relImage_def relInvImage_def by fast
+  assumes "R \<subseteq> f ` A <*> f ` A"
+  shows "relImage (relInvImage A R f) f = R"
+  using assms unfolding relImage_def relInvImage_def by fast
 
 lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)"
-by simp
+  by simp
 
 lemma fst_diag_id: "(fst \<circ> (%x. (x, x))) z = id z" by simp
 lemma snd_diag_id: "(snd \<circ> (%x. (x, x))) z = id z" by simp
@@ -159,76 +159,75 @@
 definition shift where "shift lab k = (\<lambda>kl. lab (k # kl))"
 
 lemma empty_Shift: "\<lbrakk>[] \<in> Kl; k \<in> Succ Kl []\<rbrakk> \<Longrightarrow> [] \<in> Shift Kl k"
-unfolding Shift_def Succ_def by simp
+  unfolding Shift_def Succ_def by simp
 
 lemma SuccD: "k \<in> Succ Kl kl \<Longrightarrow> kl @ [k] \<in> Kl"
-unfolding Succ_def by simp
+  unfolding Succ_def by simp
 
 lemmas SuccE = SuccD[elim_format]
 
 lemma SuccI: "kl @ [k] \<in> Kl \<Longrightarrow> k \<in> Succ Kl kl"
-unfolding Succ_def by simp
+  unfolding Succ_def by simp
 
 lemma ShiftD: "kl \<in> Shift Kl k \<Longrightarrow> k # kl \<in> Kl"
-unfolding Shift_def by simp
+  unfolding Shift_def by simp
 
 lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)"
-unfolding Succ_def Shift_def by auto
+  unfolding Succ_def Shift_def by auto
 
 lemma length_Cons: "length (x # xs) = Suc (length xs)"
-by simp
+  by simp
 
 lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)"
-by simp
+  by simp
 
 (*injection into the field of a cardinal*)
 definition "toCard_pred A r f \<equiv> inj_on f A \<and> f ` A \<subseteq> Field r \<and> Card_order r"
 definition "toCard A r \<equiv> SOME f. toCard_pred A r f"
 
 lemma ex_toCard_pred:
-"\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> \<exists> f. toCard_pred A r f"
-unfolding toCard_pred_def
-using card_of_ordLeq[of A "Field r"]
-      ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"]
-by blast
+  "\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> \<exists> f. toCard_pred A r f"
+  unfolding toCard_pred_def
+  using card_of_ordLeq[of A "Field r"]
+    ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"]
+  by blast
 
 lemma toCard_pred_toCard:
   "\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> toCard_pred A r (toCard A r)"
-unfolding toCard_def using someI_ex[OF ex_toCard_pred] .
+  unfolding toCard_def using someI_ex[OF ex_toCard_pred] .
 
-lemma toCard_inj: "\<lbrakk>|A| \<le>o r; Card_order r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow>
-  toCard A r x = toCard A r y \<longleftrightarrow> x = y"
-using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast
+lemma toCard_inj: "\<lbrakk>|A| \<le>o r; Card_order r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> toCard A r x = toCard A r y \<longleftrightarrow> x = y"
+  using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast
 
 definition "fromCard A r k \<equiv> SOME b. b \<in> A \<and> toCard A r b = k"
 
 lemma fromCard_toCard:
-"\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> fromCard A r (toCard A r b) = b"
-unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj)
+  "\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> fromCard A r (toCard A r b) = b"
+  unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj)
 
 lemma Inl_Field_csum: "a \<in> Field r \<Longrightarrow> Inl a \<in> Field (r +c s)"
-unfolding Field_card_of csum_def by auto
+  unfolding Field_card_of csum_def by auto
 
 lemma Inr_Field_csum: "a \<in> Field s \<Longrightarrow> Inr a \<in> Field (r +c s)"
-unfolding Field_card_of csum_def by auto
+  unfolding Field_card_of csum_def by auto
 
 lemma rec_nat_0_imp: "f = rec_nat f1 (%n rec. f2 n rec) \<Longrightarrow> f 0 = f1"
-by auto
+  by auto
 
 lemma rec_nat_Suc_imp: "f = rec_nat f1 (%n rec. f2 n rec) \<Longrightarrow> f (Suc n) = f2 n (f n)"
-by auto
+  by auto
 
 lemma rec_list_Nil_imp: "f = rec_list f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f [] = f1"
-by auto
+  by auto
 
 lemma rec_list_Cons_imp: "f = rec_list f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f (x # xs) = f2 x xs (f xs)"
-by auto
+  by auto
 
 lemma not_arg_cong_Inr: "x \<noteq> y \<Longrightarrow> Inr x \<noteq> Inr y"
-by simp
+  by simp
 
 lemma Collect_splitD: "x \<in> Collect (split A) \<Longrightarrow> A (fst x) (snd x)"
-by auto
+  by auto
 
 definition image2p where
   "image2p f g R = (\<lambda>x y. \<exists>x' y'. R x' y' \<and> f x' = x \<and> g y' = y)"
@@ -250,20 +249,21 @@
 
 lemma equiv_Eps_in:
 "\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> Eps (%x. x \<in> X) \<in> X"
-apply (rule someI2_ex)
-using in_quotient_imp_non_empty by blast
+  apply (rule someI2_ex)
+  using in_quotient_imp_non_empty by blast
 
 lemma equiv_Eps_preserves:
-assumes ECH: "equiv A r" and X: "X \<in> A//r"
-shows "Eps (%x. x \<in> X) \<in> A"
-apply (rule in_mono[rule_format])
- using assms apply (rule in_quotient_imp_subset)
-by (rule equiv_Eps_in) (rule assms)+
+  assumes ECH: "equiv A r" and X: "X \<in> A//r"
+  shows "Eps (%x. x \<in> X) \<in> A"
+  apply (rule in_mono[rule_format])
+   using assms apply (rule in_quotient_imp_subset)
+  by (rule equiv_Eps_in) (rule assms)+
 
 lemma proj_Eps:
-assumes "equiv A r" and "X \<in> A//r"
-shows "proj r (Eps (%x. x \<in> X)) = X"
-unfolding proj_def proof auto
+  assumes "equiv A r" and "X \<in> A//r"
+  shows "proj r (Eps (%x. x \<in> X)) = X"
+unfolding proj_def
+proof auto
   fix x assume x: "x \<in> X"
   thus "(Eps (%x. x \<in> X), x) \<in> r" using assms equiv_Eps_in in_quotient_imp_in_rel by fast
 next
@@ -276,7 +276,7 @@
 lemma univ_commute:
 assumes ECH: "equiv A r" and RES: "f respects r" and x: "x \<in> A"
 shows "(univ f) (proj r x) = f x"
-unfolding univ_def proof -
+proof (unfold univ_def)
   have prj: "proj r x \<in> A//r" using x proj_preserves by fast
   hence "Eps (%y. y \<in> proj r x) \<in> A" using ECH equiv_Eps_preserves by fast
   moreover have "proj r (Eps (%y. y \<in> proj r x)) = proj r x" using ECH prj proj_Eps by fast
@@ -285,13 +285,12 @@
 qed
 
 lemma univ_preserves:
-assumes ECH: "equiv A r" and RES: "f respects r" and
-        PRES: "\<forall> x \<in> A. f x \<in> B"
-shows "\<forall>X \<in> A//r. univ f X \<in> B"
+  assumes ECH: "equiv A r" and RES: "f respects r" and PRES: "\<forall>x \<in> A. f x \<in> B"
+  shows "\<forall>X \<in> A//r. univ f X \<in> B"
 proof
   fix X assume "X \<in> A//r"
   then obtain x where x: "x \<in> A" and X: "X = proj r x" using ECH proj_image[of r A] by blast
-  hence "univ f X = f x" using assms univ_commute by fastforce
+  hence "univ f X = f x" using ECH RES univ_commute by fastforce
   thus "univ f X \<in> B" using x PRES by simp
 qed
 
--- a/src/HOL/BNF_LFP.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/BNF_LFP.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1,9 +1,8 @@
-
 (*  Title:      HOL/BNF_LFP.thy
     Author:     Dmitriy Traytel, TU Muenchen
     Author:     Lorenz Panny, TU Muenchen
     Author:     Jasmin Blanchette, TU Muenchen
-    Copyright   2012, 2013
+    Copyright   2012, 2013, 2014
 
 Least fixed point operation on bounded natural functors.
 *)
@@ -18,37 +17,37 @@
 begin
 
 lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
-by blast
+  by blast
 
 lemma image_Collect_subsetI: "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
-by blast
+  by blast
 
 lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
-by auto
+  by auto
 
 lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x"
-by auto
+  by auto
 
 lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> underS R j"
-unfolding underS_def by simp
+  unfolding underS_def by simp
 
 lemma underS_E: "i \<in> underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R"
-unfolding underS_def by simp
+  unfolding underS_def by simp
 
 lemma underS_Field: "i \<in> underS R j \<Longrightarrow> i \<in> Field R"
-unfolding underS_def Field_def by auto
+  unfolding underS_def Field_def by auto
 
 lemma FieldI2: "(i, j) \<in> R \<Longrightarrow> j \<in> Field R"
-unfolding Field_def by auto
+  unfolding Field_def by auto
 
 lemma fst_convol': "fst (\<langle>f, g\<rangle> x) = f x"
-using fst_convol unfolding convol_def by simp
+  using fst_convol unfolding convol_def by simp
 
 lemma snd_convol': "snd (\<langle>f, g\<rangle> x) = g x"
-using snd_convol unfolding convol_def by simp
+  using snd_convol unfolding convol_def by simp
 
 lemma convol_expand_snd: "fst o f = g \<Longrightarrow> \<langle>g, snd o f\<rangle> = f"
-unfolding convol_def by auto
+  unfolding convol_def by auto
 
 lemma convol_expand_snd':
   assumes "(fst o f = g)"
@@ -60,11 +59,12 @@
   moreover have "\<dots> \<longleftrightarrow> \<langle>g, h\<rangle> = f" by (subst (2) *[symmetric]) (auto simp: convol_def fun_eq_iff)
   ultimately show ?thesis by simp
 qed
+
 lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B"
-unfolding bij_betw_def by auto
+  unfolding bij_betw_def by auto
 
 lemma bij_betw_imageE: "bij_betw f A B \<Longrightarrow> f ` A = B"
-unfolding bij_betw_def by auto
+  unfolding bij_betw_def by auto
 
 lemma f_the_inv_into_f_bij_betw: "bij_betw f A B \<Longrightarrow>
   (bij_betw f A B \<Longrightarrow> x \<in> B) \<Longrightarrow> f (the_inv_into A f x) = x"
@@ -78,7 +78,7 @@
   "\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y);
     \<And>x. x \<in> X \<Longrightarrow> f x \<in> Y;
     \<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y"
-unfolding bij_betw_def inj_on_def by blast
+  unfolding bij_betw_def inj_on_def by blast
 
 lemma surj_fun_eq:
   assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x"
--- a/src/HOL/Bali/Basis.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Bali/Basis.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -13,7 +13,7 @@
 
 declare split_if_asm  [split] option.split [split] option.split_asm [split]
 setup {* map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac)) *}
-declare if_weak_cong [cong del] option.weak_case_cong [cong del]
+declare if_weak_cong [cong del] option.case_cong_weak [cong del]
 declare length_Suc_conv [iff]
 
 lemma Collect_split_eq: "{p. P (split f p)} = {(a,b). P (f a b)}"
--- a/src/HOL/Bali/Conform.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Bali/Conform.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -521,7 +521,7 @@
 apply auto
 apply (simp only: obj_ty_cong) 
 apply (force dest: conforms_globsD intro!: lconf_upd 
-       simp add: oconf_def cong del: sum.weak_case_cong)
+       simp add: oconf_def cong del: sum.case_cong_weak)
 done
 
 lemma conforms_set_locals: 
--- a/src/HOL/Code_Evaluation.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Code_Evaluation.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -128,6 +128,19 @@
   constant "term_of \<Colon> integer \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_number/ HOLogic.code'_integerT"
 | constant "term_of \<Colon> String.literal \<Rightarrow> term" \<rightharpoonup> (Eval) "HOLogic.mk'_literal"
 
+declare [[code drop: "term_of :: integer \<Rightarrow> _"]]
+
+lemma term_of_integer [unfolded typerep_fun_def typerep_num_def typerep_integer_def, code]:
+  "term_of (i :: integer) =
+  (if i > 0 then 
+     App (Const (STR ''Num.numeral_class.numeral'') (TYPEREP(num \<Rightarrow> integer)))
+      (term_of (num_of_integer i))
+   else if i = 0 then Const (STR ''Groups.zero_class.zero'') TYPEREP(integer)
+   else
+     App (Const (STR ''Groups.uminus_class.uminus'') TYPEREP(integer \<Rightarrow> integer))
+       (term_of (- i)))"
+by(rule term_of_anything[THEN meta_eq_to_obj_eq])
+
 code_reserved Eval HOLogic
 
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,148 @@
+(*  Title:      Code_Test.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test infrastructure for the code generator
+*)
+
+theory Code_Test
+imports Main
+keywords "test_code" "eval_term" :: diag
+begin
+
+subsection {* YXML encoding for @{typ Code_Evaluation.term} *}
+
+datatype yxml_of_term = YXML
+
+lemma yot_anything: "x = (y :: yxml_of_term)"
+by(cases x y rule: yxml_of_term.exhaust[case_product yxml_of_term.exhaust])(simp)
+
+definition yot_empty :: yxml_of_term where [code del]: "yot_empty = YXML"
+definition yot_literal :: "String.literal \<Rightarrow> yxml_of_term"
+  where [code del]: "yot_literal _ = YXML"
+definition yot_append :: "yxml_of_term \<Rightarrow> yxml_of_term \<Rightarrow> yxml_of_term"
+  where [code del]: "yot_append _ _ = YXML"
+definition yot_concat :: "yxml_of_term list \<Rightarrow> yxml_of_term"
+  where [code del]: "yot_concat _ = YXML"
+
+text {* Serialise @{typ yxml_of_term} to native string of target language *}
+
+code_printing type_constructor yxml_of_term
+  \<rightharpoonup>  (SML) "string"
+  and (OCaml) "string"
+  and (Haskell) "String"
+  and (Scala) "String"
+| constant yot_empty
+  \<rightharpoonup>  (SML) "\"\""
+  and (OCaml) "\"\""
+  and (Haskell) "\"\""
+  and (Scala) "\"\""
+| constant yot_literal
+  \<rightharpoonup>  (SML) "_"
+  and (OCaml) "_"
+  and (Haskell) "_"
+  and (Scala) "_"
+| constant yot_append
+  \<rightharpoonup> (SML) "String.concat [(_), (_)]"
+  and (OCaml) "String.concat \"\" [(_); (_)]"
+  and (Haskell) infixr 5 "++"
+  and (Scala) infixl 5 "+"
+| constant yot_concat
+  \<rightharpoonup> (SML) "String.concat"
+  and (OCaml) "String.concat \"\""
+  and (Haskell) "Prelude.concat"
+  and (Scala) "_.mkString(\"\")"
+
+text {*
+  Stripped-down implementations of Isabelle's XML tree with YXML encoding
+  as defined in @{file "~~/src/Pure/PIDE/xml.ML"}, @{file "~~/src/Pure/PIDE/yxml.ML"}
+  sufficient to encode @{typ "Code_Evaluation.term"} as in @{file "~~/src/Pure/term_xml.ML"}.
+*}
+
+datatype xml_tree = XML_Tree
+
+lemma xml_tree_anything: "x = (y :: xml_tree)"
+by(cases x y rule: xml_tree.exhaust[case_product xml_tree.exhaust])(simp)
+
+context begin
+local_setup {* Local_Theory.map_naming (Name_Space.mandatory_path "xml") *}
+
+type_synonym attributes = "(String.literal \<times> String.literal) list"
+type_synonym body = "xml_tree list"
+
+definition Elem :: "String.literal \<Rightarrow> attributes \<Rightarrow> xml_tree list \<Rightarrow> xml_tree"
+where [code del]: "Elem _ _ _ = XML_Tree"
+
+definition Text :: "String.literal \<Rightarrow> xml_tree"
+where [code del]: "Text _ = XML_Tree"
+
+definition node :: "xml_tree list \<Rightarrow> xml_tree"
+where "node ts = Elem (STR '':'') [] ts"
+
+definition tagged :: "String.literal \<Rightarrow> String.literal option \<Rightarrow> xml_tree list \<Rightarrow> xml_tree"
+where "tagged tag x ts = Elem tag (case x of None \<Rightarrow> [] | Some x' \<Rightarrow> [(STR ''0'', x')]) ts"
+
+definition list where "list f xs = map (node \<circ> f) xs"
+
+definition X :: yxml_of_term where "X = yot_literal (STR [Char Nibble0 Nibble5])"
+definition Y :: yxml_of_term where "Y = yot_literal (STR [Char Nibble0 Nibble6])"
+definition XY :: yxml_of_term where "XY = yot_append X Y"
+definition XYX :: yxml_of_term where "XYX = yot_append XY X"
+
+end
+
+code_datatype xml.Elem xml.Text
+
+definition yxml_string_of_xml_tree :: "xml_tree \<Rightarrow> yxml_of_term \<Rightarrow> yxml_of_term"
+where [code del]: "yxml_string_of_xml_tree _ _ = YXML"
+
+lemma yxml_string_of_xml_tree_code [code]:
+  "yxml_string_of_xml_tree (xml.Elem name atts ts) rest =
+   yot_append xml.XY (
+   yot_append (yot_literal name) (
+   foldr (\<lambda>(a, x) rest. 
+     yot_append xml.Y (
+     yot_append (yot_literal a) (
+     yot_append (yot_literal (STR ''='')) (
+     yot_append (yot_literal x) rest)))) atts (
+   foldr yxml_string_of_xml_tree ts (
+   yot_append xml.XYX rest))))"
+  "yxml_string_of_xml_tree (xml.Text s) rest = yot_append (yot_literal s) rest"
+by(rule yot_anything)+
+
+definition yxml_string_of_body :: "xml.body \<Rightarrow> yxml_of_term"
+where "yxml_string_of_body ts = foldr yxml_string_of_xml_tree ts yot_empty"
+
+text {*
+  Encoding @{typ Code_Evaluation.term} into XML trees
+  as defined in @{file "~~/src/Pure/term_xml.ML"}
+*}
+
+definition xml_of_typ :: "Typerep.typerep \<Rightarrow> xml.body"
+where [code del]: "xml_of_typ _ = [XML_Tree]"
+
+definition xml_of_term :: "Code_Evaluation.term \<Rightarrow> xml.body"
+where [code del]: "xml_of_term _ = [XML_Tree]"
+
+lemma xml_of_typ_code [code]:
+  "xml_of_typ (typerep.Typerep t args) = [xml.tagged (STR ''0'') (Some t) (xml.list xml_of_typ args)]"
+by(simp add: xml_of_typ_def xml_tree_anything)
+
+lemma xml_of_term_code [code]:
+  "xml_of_term (Code_Evaluation.Const x ty) = [xml.tagged (STR ''0'') (Some x) (xml_of_typ ty)]"
+  "xml_of_term (Code_Evaluation.App t1 t2)  = [xml.tagged (STR ''5'') None [xml.node (xml_of_term t1), xml.node (xml_of_term t2)]]"
+  "xml_of_term (Code_Evaluation.Abs x ty t) = [xml.tagged (STR ''4'') (Some x) [xml.node (xml_of_typ ty), xml.node (xml_of_term t)]]"
+  -- {*
+    FIXME: @{const Code_Evaluation.Free} is used only in Quickcheck_Narrowing to represent
+    uninstantiated parameters in constructors. Here, we always translate them to @{ML Free} variables.
+  *}
+  "xml_of_term (Code_Evaluation.Free x ty)  = [xml.tagged (STR ''1'') (Some x) (xml_of_typ ty)]"
+by(simp_all add: xml_of_term_def xml_tree_anything)
+
+definition yxml_string_of_term :: "Code_Evaluation.term \<Rightarrow> yxml_of_term"
+where "yxml_string_of_term = yxml_string_of_body \<circ> xml_of_term"
+
+subsection {* Test engine and drivers *}
+
+ML_file "code_test.ML"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test_GHC.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,15 @@
+(*  Title:      Code_Test_GHC.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test case for test_code on GHC
+*)
+
+theory Code_Test_GHC imports Code_Test begin
+
+definition id_integer :: "integer \<Rightarrow> integer" where "id_integer = id"
+
+test_code "id_integer (14 + 7 * -12) = 140 div -2" in GHC
+
+eval_term "14 + 7 * -12 :: integer" in GHC
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test_MLton.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,13 @@
+(*  Title:      Code_Test_MLtonL.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test case for test_code on MLton
+*)
+
+theory Code_Test_MLton imports Code_Test begin
+
+test_code "14 + 7 * -12 = (140 div -2 :: integer)" in MLton
+
+eval_term "14 + 7 * -12 :: integer" in MLton
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test_OCaml.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,13 @@
+(*  Title:      Code_Test_OCaml.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test case for test_code on OCaml
+*)
+
+theory Code_Test_OCaml imports Code_Test begin
+
+test_code "14 + 7 * -12 = (140 div -2 :: integer)" in OCaml
+
+eval_term "14 + 7 * -12 :: integer" in OCaml
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test_PolyML.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,13 @@
+(*  Title:      Code_Test_PolyML.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test case for test_code on PolyML
+*)
+
+theory Code_Test_PolyML imports Code_Test begin
+
+test_code "14 + 7 * -12 = (140 div -2 :: integer)" in PolyML
+
+eval_term "14 + 7 * -12 :: integer" in PolyML
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test_SMLNJ.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,13 @@
+(*  Title:      Code_Test_SMLNJ.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test case for test_code on SMLNJ
+*)
+
+theory Code_Test_SMLNJ imports Code_Test begin
+
+test_code "14 + 7 * -12 = (140 div -2 :: integer)" in SMLNJ
+
+eval_term "14 + 7 * -12 :: integer" in SMLNJ
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/Code_Test_Scala.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,13 @@
+(*  Title:      Code_Test_Scala.thy
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test case for test_code on Scala
+*)
+
+theory Code_Test_Scala imports Code_Test begin 
+
+test_code "14 + 7 * -12 = (140 div -2 :: integer)" in Scala
+
+eval_term "14 + 7 * -12 :: integer" in Scala
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codegenerator_Test/code_test.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,594 @@
+(*  Title:      Code_Test.ML
+    Author:     Andreas Lochbihler, ETH Zurich
+
+Test infrastructure for the code generator
+*)
+
+signature CODE_TEST = sig
+  val add_driver : string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) -> theory -> theory
+  val get_driver : theory -> string -> ((Proof.context -> (string * string) list * string -> Path.T -> string) * string) option
+  val overlord : bool Config.T
+  val successN : string
+  val failureN : string
+  val start_markerN : string
+  val end_markerN : string
+  val test_terms : Proof.context -> term list -> string -> unit
+  val test_targets : Proof.context -> term list -> string list -> unit list
+  val test_code_cmd : string list -> string list -> Toplevel.state -> unit
+
+  val eval_term : Proof.context -> term -> string -> unit
+
+  val gen_driver :
+   (theory -> Path.T -> string list -> string ->
+    {files : (Path.T * string) list,
+     compile_cmd : string option, run_cmd : string, mk_code_file : string -> Path.T})
+   -> string -> string -> string
+   -> theory -> (string * string) list * string -> Path.T -> string
+
+  val ISABELLE_POLYML_PATH : string
+  val polymlN : string
+  val evaluate_in_polyml : Proof.context -> (string * string) list * string -> Path.T -> string
+
+  val mltonN : string
+  val ISABELLE_MLTON : string
+  val evaluate_in_mlton : Proof.context -> (string * string) list * string -> Path.T -> string
+
+  val smlnjN : string
+  val ISABELLE_SMLNJ : string
+  val evaluate_in_smlnj : Proof.context -> (string * string) list * string -> Path.T -> string
+
+  val ocamlN : string
+  val ISABELLE_OCAMLC : string
+  val evaluate_in_ocaml : Proof.context -> (string * string) list * string -> Path.T -> string
+
+  val ghcN : string
+  val ISABELLE_GHC : string
+  val ghc_options : string Config.T
+  val evaluate_in_ghc : Proof.context -> (string * string) list * string -> Path.T -> string
+
+  val scalaN : string
+  val ISABELLE_SCALA : string
+  val evaluate_in_scala : Proof.context -> (string * string) list * string -> Path.T -> string
+end
+
+structure Code_Test : CODE_TEST = struct
+
+(* convert a list of terms into nested tuples and back *)
+fun mk_tuples [] = @{term "()"}
+  | mk_tuples [t] = t
+  | mk_tuples (t :: ts) = HOLogic.mk_prod (t, mk_tuples ts)
+
+fun dest_tuples (Const (@{const_name Pair}, _) $ l $ r) = l :: dest_tuples r
+  | dest_tuples t = [t]
+
+
+fun map_option _ NONE = NONE
+  | map_option f (SOME x) = SOME (f x)
+
+fun last_field sep str =
+  let
+    val n = size sep;
+    val len = size str;
+    fun find i =
+      if i < 0 then NONE
+      else if String.substring (str, i, n) = sep then SOME i
+      else find (i - 1);
+  in
+    (case find (len - n) of
+      NONE => NONE
+    | SOME i => SOME (String.substring (str, 0, i), String.extract (str, i + n, NONE)))
+  end;
+
+fun split_first_last start stop s =
+  case first_field start s
+   of NONE => NONE
+    | SOME (initial, rest) =>
+      case last_field stop rest
+       of NONE => NONE
+        | SOME (middle, tail) => SOME (initial, middle, tail);
+
+(* Data slot for drivers *)
+
+structure Drivers = Theory_Data
+(
+  type T = (string * ((Proof.context -> (string * string) list * string -> Path.T -> string) * string)) list;
+  val empty = [];
+  val extend = I;
+  fun merge data : T = AList.merge (op =) (K true) data;
+)
+
+val add_driver = Drivers.map o AList.update (op =);
+val get_driver = AList.lookup (op =) o Drivers.get;
+
+(*
+  Test drivers must produce output of the following format:
+  
+  The start of the relevant data is marked with start_markerN,
+  its end with end_markerN.
+
+  Between these two markers, every line corresponds to one test.
+  Lines of successful tests start with successN, failures start with failureN.
+  The failure failureN may continue with the YXML encoding of the evaluated term.
+  There must not be any additional whitespace in between.
+*)
+
+(* Parsing of results *)
+
+val successN = "True"
+val failureN = "False"
+val start_markerN = "*@*Isabelle/Code_Test-start*@*"
+val end_markerN = "*@*Isabelle/Code_Test-end*@*"
+
+fun parse_line line =
+  if String.isPrefix successN line then (true, NONE)
+  else if String.isPrefix failureN line then (false, 
+    if size line > size failureN then
+      String.extract (line, size failureN, NONE)
+      |> YXML.parse_body
+      |> Term_XML.Decode.term
+      |> dest_tuples
+      |> SOME
+    else NONE)
+  else raise Fail ("Cannot parse result of evaluation:\n" ^ line)
+
+fun parse_result target out =
+  case split_first_last start_markerN end_markerN out
+    of NONE => error ("Evaluation failed for " ^ target ^ "!\nBash output:\n" ^ out)
+     | SOME (_, middle, _) => middle |> trim_line |> split_lines |> map parse_line
+
+(* Pretty printing of test results *)
+
+fun pretty_eval _ NONE _ = []
+  | pretty_eval ctxt (SOME evals) ts = 
+    [Pretty.fbrk,
+     Pretty.big_list "Evaluated terms"
+       (map (fn (t, eval) => Pretty.block 
+         [Syntax.pretty_term ctxt t, Pretty.brk 1, Pretty.str "=", Pretty.brk 1,
+          Syntax.pretty_term ctxt eval])
+       (ts ~~ evals))]
+
+fun pretty_failure ctxt target (((_, evals), query), eval_ts) =
+  Pretty.block (Pretty.text ("Test in " ^ target ^ " failed for") @ [Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt query)]
+    @ pretty_eval ctxt evals eval_ts)
+
+fun pretty_failures ctxt target failures =
+  Pretty.blk (0, Pretty.fbreaks (map (pretty_failure ctxt target) failures))
+
+(* Driver invocation *)
+
+val overlord = Attrib.setup_config_bool @{binding "code_test_overlord"} (K false);
+
+fun with_overlord_dir name f =
+  let
+    val path = Path.append (Path.explode "$ISABELLE_HOME_USER") (Path.basic (name ^ serial_string ()))
+    val _ = Isabelle_System.mkdirs path;
+  in
+    Exn.release (Exn.capture f path)
+  end;
+
+fun dynamic_value_strict ctxt t compiler =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    val (driver, target) = case get_driver thy compiler
+     of NONE => error ("No driver for target " ^ compiler)
+      | SOME f => f;
+    val debug = Config.get (Proof_Context.init_global thy) overlord
+    val with_dir = if debug then with_overlord_dir else Isabelle_System.with_tmp_dir
+    fun evaluate f = with_dir "Code_Test" (driver ctxt f) |> parse_result compiler
+    fun evaluator program _ vs_ty deps =
+      Exn.interruptible_capture evaluate (Code_Target.evaluator ctxt target program deps true vs_ty);
+    fun postproc f = map (apsnd (map_option (map f)))
+  in
+    Exn.release (Code_Thingol.dynamic_value ctxt (Exn.map_result o postproc) evaluator t)
+  end;
+
+(* Term preprocessing *)
+
+fun add_eval (Const (@{const_name Trueprop}, _) $ t) = add_eval t
+  | add_eval (Const (@{const_name "HOL.eq"}, _) $ lhs $ rhs) = (fn acc =>
+    acc
+    |> add_eval rhs
+    |> add_eval lhs
+    |> cons rhs
+    |> cons lhs)
+  | add_eval (Const (@{const_name "Not"}, _) $ t) = add_eval t
+  | add_eval (Const (@{const_name "Orderings.ord_class.less_eq"}, _) $ lhs $ rhs) = (fn acc =>
+    lhs :: rhs :: acc)
+  | add_eval (Const (@{const_name "Orderings.ord_class.less"}, _) $ lhs $ rhs) = (fn acc =>
+    lhs :: rhs :: acc)
+  | add_eval _ = I
+
+fun mk_term_of [] = @{term "None :: (unit \<Rightarrow> yxml_of_term) option"}
+  | mk_term_of ts =
+  let
+    val tuple = mk_tuples ts
+    val T = fastype_of tuple
+  in
+    @{term "Some :: (unit \<Rightarrow> yxml_of_term) \<Rightarrow> (unit \<Rightarrow> yxml_of_term) option"} $
+      (absdummy @{typ unit} (@{const yxml_string_of_term} $
+        (Const (@{const_name Code_Evaluation.term_of}, T --> @{typ term}) $ tuple)))
+  end
+
+fun test_terms ctxt ts target =
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    fun term_of t = Sign.of_sort thy (fastype_of t, @{sort term_of})
+
+    fun ensure_bool t = case fastype_of t of @{typ bool} => ()
+      | _ => error ("Test case not of type bool: " ^ Pretty.string_of (Syntax.pretty_term ctxt t))
+
+    val _ = map ensure_bool ts
+
+    val evals = map (fn t => filter term_of (add_eval t [])) ts
+    val eval = map mk_term_of evals
+
+    val T = HOLogic.mk_prodT (@{typ bool}, Type (@{type_name option}, [@{typ unit} --> @{typ yxml_of_term}]))
+    val t = HOLogic.mk_list T (map HOLogic.mk_prod (ts ~~ eval))
+
+    val result = dynamic_value_strict ctxt t target;
+
+    val failed =
+      filter_out (fst o fst o fst) (result ~~ ts ~~ evals)
+      handle ListPair.UnequalLengths => 
+        error ("Evaluation failed!\nWrong number of test results: " ^ Int.toString (length result))
+    val _ = case failed of [] => () 
+      | _ => error (Pretty.string_of (pretty_failures ctxt target failed))
+  in
+    ()
+  end
+
+fun test_targets ctxt = map o test_terms ctxt
+
+fun test_code_cmd raw_ts targets state =
+  let
+    val ctxt = Toplevel.context_of state;
+    val ts = Syntax.read_terms ctxt raw_ts;
+    val frees = fold Term.add_free_names ts []
+    val _ = if frees = [] then () else
+      error ("Terms contain free variables: " ^
+      Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees))))
+  in
+    test_targets ctxt ts targets; ()
+  end
+
+
+fun eval_term ctxt t target =
+  let
+    val thy = Proof_Context.theory_of ctxt
+
+    val T_t = fastype_of t
+    val _ = if Sign.of_sort thy (T_t, @{sort term_of}) then () else error 
+      ("Type " ^ Pretty.string_of (Syntax.pretty_typ ctxt T_t) ^ 
+       " of term not of sort " ^ Pretty.string_of (Syntax.pretty_sort ctxt @{sort term_of}))
+
+    val T = HOLogic.mk_prodT (@{typ bool}, Type (@{type_name option}, [@{typ unit} --> @{typ yxml_of_term}]))
+    val t' = HOLogic.mk_list T [HOLogic.mk_prod (@{term "False"}, mk_term_of [t])]
+
+    val result = dynamic_value_strict ctxt t' target;
+    val t_eval = case result of [(_, SOME [t])] => t | _ => error "Evaluation failed"
+  in
+    Pretty.writeln (Syntax.pretty_term ctxt t_eval)
+  end
+
+fun eval_term_cmd raw_t target state =
+  let
+    val ctxt = Toplevel.context_of state;
+    val t = Syntax.read_term ctxt raw_t;
+    val frees = Term.add_free_names t []
+    val _ = if frees = [] then () else
+      error ("Term contains free variables: " ^
+      Pretty.string_of (Pretty.block (Pretty.commas (map Pretty.str frees))))
+  in
+    eval_term ctxt t target
+  end
+
+
+(* Generic driver *)
+
+fun gen_driver mk_driver env_var env_var_dest compilerN ctxt (code_files, value_name) =
+  let
+    val compiler = getenv env_var
+    val _ = if compiler <> "" then () else error (Pretty.string_of (Pretty.para 
+         ("Environment variable " ^ env_var ^ " is not set. To test code generation with " ^
+         compilerN ^ ", set this variable to your " ^ env_var_dest ^ " in the settings file.")))
+
+    fun compile NONE = ()
+      | compile (SOME cmd) =
+        let
+          val (out, ret) = Isabelle_System.bash_output cmd
+        in
+          if ret = 0 then () else error
+            ("Compilation with " ^ compilerN ^ " failed:\n" ^ cmd ^ "\n" ^ out)
+        end
+
+    fun run (path : Path.T)= 
+      let
+        val modules = map fst code_files
+        val {files, compile_cmd, run_cmd, mk_code_file}
+          =  mk_driver ctxt path modules value_name
+
+        val _ = map (fn (name, code) => File.write (mk_code_file name) code) code_files
+        val _ = map (fn (name, content) => File.write name content) files
+
+        val _ = compile compile_cmd
+
+        val (out, res) = Isabelle_System.bash_output run_cmd
+        val _ = if res = 0 then () else error
+          ("Evaluation for " ^ compilerN ^ " terminated with error code " ^ Int.toString res ^
+           "\nBash output:\n" ^ out)
+      in
+        out
+      end
+  in
+    run
+  end
+
+(* Driver for PolyML *)
+
+val ISABELLE_POLYML_PATH = "ISABELLE_POLYML_PATH"
+val polymlN = "PolyML";
+
+fun mk_driver_polyml _ path _ value_name =
+  let
+    val generatedN = "generated.sml"
+    val driverN = "driver.sml"
+
+    val code_path = Path.append path (Path.basic generatedN)
+    val driver_path = Path.append path (Path.basic driverN)
+    val driver = 
+      "fun main prog_name = \n" ^
+      "  let\n" ^
+      "    fun format_term NONE = \"\"\n" ^ 
+      "      | format_term (SOME t) = t ();\n" ^
+      "    fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
+      "      | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
+      "    val result = " ^ value_name ^ " ();\n" ^
+      "    val _ = print \"" ^ start_markerN ^ "\";\n" ^
+      "    val _ = map (print o format) result;\n" ^
+      "    val _ = print \"" ^ end_markerN ^ "\";\n" ^
+      "  in\n" ^
+      "    ()\n" ^
+      "  end;\n"
+    val cmd =
+      "echo \"use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^ 
+      Path.implode driver_path ^ "\\\"; main ();\" | " ^ 
+      Path.implode (Path.variable ISABELLE_POLYML_PATH)
+  in
+    {files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path}
+  end
+
+val evaluate_in_polyml = gen_driver mk_driver_polyml ISABELLE_POLYML_PATH "PolyML executable" polymlN
+
+(* Driver for mlton *)
+
+val mltonN = "MLton"
+val ISABELLE_MLTON = "ISABELLE_MLTON"
+
+fun mk_driver_mlton _ path _ value_name =
+  let
+    val generatedN = "generated.sml"
+    val driverN = "driver.sml"
+    val projectN = "test"
+    val ml_basisN = projectN ^ ".mlb"
+
+    val code_path = Path.append path (Path.basic generatedN)
+    val driver_path = Path.append path (Path.basic driverN)
+    val ml_basis_path = Path.append path (Path.basic ml_basisN)
+    val driver = 
+      "fun format_term NONE = \"\"\n" ^ 
+      "  | format_term (SOME t) = t ();\n" ^
+      "fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
+      "  | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
+      "val result = " ^ value_name ^ " ();\n" ^
+      "val _ = print \"" ^ start_markerN ^ "\";\n" ^
+      "val _ = map (print o format) result;\n" ^
+      "val _ = print \"" ^ end_markerN ^ "\";\n"
+    val ml_basis =
+      "$(SML_LIB)/basis/basis.mlb\n" ^
+      generatedN ^ "\n" ^
+      driverN
+
+    val compile_cmd =
+      File.shell_path (Path.variable ISABELLE_MLTON) ^
+      " -default-type intinf " ^ File.shell_path ml_basis_path
+    val run_cmd = File.shell_path (Path.append path (Path.basic projectN))
+  in
+    {files = [(driver_path, driver), (ml_basis_path, ml_basis)],
+     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
+  end
+
+val evaluate_in_mlton = gen_driver mk_driver_mlton ISABELLE_MLTON "MLton executable" mltonN
+
+(* Driver for SML/NJ *)
+
+val smlnjN = "SMLNJ"
+val ISABELLE_SMLNJ = "ISABELLE_SMLNJ"
+
+fun mk_driver_smlnj _ path _ value_name =
+  let
+    val generatedN = "generated.sml"
+    val driverN = "driver.sml"
+
+    val code_path = Path.append path (Path.basic generatedN)
+    val driver_path = Path.append path (Path.basic driverN)
+    val driver = 
+      "structure Test = struct\n" ^
+      "fun main prog_name =\n" ^
+      "  let\n" ^
+      "    fun format_term NONE = \"\"\n" ^ 
+      "      | format_term (SOME t) = t ();\n" ^
+      "    fun format (true, _) = \"" ^ successN ^ "\\n\"\n" ^
+      "      | format (false, to) = \"" ^ failureN ^ "\" ^ format_term to ^ \"\\n\";\n" ^
+      "    val result = " ^ value_name ^ " ();\n" ^
+      "    val _ = print \"" ^ start_markerN ^ "\";\n" ^
+      "    val _ = map (print o format) result;\n" ^
+      "    val _ = print \"" ^ end_markerN ^ "\";\n" ^
+      "  in\n" ^
+      "    0\n" ^
+      "  end;\n" ^
+      "end;"
+    val cmd =
+      "echo \"Control.MC.matchRedundantError := false; Control.MC.matchRedundantWarn := false;" ^
+      "use \\\"" ^ Path.implode code_path ^ "\\\"; use \\\"" ^ Path.implode driver_path ^ "\\\";" ^
+      "Test.main ();\" | " ^ Path.implode (Path.variable ISABELLE_SMLNJ)
+  in
+    {files = [(driver_path, driver)], compile_cmd = NONE, run_cmd = cmd, mk_code_file = K code_path}
+  end
+
+val evaluate_in_smlnj = gen_driver mk_driver_smlnj ISABELLE_SMLNJ "SMLNJ executable" smlnjN
+
+(* Driver for OCaml *)
+
+val ocamlN = "OCaml"
+val ISABELLE_OCAMLC = "ISABELLE_OCAMLC"
+
+fun mk_driver_ocaml _ path _ value_name =
+  let
+    val generatedN = "generated.ml"
+    val driverN = "driver.ml"
+
+    val code_path = Path.append path (Path.basic generatedN)
+    val driver_path = Path.append path (Path.basic driverN)
+    val driver = 
+      "let format_term = function\n" ^
+      "  | None -> \"\"\n" ^ 
+      "  | Some t -> t ();;\n" ^
+      "let format = function\n" ^
+      "  | (true, _) -> \"" ^ successN ^ "\\n\"\n" ^
+      "  | (false, x) -> \"" ^ failureN ^ "\" ^ format_term x ^ \"\\n\";;\n" ^
+      "let result = " ^ ("Generated." ^ value_name) ^ " ();;\n" ^
+      "let main x =\n" ^
+      "  let _ = print_string \"" ^ start_markerN ^ "\" in\n" ^
+      "  let _ = List.map (fun x -> print_string (format x)) result in\n" ^
+      "  print_string \"" ^ end_markerN ^ "\";;\n" ^
+      "main ();;"
+
+    val compiled_path = Path.append path (Path.basic "test")
+    val compile_cmd =
+      Path.implode (Path.variable ISABELLE_OCAMLC) ^ " -w pu -o " ^ Path.implode compiled_path ^
+      " -I " ^ Path.implode path ^
+      " nums.cma " ^ Path.implode code_path ^ " " ^ Path.implode driver_path
+
+    val run_cmd = File.shell_path compiled_path
+  in
+    {files = [(driver_path, driver)],
+     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
+  end
+
+val evaluate_in_ocaml = gen_driver mk_driver_ocaml ISABELLE_OCAMLC "ocamlc executable" ocamlN
+
+(* Driver for GHC *)
+
+val ghcN = "GHC"
+val ISABELLE_GHC = "ISABELLE_GHC"
+
+val ghc_options = Attrib.setup_config_string @{binding code_test_ghc} (K "")
+
+fun mk_driver_ghc ctxt path modules value_name =
+  let
+    val driverN = "Main.hs"
+
+    fun mk_code_file name = Path.append path (Path.basic (name ^ ".hs"))
+    val driver_path = Path.append path (Path.basic driverN)
+    val driver = 
+      "module Main where {\n" ^
+      String.concat (map (fn module => "import qualified " ^ module ^ ";\n") modules) ^
+      "main = do {\n" ^
+      "    let {\n" ^
+      "      format_term Nothing = \"\";\n" ^ 
+      "      format_term (Just t) = t ();\n" ^
+      "      format (True, _) = \"" ^ successN ^ "\\n\";\n" ^
+      "      format (False, to) = \"" ^ failureN ^ "\" ++ format_term to ++ \"\\n\";\n" ^
+      "      result = " ^ value_name ^ " ();\n" ^
+      "    };\n" ^
+      "    Prelude.putStr \"" ^ start_markerN ^ "\";\n" ^
+      "    Prelude.mapM_ (putStr . format) result;\n" ^
+      "    Prelude.putStr \"" ^ end_markerN ^ "\";\n" ^
+      "  }\n" ^
+      "}\n"
+
+    val compiled_path = Path.append path (Path.basic "test")
+    val compile_cmd =
+      Path.implode (Path.variable ISABELLE_GHC) ^ " " ^ Code_Haskell.language_params ^ " " ^
+      Config.get ctxt ghc_options ^ " -o " ^ Path.implode compiled_path ^ " " ^
+      Path.implode driver_path ^ " -i" ^ Path.implode path
+
+    val run_cmd = File.shell_path compiled_path
+  in
+    {files = [(driver_path, driver)],
+     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = mk_code_file}
+  end
+
+val evaluate_in_ghc = gen_driver mk_driver_ghc ISABELLE_GHC "GHC executable" ghcN
+
+(* Driver for Scala *)
+
+val scalaN = "Scala"
+val ISABELLE_SCALA = "ISABELLE_SCALA"
+
+fun mk_driver_scala _ path _ value_name =
+  let
+    val generatedN = "Generated_Code"
+    val driverN = "Driver.scala"
+
+    val code_path = Path.append path (Path.basic (generatedN ^ ".scala"))
+    val driver_path = Path.append path (Path.basic driverN)
+    val driver = 
+      "import " ^ generatedN ^ "._\n" ^
+      "object Test {\n" ^
+      "  def format_term(x : Option[Unit => String]) : String = x match {\n" ^
+      "    case None => \"\"\n" ^
+      "    case Some(x) => x(())\n" ^
+      "  }\n" ^
+      "  def format(term : (Boolean, Option[Unit => String])) : String = term match {\n" ^
+      "      case (true, _) => \"True\\n\"\n" ^
+      "      case (false, x) => \"False\" + format_term(x) + \"\\n\"\n" ^
+      "  }\n" ^
+      "  def main(args:Array[String]) = {\n" ^
+      "    val result = " ^ value_name ^ "(());\n" ^
+      "    print(\"" ^ start_markerN ^ "\");\n" ^
+      "    result.map{test:(Boolean, Option[Unit => String]) => print(format(test))};\n" ^
+      "    print(\"" ^ end_markerN ^ "\");\n" ^
+      "  }\n" ^
+      "}\n"
+
+    val compile_cmd =
+      Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scalac")) ^
+      " -d " ^ File.shell_path path ^ " -classpath " ^ File.shell_path path ^ " " ^
+      File.shell_path code_path ^ " " ^ File.shell_path driver_path
+
+    val run_cmd =
+      Path.implode (Path.append (Path.variable ISABELLE_SCALA) (Path.basic "scala")) ^
+      " -cp " ^ File.shell_path path ^ " Test"
+  in
+    {files = [(driver_path, driver)],
+     compile_cmd = SOME compile_cmd, run_cmd = run_cmd, mk_code_file = K code_path}
+  end
+
+val evaluate_in_scala = gen_driver mk_driver_scala ISABELLE_SCALA "Scala directory" scalaN
+
+val test_codeP = Scan.repeat1 Parse.prop -- (@{keyword "in"} |-- Scan.repeat1 Parse.name)
+
+val _ = 
+  Outer_Syntax.command @{command_spec "test_code"}
+    "compile test cases to target languages, execute them and report results"
+      (test_codeP >> (fn (raw_ts, targets) => Toplevel.keep (test_code_cmd raw_ts targets)))
+
+val eval_termP = Parse.term -- (@{keyword "in"} |-- Parse.name)
+
+val _ = 
+  Outer_Syntax.command @{command_spec "eval_term"}
+    "evaluate term in target language"
+      (eval_termP >> (fn (raw_t, target) => Toplevel.keep (eval_term_cmd raw_t target)))
+
+val _ = Context.>> (Context.map_theory (
+  fold add_driver
+    [(polymlN, (evaluate_in_polyml, Code_ML.target_SML)),
+     (mltonN, (evaluate_in_mlton, Code_ML.target_SML)),
+     (smlnjN, (evaluate_in_smlnj, Code_ML.target_SML)),
+     (ocamlN, (evaluate_in_ocaml, Code_ML.target_OCaml)),
+     (ghcN, (evaluate_in_ghc, Code_Haskell.target)),
+     (scalaN, (evaluate_in_scala, Code_Scala.target))]))
+end
+
--- a/src/HOL/Decision_Procs/Cooper.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -2052,7 +2052,7 @@
     let ?v = "Neg e"
     have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))"
       by simp
-    from 7(5)[simplified simp_thms Inum.simps \<beta>.simps set_simps bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]]
+    from 7(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]]
     have nob: "\<not> (\<exists>j\<in> {1 ..d}. x = - ?e + j)"
       by auto
     from H p have "x + ?e > 0 \<and> x + ?e \<le> d"
@@ -2085,7 +2085,7 @@
     let ?v = "Sub (C -1) e"
     have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))"
       by simp
-    from 8(5)[simplified simp_thms Inum.simps \<beta>.simps set_simps bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]]
+    from 8(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]]
     have nob: "\<not> (\<exists>j\<in> {1 ..d}. x =  - ?e - 1 + j)"
       by auto
     from H p have "x + ?e \<ge> 0 \<and> x + ?e < d"
--- a/src/HOL/Decision_Procs/MIR.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -2612,7 +2612,7 @@
     {assume H: "\<not> real (x-d) + ?e > 0" 
       let ?v="Neg e"
       have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp
-      from 7(5)[simplified simp_thms Inum.simps \<beta>.simps set_simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] 
+      from 7(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] 
       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e + real j)" by auto 
       from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1)
       hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d"
@@ -2638,7 +2638,7 @@
     {assume H: "\<not> real (x-d) + ?e \<ge> 0" 
       let ?v="Sub (C -1) e"
       have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp
-      from 8(5)[simplified simp_thms Inum.simps \<beta>.simps set_simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] 
+      from 8(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] 
       have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x =  - ?e - 1 + real j)" by auto 
       from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1)
       hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d"
@@ -3394,7 +3394,7 @@
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
-    by (simp only: set_map set_upto set_simps)
+    by (simp only: set_map set_upto list.set)
   also have "\<dots> =   
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un 
@@ -3548,7 +3548,7 @@
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un 
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
-    by (simp only: set_map set_upto set_simps)
+    by (simp only: set_map set_upto list.set)
   also have "\<dots> =   
     ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un 
     (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un 
--- a/src/HOL/Deriv.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Deriv.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -50,24 +50,17 @@
 lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \<Longrightarrow> X = Y \<Longrightarrow> (f has_vector_derivative Y) F"
   by simp
 
-ML {*
-
-structure Derivative_Intros = Named_Thms
-(
-  val name = @{binding derivative_intros}
-  val description = "structural introduction rules for derivatives"
-)
-
-*}
-
+named_theorems derivative_intros "structural introduction rules for derivatives"
 setup {*
   let
-    val eq_thms = [@{thm has_derivative_eq_rhs}, @{thm DERIV_cong}, @{thm has_vector_derivative_eq_rhs}]
+    val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs}
     fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms
   in
-    Derivative_Intros.setup #>
     Global_Theory.add_thms_dynamic
-      (@{binding derivative_eq_intros}, map_filter eq_rule o Derivative_Intros.get o Context.proof_of)
+      (@{binding derivative_eq_intros},
+        fn context =>
+          Named_Theorems.get (Context.proof_of context) @{named_theorems derivative_intros}
+          |> map_filter eq_rule)
   end;
 *}
 
--- a/src/HOL/Enum.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Enum.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -537,6 +537,62 @@
 
 end
 
+instance finite_1 :: "{dense_linorder, wellorder}"
+by intro_classes (simp_all add: less_finite_1_def)
+
+instantiation finite_1 :: complete_lattice
+begin
+
+definition [simp]: "Inf = (\<lambda>_. a\<^sub>1)"
+definition [simp]: "Sup = (\<lambda>_. a\<^sub>1)"
+definition [simp]: "bot = a\<^sub>1"
+definition [simp]: "top = a\<^sub>1"
+definition [simp]: "inf = (\<lambda>_ _. a\<^sub>1)"
+definition [simp]: "sup = (\<lambda>_ _. a\<^sub>1)"
+
+instance by intro_classes(simp_all add: less_eq_finite_1_def)
+end
+
+instance finite_1 :: complete_distrib_lattice
+by intro_classes(simp_all add: INF_def SUP_def)
+
+instance finite_1 :: complete_linorder ..
+
+lemma finite_1_eq: "x = a\<^sub>1"
+by(cases x) simp
+
+simproc_setup finite_1_eq ("x::finite_1") = {*
+  fn _ => fn _ => fn ct => case term_of ct of
+    Const (@{const_name a\<^sub>1}, _) => NONE
+  | _ => SOME (mk_meta_eq @{thm finite_1_eq})
+*}
+
+instantiation finite_1 :: complete_boolean_algebra begin
+definition [simp]: "op - = (\<lambda>_ _. a\<^sub>1)"
+definition [simp]: "uminus = (\<lambda>_. a\<^sub>1)"
+instance by intro_classes simp_all
+end
+
+instantiation finite_1 :: 
+  "{linordered_ring_strict, linordered_comm_semiring_strict, ordered_comm_ring,
+    ordered_cancel_comm_monoid_diff, comm_monoid_mult, ordered_ring_abs,
+    one, Divides.div, sgn_if, inverse}"
+begin
+definition [simp]: "Groups.zero = a\<^sub>1"
+definition [simp]: "Groups.one = a\<^sub>1"
+definition [simp]: "op + = (\<lambda>_ _. a\<^sub>1)"
+definition [simp]: "op * = (\<lambda>_ _. a\<^sub>1)"
+definition [simp]: "op div = (\<lambda>_ _. a\<^sub>1)" 
+definition [simp]: "op mod = (\<lambda>_ _. a\<^sub>1)" 
+definition [simp]: "abs = (\<lambda>_. a\<^sub>1)"
+definition [simp]: "sgn = (\<lambda>_. a\<^sub>1)"
+definition [simp]: "inverse = (\<lambda>_. a\<^sub>1)"
+definition [simp]: "op / = (\<lambda>_ _. a\<^sub>1)"
+
+instance by intro_classes(simp_all add: less_finite_1_def)
+end
+
+declare [[simproc del: finite_1_eq]]
 hide_const (open) a\<^sub>1
 
 datatype finite_2 = a\<^sub>1 | a\<^sub>2
@@ -584,6 +640,65 @@
 
 end
 
+instance finite_2 :: wellorder
+by(rule wf_wellorderI)(simp add: less_finite_2_def, intro_classes)
+
+instantiation finite_2 :: complete_lattice
+begin
+
+definition "\<Sqinter>A = (if a\<^sub>1 \<in> A then a\<^sub>1 else a\<^sub>2)"
+definition "\<Squnion>A = (if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>1)"
+definition [simp]: "bot = a\<^sub>1"
+definition [simp]: "top = a\<^sub>2"
+definition "x \<sqinter> y = (if x = a\<^sub>1 \<or> y = a\<^sub>1 then a\<^sub>1 else a\<^sub>2)"
+definition "x \<squnion> y = (if x = a\<^sub>2 \<or> y = a\<^sub>2 then a\<^sub>2 else a\<^sub>1)"
+
+lemma neq_finite_2_a\<^sub>1_iff [simp]: "x \<noteq> a\<^sub>1 \<longleftrightarrow> x = a\<^sub>2"
+by(cases x) simp_all
+
+lemma neq_finite_2_a\<^sub>1_iff' [simp]: "a\<^sub>1 \<noteq> x \<longleftrightarrow> x = a\<^sub>2"
+by(cases x) simp_all
+
+lemma neq_finite_2_a\<^sub>2_iff [simp]: "x \<noteq> a\<^sub>2 \<longleftrightarrow> x = a\<^sub>1"
+by(cases x) simp_all
+
+lemma neq_finite_2_a\<^sub>2_iff' [simp]: "a\<^sub>2 \<noteq> x \<longleftrightarrow> x = a\<^sub>1"
+by(cases x) simp_all
+
+instance
+proof
+  fix x :: finite_2 and A
+  assume "x \<in> A"
+  then show "\<Sqinter>A \<le> x" "x \<le> \<Squnion>A"
+    by(case_tac [!] x)(auto simp add: less_eq_finite_2_def less_finite_2_def Inf_finite_2_def Sup_finite_2_def)
+qed(auto simp add: less_eq_finite_2_def less_finite_2_def inf_finite_2_def sup_finite_2_def Inf_finite_2_def Sup_finite_2_def)
+end
+
+instance finite_2 :: complete_distrib_lattice
+by(intro_classes)(auto simp add: INF_def SUP_def sup_finite_2_def inf_finite_2_def Inf_finite_2_def Sup_finite_2_def)
+
+instance finite_2 :: complete_linorder ..
+
+instantiation finite_2 :: "{field_inverse_zero, abs_if, ring_div, semiring_div_parity, sgn_if}" begin
+definition [simp]: "0 = a\<^sub>1"
+definition [simp]: "1 = a\<^sub>2"
+definition "x + y = (case (x, y) of (a\<^sub>1, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>1 | _ \<Rightarrow> a\<^sub>2)"
+definition "uminus = (\<lambda>x :: finite_2. x)"
+definition "op - = (op + :: finite_2 \<Rightarrow> _)"
+definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
+definition "inverse = (\<lambda>x :: finite_2. x)"
+definition "op / = (op * :: finite_2 \<Rightarrow> _)"
+definition "abs = (\<lambda>x :: finite_2. x)"
+definition "op div = (op / :: finite_2 \<Rightarrow> _)"
+definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | _ \<Rightarrow> a\<^sub>1)"
+definition "sgn = (\<lambda>x :: finite_2. x)"
+instance
+by intro_classes
+  (simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def times_finite_2_def
+       inverse_finite_2_def divide_finite_2_def abs_finite_2_def div_finite_2_def mod_finite_2_def sgn_finite_2_def
+     split: finite_2.splits)
+end
+
 hide_const (open) a\<^sub>1 a\<^sub>2
 
 datatype finite_3 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3
@@ -629,6 +744,85 @@
 
 end
 
+instance finite_3 :: wellorder
+proof(rule wf_wellorderI)
+  have "inv_image less_than (case_finite_3 0 1 2) = {(x, y). x < y}"
+    by(auto simp add: less_finite_3_def split: finite_3.splits)
+  from this[symmetric] show "wf \<dots>" by simp
+qed intro_classes
+
+instantiation finite_3 :: complete_lattice
+begin
+
+definition "\<Sqinter>A = (if a\<^sub>1 \<in> A then a\<^sub>1 else if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>3)"
+definition "\<Squnion>A = (if a\<^sub>3 \<in> A then a\<^sub>3 else if a\<^sub>2 \<in> A then a\<^sub>2 else a\<^sub>1)"
+definition [simp]: "bot = a\<^sub>1"
+definition [simp]: "top = a\<^sub>3"
+definition [simp]: "inf = (min :: finite_3 \<Rightarrow> _)"
+definition [simp]: "sup = (max :: finite_3 \<Rightarrow> _)"
+
+instance
+proof
+  fix x :: finite_3 and A
+  assume "x \<in> A"
+  then show "\<Sqinter>A \<le> x" "x \<le> \<Squnion>A"
+    by(case_tac [!] x)(auto simp add: Inf_finite_3_def Sup_finite_3_def less_eq_finite_3_def less_finite_3_def)
+next
+  fix A and z :: finite_3
+  assume "\<And>x. x \<in> A \<Longrightarrow> z \<le> x"
+  then show "z \<le> \<Sqinter>A"
+    by(cases z)(auto simp add: Inf_finite_3_def less_eq_finite_3_def less_finite_3_def)
+next
+  fix A and z :: finite_3
+  assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z"
+  show "\<Squnion>A \<le> z"
+    by(auto simp add: Sup_finite_3_def less_eq_finite_3_def less_finite_3_def dest: *)
+qed(auto simp add: Inf_finite_3_def Sup_finite_3_def)
+end
+
+instance finite_3 :: complete_distrib_lattice
+proof
+  fix a :: finite_3 and B
+  show "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
+  proof(cases a "\<Sqinter>B" rule: finite_3.exhaust[case_product finite_3.exhaust])
+    case a\<^sub>2_a\<^sub>3
+    then have "\<And>x. x \<in> B \<Longrightarrow> x = a\<^sub>3"
+      by(case_tac x)(auto simp add: Inf_finite_3_def split: split_if_asm)
+    then show ?thesis using a\<^sub>2_a\<^sub>3
+      by(auto simp add: INF_def Inf_finite_3_def max_def less_eq_finite_3_def less_finite_3_def split: split_if_asm)
+  qed(auto simp add: INF_def Inf_finite_3_def max_def less_finite_3_def less_eq_finite_3_def split: split_if_asm)
+  show "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
+    by(cases a "\<Squnion>B" rule: finite_3.exhaust[case_product finite_3.exhaust])
+      (auto simp add: SUP_def Sup_finite_3_def min_def less_finite_3_def less_eq_finite_3_def split: split_if_asm)
+qed
+
+instance finite_3 :: complete_linorder ..
+
+instantiation finite_3 :: "{field_inverse_zero, abs_if, ring_div, semiring_div, sgn_if}" begin
+definition [simp]: "0 = a\<^sub>1"
+definition [simp]: "1 = a\<^sub>2"
+definition
+  "x + y = (case (x, y) of
+     (a\<^sub>1, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>1 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>1
+   | (a\<^sub>1, a\<^sub>2) \<Rightarrow> a\<^sub>2 | (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \<Rightarrow> a\<^sub>2
+   | _ \<Rightarrow> a\<^sub>3)"
+definition "- x = (case x of a\<^sub>1 \<Rightarrow> a\<^sub>1 | a\<^sub>2 \<Rightarrow> a\<^sub>3 | a\<^sub>3 \<Rightarrow> a\<^sub>2)"
+definition "x - y = x + (- y :: finite_3)"
+definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \<Rightarrow> a\<^sub>2 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>3 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"
+definition "inverse = (\<lambda>x :: finite_3. x)" 
+definition "x / y = x * inverse (y :: finite_3)"
+definition "abs = (\<lambda>x :: finite_3. x)"
+definition "op div = (op / :: finite_3 \<Rightarrow> _)"
+definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \<Rightarrow> a\<^sub>2 | (a\<^sub>3, a\<^sub>1) \<Rightarrow> a\<^sub>3 | _ \<Rightarrow> a\<^sub>1)"
+definition "sgn = (\<lambda>x. case x of a\<^sub>1 \<Rightarrow> a\<^sub>1 | _ \<Rightarrow> a\<^sub>2)"
+instance
+by intro_classes
+  (simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def times_finite_3_def
+       inverse_finite_3_def divide_finite_3_def abs_finite_3_def div_finite_3_def mod_finite_3_def sgn_finite_3_def
+       less_finite_3_def
+     split: finite_3.splits)
+end
+
 hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3
 
 datatype finite_4 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4
@@ -659,6 +853,77 @@
 
 end
 
+instantiation finite_4 :: complete_lattice begin
+
+text {* @{term a\<^sub>1} $<$ @{term a\<^sub>2},@{term a\<^sub>3} $<$ @{term a\<^sub>4},
+  but @{term a\<^sub>2} and @{term a\<^sub>3} are incomparable. *}
+
+definition
+  "x < y \<longleftrightarrow> (case (x, y) of
+     (a\<^sub>1, a\<^sub>1) \<Rightarrow> False | (a\<^sub>1, _) \<Rightarrow> True
+   |  (a\<^sub>2, a\<^sub>4) \<Rightarrow> True
+   |  (a\<^sub>3, a\<^sub>4) \<Rightarrow> True  | _ \<Rightarrow> False)"
+
+definition 
+  "x \<le> y \<longleftrightarrow> (case (x, y) of
+     (a\<^sub>1, _) \<Rightarrow> True
+   | (a\<^sub>2, a\<^sub>2) \<Rightarrow> True | (a\<^sub>2, a\<^sub>4) \<Rightarrow> True
+   | (a\<^sub>3, a\<^sub>3) \<Rightarrow> True | (a\<^sub>3, a\<^sub>4) \<Rightarrow> True
+   | (a\<^sub>4, a\<^sub>4) \<Rightarrow> True | _ \<Rightarrow> False)"
+
+definition
+  "\<Sqinter>A = (if a\<^sub>1 \<in> A \<or> a\<^sub>2 \<in> A \<and> a\<^sub>3 \<in> A then a\<^sub>1 else if a\<^sub>2 \<in> A then a\<^sub>2 else if a\<^sub>3 \<in> A then a\<^sub>3 else a\<^sub>4)"
+definition
+  "\<Squnion>A = (if a\<^sub>4 \<in> A \<or> a\<^sub>2 \<in> A \<and> a\<^sub>3 \<in> A then a\<^sub>4 else if a\<^sub>2 \<in> A then a\<^sub>2 else if a\<^sub>3 \<in> A then a\<^sub>3 else a\<^sub>1)"
+definition [simp]: "bot = a\<^sub>1"
+definition [simp]: "top = a\<^sub>4"
+definition
+  "x \<sqinter> y = (case (x, y) of
+     (a\<^sub>1, _) \<Rightarrow> a\<^sub>1 | (_, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>1 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>1
+   | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
+   | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
+   | _ \<Rightarrow> a\<^sub>4)"
+definition
+  "x \<squnion> y = (case (x, y) of
+     (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4 | (a\<^sub>2, a\<^sub>3) \<Rightarrow> a\<^sub>4 | (a\<^sub>3, a\<^sub>2) \<Rightarrow> a\<^sub>4
+  | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
+  | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
+  | _ \<Rightarrow> a\<^sub>1)"
+
+instance
+proof
+  fix A and z :: finite_4
+  assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z"
+  show "\<Squnion>A \<le> z"
+    by(auto simp add: Sup_finite_4_def less_eq_finite_4_def dest!: * split: finite_4.splits)
+next
+  fix A and z :: finite_4
+  assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x"
+  show "z \<le> \<Sqinter>A"
+    by(auto simp add: Inf_finite_4_def less_eq_finite_4_def dest!: * split: finite_4.splits)
+qed(auto simp add: less_finite_4_def less_eq_finite_4_def Inf_finite_4_def Sup_finite_4_def inf_finite_4_def sup_finite_4_def split: finite_4.splits)
+
+end
+
+instance finite_4 :: complete_distrib_lattice
+proof
+  fix a :: finite_4 and B
+  show "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
+    by(cases a "\<Sqinter>B" rule: finite_4.exhaust[case_product finite_4.exhaust])
+      (auto simp add: sup_finite_4_def Inf_finite_4_def INF_def split: finite_4.splits split_if_asm)
+  show "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
+    by(cases a "\<Squnion>B" rule: finite_4.exhaust[case_product finite_4.exhaust])
+      (auto simp add: inf_finite_4_def Sup_finite_4_def SUP_def split: finite_4.splits split_if_asm)
+qed
+
+instantiation finite_4 :: complete_boolean_algebra begin
+definition "- x = (case x of a\<^sub>1 \<Rightarrow> a\<^sub>4 | a\<^sub>2 \<Rightarrow> a\<^sub>3 | a\<^sub>3 \<Rightarrow> a\<^sub>2 | a\<^sub>4 \<Rightarrow> a\<^sub>1)"
+definition "x - y = x \<sqinter> - (y :: finite_4)"
+instance
+by intro_classes
+  (simp_all add: inf_finite_4_def sup_finite_4_def uminus_finite_4_def minus_finite_4_def split: finite_4.splits)
+end
+
 hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4
 
 
@@ -691,6 +956,72 @@
 
 end
 
+instantiation finite_5 :: complete_lattice
+begin
+
+text {* The non-distributive pentagon lattice $N_5$ *}
+
+definition
+  "x < y \<longleftrightarrow> (case (x, y) of
+     (a\<^sub>1, a\<^sub>1) \<Rightarrow> False | (a\<^sub>1, _) \<Rightarrow> True
+   | (a\<^sub>2, a\<^sub>3) \<Rightarrow> True  | (a\<^sub>2, a\<^sub>5) \<Rightarrow> True
+   | (a\<^sub>3, a\<^sub>5) \<Rightarrow> True
+   | (a\<^sub>4, a\<^sub>5) \<Rightarrow> True  | _ \<Rightarrow> False)"
+
+definition
+  "x \<le> y \<longleftrightarrow> (case (x, y) of
+     (a\<^sub>1, _) \<Rightarrow> True
+   | (a\<^sub>2, a\<^sub>2) \<Rightarrow> True | (a\<^sub>2, a\<^sub>3) \<Rightarrow> True | (a\<^sub>2, a\<^sub>5) \<Rightarrow> True
+   | (a\<^sub>3, a\<^sub>3) \<Rightarrow> True | (a\<^sub>3, a\<^sub>5) \<Rightarrow> True
+   | (a\<^sub>4, a\<^sub>4) \<Rightarrow> True | (a\<^sub>4, a\<^sub>5) \<Rightarrow> True
+   | (a\<^sub>5, a\<^sub>5) \<Rightarrow> True | _ \<Rightarrow> False)"
+
+definition
+  "\<Sqinter>A = 
+  (if a\<^sub>1 \<in> A \<or> a\<^sub>4 \<in> A \<and> (a\<^sub>2 \<in> A \<or> a\<^sub>3 \<in> A) then a\<^sub>1
+   else if a\<^sub>2 \<in> A then a\<^sub>2
+   else if a\<^sub>3 \<in> A then a\<^sub>3
+   else if a\<^sub>4 \<in> A then a\<^sub>4
+   else a\<^sub>5)"
+definition
+  "\<Squnion>A = 
+  (if a\<^sub>5 \<in> A \<or> a\<^sub>4 \<in> A \<and> (a\<^sub>2 \<in> A \<or> a\<^sub>3 \<in> A) then a\<^sub>5
+   else if a\<^sub>3 \<in> A then a\<^sub>3
+   else if a\<^sub>2 \<in> A then a\<^sub>2
+   else if a\<^sub>4 \<in> A then a\<^sub>4
+   else a\<^sub>1)"
+definition [simp]: "bot = a\<^sub>1"
+definition [simp]: "top = a\<^sub>5"
+definition
+  "x \<sqinter> y = (case (x, y) of
+     (a\<^sub>1, _) \<Rightarrow> a\<^sub>1 | (_, a\<^sub>1) \<Rightarrow> a\<^sub>1 | (a\<^sub>2, a\<^sub>4) \<Rightarrow> a\<^sub>1 | (a\<^sub>4, a\<^sub>2) \<Rightarrow> a\<^sub>1 | (a\<^sub>3, a\<^sub>4) \<Rightarrow> a\<^sub>1 | (a\<^sub>4, a\<^sub>3) \<Rightarrow> a\<^sub>1
+   | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
+   | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
+   | (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4
+   | _ \<Rightarrow> a\<^sub>5)"
+definition
+  "x \<squnion> y = (case (x, y) of
+     (a\<^sub>5, _) \<Rightarrow> a\<^sub>5 | (_, a\<^sub>5) \<Rightarrow> a\<^sub>5 | (a\<^sub>2, a\<^sub>4) \<Rightarrow> a\<^sub>5 | (a\<^sub>4, a\<^sub>2) \<Rightarrow> a\<^sub>5 | (a\<^sub>3, a\<^sub>4) \<Rightarrow> a\<^sub>5 | (a\<^sub>4, a\<^sub>3) \<Rightarrow> a\<^sub>5
+   | (a\<^sub>3, _) \<Rightarrow> a\<^sub>3 | (_, a\<^sub>3) \<Rightarrow> a\<^sub>3
+   | (a\<^sub>2, _) \<Rightarrow> a\<^sub>2 | (_, a\<^sub>2) \<Rightarrow> a\<^sub>2
+   | (a\<^sub>4, _) \<Rightarrow> a\<^sub>4 | (_, a\<^sub>4) \<Rightarrow> a\<^sub>4
+   | _ \<Rightarrow> a\<^sub>1)"
+
+instance 
+proof intro_classes
+  fix A and z :: finite_5
+  assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x"
+  show "z \<le> \<Sqinter>A"
+    by(auto simp add: less_eq_finite_5_def Inf_finite_5_def split: finite_5.splits split_if_asm dest!: *)
+next
+  fix A and z :: finite_5
+  assume *: "\<And>x. x \<in> A \<Longrightarrow> x \<le> z"
+  show "\<Squnion>A \<le> z"
+    by(auto simp add: less_eq_finite_5_def Sup_finite_5_def split: finite_5.splits split_if_asm dest!: *)
+qed(auto simp add: less_eq_finite_5_def less_finite_5_def inf_finite_5_def sup_finite_5_def Inf_finite_5_def Sup_finite_5_def split: finite_5.splits split_if_asm)
+
+end
+
 hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5
 
 
--- a/src/HOL/Fields.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Fields.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -25,15 +25,7 @@
 
 text{* Lemmas @{text divide_simps} move division to the outside and eliminates them on (in)equalities. *}
 
-ML {*
-structure Divide_Simps = Named_Thms
-(
-  val name = @{binding divide_simps}
-  val description = "rewrite rules to eliminate divisions"
-)
-*}
-
-setup Divide_Simps.setup
+named_theorems divide_simps "rewrite rules to eliminate divisions"
 
 
 class division_ring = ring_1 + inverse +
--- a/src/HOL/Fun_Def.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Fun_Def.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -103,7 +103,7 @@
 ML_file "Tools/Function/induction_schema.ML"
 
 method_setup induction_schema = {*
-  Scan.succeed (RAW_METHOD o Induction_Schema.induction_schema_tac)
+  Scan.succeed (NO_CASES oo Induction_Schema.induction_schema_tac)
 *} "prove an induction principle"
 
 setup {*
@@ -117,8 +117,8 @@
 inductive is_measure :: "('a \<Rightarrow> nat) \<Rightarrow> bool"
 where is_measure_trivial: "is_measure f"
 
+named_theorems measure_function "rules that guide the heuristic generation of measure functions"
 ML_file "Tools/Function/measure_functions.ML"
-setup MeasureFunctions.setup
 
 lemma measure_size[measure_function]: "is_measure size"
 by (rule is_measure_trivial)
--- a/src/HOL/Fun_Def_Base.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Fun_Def_Base.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -9,6 +9,7 @@
 begin
 
 ML_file "Tools/Function/function_lib.ML"
+named_theorems termination_simp "simplification rules for termination proofs"
 ML_file "Tools/Function/function_common.ML"
 ML_file "Tools/Function/context_tree.ML"
 setup Function_Ctx_Tree.setup
--- a/src/HOL/Groebner_Basis.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Groebner_Basis.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -33,16 +33,7 @@
     "\<not> P \<Longrightarrow> (P \<equiv> False)"
   by auto
 
-ML {*
-structure Algebra_Simplification = Named_Thms
-(
-  val name = @{binding algebra}
-  val description = "pre-simplification rules for algebraic methods"
-)
-*}
-
-setup Algebra_Simplification.setup
-
+named_theorems algebra "pre-simplification rules for algebraic methods"
 ML_file "Tools/groebner.ML"
 
 method_setup algebra = {*
--- a/src/HOL/Groups.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Groups.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -8,17 +8,10 @@
 imports Orderings
 begin
 
-subsection {* Fact collections *}
+subsection {* Dynamic facts *}
 
-ML {*
-structure Ac_Simps = Named_Thms
-(
-  val name = @{binding ac_simps}
-  val description = "associativity and commutativity simplification rules"
-)
-*}
+named_theorems ac_simps "associativity and commutativity simplification rules"
 
-setup Ac_Simps.setup
 
 text{* The rewrites accumulated in @{text algebra_simps} deal with the
 classical algebraic structures of groups, rings and family. They simplify
@@ -29,30 +22,15 @@
 Of course it also works for fields, but it knows nothing about multiplicative
 inverses or division. This is catered for by @{text field_simps}. *}
 
-ML {*
-structure Algebra_Simps = Named_Thms
-(
-  val name = @{binding algebra_simps}
-  val description = "algebra simplification rules"
-)
-*}
+named_theorems algebra_simps "algebra simplification rules"
 
-setup Algebra_Simps.setup
 
 text{* Lemmas @{text field_simps} multiply with denominators in (in)equations
 if they can be proved to be non-zero (for equations) or positive/negative
 (for inequations). Can be too aggressive and is therefore separate from the
 more benign @{text algebra_simps}. *}
 
-ML {*
-structure Field_Simps = Named_Thms
-(
-  val name = @{binding field_simps}
-  val description = "algebra simplification rules for fields"
-)
-*}
-
-setup Field_Simps.setup
+named_theorems field_simps "algebra simplification rules for fields"
 
 
 subsection {* Abstract structures *}
--- a/src/HOL/HOL.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOL.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -763,8 +763,6 @@
 
 subsubsection {* Atomizing elimination rules *}
 
-setup AtomizeElim.setup
-
 lemma atomize_exL[atomize_elim]: "(!!x. P x ==> Q) == ((EX x. P x) ==> Q)"
   by rule iprover+
 
@@ -790,15 +788,7 @@
 seldom-used facts. Some duplicate other rules.
 *}
 
-ML {*
-structure No_ATPs = Named_Thms
-(
-  val name = @{binding no_atp}
-  val description = "theorems that should be filtered out by Sledgehammer"
-)
-*}
-
-setup {* No_ATPs.setup *}
+named_theorems no_atp "theorems that should be filtered out by Sledgehammer"
 
 
 subsubsection {* Classical Reasoner setup *}
@@ -1931,35 +1921,14 @@
 
 subsubsection {* Nitpick setup *}
 
-ML {*
-structure Nitpick_Unfolds = Named_Thms
-(
-  val name = @{binding nitpick_unfold}
-  val description = "alternative definitions of constants as needed by Nitpick"
-)
-structure Nitpick_Simps = Named_Thms
-(
-  val name = @{binding nitpick_simp}
-  val description = "equational specification of constants as needed by Nitpick"
-)
-structure Nitpick_Psimps = Named_Thms
-(
-  val name = @{binding nitpick_psimp}
-  val description = "partial equational specification of constants as needed by Nitpick"
-)
-structure Nitpick_Choice_Specs = Named_Thms
-(
-  val name = @{binding nitpick_choice_spec}
-  val description = "choice specification of constants as needed by Nitpick"
-)
-*}
-
-setup {*
-  Nitpick_Unfolds.setup
-  #> Nitpick_Simps.setup
-  #> Nitpick_Psimps.setup
-  #> Nitpick_Choice_Specs.setup
-*}
+named_theorems nitpick_unfold
+  "alternative definitions of constants as needed by Nitpick"
+named_theorems nitpick_simp
+  "equational specification of constants as needed by Nitpick"
+named_theorems nitpick_psimp
+  "partial equational specification of constants as needed by Nitpick"
+named_theorems nitpick_choice_spec
+  "choice specification of constants as needed by Nitpick"
 
 declare if_bool_eq_conj [nitpick_unfold, no_atp]
         if_bool_eq_disj [no_atp]
@@ -1967,29 +1936,12 @@
 
 subsection {* Preprocessing for the predicate compiler *}
 
-ML {*
-structure Predicate_Compile_Alternative_Defs = Named_Thms
-(
-  val name = @{binding code_pred_def}
-  val description = "alternative definitions of constants for the Predicate Compiler"
-)
-structure Predicate_Compile_Inline_Defs = Named_Thms
-(
-  val name = @{binding code_pred_inline}
-  val description = "inlining definitions for the Predicate Compiler"
-)
-structure Predicate_Compile_Simps = Named_Thms
-(
-  val name = @{binding code_pred_simp}
-  val description = "simplification rules for the optimisations in the Predicate Compiler"
-)
-*}
-
-setup {*
-  Predicate_Compile_Alternative_Defs.setup
-  #> Predicate_Compile_Inline_Defs.setup
-  #> Predicate_Compile_Simps.setup
-*}
+named_theorems code_pred_def
+  "alternative definitions of constants for the Predicate Compiler"
+named_theorems code_pred_inline
+  "inlining definitions for the Predicate Compiler"
+named_theorems code_pred_simp
+  "simplification rules for the optimisations in the Predicate Compiler"
 
 
 subsection {* Legacy tactics and ML bindings *}
--- a/src/HOL/HOLCF/Cfun.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Cfun.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -147,8 +147,8 @@
       val [T, U] = Thm.dest_ctyp (ctyp_of_term f);
       val tr = instantiate' [SOME T, SOME U] [SOME f]
           (mk_meta_eq @{thm Abs_cfun_inverse2});
-      val rules = Cont2ContData.get ctxt;
-      val tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules));
+      val rules = Named_Theorems.get ctxt @{named_theorems cont2cont};
+      val tac = SOLVED' (REPEAT_ALL_NEW (match_tac (rev rules)));
     in SOME (perhaps (SINGLE (tac 1)) tr) end
 *}
 
--- a/src/HOL/HOLCF/Cont.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Cont.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -120,15 +120,8 @@
 
 subsection {* Collection of continuity rules *}
 
-ML {*
-structure Cont2ContData = Named_Thms
-(
-  val name = @{binding cont2cont}
-  val description = "continuity intro rule"
-)
-*}
+named_theorems cont2cont "continuity intro rule"
 
-setup Cont2ContData.setup
 
 subsection {* Continuity of basic functions *}
 
--- a/src/HOL/HOLCF/Domain.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Domain.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -316,12 +316,13 @@
 
 subsection {* Setting up the domain package *}
 
+named_theorems domain_defl_simps "theorems like DEFL('a t) = t_defl$DEFL('a)"
+named_theorems domain_isodefl "theorems like isodefl d t ==> isodefl (foo_map$d) (foo_defl$t)"
+
 ML_file "Tools/Domain/domain_isomorphism.ML"
 ML_file "Tools/Domain/domain_axioms.ML"
 ML_file "Tools/Domain/domain.ML"
 
-setup Domain_Isomorphism.setup
-
 lemmas [domain_defl_simps] =
   DEFL_cfun DEFL_sfun DEFL_ssum DEFL_sprod DEFL_prod DEFL_u
   liftdefl_eq LIFTDEFL_prod u_liftdefl_liftdefl_of
--- a/src/HOL/HOLCF/Domain_Aux.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Domain_Aux.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -344,6 +344,9 @@
 
 subsection {* ML setup *}
 
+named_theorems domain_deflation "theorems like deflation a ==> deflation (foo_map$a)"
+named_theorems domain_map_ID "theorems like foo_map$ID = ID"
+
 ML_file "Tools/Domain/domain_take_proofs.ML"
 ML_file "Tools/cont_consts.ML"
 ML_file "Tools/cont_proc.ML"
--- a/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_isomorphism.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -28,8 +28,6 @@
   val domain_isomorphism_cmd :
     (string list * binding * mixfix * string * (binding * binding) option) list
       -> theory -> theory
-
-  val setup : theory -> theory
 end
 
 structure Domain_Isomorphism : DOMAIN_ISOMORPHISM =
@@ -41,24 +39,6 @@
 
 fun is_cpo thy T = Sign.of_sort thy (T, @{sort cpo})
 
-(******************************************************************************)
-(******************************** theory data *********************************)
-(******************************************************************************)
-
-structure RepData = Named_Thms
-(
-  val name = @{binding domain_defl_simps}
-  val description = "theorems like DEFL('a t) = t_defl$DEFL('a)"
-)
-
-structure IsodeflData = Named_Thms
-(
-  val name = @{binding domain_isodefl}
-  val description = "theorems like isodefl d t ==> isodefl (foo_map$d) (foo_defl$t)"
-)
-
-val setup = RepData.setup #> IsodeflData.setup
-
 
 (******************************************************************************)
 (************************** building types and terms **************************)
@@ -170,8 +150,8 @@
     val cont_thm =
       let
         val prop = mk_trp (mk_cont functional)
-        val rules = Cont2ContData.get (Proof_Context.init_global thy)
-        val tac = REPEAT_ALL_NEW (match_tac rules) 1
+        val rules = Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems cont2cont}
+        val tac = REPEAT_ALL_NEW (match_tac (rev rules)) 1
       in
         Goal.prove_global thy [] [] prop (K tac)
       end
@@ -207,8 +187,9 @@
     (tab2 : (typ * term) list)
     (T : typ) : term =
   let
-    val defl_simps = RepData.get (Proof_Context.init_global thy)
-    val rules = map (Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) defl_simps
+    val defl_simps =
+      Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems domain_defl_simps}
+    val rules = map (Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) (rev defl_simps)
     val rules' = map (apfst mk_DEFL) tab1 @ map (apfst mk_LIFTDEFL) tab2
     fun proc1 t =
       (case dest_DEFL t of
@@ -522,7 +503,8 @@
         val ((_, _, _, {DEFL, ...}), thy) =
           Domaindef.add_domaindef spec defl NONE thy
         (* declare domain_defl_simps rules *)
-        val thy = Context.theory_map (RepData.add_thm DEFL) thy
+        val thy =
+          Context.theory_map (Named_Theorems.add_thm @{named_theorems domain_defl_simps} DEFL) thy
       in
         (DEFL, thy)
       end
@@ -532,9 +514,10 @@
     fun mk_DEFL_eq_thm (lhsT, rhsT) =
       let
         val goal = mk_eqs (mk_DEFL lhsT, mk_DEFL rhsT)
-        val DEFL_simps = RepData.get (Proof_Context.init_global thy)
+        val DEFL_simps =
+          Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems domain_defl_simps}
         fun tac ctxt =
-          rewrite_goals_tac ctxt (map mk_meta_eq DEFL_simps)
+          rewrite_goals_tac ctxt (map mk_meta_eq (rev DEFL_simps))
           THEN TRY (resolve_tac defl_unfold_thms 1)
       in
         Goal.prove_global thy [] [] goal (tac o #context)
@@ -637,7 +620,7 @@
         val isodefl_rules =
           @{thms conjI isodefl_ID_DEFL isodefl_LIFTDEFL}
           @ isodefl_abs_rep_thms
-          @ IsodeflData.get (Proof_Context.init_global thy)
+          @ rev (Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems domain_isodefl})
       in
         Goal.prove_global thy [] assms goal (fn {prems, context = ctxt} =>
          EVERY
@@ -661,7 +644,9 @@
     val (isodefl_thms, thy) = thy |>
       (Global_Theory.add_thms o map (Thm.no_attributes o apsnd Drule.zero_var_indexes))
         (conjuncts isodefl_binds isodefl_thm)
-    val thy = fold (Context.theory_map o IsodeflData.add_thm) isodefl_thms thy
+    val thy =
+      fold (Context.theory_map o Named_Theorems.add_thm @{named_theorems domain_isodefl})
+        isodefl_thms thy
 
     (* prove map_ID theorems *)
     fun prove_map_ID_thm
--- a/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Tools/Domain/domain_take_proofs.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -123,31 +123,20 @@
   fun merge data = Symtab.merge (K true) data
 )
 
-structure DeflMapData = Named_Thms
-(
-  val name = @{binding domain_deflation}
-  val description = "theorems like deflation a ==> deflation (foo_map$a)"
-)
-
-structure Map_Id_Data = Named_Thms
-(
-  val name = @{binding domain_map_ID}
-  val description = "theorems like foo_map$ID = ID"
-)
-
 fun add_rec_type (tname, bs) =
     Rec_Data.map (Symtab.insert (K true) (tname, bs))
 
 fun add_deflation_thm thm =
-    Context.theory_map (DeflMapData.add_thm thm)
+    Context.theory_map (Named_Theorems.add_thm @{named_theorems domain_deflation} thm)
 
 val get_rec_tab = Rec_Data.get
-fun get_deflation_thms thy = DeflMapData.get (Proof_Context.init_global thy)
+fun get_deflation_thms thy =
+  rev (Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems domain_deflation})
 
-val map_ID_add = Map_Id_Data.add
-val get_map_ID_thms = Map_Id_Data.get o Proof_Context.init_global
+val map_ID_add = Named_Theorems.add @{named_theorems domain_map_ID}
+fun get_map_ID_thms thy =
+  rev (Named_Theorems.get (Proof_Context.init_global thy) @{named_theorems domain_map_ID})
 
-val _ = Theory.setup (DeflMapData.setup #> Map_Id_Data.setup)
 
 (******************************************************************************)
 (************************** building types and terms **************************)
--- a/src/HOL/HOLCF/Tools/fixrec.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/HOLCF/Tools/fixrec.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -130,8 +130,8 @@
           "or simp rules are configured for all non-HOLCF constants.\n" ^
           "The error occurred for the goal statement:\n" ^
           Syntax.string_of_term lthy prop)
-        val rules = Cont2ContData.get lthy
-        val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac rules))
+        val rules = Named_Theorems.get lthy @{named_theorems cont2cont}
+        val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac (rev rules)))
         val slow_tac = SOLVED' (simp_tac lthy)
         val tac = fast_tac 1 ORELSE slow_tac 1 ORELSE err
       in
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -39,13 +39,7 @@
   "(\<And>h. execute f h = execute g h) \<Longrightarrow> f = g"
     by (cases f, cases g) (auto simp: fun_eq_iff)
 
-ML {* structure Execute_Simps = Named_Thms
-(
-  val name = @{binding execute_simps}
-  val description = "simplification rules for execute"
-) *}
-
-setup Execute_Simps.setup
+named_theorems execute_simps "simplification rules for execute"
 
 lemma execute_Let [execute_simps]:
   "execute (let x = t in f x) = (let x = t in execute (f x))"
@@ -93,13 +87,7 @@
     and "execute f h \<noteq> None"
   using assms by (simp add: success_def)
 
-ML {* structure Success_Intros = Named_Thms
-(
-  val name = @{binding success_intros}
-  val description = "introduction rules for success"
-) *}
-
-setup Success_Intros.setup
+named_theorems success_intros "introduction rules for success"
 
 lemma success_tapI [success_intros]:
   "success (tap f) h"
@@ -167,19 +155,8 @@
   shows "a = b" and "h' = h''"
   using assms unfolding effect_def by auto
 
-ML {* structure Effect_Intros = Named_Thms
-(
-  val name = @{binding effect_intros}
-  val description = "introduction rules for effect"
-) *}
-
-ML {* structure Effect_Elims = Named_Thms
-(
-  val name = @{binding effect_elims}
-  val description = "elimination rules for effect"
-) *}
-
-setup "Effect_Intros.setup #> Effect_Elims.setup"
+named_theorems effect_intros "introduction rules for effect"
+named_theorems effect_elims "elimination rules for effect"
 
 lemma effect_LetI [effect_intros]:
   assumes "x = t" "effect (f x) h h' r"
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -642,7 +642,7 @@
   with init all_ref_present have q_is_new: "q \<notin> set (p#refs)"
     by (auto elim!: effect_refE intro!: Ref.noteq_I)
   from refs_of_p refs_of_q q_is_new have a3: "\<forall>qrs prs. refs_of' h2 q qrs \<and> refs_of' h2 p prs \<longrightarrow> set prs \<inter> set qrs = {}"
-    by (fastforce simp only: set_simps dest: refs_of'_is_fun)
+    by (fastforce simp only: list.set dest: refs_of'_is_fun)
   from rev'_invariant [OF effect_rev' a1 a2 a3] have "list_of h3 (Ref.get h3 v) (List.rev xs)" 
     unfolding list_of'_def by auto
   with lookup show ?thesis
--- a/src/HOL/Lattices_Big.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Lattices_Big.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -633,6 +633,16 @@
 
 end
 
+lemma Max_eq_if:
+  assumes "finite A"  "finite B"  "\<forall>a\<in>A. \<exists>b\<in>B. a \<le> b"  "\<forall>b\<in>B. \<exists>a\<in>A. b \<le> a"
+  shows "Max A = Max B"
+proof cases
+  assume "A = {}" thus ?thesis using assms by simp
+next
+  assume "A \<noteq> {}" thus ?thesis using assms
+    by(blast intro: antisym Max_in Max_ge_iff[THEN iffD2])
+qed
+
 lemma Min_antimono:
   assumes "M \<subseteq> N" and "M \<noteq> {}" and "finite N"
   shows "Min N \<le> Min M"
--- a/src/HOL/Library/Extended_Real.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/Extended_Real.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -91,21 +91,22 @@
   shows "-a = -b \<longleftrightarrow> a = b"
   by (cases rule: ereal2_cases[of a b]) simp_all
 
-function of_ereal :: "ereal \<Rightarrow> real" where
-  "of_ereal (ereal r) = r"
-| "of_ereal \<infinity> = 0"
-| "of_ereal (-\<infinity>) = 0"
+instantiation ereal :: real_of
+begin
+
+function real_ereal :: "ereal \<Rightarrow> real" where
+  "real_ereal (ereal r) = r"
+| "real_ereal \<infinity> = 0"
+| "real_ereal (-\<infinity>) = 0"
   by (auto intro: ereal_cases)
 termination by default (rule wf_empty)
 
-defs (overloaded)
-  real_of_ereal_def [code_unfold]: "real \<equiv> of_ereal"
+instance ..
+end
 
 lemma real_of_ereal[simp]:
   "real (- x :: ereal) = - (real x)"
-  "real (ereal r) = r"
-  "real (\<infinity>::ereal) = 0"
-  by (cases x) (simp_all add: real_of_ereal_def)
+  by (cases x) simp_all
 
 lemma range_ereal[simp]: "range ereal = UNIV - {\<infinity>, -\<infinity>}"
 proof safe
@@ -216,7 +217,7 @@
 instance ereal :: numeral ..
 
 lemma real_of_ereal_0[simp]: "real (0::ereal) = 0"
-  unfolding real_of_ereal_def zero_ereal_def by simp
+  unfolding zero_ereal_def by simp
 
 lemma abs_ereal_zero[simp]: "\<bar>0\<bar> = (0::ereal)"
   unfolding zero_ereal_def abs_ereal.simps by simp
--- a/src/HOL/Library/Float.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/Float.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -15,9 +15,15 @@
   morphisms real_of_float float_of
   unfolding float_def by auto
 
-defs (overloaded)
+instantiation float :: real_of
+begin
+
+definition real_float :: "float \<Rightarrow> real" where
   real_of_float_def[code_unfold]: "real \<equiv> real_of_float"
 
+instance ..
+end
+
 lemma type_definition_float': "type_definition real float_of float"
   using type_definition_float unfolding real_of_float_def .
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Lattice_Constructions.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,417 @@
+(*  Title:      HOL/Library/Lattice_Constructions.thy
+    Author:     Lukas Bulwahn
+    Copyright   2010 TU Muenchen
+*)
+
+theory Lattice_Constructions
+imports Main
+begin
+
+subsection {* Values extended by a bottom element *}
+
+datatype 'a bot = Value 'a | Bot
+
+instantiation bot :: (preorder) preorder
+begin
+
+definition less_eq_bot where
+  "x \<le> y \<longleftrightarrow> (case x of Bot \<Rightarrow> True | Value x \<Rightarrow> (case y of Bot \<Rightarrow> False | Value y \<Rightarrow> x \<le> y))"
+
+definition less_bot where
+  "x < y \<longleftrightarrow> (case y of Bot \<Rightarrow> False | Value y \<Rightarrow> (case x of Bot \<Rightarrow> True | Value x \<Rightarrow> x < y))"
+
+lemma less_eq_bot_Bot [simp]: "Bot \<le> x"
+  by (simp add: less_eq_bot_def)
+
+lemma less_eq_bot_Bot_code [code]: "Bot \<le> x \<longleftrightarrow> True"
+  by simp
+
+lemma less_eq_bot_Bot_is_Bot: "x \<le> Bot \<Longrightarrow> x = Bot"
+  by (cases x) (simp_all add: less_eq_bot_def)
+
+lemma less_eq_bot_Value_Bot [simp, code]: "Value x \<le> Bot \<longleftrightarrow> False"
+  by (simp add: less_eq_bot_def)
+
+lemma less_eq_bot_Value [simp, code]: "Value x \<le> Value y \<longleftrightarrow> x \<le> y"
+  by (simp add: less_eq_bot_def)
+
+lemma less_bot_Bot [simp, code]: "x < Bot \<longleftrightarrow> False"
+  by (simp add: less_bot_def)
+
+lemma less_bot_Bot_is_Value: "Bot < x \<Longrightarrow> \<exists>z. x = Value z"
+  by (cases x) (simp_all add: less_bot_def)
+
+lemma less_bot_Bot_Value [simp]: "Bot < Value x"
+  by (simp add: less_bot_def)
+
+lemma less_bot_Bot_Value_code [code]: "Bot < Value x \<longleftrightarrow> True"
+  by simp
+
+lemma less_bot_Value [simp, code]: "Value x < Value y \<longleftrightarrow> x < y"
+  by (simp add: less_bot_def)
+
+instance proof
+qed (auto simp add: less_eq_bot_def less_bot_def less_le_not_le elim: order_trans split: bot.splits)
+
+end 
+
+instance bot :: (order) order proof
+qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits)
+
+instance bot :: (linorder) linorder proof
+qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits)
+
+instantiation bot :: (order) bot
+begin
+
+definition "bot = Bot"
+
+instance ..
+
+end
+
+instantiation bot :: (top) top
+begin
+
+definition "top = Value top"
+
+instance ..
+
+end
+
+instantiation bot :: (semilattice_inf) semilattice_inf
+begin
+
+definition inf_bot
+where
+  "inf x y = (case x of Bot => Bot | Value v => (case y of Bot => Bot | Value v' => Value (inf v v')))"
+
+instance proof
+qed (auto simp add: inf_bot_def less_eq_bot_def split: bot.splits)
+
+end
+
+instantiation bot :: (semilattice_sup) semilattice_sup
+begin
+
+definition sup_bot
+where
+  "sup x y = (case x of Bot => y | Value v => (case y of Bot => x | Value v' => Value (sup v v')))"
+
+instance proof
+qed (auto simp add: sup_bot_def less_eq_bot_def split: bot.splits)
+
+end
+
+instance bot :: (lattice) bounded_lattice_bot
+by(intro_classes)(simp add: bot_bot_def)
+
+section {* Values extended by a top element *}
+
+datatype 'a top = Value 'a | Top
+
+instantiation top :: (preorder) preorder
+begin
+
+definition less_eq_top where
+  "x \<le> y \<longleftrightarrow> (case y of Top \<Rightarrow> True | Value y \<Rightarrow> (case x of Top \<Rightarrow> False | Value x \<Rightarrow> x \<le> y))"
+
+definition less_top where
+  "x < y \<longleftrightarrow> (case x of Top \<Rightarrow> False | Value x \<Rightarrow> (case y of Top \<Rightarrow> True | Value y \<Rightarrow> x < y))"
+
+lemma less_eq_top_Top [simp]: "x <= Top"
+  by (simp add: less_eq_top_def)
+
+lemma less_eq_top_Top_code [code]: "x \<le> Top \<longleftrightarrow> True"
+  by simp
+
+lemma less_eq_top_is_Top: "Top \<le> x \<Longrightarrow> x = Top"
+  by (cases x) (simp_all add: less_eq_top_def)
+
+lemma less_eq_top_Top_Value [simp, code]: "Top \<le> Value x \<longleftrightarrow> False"
+  by (simp add: less_eq_top_def)
+
+lemma less_eq_top_Value_Value [simp, code]: "Value x \<le> Value y \<longleftrightarrow> x \<le> y"
+  by (simp add: less_eq_top_def)
+
+lemma less_top_Top [simp, code]: "Top < x \<longleftrightarrow> False"
+  by (simp add: less_top_def)
+
+lemma less_top_Top_is_Value: "x < Top \<Longrightarrow> \<exists>z. x = Value z"
+  by (cases x) (simp_all add: less_top_def)
+
+lemma less_top_Value_Top [simp]: "Value x < Top"
+  by (simp add: less_top_def)
+
+lemma less_top_Value_Top_code [code]: "Value x < Top \<longleftrightarrow> True"
+  by simp
+
+lemma less_top_Value [simp, code]: "Value x < Value y \<longleftrightarrow> x < y"
+  by (simp add: less_top_def)
+
+instance proof
+qed (auto simp add: less_eq_top_def less_top_def less_le_not_le elim: order_trans split: top.splits)
+
+end 
+
+instance top :: (order) order proof
+qed (auto simp add: less_eq_top_def less_top_def split: top.splits)
+
+instance top :: (linorder) linorder proof
+qed (auto simp add: less_eq_top_def less_top_def split: top.splits)
+
+instantiation top :: (order) top
+begin
+
+definition "top = Top"
+
+instance ..
+
+end
+
+instantiation top :: (bot) bot
+begin
+
+definition "bot = Value bot"
+
+instance ..
+
+end
+
+instantiation top :: (semilattice_inf) semilattice_inf
+begin
+
+definition inf_top
+where
+  "inf x y = (case x of Top => y | Value v => (case y of Top => x | Value v' => Value (inf v v')))"
+
+instance proof
+qed (auto simp add: inf_top_def less_eq_top_def split: top.splits)
+
+end
+
+instantiation top :: (semilattice_sup) semilattice_sup
+begin
+
+definition sup_top
+where
+  "sup x y = (case x of Top => Top | Value v => (case y of Top => Top | Value v' => Value (sup v v')))"
+
+instance proof
+qed (auto simp add: sup_top_def less_eq_top_def split: top.splits)
+
+end
+
+instance top :: (lattice) bounded_lattice_top
+by(intro_classes)(simp add: top_top_def)
+
+subsection {* Values extended by a top and a bottom element *}
+
+datatype 'a flat_complete_lattice = Value 'a | Bot | Top
+
+instantiation flat_complete_lattice :: (type) order
+begin
+
+definition less_eq_flat_complete_lattice where
+  "x \<le> y == (case x of Bot => True | Value v1 => (case y of Bot => False | Value v2 => (v1 = v2) | Top => True) | Top => (y = Top))"
+
+definition less_flat_complete_lattice where
+  "x < y = (case x of Bot => \<not> (y = Bot) | Value v1 => (y = Top) | Top => False)"
+
+lemma [simp]: "Bot <= y"
+unfolding less_eq_flat_complete_lattice_def by auto
+
+lemma [simp]: "y <= Top"
+unfolding less_eq_flat_complete_lattice_def by (auto split: flat_complete_lattice.splits)
+
+lemma greater_than_two_values:
+  assumes "a ~= aa" "Value a <= z" "Value aa <= z"
+  shows "z = Top"
+using assms
+by (cases z) (auto simp add: less_eq_flat_complete_lattice_def)
+
+lemma lesser_than_two_values:
+  assumes "a ~= aa" "z <= Value a" "z <= Value aa"
+  shows "z = Bot"
+using assms
+by (cases z) (auto simp add: less_eq_flat_complete_lattice_def)
+
+instance proof
+qed (auto simp add: less_eq_flat_complete_lattice_def less_flat_complete_lattice_def split: flat_complete_lattice.splits)
+
+end
+
+instantiation flat_complete_lattice :: (type) bot
+begin
+
+definition "bot = Bot"
+
+instance ..
+
+end
+
+instantiation flat_complete_lattice :: (type) top
+begin
+
+definition "top = Top"
+
+instance ..
+
+end
+
+instantiation flat_complete_lattice :: (type) lattice
+begin
+
+definition inf_flat_complete_lattice
+where
+  "inf x y = (case x of Bot => Bot | Value v1 => (case y of Bot => Bot | Value v2 => if (v1 = v2) then x else Bot | Top => x) | Top => y)"
+
+definition sup_flat_complete_lattice
+where
+  "sup x y = (case x of Bot => y | Value v1 => (case y of Bot => x | Value v2 => if v1 = v2 then x else Top | Top => Top) | Top => Top)"
+
+instance proof
+qed (auto simp add: inf_flat_complete_lattice_def sup_flat_complete_lattice_def less_eq_flat_complete_lattice_def split: flat_complete_lattice.splits)
+
+end
+
+instantiation flat_complete_lattice :: (type) complete_lattice
+begin
+
+definition Sup_flat_complete_lattice
+where
+  "Sup A = (if (A = {} \<or> A = {Bot}) then Bot else (if (\<exists> v. A - {Bot} = {Value v}) then Value (THE v. A - {Bot} = {Value v}) else Top))"
+
+definition Inf_flat_complete_lattice
+where
+  "Inf A = (if (A = {} \<or> A = {Top}) then Top else (if (\<exists> v. A - {Top} = {Value v}) then Value (THE v. A - {Top} = {Value v}) else Bot))"
+ 
+instance
+proof
+  fix x A
+  assume "(x :: 'a flat_complete_lattice) : A"
+  {
+    fix v
+    assume "A - {Top} = {Value v}"
+    from this have "(THE v. A - {Top} = {Value v}) = v"
+      by (auto intro!: the1_equality)
+    moreover
+    from `x : A` `A - {Top} = {Value v}` have "x = Top \<or> x = Value v"
+      by auto
+    ultimately have "Value (THE v. A - {Top} = {Value v}) <= x"
+      by auto
+  }
+  from `x : A` this show "Inf A <= x"
+    unfolding Inf_flat_complete_lattice_def
+    by fastforce
+next
+  fix z A
+  assume z: "\<And>x. x : A ==> z <= (x :: 'a flat_complete_lattice)"
+  {
+    fix v
+    assume "A - {Top} = {Value v}"
+    moreover
+    from this have "(THE v. A - {Top} = {Value v}) = v"
+      by (auto intro!: the1_equality)
+    moreover
+    note z
+    moreover
+    ultimately have "z <= Value (THE v::'a. A - {Top} = {Value v})"
+      by auto
+  } moreover
+  {
+    assume not_one_value: "A ~= {}" "A ~= {Top}" "~ (EX v::'a. A - {Top} = {Value v})"
+    have "z <= Bot"
+    proof (cases "A - {Top} = {Bot}")
+      case True
+      from this z show ?thesis
+        by auto
+    next
+      case False
+      from not_one_value
+      obtain a1 where a1: "a1 : A - {Top}" by auto
+      from not_one_value False a1
+      obtain a2 where "a2 : A - {Top} \<and> a1 \<noteq> a2"
+        by (cases a1) auto
+      from this a1 z[of "a1"] z[of "a2"] show ?thesis
+        apply (cases a1)
+        apply auto
+        apply (cases a2)
+        apply auto
+        apply (auto dest!: lesser_than_two_values)
+        done
+    qed
+  } moreover
+  note z moreover
+  ultimately show "z <= Inf A"
+    unfolding Inf_flat_complete_lattice_def
+    by auto
+next
+  fix x A
+  assume "(x :: 'a flat_complete_lattice) : A"
+  {
+    fix v
+    assume "A - {Bot} = {Value v}"
+    from this have "(THE v. A - {Bot} = {Value v}) = v"
+      by (auto intro!: the1_equality)
+    moreover
+    from `x : A` `A - {Bot} = {Value v}` have "x = Bot \<or> x = Value v"
+      by auto
+    ultimately have "x <= Value (THE v. A - {Bot} = {Value v})"
+      by auto
+  }
+  from `x : A` this show "x <= Sup A"
+    unfolding Sup_flat_complete_lattice_def
+    by fastforce
+next
+  fix z A
+  assume z: "\<And>x. x : A ==> x <= (z :: 'a flat_complete_lattice)"
+  {
+    fix v
+    assume "A - {Bot} = {Value v}"
+    moreover
+    from this have "(THE v. A - {Bot} = {Value v}) = v"
+      by (auto intro!: the1_equality)
+    moreover
+    note z
+    moreover
+    ultimately have "Value (THE v::'a. A - {Bot} = {Value v}) <= z"
+      by auto
+  } moreover
+  {
+    assume not_one_value: "A ~= {}" "A ~= {Bot}" "~ (EX v::'a. A - {Bot} = {Value v})"
+    have "Top <= z"
+    proof (cases "A - {Bot} = {Top}")
+      case True
+      from this z show ?thesis
+        by auto
+    next
+      case False
+      from not_one_value
+      obtain a1 where a1: "a1 : A - {Bot}" by auto
+      from not_one_value False a1
+      obtain a2 where "a2 : A - {Bot} \<and> a1 \<noteq> a2"
+        by (cases a1) auto
+      from this a1 z[of "a1"] z[of "a2"] show ?thesis
+        apply (cases a1)
+        apply auto
+        apply (cases a2)
+        apply (auto dest!: greater_than_two_values)
+        done
+    qed
+  } moreover
+  note z moreover
+  ultimately show "Sup A <= z"
+    unfolding Sup_flat_complete_lattice_def
+    by auto
+next
+  show "Inf {} = (top :: 'a flat_complete_lattice)"
+    by(simp add: Inf_flat_complete_lattice_def top_flat_complete_lattice_def)
+next
+  show "Sup {} = (bot :: 'a flat_complete_lattice)"
+    by(simp add: Sup_flat_complete_lattice_def bot_flat_complete_lattice_def)
+qed
+
+end
+
+end
\ No newline at end of file
--- a/src/HOL/Library/Library.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/Library.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -32,6 +32,7 @@
   IArray
   Lattice_Algebras
   Lattice_Syntax
+  Lattice_Constructions
   ListVector
   Lubs_Glbs
   Mapping
--- a/src/HOL/Library/Multiset.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/Multiset.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -530,6 +530,17 @@
   "Multiset.filter P (M #\<inter> N) = Multiset.filter P M #\<inter> Multiset.filter P N"
   by (rule multiset_eqI) simp
 
+lemma multiset_filter_subset[simp]: "Multiset.filter f M \<le> M"
+  unfolding less_eq_multiset.rep_eq by auto
+
+lemma multiset_filter_mono: assumes "A \<le> B"
+  shows "Multiset.filter f A \<le> Multiset.filter f B"
+proof -
+  from assms[unfolded mset_le_exists_conv]
+  obtain C where B: "B = A + C" by auto
+  show ?thesis unfolding B by auto
+qed
+
 syntax
   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ :# _./ _#})")
 syntax (xsymbol)
@@ -1325,6 +1336,17 @@
   "mcard (multiset_of xs) = length xs"
   by (induct xs) simp_all
 
+lemma mcard_mono: assumes "A \<le> B"
+  shows "mcard A \<le> mcard B"
+proof -
+  from assms[unfolded mset_le_exists_conv]
+  obtain C where B: "B = A + C" by auto
+  show ?thesis unfolding B by (induct C, auto)
+qed
+
+lemma mcard_filter_lesseq[simp]: "mcard (Multiset.filter f M) \<le> mcard M"
+  by (rule mcard_mono[OF multiset_filter_subset])
+
 
 subsection {* Alternative representations *}
 
@@ -2224,697 +2246,214 @@
 
 subsection {* BNF setup *}
 
-lemma setsum_gt_0_iff:
-fixes f :: "'a \<Rightarrow> nat" assumes "finite A"
-shows "setsum f A > 0 \<longleftrightarrow> (\<exists> a \<in> A. f a > 0)"
-(is "?L \<longleftrightarrow> ?R")
-proof-
-  have "?L \<longleftrightarrow> \<not> setsum f A = 0" by fast
-  also have "... \<longleftrightarrow> (\<exists> a \<in> A. f a \<noteq> 0)" using assms by simp
-  also have "... \<longleftrightarrow> ?R" by simp
-  finally show ?thesis .
-qed
-
-lift_definition mmap :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" is
-  "\<lambda>h f b. setsum f {a. h a = b \<and> f a > 0} :: nat"
-unfolding multiset_def proof safe
-  fix h :: "'a \<Rightarrow> 'b" and f :: "'a \<Rightarrow> nat"
-  assume fin: "finite {a. 0 < f a}"  (is "finite ?A")
-  show "finite {b. 0 < setsum f {a. h a = b \<and> 0 < f a}}"
-  (is "finite {b. 0 < setsum f (?As b)}")
-  proof- let ?B = "{b. 0 < setsum f (?As b)}"
-    have "\<And> b. finite (?As b)" using fin by simp
-    hence B: "?B = {b. ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
-    hence "?B \<subseteq> h ` ?A" by auto
-    thus ?thesis using finite_surj[OF fin] by auto
-  qed
-qed
-
-lemma mmap_id0: "mmap id = id"
-proof (intro ext multiset_eqI)
-  fix f a show "count (mmap id f) a = count (id f) a"
-  proof (cases "count f a = 0")
-    case False
-    hence 1: "{aa. aa = a \<and> aa \<in># f} = {a}" by auto
-    thus ?thesis by transfer auto
-  qed (transfer, simp)
-qed
-
-lemma inj_on_setsum_inv:
-assumes 1: "(0::nat) < setsum (count f) {a. h a = b' \<and> a \<in># f}" (is "0 < setsum (count f) ?A'")
-and     2: "{a. h a = b \<and> a \<in># f} = {a. h a = b' \<and> a \<in># f}" (is "?A = ?A'")
-shows "b = b'"
-using assms by (auto simp add: setsum_gt_0_iff)
-
-lemma mmap_comp:
-fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
-shows "mmap (h2 o h1) = mmap h2 o mmap h1"
-proof (intro ext multiset_eqI)
-  fix f :: "'a multiset" fix c :: 'c
-  let ?A = "{a. h2 (h1 a) = c \<and> a \<in># f}"
-  let ?As = "\<lambda> b. {a. h1 a = b \<and> a \<in># f}"
-  let ?B = "{b. h2 b = c \<and> 0 < setsum (count f) (?As b)}"
-  have 0: "{?As b | b.  b \<in> ?B} = ?As ` ?B" by auto
-  have "\<And> b. finite (?As b)" by transfer (simp add: multiset_def)
-  hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
-  hence A: "?A = \<Union> {?As b | b.  b \<in> ?B}" by auto
-  have "setsum (count f) ?A = setsum (setsum (count f)) {?As b | b.  b \<in> ?B}"
-    unfolding A by transfer (intro setsum.Union_disjoint [simplified], auto simp: multiset_def setsum.Union_disjoint)
-  also have "... = setsum (setsum (count f)) (?As ` ?B)" unfolding 0 ..
-  also have "... = setsum (setsum (count f) o ?As) ?B"
-    by (intro setsum.reindex) (auto simp add: setsum_gt_0_iff inj_on_def)
-  also have "... = setsum (\<lambda> b. setsum (count f) (?As b)) ?B" unfolding comp_def ..
-  finally have "setsum (count f) ?A = setsum (\<lambda> b. setsum (count f) (?As b)) ?B" .
-  thus "count (mmap (h2 \<circ> h1) f) c = count ((mmap h2 \<circ> mmap h1) f) c"
-    by transfer (unfold comp_apply, blast)
-qed
-
-lemma mmap_cong:
-assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a"
-shows "mmap f M = mmap g M"
-using assms by transfer (auto intro!: setsum.cong)
-
-context
-begin
-interpretation lifting_syntax .
-
-lemma set_of_transfer[transfer_rule]: "(pcr_multiset op = ===> op =) (\<lambda>f. {a. 0 < f a}) set_of"
-  unfolding set_of_def pcr_multiset_def cr_multiset_def rel_fun_def by auto
-
-end
-
-lemma set_of_mmap: "set_of o mmap h = image h o set_of"
-proof (rule ext, unfold comp_apply)
-  fix M show "set_of (mmap h M) = h ` set_of M"
-    by transfer (auto simp add: multiset_def setsum_gt_0_iff)
-qed
-
-lemma multiset_of_surj:
-  "multiset_of ` {as. set as \<subseteq> A} = {M. set_of M \<subseteq> A}"
-proof safe
-  fix M assume M: "set_of M \<subseteq> A"
-  obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto
-  hence "set as \<subseteq> A" using M by auto
-  thus "M \<in> multiset_of ` {as. set as \<subseteq> A}" using eq by auto
+definition rel_mset where
+  "rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. multiset_of xs = X \<and> multiset_of ys = Y \<and> list_all2 R xs ys)"
+
+lemma multiset_of_zip_take_Cons_drop_twice:
+  assumes "length xs = length ys" "j \<le> length xs"
+  shows "multiset_of (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) =
+    multiset_of (zip xs ys) + {#(x, y)#}"
+using assms
+proof (induct xs ys arbitrary: x y j rule: list_induct2)
+  case Nil
+  thus ?case
+    by simp
 next
-  show "\<And>x xa xb. \<lbrakk>set xa \<subseteq> A; xb \<in> set_of (multiset_of xa)\<rbrakk> \<Longrightarrow> xb \<in> A"
-  by (erule set_mp) (unfold set_of_multiset_of)
-qed
-
-lemma card_of_set_of:
-"(card_of {M. set_of M \<subseteq> A}, card_of {as. set as \<subseteq> A}) \<in> ordLeq"
-apply(rule surj_imp_ordLeq[of _ multiset_of]) using multiset_of_surj by auto
-
-lemma nat_sum_induct:
-assumes "\<And>n1 n2. (\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow> phi m1 m2) \<Longrightarrow> phi n1 n2"
-shows "phi (n1::nat) (n2::nat)"
-proof-
-  let ?chi = "\<lambda> n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)"
-  have "?chi (n1,n2)"
-  apply(induct rule: measure_induct[of "\<lambda> n1n2. fst n1n2 + snd n1n2" ?chi])
-  using assms by (metis fstI sndI)
-  thus ?thesis by simp
-qed
-
-lemma matrix_count:
-fixes ct1 ct2 :: "nat \<Rightarrow> nat"
-assumes "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
-shows
-"\<exists> ct. (\<forall> i1 \<le> n1. setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ct1 i1) \<and>
-       (\<forall> i2 \<le> n2. setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ct2 i2)"
-(is "?phi ct1 ct2 n1 n2")
-proof-
-  have "\<forall> ct1 ct2 :: nat \<Rightarrow> nat.
-        setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"
-  proof(induct rule: nat_sum_induct[of
-"\<lambda> n1 n2. \<forall> ct1 ct2 :: nat \<Rightarrow> nat.
-     setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"],
-      clarify)
-  fix n1 n2 :: nat and ct1 ct2 :: "nat \<Rightarrow> nat"
-  assume IH: "\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow>
-                \<forall> dt1 dt2 :: nat \<Rightarrow> nat.
-                setsum dt1 {..<Suc m1} = setsum dt2 {..<Suc m2} \<longrightarrow> ?phi dt1 dt2 m1 m2"
-  and ss: "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
-  show "?phi ct1 ct2 n1 n2"
-  proof(cases n1)
-    case 0 note n1 = 0
-    show ?thesis
-    proof(cases n2)
-      case 0 note n2 = 0
-      let ?ct = "\<lambda> i1 i2. ct2 0"
-      show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by simp
-    next
-      case (Suc m2) note n2 = Suc
-      let ?ct = "\<lambda> i1 i2. ct2 i2"
-      show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
-    qed
+  case (Cons x xs y ys)
+  thus ?case
+  proof (cases "j = 0")
+    case True
+    thus ?thesis
+      by simp
   next
-    case (Suc m1) note n1 = Suc
-    show ?thesis
-    proof(cases n2)
-      case 0 note n2 = 0
-      let ?ct = "\<lambda> i1 i2. ct1 i1"
-      show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
-    next
-      case (Suc m2) note n2 = Suc
-      show ?thesis
-      proof(cases "ct1 n1 \<le> ct2 n2")
-        case True
-        def dt2 \<equiv> "\<lambda> i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2"
-        have "setsum ct1 {..<Suc m1} = setsum dt2 {..<Suc n2}"
-        unfolding dt2_def using ss n1 True by auto
-        hence "?phi ct1 dt2 m1 n2" using IH[of m1 n2] n1 by simp
-        then obtain dt where
-        1: "\<And> i1. i1 \<le> m1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc n2} = ct1 i1" and
-        2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc m1} = dt2 i2" by auto
-        let ?ct = "\<lambda> i1 i2. if i1 = n1 then (if i2 = n2 then ct1 n1 else 0)
-                                       else dt i1 i2"
-        show ?thesis apply(rule exI[of _ ?ct])
-        using n1 n2 1 2 True unfolding dt2_def by simp
-      next
-        case False
-        hence False: "ct2 n2 < ct1 n1" by simp
-        def dt1 \<equiv> "\<lambda> i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1"
-        have "setsum dt1 {..<Suc n1} = setsum ct2 {..<Suc m2}"
-        unfolding dt1_def using ss n2 False by auto
-        hence "?phi dt1 ct2 n1 m2" using IH[of n1 m2] n2 by simp
-        then obtain dt where
-        1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc m2} = dt1 i1" and
-        2: "\<And> i2. i2 \<le> m2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc n1} = ct2 i2" by force
-        let ?ct = "\<lambda> i1 i2. if i2 = n2 then (if i1 = n1 then ct2 n2 else 0)
-                                       else dt i1 i2"
-        show ?thesis apply(rule exI[of _ ?ct])
-        using n1 n2 1 2 False unfolding dt1_def by simp
-      qed
-    qed
-  qed
-  qed
-  thus ?thesis using assms by auto
-qed
-
-definition
-"inj2 u B1 B2 \<equiv>
- \<forall> b1 b1' b2 b2'. {b1,b1'} \<subseteq> B1 \<and> {b2,b2'} \<subseteq> B2 \<and> u b1 b2 = u b1' b2'
-                  \<longrightarrow> b1 = b1' \<and> b2 = b2'"
-
-lemma matrix_setsum_finite:
-assumes B1: "B1 \<noteq> {}" "finite B1" and B2: "B2 \<noteq> {}" "finite B2" and u: "inj2 u B1 B2"
-and ss: "setsum N1 B1 = setsum N2 B2"
-shows "\<exists> M :: 'a \<Rightarrow> nat.
-            (\<forall> b1 \<in> B1. setsum (\<lambda> b2. M (u b1 b2)) B2 = N1 b1) \<and>
-            (\<forall> b2 \<in> B2. setsum (\<lambda> b1. M (u b1 b2)) B1 = N2 b2)"
-proof-
-  obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps)
-  then obtain e1 where e1: "bij_betw e1 {..<Suc n1} B1"
-  using ex_bij_betw_finite_nat[OF B1(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
-  hence e1_inj: "inj_on e1 {..<Suc n1}" and e1_surj: "e1 ` {..<Suc n1} = B1"
-  unfolding bij_betw_def by auto
-  def f1 \<equiv> "inv_into {..<Suc n1} e1"
-  have f1: "bij_betw f1 B1 {..<Suc n1}"
-  and f1e1[simp]: "\<And> i1. i1 < Suc n1 \<Longrightarrow> f1 (e1 i1) = i1"
-  and e1f1[simp]: "\<And> b1. b1 \<in> B1 \<Longrightarrow> e1 (f1 b1) = b1" unfolding f1_def
-  apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff)
-  by (metis e1_surj f_inv_into_f)
-  (*  *)
-  obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps)
-  then obtain e2 where e2: "bij_betw e2 {..<Suc n2} B2"
-  using ex_bij_betw_finite_nat[OF B2(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
-  hence e2_inj: "inj_on e2 {..<Suc n2}" and e2_surj: "e2 ` {..<Suc n2} = B2"
-  unfolding bij_betw_def by auto
-  def f2 \<equiv> "inv_into {..<Suc n2} e2"
-  have f2: "bij_betw f2 B2 {..<Suc n2}"
-  and f2e2[simp]: "\<And> i2. i2 < Suc n2 \<Longrightarrow> f2 (e2 i2) = i2"
-  and e2f2[simp]: "\<And> b2. b2 \<in> B2 \<Longrightarrow> e2 (f2 b2) = b2" unfolding f2_def
-  apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff)
-  by (metis e2_surj f_inv_into_f)
-  (*  *)
-  let ?ct1 = "N1 o e1"  let ?ct2 = "N2 o e2"
-  have ss: "setsum ?ct1 {..<Suc n1} = setsum ?ct2 {..<Suc n2}"
-  unfolding setsum.reindex[OF e1_inj, symmetric] setsum.reindex[OF e2_inj, symmetric]
-  e1_surj e2_surj using ss .
-  obtain ct where
-  ct1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ?ct1 i1" and
-  ct2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ?ct2 i2"
-  using matrix_count[OF ss] by blast
-  (*  *)
-  def A \<equiv> "{u b1 b2 | b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"
-  have "\<forall> a \<in> A. \<exists> b1b2 \<in> B1 <*> B2. u (fst b1b2) (snd b1b2) = a"
-  unfolding A_def Ball_def mem_Collect_eq by auto
-  then obtain h1h2 where h12:
-  "\<And>a. a \<in> A \<Longrightarrow> u (fst (h1h2 a)) (snd (h1h2 a)) = a \<and> h1h2 a \<in> B1 <*> B2" by metis
-  def h1 \<equiv> "fst o h1h2"  def h2 \<equiv> "snd o h1h2"
-  have h12[simp]: "\<And>a. a \<in> A \<Longrightarrow> u (h1 a) (h2 a) = a"
-                  "\<And> a. a \<in> A \<Longrightarrow> h1 a \<in> B1"  "\<And> a. a \<in> A \<Longrightarrow> h2 a \<in> B2"
-  using h12 unfolding h1_def h2_def by force+
-  {fix b1 b2 assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2"
-   hence inA: "u b1 b2 \<in> A" unfolding A_def by auto
-   hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto
-   moreover have "h1 (u b1 b2) \<in> B1" "h2 (u b1 b2) \<in> B2" using inA by auto
-   ultimately have "h1 (u b1 b2) = b1 \<and> h2 (u b1 b2) = b2"
-   using u b1 b2 unfolding inj2_def by fastforce
-  }
-  hence h1[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h1 (u b1 b2) = b1" and
-        h2[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h2 (u b1 b2) = b2" by auto
-  def M \<equiv> "\<lambda> a. ct (f1 (h1 a)) (f2 (h2 a))"
-  show ?thesis
-  apply(rule exI[of _ M]) proof safe
-    fix b1 assume b1: "b1 \<in> B1"
-    hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def
-    by (metis image_eqI lessThan_iff less_Suc_eq_le)
-    have "(\<Sum>b2\<in>B2. M (u b1 b2)) = (\<Sum>i2<Suc n2. ct (f1 b1) (f2 (e2 i2)))"
-    unfolding e2_surj[symmetric] setsum.reindex[OF e2_inj]
-    unfolding M_def comp_def apply(intro setsum.cong) apply force
-    by (metis e2_surj b1 h1 h2 imageI)
-    also have "... = N1 b1" using b1 ct1[OF f1b1] by simp
-    finally show "(\<Sum>b2\<in>B2. M (u b1 b2)) = N1 b1" .
-  next
-    fix b2 assume b2: "b2 \<in> B2"
-    hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def
-    by (metis image_eqI lessThan_iff less_Suc_eq_le)
-    have "(\<Sum>b1\<in>B1. M (u b1 b2)) = (\<Sum>i1<Suc n1. ct (f1 (e1 i1)) (f2 b2))"
-    unfolding e1_surj[symmetric] setsum.reindex[OF e1_inj]
-    unfolding M_def comp_def apply(intro setsum.cong) apply force
-    by (metis e1_surj b2 h1 h2 imageI)
-    also have "... = N2 b2" using b2 ct2[OF f2b2] by simp
-    finally show "(\<Sum>b1\<in>B1. M (u b1 b2)) = N2 b2" .
-  qed
-qed
-
-lemma supp_vimage_mmap: "set_of M \<subseteq> f -` (set_of (mmap f M))"
-  by transfer (auto simp: multiset_def setsum_gt_0_iff)
-
-lemma mmap_ge_0: "b \<in># mmap f M \<longleftrightarrow> (\<exists>a. a \<in># M \<and> f a = b)"
-  by transfer (auto simp: multiset_def setsum_gt_0_iff)
-
-lemma finite_twosets:
-assumes "finite B1" and "finite B2"
-shows "finite {u b1 b2 |b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"  (is "finite ?A")
-proof-
-  have A: "?A = (\<lambda> b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force
-  show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto
+    case False
+    then obtain k where k: "j = Suc k"
+      by (case_tac j) simp
+    hence "k \<le> length xs"
+      using Cons.prems by auto
+    hence "multiset_of (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) =
+      multiset_of (zip xs ys) + {#(x, y)#}"
+      by (rule Cons.hyps(2))
+    thus ?thesis
+      unfolding k by (auto simp: add.commute union_lcomm)
+  qed      
 qed
 
-(* Weak pullbacks: *)
-definition wpull where
-"wpull A B1 B2 f1 f2 p1 p2 \<longleftrightarrow>
- (\<forall> b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<longrightarrow> (\<exists> a \<in> A. p1 a = b1 \<and> p2 a = b2))"
-
-(* Weak pseudo-pullbacks *)
-definition wppull where
-"wppull A B1 B2 f1 f2 e1 e2 p1 p2 \<longleftrightarrow>
- (\<forall> b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<longrightarrow>
-           (\<exists> a \<in> A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2))"
-
-
-(* The pullback of sets *)
-definition thePull where
-"thePull B1 B2 f1 f2 = {(b1,b2). b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2}"
-
-lemma wpull_thePull:
-"wpull (thePull B1 B2 f1 f2) B1 B2 f1 f2 fst snd"
-unfolding wpull_def thePull_def by auto
-
-lemma wppull_thePull:
-assumes "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
-shows
-"\<exists> j. \<forall> a' \<in> thePull B1 B2 f1 f2.
-   j a' \<in> A \<and>
-   e1 (p1 (j a')) = e1 (fst a') \<and> e2 (p2 (j a')) = e2 (snd a')"
-(is "\<exists> j. \<forall> a' \<in> ?A'. ?phi a' (j a')")
-proof(rule bchoice[of ?A' ?phi], default)
-  fix a' assume a': "a' \<in> ?A'"
-  hence "fst a' \<in> B1" unfolding thePull_def by auto
-  moreover
-  from a' have "snd a' \<in> B2" unfolding thePull_def by auto
-  moreover have "f1 (fst a') = f2 (snd a')"
-  using a' unfolding csquare_def thePull_def by auto
-  ultimately show "\<exists> ja'. ?phi a' ja'"
-  using assms unfolding wppull_def by blast
-qed
-
-lemma wpull_wppull:
-assumes wp: "wpull A' B1 B2 f1 f2 p1' p2'" and
-1: "\<forall> a' \<in> A'. j a' \<in> A \<and> e1 (p1 (j a')) = e1 (p1' a') \<and> e2 (p2 (j a')) = e2 (p2' a')"
-shows "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
-unfolding wppull_def proof safe
-  fix b1 b2
-  assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2" and f: "f1 b1 = f2 b2"
-  then obtain a' where a': "a' \<in> A'" and b1: "b1 = p1' a'" and b2: "b2 = p2' a'"
-  using wp unfolding wpull_def by blast
-  show "\<exists>a\<in>A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2"
-  apply (rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto
+lemma ex_multiset_of_zip_left:
+  assumes "length xs = length ys" "multiset_of xs' = multiset_of xs"
+  shows "\<exists>ys'. length ys' = length xs' \<and> multiset_of (zip xs' ys') = multiset_of (zip xs ys)"
+using assms 
+proof (induct xs ys arbitrary: xs' rule: list_induct2)
+  case Nil
+  thus ?case
+    by auto
+next
+  case (Cons x xs y ys xs')
+  obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x"
+  proof -
+    assume "\<And>j. \<lbrakk>j < length xs'; xs' ! j = x\<rbrakk> \<Longrightarrow> ?thesis"
+    moreover have "\<And>k m n. (m\<Colon>nat) + n < m + k \<or> \<not> n < k" by linarith
+    moreover have "\<And>n a as. n - n < length (a # as) \<or> n < n"
+      by (metis Nat.add_diff_inverse diff_add_inverse2 impossible_Cons le_add1
+        less_diff_conv not_add_less2)
+    moreover have "\<not> length xs' < length xs'" by blast
+    ultimately show ?thesis
+      by (metis (no_types) Cons.prems Nat.add_diff_inverse diff_add_inverse2 length_append
+        less_diff_conv list.set_intros(1) multiset_of_eq_setD nth_append_length split_list)
+  qed
+
+  def xsa \<equiv> "take j xs' @ drop (Suc j) xs'" 
+  have "multiset_of xs' = {#x#} + multiset_of xsa"
+    unfolding xsa_def using j_len nth_j
+    by (metis (no_types) ab_semigroup_add_class.add_ac(1) append_take_drop_id drop_Suc_conv_tl
+      multiset_of.simps(2) union_code union_commute)
+  hence ms_x: "multiset_of xsa = multiset_of xs"
+    by (metis Cons.prems add.commute add_right_imp_eq multiset_of.simps(2))
+  then obtain ysa where
+    len_a: "length ysa = length xsa" and ms_a: "multiset_of (zip xsa ysa) = multiset_of (zip xs ys)"
+    using Cons.hyps(2) by blast
+
+  def ys' \<equiv> "take j ysa @ y # drop j ysa"
+  have xs': "xs' = take j xsa @ x # drop j xsa"
+    using ms_x j_len nth_j Cons.prems xsa_def
+    by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc drop_Suc_conv_tl length_Cons
+      length_drop mcard_multiset_of)
+  have j_len': "j \<le> length xsa"
+    using j_len xs' xsa_def
+    by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less)
+  have "length ys' = length xs'"
+    unfolding ys'_def using Cons.prems len_a ms_x
+    by (metis add_Suc_right append_take_drop_id length_Cons length_append multiset_of_eq_length)
+  moreover have "multiset_of (zip xs' ys') = multiset_of (zip (x # xs) (y # ys))"
+    unfolding xs' ys'_def
+    by (rule trans[OF multiset_of_zip_take_Cons_drop_twice])
+      (auto simp: len_a ms_a j_len' add.commute)
+  ultimately show ?case
+    by blast
 qed
 
-lemma wppull_fstOp_sndOp:
-shows "wppull (Collect (split (P OO Q))) (Collect (split P)) (Collect (split Q))
-  snd fst fst snd (BNF_Def.fstOp P Q) (BNF_Def.sndOp P Q)"
-using pick_middlep unfolding wppull_def fstOp_def sndOp_def relcompp.simps by auto
-
-lemma wpull_mmap:
-fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set"
-assumes wp: "wpull A B1 B2 f1 f2 p1 p2"
-shows
-"wpull {M. set_of M \<subseteq> A}
-       {N1. set_of N1 \<subseteq> B1} {N2. set_of N2 \<subseteq> B2}
-       (mmap f1) (mmap f2) (mmap p1) (mmap p2)"
-unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq)
-  fix N1 :: "'b1 multiset" and N2 :: "'b2 multiset"
-  assume mmap': "mmap f1 N1 = mmap f2 N2"
-  and N1[simp]: "set_of N1 \<subseteq> B1"
-  and N2[simp]: "set_of N2 \<subseteq> B2"
-  def P \<equiv> "mmap f1 N1"
-  have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto
-  note P = P1 P2
-  have fin_N1[simp]: "finite (set_of N1)"
-   and fin_N2[simp]: "finite (set_of N2)"
-   and fin_P[simp]: "finite (set_of P)" by auto
-
-  def set1 \<equiv> "\<lambda> c. {b1 \<in> set_of N1. f1 b1 = c}"
-  have set1[simp]: "\<And> c b1. b1 \<in> set1 c \<Longrightarrow> f1 b1 = c" unfolding set1_def by auto
-  have fin_set1: "\<And> c. c \<in> set_of P \<Longrightarrow> finite (set1 c)"
-    using N1(1) unfolding set1_def multiset_def by auto
-  have set1_NE: "\<And> c. c \<in> set_of P \<Longrightarrow> set1 c \<noteq> {}"
-   unfolding set1_def set_of_def P mmap_ge_0 by auto
-  have supp_N1_set1: "set_of N1 = (\<Union> c \<in> set_of P. set1 c)"
-    using supp_vimage_mmap[of N1 f1] unfolding set1_def P1 by auto
-  hence set1_inclN1: "\<And>c. c \<in> set_of P \<Longrightarrow> set1 c \<subseteq> set_of N1" by auto
-  hence set1_incl: "\<And> c. c \<in> set_of P \<Longrightarrow> set1 c \<subseteq> B1" using N1 by blast
-  have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
-    unfolding set1_def by auto
-  have setsum_set1: "\<And> c. setsum (count N1) (set1 c) = count P c"
-    unfolding P1 set1_def by transfer (auto intro: setsum.cong)
-
-  def set2 \<equiv> "\<lambda> c. {b2 \<in> set_of N2. f2 b2 = c}"
-  have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto
-  have fin_set2: "\<And> c. c \<in> set_of P \<Longrightarrow> finite (set2 c)"
-  using N2(1) unfolding set2_def multiset_def by auto
-  have set2_NE: "\<And> c. c \<in> set_of P \<Longrightarrow> set2 c \<noteq> {}"
-    unfolding set2_def P2 mmap_ge_0 set_of_def by auto
-  have supp_N2_set2: "set_of N2 = (\<Union> c \<in> set_of P. set2 c)"
-    using supp_vimage_mmap[of N2 f2] unfolding set2_def P2 by auto
-  hence set2_inclN2: "\<And>c. c \<in> set_of P \<Longrightarrow> set2 c \<subseteq> set_of N2" by auto
-  hence set2_incl: "\<And> c. c \<in> set_of P \<Longrightarrow> set2 c \<subseteq> B2" using N2 by blast
-  have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
-    unfolding set2_def by auto
-  have setsum_set2: "\<And> c. setsum (count N2) (set2 c) = count P c"
-    unfolding P2 set2_def by transfer (auto intro: setsum.cong)
-
-  have ss: "\<And> c. c \<in> set_of P \<Longrightarrow> setsum (count N1) (set1 c) = setsum (count N2) (set2 c)"
-    unfolding setsum_set1 setsum_set2 ..
-  have "\<forall> c \<in> set_of P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
-          \<exists> a \<in> A. p1 a = fst b1b2 \<and> p2 a = snd b1b2"
-    using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq
-    by simp (metis set1 set2 set_rev_mp)
-  then obtain uu where uu:
-  "\<forall> c \<in> set_of P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
-     uu c b1b2 \<in> A \<and> p1 (uu c b1b2) = fst b1b2 \<and> p2 (uu c b1b2) = snd b1b2" by metis
-  def u \<equiv> "\<lambda> c b1 b2. uu c (b1,b2)"
-  have u[simp]:
-  "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> A"
-  "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p1 (u c b1 b2) = b1"
-  "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p2 (u c b1 b2) = b2"
-    using uu unfolding u_def by auto
-  {fix c assume c: "c \<in> set_of P"
-   have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify
-     fix b1 b1' b2 b2'
-     assume "{b1, b1'} \<subseteq> set1 c" "{b2, b2'} \<subseteq> set2 c" and 0: "u c b1 b2 = u c b1' b2'"
-     hence "p1 (u c b1 b2) = b1 \<and> p2 (u c b1 b2) = b2 \<and>
-            p1 (u c b1' b2') = b1' \<and> p2 (u c b1' b2') = b2'"
-     using u(2)[OF c] u(3)[OF c] by simp metis
-     thus "b1 = b1' \<and> b2 = b2'" using 0 by auto
-   qed
-  } note inj = this
-  def sset \<equiv> "\<lambda> c. {u c b1 b2 | b1 b2. b1 \<in> set1 c \<and> b2 \<in> set2 c}"
-  have fin_sset[simp]: "\<And> c. c \<in> set_of P \<Longrightarrow> finite (sset c)" unfolding sset_def
-    using fin_set1 fin_set2 finite_twosets by blast
-  have sset_A: "\<And> c. c \<in> set_of P \<Longrightarrow> sset c \<subseteq> A" unfolding sset_def by auto
-  {fix c a assume c: "c \<in> set_of P" and ac: "a \<in> sset c"
-   then obtain b1 b2 where b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
-   and a: "a = u c b1 b2" unfolding sset_def by auto
-   have "p1 a \<in> set1 c" and p2a: "p2 a \<in> set2 c"
-   using ac a b1 b2 c u(2) u(3) by simp+
-   hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c]
-   unfolding inj2_def by (metis c u(2) u(3))
-  } note u_p12[simp] = this
-  {fix c a assume c: "c \<in> set_of P" and ac: "a \<in> sset c"
-   hence "p1 a \<in> set1 c" unfolding sset_def by auto
-  }note p1[simp] = this
-  {fix c a assume c: "c \<in> set_of P" and ac: "a \<in> sset c"
-   hence "p2 a \<in> set2 c" unfolding sset_def by auto
-  }note p2[simp] = this
-
-  {fix c assume c: "c \<in> set_of P"
-   hence "\<exists> M. (\<forall> b1 \<in> set1 c. setsum (\<lambda> b2. M (u c b1 b2)) (set2 c) = count N1 b1) \<and>
-               (\<forall> b2 \<in> set2 c. setsum (\<lambda> b1. M (u c b1 b2)) (set1 c) = count N2 b2)"
-   unfolding sset_def
-   using matrix_setsum_finite[OF set1_NE[OF c] fin_set1[OF c]
-                                 set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto
-  }
-  then obtain Ms where
-  ss1: "\<And> c b1. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c\<rbrakk> \<Longrightarrow>
-                   setsum (\<lambda> b2. Ms c (u c b1 b2)) (set2 c) = count N1 b1" and
-  ss2: "\<And> c b2. \<lbrakk>c \<in> set_of P; b2 \<in> set2 c\<rbrakk> \<Longrightarrow>
-                   setsum (\<lambda> b1. Ms c (u c b1 b2)) (set1 c) = count N2 b2"
-  by metis
-  def SET \<equiv> "\<Union> c \<in> set_of P. sset c"
-  have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto
-  have SET_A: "SET \<subseteq> A" unfolding SET_def using sset_A by blast
-  have u_SET[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> SET"
-    unfolding SET_def sset_def by blast
-  {fix c a assume c: "c \<in> set_of P" and a: "a \<in> SET" and p1a: "p1 a \<in> set1 c"
-   then obtain c' where c': "c' \<in> set_of P" and ac': "a \<in> sset c'"
-    unfolding SET_def by auto
-   hence "p1 a \<in> set1 c'" unfolding sset_def by auto
-   hence eq: "c = c'" using p1a c c' set1_disj by auto
-   hence "a \<in> sset c" using ac' by simp
-  } note p1_rev = this
-  {fix c a assume c: "c \<in> set_of P" and a: "a \<in> SET" and p2a: "p2 a \<in> set2 c"
-   then obtain c' where c': "c' \<in> set_of P" and ac': "a \<in> sset c'"
-   unfolding SET_def by auto
-   hence "p2 a \<in> set2 c'" unfolding sset_def by auto
-   hence eq: "c = c'" using p2a c c' set2_disj by auto
-   hence "a \<in> sset c" using ac' by simp
-  } note p2_rev = this
-
-  have "\<forall> a \<in> SET. \<exists> c \<in> set_of P. a \<in> sset c" unfolding SET_def by auto
-  then obtain h where h: "\<forall> a \<in> SET. h a \<in> set_of P \<and> a \<in> sset (h a)" by metis
-  have h_u[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
-                      \<Longrightarrow> h (u c b1 b2) = c"
-  by (metis h p2 set2 u(3) u_SET)
-  have h_u1: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
-                      \<Longrightarrow> h (u c b1 b2) = f1 b1"
-  using h unfolding sset_def by auto
-  have h_u2: "\<And> c b1 b2. \<lbrakk>c \<in> set_of P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
-                      \<Longrightarrow> h (u c b1 b2) = f2 b2"
-  using h unfolding sset_def by auto
-  def M \<equiv>
-    "Abs_multiset (\<lambda> a. if a \<in> SET \<and> p1 a \<in> set_of N1 \<and> p2 a \<in> set_of N2 then Ms (h a) a else 0)"
-  have "(\<lambda> a. if a \<in> SET \<and> p1 a \<in> set_of N1 \<and> p2 a \<in> set_of N2 then Ms (h a) a else 0) \<in> multiset"
-    unfolding multiset_def by auto
-  hence [transfer_rule]: "pcr_multiset op = (\<lambda> a. if a \<in> SET \<and> p1 a \<in> set_of N1 \<and> p2 a \<in> set_of N2 then Ms (h a) a else 0) M"
-    unfolding M_def pcr_multiset_def cr_multiset_def by (auto simp: Abs_multiset_inverse)
-  have sM: "set_of M \<subseteq> SET" "set_of M \<subseteq> p1 -` (set_of N1)" "set_of M \<subseteq> p2 -` set_of N2"
-    by (transfer, auto split: split_if_asm)+
-  show "\<exists>M. set_of M \<subseteq> A \<and> mmap p1 M = N1 \<and> mmap p2 M = N2"
-  proof(rule exI[of _ M], safe)
-    fix a assume *: "a \<in> set_of M"
-    from SET_A show "a \<in> A"
-    proof (cases "a \<in> SET")
-      case False thus ?thesis using * by transfer' auto
-    qed blast
-  next
-    show "mmap p1 M = N1"
-    proof(intro multiset_eqI)
-      fix b1
-      let ?K = "{a. p1 a = b1 \<and> a \<in># M}"
-      have "setsum (count M) ?K = count N1 b1"
-      proof(cases "b1 \<in> set_of N1")
-        case False
-        hence "?K = {}" using sM(2) by auto
-        thus ?thesis using False by auto
-      next
-        case True
-        def c \<equiv> "f1 b1"
-        have c: "c \<in> set_of P" and b1: "b1 \<in> set1 c"
-          unfolding set1_def c_def P1 using True by (auto simp: comp_eq_dest[OF set_of_mmap])
-        with sM(1) have "setsum (count M) ?K = setsum (count M) {a. p1 a = b1 \<and> a \<in> SET}"
-          by transfer (force intro: setsum.mono_neutral_cong_left split: split_if_asm)
-        also have "... = setsum (count M) ((\<lambda> b2. u c b1 b2) ` (set2 c))"
-          apply(rule setsum.cong) using c b1 proof safe
-          fix a assume p1a: "p1 a \<in> set1 c" and "c \<in> set_of P" and "a \<in> SET"
-          hence ac: "a \<in> sset c" using p1_rev by auto
-          hence "a = u c (p1 a) (p2 a)" using c by auto
-          moreover have "p2 a \<in> set2 c" using ac c by auto
-          ultimately show "a \<in> u c (p1 a) ` set2 c" by auto
-        qed auto
-        also have "... = setsum (\<lambda> b2. count M (u c b1 b2)) (set2 c)"
-          unfolding comp_def[symmetric] apply(rule setsum.reindex)
-          using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast
-        also have "... = count N1 b1" unfolding ss1[OF c b1, symmetric]
-          apply(rule setsum.cong[OF refl]) apply (transfer fixing: Ms u c b1 set2)
-          using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1]
-            [[hypsubst_thin = true]]
-          by fastforce
-        finally show ?thesis .
-      qed
-      thus "count (mmap p1 M) b1 = count N1 b1" by transfer
-    qed
-  next
-    show "mmap p2 M = N2"
-    proof(intro multiset_eqI)
-      fix b2
-      let ?K = "{a. p2 a = b2 \<and> a \<in># M}"
-      have "setsum (count M) ?K = count N2 b2"
-      proof(cases "b2 \<in> set_of N2")
-        case False
-        hence "?K = {}" using sM(3) by auto
-        thus ?thesis using False by auto
-      next
-        case True
-        def c \<equiv> "f2 b2"
-        have c: "c \<in> set_of P" and b2: "b2 \<in> set2 c"
-          unfolding set2_def c_def P2 using True by (auto simp: comp_eq_dest[OF set_of_mmap])
-        with sM(1) have "setsum (count M) ?K = setsum (count M) {a. p2 a = b2 \<and> a \<in> SET}"
-          by transfer (force intro: setsum.mono_neutral_cong_left split: split_if_asm)
-        also have "... = setsum (count M) ((\<lambda> b1. u c b1 b2) ` (set1 c))"
-          apply(rule setsum.cong) using c b2 proof safe
-          fix a assume p2a: "p2 a \<in> set2 c" and "c \<in> set_of P" and "a \<in> SET"
-          hence ac: "a \<in> sset c" using p2_rev by auto
-          hence "a = u c (p1 a) (p2 a)" using c by auto
-          moreover have "p1 a \<in> set1 c" using ac c by auto
-          ultimately show "a \<in> (\<lambda>x. u c x (p2 a)) ` set1 c" by auto
-        qed auto
-        also have "... = setsum (count M o (\<lambda> b1. u c b1 b2)) (set1 c)"
-          apply(rule setsum.reindex)
-          using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast
-        also have "... = setsum (\<lambda> b1. count M (u c b1 b2)) (set1 c)" by simp
-        also have "... = count N2 b2" unfolding ss2[OF c b2, symmetric] comp_def
-          apply(rule setsum.cong[OF refl]) apply (transfer fixing: Ms u c b2 set1)
-          using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2] set1_def
-            [[hypsubst_thin = true]]
-          by fastforce
-        finally show ?thesis .
-      qed
-      thus "count (mmap p2 M) b2 = count N2 b2" by transfer
-    qed
-  qed
+lemma list_all2_reorder_left_invariance:
+  assumes rel: "list_all2 R xs ys" and ms_x: "multiset_of xs' = multiset_of xs"
+  shows "\<exists>ys'. list_all2 R xs' ys' \<and> multiset_of ys' = multiset_of ys"
+proof -
+  have len: "length xs = length ys"
+    using rel list_all2_conv_all_nth by auto
+  obtain ys' where
+    len': "length xs' = length ys'" and ms_xy: "multiset_of (zip xs' ys') = multiset_of (zip xs ys)"
+    using len ms_x by (metis ex_multiset_of_zip_left)
+  have "list_all2 R xs' ys'"
+    using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: multiset_of_eq_setD)
+  moreover have "multiset_of ys' = multiset_of ys"
+    using len len' ms_xy map_snd_zip multiset_of_map by metis
+  ultimately show ?thesis
+    by blast
 qed
 
-lemma set_of_bd: "(card_of (set_of x), natLeq) \<in> ordLeq"
-  by transfer
-    (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def)
-
-lemma wppull_mmap:
-  assumes "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
-  shows "wppull {M. set_of M \<subseteq> A} {N1. set_of N1 \<subseteq> B1} {N2. set_of N2 \<subseteq> B2}
-    (mmap f1) (mmap f2) (mmap e1) (mmap e2) (mmap p1) (mmap p2)"
-proof -
-  from assms obtain j where j: "\<forall>a'\<in>thePull B1 B2 f1 f2.
-    j a' \<in> A \<and> e1 (p1 (j a')) = e1 (fst a') \<and> e2 (p2 (j a')) = e2 (snd a')" 
-    by (blast dest: wppull_thePull)
-  then show ?thesis
-    by (intro wpull_wppull[OF wpull_mmap[OF wpull_thePull], of _ _ _ _ "mmap j"])
-      (auto simp: comp_eq_dest_lhs[OF mmap_comp[symmetric]] comp_eq_dest[OF set_of_mmap]
-        intro!: mmap_cong simp del: mem_set_of_iff simp: mem_set_of_iff[symmetric])
-qed
+lemma ex_multiset_of: "\<exists>xs. multiset_of xs = X"
+  by (induct X) (simp, metis multiset_of.simps(2))
 
 bnf "'a multiset"
-  map: mmap
+  map: image_mset
   sets: set_of 
   bd: natLeq
   wits: "{#}"
-by (auto simp add: mmap_id0 mmap_comp set_of_mmap natLeq_card_order natLeq_cinfinite set_of_bd
-  Grp_def relcompp.simps intro: mmap_cong)
-  (metis wppull_mmap[OF wppull_fstOp_sndOp, unfolded wppull_def
-    o_eq_dest_lhs[OF mmap_comp[symmetric]] fstOp_def sndOp_def comp_def, simplified])
-
-inductive rel_multiset' where
-  Zero[intro]: "rel_multiset' R {#} {#}"
-| Plus[intro]: "\<lbrakk>R a b; rel_multiset' R M N\<rbrakk> \<Longrightarrow> rel_multiset' R (M + {#a#}) (N + {#b#})"
-
-lemma map_multiset_Zero_iff[simp]: "mmap f M = {#} \<longleftrightarrow> M = {#}"
-by (metis image_is_empty multiset.set_map set_of_eq_empty_iff)
-
-lemma map_multiset_Zero[simp]: "mmap f {#} = {#}" by simp
-
-lemma rel_multiset_Zero: "rel_multiset R {#} {#}"
-unfolding rel_multiset_def Grp_def by auto
+  rel: rel_mset
+proof -
+  show "image_mset id = id"
+    by (rule image_mset.id)
+next
+  show "\<And>f g. image_mset (g \<circ> f) = image_mset g \<circ> image_mset f"
+    unfolding comp_def by (rule ext) (simp add: image_mset.compositionality comp_def)
+next
+  fix X :: "'a multiset"
+  show "\<And>f g. (\<And>z. z \<in> set_of X \<Longrightarrow> f z = g z) \<Longrightarrow> image_mset f X = image_mset g X"
+    by (induct X, (simp (no_asm))+,
+      metis One_nat_def Un_iff count_single mem_set_of_iff set_of_union zero_less_Suc)
+next
+  show "\<And>f. set_of \<circ> image_mset f = op ` f \<circ> set_of"
+    by auto
+next
+  show "card_order natLeq"
+    by (rule natLeq_card_order)
+next
+  show "BNF_Cardinal_Arithmetic.cinfinite natLeq"
+    by (rule natLeq_cinfinite)
+next
+  show "\<And>X. ordLeq3 (card_of (set_of X)) natLeq"
+    by transfer
+      (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def)
+next
+  show "\<And>R S. rel_mset R OO rel_mset S \<le> rel_mset (R OO S)"
+    unfolding rel_mset_def[abs_def] OO_def
+    apply clarify
+    apply (rename_tac X Z Y xs ys' ys zs)
+    apply (drule_tac xs = ys' and ys = zs and xs' = ys in list_all2_reorder_left_invariance)
+    by (auto intro: list_all2_trans)
+next
+  show "\<And>R. rel_mset R =
+    (BNF_Def.Grp {x. set_of x \<subseteq> {(x, y). R x y}} (image_mset fst))\<inverse>\<inverse> OO
+    BNF_Def.Grp {x. set_of x \<subseteq> {(x, y). R x y}} (image_mset snd)"
+    unfolding rel_mset_def[abs_def] BNF_Def.Grp_def OO_def
+    apply (rule ext)+
+    apply auto
+     apply (rule_tac x = "multiset_of (zip xs ys)" in exI)
+     apply auto[1]
+        apply (metis list_all2_lengthD map_fst_zip multiset_of_map)
+       apply (auto simp: list_all2_iff)[1]
+      apply (metis list_all2_lengthD map_snd_zip multiset_of_map)
+     apply (auto simp: list_all2_iff)[1]
+    apply (rename_tac XY)
+    apply (cut_tac X = XY in ex_multiset_of)
+    apply (erule exE)
+    apply (rename_tac xys)
+    apply (rule_tac x = "map fst xys" in exI)
+    apply (auto simp: multiset_of_map)
+    apply (rule_tac x = "map snd xys" in exI)
+    by (auto simp: multiset_of_map list_all2I subset_eq zip_map_fst_snd)
+next
+  show "\<And>z. z \<in> set_of {#} \<Longrightarrow> False"
+    by auto
+qed
+
+inductive rel_mset' where
+  Zero[intro]: "rel_mset' R {#} {#}"
+| Plus[intro]: "\<lbrakk>R a b; rel_mset' R M N\<rbrakk> \<Longrightarrow> rel_mset' R (M + {#a#}) (N + {#b#})"
+
+lemma rel_mset_Zero: "rel_mset R {#} {#}"
+unfolding rel_mset_def Grp_def by auto
 
 declare multiset.count[simp]
 declare Abs_multiset_inverse[simp]
 declare multiset.count_inverse[simp]
 declare union_preserves_multiset[simp]
 
-lemma map_multiset_Plus[simp]: "mmap f (M1 + M2) = mmap f M1 + mmap f M2"
-proof (intro multiset_eqI, transfer fixing: f)
-  fix x :: 'a and M1 M2 :: "'b \<Rightarrow> nat"
-  assume "M1 \<in> multiset" "M2 \<in> multiset"
-  hence "setsum M1 {a. f a = x \<and> 0 < M1 a} = setsum M1 {a. f a = x \<and> 0 < M1 a + M2 a}"
-        "setsum M2 {a. f a = x \<and> 0 < M2 a} = setsum M2 {a. f a = x \<and> 0 < M1 a + M2 a}"
-    by (auto simp: multiset_def intro!: setsum.mono_neutral_cong_left)
-  then show "(\<Sum>a | f a = x \<and> 0 < M1 a + M2 a. M1 a + M2 a) =
-       setsum M1 {a. f a = x \<and> 0 < M1 a} +
-       setsum M2 {a. f a = x \<and> 0 < M2 a}"
-    by (auto simp: setsum.distrib[symmetric])
-qed
-
-lemma map_multiset_single[simp]: "mmap f {#a#} = {#f a#}"
-  by transfer auto
-
-lemma rel_multiset_Plus:
-assumes ab: "R a b" and MN: "rel_multiset R M N"
-shows "rel_multiset R (M + {#a#}) (N + {#b#})"
+lemma rel_mset_Plus:
+assumes ab: "R a b" and MN: "rel_mset R M N"
+shows "rel_mset R (M + {#a#}) (N + {#b#})"
 proof-
   {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
-   hence "\<exists>ya. mmap fst y + {#a#} = mmap fst ya \<and>
-               mmap snd y + {#b#} = mmap snd ya \<and>
+   hence "\<exists>ya. image_mset fst y + {#a#} = image_mset fst ya \<and>
+               image_mset snd y + {#b#} = image_mset snd ya \<and>
                set_of ya \<subseteq> {(x, y). R x y}"
    apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
   }
   thus ?thesis
   using assms
-  unfolding rel_multiset_def Grp_def by force
+  unfolding multiset.rel_compp_Grp Grp_def by blast
 qed
 
-lemma rel_multiset'_imp_rel_multiset:
-"rel_multiset' R M N \<Longrightarrow> rel_multiset R M N"
-apply(induct rule: rel_multiset'.induct)
-using rel_multiset_Zero rel_multiset_Plus by auto
-
-lemma mcard_mmap[simp]: "mcard (mmap f M) = mcard M"
-proof -
-  def A \<equiv> "\<lambda> b. {a. f a = b \<and> a \<in># M}"
-  let ?B = "{b. 0 < setsum (count M) (A b)}"
-  have "{b. \<exists>a. f a = b \<and> a \<in># M} \<subseteq> f ` {a. a \<in># M}" by auto
-  moreover have "finite (f ` {a. a \<in># M})" apply(rule finite_imageI)
-  using finite_Collect_mem .
-  ultimately have fin: "finite {b. \<exists>a. f a = b \<and> a \<in># M}" by(rule finite_subset)
-  have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp
-    by (metis (lifting, full_types) mem_Collect_eq neq0_conv setsum.neutral)
-  have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)"
-  apply safe
-    apply (metis less_not_refl setsum_gt_0_iff setsum.infinite)
-    by (metis A_def finite_Collect_conjI finite_Collect_mem setsum_gt_0_iff)
-  hence AB: "A ` ?B = {A b | b. \<exists> a \<in> A b. count M a > 0}" by auto
-
-  have "setsum (\<lambda> x. setsum (count M) (A x)) ?B = setsum (setsum (count M) o A) ?B"
-  unfolding comp_def ..
-  also have "... = (\<Sum>x\<in> A ` ?B. setsum (count M) x)"
-  unfolding setsum.reindex [OF i, symmetric] ..
-  also have "... = setsum (count M) (\<Union>x\<in>A ` {b. 0 < setsum (count M) (A b)}. x)"
-  (is "_ = setsum (count M) ?J")
-  apply(rule setsum.UNION_disjoint[symmetric])
-  using 0 fin unfolding A_def by auto
-  also have "?J = {a. a \<in># M}" unfolding AB unfolding A_def by auto
-  finally have "setsum (\<lambda> x. setsum (count M) (A x)) ?B =
-                setsum (count M) {a. a \<in># M}" .
-  then show ?thesis unfolding mcard_unfold_setsum A_def by transfer
-qed
-
-lemma rel_multiset_mcard:
-assumes "rel_multiset R M N"
-shows "mcard M = mcard N"
-using assms unfolding rel_multiset_def Grp_def by auto
+lemma rel_mset'_imp_rel_mset:
+"rel_mset' R M N \<Longrightarrow> rel_mset R M N"
+apply(induct rule: rel_mset'.induct)
+using rel_mset_Zero rel_mset_Plus by auto
+
+lemma mcard_image_mset[simp]: "mcard (image_mset f M) = mcard M"
+  unfolding size_eq_mcard[symmetric] by (rule size_image_mset)
+
+lemma rel_mset_mcard:
+  assumes "rel_mset R M N"
+  shows "mcard M = mcard N"
+using assms unfolding multiset.rel_compp_Grp Grp_def by auto
 
 lemma multiset_induct2[case_names empty addL addR]:
 assumes empty: "P {#} {#}"
@@ -2946,100 +2485,96 @@
 qed
 
 lemma msed_map_invL:
-assumes "mmap f (M + {#a#}) = N"
-shows "\<exists> N1. N = N1 + {#f a#} \<and> mmap f M = N1"
+assumes "image_mset f (M + {#a#}) = N"
+shows "\<exists>N1. N = N1 + {#f a#} \<and> image_mset f M = N1"
 proof-
   have "f a \<in># N"
   using assms multiset.set_map[of f "M + {#a#}"] by auto
   then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
-  have "mmap f M = N1" using assms unfolding N by simp
+  have "image_mset f M = N1" using assms unfolding N by simp
   thus ?thesis using N by blast
 qed
 
 lemma msed_map_invR:
-assumes "mmap f M = N + {#b#}"
-shows "\<exists> M1 a. M = M1 + {#a#} \<and> f a = b \<and> mmap f M1 = N"
+assumes "image_mset f M = N + {#b#}"
+shows "\<exists>M1 a. M = M1 + {#a#} \<and> f a = b \<and> image_mset f M1 = N"
 proof-
   obtain a where a: "a \<in># M" and fa: "f a = b"
   using multiset.set_map[of f M] unfolding assms
   by (metis image_iff mem_set_of_iff union_single_eq_member)
   then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
-  have "mmap f M1 = N" using assms unfolding M fa[symmetric] by simp
+  have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp
   thus ?thesis using M fa by blast
 qed
 
 lemma msed_rel_invL:
-assumes "rel_multiset R (M + {#a#}) N"
-shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> rel_multiset R M N1"
+assumes "rel_mset R (M + {#a#}) N"
+shows "\<exists>N1 b. N = N1 + {#b#} \<and> R a b \<and> rel_mset R M N1"
 proof-
-  obtain K where KM: "mmap fst K = M + {#a#}"
-  and KN: "mmap snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
+  obtain K where KM: "image_mset fst K = M + {#a#}"
+  and KN: "image_mset snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
   using assms
-  unfolding rel_multiset_def Grp_def by auto
+  unfolding multiset.rel_compp_Grp Grp_def by auto
   obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
-  and K1M: "mmap fst K1 = M" using msed_map_invR[OF KM] by auto
-  obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "mmap snd K1 = N1"
+  and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto
+  obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "image_mset snd K1 = N1"
   using msed_map_invL[OF KN[unfolded K]] by auto
   have Rab: "R a (snd ab)" using sK a unfolding K by auto
-  have "rel_multiset R M N1" using sK K1M K1N1
-  unfolding K rel_multiset_def Grp_def by auto
+  have "rel_mset R M N1" using sK K1M K1N1
+  unfolding K multiset.rel_compp_Grp Grp_def by auto
   thus ?thesis using N Rab by auto
 qed
 
 lemma msed_rel_invR:
-assumes "rel_multiset R M (N + {#b#})"
-shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> rel_multiset R M1 N"
+assumes "rel_mset R M (N + {#b#})"
+shows "\<exists>M1 a. M = M1 + {#a#} \<and> R a b \<and> rel_mset R M1 N"
 proof-
-  obtain K where KN: "mmap snd K = N + {#b#}"
-  and KM: "mmap fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
+  obtain K where KN: "image_mset snd K = N + {#b#}"
+  and KM: "image_mset fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
   using assms
-  unfolding rel_multiset_def Grp_def by auto
+  unfolding multiset.rel_compp_Grp Grp_def by auto
   obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
-  and K1N: "mmap snd K1 = N" using msed_map_invR[OF KN] by auto
-  obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "mmap fst K1 = M1"
+  and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto
+  obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "image_mset fst K1 = M1"
   using msed_map_invL[OF KM[unfolded K]] by auto
   have Rab: "R (fst ab) b" using sK b unfolding K by auto
-  have "rel_multiset R M1 N" using sK K1N K1M1
-  unfolding K rel_multiset_def Grp_def by auto
+  have "rel_mset R M1 N" using sK K1N K1M1
+  unfolding K multiset.rel_compp_Grp Grp_def by auto
   thus ?thesis using M Rab by auto
 qed
 
-lemma rel_multiset_imp_rel_multiset':
-assumes "rel_multiset R M N"
-shows "rel_multiset' R M N"
+lemma rel_mset_imp_rel_mset':
+assumes "rel_mset R M N"
+shows "rel_mset' R M N"
 using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
   case (less M)
-  have c: "mcard M = mcard N" using rel_multiset_mcard[OF less.prems] .
+  have c: "mcard M = mcard N" using rel_mset_mcard[OF less.prems] .
   show ?case
   proof(cases "M = {#}")
     case True hence "N = {#}" using c by simp
-    thus ?thesis using True rel_multiset'.Zero by auto
+    thus ?thesis using True rel_mset'.Zero by auto
   next
     case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
-    obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_multiset R M1 N1"
+    obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_mset R M1 N1"
     using msed_rel_invL[OF less.prems[unfolded M]] by auto
-    have "rel_multiset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
-    thus ?thesis using rel_multiset'.Plus[of R a b, OF R] unfolding M N by simp
+    have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
+    thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp
   qed
 qed
 
-lemma rel_multiset_rel_multiset':
-"rel_multiset R M N = rel_multiset' R M N"
-using  rel_multiset_imp_rel_multiset' rel_multiset'_imp_rel_multiset by auto
-
-(* The main end product for rel_multiset: inductive characterization *)
-theorems rel_multiset_induct[case_names empty add, induct pred: rel_multiset] =
-         rel_multiset'.induct[unfolded rel_multiset_rel_multiset'[symmetric]]
+lemma rel_mset_rel_mset':
+"rel_mset R M N = rel_mset' R M N"
+using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto
+
+(* The main end product for rel_mset: inductive characterization *)
+theorems rel_mset_induct[case_names empty add, induct pred: rel_mset] =
+         rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]]
 
 
 subsection {* Size setup *}
 
-lemma multiset_size_o_map: "size_multiset g \<circ> mmap f = size_multiset (g \<circ> f)"
-apply (rule ext)
-apply (unfold o_apply)
-apply (induct_tac x)
-apply auto
-done
+lemma multiset_size_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)"
+  unfolding o_apply by (rule ext) (induct_tac, auto)
 
 setup {*
 BNF_LFP_Size.register_size_global @{type_name multiset} @{const_name size_multiset}
--- a/src/HOL/Library/Permutation.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/Permutation.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -162,7 +162,7 @@
   apply (case_tac "remdups xs")
    apply simp_all
   apply (subgoal_tac "a \<in> set (remdups ys)")
-   prefer 2 apply (metis set_simps(2) insert_iff set_remdups)
+   prefer 2 apply (metis list.set(2) insert_iff set_remdups)
   apply (drule split_list) apply (elim exE conjE)
   apply (drule_tac x = list in spec) apply (erule impE) prefer 2
    apply (drule_tac x = "ysa @ zs" in spec) apply (erule impE) prefer 2
--- a/src/HOL/Library/Quickcheck_Types.thy	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,469 +0,0 @@
-(*  Title:      HOL/Library/Quickcheck_Types.thy
-    Author:     Lukas Bulwahn
-    Copyright   2010 TU Muenchen
-*)
-
-theory Quickcheck_Types
-imports Main
-begin
-
-text {*
-This theory provides some default types for the quickcheck execution.
-In most cases, the default type @{typ "int"} meets the sort constraints
-of the proposition.
-But for the type classes bot and top, the type @{typ "int"} is insufficient.
-Hence, we provide other types than @{typ "int"} as further default types.  
-*}
-
-subsection {* A non-distributive lattice *}
-
-datatype non_distrib_lattice = Zero | A | B | C | One
-
-instantiation non_distrib_lattice :: order
-begin
-
-definition less_eq_non_distrib_lattice
-where
-  "a <= b = (case a of Zero => True | A => (b = A) \<or> (b = One) | B => (b = B) \<or> (b = One) | C => (b = C) \<or> (b = One) | One => (b = One))"
-
-definition less_non_distrib_lattice
-where
-  "a < b = (case a of Zero => (b \<noteq> Zero) | A => (b = One) | B => (b = One) | C => (b = One) | One => False)"
-
-instance proof
-qed (auto simp add: less_eq_non_distrib_lattice_def less_non_distrib_lattice_def split: non_distrib_lattice.split non_distrib_lattice.split_asm)
-
-end
-
-instantiation non_distrib_lattice :: lattice
-begin
-
-
-definition sup_non_distrib_lattice
-where
-  "sup a b = (if a = b then a else (if a = Zero then b else (if b = Zero then a else One)))"
-
-definition inf_non_distrib_lattice
-where
-  "inf a b = (if a = b then a else (if a = One then b else (if b = One then a else Zero)))"
-
-instance proof
-qed (auto simp add: inf_non_distrib_lattice_def sup_non_distrib_lattice_def less_eq_non_distrib_lattice_def split: split_if non_distrib_lattice.split non_distrib_lattice.split_asm)
-
-end
-
-hide_const Zero A B C One
-
-subsection {* Values extended by a bottom element *}
-
-datatype 'a bot = Value 'a | Bot
-
-instantiation bot :: (preorder) preorder
-begin
-
-definition less_eq_bot where
-  "x \<le> y \<longleftrightarrow> (case x of Bot \<Rightarrow> True | Value x \<Rightarrow> (case y of Bot \<Rightarrow> False | Value y \<Rightarrow> x \<le> y))"
-
-definition less_bot where
-  "x < y \<longleftrightarrow> (case y of Bot \<Rightarrow> False | Value y \<Rightarrow> (case x of Bot \<Rightarrow> True | Value x \<Rightarrow> x < y))"
-
-lemma less_eq_bot_Bot [simp]: "Bot \<le> x"
-  by (simp add: less_eq_bot_def)
-
-lemma less_eq_bot_Bot_code [code]: "Bot \<le> x \<longleftrightarrow> True"
-  by simp
-
-lemma less_eq_bot_Bot_is_Bot: "x \<le> Bot \<Longrightarrow> x = Bot"
-  by (cases x) (simp_all add: less_eq_bot_def)
-
-lemma less_eq_bot_Value_Bot [simp, code]: "Value x \<le> Bot \<longleftrightarrow> False"
-  by (simp add: less_eq_bot_def)
-
-lemma less_eq_bot_Value [simp, code]: "Value x \<le> Value y \<longleftrightarrow> x \<le> y"
-  by (simp add: less_eq_bot_def)
-
-lemma less_bot_Bot [simp, code]: "x < Bot \<longleftrightarrow> False"
-  by (simp add: less_bot_def)
-
-lemma less_bot_Bot_is_Value: "Bot < x \<Longrightarrow> \<exists>z. x = Value z"
-  by (cases x) (simp_all add: less_bot_def)
-
-lemma less_bot_Bot_Value [simp]: "Bot < Value x"
-  by (simp add: less_bot_def)
-
-lemma less_bot_Bot_Value_code [code]: "Bot < Value x \<longleftrightarrow> True"
-  by simp
-
-lemma less_bot_Value [simp, code]: "Value x < Value y \<longleftrightarrow> x < y"
-  by (simp add: less_bot_def)
-
-instance proof
-qed (auto simp add: less_eq_bot_def less_bot_def less_le_not_le elim: order_trans split: bot.splits)
-
-end 
-
-instance bot :: (order) order proof
-qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits)
-
-instance bot :: (linorder) linorder proof
-qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits)
-
-instantiation bot :: (order) bot
-begin
-
-definition "bot = Bot"
-
-instance ..
-
-end
-
-instantiation bot :: (top) top
-begin
-
-definition "top = Value top"
-
-instance ..
-
-end
-
-instantiation bot :: (semilattice_inf) semilattice_inf
-begin
-
-definition inf_bot
-where
-  "inf x y = (case x of Bot => Bot | Value v => (case y of Bot => Bot | Value v' => Value (inf v v')))"
-
-instance proof
-qed (auto simp add: inf_bot_def less_eq_bot_def split: bot.splits)
-
-end
-
-instantiation bot :: (semilattice_sup) semilattice_sup
-begin
-
-definition sup_bot
-where
-  "sup x y = (case x of Bot => y | Value v => (case y of Bot => x | Value v' => Value (sup v v')))"
-
-instance proof
-qed (auto simp add: sup_bot_def less_eq_bot_def split: bot.splits)
-
-end
-
-instance bot :: (lattice) bounded_lattice_bot
-by(intro_classes)(simp add: bot_bot_def)
-
-section {* Values extended by a top element *}
-
-datatype 'a top = Value 'a | Top
-
-instantiation top :: (preorder) preorder
-begin
-
-definition less_eq_top where
-  "x \<le> y \<longleftrightarrow> (case y of Top \<Rightarrow> True | Value y \<Rightarrow> (case x of Top \<Rightarrow> False | Value x \<Rightarrow> x \<le> y))"
-
-definition less_top where
-  "x < y \<longleftrightarrow> (case x of Top \<Rightarrow> False | Value x \<Rightarrow> (case y of Top \<Rightarrow> True | Value y \<Rightarrow> x < y))"
-
-lemma less_eq_top_Top [simp]: "x <= Top"
-  by (simp add: less_eq_top_def)
-
-lemma less_eq_top_Top_code [code]: "x \<le> Top \<longleftrightarrow> True"
-  by simp
-
-lemma less_eq_top_is_Top: "Top \<le> x \<Longrightarrow> x = Top"
-  by (cases x) (simp_all add: less_eq_top_def)
-
-lemma less_eq_top_Top_Value [simp, code]: "Top \<le> Value x \<longleftrightarrow> False"
-  by (simp add: less_eq_top_def)
-
-lemma less_eq_top_Value_Value [simp, code]: "Value x \<le> Value y \<longleftrightarrow> x \<le> y"
-  by (simp add: less_eq_top_def)
-
-lemma less_top_Top [simp, code]: "Top < x \<longleftrightarrow> False"
-  by (simp add: less_top_def)
-
-lemma less_top_Top_is_Value: "x < Top \<Longrightarrow> \<exists>z. x = Value z"
-  by (cases x) (simp_all add: less_top_def)
-
-lemma less_top_Value_Top [simp]: "Value x < Top"
-  by (simp add: less_top_def)
-
-lemma less_top_Value_Top_code [code]: "Value x < Top \<longleftrightarrow> True"
-  by simp
-
-lemma less_top_Value [simp, code]: "Value x < Value y \<longleftrightarrow> x < y"
-  by (simp add: less_top_def)
-
-instance proof
-qed (auto simp add: less_eq_top_def less_top_def less_le_not_le elim: order_trans split: top.splits)
-
-end 
-
-instance top :: (order) order proof
-qed (auto simp add: less_eq_top_def less_top_def split: top.splits)
-
-instance top :: (linorder) linorder proof
-qed (auto simp add: less_eq_top_def less_top_def split: top.splits)
-
-instantiation top :: (order) top
-begin
-
-definition "top = Top"
-
-instance ..
-
-end
-
-instantiation top :: (bot) bot
-begin
-
-definition "bot = Value bot"
-
-instance ..
-
-end
-
-instantiation top :: (semilattice_inf) semilattice_inf
-begin
-
-definition inf_top
-where
-  "inf x y = (case x of Top => y | Value v => (case y of Top => x | Value v' => Value (inf v v')))"
-
-instance proof
-qed (auto simp add: inf_top_def less_eq_top_def split: top.splits)
-
-end
-
-instantiation top :: (semilattice_sup) semilattice_sup
-begin
-
-definition sup_top
-where
-  "sup x y = (case x of Top => Top | Value v => (case y of Top => Top | Value v' => Value (sup v v')))"
-
-instance proof
-qed (auto simp add: sup_top_def less_eq_top_def split: top.splits)
-
-end
-
-instance top :: (lattice) bounded_lattice_top
-by(intro_classes)(simp add: top_top_def)
-
-
-datatype 'a flat_complete_lattice = Value 'a | Bot | Top
-
-instantiation flat_complete_lattice :: (type) order
-begin
-
-definition less_eq_flat_complete_lattice where
-  "x \<le> y == (case x of Bot => True | Value v1 => (case y of Bot => False | Value v2 => (v1 = v2) | Top => True) | Top => (y = Top))"
-
-definition less_flat_complete_lattice where
-  "x < y = (case x of Bot => \<not> (y = Bot) | Value v1 => (y = Top) | Top => False)"
-
-lemma [simp]: "Bot <= y"
-unfolding less_eq_flat_complete_lattice_def by auto
-
-lemma [simp]: "y <= Top"
-unfolding less_eq_flat_complete_lattice_def by (auto split: flat_complete_lattice.splits)
-
-lemma greater_than_two_values:
-  assumes "a ~= aa" "Value a <= z" "Value aa <= z"
-  shows "z = Top"
-using assms
-by (cases z) (auto simp add: less_eq_flat_complete_lattice_def)
-
-lemma lesser_than_two_values:
-  assumes "a ~= aa" "z <= Value a" "z <= Value aa"
-  shows "z = Bot"
-using assms
-by (cases z) (auto simp add: less_eq_flat_complete_lattice_def)
-
-instance proof
-qed (auto simp add: less_eq_flat_complete_lattice_def less_flat_complete_lattice_def split: flat_complete_lattice.splits)
-
-end
-
-instantiation flat_complete_lattice :: (type) bot
-begin
-
-definition "bot = Bot"
-
-instance ..
-
-end
-
-instantiation flat_complete_lattice :: (type) top
-begin
-
-definition "top = Top"
-
-instance ..
-
-end
-
-instantiation flat_complete_lattice :: (type) lattice
-begin
-
-definition inf_flat_complete_lattice
-where
-  "inf x y = (case x of Bot => Bot | Value v1 => (case y of Bot => Bot | Value v2 => if (v1 = v2) then x else Bot | Top => x) | Top => y)"
-
-definition sup_flat_complete_lattice
-where
-  "sup x y = (case x of Bot => y | Value v1 => (case y of Bot => x | Value v2 => if v1 = v2 then x else Top | Top => Top) | Top => Top)"
-
-instance proof
-qed (auto simp add: inf_flat_complete_lattice_def sup_flat_complete_lattice_def less_eq_flat_complete_lattice_def split: flat_complete_lattice.splits)
-
-end
-
-instantiation flat_complete_lattice :: (type) complete_lattice
-begin
-
-definition Sup_flat_complete_lattice
-where
-  "Sup A = (if (A = {} \<or> A = {Bot}) then Bot else (if (\<exists> v. A - {Bot} = {Value v}) then Value (THE v. A - {Bot} = {Value v}) else Top))"
-
-definition Inf_flat_complete_lattice
-where
-  "Inf A = (if (A = {} \<or> A = {Top}) then Top else (if (\<exists> v. A - {Top} = {Value v}) then Value (THE v. A - {Top} = {Value v}) else Bot))"
- 
-instance
-proof
-  fix x A
-  assume "(x :: 'a flat_complete_lattice) : A"
-  {
-    fix v
-    assume "A - {Top} = {Value v}"
-    from this have "(THE v. A - {Top} = {Value v}) = v"
-      by (auto intro!: the1_equality)
-    moreover
-    from `x : A` `A - {Top} = {Value v}` have "x = Top \<or> x = Value v"
-      by auto
-    ultimately have "Value (THE v. A - {Top} = {Value v}) <= x"
-      by auto
-  }
-  from `x : A` this show "Inf A <= x"
-    unfolding Inf_flat_complete_lattice_def
-    by fastforce
-next
-  fix z A
-  assume z: "\<And>x. x : A ==> z <= (x :: 'a flat_complete_lattice)"
-  {
-    fix v
-    assume "A - {Top} = {Value v}"
-    moreover
-    from this have "(THE v. A - {Top} = {Value v}) = v"
-      by (auto intro!: the1_equality)
-    moreover
-    note z
-    moreover
-    ultimately have "z <= Value (THE v::'a. A - {Top} = {Value v})"
-      by auto
-  } moreover
-  {
-    assume not_one_value: "A ~= {}" "A ~= {Top}" "~ (EX v::'a. A - {Top} = {Value v})"
-    have "z <= Bot"
-    proof (cases "A - {Top} = {Bot}")
-      case True
-      from this z show ?thesis
-        by auto
-    next
-      case False
-      from not_one_value
-      obtain a1 where a1: "a1 : A - {Top}" by auto
-      from not_one_value False a1
-      obtain a2 where "a2 : A - {Top} \<and> a1 \<noteq> a2"
-        by (cases a1) auto
-      from this a1 z[of "a1"] z[of "a2"] show ?thesis
-        apply (cases a1)
-        apply auto
-        apply (cases a2)
-        apply auto
-        apply (auto dest!: lesser_than_two_values)
-        done
-    qed
-  } moreover
-  note z moreover
-  ultimately show "z <= Inf A"
-    unfolding Inf_flat_complete_lattice_def
-    by auto
-next
-  fix x A
-  assume "(x :: 'a flat_complete_lattice) : A"
-  {
-    fix v
-    assume "A - {Bot} = {Value v}"
-    from this have "(THE v. A - {Bot} = {Value v}) = v"
-      by (auto intro!: the1_equality)
-    moreover
-    from `x : A` `A - {Bot} = {Value v}` have "x = Bot \<or> x = Value v"
-      by auto
-    ultimately have "x <= Value (THE v. A - {Bot} = {Value v})"
-      by auto
-  }
-  from `x : A` this show "x <= Sup A"
-    unfolding Sup_flat_complete_lattice_def
-    by fastforce
-next
-  fix z A
-  assume z: "\<And>x. x : A ==> x <= (z :: 'a flat_complete_lattice)"
-  {
-    fix v
-    assume "A - {Bot} = {Value v}"
-    moreover
-    from this have "(THE v. A - {Bot} = {Value v}) = v"
-      by (auto intro!: the1_equality)
-    moreover
-    note z
-    moreover
-    ultimately have "Value (THE v::'a. A - {Bot} = {Value v}) <= z"
-      by auto
-  } moreover
-  {
-    assume not_one_value: "A ~= {}" "A ~= {Bot}" "~ (EX v::'a. A - {Bot} = {Value v})"
-    have "Top <= z"
-    proof (cases "A - {Bot} = {Top}")
-      case True
-      from this z show ?thesis
-        by auto
-    next
-      case False
-      from not_one_value
-      obtain a1 where a1: "a1 : A - {Bot}" by auto
-      from not_one_value False a1
-      obtain a2 where "a2 : A - {Bot} \<and> a1 \<noteq> a2"
-        by (cases a1) auto
-      from this a1 z[of "a1"] z[of "a2"] show ?thesis
-        apply (cases a1)
-        apply auto
-        apply (cases a2)
-        apply (auto dest!: greater_than_two_values)
-        done
-    qed
-  } moreover
-  note z moreover
-  ultimately show "Sup A <= z"
-    unfolding Sup_flat_complete_lattice_def
-    by auto
-next
-  show "Inf {} = (top :: 'a flat_complete_lattice)"
-    by(simp add: Inf_flat_complete_lattice_def top_flat_complete_lattice_def)
-next
-  show "Sup {} = (bot :: 'a flat_complete_lattice)"
-    by(simp add: Sup_flat_complete_lattice_def bot_flat_complete_lattice_def)
-qed
-
-end
-
-section {* Quickcheck configuration *}
-
-quickcheck_params[finite_types = false, default_type = ["int", "non_distrib_lattice", "int bot", "int top", "Enum.finite_4 flat_complete_lattice"]]
-
-hide_type non_distrib_lattice flat_complete_lattice bot top
-
-end
\ No newline at end of file
--- a/src/HOL/Library/RBT_Impl.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/RBT_Impl.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1754,7 +1754,7 @@
   Abs_compare_cases Abs_compare_induct Abs_compare_inject Abs_compare_inverse
   Rep_compare Rep_compare_cases Rep_compare_induct Rep_compare_inject Rep_compare_inverse
   compare.simps compare.exhaust compare.induct compare.rec compare.simps
-  compare.size compare.case_cong compare.weak_case_cong compare.case
+  compare.size compare.case_cong compare.case_cong_weak compare.case
   compare.nchotomy compare.split compare.split_asm rec_compare_def
   compare.eq.refl compare.eq.simps
   compare.EQ_def compare.GT_def compare.LT_def
--- a/src/HOL/Library/RBT_Set.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/RBT_Set.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -522,7 +522,7 @@
 
 code_datatype Set Coset
 
-declare set_simps[code]
+declare list.set[code] (* needed? *)
 
 lemma empty_Set [code]:
   "Set.empty = Set RBT.empty"
--- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -1048,7 +1048,7 @@
 
 fun sos_tac print_cert prover ctxt =
   (* The SOS prover breaks if mult_nonneg_nonneg is in the simpset *)
-  let val ctxt' = ctxt delsimps [@{thm mult_nonneg_nonneg}]
+  let val ctxt' = Context_Position.set_visible false ctxt delsimps @{thms mult_nonneg_nonneg}
   in Object_Logic.full_atomize_tac ctxt' THEN'
      elim_denom_tac ctxt' THEN'
      core_sos_tac print_cert prover ctxt'
--- a/src/HOL/Library/refute.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/refute.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -2909,7 +2909,7 @@
           Node xs => xs
         | _       => raise REFUTE ("set_printer",
           "interpretation for set type is a leaf"))
-      val elements = List.mapPartial (fn (arg, result) =>
+      val elements = map_filter (fn (arg, result) =>
         case result of
           Leaf [fmTrue, (* fmFalse *) _] =>
           if Prop_Logic.eval assignment fmTrue then
--- a/src/HOL/Library/simps_case_conv.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Library/simps_case_conv.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -221,15 +221,15 @@
 val _ =
   Outer_Syntax.local_theory @{command_spec "case_of_simps"}
     "turn a list of equations into a case expression"
-    (Parse_Spec.opt_thm_name ":"  -- Parse_Spec.xthms1 >> case_of_simps_cmd)
+    (Parse_Spec.opt_thm_name ":"  -- Parse.xthms1 >> case_of_simps_cmd)
 
 val parse_splits = @{keyword "("} |-- Parse.reserved "splits" |-- @{keyword ":"} |--
-  Parse_Spec.xthms1 --| @{keyword ")"}
+  Parse.xthms1 --| @{keyword ")"}
 
 val _ =
   Outer_Syntax.local_theory @{command_spec "simps_of_case"}
     "perform case split on rule"
-    (Parse_Spec.opt_thm_name ":"  -- Parse_Spec.xthm --
+    (Parse_Spec.opt_thm_name ":"  -- Parse.xthm --
       Scan.optional parse_splits [] >> simps_of_case_cmd)
 
 end
--- a/src/HOL/Lifting.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Lifting.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -545,6 +545,8 @@
 
 ML_file "Tools/Lifting/lifting_util.ML"
 
+named_theorems relator_eq_onp
+  "theorems that a relator of an eq_onp is an eq_onp of the corresponding predicate"
 ML_file "Tools/Lifting/lifting_info.ML"
 setup Lifting_Info.setup
 
--- a/src/HOL/List.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/List.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -39,6 +39,8 @@
 
 setup {* Sign.parent_path *}
 
+lemmas set_simps = list.set (* legacy *)
+
 syntax
   -- {* list Enumeration *}
   "_list" :: "args => 'a list"    ("[(_)]")
@@ -54,16 +56,9 @@
 "last (x # xs) = (if xs = [] then x else last xs)"
 
 primrec butlast :: "'a list \<Rightarrow> 'a list" where
-"butlast []= []" |
+"butlast [] = []" |
 "butlast (x # xs) = (if xs = [] then [] else x # butlast xs)"
 
-declare list.set[simp del, code del]
-
-lemma set_simps[simp, code, code_post]:
-  "set [] = {}"
-  "set (x # xs) = insert x (set xs)"
-by (simp_all add: list.set)
-
 lemma set_rec: "set xs = rec_list {} (\<lambda>x _. insert x) xs"
   by (induct xs) auto
 
@@ -575,7 +570,7 @@
 
 fun simproc ctxt redex =
   let
-    val set_Nil_I = @{thm trans} OF [@{thm set_simps(1)}, @{thm empty_def}]
+    val set_Nil_I = @{thm trans} OF [@{thm list.set(1)}, @{thm empty_def}]
     val set_singleton = @{lemma "set [a] = {x. x = a}" by simp}
     val inst_Collect_mem_eq = @{lemma "set A = {x. x : set A}" by simp}
     val del_refl_eq = @{lemma "(t = t & P) == P" by simp}
@@ -1255,6 +1250,8 @@
 
 subsubsection {* @{const set} *}
 
+declare list.set[code_post]  --"pretty output"
+
 lemma finite_set [iff]: "finite (set xs)"
 by (induct xs) auto
 
@@ -1404,7 +1401,7 @@
 
 
 lemma finite_list: "finite A ==> EX xs. set xs = A"
-  by (erule finite_induct) (auto simp add: set_simps(2) [symmetric] simp del: set_simps(2))
+  by (erule finite_induct) (auto simp add: list.set(2)[symmetric] simp del: list.set(2))
 
 lemma card_length: "card (set xs) \<le> length xs"
 by (induct xs) (auto simp add: card_insert_if)
@@ -3439,6 +3436,9 @@
 "distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
 by force
 
+lemma hd_remdups_adj[simp]: "hd (remdups_adj xs) = hd xs"
+  by (induction xs rule: remdups_adj.induct) simp_all
+
 lemma remdups_adj_Cons: "remdups_adj (x # xs) =
   (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
   by (induct xs arbitrary: x) (auto split: list.splits)
@@ -3447,6 +3447,13 @@
   "remdups_adj (xs @ [x,y]) = remdups_adj (xs @ [x]) @ (if x = y then [] else [y])"
   by (induct xs rule: remdups_adj.induct, simp_all)
 
+lemma remdups_adj_adjacent:
+  "Suc i < length (remdups_adj xs) \<Longrightarrow> remdups_adj xs ! i \<noteq> remdups_adj xs ! Suc i"
+proof (induction xs arbitrary: i rule: remdups_adj.induct)
+  case (3 x y xs i)
+  thus ?case by (cases i, cases "x = y") (simp, auto simp: hd_conv_nth[symmetric])
+qed simp_all
+
 lemma remdups_adj_rev[simp]: "remdups_adj (rev xs) = rev (remdups_adj xs)"
   by (induct xs rule: remdups_adj.induct, simp_all add: remdups_adj_append_two)
 
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1392,7 +1392,7 @@
 
   apply (simp (no_asm_simp) add: max_ssize_def del: max_of_list_append)
     apply (rule max_of_list_sublist)
-    apply (simp (no_asm_simp) only: set_append set_simps list.map) apply blast
+    apply (simp (no_asm_simp) only: set_append list.set list.map) apply blast
   apply (simp (no_asm_simp))
   apply simp                    (* subgoal bc3 = [] *)
   apply (simp add: comb_nil_def) (* subgoal mt3 = [] \<and> sttp2 = sttp3 *)
@@ -1419,7 +1419,7 @@
      (* (some) preconditions of  wt_instr_offset *)
   apply (simp (no_asm_simp) add: max_ssize_def del: max_of_list_append)
   apply (rule max_of_list_sublist)
-    apply (simp (no_asm_simp) only: set_append set_simps list.map) apply blast
+    apply (simp (no_asm_simp) only: set_append list.set list.map) apply blast
   apply (simp (no_asm_simp))
 
 apply (drule_tac x=sttp2 in spec, simp) (* subgoal \<exists>mt3_rest. \<dots> *)
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -15,19 +15,15 @@
 val proverK = "prover" (*=NAME: name of the external prover to call*)
 val prover_timeoutK = "prover_timeout" (*=TIME: timeout for invoked ATP (seconds of process time)*)
 val keepK = "keep" (*=PATH: path where to keep temporary files created by sledgehammer*)
-val minimizeK = "minimize" (*: enable minimization of theorem set found by sledgehammer*)
-                           (*refers to minimization attempted by Mirabelle*)
-val minimize_timeoutK = "minimize_timeout" (*=TIME: timeout for each minimization step (seconds of*)
 
 val proof_methodK = "proof_method" (*=NAME: how to reconstruct proofs (ie. using metis/smt)*)
-val metis_ftK = "metis_ft" (*: apply metis with fully-typed encoding to the theorems found by sledgehammer*)
 
 val max_factsK = "max_facts" (*=NUM: max. relevant clauses to use*)
 val max_relevantK = "max_relevant" (*=NUM: max. relevant clauses to use*)
 val max_callsK = "max_calls" (*=NUM: max. no. of calls to sledgehammer*)
 val preplay_timeoutK = "preplay_timeout" (*=TIME: timeout for finding reconstructed proof*)
 val isar_proofsK = "isar_proofs" (*: enable Isar proof generation*)
-val sh_minimizeK = "sh_minimize" (*: instruct sledgehammer to run its minimizer*)
+val minimizeK = "minimize" (*: instruct sledgehammer to run its minimizer*)
 
 val check_trivialK = "check_trivial" (*: check if goals are "trivial" (false by default)*)
 val fact_filterK = "fact_filter" (*=STRING: fact filter*)
@@ -43,14 +39,13 @@
 val max_mono_itersK = "max_mono_iters" (*=NUM: max. iterations of monomorphiser*)
 
 fun sh_tag id = "#" ^ string_of_int id ^ " sledgehammer: "
-fun minimize_tag id = "#" ^ string_of_int id ^ " minimize (sledgehammer): "
 fun proof_method_tag meth id = "#" ^ string_of_int id ^ " " ^ (!meth) ^ " (sledgehammer): "
 
 val separator = "-----"
 
 (*FIXME sensible to have Mirabelle-level Sledgehammer defaults?*)
 (*defaults used in this Mirabelle action*)
-val preplay_timeout_default = "3"
+val preplay_timeout_default = "1"
 val lam_trans_default = "smart"
 val uncurried_aliases_default = "smart"
 val fact_filter_default = "smart"
@@ -60,7 +55,6 @@
 val slice_default = "true"
 val max_calls_default = "10000000"
 val trivial_default = "false"
-val minimize_timeout_default = 5
 
 (*If a key is present in args then augment a list with its pair*)
 (*This is used to avoid fixing default values at the Mirabelle level, and
@@ -93,11 +87,6 @@
   posns: (Position.T * bool) list
   }
 
-datatype min_data = MinData of {
-  succs: int,
-  ab_ratios: int
-  }
-
 fun make_sh_data
       (calls,success,nontriv_calls,nontriv_success,lemmas,max_lems,time_isa,
        time_prover,time_prover_fail) =
@@ -106,9 +95,6 @@
          time_isa=time_isa, time_prover=time_prover,
          time_prover_fail=time_prover_fail}
 
-fun make_min_data (succs, ab_ratios) =
-  MinData{succs=succs, ab_ratios=ab_ratios}
-
 fun make_re_data (calls,success,nontriv_calls,nontriv_success,proofs,time,
                   timeout,lemmas,posns) =
   ReData{calls=calls, success=success, nontriv_calls=nontriv_calls,
@@ -116,7 +102,6 @@
          timeout=timeout, lemmas=lemmas, posns=posns}
 
 val empty_sh_data = make_sh_data (0, 0, 0, 0, 0, 0, 0, 0, 0)
-val empty_min_data = make_min_data (0, 0)
 val empty_re_data = make_re_data (0, 0, 0, 0, 0, 0, 0, (0,0,0), [])
 
 fun tuple_of_sh_data (ShData {calls, success, nontriv_calls, nontriv_success,
@@ -124,55 +109,28 @@
   time_prover, time_prover_fail}) = (calls, success, nontriv_calls,
   nontriv_success, lemmas, max_lems, time_isa, time_prover, time_prover_fail)
 
-fun tuple_of_min_data (MinData {succs, ab_ratios}) = (succs, ab_ratios)
-
 fun tuple_of_re_data (ReData {calls, success, nontriv_calls, nontriv_success,
   proofs, time, timeout, lemmas, posns}) = (calls, success, nontriv_calls,
   nontriv_success, proofs, time, timeout, lemmas, posns)
 
-datatype proof_method_mode =
-  Unminimized | Minimized | UnminimizedFT | MinimizedFT
-
 datatype data = Data of {
   sh: sh_data,
-  min: min_data,
-  re_u: re_data, (* proof method with unminimized set of lemmas *)
-  re_m: re_data, (* proof method with minimized set of lemmas *)
-  re_uft: re_data, (* proof method with unminimized set of lemmas and fully-typed *)
-  re_mft: re_data, (* proof method with minimized set of lemmas and fully-typed *)
-  mini: bool   (* with minimization *)
+  re_u: re_data (* proof method with unminimized set of lemmas *)
   }
 
-fun make_data (sh, min, re_u, re_m, re_uft, re_mft, mini) =
-  Data {sh=sh, min=min, re_u=re_u, re_m=re_m, re_uft=re_uft, re_mft=re_mft,
-    mini=mini}
-
-val empty_data = make_data (empty_sh_data, empty_min_data,
-  empty_re_data, empty_re_data, empty_re_data, empty_re_data, false)
+fun make_data (sh, re_u) = Data {sh=sh, re_u=re_u}
 
-fun map_sh_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
-  let val sh' = make_sh_data (f (tuple_of_sh_data sh))
-  in make_data (sh', min, re_u, re_m, re_uft, re_mft, mini) end
-
-fun map_min_data f (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
-  let val min' = make_min_data (f (tuple_of_min_data min))
-  in make_data (sh, min', re_u, re_m, re_uft, re_mft, mini) end
+val empty_data = make_data (empty_sh_data, empty_re_data)
 
-fun map_re_data f m (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
+fun map_sh_data f (Data {sh, re_u}) =
+  let val sh' = make_sh_data (f (tuple_of_sh_data sh))
+  in make_data (sh', re_u) end
+
+fun map_re_data f (Data {sh, re_u}) =
   let
-    fun map_me g Unminimized   (u, m, uft, mft) = (g u, m, uft, mft)
-      | map_me g Minimized     (u, m, uft, mft) = (u, g m, uft, mft)
-      | map_me g UnminimizedFT (u, m, uft, mft) = (u, m, g uft, mft)
-      | map_me g MinimizedFT   (u, m, uft, mft) = (u, m, uft, g mft)
-
     val f' = make_re_data o f o tuple_of_re_data
-
-    val (re_u', re_m', re_uft', re_mft') =
-      map_me f' m (re_u, re_m, re_uft, re_mft)
-  in make_data (sh, min, re_u', re_m', re_uft', re_mft', mini) end
-
-fun set_mini mini (Data {sh, min, re_u, re_m, re_uft, re_mft, ...}) =
-  make_data (sh, min, re_u, re_m, re_uft, re_mft, mini)
+    val re_u' = f' re_u
+  in make_data (sh, re_u') end
 
 fun inc_max (n:int) (s,sos,m) = (s+n, sos + n*n, Int.max(m,n));
 
@@ -212,12 +170,6 @@
   (fn (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail)
     => (calls,success,nontriv_calls, nontriv_success, lemmas,max_lems,time_isa,time_prover,time_prover_fail + t))
 
-val inc_min_succs = map_min_data
-  (fn (succs,ab_ratios) => (succs+1, ab_ratios))
-
-fun inc_min_ab_ratios r = map_min_data
-  (fn (succs, ab_ratios) => (succs, ab_ratios+r))
-
 val inc_proof_method_calls = map_re_data
   (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
     => (calls + 1, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas,posns))
@@ -238,21 +190,21 @@
   (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
     => (calls, success, nontriv_calls, nontriv_success, proofs + 1, time, timeout, lemmas,posns))
 
-fun inc_proof_method_time m t = map_re_data
+fun inc_proof_method_time t = map_re_data
  (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
-  => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns)) m
+  => (calls, success, nontriv_calls, nontriv_success, proofs, time + t, timeout, lemmas,posns))
 
 val inc_proof_method_timeout = map_re_data
   (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
     => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout + 1, lemmas,posns))
 
-fun inc_proof_method_lemmas m n = map_re_data
+fun inc_proof_method_lemmas n = map_re_data
   (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
-    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns)) m
+    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, inc_max n lemmas, posns))
 
-fun inc_proof_method_posns m pos = map_re_data
+fun inc_proof_method_posns pos = map_re_data
   (fn (calls,success,nontriv_calls, nontriv_success, proofs,time,timeout,lemmas,posns)
-    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns)) m
+    => (calls, success, nontriv_calls, nontriv_success, proofs, time, timeout, lemmas, pos::posns))
 
 val str0 = string_of_int o the_default 0
 
@@ -311,37 +263,23 @@
   else ()
  )
 
-fun log_min_data log (succs, ab_ratios) =
-  (log ("Number of successful minimizations: " ^ string_of_int succs);
-   log ("After/before ratios: " ^ string_of_int ab_ratios)
-  )
-
 in
 
-fun log_data id log (Data {sh, min, re_u, re_m, re_uft, re_mft, mini}) =
+fun log_data id log (Data {sh, re_u}) =
   let
     val ShData {calls=sh_calls, ...} = sh
 
     fun app_if (ReData {calls, ...}) f = if calls > 0 then f () else ()
     fun log_re tag m =
       log_re_data log tag sh_calls (tuple_of_re_data m)
-    fun log_proof_method (tag1, m1) (tag2, m2) = app_if m1 (fn () =>
-      (log_re tag1 m1; log ""; app_if m2 (fn () => log_re tag2 m2)))
+    fun log_proof_method (tag1, m1) = app_if m1 (fn () => (log_re tag1 m1; log ""))
   in
     if sh_calls > 0
     then
      (log ("\n\n\nReport #" ^ string_of_int id ^ ":\n");
       log_sh_data log (tuple_of_sh_data sh);
       log "";
-      if not mini
-      then log_proof_method ("", re_u) ("fully-typed ", re_uft)
-      else
-        app_if re_u (fn () =>
-         (log_proof_method ("unminimized ", re_u) ("unminimized fully-typed ", re_uft);
-          log "";
-          app_if re_m (fn () =>
-            (log_min_data log (tuple_of_min_data min); log "";
-             log_proof_method ("", re_m) ("fully-typed ", re_mft))))))
+      log_proof_method ("", re_u))
     else ()
   end
 
@@ -411,7 +349,7 @@
 
 fun run_sh prover_name fact_filter type_enc strict max_facts slice
       lam_trans uncurried_aliases e_selection_heuristic term_order force_sos
-      hard_timeout timeout preplay_timeout isar_proofsLST sh_minimizeLST
+      hard_timeout timeout preplay_timeout isar_proofsLST minimizeLST
       max_new_mono_instancesLST max_mono_itersLST dir pos st =
   let
     val thy = Proof.theory_of st
@@ -421,7 +359,7 @@
         Config.put Sledgehammer_Prover_ATP.atp_dest_dir dir
         #> Config.put Sledgehammer_Prover_ATP.atp_problem_prefix
           ("prob_" ^ str0 (Position.line_of pos) ^ "__")
-        #> Config.put SMT_Config.debug_files
+        #> Config.put SMT2_Config.debug_files
           (dir ^ "/" ^ Name.desymbolize (SOME false) (ATP_Util.timestamp ()) ^ "_"
           ^ serial_string ())
       | set_file_name NONE = I
@@ -435,7 +373,7 @@
                   term_order |> the_default I)
             #> (Option.map (Config.put ATP_Systems.force_sos)
                   force_sos |> the_default I))
-    val params as {max_facts, ...} =
+    val params as {max_facts, minimize, preplay_timeout, ...} =
       Sledgehammer_Commands.default_params thy
          ([("verbose", "true"),
            ("fact_filter", fact_filter),
@@ -448,7 +386,7 @@
            ("timeout", string_of_int timeout),
            ("preplay_timeout", preplay_timeout)]
           |> isar_proofsLST
-          |> sh_minimizeLST (*don't confuse the two minimization flags*)
+          |> minimizeLST (*don't confuse the two minimization flags*)
           |> max_new_mono_instancesLST
           |> max_mono_itersLST)
     val default_max_facts =
@@ -460,11 +398,9 @@
       | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
     fun failed failure =
       ({outcome = SOME failure, used_facts = [], used_from = [],
-        run_time = Time.zeroTime,
-        preplay = Lazy.value (Sledgehammer_Proof_Methods.Metis_Method (NONE, NONE),
-          Sledgehammer_Proof_Methods.Play_Failed),
-        message = K "", message_tail = ""}, ~1)
-    val ({outcome, used_facts, run_time, preplay, message, message_tail, ...}
+        preferred_methss = (Sledgehammer_Proof_Methods.Auto_Method, []), run_time = Time.zeroTime,
+        message = K ""}, ~1)
+    val ({outcome, used_facts, preferred_methss, run_time, message, ...}
          : Sledgehammer_Prover.prover_result,
          time_isa) = time_limit (Mirabelle.cpu_time (fn () =>
       let
@@ -488,11 +424,12 @@
         val problem =
           {comment = "", state = st', goal = goal, subgoal = i,
            subgoal_count = Sledgehammer_Util.subgoal_count st, factss = factss}
-      in prover params (K (K (K ""))) problem end)) ()
+      in prover params problem end)) ()
       handle TimeLimit.TimeOut => failed ATP_Proof.TimedOut
            | Fail "inappropriate" => failed ATP_Proof.Inappropriate
     val time_prover = run_time |> Time.toMilliseconds
-    val msg = message (Lazy.force preplay) ^ message_tail
+    val msg = message (fn () => Sledgehammer.play_one_line_proof minimize preplay_timeout used_facts
+      st' i preferred_methss)
   in
     (case outcome of
       NONE => (msg, SH_OK (time_isa, time_prover, used_facts))
@@ -534,7 +471,7 @@
     val preplay_timeout = AList.lookup (op =) args preplay_timeoutK
       |> the_default preplay_timeout_default
     val isar_proofsLST = available_parameter args isar_proofsK "isar_proofs"
-    val sh_minimizeLST = available_parameter args sh_minimizeK "minimize"
+    val minimizeLST = available_parameter args minimizeK "minimize"
     val max_new_mono_instancesLST =
       available_parameter args max_new_mono_instancesK max_new_mono_instancesK
     val max_mono_itersLST = available_parameter args max_mono_itersK max_mono_itersK
@@ -542,7 +479,7 @@
     val (msg, result) =
       run_sh prover_name fact_filter type_enc strict max_facts slice lam_trans
         uncurried_aliases e_selection_heuristic term_order force_sos
-        hard_timeout timeout preplay_timeout isar_proofsLST sh_minimizeLST
+        hard_timeout timeout preplay_timeout isar_proofsLST minimizeLST
         max_new_mono_instancesLST max_mono_itersLST dir pos st
   in
     (case result of
@@ -574,57 +511,6 @@
 
 end
 
-fun run_minimize args meth named_thms id ({pre = st, log, ...} : Mirabelle.run_args) =
-  let
-    val thy = Proof.theory_of st
-    val {goal, ...} = Proof.goal st
-    val n0 = length (these (!named_thms))
-    val prover_name = get_prover_name thy args
-    val type_enc = AList.lookup (op =) args type_encK |> the_default type_enc_default
-    val strict = AList.lookup (op =) args strictK |> the_default strict_default
-    val timeout =
-      AList.lookup (op =) args minimize_timeoutK
-      |> Option.map (fst o read_int o raw_explode)  (* FIXME Symbol.explode (?) *)
-      |> the_default minimize_timeout_default
-    val preplay_timeout = AList.lookup (op =) args preplay_timeoutK
-      |> the_default preplay_timeout_default
-    val isar_proofsLST = available_parameter args isar_proofsK "isar_proofs"
-    val sh_minimizeLST = available_parameter args sh_minimizeK "minimize"
-    val max_new_mono_instancesLST =
-      available_parameter args max_new_mono_instancesK max_new_mono_instancesK
-    val max_mono_itersLST = available_parameter args max_mono_itersK max_mono_itersK
-    val params = Sledgehammer_Commands.default_params thy
-     ([("provers", prover_name),
-       ("verbose", "true"),
-       ("type_enc", type_enc),
-       ("strict", strict),
-       ("timeout", string_of_int timeout),
-       ("preplay_timeout", preplay_timeout)]
-      |> isar_proofsLST
-      |> sh_minimizeLST (*don't confuse the two minimization flags*)
-      |> max_new_mono_instancesLST
-      |> max_mono_itersLST)
-    val minimize =
-      Sledgehammer_Prover_Minimize.minimize_facts (K ()) prover_name params true 1
-        (Sledgehammer_Util.subgoal_count st)
-    val _ = log separator
-    val (used_facts, (preplay, message, message_tail)) =
-      minimize st goal NONE (these (!named_thms))
-    val msg = message (Lazy.force preplay) ^ message_tail
-  in
-    (case used_facts of
-      SOME named_thms' =>
-        (change_data id inc_min_succs;
-         change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0));
-         if length named_thms' = n0
-         then log (minimize_tag id ^ "already minimal")
-         else (meth := proof_method_from_msg args msg;
-               named_thms := SOME named_thms';
-               log (minimize_tag id ^ "succeeded:\n" ^ msg))
-        )
-    | NONE => log (minimize_tag id ^ "failed: " ^ msg))
-  end
-
 fun override_params prover type_enc timeout =
   [("provers", prover),
    ("max_facts", "0"),
@@ -633,13 +519,13 @@
    ("slice", "false"),
    ("timeout", timeout |> Time.toSeconds |> string_of_int)]
 
-fun run_proof_method trivial full m name meth named_thms id
+fun run_proof_method trivial full name meth named_thms id
     ({pre=st, timeout, log, pos, ...}: Mirabelle.run_args) =
   let
     fun do_method named_thms ctxt =
       let
         val ref_of_str =
-          suffix ";" #> Outer_Syntax.scan Position.none #> Parse_Spec.xthm
+          suffix ";" #> Outer_Syntax.scan (Keyword.get_lexicons ()) Position.none #> Parse.xthm
           #> fst
         val thms = named_thms |> maps snd
         val facts = named_thms |> map (ref_of_str o fst o fst)
@@ -648,16 +534,16 @@
           timeout |> Time.toReal |> curry Real.* time_slice |> Time.fromReal
         fun sledge_tac time_slice prover type_enc =
           Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt
-            (override_params prover type_enc (my_timeout time_slice)) fact_override
+            (override_params prover type_enc (my_timeout time_slice)) fact_override []
       in
         if !meth = "sledgehammer_tac" then
           sledge_tac 0.2 ATP_Proof.vampireN "mono_native"
           ORELSE' sledge_tac 0.2 ATP_Proof.eN "poly_guards??"
           ORELSE' sledge_tac 0.2 ATP_Proof.spassN "mono_native"
           ORELSE' sledge_tac 0.2 ATP_Proof.z3_tptpN "poly_tags??"
-          ORELSE' SMT_Solver.smt_tac ctxt thms
+          ORELSE' SMT2_Solver.smt2_tac ctxt thms
         else if !meth = "smt" then
-          SMT_Solver.smt_tac ctxt thms
+          SMT2_Solver.smt2_tac ctxt thms
         else if full then
           Metis_Tactic.metis_tac [ATP_Proof_Reconstruct.full_typesN]
             ATP_Proof_Reconstruct.default_metis_lam_trans ctxt thms
@@ -665,7 +551,7 @@
           let
             val (type_encs, lam_trans) =
               !meth
-              |> Outer_Syntax.scan Position.start
+              |> Outer_Syntax.scan (Keyword.get_lexicons ()) Position.start
               |> filter Token.is_proper |> tl
               |> Metis_Tactic.parse_metis_options |> fst
               |>> the_default [ATP_Proof_Reconstruct.partial_typesN]
@@ -680,22 +566,22 @@
       Mirabelle.can_apply timeout (do_method named_thms) st
 
     fun with_time (false, t) = "failed (" ^ string_of_int t ^ ")"
-      | with_time (true, t) = (change_data id (inc_proof_method_success m);
+      | with_time (true, t) = (change_data id inc_proof_method_success;
           if trivial then ()
-          else change_data id (inc_proof_method_nontriv_success m);
-          change_data id (inc_proof_method_lemmas m (length named_thms));
-          change_data id (inc_proof_method_time m t);
-          change_data id (inc_proof_method_posns m (pos, trivial));
-          if name = "proof" then change_data id (inc_proof_method_proofs m) else ();
+          else change_data id inc_proof_method_nontriv_success;
+          change_data id (inc_proof_method_lemmas (length named_thms));
+          change_data id (inc_proof_method_time t);
+          change_data id (inc_proof_method_posns (pos, trivial));
+          if name = "proof" then change_data id inc_proof_method_proofs else ();
           "succeeded (" ^ string_of_int t ^ ")")
     fun timed_method named_thms =
       (with_time (Mirabelle.cpu_time apply_method named_thms), true)
-      handle TimeLimit.TimeOut => (change_data id (inc_proof_method_timeout m); ("timeout", false))
+      handle TimeLimit.TimeOut => (change_data id inc_proof_method_timeout; ("timeout", false))
            | ERROR msg => ("error: " ^ msg, false)
 
     val _ = log separator
-    val _ = change_data id (inc_proof_method_calls m)
-    val _ = if trivial then () else change_data id (inc_proof_method_nontriv_calls m)
+    val _ = change_data id inc_proof_method_calls
+    val _ = if trivial then () else change_data id inc_proof_method_nontriv_calls
   in
     named_thms
     |> timed_method
@@ -724,38 +610,18 @@
           val meth = Unsynchronized.ref ""
           val named_thms =
             Unsynchronized.ref (NONE : ((string * stature) * thm list) list option)
-          val minimize = AList.defined (op =) args minimizeK
-          val metis_ft = AList.defined (op =) args metis_ftK
           val trivial =
             if AList.lookup (op =) args check_trivialK |> the_default trivial_default
                             |> Markup.parse_bool then
               Try0.try0 (SOME try_timeout) ([], [], [], []) pre
               handle TimeLimit.TimeOut => false
             else false
-          fun apply_method m1 m2 =
-            if metis_ft
-            then
-              if not (Mirabelle.catch_result (proof_method_tag meth) false
-                  (run_proof_method trivial false m1 name meth (these (!named_thms))) id st)
-              then
-                (Mirabelle.catch_result (proof_method_tag meth) false
-                  (run_proof_method trivial true m2 name meth (these (!named_thms))) id st; ())
-              else ()
-            else
-              (Mirabelle.catch_result (proof_method_tag meth) false
-                (run_proof_method trivial false m1 name meth (these (!named_thms))) id st; ())
+          fun apply_method () =
+            (Mirabelle.catch_result (proof_method_tag meth) false
+              (run_proof_method trivial false name meth (these (!named_thms))) id st; ())
         in
-          change_data id (set_mini minimize);
           Mirabelle.catch sh_tag (run_sledgehammer trivial args meth named_thms) id st;
-          if is_some (!named_thms)
-          then
-           (apply_method Unminimized UnminimizedFT;
-            if minimize andalso not (null (these (!named_thms)))
-            then
-             (Mirabelle.catch minimize_tag (run_minimize args meth named_thms) id st;
-              apply_method Minimized MinimizedFT)
-            else ())
-          else ()
+          if is_some (!named_thms) then apply_method () else ()
         end
     end
   end
--- a/src/HOL/Nat.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Nat.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -12,6 +12,8 @@
 begin
 
 ML_file "~~/src/Tools/rat.ML"
+
+named_theorems arith "arith facts -- only ground formulas"
 ML_file "Tools/arith_data.ML"
 ML_file "~~/src/Provers/Arith/fast_lin_arith.ML"
 
@@ -130,9 +132,9 @@
   nat.collapse
   nat.expand
   nat.sel
-  nat.sel_exhaust
-  nat.sel_split
-  nat.sel_split_asm
+  nat.exhaust_sel
+  nat.split_sel
+  nat.split_sel_asm
 
 lemma nat_exhaust [case_names 0 Suc, cases type: nat]:
   -- {* for backward compatibility -- names of variables differ *}
--- a/src/HOL/Nitpick.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Nitpick.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -14,109 +14,105 @@
   "nitpick_params" :: thy_decl
 begin
 
-typedecl bisim_iterator
+datatype ('a, 'b) fun_box = FunBox "'a \<Rightarrow> 'b"
+datatype ('a, 'b) pair_box = PairBox 'a 'b
+datatype 'a word = Word "'a set"
 
-axiomatization unknown :: 'a
-           and is_unknown :: "'a \<Rightarrow> bool"
-           and bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
-           and bisim_iterator_max :: bisim_iterator
-           and Quot :: "'a \<Rightarrow> 'b"
-           and safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
-
-datatype ('a, 'b) fun_box = FunBox "('a \<Rightarrow> 'b)"
-datatype ('a, 'b) pair_box = PairBox 'a 'b
-
+typedecl bisim_iterator
 typedecl unsigned_bit
 typedecl signed_bit
 
-datatype 'a word = Word "('a set)"
+consts
+  unknown :: 'a
+  is_unknown :: "'a \<Rightarrow> bool"
+  bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+  bisim_iterator_max :: bisim_iterator
+  Quot :: "'a \<Rightarrow> 'b"
+  safe_The :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
 
 text {*
 Alternative definitions.
 *}
 
-lemma Ex1_unfold [nitpick_unfold]:
-"Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
-apply (rule eq_reflection)
-apply (simp add: Ex1_def set_eq_iff)
-apply (rule iffI)
- apply (erule exE)
- apply (erule conjE)
- apply (rule_tac x = x in exI)
- apply (rule allI)
- apply (rename_tac y)
- apply (erule_tac x = y in allE)
-by auto
+lemma Ex1_unfold[nitpick_unfold]: "Ex1 P \<equiv> \<exists>x. {x. P x} = {x}"
+  apply (rule eq_reflection)
+  apply (simp add: Ex1_def set_eq_iff)
+  apply (rule iffI)
+   apply (erule exE)
+   apply (erule conjE)
+   apply (rule_tac x = x in exI)
+   apply (rule allI)
+   apply (rename_tac y)
+   apply (erule_tac x = y in allE)
+  by auto
 
-lemma rtrancl_unfold [nitpick_unfold]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
+lemma rtrancl_unfold[nitpick_unfold]: "r\<^sup>* \<equiv> (r\<^sup>+)\<^sup>="
   by (simp only: rtrancl_trancl_reflcl)
 
-lemma rtranclp_unfold [nitpick_unfold]:
-"rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
-by (rule eq_reflection) (auto dest: rtranclpD)
+lemma rtranclp_unfold[nitpick_unfold]: "rtranclp r a b \<equiv> (a = b \<or> tranclp r a b)"
+  by (rule eq_reflection) (auto dest: rtranclpD)
 
-lemma tranclp_unfold [nitpick_unfold]:
-"tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
-by (simp add: trancl_def)
+lemma tranclp_unfold[nitpick_unfold]:
+  "tranclp r a b \<equiv> (a, b) \<in> trancl {(x, y). r x y}"
+  by (simp add: trancl_def)
 
 lemma [nitpick_simp]:
-"of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))"
-by (cases n) auto
+  "of_nat n = (if n = 0 then 0 else 1 + of_nat (n - 1))"
+  by (cases n) auto
 
 definition prod :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<times> 'b) set" where
-"prod A B = {(a, b). a \<in> A \<and> b \<in> B}"
+  "prod A B = {(a, b). a \<in> A \<and> b \<in> B}"
 
 definition refl' :: "('a \<times> 'a) set \<Rightarrow> bool" where
-"refl' r \<equiv> \<forall>x. (x, x) \<in> r"
+  "refl' r \<equiv> \<forall>x. (x, x) \<in> r"
 
 definition wf' :: "('a \<times> 'a) set \<Rightarrow> bool" where
-"wf' r \<equiv> acyclic r \<and> (finite r \<or> unknown)"
+  "wf' r \<equiv> acyclic r \<and> (finite r \<or> unknown)"
 
 definition card' :: "'a set \<Rightarrow> nat" where
-"card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
+  "card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
 
 definition setsum' :: "('a \<Rightarrow> 'b\<Colon>comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
-"setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
+  "setsum' f A \<equiv> if finite A then listsum (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
 
 inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
-"fold_graph' f z {} z" |
-"\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
+  "fold_graph' f z {} z" |
+  "\<lbrakk>x \<in> A; fold_graph' f z (A - {x}) y\<rbrakk> \<Longrightarrow> fold_graph' f z A (f x y)"
 
 text {*
 The following lemmas are not strictly necessary but they help the
 \textit{specialize} optimization.
 *}
 
-lemma The_psimp [nitpick_psimp]:
-  "P = (op =) x \<Longrightarrow> The P = x"
+lemma The_psimp[nitpick_psimp]: "P = (op =) x \<Longrightarrow> The P = x"
   by auto
 
-lemma Eps_psimp [nitpick_psimp]:
-"\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
-apply (cases "P (Eps P)")
- apply auto
-apply (erule contrapos_np)
-by (rule someI)
+lemma Eps_psimp[nitpick_psimp]:
+  "\<lbrakk>P x; \<not> P y; Eps P = y\<rbrakk> \<Longrightarrow> Eps P = x"
+  apply (cases "P (Eps P)")
+   apply auto
+  apply (erule contrapos_np)
+  by (rule someI)
 
-lemma case_unit_unfold [nitpick_unfold]:
-"case_unit x u \<equiv> x"
-apply (subgoal_tac "u = ()")
- apply (simp only: unit.case)
-by simp
+lemma case_unit_unfold[nitpick_unfold]:
+  "case_unit x u \<equiv> x"
+  apply (subgoal_tac "u = ()")
+   apply (simp only: unit.case)
+  by simp
 
-declare unit.case [nitpick_simp del]
+declare unit.case[nitpick_simp del]
 
-lemma case_nat_unfold [nitpick_unfold]:
-"case_nat x f n \<equiv> if n = 0 then x else f (n - 1)"
-apply (rule eq_reflection)
-by (cases n) auto
+lemma case_nat_unfold[nitpick_unfold]:
+  "case_nat x f n \<equiv> if n = 0 then x else f (n - 1)"
+  apply (rule eq_reflection)
+  by (cases n) auto
 
-declare nat.case [nitpick_simp del]
+declare nat.case[nitpick_simp del]
 
-lemma size_list_simp [nitpick_simp]:
-"size_list f xs = (if xs = [] then 0 else Suc (f (hd xs) + size_list f (tl xs)))"
-"size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
-by (cases xs) auto
+lemma size_list_simp[nitpick_simp]:
+  "size_list f xs = (if xs = [] then 0 else Suc (f (hd xs) + size_list f (tl xs)))"
+  "size xs = (if xs = [] then 0 else Suc (size (tl xs)))"
+  by (cases xs) auto
 
 text {*
 Auxiliary definitions used to provide an alternative representation for
@@ -124,89 +120,89 @@
 *}
 
 function nat_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
-[simp del]: "nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))"
-by auto
-termination
-apply (relation "measure (\<lambda>(x, y). x + y + (if y > x then 1 else 0))")
- apply auto
- apply (metis mod_less_divisor xt1(9))
-by (metis mod_mod_trivial mod_self nat_neq_iff xt1(10))
+  "nat_gcd x y = (if y = 0 then x else nat_gcd y (x mod y))"
+  by auto
+  termination
+  apply (relation "measure (\<lambda>(x, y). x + y + (if y > x then 1 else 0))")
+   apply auto
+   apply (metis mod_less_divisor xt1(9))
+  by (metis mod_mod_trivial mod_self nat_neq_iff xt1(10))
+
+declare nat_gcd.simps[simp del]
 
 definition nat_lcm :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
-"nat_lcm x y = x * y div (nat_gcd x y)"
+  "nat_lcm x y = x * y div (nat_gcd x y)"
 
 definition int_gcd :: "int \<Rightarrow> int \<Rightarrow> int" where
-"int_gcd x y = int (nat_gcd (nat (abs x)) (nat (abs y)))"
+  "int_gcd x y = int (nat_gcd (nat (abs x)) (nat (abs y)))"
 
 definition int_lcm :: "int \<Rightarrow> int \<Rightarrow> int" where
-"int_lcm x y = int (nat_lcm (nat (abs x)) (nat (abs y)))"
+  "int_lcm x y = int (nat_lcm (nat (abs x)) (nat (abs y)))"
 
 definition Frac :: "int \<times> int \<Rightarrow> bool" where
-"Frac \<equiv> \<lambda>(a, b). b > 0 \<and> int_gcd a b = 1"
+  "Frac \<equiv> \<lambda>(a, b). b > 0 \<and> int_gcd a b = 1"
 
-axiomatization
-  Abs_Frac :: "int \<times> int \<Rightarrow> 'a" and
+consts
+  Abs_Frac :: "int \<times> int \<Rightarrow> 'a"
   Rep_Frac :: "'a \<Rightarrow> int \<times> int"
 
 definition zero_frac :: 'a where
-"zero_frac \<equiv> Abs_Frac (0, 1)"
+  "zero_frac \<equiv> Abs_Frac (0, 1)"
 
 definition one_frac :: 'a where
-"one_frac \<equiv> Abs_Frac (1, 1)"
+  "one_frac \<equiv> Abs_Frac (1, 1)"
 
 definition num :: "'a \<Rightarrow> int" where
-"num \<equiv> fst o Rep_Frac"
+  "num \<equiv> fst o Rep_Frac"
 
 definition denom :: "'a \<Rightarrow> int" where
-"denom \<equiv> snd o Rep_Frac"
+  "denom \<equiv> snd o Rep_Frac"
 
 function norm_frac :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
-[simp del]: "norm_frac a b = (if b < 0 then norm_frac (- a) (- b)
-                              else if a = 0 \<or> b = 0 then (0, 1)
-                              else let c = int_gcd a b in (a div c, b div c))"
-by pat_completeness auto
-termination by (relation "measure (\<lambda>(_, b). if b < 0 then 1 else 0)") auto
+  "norm_frac a b =
+    (if b < 0 then norm_frac (- a) (- b)
+     else if a = 0 \<or> b = 0 then (0, 1)
+     else let c = int_gcd a b in (a div c, b div c))"
+  by pat_completeness auto
+  termination by (relation "measure (\<lambda>(_, b). if b < 0 then 1 else 0)") auto
+
+declare norm_frac.simps[simp del]
 
 definition frac :: "int \<Rightarrow> int \<Rightarrow> 'a" where
-"frac a b \<equiv> Abs_Frac (norm_frac a b)"
+  "frac a b \<equiv> Abs_Frac (norm_frac a b)"
 
 definition plus_frac :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
-[nitpick_simp]:
-"plus_frac q r = (let d = int_lcm (denom q) (denom r) in
-                    frac (num q * (d div denom q) + num r * (d div denom r)) d)"
+  [nitpick_simp]: "plus_frac q r = (let d = int_lcm (denom q) (denom r) in
+    frac (num q * (d div denom q) + num r * (d div denom r)) d)"
 
 definition times_frac :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
-[nitpick_simp]:
-"times_frac q r = frac (num q * num r) (denom q * denom r)"
+  [nitpick_simp]: "times_frac q r = frac (num q * num r) (denom q * denom r)"
 
 definition uminus_frac :: "'a \<Rightarrow> 'a" where
-"uminus_frac q \<equiv> Abs_Frac (- num q, denom q)"
+  "uminus_frac q \<equiv> Abs_Frac (- num q, denom q)"
 
 definition number_of_frac :: "int \<Rightarrow> 'a" where
-"number_of_frac n \<equiv> Abs_Frac (n, 1)"
+  "number_of_frac n \<equiv> Abs_Frac (n, 1)"
 
 definition inverse_frac :: "'a \<Rightarrow> 'a" where
-"inverse_frac q \<equiv> frac (denom q) (num q)"
+  "inverse_frac q \<equiv> frac (denom q) (num q)"
 
 definition less_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
-[nitpick_simp]:
-"less_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) < 0"
+  [nitpick_simp]: "less_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) < 0"
 
 definition less_eq_frac :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
-[nitpick_simp]:
-"less_eq_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) \<le> 0"
+  [nitpick_simp]: "less_eq_frac q r \<longleftrightarrow> num (plus_frac q (uminus_frac r)) \<le> 0"
 
 definition of_frac :: "'a \<Rightarrow> 'b\<Colon>{inverse,ring_1}" where
-"of_frac q \<equiv> of_int (num q) / of_int (denom q)"
+  "of_frac q \<equiv> of_int (num q) / of_int (denom q)"
 
 axiomatization wf_wfrec :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
 
 definition wf_wfrec' :: "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
-[nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x"
+  [nitpick_simp]: "wf_wfrec' R F x = F (cut (wf_wfrec R F) R x) x"
 
 definition wfrec' ::  "('a \<times> 'a) set \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
-"wfrec' R F x \<equiv> if wf R then wf_wfrec' R F x
-                else THE y. wfrec_rel R (%f x. F (cut f R x) x) x y"
+  "wfrec' R F x \<equiv> if wf R then wf_wfrec' R F x else THE y. wfrec_rel R (\<lambda>f x. F (cut f R x) x) x y"
 
 ML_file "Tools/Nitpick/kodkod.ML"
 ML_file "Tools/Nitpick/kodkod_sat.ML"
@@ -234,20 +230,18 @@
      (@{const_name wfrec}, @{const_name wfrec'})]
 *}
 
-hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The
-    FunBox PairBox Word prod refl' wf' card' setsum'
-    fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac zero_frac
-    one_frac num denom norm_frac frac plus_frac times_frac uminus_frac
-    number_of_frac inverse_frac less_frac less_eq_frac of_frac wf_wfrec wf_wfrec
-    wfrec'
+hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The FunBox PairBox Word prod
+  refl' wf' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac
+  zero_frac one_frac num denom norm_frac frac plus_frac times_frac uminus_frac number_of_frac
+  inverse_frac less_frac less_eq_frac of_frac wf_wfrec wf_wfrec wfrec'
+
 hide_type (open) bisim_iterator fun_box pair_box unsigned_bit signed_bit word
-hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold
-    prod_def refl'_def wf'_def card'_def setsum'_def
-    fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
-    size_list_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def
-    zero_frac_def one_frac_def num_def denom_def norm_frac_def frac_def
-    plus_frac_def times_frac_def uminus_frac_def number_of_frac_def
-    inverse_frac_def less_frac_def less_eq_frac_def of_frac_def wf_wfrec'_def
-    wfrec'_def
+
+hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold prod_def refl'_def wf'_def
+  card'_def setsum'_def fold_graph'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
+  size_list_simp nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def one_frac_def
+  num_def denom_def norm_frac_def frac_def plus_frac_def times_frac_def uminus_frac_def
+  number_of_frac_def inverse_frac_def less_frac_def less_eq_frac_def of_frac_def wf_wfrec'_def
+  wfrec'_def
 
 end
--- a/src/HOL/Nominal/nominal_induct.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Nominal/nominal_induct.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -172,9 +172,8 @@
   Scan.lift (Args.mode Induct.no_simpN) --
   (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
     avoiding -- fixing -- rule_spec) >>
-  (fn (no_simp, (((x, y), z), w)) => fn ctxt =>
-    RAW_METHOD_CASES (fn facts =>
-      HEADGOAL (nominal_induct_tac ctxt (not no_simp) x y z w facts)));
+  (fn (no_simp, (((x, y), z), w)) => fn ctxt => fn facts =>
+    HEADGOAL (nominal_induct_tac ctxt (not no_simp) x y z w facts));
 
 end;
 
--- a/src/HOL/Nominal/nominal_inductive.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Nominal/nominal_inductive.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -167,8 +167,8 @@
     val _ = (case duplicates (op = o pairself fst) avoids of
         [] => ()
       | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs)));
-    val _ = assert_all (null o duplicates op = o snd) avoids
-      (fn (a, _) => error ("Duplicate variable names for case " ^ quote a));
+    val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse
+      error ("Duplicate variable names for case " ^ quote a));
     val _ = (case subtract (op =) induct_cases (map fst avoids) of
         [] => ()
       | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs));
--- a/src/HOL/Nominal/nominal_primrec.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Nominal/nominal_primrec.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -373,9 +373,10 @@
           |> snd
         end)
       [goals] |>
-    Proof.apply (Method.Basic (fn ctxt => RAW_METHOD (fn _ =>
-      rewrite_goals_tac ctxt defs_thms THEN
-      compose_tac (false, rule, length rule_prems) 1))) |>
+    Proof.apply (Method.Basic (fn ctxt => fn _ =>
+      NO_CASES
+       (rewrite_goals_tac ctxt defs_thms THEN
+        compose_tac (false, rule, length rule_prems) 1))) |>
     Seq.hd
   end;
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,1819 @@
+(* Author: Manuel Eberl *)
+
+header {* Abstract euclidean algorithm *}
+
+theory Euclidean_Algorithm
+imports Complex_Main
+begin
+
+lemma finite_int_set_iff_bounded_le:
+  "finite (N::int set) = (\<exists>m\<ge>0. \<forall>n\<in>N. abs n \<le> m)"
+proof
+  assume "finite (N::int set)"
+  hence "finite (nat ` abs ` N)" by (intro finite_imageI)
+  hence "\<exists>m. \<forall>n\<in>nat`abs`N. n \<le> m" by (simp add: finite_nat_set_iff_bounded_le)
+  then obtain m :: nat where "\<forall>n\<in>N. nat (abs n) \<le> nat (int m)" by auto
+  then show "\<exists>m\<ge>0. \<forall>n\<in>N. abs n \<le> m" by (intro exI[of _ "int m"]) (auto simp: nat_le_eq_zle)
+next
+  assume "\<exists>m\<ge>0. \<forall>n\<in>N. abs n \<le> m"
+  then obtain m where "m \<ge> 0" and "\<forall>n\<in>N. abs n \<le> m" by blast
+  hence "\<forall>n\<in>N. nat (abs n) \<le> nat m" by (auto simp: nat_le_eq_zle)
+  hence "\<forall>n\<in>nat`abs`N. n \<le> nat m" by (auto simp: nat_le_eq_zle)
+  hence A: "finite ((nat \<circ> abs)`N)" unfolding o_def 
+      by (subst finite_nat_set_iff_bounded_le) blast
+  {
+    assume "\<not>finite N"
+    from pigeonhole_infinite[OF this A] obtain x 
+       where "x \<in> N" and B: "~finite {a\<in>N. nat (abs a) = nat (abs x)}" 
+       unfolding o_def by blast
+    have "{a\<in>N. nat (abs a) = nat (abs x)} \<subseteq> {x, -x}" by auto
+    hence "finite {a\<in>N. nat (abs a) = nat (abs x)}" by (rule finite_subset) simp
+    with B have False by contradiction
+  }
+  then show "finite N" by blast
+qed
+
+context semiring_div
+begin
+
+lemma dvd_setprod [intro]:
+  assumes "finite A" and "x \<in> A"
+  shows "f x dvd setprod f A"
+proof
+  from `finite A` have "setprod f (insert x (A - {x})) = f x * setprod f (A - {x})"
+    by (intro setprod.insert) auto
+  also from `x \<in> A` have "insert x (A - {x}) = A" by blast
+  finally show "setprod f A = f x * setprod f (A - {x})" .
+qed
+
+lemma dvd_mult_cancel_left:
+  assumes "a \<noteq> 0" and "a * b dvd a * c"
+  shows "b dvd c"
+proof-
+  from assms(2) obtain k where "a * c = a * b * k" unfolding dvd_def by blast
+  hence "c * a = b * k * a" by (simp add: ac_simps)
+  hence "c * (a div a) = b * k * (a div a)" by (simp add: div_mult_swap)
+  also from `a \<noteq> 0` have "a div a = 1" by simp
+  finally show ?thesis by simp
+qed
+
+lemma dvd_mult_cancel_right:
+  "a \<noteq> 0 \<Longrightarrow> b * a dvd c * a \<Longrightarrow> b dvd c"
+  by (subst (asm) (1 2) ac_simps, rule dvd_mult_cancel_left)
+
+lemma nonzero_pow_nonzero:
+  "a \<noteq> 0 \<Longrightarrow> a ^ n \<noteq> 0"
+  by (induct n) (simp_all add: no_zero_divisors)
+
+lemma zero_pow_zero: "n \<noteq> 0 \<Longrightarrow> 0 ^ n = 0"
+  by (cases n, simp_all)
+
+lemma pow_zero_iff:
+  "n \<noteq> 0 \<Longrightarrow> a^n = 0 \<longleftrightarrow> a = 0"
+  using nonzero_pow_nonzero zero_pow_zero by auto
+
+end
+
+context semiring_div
+begin 
+
+definition ring_inv :: "'a \<Rightarrow> 'a"
+where
+  "ring_inv x = 1 div x"
+
+definition is_unit :: "'a \<Rightarrow> bool"
+where
+  "is_unit x \<longleftrightarrow> x dvd 1"
+
+definition associated :: "'a \<Rightarrow> 'a \<Rightarrow> bool" 
+where
+  "associated x y \<longleftrightarrow> x dvd y \<and> y dvd x"
+
+lemma unit_prod [intro]:
+  "is_unit x \<Longrightarrow> is_unit y \<Longrightarrow> is_unit (x * y)"
+  unfolding is_unit_def by (subst mult_1_left [of 1, symmetric], rule mult_dvd_mono) 
+
+lemma unit_ring_inv:
+  "is_unit y \<Longrightarrow> x div y = x * ring_inv y"
+  by (simp add: div_mult_swap ring_inv_def is_unit_def)
+
+lemma unit_ring_inv_ring_inv [simp]:
+  "is_unit x \<Longrightarrow> ring_inv (ring_inv x) = x"
+  unfolding is_unit_def ring_inv_def
+  by (metis div_mult_mult1_if div_mult_self1_is_id dvd_mult_div_cancel mult_1_right)
+
+lemma inv_imp_eq_ring_inv:
+  "a * b = 1 \<Longrightarrow> ring_inv a = b"
+  by (metis dvd_mult_div_cancel dvd_mult_right mult_1_right mult.left_commute one_dvd ring_inv_def)
+
+lemma ring_inv_is_inv1 [simp]:
+  "is_unit a \<Longrightarrow> a * ring_inv a = 1"
+  unfolding is_unit_def ring_inv_def by (simp add: dvd_mult_div_cancel)
+
+lemma ring_inv_is_inv2 [simp]:
+  "is_unit a \<Longrightarrow> ring_inv a * a = 1"
+  by (simp add: ac_simps)
+
+lemma unit_ring_inv_unit [simp, intro]:
+  assumes "is_unit x"
+  shows "is_unit (ring_inv x)"
+proof -
+  from assms have "1 = ring_inv x * x" by simp
+  then show "is_unit (ring_inv x)" unfolding is_unit_def by (rule dvdI)
+qed
+
+lemma mult_unit_dvd_iff:
+  "is_unit y \<Longrightarrow> x * y dvd z \<longleftrightarrow> x dvd z"
+proof
+  assume "is_unit y" "x * y dvd z"
+  then show "x dvd z" by (simp add: dvd_mult_left)
+next
+  assume "is_unit y" "x dvd z"
+  then obtain k where "z = x * k" unfolding dvd_def by blast
+  with `is_unit y` have "z = (x * y) * (ring_inv y * k)" 
+      by (simp add: mult_ac)
+  then show "x * y dvd z" by (rule dvdI)
+qed
+
+lemma div_unit_dvd_iff:
+  "is_unit y \<Longrightarrow> x div y dvd z \<longleftrightarrow> x dvd z"
+  by (subst unit_ring_inv) (assumption, simp add: mult_unit_dvd_iff)
+
+lemma dvd_mult_unit_iff:
+  "is_unit y \<Longrightarrow> x dvd z * y \<longleftrightarrow> x dvd z"
+proof
+  assume "is_unit y" and "x dvd z * y"
+  have "z * y dvd z * (y * ring_inv y)" by (subst mult_assoc [symmetric]) simp
+  also from `is_unit y` have "y * ring_inv y = 1" by simp
+  finally have "z * y dvd z" by simp
+  with `x dvd z * y` show "x dvd z" by (rule dvd_trans)
+next
+  assume "x dvd z"
+  then show "x dvd z * y" by simp
+qed
+
+lemma dvd_div_unit_iff:
+  "is_unit y \<Longrightarrow> x dvd z div y \<longleftrightarrow> x dvd z"
+  by (subst unit_ring_inv) (assumption, simp add: dvd_mult_unit_iff)
+
+lemmas unit_dvd_iff = mult_unit_dvd_iff div_unit_dvd_iff dvd_mult_unit_iff dvd_div_unit_iff
+
+lemma unit_div [intro]:
+  "is_unit x \<Longrightarrow> is_unit y \<Longrightarrow> is_unit (x div y)"
+  by (subst unit_ring_inv) (assumption, rule unit_prod, simp_all)
+
+lemma unit_div_mult_swap:
+  "is_unit z \<Longrightarrow> x * (y div z) = x * y div z"
+  by (simp only: unit_ring_inv [of _ y] unit_ring_inv [of _ "x*y"] ac_simps)
+
+lemma unit_div_commute:
+  "is_unit y \<Longrightarrow> x div y * z = x * z div y"
+  by (simp only: unit_ring_inv [of _ x] unit_ring_inv [of _ "x*z"] ac_simps)
+
+lemma unit_imp_dvd [dest]:
+  "is_unit y \<Longrightarrow> y dvd x"
+  by (rule dvd_trans [of _ 1]) (simp_all add: is_unit_def)
+
+lemma dvd_unit_imp_unit:
+  "is_unit y \<Longrightarrow> x dvd y \<Longrightarrow> is_unit x"
+  by (unfold is_unit_def) (rule dvd_trans)
+
+lemma ring_inv_0 [simp]:
+  "ring_inv 0 = 0"
+  unfolding ring_inv_def by simp
+
+lemma unit_ring_inv'1:
+  assumes "is_unit y"
+  shows "x div (y * z) = x * ring_inv y div z" 
+proof -
+  from assms have "x div (y * z) = x * (ring_inv y * y) div (y * z)"
+    by simp
+  also have "... = y * (x * ring_inv y) div (y * z)"
+    by (simp only: mult_ac)
+  also have "... = x * ring_inv y div z"
+    by (cases "y = 0", simp, rule div_mult_mult1)
+  finally show ?thesis .
+qed
+
+lemma associated_comm:
+  "associated x y \<Longrightarrow> associated y x"
+  by (simp add: associated_def)
+
+lemma associated_0 [simp]:
+  "associated 0 b \<longleftrightarrow> b = 0"
+  "associated a 0 \<longleftrightarrow> a = 0"
+  unfolding associated_def by simp_all
+
+lemma associated_unit:
+  "is_unit x \<Longrightarrow> associated x y \<Longrightarrow> is_unit y"
+  unfolding associated_def by (fast dest: dvd_unit_imp_unit)
+
+lemma is_unit_1 [simp]:
+  "is_unit 1"
+  unfolding is_unit_def by simp
+
+lemma not_is_unit_0 [simp]:
+  "\<not> is_unit 0"
+  unfolding is_unit_def by auto
+
+lemma unit_mult_left_cancel:
+  assumes "is_unit x"
+  shows "(x * y) = (x * z) \<longleftrightarrow> y = z"
+proof -
+  from assms have "x \<noteq> 0" by auto
+  then show ?thesis by (metis div_mult_self1_is_id)
+qed
+
+lemma unit_mult_right_cancel:
+  "is_unit x \<Longrightarrow> (y * x) = (z * x) \<longleftrightarrow> y = z"
+  by (simp add: ac_simps unit_mult_left_cancel)
+
+lemma unit_div_cancel:
+  "is_unit x \<Longrightarrow> (y div x) = (z div x) \<longleftrightarrow> y = z"
+  apply (subst unit_ring_inv[of _ y], assumption)
+  apply (subst unit_ring_inv[of _ z], assumption)
+  apply (rule unit_mult_right_cancel, erule unit_ring_inv_unit)
+  done
+
+lemma unit_eq_div1:
+  "is_unit y \<Longrightarrow> x div y = z \<longleftrightarrow> x = z * y"
+  apply (subst unit_ring_inv, assumption)
+  apply (subst unit_mult_right_cancel[symmetric], assumption)
+  apply (subst mult_assoc, subst ring_inv_is_inv2, assumption, simp)
+  done
+
+lemma unit_eq_div2:
+  "is_unit y \<Longrightarrow> x = z div y \<longleftrightarrow> x * y = z"
+  by (subst (1 2) eq_commute, simp add: unit_eq_div1, subst eq_commute, rule refl)
+
+lemma associated_iff_div_unit:
+  "associated x y \<longleftrightarrow> (\<exists>z. is_unit z \<and> x = z * y)"
+proof
+  assume "associated x y"
+  show "\<exists>z. is_unit z \<and> x = z * y"
+  proof (cases "x = 0")
+    assume "x = 0"
+    then show "\<exists>z. is_unit z \<and> x = z * y" using `associated x y`
+        by (intro exI[of _ 1], simp add: associated_def)
+  next
+    assume [simp]: "x \<noteq> 0"
+    hence [simp]: "x dvd y" "y dvd x" using `associated x y`
+        unfolding associated_def by simp_all
+    hence "1 = x div y * (y div x)"
+      by (simp add: div_mult_swap dvd_div_mult_self)
+    hence "is_unit (x div y)" unfolding is_unit_def by (rule dvdI)
+    moreover have "x = (x div y) * y" by (simp add: dvd_div_mult_self)
+    ultimately show ?thesis by blast
+  qed
+next
+  assume "\<exists>z. is_unit z \<and> x = z * y"
+  then obtain z where "is_unit z" and "x = z * y" by blast
+  hence "y = x * ring_inv z" by (simp add: algebra_simps)
+  hence "x dvd y" by simp
+  moreover from `x = z * y` have "y dvd x" by simp
+  ultimately show "associated x y" unfolding associated_def by simp
+qed
+
+lemmas unit_simps = mult_unit_dvd_iff div_unit_dvd_iff dvd_mult_unit_iff 
+  dvd_div_unit_iff unit_div_mult_swap unit_div_commute
+  unit_mult_left_cancel unit_mult_right_cancel unit_div_cancel 
+  unit_eq_div1 unit_eq_div2
+
+end
+
+context ring_div
+begin
+
+lemma is_unit_neg [simp]:
+  "is_unit (- x) \<Longrightarrow> is_unit x"
+  unfolding is_unit_def by simp
+
+lemma is_unit_neg_1 [simp]:
+  "is_unit (-1)"
+  unfolding is_unit_def by simp
+
+end
+
+lemma is_unit_nat [simp]:
+  "is_unit (x::nat) \<longleftrightarrow> x = 1"
+  unfolding is_unit_def by simp
+
+lemma is_unit_int:
+  "is_unit (x::int) \<longleftrightarrow> x = 1 \<or> x = -1"
+  unfolding is_unit_def by auto
+
+text {*
+  A Euclidean semiring is a semiring upon which the Euclidean algorithm can be
+  implemented. It must provide:
+  \begin{itemize}
+  \item division with remainder
+  \item a size function such that @{term "size (a mod b) < size b"} 
+        for any @{term "b \<noteq> 0"}
+  \item a normalisation factor such that two associated numbers are equal iff 
+        they are the same when divided by their normalisation factors.
+  \end{itemize}
+  The existence of these functions makes it possible to derive gcd and lcm functions 
+  for any Euclidean semiring.
+*} 
+class euclidean_semiring = semiring_div + 
+  fixes euclidean_size :: "'a \<Rightarrow> nat"
+  fixes normalisation_factor :: "'a \<Rightarrow> 'a"
+  assumes mod_size_less [simp]: 
+    "b \<noteq> 0 \<Longrightarrow> euclidean_size (a mod b) < euclidean_size b"
+  assumes size_mult_mono:
+    "b \<noteq> 0 \<Longrightarrow> euclidean_size (a * b) \<ge> euclidean_size a"
+  assumes normalisation_factor_is_unit [intro,simp]: 
+    "a \<noteq> 0 \<Longrightarrow> is_unit (normalisation_factor a)"
+  assumes normalisation_factor_mult: "normalisation_factor (a * b) = 
+    normalisation_factor a * normalisation_factor b"
+  assumes normalisation_factor_unit: "is_unit x \<Longrightarrow> normalisation_factor x = x"
+  assumes normalisation_factor_0 [simp]: "normalisation_factor 0 = 0"
+begin
+
+lemma normalisation_factor_dvd [simp]:
+  "a \<noteq> 0 \<Longrightarrow> normalisation_factor a dvd b"
+  by (rule unit_imp_dvd, simp)
+    
+lemma normalisation_factor_1 [simp]:
+  "normalisation_factor 1 = 1"
+  by (simp add: normalisation_factor_unit)
+
+lemma normalisation_factor_0_iff [simp]:
+  "normalisation_factor x = 0 \<longleftrightarrow> x = 0"
+proof
+  assume "normalisation_factor x = 0"
+  hence "\<not> is_unit (normalisation_factor x)"
+    by (metis not_is_unit_0)
+  then show "x = 0" by force
+next
+  assume "x = 0"
+  then show "normalisation_factor x = 0" by simp
+qed
+
+lemma normalisation_factor_pow:
+  "normalisation_factor (x ^ n) = normalisation_factor x ^ n"
+  by (induct n) (simp_all add: normalisation_factor_mult power_Suc2)
+
+lemma normalisation_correct [simp]:
+  "normalisation_factor (x div normalisation_factor x) = (if x = 0 then 0 else 1)"
+proof (cases "x = 0", simp)
+  assume "x \<noteq> 0"
+  let ?nf = "normalisation_factor"
+  from normalisation_factor_is_unit[OF `x \<noteq> 0`] have "?nf x \<noteq> 0"
+    by (metis not_is_unit_0) 
+  have "?nf (x div ?nf x) * ?nf (?nf x) = ?nf (x div ?nf x * ?nf x)" 
+    by (simp add: normalisation_factor_mult)
+  also have "x div ?nf x * ?nf x = x" using `x \<noteq> 0`
+    by (simp add: dvd_div_mult_self)
+  also have "?nf (?nf x) = ?nf x" using `x \<noteq> 0` 
+    normalisation_factor_is_unit normalisation_factor_unit by simp
+  finally show ?thesis using `x \<noteq> 0` and `?nf x \<noteq> 0` 
+    by (metis div_mult_self2_is_id div_self)
+qed
+
+lemma normalisation_0_iff [simp]:
+  "x div normalisation_factor x = 0 \<longleftrightarrow> x = 0"
+  by (cases "x = 0", simp, subst unit_eq_div1, blast, simp)
+
+lemma associated_iff_normed_eq:
+  "associated a b \<longleftrightarrow> a div normalisation_factor a = b div normalisation_factor b"
+proof (cases "b = 0", simp, cases "a = 0", metis associated_0(1) normalisation_0_iff, rule iffI)
+  let ?nf = normalisation_factor
+  assume "a \<noteq> 0" "b \<noteq> 0" "a div ?nf a = b div ?nf b"
+  hence "a = b * (?nf a div ?nf b)"
+    apply (subst (asm) unit_eq_div1, blast, subst (asm) unit_div_commute, blast)
+    apply (subst div_mult_swap, simp, simp)
+    done
+  with `a \<noteq> 0` `b \<noteq> 0` have "\<exists>z. is_unit z \<and> a = z * b"
+    by (intro exI[of _ "?nf a div ?nf b"], force simp: mult_ac)
+  with associated_iff_div_unit show "associated a b" by simp
+next
+  let ?nf = normalisation_factor
+  assume "a \<noteq> 0" "b \<noteq> 0" "associated a b"
+  with associated_iff_div_unit obtain z where "is_unit z" and "a = z * b" by blast
+  then show "a div ?nf a = b div ?nf b"
+    apply (simp only: `a = z * b` normalisation_factor_mult normalisation_factor_unit)
+    apply (rule div_mult_mult1, force)
+    done
+  qed
+
+lemma normed_associated_imp_eq:
+  "associated a b \<Longrightarrow> normalisation_factor a \<in> {0, 1} \<Longrightarrow> normalisation_factor b \<in> {0, 1} \<Longrightarrow> a = b"
+  by (simp add: associated_iff_normed_eq, elim disjE, simp_all)
+    
+lemmas normalisation_factor_dvd_iff [simp] =
+  unit_dvd_iff [OF normalisation_factor_is_unit]
+
+lemma euclidean_division:
+  fixes a :: 'a and b :: 'a
+  assumes "b \<noteq> 0"
+  obtains s and t where "a = s * b + t" 
+    and "euclidean_size t < euclidean_size b"
+proof -
+  from div_mod_equality[of a b 0] 
+     have "a = a div b * b + a mod b" by simp
+  with that and assms show ?thesis by force
+qed
+
+lemma dvd_euclidean_size_eq_imp_dvd:
+  assumes "a \<noteq> 0" and b_dvd_a: "b dvd a" and size_eq: "euclidean_size a = euclidean_size b"
+  shows "a dvd b"
+proof (subst dvd_eq_mod_eq_0, rule ccontr)
+  assume "b mod a \<noteq> 0"
+  from b_dvd_a have b_dvd_mod: "b dvd b mod a" by (simp add: dvd_mod_iff)
+  from b_dvd_mod obtain c where "b mod a = b * c" unfolding dvd_def by blast
+    with `b mod a \<noteq> 0` have "c \<noteq> 0" by auto
+  with `b mod a = b * c` have "euclidean_size (b mod a) \<ge> euclidean_size b"
+      using size_mult_mono by force
+  moreover from `a \<noteq> 0` have "euclidean_size (b mod a) < euclidean_size a"
+      using mod_size_less by blast
+  ultimately show False using size_eq by simp
+qed
+
+function gcd_eucl :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+where
+  "gcd_eucl a b = (if b = 0 then a div normalisation_factor a else gcd_eucl b (a mod b))"
+  by (pat_completeness, simp)
+termination by (relation "measure (euclidean_size \<circ> snd)", simp_all)
+
+declare gcd_eucl.simps [simp del]
+
+lemma gcd_induct: "\<lbrakk>\<And>b. P b 0; \<And>a b. 0 \<noteq> b \<Longrightarrow> P b (a mod b) \<Longrightarrow> P a b\<rbrakk> \<Longrightarrow> P a b"
+proof (induct a b rule: gcd_eucl.induct)
+  case ("1" m n)
+    then show ?case by (cases "n = 0") auto
+qed
+
+definition lcm_eucl :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+where
+  "lcm_eucl a b = a * b div (gcd_eucl a b * normalisation_factor (a * b))"
+
+  (* Somewhat complicated definition of Lcm that has the advantage of working
+     for infinite sets as well *)
+
+definition Lcm_eucl :: "'a set \<Rightarrow> 'a"
+where
+  "Lcm_eucl A = (if \<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) then
+     let l = SOME l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l =
+       (LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n)
+       in l div normalisation_factor l
+      else 0)"
+
+definition Gcd_eucl :: "'a set \<Rightarrow> 'a"
+where
+  "Gcd_eucl A = Lcm_eucl {d. \<forall>a\<in>A. d dvd a}"
+
+end
+
+class euclidean_semiring_gcd = euclidean_semiring + gcd + Gcd +
+  assumes gcd_gcd_eucl: "gcd = gcd_eucl" and lcm_lcm_eucl: "lcm = lcm_eucl"
+  assumes Gcd_Gcd_eucl: "Gcd = Gcd_eucl" and Lcm_Lcm_eucl: "Lcm = Lcm_eucl"
+begin
+
+lemma gcd_red:
+  "gcd x y = gcd y (x mod y)"
+  by (metis gcd_eucl.simps mod_0 mod_by_0 gcd_gcd_eucl)
+
+lemma gcd_non_0:
+  "y \<noteq> 0 \<Longrightarrow> gcd x y = gcd y (x mod y)"
+  by (rule gcd_red)
+
+lemma gcd_0_left:
+  "gcd 0 x = x div normalisation_factor x"
+   by (simp only: gcd_gcd_eucl, subst gcd_eucl.simps, subst gcd_eucl.simps, simp add: Let_def)
+
+lemma gcd_0:
+  "gcd x 0 = x div normalisation_factor x"
+  by (simp only: gcd_gcd_eucl, subst gcd_eucl.simps, simp add: Let_def)
+
+lemma gcd_dvd1 [iff]: "gcd x y dvd x"
+  and gcd_dvd2 [iff]: "gcd x y dvd y"
+proof (induct x y rule: gcd_eucl.induct)
+  fix x y :: 'a
+  assume IH1: "y \<noteq> 0 \<Longrightarrow> gcd y (x mod y) dvd y"
+  assume IH2: "y \<noteq> 0 \<Longrightarrow> gcd y (x mod y) dvd (x mod y)"
+  
+  have "gcd x y dvd x \<and> gcd x y dvd y"
+  proof (cases "y = 0")
+    case True
+      then show ?thesis by (cases "x = 0", simp_all add: gcd_0)
+  next
+    case False
+      with IH1 and IH2 show ?thesis by (simp add: gcd_non_0 dvd_mod_iff)
+  qed
+  then show "gcd x y dvd x" "gcd x y dvd y" by simp_all
+qed
+
+lemma dvd_gcd_D1: "k dvd gcd m n \<Longrightarrow> k dvd m"
+  by (rule dvd_trans, assumption, rule gcd_dvd1)
+
+lemma dvd_gcd_D2: "k dvd gcd m n \<Longrightarrow> k dvd n"
+  by (rule dvd_trans, assumption, rule gcd_dvd2)
+
+lemma gcd_greatest:
+  fixes k x y :: 'a
+  shows "k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> k dvd gcd x y"
+proof (induct x y rule: gcd_eucl.induct)
+  case (1 x y)
+  show ?case
+    proof (cases "y = 0")
+      assume "y = 0"
+      with 1 show ?thesis by (cases "x = 0", simp_all add: gcd_0)
+    next
+      assume "y \<noteq> 0"
+      with 1 show ?thesis by (simp add: gcd_non_0 dvd_mod_iff) 
+    qed
+qed
+
+lemma dvd_gcd_iff:
+  "k dvd gcd x y \<longleftrightarrow> k dvd x \<and> k dvd y"
+  by (blast intro!: gcd_greatest intro: dvd_trans)
+
+lemmas gcd_greatest_iff = dvd_gcd_iff
+
+lemma gcd_zero [simp]:
+  "gcd x y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
+  by (metis dvd_0_left dvd_refl gcd_dvd1 gcd_dvd2 gcd_greatest)+
+
+lemma normalisation_factor_gcd [simp]:
+  "normalisation_factor (gcd x y) = (if x = 0 \<and> y = 0 then 0 else 1)" (is "?f x y = ?g x y")
+proof (induct x y rule: gcd_eucl.induct)
+  fix x y :: 'a
+  assume IH: "y \<noteq> 0 \<Longrightarrow> ?f y (x mod y) = ?g y (x mod y)"
+  then show "?f x y = ?g x y" by (cases "y = 0", auto simp: gcd_non_0 gcd_0)
+qed
+
+lemma gcdI:
+  "k dvd x \<Longrightarrow> k dvd y \<Longrightarrow> (\<And>l. l dvd x \<Longrightarrow> l dvd y \<Longrightarrow> l dvd k)
+    \<Longrightarrow> normalisation_factor k = (if k = 0 then 0 else 1) \<Longrightarrow> k = gcd x y"
+  by (intro normed_associated_imp_eq) (auto simp: associated_def intro: gcd_greatest)
+
+sublocale gcd!: abel_semigroup gcd
+proof
+  fix x y z 
+  show "gcd (gcd x y) z = gcd x (gcd y z)"
+  proof (rule gcdI)
+    have "gcd (gcd x y) z dvd gcd x y" "gcd x y dvd x" by simp_all
+    then show "gcd (gcd x y) z dvd x" by (rule dvd_trans)
+    have "gcd (gcd x y) z dvd gcd x y" "gcd x y dvd y" by simp_all
+    hence "gcd (gcd x y) z dvd y" by (rule dvd_trans)
+    moreover have "gcd (gcd x y) z dvd z" by simp
+    ultimately show "gcd (gcd x y) z dvd gcd y z"
+      by (rule gcd_greatest)
+    show "normalisation_factor (gcd (gcd x y) z) =  (if gcd (gcd x y) z = 0 then 0 else 1)"
+      by auto
+    fix l assume "l dvd x" and "l dvd gcd y z"
+    with dvd_trans[OF _ gcd_dvd1] and dvd_trans[OF _ gcd_dvd2]
+      have "l dvd y" and "l dvd z" by blast+
+    with `l dvd x` show "l dvd gcd (gcd x y) z"
+      by (intro gcd_greatest)
+  qed
+next
+  fix x y
+  show "gcd x y = gcd y x"
+    by (rule gcdI) (simp_all add: gcd_greatest)
+qed
+
+lemma gcd_unique: "d dvd a \<and> d dvd b \<and> 
+    normalisation_factor d = (if d = 0 then 0 else 1) \<and>
+    (\<forall>e. e dvd a \<and> e dvd b \<longrightarrow> e dvd d) \<longleftrightarrow> d = gcd a b"
+  by (rule, auto intro: gcdI simp: gcd_greatest)
+
+lemma gcd_dvd_prod: "gcd a b dvd k * b"
+  using mult_dvd_mono [of 1] by auto
+
+lemma gcd_1_left [simp]: "gcd 1 x = 1"
+  by (rule sym, rule gcdI, simp_all)
+
+lemma gcd_1 [simp]: "gcd x 1 = 1"
+  by (rule sym, rule gcdI, simp_all)
+
+lemma gcd_proj2_if_dvd: 
+  "y dvd x \<Longrightarrow> gcd x y = y div normalisation_factor y"
+  by (cases "y = 0", simp_all add: dvd_eq_mod_eq_0 gcd_non_0 gcd_0)
+
+lemma gcd_proj1_if_dvd: 
+  "x dvd y \<Longrightarrow> gcd x y = x div normalisation_factor x"
+  by (subst gcd.commute, simp add: gcd_proj2_if_dvd)
+
+lemma gcd_proj1_iff: "gcd m n = m div normalisation_factor m \<longleftrightarrow> m dvd n"
+proof
+  assume A: "gcd m n = m div normalisation_factor m"
+  show "m dvd n"
+  proof (cases "m = 0")
+    assume [simp]: "m \<noteq> 0"
+    from A have B: "m = gcd m n * normalisation_factor m"
+      by (simp add: unit_eq_div2)
+    show ?thesis by (subst B, simp add: mult_unit_dvd_iff)
+  qed (insert A, simp)
+next
+  assume "m dvd n"
+  then show "gcd m n = m div normalisation_factor m" by (rule gcd_proj1_if_dvd)
+qed
+  
+lemma gcd_proj2_iff: "gcd m n = n div normalisation_factor n \<longleftrightarrow> n dvd m"
+  by (subst gcd.commute, simp add: gcd_proj1_iff)
+
+lemma gcd_mod1 [simp]:
+  "gcd (x mod y) y = gcd x y"
+  by (rule gcdI, metis dvd_mod_iff gcd_dvd1 gcd_dvd2, simp_all add: gcd_greatest dvd_mod_iff)
+
+lemma gcd_mod2 [simp]:
+  "gcd x (y mod x) = gcd x y"
+  by (rule gcdI, simp, metis dvd_mod_iff gcd_dvd1 gcd_dvd2, simp_all add: gcd_greatest dvd_mod_iff)
+         
+lemma normalisation_factor_dvd' [simp]:
+  "normalisation_factor x dvd x"
+  by (cases "x = 0", simp_all)
+
+lemma gcd_mult_distrib': 
+  "k div normalisation_factor k * gcd x y = gcd (k*x) (k*y)"
+proof (induct x y rule: gcd_eucl.induct)
+  case (1 x y)
+  show ?case
+  proof (cases "y = 0")
+    case True
+    then show ?thesis by (simp add: normalisation_factor_mult gcd_0 algebra_simps div_mult_div_if_dvd)
+  next
+    case False
+    hence "k div normalisation_factor k * gcd x y =  gcd (k * y) (k * (x mod y))" 
+      using 1 by (subst gcd_red, simp)
+    also have "... = gcd (k * x) (k * y)"
+      by (simp add: mult_mod_right gcd.commute)
+    finally show ?thesis .
+  qed
+qed
+
+lemma gcd_mult_distrib:
+  "k * gcd x y = gcd (k*x) (k*y) * normalisation_factor k"
+proof-
+  let ?nf = "normalisation_factor"
+  from gcd_mult_distrib' 
+    have "gcd (k*x) (k*y) = k div ?nf k * gcd x y" ..
+  also have "... = k * gcd x y div ?nf k"
+    by (metis dvd_div_mult dvd_eq_mod_eq_0 mod_0 normalisation_factor_dvd)
+  finally show ?thesis
+    by (simp add: ac_simps dvd_mult_div_cancel)
+qed
+
+lemma euclidean_size_gcd_le1 [simp]:
+  assumes "a \<noteq> 0"
+  shows "euclidean_size (gcd a b) \<le> euclidean_size a"
+proof -
+   have "gcd a b dvd a" by (rule gcd_dvd1)
+   then obtain c where A: "a = gcd a b * c" unfolding dvd_def by blast
+   with `a \<noteq> 0` show ?thesis by (subst (2) A, intro size_mult_mono) auto
+qed
+
+lemma euclidean_size_gcd_le2 [simp]:
+  "b \<noteq> 0 \<Longrightarrow> euclidean_size (gcd a b) \<le> euclidean_size b"
+  by (subst gcd.commute, rule euclidean_size_gcd_le1)
+
+lemma euclidean_size_gcd_less1:
+  assumes "a \<noteq> 0" and "\<not>a dvd b"
+  shows "euclidean_size (gcd a b) < euclidean_size a"
+proof (rule ccontr)
+  assume "\<not>euclidean_size (gcd a b) < euclidean_size a"
+  with `a \<noteq> 0` have "euclidean_size (gcd a b) = euclidean_size a"
+    by (intro le_antisym, simp_all)
+  with assms have "a dvd gcd a b" by (auto intro: dvd_euclidean_size_eq_imp_dvd)
+  hence "a dvd b" using dvd_gcd_D2 by blast
+  with `\<not>a dvd b` show False by contradiction
+qed
+
+lemma euclidean_size_gcd_less2:
+  assumes "b \<noteq> 0" and "\<not>b dvd a"
+  shows "euclidean_size (gcd a b) < euclidean_size b"
+  using assms by (subst gcd.commute, rule euclidean_size_gcd_less1)
+
+lemma gcd_mult_unit1: "is_unit a \<Longrightarrow> gcd (x*a) y = gcd x y"
+  apply (rule gcdI)
+  apply (rule dvd_trans, rule gcd_dvd1, simp add: unit_simps)
+  apply (rule gcd_dvd2)
+  apply (rule gcd_greatest, simp add: unit_simps, assumption)
+  apply (subst normalisation_factor_gcd, simp add: gcd_0)
+  done
+
+lemma gcd_mult_unit2: "is_unit a \<Longrightarrow> gcd x (y*a) = gcd x y"
+  by (subst gcd.commute, subst gcd_mult_unit1, assumption, rule gcd.commute)
+
+lemma gcd_div_unit1: "is_unit a \<Longrightarrow> gcd (x div a) y = gcd x y"
+  by (simp add: unit_ring_inv gcd_mult_unit1)
+
+lemma gcd_div_unit2: "is_unit a \<Longrightarrow> gcd x (y div a) = gcd x y"
+  by (simp add: unit_ring_inv gcd_mult_unit2)
+
+lemma gcd_idem: "gcd x x = x div normalisation_factor x"
+  by (cases "x = 0") (simp add: gcd_0_left, rule sym, rule gcdI, simp_all)
+
+lemma gcd_right_idem: "gcd (gcd p q) q = gcd p q"
+  apply (rule gcdI)
+  apply (simp add: ac_simps)
+  apply (rule gcd_dvd2)
+  apply (rule gcd_greatest, erule (1) gcd_greatest, assumption)
+  apply (simp add: gcd_zero)
+  done
+
+lemma gcd_left_idem: "gcd p (gcd p q) = gcd p q"
+  apply (rule gcdI)
+  apply simp
+  apply (rule dvd_trans, rule gcd_dvd2, rule gcd_dvd2)
+  apply (rule gcd_greatest, assumption, erule gcd_greatest, assumption)
+  apply (simp add: gcd_zero)
+  done
+
+lemma comp_fun_idem_gcd: "comp_fun_idem gcd"
+proof
+  fix a b show "gcd a \<circ> gcd b = gcd b \<circ> gcd a"
+    by (simp add: fun_eq_iff ac_simps)
+next
+  fix a show "gcd a \<circ> gcd a = gcd a"
+    by (simp add: fun_eq_iff gcd_left_idem)
+qed
+
+lemma coprime_dvd_mult:
+  assumes "gcd k n = 1" and "k dvd m * n"
+  shows "k dvd m"
+proof -
+  let ?nf = "normalisation_factor"
+  from assms gcd_mult_distrib [of m k n] 
+    have A: "m = gcd (m * k) (m * n) * ?nf m" by simp
+  from `k dvd m * n` show ?thesis by (subst A, simp_all add: gcd_greatest)
+qed
+
+lemma coprime_dvd_mult_iff:
+  "gcd k n = 1 \<Longrightarrow> (k dvd m * n) = (k dvd m)"
+  by (rule, rule coprime_dvd_mult, simp_all)
+
+lemma gcd_dvd_antisym:
+  "gcd a b dvd gcd c d \<Longrightarrow> gcd c d dvd gcd a b \<Longrightarrow> gcd a b = gcd c d"
+proof (rule gcdI)
+  assume A: "gcd a b dvd gcd c d" and B: "gcd c d dvd gcd a b"
+  have "gcd c d dvd c" by simp
+  with A show "gcd a b dvd c" by (rule dvd_trans)
+  have "gcd c d dvd d" by simp
+  with A show "gcd a b dvd d" by (rule dvd_trans)
+  show "normalisation_factor (gcd a b) = (if gcd a b = 0 then 0 else 1)"
+    by (simp add: gcd_zero)
+  fix l assume "l dvd c" and "l dvd d"
+  hence "l dvd gcd c d" by (rule gcd_greatest)
+  from this and B show "l dvd gcd a b" by (rule dvd_trans)
+qed
+
+lemma gcd_mult_cancel:
+  assumes "gcd k n = 1"
+  shows "gcd (k * m) n = gcd m n"
+proof (rule gcd_dvd_antisym)
+  have "gcd (gcd (k * m) n) k = gcd (gcd k n) (k * m)" by (simp add: ac_simps)
+  also note `gcd k n = 1`
+  finally have "gcd (gcd (k * m) n) k = 1" by simp
+  hence "gcd (k * m) n dvd m" by (rule coprime_dvd_mult, simp add: ac_simps)
+  moreover have "gcd (k * m) n dvd n" by simp
+  ultimately show "gcd (k * m) n dvd gcd m n" by (rule gcd_greatest)
+  have "gcd m n dvd (k * m)" and "gcd m n dvd n" by simp_all
+  then show "gcd m n dvd gcd (k * m) n" by (rule gcd_greatest)
+qed
+
+lemma coprime_crossproduct:
+  assumes [simp]: "gcd a d = 1" "gcd b c = 1"
+  shows "associated (a * c) (b * d) \<longleftrightarrow> associated a b \<and> associated c d" (is "?lhs \<longleftrightarrow> ?rhs")
+proof
+  assume ?rhs then show ?lhs unfolding associated_def by (fast intro: mult_dvd_mono)
+next
+  assume ?lhs
+  from `?lhs` have "a dvd b * d" unfolding associated_def by (metis dvd_mult_left) 
+  hence "a dvd b" by (simp add: coprime_dvd_mult_iff)
+  moreover from `?lhs` have "b dvd a * c" unfolding associated_def by (metis dvd_mult_left) 
+  hence "b dvd a" by (simp add: coprime_dvd_mult_iff)
+  moreover from `?lhs` have "c dvd d * b" 
+    unfolding associated_def by (metis dvd_mult_right ac_simps)
+  hence "c dvd d" by (simp add: coprime_dvd_mult_iff gcd.commute)
+  moreover from `?lhs` have "d dvd c * a"
+    unfolding associated_def by (metis dvd_mult_right ac_simps)
+  hence "d dvd c" by (simp add: coprime_dvd_mult_iff gcd.commute)
+  ultimately show ?rhs unfolding associated_def by simp
+qed
+
+lemma gcd_add1 [simp]:
+  "gcd (m + n) n = gcd m n"
+  by (cases "n = 0", simp_all add: gcd_non_0)
+
+lemma gcd_add2 [simp]:
+  "gcd m (m + n) = gcd m n"
+  using gcd_add1 [of n m] by (simp add: ac_simps)
+
+lemma gcd_add_mult: "gcd m (k * m + n) = gcd m n"
+  by (subst gcd.commute, subst gcd_red, simp)
+
+lemma coprimeI: "(\<And>l. \<lbrakk>l dvd x; l dvd y\<rbrakk> \<Longrightarrow> l dvd 1) \<Longrightarrow> gcd x y = 1"
+  by (rule sym, rule gcdI, simp_all)
+
+lemma coprime: "gcd a b = 1 \<longleftrightarrow> (\<forall>d. d dvd a \<and> d dvd b \<longleftrightarrow> is_unit d)"
+  by (auto simp: is_unit_def intro: coprimeI gcd_greatest dvd_gcd_D1 dvd_gcd_D2)
+
+lemma div_gcd_coprime:
+  assumes nz: "a \<noteq> 0 \<or> b \<noteq> 0"
+  defines [simp]: "d \<equiv> gcd a b"
+  defines [simp]: "a' \<equiv> a div d" and [simp]: "b' \<equiv> b div d"
+  shows "gcd a' b' = 1"
+proof (rule coprimeI)
+  fix l assume "l dvd a'" "l dvd b'"
+  then obtain s t where "a' = l * s" "b' = l * t" unfolding dvd_def by blast
+  moreover have "a = a' * d" "b = b' * d" by (simp_all add: dvd_div_mult_self)
+  ultimately have "a = (l * d) * s" "b = (l * d) * t"
+    by (metis ac_simps)+
+  hence "l*d dvd a" and "l*d dvd b" by (simp_all only: dvd_triv_left)
+  hence "l*d dvd d" by (simp add: gcd_greatest)
+  then obtain u where "u * l * d = d" unfolding dvd_def
+    by (metis ac_simps mult_assoc)
+  moreover from nz have "d \<noteq> 0" by (simp add: gcd_zero)
+  ultimately have "u * l = 1" 
+    by (metis div_mult_self1_is_id div_self ac_simps)
+  then show "l dvd 1" by force
+qed
+
+lemma coprime_mult: 
+  assumes da: "gcd d a = 1" and db: "gcd d b = 1"
+  shows "gcd d (a * b) = 1"
+  apply (subst gcd.commute)
+  using da apply (subst gcd_mult_cancel)
+  apply (subst gcd.commute, assumption)
+  apply (subst gcd.commute, rule db)
+  done
+
+lemma coprime_lmult:
+  assumes dab: "gcd d (a * b) = 1" 
+  shows "gcd d a = 1"
+proof (rule coprimeI)
+  fix l assume "l dvd d" and "l dvd a"
+  hence "l dvd a * b" by simp
+  with `l dvd d` and dab show "l dvd 1" by (auto intro: gcd_greatest)
+qed
+
+lemma coprime_rmult:
+  assumes dab: "gcd d (a * b) = 1"
+  shows "gcd d b = 1"
+proof (rule coprimeI)
+  fix l assume "l dvd d" and "l dvd b"
+  hence "l dvd a * b" by simp
+  with `l dvd d` and dab show "l dvd 1" by (auto intro: gcd_greatest)
+qed
+
+lemma coprime_mul_eq: "gcd d (a * b) = 1 \<longleftrightarrow> gcd d a = 1 \<and> gcd d b = 1"
+  using coprime_rmult[of d a b] coprime_lmult[of d a b] coprime_mult[of d a b] by blast
+
+lemma gcd_coprime:
+  assumes z: "gcd a b \<noteq> 0" and a: "a = a' * gcd a b" and b: "b = b' * gcd a b"
+  shows "gcd a' b' = 1"
+proof -
+  from z have "a \<noteq> 0 \<or> b \<noteq> 0" by (simp add: gcd_zero)
+  with div_gcd_coprime have "gcd (a div gcd a b) (b div gcd a b) = 1" .
+  also from assms have "a div gcd a b = a'" by (metis div_mult_self2_is_id)+
+  also from assms have "b div gcd a b = b'" by (metis div_mult_self2_is_id)+
+  finally show ?thesis .
+qed
+
+lemma coprime_power:
+  assumes "0 < n"
+  shows "gcd a (b ^ n) = 1 \<longleftrightarrow> gcd a b = 1"
+using assms proof (induct n)
+  case (Suc n) then show ?case
+    by (cases n) (simp_all add: coprime_mul_eq)
+qed simp
+
+lemma gcd_coprime_exists:
+  assumes nz: "gcd a b \<noteq> 0"
+  shows "\<exists>a' b'. a = a' * gcd a b \<and> b = b' * gcd a b \<and> gcd a' b' = 1"
+  apply (rule_tac x = "a div gcd a b" in exI)
+  apply (rule_tac x = "b div gcd a b" in exI)
+  apply (insert nz, auto simp add: dvd_div_mult gcd_0_left  gcd_zero intro: div_gcd_coprime)
+  done
+
+lemma coprime_exp:
+  "gcd d a = 1 \<Longrightarrow> gcd d (a^n) = 1"
+  by (induct n, simp_all add: coprime_mult)
+
+lemma coprime_exp2 [intro]:
+  "gcd a b = 1 \<Longrightarrow> gcd (a^n) (b^m) = 1"
+  apply (rule coprime_exp)
+  apply (subst gcd.commute)
+  apply (rule coprime_exp)
+  apply (subst gcd.commute)
+  apply assumption
+  done
+
+lemma gcd_exp:
+  "gcd (a^n) (b^n) = (gcd a b) ^ n"
+proof (cases "a = 0 \<and> b = 0")
+  assume "a = 0 \<and> b = 0"
+  then show ?thesis by (cases n, simp_all add: gcd_0_left)
+next
+  assume A: "\<not>(a = 0 \<and> b = 0)"
+  hence "1 = gcd ((a div gcd a b)^n) ((b div gcd a b)^n)"
+    using div_gcd_coprime by (subst sym, auto simp: div_gcd_coprime)
+  hence "(gcd a b) ^ n = (gcd a b) ^ n * ..." by simp
+  also note gcd_mult_distrib
+  also have "normalisation_factor ((gcd a b)^n) = 1"
+    by (simp add: normalisation_factor_pow A)
+  also have "(gcd a b)^n * (a div gcd a b)^n = a^n"
+    by (subst ac_simps, subst div_power, simp, rule dvd_div_mult_self, rule dvd_power_same, simp)
+  also have "(gcd a b)^n * (b div gcd a b)^n = b^n"
+    by (subst ac_simps, subst div_power, simp, rule dvd_div_mult_self, rule dvd_power_same, simp)
+  finally show ?thesis by simp
+qed
+
+lemma coprime_common_divisor: 
+  "gcd a b = 1 \<Longrightarrow> x dvd a \<Longrightarrow> x dvd b \<Longrightarrow> is_unit x"
+  apply (subgoal_tac "x dvd gcd a b")
+  apply (simp add: is_unit_def)
+  apply (erule (1) gcd_greatest)
+  done
+
+lemma division_decomp: 
+  assumes dc: "a dvd b * c"
+  shows "\<exists>b' c'. a = b' * c' \<and> b' dvd b \<and> c' dvd c"
+proof (cases "gcd a b = 0")
+  assume "gcd a b = 0"
+  hence "a = 0 \<and> b = 0" by (simp add: gcd_zero)
+  hence "a = 0 * c \<and> 0 dvd b \<and> c dvd c" by simp
+  then show ?thesis by blast
+next
+  let ?d = "gcd a b"
+  assume "?d \<noteq> 0"
+  from gcd_coprime_exists[OF this]
+    obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "gcd a' b' = 1"
+    by blast
+  from ab'(1) have "a' dvd a" unfolding dvd_def by blast
+  with dc have "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp
+  from dc ab'(1,2) have "a'*?d dvd (b'*?d) * c" by simp
+  hence "?d * a' dvd ?d * (b' * c)" by (simp add: mult_ac)
+  with `?d \<noteq> 0` have "a' dvd b' * c" by (rule dvd_mult_cancel_left)
+  with coprime_dvd_mult[OF ab'(3)] 
+    have "a' dvd c" by (subst (asm) ac_simps, blast)
+  with ab'(1) have "a = ?d * a' \<and> ?d dvd b \<and> a' dvd c" by (simp add: mult_ac)
+  then show ?thesis by blast
+qed
+
+lemma pow_divides_pow:
+  assumes ab: "a ^ n dvd b ^ n" and n: "n \<noteq> 0"
+  shows "a dvd b"
+proof (cases "gcd a b = 0")
+  assume "gcd a b = 0"
+  then show ?thesis by (simp add: gcd_zero)
+next
+  let ?d = "gcd a b"
+  assume "?d \<noteq> 0"
+  from n obtain m where m: "n = Suc m" by (cases n, simp_all)
+  from `?d \<noteq> 0` have zn: "?d ^ n \<noteq> 0" by (rule nonzero_pow_nonzero)
+  from gcd_coprime_exists[OF `?d \<noteq> 0`]
+    obtain a' b' where ab': "a = a' * ?d" "b = b' * ?d" "gcd a' b' = 1"
+    by blast
+  from ab have "(a' * ?d) ^ n dvd (b' * ?d) ^ n"
+    by (simp add: ab'(1,2)[symmetric])
+  hence "?d^n * a'^n dvd ?d^n * b'^n"
+    by (simp only: power_mult_distrib ac_simps)
+  with zn have "a'^n dvd b'^n" by (rule dvd_mult_cancel_left)
+  hence "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by (simp add: m)
+  hence "a' dvd b'^m * b'" by (simp add: m ac_simps)
+  with coprime_dvd_mult[OF coprime_exp[OF ab'(3), of m]]
+    have "a' dvd b'" by (subst (asm) ac_simps, blast)
+  hence "a'*?d dvd b'*?d" by (rule mult_dvd_mono, simp)
+  with ab'(1,2) show ?thesis by simp
+qed
+
+lemma pow_divides_eq [simp]:
+  "n \<noteq> 0 \<Longrightarrow> a ^ n dvd b ^ n \<longleftrightarrow> a dvd b"
+  by (auto intro: pow_divides_pow dvd_power_same)
+
+lemma divides_mult:
+  assumes mr: "m dvd r" and nr: "n dvd r" and mn: "gcd m n = 1"
+  shows "m * n dvd r"
+proof -
+  from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'"
+    unfolding dvd_def by blast
+  from mr n' have "m dvd n'*n" by (simp add: ac_simps)
+  hence "m dvd n'" using coprime_dvd_mult_iff[OF mn] by simp
+  then obtain k where k: "n' = m*k" unfolding dvd_def by blast
+  with n' have "r = m * n * k" by (simp add: mult_ac)
+  then show ?thesis unfolding dvd_def by blast
+qed
+
+lemma coprime_plus_one [simp]: "gcd (n + 1) n = 1"
+  by (subst add_commute, simp)
+
+lemma setprod_coprime [rule_format]:
+  "(\<forall>i\<in>A. gcd (f i) x = 1) \<longrightarrow> gcd (\<Prod>i\<in>A. f i) x = 1"
+  apply (cases "finite A")
+  apply (induct set: finite)
+  apply (auto simp add: gcd_mult_cancel)
+  done
+
+lemma coprime_divisors: 
+  assumes "d dvd a" "e dvd b" "gcd a b = 1"
+  shows "gcd d e = 1" 
+proof -
+  from assms obtain k l where "a = d * k" "b = e * l"
+    unfolding dvd_def by blast
+  with assms have "gcd (d * k) (e * l) = 1" by simp
+  hence "gcd (d * k) e = 1" by (rule coprime_lmult)
+  also have "gcd (d * k) e = gcd e (d * k)" by (simp add: ac_simps)
+  finally have "gcd e d = 1" by (rule coprime_lmult)
+  then show ?thesis by (simp add: ac_simps)
+qed
+
+lemma invertible_coprime:
+  "x * y mod m = 1 \<Longrightarrow> gcd x m = 1"
+  by (metis coprime_lmult gcd_1 ac_simps gcd_red)
+
+lemma lcm_gcd:
+  "lcm a b = a * b div (gcd a b * normalisation_factor (a*b))"
+  by (simp only: lcm_lcm_eucl gcd_gcd_eucl lcm_eucl_def)
+
+lemma lcm_gcd_prod:
+  "lcm a b * gcd a b = a * b div normalisation_factor (a*b)"
+proof (cases "a * b = 0")
+  let ?nf = normalisation_factor
+  assume "a * b \<noteq> 0"
+  hence "gcd a b \<noteq> 0" by (auto simp add: gcd_zero)
+  from lcm_gcd have "lcm a b * gcd a b = gcd a b * (a * b div (?nf (a*b) * gcd a b))" 
+    by (simp add: mult_ac)
+  also from `a * b \<noteq> 0` have "... = a * b div ?nf (a*b)" 
+    by (simp_all add: unit_ring_inv'1 dvd_mult_div_cancel unit_ring_inv)
+  finally show ?thesis .
+qed (simp add: lcm_gcd)
+
+lemma lcm_dvd1 [iff]:
+  "x dvd lcm x y"
+proof (cases "x*y = 0")
+  assume "x * y \<noteq> 0"
+  hence "gcd x y \<noteq> 0" by (auto simp: gcd_zero)
+  let ?c = "ring_inv (normalisation_factor (x*y))"
+  from `x * y \<noteq> 0` have [simp]: "is_unit (normalisation_factor (x*y))" by simp
+  from lcm_gcd_prod[of x y] have "lcm x y * gcd x y = x * ?c * y"
+    by (simp add: mult_ac unit_ring_inv)
+  hence "lcm x y * gcd x y div gcd x y = x * ?c * y div gcd x y" by simp
+  with `gcd x y \<noteq> 0` have "lcm x y = x * ?c * y div gcd x y"
+    by (subst (asm) div_mult_self2_is_id, simp_all)
+  also have "... = x * (?c * y div gcd x y)"
+    by (metis div_mult_swap gcd_dvd2 mult_assoc)
+  finally show ?thesis by (rule dvdI)
+qed (simp add: lcm_gcd)
+
+lemma lcm_least:
+  "\<lbrakk>a dvd k; b dvd k\<rbrakk> \<Longrightarrow> lcm a b dvd k"
+proof (cases "k = 0")
+  let ?nf = normalisation_factor
+  assume "k \<noteq> 0"
+  hence "is_unit (?nf k)" by simp
+  hence "?nf k \<noteq> 0" by (metis not_is_unit_0)
+  assume A: "a dvd k" "b dvd k"
+  hence "gcd a b \<noteq> 0" using `k \<noteq> 0` by (auto simp add: gcd_zero)
+  from A obtain r s where ar: "k = a * r" and bs: "k = b * s" 
+    unfolding dvd_def by blast
+  with `k \<noteq> 0` have "r * s \<noteq> 0" 
+    by (intro notI) (drule divisors_zero, elim disjE, simp_all)
+  hence "is_unit (?nf (r * s))" by simp
+  let ?c = "?nf k div ?nf (r*s)"
+  from `is_unit (?nf k)` and `is_unit (?nf (r * s))` have "is_unit ?c" by (rule unit_div)
+  hence "?c \<noteq> 0" using not_is_unit_0 by fast 
+  from ar bs have "k * k * gcd s r = ?nf k * k * gcd (k * s) (k * r)"
+    by (subst mult_assoc, subst gcd_mult_distrib[of k s r], simp only: ac_simps mult_assoc)
+  also have "... = ?nf k * k * gcd ((r*s) * a) ((r*s) * b)"
+    by (subst (3) `k = a * r`, subst (3) `k = b * s`, simp add: algebra_simps)
+  also have "... = ?c * r*s * k * gcd a b" using `r * s \<noteq> 0`
+    by (subst gcd_mult_distrib'[symmetric], simp add: algebra_simps unit_simps)
+  finally have "(a*r) * (b*s) * gcd s r = ?c * k * r * s * gcd a b"
+    by (subst ar[symmetric], subst bs[symmetric], simp add: mult_ac)
+  hence "a * b * gcd s r * (r * s) = ?c * k * gcd a b * (r * s)"
+    by (simp add: algebra_simps)
+  hence "?c * k * gcd a b = a * b * gcd s r" using `r * s \<noteq> 0`
+    by (metis div_mult_self2_is_id)
+  also have "... = lcm a b * gcd a b * gcd s r * ?nf (a*b)"
+    by (subst lcm_gcd_prod[of a b], metis gcd_mult_distrib gcd_mult_distrib') 
+  also have "... = lcm a b * gcd s r * ?nf (a*b) * gcd a b"
+    by (simp add: algebra_simps)
+  finally have "k * ?c = lcm a b * gcd s r * ?nf (a*b)" using `gcd a b \<noteq> 0`
+    by (metis mult.commute div_mult_self2_is_id)
+  hence "k = lcm a b * (gcd s r * ?nf (a*b)) div ?c" using `?c \<noteq> 0`
+    by (metis div_mult_self2_is_id mult_assoc) 
+  also have "... = lcm a b * (gcd s r * ?nf (a*b) div ?c)" using `is_unit ?c`
+    by (simp add: unit_simps)
+  finally show ?thesis by (rule dvdI)
+qed simp
+
+lemma lcm_zero:
+  "lcm a b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
+proof -
+  let ?nf = normalisation_factor
+  {
+    assume "a \<noteq> 0" "b \<noteq> 0"
+    hence "a * b div ?nf (a * b) \<noteq> 0" by (simp add: no_zero_divisors)
+    moreover from `a \<noteq> 0` and `b \<noteq> 0` have "gcd a b \<noteq> 0" by (simp add: gcd_zero)
+    ultimately have "lcm a b \<noteq> 0" using lcm_gcd_prod[of a b] by (intro notI, simp)
+  } moreover {
+    assume "a = 0 \<or> b = 0"
+    hence "lcm a b = 0" by (elim disjE, simp_all add: lcm_gcd)
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemmas lcm_0_iff = lcm_zero
+
+lemma gcd_lcm: 
+  assumes "lcm a b \<noteq> 0"
+  shows "gcd a b = a * b div (lcm a b * normalisation_factor (a * b))"
+proof-
+  from assms have "gcd a b \<noteq> 0" by (simp add: gcd_zero lcm_zero)
+  let ?c = "normalisation_factor (a*b)"
+  from `lcm a b \<noteq> 0` have "?c \<noteq> 0" by (intro notI, simp add: lcm_zero no_zero_divisors)
+  hence "is_unit ?c" by simp
+  from lcm_gcd_prod [of a b] have "gcd a b = a * b div ?c div lcm a b"
+    by (subst (2) div_mult_self2_is_id[OF `lcm a b \<noteq> 0`, symmetric], simp add: mult_ac)
+  also from `is_unit ?c` have "... = a * b div (?c * lcm a b)"
+    by (simp only: unit_ring_inv'1 unit_ring_inv)
+  finally show ?thesis by (simp only: ac_simps)
+qed
+
+lemma normalisation_factor_lcm [simp]:
+  "normalisation_factor (lcm a b) = (if a = 0 \<or> b = 0 then 0 else 1)"
+proof (cases "a = 0 \<or> b = 0")
+  case True then show ?thesis
+    by (simp add: lcm_gcd) (metis div_0 ac_simps mult_zero_left normalisation_factor_0)
+next
+  case False
+  let ?nf = normalisation_factor
+  from lcm_gcd_prod[of a b] 
+    have "?nf (lcm a b) * ?nf (gcd a b) = ?nf (a*b) div ?nf (a*b)"
+    by (metis div_by_0 div_self normalisation_correct normalisation_factor_0 normalisation_factor_mult)
+  also have "... = (if a*b = 0 then 0 else 1)"
+    by (cases "a*b = 0", simp, subst div_self, metis dvd_0_left normalisation_factor_dvd, simp)
+  finally show ?thesis using False by (simp add: no_zero_divisors)
+qed
+
+lemma lcm_dvd2 [iff]: "y dvd lcm x y"
+  using lcm_dvd1 [of y x] by (simp add: lcm_gcd ac_simps)
+
+lemma lcmI:
+  "\<lbrakk>x dvd k; y dvd k; \<And>l. x dvd l \<Longrightarrow> y dvd l \<Longrightarrow> k dvd l;
+    normalisation_factor k = (if k = 0 then 0 else 1)\<rbrakk> \<Longrightarrow> k = lcm x y"
+  by (intro normed_associated_imp_eq) (auto simp: associated_def intro: lcm_least)
+
+sublocale lcm!: abel_semigroup lcm
+proof
+  fix x y z
+  show "lcm (lcm x y) z = lcm x (lcm y z)"
+  proof (rule lcmI)
+    have "x dvd lcm x y" and "lcm x y dvd lcm (lcm x y) z" by simp_all
+    then show "x dvd lcm (lcm x y) z" by (rule dvd_trans)
+    
+    have "y dvd lcm x y" and "lcm x y dvd lcm (lcm x y) z" by simp_all
+    hence "y dvd lcm (lcm x y) z" by (rule dvd_trans)
+    moreover have "z dvd lcm (lcm x y) z" by simp
+    ultimately show "lcm y z dvd lcm (lcm x y) z" by (rule lcm_least)
+
+    fix l assume "x dvd l" and "lcm y z dvd l"
+    have "y dvd lcm y z" by simp
+    from this and `lcm y z dvd l` have "y dvd l" by (rule dvd_trans)
+    have "z dvd lcm y z" by simp
+    from this and `lcm y z dvd l` have "z dvd l" by (rule dvd_trans)
+    from `x dvd l` and `y dvd l` have "lcm x y dvd l" by (rule lcm_least)
+    from this and `z dvd l` show "lcm (lcm x y) z dvd l" by (rule lcm_least)
+  qed (simp add: lcm_zero)
+next
+  fix x y
+  show "lcm x y = lcm y x"
+    by (simp add: lcm_gcd ac_simps)
+qed
+
+lemma dvd_lcm_D1:
+  "lcm m n dvd k \<Longrightarrow> m dvd k"
+  by (rule dvd_trans, rule lcm_dvd1, assumption)
+
+lemma dvd_lcm_D2:
+  "lcm m n dvd k \<Longrightarrow> n dvd k"
+  by (rule dvd_trans, rule lcm_dvd2, assumption)
+
+lemma gcd_dvd_lcm [simp]:
+  "gcd a b dvd lcm a b"
+  by (metis dvd_trans gcd_dvd2 lcm_dvd2)
+
+lemma lcm_1_iff:
+  "lcm a b = 1 \<longleftrightarrow> is_unit a \<and> is_unit b"
+proof
+  assume "lcm a b = 1"
+  then show "is_unit a \<and> is_unit b" unfolding is_unit_def by auto
+next
+  assume "is_unit a \<and> is_unit b"
+  hence "a dvd 1" and "b dvd 1" unfolding is_unit_def by simp_all
+  hence "is_unit (lcm a b)" unfolding is_unit_def by (rule lcm_least)
+  hence "lcm a b = normalisation_factor (lcm a b)"
+    by (subst normalisation_factor_unit, simp_all)
+  also have "\<dots> = 1" using `is_unit a \<and> is_unit b` by (auto simp add: is_unit_def)
+  finally show "lcm a b = 1" .
+qed
+
+lemma lcm_0_left [simp]:
+  "lcm 0 x = 0"
+  by (rule sym, rule lcmI, simp_all)
+
+lemma lcm_0 [simp]:
+  "lcm x 0 = 0"
+  by (rule sym, rule lcmI, simp_all)
+
+lemma lcm_unique:
+  "a dvd d \<and> b dvd d \<and> 
+  normalisation_factor d = (if d = 0 then 0 else 1) \<and>
+  (\<forall>e. a dvd e \<and> b dvd e \<longrightarrow> d dvd e) \<longleftrightarrow> d = lcm a b"
+  by (rule, auto intro: lcmI simp: lcm_least lcm_zero)
+
+lemma dvd_lcm_I1 [simp]:
+  "k dvd m \<Longrightarrow> k dvd lcm m n"
+  by (metis lcm_dvd1 dvd_trans)
+
+lemma dvd_lcm_I2 [simp]:
+  "k dvd n \<Longrightarrow> k dvd lcm m n"
+  by (metis lcm_dvd2 dvd_trans)
+
+lemma lcm_1_left [simp]:
+  "lcm 1 x = x div normalisation_factor x"
+  by (cases "x = 0") (simp, rule sym, rule lcmI, simp_all)
+
+lemma lcm_1_right [simp]:
+  "lcm x 1 = x div normalisation_factor x"
+  by (simp add: ac_simps)
+
+lemma lcm_coprime:
+  "gcd a b = 1 \<Longrightarrow> lcm a b = a * b div normalisation_factor (a*b)"
+  by (subst lcm_gcd) simp
+
+lemma lcm_proj1_if_dvd: 
+  "y dvd x \<Longrightarrow> lcm x y = x div normalisation_factor x"
+  by (cases "x = 0") (simp, rule sym, rule lcmI, simp_all)
+
+lemma lcm_proj2_if_dvd: 
+  "x dvd y \<Longrightarrow> lcm x y = y div normalisation_factor y"
+  using lcm_proj1_if_dvd [of x y] by (simp add: ac_simps)
+
+lemma lcm_proj1_iff:
+  "lcm m n = m div normalisation_factor m \<longleftrightarrow> n dvd m"
+proof
+  assume A: "lcm m n = m div normalisation_factor m"
+  show "n dvd m"
+  proof (cases "m = 0")
+    assume [simp]: "m \<noteq> 0"
+    from A have B: "m = lcm m n * normalisation_factor m"
+      by (simp add: unit_eq_div2)
+    show ?thesis by (subst B, simp)
+  qed simp
+next
+  assume "n dvd m"
+  then show "lcm m n = m div normalisation_factor m" by (rule lcm_proj1_if_dvd)
+qed
+
+lemma lcm_proj2_iff:
+  "lcm m n = n div normalisation_factor n \<longleftrightarrow> m dvd n"
+  using lcm_proj1_iff [of n m] by (simp add: ac_simps)
+
+lemma euclidean_size_lcm_le1: 
+  assumes "a \<noteq> 0" and "b \<noteq> 0"
+  shows "euclidean_size a \<le> euclidean_size (lcm a b)"
+proof -
+  have "a dvd lcm a b" by (rule lcm_dvd1)
+  then obtain c where A: "lcm a b = a * c" unfolding dvd_def by blast
+  with `a \<noteq> 0` and `b \<noteq> 0` have "c \<noteq> 0" by (auto simp: lcm_zero)
+  then show ?thesis by (subst A, intro size_mult_mono)
+qed
+
+lemma euclidean_size_lcm_le2:
+  "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> euclidean_size b \<le> euclidean_size (lcm a b)"
+  using euclidean_size_lcm_le1 [of b a] by (simp add: ac_simps)
+
+lemma euclidean_size_lcm_less1:
+  assumes "b \<noteq> 0" and "\<not>b dvd a"
+  shows "euclidean_size a < euclidean_size (lcm a b)"
+proof (rule ccontr)
+  from assms have "a \<noteq> 0" by auto
+  assume "\<not>euclidean_size a < euclidean_size (lcm a b)"
+  with `a \<noteq> 0` and `b \<noteq> 0` have "euclidean_size (lcm a b) = euclidean_size a"
+    by (intro le_antisym, simp, intro euclidean_size_lcm_le1)
+  with assms have "lcm a b dvd a" 
+    by (rule_tac dvd_euclidean_size_eq_imp_dvd) (auto simp: lcm_zero)
+  hence "b dvd a" by (rule dvd_lcm_D2)
+  with `\<not>b dvd a` show False by contradiction
+qed
+
+lemma euclidean_size_lcm_less2:
+  assumes "a \<noteq> 0" and "\<not>a dvd b"
+  shows "euclidean_size b < euclidean_size (lcm a b)"
+  using assms euclidean_size_lcm_less1 [of a b] by (simp add: ac_simps)
+
+lemma lcm_mult_unit1:
+  "is_unit a \<Longrightarrow> lcm (x*a) y = lcm x y"
+  apply (rule lcmI)
+  apply (rule dvd_trans[of _ "x*a"], simp, rule lcm_dvd1)
+  apply (rule lcm_dvd2)
+  apply (rule lcm_least, simp add: unit_simps, assumption)
+  apply (subst normalisation_factor_lcm, simp add: lcm_zero)
+  done
+
+lemma lcm_mult_unit2:
+  "is_unit a \<Longrightarrow> lcm x (y*a) = lcm x y"
+  using lcm_mult_unit1 [of a y x] by (simp add: ac_simps)
+
+lemma lcm_div_unit1:
+  "is_unit a \<Longrightarrow> lcm (x div a) y = lcm x y"
+  by (simp add: unit_ring_inv lcm_mult_unit1)
+
+lemma lcm_div_unit2:
+  "is_unit a \<Longrightarrow> lcm x (y div a) = lcm x y"
+  by (simp add: unit_ring_inv lcm_mult_unit2)
+
+lemma lcm_left_idem:
+  "lcm p (lcm p q) = lcm p q"
+  apply (rule lcmI)
+  apply simp
+  apply (subst lcm.assoc [symmetric], rule lcm_dvd2)
+  apply (rule lcm_least, assumption)
+  apply (erule (1) lcm_least)
+  apply (auto simp: lcm_zero)
+  done
+
+lemma lcm_right_idem:
+  "lcm (lcm p q) q = lcm p q"
+  apply (rule lcmI)
+  apply (subst lcm.assoc, rule lcm_dvd1)
+  apply (rule lcm_dvd2)
+  apply (rule lcm_least, erule (1) lcm_least, assumption)
+  apply (auto simp: lcm_zero)
+  done
+
+lemma comp_fun_idem_lcm: "comp_fun_idem lcm"
+proof
+  fix a b show "lcm a \<circ> lcm b = lcm b \<circ> lcm a"
+    by (simp add: fun_eq_iff ac_simps)
+next
+  fix a show "lcm a \<circ> lcm a = lcm a" unfolding o_def
+    by (intro ext, simp add: lcm_left_idem)
+qed
+
+lemma dvd_Lcm [simp]: "x \<in> A \<Longrightarrow> x dvd Lcm A"
+  and Lcm_dvd [simp]: "(\<forall>x\<in>A. x dvd l') \<Longrightarrow> Lcm A dvd l'"
+  and normalisation_factor_Lcm [simp]: 
+          "normalisation_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)"
+proof -
+  have "(\<forall>x\<in>A. x dvd Lcm A) \<and> (\<forall>l'. (\<forall>x\<in>A. x dvd l') \<longrightarrow> Lcm A dvd l') \<and>
+    normalisation_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)" (is ?thesis)
+  proof (cases "\<exists>l. l \<noteq>  0 \<and> (\<forall>x\<in>A. x dvd l)")
+    case False
+    hence "Lcm A = 0" by (auto simp: Lcm_Lcm_eucl Lcm_eucl_def)
+    with False show ?thesis by auto
+  next
+    case True
+    then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l\<^sub>0)" by blast
+    def n \<equiv> "LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n"
+    def l \<equiv> "SOME l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n"
+    have "\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n"
+      apply (subst n_def)
+      apply (rule LeastI[of _ "euclidean_size l\<^sub>0"])
+      apply (rule exI[of _ l\<^sub>0])
+      apply (simp add: l\<^sub>0_props)
+      done
+    from someI_ex[OF this] have "l \<noteq> 0" and "\<forall>x\<in>A. x dvd l" and "euclidean_size l = n" 
+      unfolding l_def by simp_all
+    {
+      fix l' assume "\<forall>x\<in>A. x dvd l'"
+      with `\<forall>x\<in>A. x dvd l` have "\<forall>x\<in>A. x dvd gcd l l'" by (auto intro: gcd_greatest)
+      moreover from `l \<noteq> 0` have "gcd l l' \<noteq> 0" by (simp add: gcd_zero)
+      ultimately have "\<exists>b. b \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd b) \<and> euclidean_size b = euclidean_size (gcd l l')"
+        by (intro exI[of _ "gcd l l'"], auto)
+      hence "euclidean_size (gcd l l') \<ge> n" by (subst n_def) (rule Least_le)
+      moreover have "euclidean_size (gcd l l') \<le> n"
+      proof -
+        have "gcd l l' dvd l" by simp
+        then obtain a where "l = gcd l l' * a" unfolding dvd_def by blast
+        with `l \<noteq> 0` have "a \<noteq> 0" by auto
+        hence "euclidean_size (gcd l l') \<le> euclidean_size (gcd l l' * a)"
+          by (rule size_mult_mono)
+        also have "gcd l l' * a = l" using `l = gcd l l' * a` ..
+        also note `euclidean_size l = n`
+        finally show "euclidean_size (gcd l l') \<le> n" .
+      qed
+      ultimately have "euclidean_size l = euclidean_size (gcd l l')" 
+        by (intro le_antisym, simp_all add: `euclidean_size l = n`)
+      with `l \<noteq> 0` have "l dvd gcd l l'" by (blast intro: dvd_euclidean_size_eq_imp_dvd)
+      hence "l dvd l'" by (blast dest: dvd_gcd_D2)
+    }
+
+    with `(\<forall>x\<in>A. x dvd l)` and normalisation_factor_is_unit[OF `l \<noteq> 0`] and `l \<noteq> 0`
+      have "(\<forall>x\<in>A. x dvd l div normalisation_factor l) \<and> 
+        (\<forall>l'. (\<forall>x\<in>A. x dvd l') \<longrightarrow> l div normalisation_factor l dvd l') \<and>
+        normalisation_factor (l div normalisation_factor l) = 
+        (if l div normalisation_factor l = 0 then 0 else 1)"
+      by (auto simp: unit_simps)
+    also from True have "l div normalisation_factor l = Lcm A"
+      by (simp add: Lcm_Lcm_eucl Lcm_eucl_def Let_def n_def l_def)
+    finally show ?thesis .
+  qed
+  note A = this
+
+  {fix x assume "x \<in> A" then show "x dvd Lcm A" using A by blast}
+  {fix l' assume "\<forall>x\<in>A. x dvd l'" then show "Lcm A dvd l'" using A by blast}
+  from A show "normalisation_factor (Lcm A) = (if Lcm A = 0 then 0 else 1)" by blast
+qed
+    
+lemma LcmI:
+  "(\<And>x. x\<in>A \<Longrightarrow> x dvd l) \<Longrightarrow> (\<And>l'. (\<forall>x\<in>A. x dvd l') \<Longrightarrow> l dvd l') \<Longrightarrow>
+      normalisation_factor l = (if l = 0 then 0 else 1) \<Longrightarrow> l = Lcm A"
+  by (intro normed_associated_imp_eq)
+    (auto intro: Lcm_dvd dvd_Lcm simp: associated_def)
+
+lemma Lcm_subset:
+  "A \<subseteq> B \<Longrightarrow> Lcm A dvd Lcm B"
+  by (blast intro: Lcm_dvd dvd_Lcm)
+
+lemma Lcm_Un:
+  "Lcm (A \<union> B) = lcm (Lcm A) (Lcm B)"
+  apply (rule lcmI)
+  apply (blast intro: Lcm_subset)
+  apply (blast intro: Lcm_subset)
+  apply (intro Lcm_dvd ballI, elim UnE)
+  apply (rule dvd_trans, erule dvd_Lcm, assumption)
+  apply (rule dvd_trans, erule dvd_Lcm, assumption)
+  apply simp
+  done
+
+lemma Lcm_1_iff:
+  "Lcm A = 1 \<longleftrightarrow> (\<forall>x\<in>A. is_unit x)"
+proof
+  assume "Lcm A = 1"
+  then show "\<forall>x\<in>A. is_unit x" unfolding is_unit_def by auto
+qed (rule LcmI [symmetric], auto)
+
+lemma Lcm_no_units:
+  "Lcm A = Lcm (A - {x. is_unit x})"
+proof -
+  have "(A - {x. is_unit x}) \<union> {x\<in>A. is_unit x} = A" by blast
+  hence "Lcm A = lcm (Lcm (A - {x. is_unit x})) (Lcm {x\<in>A. is_unit x})"
+    by (simp add: Lcm_Un[symmetric])
+  also have "Lcm {x\<in>A. is_unit x} = 1" by (simp add: Lcm_1_iff)
+  finally show ?thesis by simp
+qed
+
+lemma Lcm_empty [simp]:
+  "Lcm {} = 1"
+  by (simp add: Lcm_1_iff)
+
+lemma Lcm_eq_0 [simp]:
+  "0 \<in> A \<Longrightarrow> Lcm A = 0"
+  by (drule dvd_Lcm) simp
+
+lemma Lcm0_iff':
+  "Lcm A = 0 \<longleftrightarrow> \<not>(\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l))"
+proof
+  assume "Lcm A = 0"
+  show "\<not>(\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l))"
+  proof
+    assume ex: "\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l)"
+    then obtain l\<^sub>0 where l\<^sub>0_props: "l\<^sub>0 \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l\<^sub>0)" by blast
+    def n \<equiv> "LEAST n. \<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n"
+    def l \<equiv> "SOME l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n"
+    have "\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l) \<and> euclidean_size l = n"
+      apply (subst n_def)
+      apply (rule LeastI[of _ "euclidean_size l\<^sub>0"])
+      apply (rule exI[of _ l\<^sub>0])
+      apply (simp add: l\<^sub>0_props)
+      done
+    from someI_ex[OF this] have "l \<noteq> 0" unfolding l_def by simp_all
+    hence "l div normalisation_factor l \<noteq> 0" by simp
+    also from ex have "l div normalisation_factor l = Lcm A"
+       by (simp only: Lcm_Lcm_eucl Lcm_eucl_def n_def l_def if_True Let_def)
+    finally show False using `Lcm A = 0` by contradiction
+  qed
+qed (simp only: Lcm_Lcm_eucl Lcm_eucl_def if_False)
+
+lemma Lcm0_iff [simp]:
+  "finite A \<Longrightarrow> Lcm A = 0 \<longleftrightarrow> 0 \<in> A"
+proof -
+  assume "finite A"
+  have "0 \<in> A \<Longrightarrow> Lcm A = 0"  by (intro dvd_0_left dvd_Lcm)
+  moreover {
+    assume "0 \<notin> A"
+    hence "\<Prod>A \<noteq> 0" 
+      apply (induct rule: finite_induct[OF `finite A`]) 
+      apply simp
+      apply (subst setprod.insert, assumption, assumption)
+      apply (rule no_zero_divisors)
+      apply blast+
+      done
+    moreover from `finite A` have "\<forall>x\<in>A. x dvd \<Prod>A" by (intro ballI dvd_setprod)
+    ultimately have "\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l)" by blast
+    with Lcm0_iff' have "Lcm A \<noteq> 0" by simp
+  }
+  ultimately show "Lcm A = 0 \<longleftrightarrow> 0 \<in> A" by blast
+qed
+
+lemma Lcm_no_multiple:
+  "(\<forall>m. m \<noteq> 0 \<longrightarrow> (\<exists>x\<in>A. \<not>x dvd m)) \<Longrightarrow> Lcm A = 0"
+proof -
+  assume "\<forall>m. m \<noteq> 0 \<longrightarrow> (\<exists>x\<in>A. \<not>x dvd m)"
+  hence "\<not>(\<exists>l. l \<noteq> 0 \<and> (\<forall>x\<in>A. x dvd l))" by blast
+  then show "Lcm A = 0" by (simp only: Lcm_Lcm_eucl Lcm_eucl_def if_False)
+qed
+
+lemma Lcm_insert [simp]:
+  "Lcm (insert a A) = lcm a (Lcm A)"
+proof (rule lcmI)
+  fix l assume "a dvd l" and "Lcm A dvd l"
+  hence "\<forall>x\<in>A. x dvd l" by (blast intro: dvd_trans dvd_Lcm)
+  with `a dvd l` show "Lcm (insert a A) dvd l" by (force intro: Lcm_dvd)
+qed (auto intro: Lcm_dvd dvd_Lcm)
+ 
+lemma Lcm_finite:
+  assumes "finite A"
+  shows "Lcm A = Finite_Set.fold lcm 1 A"
+  by (induct rule: finite.induct[OF `finite A`])
+    (simp_all add: comp_fun_idem.fold_insert_idem[OF comp_fun_idem_lcm])
+
+lemma Lcm_set [code, code_unfold]:
+  "Lcm (set xs) = fold lcm xs 1"
+  using comp_fun_idem.fold_set_fold[OF comp_fun_idem_lcm] Lcm_finite by (simp add: ac_simps)
+
+lemma Lcm_singleton [simp]:
+  "Lcm {a} = a div normalisation_factor a"
+  by simp
+
+lemma Lcm_2 [simp]:
+  "Lcm {a,b} = lcm a b"
+  by (simp only: Lcm_insert Lcm_empty lcm_1_right)
+    (cases "b = 0", simp, rule lcm_div_unit2, simp)
+
+lemma Lcm_coprime:
+  assumes "finite A" and "A \<noteq> {}" 
+  assumes "\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> gcd a b = 1"
+  shows "Lcm A = \<Prod>A div normalisation_factor (\<Prod>A)"
+using assms proof (induct rule: finite_ne_induct)
+  case (insert a A)
+  have "Lcm (insert a A) = lcm a (Lcm A)" by simp
+  also from insert have "Lcm A = \<Prod>A div normalisation_factor (\<Prod>A)" by blast
+  also have "lcm a \<dots> = lcm a (\<Prod>A)" by (cases "\<Prod>A = 0") (simp_all add: lcm_div_unit2)
+  also from insert have "gcd a (\<Prod>A) = 1" by (subst gcd.commute, intro setprod_coprime) auto
+  with insert have "lcm a (\<Prod>A) = \<Prod>(insert a A) div normalisation_factor (\<Prod>(insert a A))"
+    by (simp add: lcm_coprime)
+  finally show ?case .
+qed simp
+      
+lemma Lcm_coprime':
+  "card A \<noteq> 0 \<Longrightarrow> (\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<noteq> b \<Longrightarrow> gcd a b = 1)
+    \<Longrightarrow> Lcm A = \<Prod>A div normalisation_factor (\<Prod>A)"
+  by (rule Lcm_coprime) (simp_all add: card_eq_0_iff)
+
+lemma Gcd_Lcm:
+  "Gcd A = Lcm {d. \<forall>x\<in>A. d dvd x}"
+  by (simp add: Gcd_Gcd_eucl Lcm_Lcm_eucl Gcd_eucl_def)
+
+lemma Gcd_dvd [simp]: "x \<in> A \<Longrightarrow> Gcd A dvd x"
+  and dvd_Gcd [simp]: "(\<forall>x\<in>A. g' dvd x) \<Longrightarrow> g' dvd Gcd A"
+  and normalisation_factor_Gcd [simp]: 
+    "normalisation_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)"
+proof -
+  fix x assume "x \<in> A"
+  hence "Lcm {d. \<forall>x\<in>A. d dvd x} dvd x" by (intro Lcm_dvd) blast
+  then show "Gcd A dvd x" by (simp add: Gcd_Lcm)
+next
+  fix g' assume "\<forall>x\<in>A. g' dvd x"
+  hence "g' dvd Lcm {d. \<forall>x\<in>A. d dvd x}" by (intro dvd_Lcm) blast
+  then show "g' dvd Gcd A" by (simp add: Gcd_Lcm)
+next
+  show "normalisation_factor (Gcd A) = (if Gcd A = 0 then 0 else 1)"
+    by (simp add: Gcd_Lcm normalisation_factor_Lcm)
+qed
+
+lemma GcdI:
+  "(\<And>x. x\<in>A \<Longrightarrow> l dvd x) \<Longrightarrow> (\<And>l'. (\<forall>x\<in>A. l' dvd x) \<Longrightarrow> l' dvd l) \<Longrightarrow>
+    normalisation_factor l = (if l = 0 then 0 else 1) \<Longrightarrow> l = Gcd A"
+  by (intro normed_associated_imp_eq)
+    (auto intro: Gcd_dvd dvd_Gcd simp: associated_def)
+
+lemma Lcm_Gcd:
+  "Lcm A = Gcd {m. \<forall>x\<in>A. x dvd m}"
+  by (rule LcmI[symmetric]) (auto intro: dvd_Gcd Gcd_dvd)
+
+lemma Gcd_0_iff:
+  "Gcd A = 0 \<longleftrightarrow> A \<subseteq> {0}"
+  apply (rule iffI)
+  apply (rule subsetI, drule Gcd_dvd, simp)
+  apply (auto intro: GcdI[symmetric])
+  done
+
+lemma Gcd_empty [simp]:
+  "Gcd {} = 0"
+  by (simp add: Gcd_0_iff)
+
+lemma Gcd_1:
+  "1 \<in> A \<Longrightarrow> Gcd A = 1"
+  by (intro GcdI[symmetric]) (auto intro: Gcd_dvd dvd_Gcd)
+
+lemma Gcd_insert [simp]:
+  "Gcd (insert a A) = gcd a (Gcd A)"
+proof (rule gcdI)
+  fix l assume "l dvd a" and "l dvd Gcd A"
+  hence "\<forall>x\<in>A. l dvd x" by (blast intro: dvd_trans Gcd_dvd)
+  with `l dvd a` show "l dvd Gcd (insert a A)" by (force intro: Gcd_dvd)
+qed (auto intro: Gcd_dvd dvd_Gcd simp: normalisation_factor_Gcd)
+
+lemma Gcd_finite:
+  assumes "finite A"
+  shows "Gcd A = Finite_Set.fold gcd 0 A"
+  by (induct rule: finite.induct[OF `finite A`])
+    (simp_all add: comp_fun_idem.fold_insert_idem[OF comp_fun_idem_gcd])
+
+lemma Gcd_set [code, code_unfold]:
+  "Gcd (set xs) = fold gcd xs 0"
+  using comp_fun_idem.fold_set_fold[OF comp_fun_idem_gcd] Gcd_finite by (simp add: ac_simps)
+
+lemma Gcd_singleton [simp]: "Gcd {a} = a div normalisation_factor a"
+  by (simp add: gcd_0)
+
+lemma Gcd_2 [simp]: "Gcd {a,b} = gcd a b"
+  by (simp only: Gcd_insert Gcd_empty gcd_0) (cases "b = 0", simp, rule gcd_div_unit2, simp)
+
+end
+
+text {*
+  A Euclidean ring is a Euclidean semiring with additive inverses. It provides a 
+  few more lemmas; in particular, Bezout's lemma holds for any Euclidean ring.
+*}
+
+class euclidean_ring = euclidean_semiring + idom
+
+class euclidean_ring_gcd = euclidean_semiring_gcd + idom
+begin
+
+subclass euclidean_ring ..
+
+lemma gcd_neg1 [simp]:
+  "gcd (-x) y = gcd x y"
+  by (rule sym, rule gcdI, simp_all add: gcd_greatest gcd_zero)
+
+lemma gcd_neg2 [simp]:
+  "gcd x (-y) = gcd x y"
+  by (rule sym, rule gcdI, simp_all add: gcd_greatest gcd_zero)
+
+lemma gcd_neg_numeral_1 [simp]:
+  "gcd (- numeral n) x = gcd (numeral n) x"
+  by (fact gcd_neg1)
+
+lemma gcd_neg_numeral_2 [simp]:
+  "gcd x (- numeral n) = gcd x (numeral n)"
+  by (fact gcd_neg2)
+
+lemma gcd_diff1: "gcd (m - n) n = gcd m n"
+  by (subst diff_conv_add_uminus, subst gcd_neg2[symmetric],  subst gcd_add1, simp)
+
+lemma gcd_diff2: "gcd (n - m) n = gcd m n"
+  by (subst gcd_neg1[symmetric], simp only: minus_diff_eq gcd_diff1)
+
+lemma coprime_minus_one [simp]: "gcd (n - 1) n = 1"
+proof -
+  have "gcd (n - 1) n = gcd n (n - 1)" by (fact gcd.commute)
+  also have "\<dots> = gcd ((n - 1) + 1) (n - 1)" by simp
+  also have "\<dots> = 1" by (rule coprime_plus_one)
+  finally show ?thesis .
+qed
+
+lemma lcm_neg1 [simp]: "lcm (-x) y = lcm x y"
+  by (rule sym, rule lcmI, simp_all add: lcm_least lcm_zero)
+
+lemma lcm_neg2 [simp]: "lcm x (-y) = lcm x y"
+  by (rule sym, rule lcmI, simp_all add: lcm_least lcm_zero)
+
+lemma lcm_neg_numeral_1 [simp]: "lcm (- numeral n) x = lcm (numeral n) x"
+  by (fact lcm_neg1)
+
+lemma lcm_neg_numeral_2 [simp]: "lcm x (- numeral n) = lcm x (numeral n)"
+  by (fact lcm_neg2)
+
+function euclid_ext :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<times> 'a \<times> 'a" where
+  "euclid_ext a b = 
+     (if b = 0 then 
+        let x = ring_inv (normalisation_factor a) in (x, 0, a * x)
+      else 
+        case euclid_ext b (a mod b) of
+            (s,t,c) \<Rightarrow> (t, s - t * (a div b), c))"
+  by (pat_completeness, simp)
+  termination by (relation "measure (euclidean_size \<circ> snd)", simp_all)
+
+declare euclid_ext.simps [simp del]
+
+lemma euclid_ext_0: 
+  "euclid_ext a 0 = (ring_inv (normalisation_factor a), 0, a * ring_inv (normalisation_factor a))"
+  by (subst euclid_ext.simps, simp add: Let_def)
+
+lemma euclid_ext_non_0:
+  "b \<noteq> 0 \<Longrightarrow> euclid_ext a b = (case euclid_ext b (a mod b) of 
+    (s,t,c) \<Rightarrow> (t, s - t * (a div b), c))"
+  by (subst euclid_ext.simps, simp)
+
+definition euclid_ext' :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<times> 'a"
+where
+  "euclid_ext' a b = (case euclid_ext a b of (s, t, _) \<Rightarrow> (s, t))"
+
+lemma euclid_ext_gcd [simp]:
+  "(case euclid_ext a b of (_,_,t) \<Rightarrow> t) = gcd a b"
+proof (induct a b rule: euclid_ext.induct)
+  case (1 a b)
+  then show ?case
+  proof (cases "b = 0")
+    case True
+      then show ?thesis by (cases "a = 0") 
+        (simp_all add: euclid_ext_0 unit_div mult_ac unit_simps gcd_0)
+    next
+    case False with 1 show ?thesis
+      by (simp add: euclid_ext_non_0 ac_simps split: prod.split prod.split_asm)
+    qed
+qed
+
+lemma euclid_ext_gcd' [simp]:
+  "euclid_ext a b = (r, s, t) \<Longrightarrow> t = gcd a b"
+  by (insert euclid_ext_gcd[of a b], drule (1) subst, simp)
+
+lemma euclid_ext_correct:
+  "case euclid_ext x y of (s,t,c) \<Rightarrow> s*x + t*y = c"
+proof (induct x y rule: euclid_ext.induct)
+  case (1 x y)
+  show ?case
+  proof (cases "y = 0")
+    case True
+    then show ?thesis by (simp add: euclid_ext_0 mult_ac)
+  next
+    case False
+    obtain s t c where stc: "euclid_ext y (x mod y) = (s,t,c)"
+      by (cases "euclid_ext y (x mod y)", blast)
+    from 1 have "c = s * y + t * (x mod y)" by (simp add: stc False)
+    also have "... = t*((x div y)*y + x mod y) + (s - t * (x div y))*y"
+      by (simp add: algebra_simps) 
+    also have "(x div y)*y + x mod y = x" using mod_div_equality .
+    finally show ?thesis
+      by (subst euclid_ext.simps, simp add: False stc)
+    qed
+qed
+
+lemma euclid_ext'_correct:
+  "fst (euclid_ext' a b) * a + snd (euclid_ext' a b) * b = gcd a b"
+proof-
+  obtain s t c where "euclid_ext a b = (s,t,c)"
+    by (cases "euclid_ext a b", blast)
+  with euclid_ext_correct[of a b] euclid_ext_gcd[of a b]
+    show ?thesis unfolding euclid_ext'_def by simp
+qed
+
+lemma bezout: "\<exists>s t. s * x + t * y = gcd x y"
+  using euclid_ext'_correct by blast
+
+lemma euclid_ext'_0 [simp]: "euclid_ext' x 0 = (ring_inv (normalisation_factor x), 0)" 
+  by (simp add: bezw_def euclid_ext'_def euclid_ext_0)
+
+lemma euclid_ext'_non_0: "y \<noteq> 0 \<Longrightarrow> euclid_ext' x y = (snd (euclid_ext' y (x mod y)),
+  fst (euclid_ext' y (x mod y)) - snd (euclid_ext' y (x mod y)) * (x div y))"
+  by (cases "euclid_ext y (x mod y)") 
+    (simp add: euclid_ext'_def euclid_ext_non_0)
+  
+end
+
+instantiation nat :: euclidean_semiring
+begin
+
+definition [simp]:
+  "euclidean_size_nat = (id :: nat \<Rightarrow> nat)"
+
+definition [simp]:
+  "normalisation_factor_nat (n::nat) = (if n = 0 then 0 else 1 :: nat)"
+
+instance proof
+qed (simp_all add: is_unit_def)
+
+end
+
+instantiation int :: euclidean_ring
+begin
+
+definition [simp]:
+  "euclidean_size_int = (nat \<circ> abs :: int \<Rightarrow> nat)"
+
+definition [simp]:
+  "normalisation_factor_int = (sgn :: int \<Rightarrow> int)"
+
+instance proof
+  case goal2 then show ?case by (auto simp add: abs_mult nat_mult_distrib)
+next
+  case goal3 then show ?case by (simp add: zsgn_def is_unit_def)
+next
+  case goal5 then show ?case by (auto simp: zsgn_def is_unit_def)
+next
+  case goal6 then show ?case by (auto split: abs_split simp: zsgn_def is_unit_def)
+qed (auto simp: sgn_times split: abs_split)
+
+end
+
+end
+
--- a/src/HOL/Old_Number_Theory/Factorization.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Old_Number_Theory/Factorization.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -143,7 +143,7 @@
   apply (induct xs)
    apply simp
    apply (case_tac xs)
-    apply (simp_all cong del: list.weak_case_cong)
+    apply (simp_all cong del: list.case_cong_weak)
   done
 
 lemma nondec_sort: "nondec (sort xs)"
--- a/src/HOL/Partial_Function.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Partial_Function.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -9,8 +9,9 @@
 keywords "partial_function" :: thy_decl
 begin
 
+named_theorems partial_function_mono "monotonicity rules for partial function definitions"
 ML_file "Tools/Function/partial_function.ML"
-setup Partial_Function.setup
+
 
 subsection {* Axiomatic setup *}
 
--- a/src/HOL/Product_Type.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Product_Type.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -281,7 +281,7 @@
 setup {* Sign.parent_path *}
 
 declare prod.case [nitpick_simp del]
-declare prod.weak_case_cong [cong del]
+declare prod.case_cong_weak [cong del]
 
 
 subsubsection {* Tuple syntax *}
@@ -486,7 +486,7 @@
 
 lemma split_weak_cong: "p = q \<Longrightarrow> split c p = split c q"
   -- {* Prevents simplification of @{term c}: much faster *}
-  by (fact prod.weak_case_cong)
+  by (fact prod.case_cong_weak)
 
 lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g"
   by (simp add: split_eta)
--- a/src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -4,9 +4,11 @@
 *)
 
 theory Quickcheck_Lattice_Examples
-imports "~~/src/HOL/Library/Quickcheck_Types"
+imports Main
 begin
 
+declare [[quickcheck_finite_type_size=5]]
+
 text {* We show how other default types help to find counterexamples to propositions if
   the standard default type @{typ int} is insufficient. *}
 
--- a/src/HOL/Quotient.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Quotient.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -748,8 +748,12 @@
 
 text {* Auxiliary data for the quotient package *}
 
+named_theorems quot_equiv "equivalence relation theorems"
+named_theorems quot_respect "respectfulness theorems"
+named_theorems quot_preserve "preservation theorems"
+named_theorems id_simps "identity simp rules for maps"
+named_theorems quot_thm "quotient theorems"
 ML_file "Tools/Quotient/quotient_info.ML"
-setup Quotient_Info.setup
 
 declare [[mapQ3 "fun" = (rel_fun, fun_quotient3)]]
 
--- a/src/HOL/Quotient_Examples/FSet.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Quotient_Examples/FSet.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -985,7 +985,7 @@
   have b: "\<And>x' ys'. \<lbrakk>\<not> List.member ys' x'; a # xs \<approx> x' # ys'\<rbrakk> \<Longrightarrow> thesis" by fact
   have c: "xs = [] \<Longrightarrow> thesis" using b 
     apply(simp)
-    by (metis List.set_simps(1) emptyE empty_subsetI)
+    by (metis list.set(1) emptyE empty_subsetI)
   have "\<And>x ys. \<lbrakk>\<not> List.member ys x; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis"
   proof -
     fix x :: 'a
--- a/src/HOL/Quotient_Examples/Lift_FSet.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Quotient_Examples/Lift_FSet.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -151,7 +151,7 @@
   using filter_filter [Transfer.transferred] .
 
 lemma "fset (fcons x xs) = insert x (fset xs)"
-  using set_simps(2) [Transfer.transferred] .
+  using list.set(2) [Transfer.transferred] .
 
 lemma "fset (fappend xs ys) = fset xs \<union> fset ys"
   using set_append [Transfer.transferred] .
--- a/src/HOL/ROOT	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/ROOT	Wed Aug 27 15:52:58 2014 +0200
@@ -38,7 +38,6 @@
     Product_Lexorder
     Product_Order
     Finite_Lattice
-    Quickcheck_Types
     (*data refinements and dependent applications*)
     AList_Mapping
     Code_Binary_Nat
@@ -189,6 +188,7 @@
     Pocklington
     Gauss
     Number_Theory
+    Euclidean_Algorithm
   document_files
     "root.tex"
 
@@ -237,6 +237,19 @@
     Generate_Target_Nat
     Generate_Efficient_Datastructures
     Generate_Pretty_Char
+    Code_Test
+  theories[condition = ISABELLE_GHC]
+    Code_Test_GHC
+  theories[condition = ISABELLE_MLTON]
+    Code_Test_MLton
+  theories[condition = ISABELLE_OCAMLC]
+    Code_Test_OCaml
+  theories[condition = ISABELLE_POLYML_PATH]
+    Code_Test_PolyML
+  theories[condition = ISABELLE_SCALA]
+    Code_Test_Scala
+  theories[condition = ISABELLE_SMLNJ]
+    Code_Test_SMLNJ
 
 session "HOL-Metis_Examples" in Metis_Examples = HOL +
   description {*
@@ -782,7 +795,6 @@
   files
     "Boogie_Dijkstra.certs2"
     "Boogie_Max.certs2"
-    "SMT_Examples.certs"
     "SMT_Examples.certs2"
     "SMT_Word_Examples.certs2"
     "VCC_Max.certs2"
--- a/src/HOL/Real.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Real.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -1000,13 +1000,24 @@
 where
   "real_of_rat \<equiv> of_rat"
 
-consts
-  (*overloaded constant for injecting other types into "real"*)
-  real :: "'a => real"
+class real_of =
+  fixes real :: "'a \<Rightarrow> real"
+
+instantiation nat :: real_of
+begin
+
+definition real_nat :: "nat \<Rightarrow> real" where real_of_nat_def [code_unfold]: "real \<equiv> of_nat" 
 
-defs (overloaded)
-  real_of_nat_def [code_unfold]: "real == real_of_nat"
-  real_of_int_def [code_unfold]: "real == real_of_int"
+instance ..
+end
+
+instantiation int :: real_of
+begin
+
+definition real_int :: "int \<Rightarrow> real" where real_of_int_def [code_unfold]: "real \<equiv> of_int" 
+
+instance ..
+end
 
 declare [[coercion_enabled]]
 declare [[coercion "real::nat\<Rightarrow>real"]]
@@ -1463,12 +1474,14 @@
       @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one},
       @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
       @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
-      @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}]
+      @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)},
+      @{thm real_of_int_def[symmetric]}, @{thm real_of_nat_def[symmetric]}]
   #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
-  #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
+  #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"})
+  #> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat \<Rightarrow> real"})
+  #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int \<Rightarrow> real"}))
 *}
 
-
 subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
 
 lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" 
@@ -1650,78 +1663,66 @@
 lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
 unfolding real_of_int_def by (rule floor_exists)
 
-lemma lemma_floor:
-  assumes a1: "real m \<le> r" and a2: "r < real n + 1"
-  shows "m \<le> (n::int)"
-proof -
-  have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
-  also have "... = real (n + 1)" by simp
-  finally have "m < n + 1" by (simp only: real_of_int_less_iff)
-  thus ?thesis by arith
-qed
+lemma lemma_floor: "real m \<le> r \<Longrightarrow> r < real n + 1 \<Longrightarrow> m \<le> (n::int)"
+  by simp
 
 lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
 unfolding real_of_int_def by (rule of_int_floor_le)
 
 lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
-by (auto intro: lemma_floor)
+  by simp
 
 lemma real_of_int_floor_cancel [simp]:
     "(real (floor x) = x) = (\<exists>n::int. x = real n)"
   using floor_real_of_int by metis
 
 lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
-  unfolding real_of_int_def using floor_unique [of n x] by simp
+  by linarith
 
 lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
-  unfolding real_of_int_def by (rule floor_unique)
+  by linarith
 
 lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
-apply (rule inj_int [THEN injD])
-apply (simp add: real_of_nat_Suc)
-apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
-done
+  by linarith
 
 lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
-apply (drule order_le_imp_less_or_eq)
-apply (auto intro: floor_eq3)
-done
+  by linarith
 
 lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
-  unfolding real_of_int_def using floor_correct [of r] by simp
+  by linarith
 
 lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
-  unfolding real_of_int_def using floor_correct [of r] by simp
+  by linarith
 
 lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
-  unfolding real_of_int_def using floor_correct [of r] by simp
+  by linarith
 
 lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
-  unfolding real_of_int_def using floor_correct [of r] by simp
+  by linarith
 
 lemma le_floor: "real a <= x ==> a <= floor x"
-  unfolding real_of_int_def by (simp add: le_floor_iff)
+  by linarith
 
 lemma real_le_floor: "a <= floor x ==> real a <= x"
-  unfolding real_of_int_def by (simp add: le_floor_iff)
+  by linarith
 
 lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
-  unfolding real_of_int_def by (rule le_floor_iff)
+  by linarith
 
 lemma floor_less_eq: "(floor x < a) = (x < real a)"
-  unfolding real_of_int_def by (rule floor_less_iff)
+  by linarith
 
 lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
-  unfolding real_of_int_def by (rule less_floor_iff)
+  by linarith
 
 lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
-  unfolding real_of_int_def by (rule floor_le_iff)
+  by linarith
 
 lemma floor_add [simp]: "floor (x + real a) = floor x + a"
-  unfolding real_of_int_def by (rule floor_add_of_int)
+  by linarith
 
 lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
-  unfolding real_of_int_def by (rule floor_diff_of_int)
+  by linarith
 
 lemma le_mult_floor:
   assumes "0 \<le> (a :: real)" and "0 \<le> b"
@@ -1746,56 +1747,56 @@
 qed (auto simp: real_of_int_div)
 
 lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
-  unfolding real_of_nat_def by simp
+  by linarith
 
 lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
-  unfolding real_of_int_def by (rule le_of_int_ceiling)
+  by linarith
 
 lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
-  unfolding real_of_int_def by simp
+  by linarith
 
 lemma real_of_int_ceiling_cancel [simp]:
      "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
   using ceiling_real_of_int by metis
 
 lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
-  unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
+  by linarith
 
 lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
-  unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp
+  by linarith
 
 lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n  |] ==> ceiling x = n"
-  unfolding real_of_int_def using ceiling_unique [of n x] by simp
+  by linarith
 
 lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
-  unfolding real_of_int_def using ceiling_correct [of r] by simp
+  by linarith
 
 lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
-  unfolding real_of_int_def using ceiling_correct [of r] by simp
+  by linarith
 
 lemma ceiling_le: "x <= real a ==> ceiling x <= a"
-  unfolding real_of_int_def by (simp add: ceiling_le_iff)
+  by linarith
 
 lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
-  unfolding real_of_int_def by (simp add: ceiling_le_iff)
+  by linarith
 
 lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
-  unfolding real_of_int_def by (rule ceiling_le_iff)
+  by linarith
 
 lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
-  unfolding real_of_int_def by (rule less_ceiling_iff)
+  by linarith
 
 lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
-  unfolding real_of_int_def by (rule ceiling_less_iff)
+  by linarith
 
 lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
-  unfolding real_of_int_def by (rule le_ceiling_iff)
+  by linarith
 
 lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
-  unfolding real_of_int_def by (rule ceiling_add_of_int)
+  by linarith
 
 lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
-  unfolding real_of_int_def by (rule ceiling_diff_of_int)
+  by linarith
 
 
 subsubsection {* Versions for the natural numbers *}
@@ -1808,111 +1809,88 @@
   natceiling :: "real => nat" where
   "natceiling x = nat(ceiling x)"
 
+lemma natfloor_split[arith_split]: "P (natfloor t) \<longleftrightarrow> (t < 0 \<longrightarrow> P 0) \<and> (\<forall>n. of_nat n \<le> t \<and> t < of_nat n + 1 \<longrightarrow> P n)"
+proof -
+  have [dest]: "\<And>n m::nat. real n \<le> t \<Longrightarrow> t < real n + 1 \<Longrightarrow> real m \<le> t \<Longrightarrow> t < real m + 1 \<Longrightarrow> n = m"
+    by simp
+  show ?thesis
+    by (auto simp: natfloor_def real_of_nat_def[symmetric] split: split_nat floor_split)
+qed
+
+lemma natceiling_split[arith_split]:
+  "P (natceiling t) \<longleftrightarrow> (t \<le> - 1 \<longrightarrow> P 0) \<and> (\<forall>n. of_nat n - 1 < t \<and> t \<le> of_nat n \<longrightarrow> P n)"
+proof -
+  have [dest]: "\<And>n m::nat. real n - 1 < t \<Longrightarrow> t \<le> real n \<Longrightarrow> real m - 1 < t \<Longrightarrow> t \<le> real m \<Longrightarrow> n = m"
+    by simp
+  show ?thesis
+    by (auto simp: natceiling_def real_of_nat_def[symmetric] split: split_nat ceiling_split)
+qed
+
 lemma natfloor_zero [simp]: "natfloor 0 = 0"
-  by (unfold natfloor_def, simp)
+  by linarith
 
 lemma natfloor_one [simp]: "natfloor 1 = 1"
-  by (unfold natfloor_def, simp)
-
-lemma zero_le_natfloor [simp]: "0 <= natfloor x"
-  by (unfold natfloor_def, simp)
+  by linarith
 
 lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n"
   by (unfold natfloor_def, simp)
 
 lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
-  by (unfold natfloor_def, simp)
+  by linarith
 
 lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
-  by (unfold natfloor_def, simp)
+  by linarith
 
 lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
-  unfolding natfloor_def by simp
+  by linarith
 
 lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
-  unfolding natfloor_def by (intro nat_mono floor_mono)
+  by linarith
 
 lemma le_natfloor: "real x <= a ==> x <= natfloor a"
-  apply (unfold natfloor_def)
-  apply (subst nat_int [THEN sym])
-  apply (rule nat_mono)
-  apply (rule le_floor)
-  apply simp
-done
+  by linarith
 
 lemma natfloor_less_iff: "0 \<le> x \<Longrightarrow> natfloor x < n \<longleftrightarrow> x < real n"
-  unfolding natfloor_def real_of_nat_def
-  by (simp add: nat_less_iff floor_less_iff)
+  by linarith
 
-lemma less_natfloor:
-  assumes "0 \<le> x" and "x < real (n :: nat)"
-  shows "natfloor x < n"
-  using assms by (simp add: natfloor_less_iff)
+lemma less_natfloor: "0 \<le> x \<Longrightarrow> x < real (n :: nat) \<Longrightarrow> natfloor x < n"
+  by linarith
 
 lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
-  apply (rule iffI)
-  apply (rule order_trans)
-  prefer 2
-  apply (erule real_natfloor_le)
-  apply (subst real_of_nat_le_iff)
-  apply assumption
-  apply (erule le_natfloor)
-done
+  by linarith
 
 lemma le_natfloor_eq_numeral [simp]:
-    "~ neg((numeral n)::int) ==> 0 <= x ==>
-      (numeral n <= natfloor x) = (numeral n <= x)"
-  apply (subst le_natfloor_eq, assumption)
-  apply simp
-done
+    "0 \<le> x \<Longrightarrow> (numeral n \<le> natfloor x) = (numeral n \<le> x)"
+  by (subst le_natfloor_eq, assumption) simp
+
+lemma le_natfloor_eq_one [simp]: "(1 \<le> natfloor x) = (1 \<le> x)"
+  by linarith
 
-lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
-  apply (case_tac "0 <= x")
-  apply (subst le_natfloor_eq, assumption, simp)
-  apply (rule iffI)
-  apply (subgoal_tac "natfloor x <= natfloor 0")
-  apply simp
-  apply (rule natfloor_mono)
-  apply simp
-  apply simp
-done
+lemma natfloor_eq: "real n \<le> x \<Longrightarrow> x < real n + 1 \<Longrightarrow> natfloor x = n"
+  by linarith
 
-lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
-  unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"])
-
-lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
-  apply (case_tac "0 <= x")
-  apply (unfold natfloor_def)
-  apply simp
-  apply simp_all
-done
+lemma real_natfloor_add_one_gt: "x < real (natfloor x) + 1"
+  by linarith
 
 lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
-using real_natfloor_add_one_gt by (simp add: algebra_simps)
+  by linarith
 
 lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
-  apply (subgoal_tac "z < real(natfloor z) + 1")
-  apply arith
-  apply (rule real_natfloor_add_one_gt)
-done
+  by linarith
 
 lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
-  unfolding natfloor_def
-  unfolding real_of_int_of_nat_eq [symmetric] floor_add
-  by (simp add: nat_add_distrib)
+  by linarith
 
 lemma natfloor_add_numeral [simp]:
-    "~neg ((numeral n)::int) ==> 0 <= x ==>
-      natfloor (x + numeral n) = natfloor x + numeral n"
+    "0 <= x \<Longrightarrow> natfloor (x + numeral n) = natfloor x + numeral n"
   by (simp add: natfloor_add [symmetric])
 
 lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
-  by (simp add: natfloor_add [symmetric] del: One_nat_def)
+  by linarith
 
 lemma natfloor_subtract [simp]:
     "natfloor(x - real a) = natfloor x - a"
-  unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract
-  by simp
+  by linarith
 
 lemma natfloor_div_nat:
   assumes "1 <= x" and "y > 0"
@@ -1939,67 +1917,57 @@
     (auto simp add: le_natfloor_eq mult_mono' real_natfloor_le natfloor_neg)
 
 lemma natceiling_zero [simp]: "natceiling 0 = 0"
-  by (unfold natceiling_def, simp)
+  by linarith
 
 lemma natceiling_one [simp]: "natceiling 1 = 1"
-  by (unfold natceiling_def, simp)
+  by linarith
 
 lemma zero_le_natceiling [simp]: "0 <= natceiling x"
-  by (unfold natceiling_def, simp)
+  by linarith
 
 lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n"
-  by (unfold natceiling_def, simp)
+  by (simp add: natceiling_def)
 
 lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
-  by (unfold natceiling_def, simp)
+  by linarith
 
 lemma real_natceiling_ge: "x <= real(natceiling x)"
-  unfolding natceiling_def by (cases "x < 0", simp_all)
+  by linarith
 
 lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
-  unfolding natceiling_def by simp
+  by linarith
 
 lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
-  unfolding natceiling_def by (intro nat_mono ceiling_mono)
+  by linarith
 
 lemma natceiling_le: "x <= real a ==> natceiling x <= a"
-  unfolding natceiling_def real_of_nat_def
-  by (simp add: nat_le_iff ceiling_le_iff)
+  by linarith
 
 lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)"
-  unfolding natceiling_def real_of_nat_def
-  by (simp add: nat_le_iff ceiling_le_iff)
+  by linarith
 
 lemma natceiling_le_eq_numeral [simp]:
-    "~ neg((numeral n)::int) ==>
-      (natceiling x <= numeral n) = (x <= numeral n)"
+    "(natceiling x <= numeral n) = (x <= numeral n)"
   by (simp add: natceiling_le_eq)
 
 lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
-  unfolding natceiling_def
-  by (simp add: nat_le_iff ceiling_le_iff)
+  by linarith
 
 lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
-  unfolding natceiling_def
-  by (simp add: ceiling_eq2 [where n="int n"])
+  by linarith
 
-lemma natceiling_add [simp]: "0 <= x ==>
-    natceiling (x + real a) = natceiling x + a"
-  unfolding natceiling_def
-  unfolding real_of_int_of_nat_eq [symmetric] ceiling_add
-  by (simp add: nat_add_distrib)
+lemma natceiling_add [simp]: "0 <= x ==> natceiling (x + real a) = natceiling x + a"
+  by linarith
 
 lemma natceiling_add_numeral [simp]:
-    "~ neg ((numeral n)::int) ==> 0 <= x ==>
-      natceiling (x + numeral n) = natceiling x + numeral n"
+    "0 <= x ==> natceiling (x + numeral n) = natceiling x + numeral n"
   by (simp add: natceiling_add [symmetric])
 
 lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
-  by (simp add: natceiling_add [symmetric] del: One_nat_def)
+  by linarith
 
 lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a"
-  unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract
-  by simp
+  by linarith
 
 lemma Rats_no_top_le: "\<exists> q \<in> \<rat>. (x :: real) \<le> q"
   by (auto intro!: bexI[of _ "of_nat (natceiling x)"]) (metis real_natceiling_ge real_of_nat_def)
--- a/src/HOL/SMT.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/SMT.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -126,6 +126,7 @@
 ML_file "Tools/SMT/z3_proof_tools.ML"
 ML_file "Tools/SMT/z3_proof_literals.ML"
 ML_file "Tools/SMT/z3_proof_methods.ML"
+named_theorems z3_simp "simplification rules for Z3 proof reconstruction"
 ML_file "Tools/SMT/z3_proof_reconstruction.ML"
 ML_file "Tools/SMT/z3_model.ML"
 ML_file "Tools/SMT/smt_setup_solvers.ML"
@@ -135,7 +136,6 @@
   SMT_Normalize.setup #>
   SMTLIB_Interface.setup #>
   Z3_Interface.setup #>
-  Z3_Proof_Reconstruction.setup #>
   SMT_Setup_Solvers.setup
 *}
 
--- a/src/HOL/SMT2.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/SMT2.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -108,6 +108,7 @@
 ML_file "Tools/SMT2/smtlib2.ML"
 ML_file "Tools/SMT2/smtlib2_interface.ML"
 ML_file "Tools/SMT2/smtlib2_proof.ML"
+ML_file "Tools/SMT2/smtlib2_isar.ML"
 ML_file "Tools/SMT2/z3_new_proof.ML"
 ML_file "Tools/SMT2/z3_new_isar.ML"
 ML_file "Tools/SMT2/smt2_solver.ML"
@@ -117,6 +118,9 @@
 ML_file "Tools/SMT2/z3_new_replay_rules.ML"
 ML_file "Tools/SMT2/z3_new_replay_methods.ML"
 ML_file "Tools/SMT2/z3_new_replay.ML"
+ML_file "Tools/SMT2/verit_proof.ML"
+ML_file "Tools/SMT2/verit_isar.ML"
+ML_file "Tools/SMT2/verit_proof_parse.ML"
 ML_file "Tools/SMT2/smt2_systems.ML"
 
 method_setup smt2 = {*
--- a/src/HOL/SMT_Examples/SMT_Examples.certs	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-43550507f510d81bc4fb9ef8c1fd14424eaa9070 37 0
-#2 := false
-#10 := 0::Int
-decl f3 :: Int
-#7 := f3
-#12 := (<= f3 0::Int)
-#54 := (not #12)
-decl f4 :: Int
-#8 := f4
-#13 := (<= f4 0::Int)
-#9 := (* f3 f4)
-#11 := (<= #9 0::Int)
-#37 := (not #11)
-#44 := (or #37 #12 #13)
-#47 := (not #44)
-#14 := (or #12 #13)
-#15 := (implies #11 #14)
-#16 := (not #15)
-#50 := (iff #16 #47)
-#38 := (or #37 #14)
-#41 := (not #38)
-#48 := (iff #41 #47)
-#45 := (iff #38 #44)
-#46 := [rewrite]: #45
-#49 := [monotonicity #46]: #48
-#42 := (iff #16 #41)
-#39 := (iff #15 #38)
-#40 := [rewrite]: #39
-#43 := [monotonicity #40]: #42
-#51 := [trans #43 #49]: #50
-#36 := [asserted]: #16
-#52 := [mp #36 #51]: #47
-#55 := [not-or-elim #52]: #54
-#56 := (not #13)
-#57 := [not-or-elim #52]: #56
-#53 := [not-or-elim #52]: #11
-[th-lemma arith farkas 1 1 1 #53 #57 #55]: false
-unsat
--- a/src/HOL/SMT_Examples/SMT_Examples.certs2	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/SMT_Examples/SMT_Examples.certs2	Wed Aug 27 15:52:58 2014 +0200
@@ -1,10 +1,3 @@
-3aa17d1c77bc1a92bca05df291d11d81c645a931 6 0
-unsat
-((set-logic AUFLIA)
-(proof
-(let ((@x30 (rewrite (= (not true) false))))
-(mp (asserted (not true)) @x30 false))))
-
 572677daa32981bf8212986300f1362edf42a0c1 7 0
 unsat
 ((set-logic AUFLIA)
@@ -13,6 +6,13 @@
 (let ((@x40 (trans @x36 (rewrite (= (not true) false)) (= (not (or p$ (not p$))) false))))
 (mp (asserted (not (or p$ (not p$)))) @x40 false)))))
 
+3aa17d1c77bc1a92bca05df291d11d81c645a931 6 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x30 (rewrite (= (not true) false))))
+(mp (asserted (not true)) @x30 false))))
+
 dfd95b23f80baacb2acdc442487bd8121f072035 9 0
 unsat
 ((set-logic AUFLIA)
@@ -1033,7 +1033,7 @@
 (let ((@x59 (trans @x55 (rewrite (= (not true) false)) (= (not (< 5 (ite (<= 3 8) 8 3))) false))))
 (mp (asserted (not (< 5 (ite (<= 3 8) 8 3)))) @x59 false)))))))))
 
-2d63144daf211d89368e355b9b23a3b5118b7ba9 88 0
+6b0b089fbe179e8a27509c818f9a5e6847ac6bf2 88 0
 unsat
 ((set-logic AUFLIRA)
 (proof
@@ -1207,7 +1207,7 @@
 (let ((@x67 (mp (asserted (not (< a$ 0.0))) @x66 $x58)))
 ((_ th-lemma arith farkas 7 3/2 1) @x67 @x52 @x40 false)))))))))))))))))
 
-b9f61649fae66ac195bf2593181f9d76c958bfe2 22 0
+3a6df2b095b936aac9a1d533e306f2d31b4fb44e 22 0
 unsat
 ((set-logic AUFLIA)
 (proof
@@ -1390,7 +1390,7 @@
 (let ((@x433 (mp (not-or-elim @x205 (not $x57)) @x432 $x422)))
 (unit-resolution @x433 @x488 (mp @x478 @x480 $x44) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
 
-fbc59441c65d9a844da44405d06d138b55c5d187 933 0
+32286f9c5e71eb2b15c18f86f04c80931e2e307b 933 0
 unsat
 ((set-logic AUFLIA)
 (proof
@@ -2345,7 +2345,7 @@
 (let ((@x62 (monotonicity @x59 (= $x36 (not $x43)))))
 (mp (asserted $x36) (trans @x62 @x71 (= $x36 false)) false))))))))))))))))))
 
-d2037888c28cf52f7a45f66288246169de6f14ad 113 0
+faae12ee7efe4347f92e42776a0e0e57a624319c 113 0
 unsat
 ((set-logic <null>)
 (proof
@@ -2459,7 +2459,7 @@
 (let ((@x74 (mp (asserted $x36) @x73 $x67)))
 ((_ th-lemma arith farkas -1 1 1) @x74 (unit-resolution ((_ th-lemma arith) (or false $x305)) (true-axiom true) $x305) @x337 false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
 
-29e336c1b1dbb5e85401e6a2954560661ff3cadc 112 0
+57f344c9e787868c98d1e622f158c445c1899c96 112 0
 unsat
 ((set-logic <null>)
 (proof
@@ -2572,7 +2572,7 @@
 (let ((@x70 (mp (asserted (not (< (+ x$ (+ ?x29 ?x29)) (+ x$ 3)))) @x69 $x63)))
 ((_ th-lemma arith farkas -1 1 1) @x70 @x336 (unit-resolution ((_ th-lemma arith) (or false $x319)) (true-axiom true) $x319) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
 
-5bcedd8db3cccf5f1df2ef7ad1ca5e39c817a6f4 32 0
+3938db798ebafb55191dcdaf83a8615d1d59c0c3 32 0
 unsat
 ((set-logic <null>)
 (proof
@@ -2605,7 +2605,7 @@
 (let ((@x117 (unit-resolution ((_ th-lemma arith assign-bounds 1) (or $x102 (not $x100))) (unit-resolution ((_ th-lemma arith triangle-eq) (or (not $x95) $x100)) @x98 $x100) $x102)))
 (unit-resolution ((_ th-lemma arith triangle-eq) (or $x28 (not $x101) (not $x102))) @x117 @x110 @x30 false))))))))))))))))))))))))))))))
 
-97186805a3315ef9a1cc4847a2e19a6d09c77915 236 0
+353c8b65ed1b98772a89ffdae52f1cfae628dd6a 236 0
 unsat
 ((set-logic <null>)
 (proof
@@ -3224,7 +3224,7 @@
 (let ((@x66 (mp~ (mp (asserted $x33) @x60 $x56) (nnf-pos (refl (~ $x53 $x53)) (~ $x56 $x56)) $x56)))
 (unit-resolution @x66 @x464 false)))))))))))))))))))))))))
 
-a8cb4a130675f119ab8ba11cbe3a15041f18d2c6 62 0
+a02ae6c9688537bbe4c3be0d3dcebbc703417864 62 0
 unsat
 ((set-logic AUFLIA)
 (declare-fun ?v0!1 () Int)
@@ -3287,7 +3287,7 @@
 (let ((@x515 (unit-resolution (def-axiom (or z3name!0 $x220)) (unit-resolution @x131 @x238 $x88) $x220)))
 (unit-resolution ((_ th-lemma arith triangle-eq) (or (not $x220) (>= ?x96 3))) @x515 @x245 false))))))))))))))))))))))))))))))))))))))))))))))))))))))
 
-9e0e67e5bd5078ab683d440f1a73c518a403be1b 39 0
+9853592ad3514c1450e40271884a9f21f57ff85b 39 0
 unsat
 ((set-logic AUFLIA)
 (proof
@@ -3380,7 +3380,7 @@
 (let ((@x117 (and-elim (not-or-elim @x112 (and $x100 $x102)) $x102)))
 ((_ th-lemma arith farkas 1 1 1) @x117 @x116 @x118 false)))))))))))))))))))))))))))))))))))
 
-0d380fa4e68ab250e8351813b95695943794f02d 46 0
+9201a8009730b821ad6a3a2b64598e50ab5748ca 46 0
 unsat
 ((set-logic AUFLIRA)
 (declare-fun ?v1!1 () Int)
@@ -3600,6 +3600,33 @@
 (let ((@x73 (not-or-elim @x70 $x62)))
 (unit-resolution (unit-resolution ((_ th-lemma arith farkas 1 1) (or $x65 (not $x62))) @x73 $x65) @x74 false))))))))))))))))))
 
+d98ad8f668dead6f610669a52351ea0176a811b0 26 0
+unsat
+((set-logic <null>)
+(proof
+(let (($x58 (<= b$ 0)))
+(let (($x62 (or (not (and (not (<= a$ 0)) (not (<= (* a$ b$) 0)))) (not $x58))))
+(let (($x65 (not $x62)))
+(let (($x35 (not (=> (and (< 0 a$) (< 0 (* a$ b$))) (< 0 b$)))))
+(let (($x33 (< 0 b$)))
+(let (($x38 (or (not (and (< 0 a$) (< 0 (* a$ b$)))) $x33)))
+(let (($x56 (= (not (and (< 0 a$) (< 0 (* a$ b$)))) (not (and (not (<= a$ 0)) (not (<= (* a$ b$) 0)))))))
+(let ((?x30 (* a$ b$)))
+(let (($x48 (<= ?x30 0)))
+(let (($x49 (not $x48)))
+(let (($x44 (<= a$ 0)))
+(let (($x45 (not $x44)))
+(let (($x52 (and $x45 $x49)))
+(let (($x32 (and (< 0 a$) (< 0 ?x30))))
+(let ((@x54 (monotonicity (rewrite (= (< 0 a$) $x45)) (rewrite (= (< 0 ?x30) $x49)) (= $x32 $x52))))
+(let ((@x64 (monotonicity (monotonicity @x54 $x56) (rewrite (= $x33 (not $x58))) (= $x38 $x62))))
+(let ((@x43 (monotonicity (rewrite (= (=> $x32 $x33) $x38)) (= $x35 (not $x38)))))
+(let ((@x69 (trans @x43 (monotonicity @x64 (= (not $x38) $x65)) (= $x35 $x65))))
+(let ((@x74 (not-or-elim (mp (asserted $x35) @x69 $x65) $x58)))
+(let ((@x72 (and-elim (not-or-elim (mp (asserted $x35) @x69 $x65) $x52) $x45)))
+(let ((@x73 (and-elim (not-or-elim (mp (asserted $x35) @x69 $x65) $x52) $x49)))
+((_ th-lemma arith farkas 1 1 1) @x73 @x72 @x74 false))))))))))))))))))))))))
+
 271390ea915947de195c2202e30f90bb84689d60 26 0
 unsat
 ((set-logic <null>)
@@ -3627,33 +3654,6 @@
 (let ((@x92 (trans @x88 (rewrite (= (not true) false)) (= $x39 false))))
 (mp (asserted $x39) @x92 false))))))))))))))))))))))))
 
-d98ad8f668dead6f610669a52351ea0176a811b0 26 0
-unsat
-((set-logic <null>)
-(proof
-(let (($x58 (<= b$ 0)))
-(let (($x62 (or (not (and (not (<= a$ 0)) (not (<= (* a$ b$) 0)))) (not $x58))))
-(let (($x65 (not $x62)))
-(let (($x35 (not (=> (and (< 0 a$) (< 0 (* a$ b$))) (< 0 b$)))))
-(let (($x33 (< 0 b$)))
-(let (($x38 (or (not (and (< 0 a$) (< 0 (* a$ b$)))) $x33)))
-(let (($x56 (= (not (and (< 0 a$) (< 0 (* a$ b$)))) (not (and (not (<= a$ 0)) (not (<= (* a$ b$) 0)))))))
-(let ((?x30 (* a$ b$)))
-(let (($x48 (<= ?x30 0)))
-(let (($x49 (not $x48)))
-(let (($x44 (<= a$ 0)))
-(let (($x45 (not $x44)))
-(let (($x52 (and $x45 $x49)))
-(let (($x32 (and (< 0 a$) (< 0 ?x30))))
-(let ((@x54 (monotonicity (rewrite (= (< 0 a$) $x45)) (rewrite (= (< 0 ?x30) $x49)) (= $x32 $x52))))
-(let ((@x64 (monotonicity (monotonicity @x54 $x56) (rewrite (= $x33 (not $x58))) (= $x38 $x62))))
-(let ((@x43 (monotonicity (rewrite (= (=> $x32 $x33) $x38)) (= $x35 (not $x38)))))
-(let ((@x69 (trans @x43 (monotonicity @x64 (= (not $x38) $x65)) (= $x35 $x65))))
-(let ((@x74 (not-or-elim (mp (asserted $x35) @x69 $x65) $x58)))
-(let ((@x72 (and-elim (not-or-elim (mp (asserted $x35) @x69 $x65) $x52) $x45)))
-(let ((@x73 (and-elim (not-or-elim (mp (asserted $x35) @x69 $x65) $x52) $x49)))
-((_ th-lemma arith farkas 1 1 1) @x73 @x72 @x74 false))))))))))))))))))))))))
-
 b216c79478e44396acca3654b76845499fc18a04 23 0
 unsat
 ((set-logic <null>)
@@ -3944,6 +3944,21 @@
 (let ((@x53 (trans (monotonicity @x46 (= $x33 (not true))) (rewrite (= (not true) false)) (= $x33 false))))
 (mp (asserted $x33) @x53 false)))))))))))
 
+8b09776b03122aeacc9dd9526e1f0e5d41a07f14 14 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x29 (forall ((?v0 A$) )(g$ ?v0))
+))
+(let (($x30 (ite $x29 true false)))
+(let (($x31 (f$ $x30)))
+(let (($x32 (=> $x31 true)))
+(let (($x33 (not $x32)))
+(let ((@x42 (monotonicity (monotonicity (rewrite (= $x30 $x29)) (= $x31 (f$ $x29))) (= $x32 (=> (f$ $x29) true)))))
+(let ((@x46 (trans @x42 (rewrite (= (=> (f$ $x29) true) true)) (= $x32 true))))
+(let ((@x53 (trans (monotonicity @x46 (= $x33 (not true))) (rewrite (= (not true) false)) (= $x33 false))))
+(mp (asserted $x33) @x53 false)))))))))))
+
 b221de9d8dbe279344ac85e2ada07f5722636ce5 46 0
 unsat
 ((set-logic AUFLIA)
@@ -3991,21 +4006,6 @@
 (let ((@x478 (mp ((_ quant-inst 3 42) (or (not $x52) $x171)) (trans (monotonicity @x131 $x137) (rewrite (= $x134 $x134)) $x137) $x134)))
 (unit-resolution (unit-resolution @x478 @x78 $x168) (mp @x77 @x472 (not $x168)) false)))))))))))))))))))))))))))))))))))
 
-8b09776b03122aeacc9dd9526e1f0e5d41a07f14 14 0
-unsat
-((set-logic AUFLIA)
-(proof
-(let (($x29 (forall ((?v0 A$) )(g$ ?v0))
-))
-(let (($x30 (ite $x29 true false)))
-(let (($x31 (f$ $x30)))
-(let (($x32 (=> $x31 true)))
-(let (($x33 (not $x32)))
-(let ((@x42 (monotonicity (monotonicity (rewrite (= $x30 $x29)) (= $x31 (f$ $x29))) (= $x32 (=> (f$ $x29) true)))))
-(let ((@x46 (trans @x42 (rewrite (= (=> (f$ $x29) true) true)) (= $x32 true))))
-(let ((@x53 (trans (monotonicity @x46 (= $x33 (not true))) (rewrite (= (not true) false)) (= $x33 false))))
-(mp (asserted $x33) @x53 false)))))))))))
-
 5d3ccbcf168a634cad3952ad8f6d2798329d6a77 75 0
 unsat
 ((set-logic AUFLIA)
@@ -4204,7 +4204,7 @@
 (let ((@x81 (asserted $x80)))
 (unit-resolution @x81 (trans @x397 ((_ th-lemma arith eq-propagate 1 1 -4 -4) @x410 @x422 @x428 @x438 (= ?x490 6)) $x79) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
 
-fa62bf7228a50eb8c663092f87f9af7c25feaffe 336 0
+640bb6103260f74026864b86c2301c00ea737ff0 336 0
 unsat
 ((set-logic <null>)
 (proof
--- a/src/HOL/SMT_Examples/SMT_Examples.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/SMT_Examples/SMT_Examples.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -8,9 +8,6 @@
 imports Complex_Main
 begin
 
-declare [[smt_certificates = "SMT_Examples.certs"]]
-declare [[smt_read_only_certificates = true]]
-
 declare [[smt2_certificates = "SMT_Examples.certs2"]]
 declare [[smt2_read_only_certificates = true]]
 
@@ -382,16 +379,12 @@
    U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)"
   using [[z3_new_extensions]] by smt2
 
-lemma [z3_rule, z3_new_rule]:
+lemma [z3_new_rule]:
   fixes x :: "int"
   assumes "x * y \<le> 0" and "\<not> y \<le> 0" and "\<not> x \<le> 0"
   shows False
   using assms by (metis mult_le_0_iff)
 
-lemma "x * y \<le> (0 :: int) \<Longrightarrow> x \<le> 0 \<or> y \<le> 0"
-  using [[z3_with_extensions]] [[z3_new_extensions]]
-  by smt (* smt2 FIXME: "th-lemma" tactic fails *)
-
 
 section {* Pairs *}
 
--- a/src/HOL/SMT_Examples/SMT_Tests.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -8,7 +8,6 @@
 imports Complex_Main
 begin
 
-smt_status
 smt2_status
 
 text {* Most examples are taken from various Isabelle theories and from HOL4. *}
@@ -588,7 +587,7 @@
   "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
      p \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
   using point.simps bw_point.simps
-  by smt+ (* smt2 FIXME: bad Z3 4.3.x proof *)
+  sorry (* smt2 FIXME: bad Z3 4.3.x proof *)
 
 lemma
   "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> black := w \<rparr> = \<lparr> cx = 3, cy = 4, black = w \<rparr>"
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.certs2	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.certs2	Wed Aug 27 15:52:58 2014 +0200
@@ -43,14 +43,6 @@
 (let ((@x49 (trans @x45 (rewrite (= (not true) false)) (= (not (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8))) false))))
 (mp (asserted (not (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8)))) @x49 false)))))))
 
-6dd848d7b26e0521039e21a5e2bafebc1fee8c9b 7 0
-unsat
-((set-logic <null>)
-(proof
-(let ((@x35 (monotonicity (rewrite (= (= (_ bv11 5) (_ bv11 5)) true)) (= (not (= (_ bv11 5) (_ bv11 5))) (not true)))))
-(let ((@x39 (trans @x35 (rewrite (= (not true) false)) (= (not (= (_ bv11 5) (_ bv11 5))) false))))
-(mp (asserted (not (= (_ bv11 5) (_ bv11 5)))) @x39 false)))))
-
 45bf9e0a746f7dde46f8242e5851928c2671c7f2 9 0
 unsat
 ((set-logic <null>)
@@ -61,6 +53,14 @@
 (let ((@x54 (trans @x50 (rewrite (= (not true) false)) (= (not (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8)))) false))))
 (mp (asserted (not (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8))))) @x54 false)))))))
 
+6dd848d7b26e0521039e21a5e2bafebc1fee8c9b 7 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (= (_ bv11 5) (_ bv11 5)) true)) (= (not (= (_ bv11 5) (_ bv11 5))) (not true)))))
+(let ((@x39 (trans @x35 (rewrite (= (not true) false)) (= (not (= (_ bv11 5) (_ bv11 5))) false))))
+(mp (asserted (not (= (_ bv11 5) (_ bv11 5)))) @x39 false)))))
+
 00a7ff287d9c23d984163ea8e0cac8ac61008afd 11 0
 unsat
 ((set-logic <null>)
@@ -142,6 +142,15 @@
 (let ((@x48 (trans @x44 (rewrite (= (not true) false)) (= (not (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8))) false))))
 (mp (asserted (not (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8)))) @x48 false)))))))
 
+3838eb33edcd292c3a0ecbf1662b54af3940189f 8 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x36 (monotonicity (rewrite (= (bvnot (_ bv240 16)) (_ bv65295 16))) (= (= (bvnot (_ bv240 16)) (_ bv65295 16)) (= (_ bv65295 16) (_ bv65295 16))))))
+(let ((@x40 (trans @x36 (rewrite (= (= (_ bv65295 16) (_ bv65295 16)) true)) (= (= (bvnot (_ bv240 16)) (_ bv65295 16)) true))))
+(let ((@x47 (trans (monotonicity @x40 (= (not (= (bvnot (_ bv240 16)) (_ bv65295 16))) (not true))) (rewrite (= (not true) false)) (= (not (= (bvnot (_ bv240 16)) (_ bv65295 16))) false))))
+(mp (asserted (not (= (bvnot (_ bv240 16)) (_ bv65295 16)))) @x47 false))))))
+
 14431ccb33137348161eb6ca6cfb2e57942c3f9d 9 0
 unsat
 ((set-logic <null>)
@@ -152,14 +161,15 @@
 (let ((@x48 (trans @x44 (rewrite (= (not true) false)) (= (not (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12))) false))))
 (mp (asserted (not (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12)))) @x48 false)))))))
 
-3838eb33edcd292c3a0ecbf1662b54af3940189f 8 0
+880bee16a8f6469b14122277b92e87533ef6a071 9 0
 unsat
 ((set-logic <null>)
 (proof
-(let ((@x36 (monotonicity (rewrite (= (bvnot (_ bv240 16)) (_ bv65295 16))) (= (= (bvnot (_ bv240 16)) (_ bv65295 16)) (= (_ bv65295 16) (_ bv65295 16))))))
-(let ((@x40 (trans @x36 (rewrite (= (= (_ bv65295 16) (_ bv65295 16)) true)) (= (= (bvnot (_ bv240 16)) (_ bv65295 16)) true))))
-(let ((@x47 (trans (monotonicity @x40 (= (not (= (bvnot (_ bv240 16)) (_ bv65295 16))) (not true))) (rewrite (= (not true) false)) (= (not (= (bvnot (_ bv240 16)) (_ bv65295 16))) false))))
-(mp (asserted (not (= (bvnot (_ bv240 16)) (_ bv65295 16)))) @x47 false))))))
+(let ((@x37 (monotonicity (rewrite (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) (= (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)) (= (_ bv207 10) (_ bv207 10))))))
+(let ((@x41 (trans @x37 (rewrite (= (= (_ bv207 10) (_ bv207 10)) true)) (= (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)) true))))
+(let ((@x44 (monotonicity @x41 (= (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) (not true)))))
+(let ((@x48 (trans @x44 (rewrite (= (not true) false)) (= (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) false))))
+(mp (asserted (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)))) @x48 false)))))))
 
 8d29c9b5ef712a3d9d2db87383c9c25c0b553e01 8 0
 unsat
@@ -170,16 +180,6 @@
 (let ((@x47 (trans (monotonicity @x40 (= (not (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2))) (not true))) (rewrite (= (not true) false)) (= (not (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2))) false))))
 (mp (asserted (not (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2)))) @x47 false))))))
 
-880bee16a8f6469b14122277b92e87533ef6a071 9 0
-unsat
-((set-logic <null>)
-(proof
-(let ((@x37 (monotonicity (rewrite (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) (= (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)) (= (_ bv207 10) (_ bv207 10))))))
-(let ((@x41 (trans @x37 (rewrite (= (= (_ bv207 10) (_ bv207 10)) true)) (= (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)) true))))
-(let ((@x44 (monotonicity @x41 (= (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) (not true)))))
-(let ((@x48 (trans @x44 (rewrite (= (not true) false)) (= (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) false))))
-(mp (asserted (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)))) @x48 false)))))))
-
 446b3cb9d63aa1f488dc092ed3fb111d2ad50b4e 9 0
 unsat
 ((set-logic <null>)
@@ -240,6 +240,16 @@
 (let ((@x47 (trans @x43 (rewrite (= (not true) false)) (= (not (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4))) false))))
 (mp (asserted (not (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4)))) @x47 false)))))))
 
+33a52e620069e1ecebbc00aa6b0099170558c111 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x36 (monotonicity (rewrite (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) (= (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)) (= (_ bv13 4) (_ bv13 4))))))
+(let ((@x40 (trans @x36 (rewrite (= (= (_ bv13 4) (_ bv13 4)) true)) (= (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)) true))))
+(let ((@x43 (monotonicity @x40 (= (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) (not true)))))
+(let ((@x47 (trans @x43 (rewrite (= (not true) false)) (= (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) false))))
+(mp (asserted (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)))) @x47 false)))))))
+
 5c4e275bed2d91897e820ccd6744b0a775a6e26e 17 0
 unsat
 ((set-logic <null>)
@@ -258,46 +268,6 @@
 (let ((@x55 (unit-resolution ((_ th-lemma bv) $x53) @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 $x52)))
 (unit-resolution @x55 @x63 false)))))))))))))))
 
-33a52e620069e1ecebbc00aa6b0099170558c111 9 0
-unsat
-((set-logic <null>)
-(proof
-(let ((@x36 (monotonicity (rewrite (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) (= (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)) (= (_ bv13 4) (_ bv13 4))))))
-(let ((@x40 (trans @x36 (rewrite (= (= (_ bv13 4) (_ bv13 4)) true)) (= (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)) true))))
-(let ((@x43 (monotonicity @x40 (= (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) (not true)))))
-(let ((@x47 (trans @x43 (rewrite (= (not true) false)) (= (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) false))))
-(mp (asserted (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)))) @x47 false)))))))
-
-115ab22c9945d493b971e69a38d9e608c5b40a71 29 0
-unsat
-((set-logic <null>)
-(proof
-(let ((?x28 (bv2int$ (_ bv0 2))))
-(let (($x183 (<= ?x28 0)))
-(let (($x184 (not $x183)))
-(let (($x175 (forall ((?v0 (_ BitVec 2)) )(!(let ((?x47 (bv2int$ ?v0)))
-(let (($x53 (<= ?x47 0)))
-(not $x53))) :pattern ( (bv2int$ ?v0) )))
-))
-(let (($x57 (forall ((?v0 (_ BitVec 2)) )(let ((?x47 (bv2int$ ?v0)))
-(let (($x53 (<= ?x47 0)))
-(not $x53))))
-))
-(let ((@x177 (refl (= (not (<= (bv2int$ ?0) 0)) (not (<= (bv2int$ ?0) 0))))))
-(let ((@x112 (refl (~ (not (<= (bv2int$ ?0) 0)) (not (<= (bv2int$ ?0) 0))))))
-(let (($x49 (forall ((?v0 (_ BitVec 2)) )(let ((?x47 (bv2int$ ?v0)))
-(< 0 ?x47)))
-))
-(let ((@x56 (rewrite (= (< 0 (bv2int$ ?0)) (not (<= (bv2int$ ?0) 0))))))
-(let ((@x115 (mp~ (mp (asserted $x49) (quant-intro @x56 (= $x49 $x57)) $x57) (nnf-pos @x112 (~ $x57 $x57)) $x57)))
-(let ((@x180 (mp @x115 (quant-intro @x177 (= $x57 $x175)) $x175)))
-(let (($x187 (not $x175)))
-(let (($x188 (or $x187 $x184)))
-(let ((@x189 ((_ quant-inst (_ bv0 2)) $x188)))
-(let (($x29 (= ?x28 0)))
-(let ((@x30 (asserted $x29)))
-(unit-resolution ((_ th-lemma arith triangle-eq) (or (not $x29) $x183)) @x30 (unit-resolution @x189 @x180 $x184) false)))))))))))))))))))
-
 1d4a0e2a4449a8adbcf5a134bf7f2b0ee940d954 51 0
 unsat
 ((set-logic <null>)
@@ -350,6 +320,36 @@
 (let ((@x314 (unit-resolution ((_ th-lemma bv) $x312) @x26 @x26 @x26 @x26 @x26 @x26 @x26 @x26 (unit-resolution (def-axiom (or $x95 (not $x74))) @x303 (not $x74)) (unit-resolution (def-axiom (or $x95 (not $x75))) @x303 (not $x75)) (unit-resolution (def-axiom (or $x95 (not $x76))) @x303 (not $x76)) (unit-resolution (def-axiom (or $x95 (not $x77))) @x303 (not $x77)) (unit-resolution (def-axiom (or $x95 (not $x78))) @x303 (not $x78)) (unit-resolution (def-axiom (or $x95 (not $x79))) @x303 (not $x79)) (unit-resolution (def-axiom (or $x95 (not $x80))) @x303 (not $x80)) (unit-resolution (def-axiom (or $x95 $x264)) @x303 $x264) $x300)))
 (unit-resolution @x314 @x322 false)))))))))))))))))))))))))))))))))))))))))))))))))
 
+115ab22c9945d493b971e69a38d9e608c5b40a71 29 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x28 (bv2int$ (_ bv0 2))))
+(let (($x183 (<= ?x28 0)))
+(let (($x184 (not $x183)))
+(let (($x175 (forall ((?v0 (_ BitVec 2)) )(!(let ((?x47 (bv2int$ ?v0)))
+(let (($x53 (<= ?x47 0)))
+(not $x53))) :pattern ( (bv2int$ ?v0) )))
+))
+(let (($x57 (forall ((?v0 (_ BitVec 2)) )(let ((?x47 (bv2int$ ?v0)))
+(let (($x53 (<= ?x47 0)))
+(not $x53))))
+))
+(let ((@x177 (refl (= (not (<= (bv2int$ ?0) 0)) (not (<= (bv2int$ ?0) 0))))))
+(let ((@x112 (refl (~ (not (<= (bv2int$ ?0) 0)) (not (<= (bv2int$ ?0) 0))))))
+(let (($x49 (forall ((?v0 (_ BitVec 2)) )(let ((?x47 (bv2int$ ?v0)))
+(< 0 ?x47)))
+))
+(let ((@x56 (rewrite (= (< 0 (bv2int$ ?0)) (not (<= (bv2int$ ?0) 0))))))
+(let ((@x115 (mp~ (mp (asserted $x49) (quant-intro @x56 (= $x49 $x57)) $x57) (nnf-pos @x112 (~ $x57 $x57)) $x57)))
+(let ((@x180 (mp @x115 (quant-intro @x177 (= $x57 $x175)) $x175)))
+(let (($x187 (not $x175)))
+(let (($x188 (or $x187 $x184)))
+(let ((@x189 ((_ quant-inst (_ bv0 2)) $x188)))
+(let (($x29 (= ?x28 0)))
+(let ((@x30 (asserted $x29)))
+(unit-resolution ((_ th-lemma arith triangle-eq) (or (not $x29) $x183)) @x30 (unit-resolution @x189 @x180 $x184) false)))))))))))))))))))
+
 d14e7b8f0d1858316e700b4eb09e7d03e57cf9c3 16 0
 unsat
 ((set-logic <null>)
--- a/src/HOL/TPTP/ATP_Problem_Import.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/ATP_Problem_Import.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -5,10 +5,7 @@
 header {* ATP Problem Importer *}
 
 theory ATP_Problem_Import
-imports
-  Complex_Main
-  TPTP_Interpret
-  "~~/src/HOL/Library/Refute"
+imports Complex_Main TPTP_Interpret "~~/src/HOL/Library/Refute"
 begin
 
 ML_file "sledgehammer_tactics.ML"
@@ -24,8 +21,7 @@
 ML_file "atp_problem_import.ML"
 
 (*
-ML {* ATP_Problem_Import.isabelle_tptp_file @{theory} 50
-          "$TPTP/Problems/PUZ/PUZ107^5.p" *}
+ML {* ATP_Problem_Import.isabelle_tptp_file @{theory} 50 "$TPTP/Problems/PUZ/PUZ107^5.p" *}
 *)
 
 end
--- a/src/HOL/TPTP/THF_Arith.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/THF_Arith.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -51,7 +51,7 @@
 
 overloading real_to_rat \<equiv> "to_rat :: real \<Rightarrow> rat"
 begin
-  definition "real_to_rat (x\<Colon>real) = (inv real x\<Colon>rat)"
+  definition "real_to_rat (x\<Colon>real) = (inv of_rat x\<Colon>rat)"
 end
 
 overloading int_to_real \<equiv> "to_real :: int \<Rightarrow> real"
@@ -85,6 +85,6 @@
 by (metis Rats_of_rat rat_to_real_def real_is_rat_def)
 
 lemma inj_of_rat [intro, simp]: "inj (of_rat\<Colon>rat\<Rightarrow>real)"
-by (metis injI of_rat_eq_iff rat_to_real_def)
+by (metis injI of_rat_eq_iff)
 
 end
--- a/src/HOL/TPTP/TPTP_Interpret.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Interpret.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -6,12 +6,12 @@
 *)
 
 theory TPTP_Interpret
-imports Main TPTP_Parser
+imports Complex_Main TPTP_Parser
 keywords "import_tptp" :: thy_decl
 begin
 
-typedecl "ind"
+typedecl ind
 
-ML_file  "TPTP_Parser/tptp_interpret.ML"
+ML_file "TPTP_Parser/tptp_interpret.ML"
 
 end
\ No newline at end of file
--- a/src/HOL/TPTP/TPTP_Parser/README	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/README	Wed Aug 27 15:52:58 2014 +0200
@@ -29,4 +29,4 @@
 ML-Yacc's library.
 
 Nik Sultana
-9th March 2012
\ No newline at end of file
+9th March 2012
--- a/src/HOL/TPTP/TPTP_Parser/tptp.lex	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp.lex	Wed Aug 27 15:52:58 2014 +0200
@@ -86,6 +86,7 @@
 upper_word                = {upper_alpha}{alpha_numeric}*;
 dollar_dollar_word        = {ddollar}{lower_word};
 dollar_word               = {dollar}{lower_word};
+dollar_underscore         = {dollar}_;
 
 distinct_object           = {double_quote}{do_char}+{double_quote};
 
@@ -177,6 +178,7 @@
 
 {lower_word}   => (col:=yypos-(!eolpos); T.LOWER_WORD(yytext,!linep,!col));
 {dollar_word}  => (col:=yypos-(!eolpos); T.DOLLAR_WORD(yytext,!linep,!col));
+{dollar_underscore}  => (col:=yypos-(!eolpos); T.DOLLAR_WORD(yytext,!linep,!col));
 {dollar_dollar_word}  => (col:=yypos-(!eolpos); T.DOLLAR_DOLLAR_WORD(yytext,!linep,!col));
 
 "+"           => (col:=yypos-(!eolpos); T.PLUS(!linep,!col));
--- a/src/HOL/TPTP/TPTP_Parser/tptp.yacc	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp.yacc	Wed Aug 27 15:52:58 2014 +0200
@@ -491,7 +491,7 @@
 tff_unitary_type : tff_atomic_type               (( tff_atomic_type ))
                  | LPAREN tff_xprod_type RPAREN  (( tff_xprod_type ))
 
-tff_atomic_type : atomic_word   (( Atom_type atomic_word ))
+tff_atomic_type : atomic_word   (( Atom_type (atomic_word, []) ))
                 | defined_type  (( Defined_type defined_type ))
                 | atomic_word LPAREN tff_type_arguments RPAREN (( Fmla_type (Fmla (Uninterpreted atomic_word, (map Type_fmla tff_type_arguments))) ))
                 | variable_ (( Fmla_type (Pred (Interpreted_ExtraLogic Apply, [Term_Var variable_])) ))
@@ -634,6 +634,7 @@
   | "$real" => Type_Real
   | "$rat" => Type_Rat
   | "$int" => Type_Int
+  | "$_" => Type_Dummy
   | thing => raise UNRECOGNISED_SYMBOL ("defined_type", thing)
 ))
 
@@ -747,6 +748,7 @@
   | "$real" => TypeSymbol Type_Real
   | "$rat" => TypeSymbol Type_Rat
   | "$tType" => TypeSymbol Type_Type
+  | "$_" => TypeSymbol Type_Dummy
 
   | "$true" => Interpreted_Logic True
   | "$false" => Interpreted_Logic False
--- a/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_interpret.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -8,11 +8,11 @@
 sig
   (*Signature extension: typing information for variables and constants*)
   type var_types = (string * typ option) list
-  type const_map = (string * term) list
+  type const_map = (string * (term * int list)) list
 
   (*Mapping from TPTP types to Isabelle/HOL types. This map works over all
   base types. The map must be total wrt TPTP types.*)
-  type type_map = (TPTP_Syntax.tptp_type * typ) list
+  type type_map = (string * (string * int)) list
 
   (*A parsed annotated formula is represented as a 5-tuple consisting of
   the formula's label, its role, the TPTP formula, its Isabelle/HOL meaning, and
@@ -35,7 +35,7 @@
      problem_name : TPTP_Problem_Name.problem_name option}
 
   (*Generates a fresh Isabelle/HOL type for interpreting a TPTP type in a theory.*)
-  val declare_type : config -> (TPTP_Syntax.tptp_type * string) -> type_map ->
+  val declare_type : config -> string * (string * int) -> type_map ->
     theory -> type_map * theory
 
   (*Map TPTP types to Isabelle/HOL types*)
@@ -132,9 +132,9 @@
 
 (** Signatures and other type abbrevations **)
 
-type const_map = (string * term) list
+type const_map = (string * (term * int list)) list
 type var_types = (string * typ option) list
-type type_map = (TPTP_Syntax.tptp_type * typ) list
+type type_map = (string * (string * int)) list
 type tptp_formula_meaning =
   string * TPTP_Syntax.role * term * TPTP_Proof.source_info option
 type tptp_general_meaning =
@@ -189,17 +189,19 @@
 (*Returns updated theory and the name of the final type's name -- i.e. if the
 original name is already taken then the function looks for an available
 alternative. It also returns an updated type_map if one has been passed to it.*)
-fun declare_type (config : config) (thf_type, type_name) ty_map thy =
+fun declare_type (config : config) (thf_type, (type_name, arity)) ty_map thy =
   if type_exists config thy type_name then
-    raise MISINTERPRET_TYPE ("Type already exists", Atom_type type_name)
+    raise MISINTERPRET_TYPE ("Type already exists", Atom_type (type_name, []))
   else
     let
       val binding = mk_binding config type_name
       val final_fqn = Sign.full_name thy binding
+      val tyargs =
+        List.tabulate (arity, rpair @{sort type} o prefix "'" o string_of_int)
       val thy' =
-        Typedecl.typedecl_global (mk_binding config type_name, [], NoSyn) thy
+        Typedecl.typedecl_global (mk_binding config type_name, tyargs, NoSyn) thy
         |> snd
-      val typ_map_entry = (thf_type, (Type (final_fqn, [])))
+      val typ_map_entry = (thf_type, (final_fqn, arity))
       val ty_map' = typ_map_entry :: ty_map
     in (ty_map', thy') end
 
@@ -217,42 +219,56 @@
     raise MISINTERPRET_TERM
      ("Const already declared", Term_Func (Uninterpreted const_name, []))
   else
-    Theory.specify_const
-     ((mk_binding config const_name, ty), NoSyn) thy
-
-fun declare_const_ifnot config const_name ty thy =
-  if const_exists config thy const_name then
-    (Sign.mk_const thy ((Sign.full_name thy (mk_binding config const_name)), []), thy)
-  else declare_constant config const_name ty thy
+    Theory.specify_const ((mk_binding config const_name, ty), NoSyn) thy
 
 
 (** Interpretation functions **)
 
-(*Types in THF are encoded as formulas. This function translates them to type form.*)
+(*Types in TFF/THF are encoded as formulas. These functions translate them to type form.*)
+
+fun termtype_to_type (Term_Func (TypeSymbol typ, [])) =
+      Defined_type typ
+  | termtype_to_type (Term_Func (Uninterpreted str, tms)) =
+      Atom_type (str, map termtype_to_type tms)
+  | termtype_to_type (Term_Var str) = Var_type str
+
 (*FIXME possibly incomplete hack*)
-fun fmlatype_to_type (Atom (THF_Atom_term (Term_Func (TypeSymbol typ, [])))) =
-      Defined_type typ
-  | fmlatype_to_type (Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))) =
-      Atom_type str
+fun fmlatype_to_type (Atom (THF_Atom_term tm)) = termtype_to_type tm
   | fmlatype_to_type (Type_fmla (Fn_type (ty1, ty2))) =
       let
         val ty1' = case ty1 of
             Fn_type _ => fmlatype_to_type (Type_fmla ty1)
           | Fmla_type ty1 => fmlatype_to_type ty1
+          | _ => ty1
         val ty2' = case ty2 of
             Fn_type _ => fmlatype_to_type (Type_fmla ty2)
           | Fmla_type ty2 => fmlatype_to_type ty2
+          | _ => ty2
       in Fn_type (ty1', ty2') end
+  | fmlatype_to_type (Type_fmla ty) = ty
+  | fmlatype_to_type (Fmla (Uninterpreted str, fmla)) =
+      Atom_type (str, map fmlatype_to_type fmla)
+  | fmlatype_to_type (Quant (Dep_Prod, _, fmla)) = fmlatype_to_type fmla
+  | fmlatype_to_type (Pred (Interpreted_ExtraLogic Apply, [tm])) =
+      termtype_to_type tm
+
+fun tfree_name_of_var_type str = "'" ^ Name.desymbolize (SOME false) str
+fun tfree_of_var_type str = TFree (tfree_name_of_var_type str, @{sort type})
 
 fun interpret_type config thy type_map thf_ty =
   let
-    fun lookup_type ty_map thf_ty =
-      (case AList.lookup (op =) ty_map thf_ty of
+    fun lookup_type ty_map ary str =
+      (case AList.lookup (op =) ty_map str of
           NONE =>
-            raise MISINTERPRET_TYPE
+            raise MISINTERPRET_SYMBOL
               ("Could not find the interpretation of this TPTP type in the \
-               \mapping to Isabelle/HOL", thf_ty)
-        | SOME ty => ty)
+               \mapping to Isabelle/HOL", Uninterpreted str)
+        | SOME (str', ary') =>
+            if ary' = ary then
+              str'
+            else
+              raise MISINTERPRET_SYMBOL ("TPTP type used with wrong arity",
+                Uninterpreted str))
   in
     case thf_ty of
        Prod_type (thf_ty1, thf_ty2) =>
@@ -263,22 +279,19 @@
          Type (@{type_name fun},
           [interpret_type config thy type_map thf_ty1,
            interpret_type config thy type_map thf_ty2])
-     | Atom_type _ =>
-         lookup_type type_map thf_ty
+     | Atom_type (str, thf_tys) =>
+         Type (lookup_type type_map (length thf_tys) str,
+           map (interpret_type config thy type_map) thf_tys)
+     | Var_type str => tfree_of_var_type str
      | Defined_type tptp_base_type =>
          (case tptp_base_type of
             Type_Ind => @{typ ind}
           | Type_Bool => HOLogic.boolT
           | Type_Type => raise MISINTERPRET_TYPE ("No type interpretation", thf_ty)
-         (*FIXME these types are currently unsupported, so they're treated as
-         individuals*)
-          | Type_Int =>
-              interpret_type config thy type_map (Defined_type Type_Ind)
-          | Type_Rat =>
-              interpret_type config thy type_map (Defined_type Type_Ind)
-          | Type_Real =>
-              interpret_type config thy type_map (Defined_type Type_Ind)
-          )
+          | Type_Int => @{typ int}
+          | Type_Rat => @{typ rat}
+          | Type_Real => @{typ real}
+          | Type_Dummy => dummyT)
      | Sum_type _ =>
          raise MISINTERPRET_TYPE (*FIXME*)
           ("No type interpretation (sum type)", thf_ty)
@@ -290,19 +303,15 @@
           ("No type interpretation (subtype)", thf_ty)
   end
 
-fun the_const config thy language const_map_payload str =
-  if language = THF then
-    (case AList.lookup (op =) const_map_payload str of
-        NONE => raise MISINTERPRET_TERM
-          ("Could not find the interpretation of this constant in the \
-            \mapping to Isabelle/HOL", Term_Func (Uninterpreted str, []))
-      | SOME t => t)
-  else
-    if const_exists config thy str then
-       Sign.mk_const thy ((Sign.full_name thy (mk_binding config str)), [])
-    else raise MISINTERPRET_TERM
-          ("Could not find the interpretation of this constant in the \
-            \mapping to Isabelle/HOL", Term_Func (Uninterpreted str, []))
+fun permute_type_args perm Ts = map (nth Ts) perm
+
+fun the_const thy const_map str tyargs =
+  (case AList.lookup (op =) const_map str of
+      SOME (Const (s, _), tyarg_perm) =>
+      Sign.mk_const thy (s, permute_type_args tyarg_perm tyargs)
+    | _ => raise MISINTERPRET_TERM
+        ("Could not find the interpretation of this constant in the \
+          \mapping to Isabelle/HOL", Term_Func (Uninterpreted str, [])))
 
 (*Eta-expands n-ary function.
  "str" is the name of an Isabelle/HOL constant*)
@@ -346,18 +355,24 @@
   | Is_Int => 1
   | Is_Rat => 1
   | Distinct => 1
-  | Apply=> 2
+  | Apply => 2
 
-fun interpret_symbol config language const_map symb thy =
+fun type_arity_of_symbol thy config (Uninterpreted n) =
+    let val s = full_name thy config n in
+      length (Sign.const_typargs thy (s, Sign.the_const_type thy s))
+    end
+  | type_arity_of_symbol _ _ _ = 0
+
+fun interpret_symbol config const_map symb tyargs thy =
   case symb of
      Uninterpreted n =>
        (*Constant would have been added to the map at the time of
        declaration*)
-       the_const config thy language const_map n
+       the_const thy const_map n tyargs
    | Interpreted_ExtraLogic interpreted_symbol =>
        (*FIXME not interpreting*)
        Sign.mk_const thy ((Sign.full_name thy (mk_binding config
-          (string_of_interpreted_symbol interpreted_symbol))), [])
+          (string_of_interpreted_symbol interpreted_symbol))), tyargs)
    | Interpreted_Logic logic_symbol =>
        (case logic_symbol of
           Equals => HOLogic.eq_const dummyT
@@ -427,16 +442,14 @@
         in
           case symb of
              Uninterpreted const_name =>
-               declare_const_ifnot config const_name
-                (mk_fun_type (replicate (length tptp_ts) ind_type) value_type I) thy'
-               |> snd
+               perhaps (try (snd o declare_constant config const_name
+                (mk_fun_type (replicate (length tptp_ts) ind_type) value_type I))) thy'
            | _ => thy'
         end
     | Atom_App _ => thy
     | Atom_Arity (const_name, n_args) =>
-        declare_const_ifnot config const_name
-         (mk_fun_type (replicate n_args ind_type) value_type I) thy
-        |> snd
+        perhaps (try (snd o declare_constant config const_name
+         (mk_fun_type (replicate n_args ind_type) value_type I))) thy
   end
 
 (*FIXME only used until interpretation is implemented*)
@@ -499,24 +512,32 @@
                (interpret_term formula_level config language const_map
                 var_types type_map (hd tptp_ts)))
            | _ =>
-              (*apply symb to tptp_ts*)
-              if is_prod_typed thy' config symb then
-                let
-                  val (t, thy'') =
-                    mtimes'
-                      (fold_map (interpret_term false config language const_map
-                       var_types type_map) (tl tptp_ts) thy')
-                      (interpret_term false config language const_map
-                       var_types type_map (hd tptp_ts))
-                in (interpret_symbol config language const_map symb thy' $ t, thy'')
-                end
-              else
-                (
-                mapply'
-                  (fold_map (interpret_term false config language const_map
-                   var_types type_map) tptp_ts thy')
-                  (`(interpret_symbol config language const_map symb))
-                )
+              let
+                val typ_arity = type_arity_of_symbol thy config symb
+                val (tptp_type_args, tptp_term_args) = chop typ_arity tptp_ts
+                val tyargs =
+                  map (interpret_type config thy type_map o termtype_to_type)
+                    tptp_type_args
+              in
+                (*apply symb to tptp_ts*)
+                if is_prod_typed thy' config symb then
+                  let
+                    val (t, thy'') =
+                      mtimes'
+                        (fold_map (interpret_term false config language const_map
+                         var_types type_map) (tl tptp_term_args) thy')
+                        (interpret_term false config language const_map
+                         var_types type_map (hd tptp_term_args))
+                  in (interpret_symbol config const_map symb tyargs thy' $ t, thy'')
+                  end
+                else
+                  (
+                  mapply'
+                    (fold_map (interpret_term false config language const_map
+                     var_types type_map) tptp_term_args thy')
+                    (`(interpret_symbol config const_map symb tyargs))
+                  )
+              end
        end
   | Term_Var n =>
      (if language = THF orelse language = TFF then
@@ -543,14 +564,12 @@
   | Term_Num (number_kind, num) =>
       let
         (*FIXME hack*)
-        val ind_type = interpret_type config thy type_map (Defined_type Type_Ind)
-        val prefix = case number_kind of
-            Int_num => "intn_"
-          | Real_num => "realn_"
-          | Rat_num => "ratn_"
-      (*FIXME hack -- for Int_num should be
-       HOLogic.mk_number @{typ "int"} (the (Int.fromString num))*)
-      in declare_const_ifnot config (prefix ^ num) ind_type thy end
+        val tptp_type = case number_kind of
+            Int_num => Type_Int
+          | Real_num => Type_Real
+          | Rat_num => Type_Rat
+        val T = interpret_type config thy type_map (Defined_type tptp_type)
+      in (HOLogic.mk_number T (the (Int.fromString num)), thy) end
   | Term_Distinct_Object str =>
       declare_constant config ("do_" ^ str)
         (interpret_type config thy type_map (Defined_type Type_Ind)) thy
@@ -577,11 +596,9 @@
                  (Atom_Arity (const_name, length tptp_formulas)) thy'
             in
               (if is_prod_typed thy' config symbol then
-                 mtimes thy' args (interpret_symbol config language
-                  const_map symbol thy')
+                 mtimes thy' args (interpret_symbol config const_map symbol [] thy')
                else
-                mapply args
-                 (interpret_symbol config language const_map symbol thy'),
+                mapply args (interpret_symbol config const_map symbol [] thy'),
               thy')
             end
         | _ =>
@@ -593,11 +610,9 @@
                  tptp_formulas thy
             in
               (if is_prod_typed thy' config symbol then
-                 mtimes thy' args (interpret_symbol config language
-                  const_map symbol thy')
+                 mtimes thy' args (interpret_symbol config const_map symbol [] thy')
                else
-                 mapply args
-                  (interpret_symbol config language const_map symbol thy'),
+                 mapply args (interpret_symbol config const_map symbol [] thy'),
                thy')
             end)
     | Sequent _ =>
@@ -669,12 +684,12 @@
         (case tptp_atom of
           TFF_Typed_Atom (symbol, tptp_type_opt) =>
             (*FIXME ignoring type info*)
-            (interpret_symbol config language const_map symbol thy, thy)
+            (interpret_symbol config const_map symbol [] thy, thy)
         | THF_Atom_term tptp_term =>
             interpret_term true config language const_map var_types
              type_map tptp_term thy
         | THF_Atom_conn_term symbol =>
-            (interpret_symbol config language const_map symbol thy, thy))
+            (interpret_symbol config const_map symbol [] thy, thy))
     | Type_fmla _ =>
          raise MISINTERPRET_FORMULA
           ("Cannot interpret types as formulas", tptp_fmla)
@@ -684,20 +699,31 @@
 
 (*Extract the type from a typing*)
 local
+  fun type_vars_of_fmlatype (Quant (Dep_Prod, varlist, fmla)) =
+      map fst varlist @ type_vars_of_fmlatype fmla
+    | type_vars_of_fmlatype _ = []
+
   fun extract_type tptp_type =
     case tptp_type of
-        Fmla_type typ => fmlatype_to_type typ
-      | _ => tptp_type
+        Fmla_type fmla => (type_vars_of_fmlatype fmla, fmlatype_to_type fmla)
+      | _ => ([], tptp_type)
 in
   fun typeof_typing (THF_typing (_, tptp_type)) = extract_type tptp_type
   fun typeof_tff_typing (Atom (TFF_Typed_Atom (_, SOME tptp_type))) =
     extract_type tptp_type
 end
+
 fun nameof_typing
   (THF_typing
      ((Atom (THF_Atom_term (Term_Func (Uninterpreted str, [])))), _)) = str
 fun nameof_tff_typing (Atom (TFF_Typed_Atom (Uninterpreted str, _))) = str
 
+fun strip_prod_type (Prod_type (ty1, ty2)) = ty1 :: strip_prod_type ty2
+  | strip_prod_type ty = [ty]
+
+fun dest_fn_type (Fn_type (ty1, ty2)) = (strip_prod_type ty1, ty2)
+  | dest_fn_type ty = ([], ty)
+
 fun resolve_include_path path_prefixes path_suffix =
   case find_first (fn prefix => File.exists (Path.append prefix path_suffix))
          path_prefixes of
@@ -705,6 +731,15 @@
   | NONE =>
     error ("Cannot find include file " ^ quote (Path.implode path_suffix))
 
+(* Ideally, to be in sync with TFF1, we should perform full type skolemization here.
+   But the problems originating from HOL systems are restricted to top-level
+   universal quantification for types. *)
+fun remove_leading_type_quantifiers (Quant (Forall, varlist, tptp_fmla)) =
+    (case filter_out (curry (op =) (SOME (Defined_type Type_Type)) o snd) varlist of
+      [] => remove_leading_type_quantifiers tptp_fmla
+    | varlist' => Quant (Forall, varlist', tptp_fmla))
+  | remove_leading_type_quantifiers tptp_fmla = tptp_fmla
+
 fun interpret_line config inc_list type_map const_map path_prefixes line thy =
   case line of
      Include (_, quoted_path, inc_list) =>
@@ -725,7 +760,7 @@
          case role of
            Role_Type =>
              let
-               val (tptp_ty, name) =
+               val ((tptp_type_vars, tptp_ty), name) =
                  if lang = THF then
                    (typeof_typing tptp_formula (*assuming tptp_formula is a typing*),
                     nameof_typing tptp_formula (*and that the LHS is a (atom) name*))
@@ -733,22 +768,15 @@
                    (typeof_tff_typing tptp_formula,
                     nameof_tff_typing tptp_formula)
              in
-               case tptp_ty of
-                 Defined_type Type_Type => (*add new type*)
+               case dest_fn_type tptp_ty of
+                 (tptp_binders, Defined_type Type_Type) => (*add new type*)
                    (*generate an Isabelle/HOL type to interpret this TPTP type,
                    and add it to both the Isabelle/HOL theory and to the type_map*)
                     let
                       val (type_map', thy') =
-                        declare_type
-                         config
-                         (Atom_type name, name)
-                         type_map
-                         thy
-                    in ((type_map',
-                         const_map,
-                         []),
-                        thy')
-                    end
+                        declare_type config
+                          (name, (name, length tptp_binders)) type_map thy
+                    in ((type_map', const_map, []), thy') end
 
                | _ => (*declaration of constant*)
                   (*Here we populate the map from constants to the Isabelle/HOL
@@ -758,7 +786,7 @@
                     (*make sure that the theory contains all the types appearing
                     in this constant's signature. Exception is thrown if encounter
                     an undeclared types.*)
-                    val (t, thy') =
+                    val (t as Const (name', _), thy') =
                       let
                         fun analyse_type thy thf_ty =
                            if #cautious config then
@@ -766,13 +794,13 @@
                                Fn_type (thf_ty1, thf_ty2) =>
                                  (analyse_type thy thf_ty1;
                                  analyse_type thy thf_ty2)
-                             | Atom_type ty_name =>
+                             | Atom_type (ty_name, _) =>
                                  if type_exists config thy ty_name then ()
                                  else
                                   raise MISINTERPRET_TYPE
                                    ("Type (in signature of " ^
                                       name ^ ") has not been declared",
-                                     Atom_type ty_name)
+                                     Atom_type (ty_name, []))
                              | _ => ()
                            else () (*skip test if we're not being cautious.*)
                       in
@@ -784,7 +812,19 @@
                     use "::". Note that here we use a constant's name rather
                     than the possibly- new one, since all references in the
                     theory will be to this name.*)
-                    val const_map' = ((name, t) :: const_map)
+
+                    val tf_tyargs = map tfree_of_var_type tptp_type_vars
+                    val isa_tyargs = Sign.const_typargs thy' (name', ty)
+                    val _ =
+                      if length isa_tyargs < length tf_tyargs then
+                        raise MISINTERPRET_SYMBOL
+                          ("Cannot handle phantom types for constant",
+                           Uninterpreted name)
+                      else
+                        ()
+                    val tyarg_perm =
+                      map (fn T => find_index (curry (op =) T) tf_tyargs) isa_tyargs
+                    val const_map' = ((name, (t, tyarg_perm)) :: const_map)
                   in ((type_map,(*type_map unchanged, since a constant's
                                   declaration shouldn't include any new types.*)
                        const_map',(*typing table of constant was extended*)
@@ -798,7 +838,7 @@
                (*gather interpreted term, and the map of types occurring in that term*)
                val (t, thy') =
                  interpret_formula config lang
-                  const_map [] type_map tptp_formula thy
+                  const_map [] type_map (remove_leading_type_quantifiers tptp_formula) thy
                  |>> HOLogic.mk_Trueprop
                (*type_maps grow monotonically, so return the newest value (type_mapped);
                there's no need to unify it with the type_map parameter.*)
--- a/src/HOL/TPTP/TPTP_Parser/tptp_lexyacc.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_lexyacc.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -174,9 +174,9 @@
 \\000"
 ),
  (1, 
-"\000\000\000\000\000\000\000\000\000\142\144\000\000\143\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\142\138\133\000\101\089\088\083\082\081\080\078\077\072\070\057\
+"\000\000\000\000\000\000\000\000\000\143\145\000\000\144\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\143\139\134\000\101\089\088\083\082\081\080\078\077\072\070\057\
 \\048\048\048\048\048\048\048\048\048\048\045\000\039\037\036\033\
 \\030\029\029\029\029\029\029\029\029\029\029\029\029\029\029\029\
 \\029\029\029\029\029\029\029\029\029\029\029\028\000\027\026\000\
@@ -847,10 +847,10 @@
  (101, 
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\131\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\132\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\131\
 \\000\102\102\128\102\102\124\102\102\118\102\102\108\102\102\102\
 \\102\102\102\102\103\102\102\102\102\102\102\000\000\000\000\000\
 \\000"
@@ -1053,76 +1053,76 @@
 \\102\102\102\102\102\102\102\102\102\102\102\000\000\000\000\000\
 \\000"
 ),
- (131, 
-"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\132\132\132\132\132\132\132\132\132\132\132\132\132\132\132\
-\\132\132\132\132\132\132\132\132\132\132\132\000\000\000\000\000\
-\\000"
-),
  (132, 
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\132\132\132\132\132\132\132\132\132\132\000\000\000\000\000\000\
-\\000\132\132\132\132\132\132\132\132\132\132\132\132\132\132\132\
-\\132\132\132\132\132\132\132\132\132\132\132\000\000\000\000\132\
-\\000\132\132\132\132\132\132\132\132\132\132\132\132\132\132\132\
-\\132\132\132\132\132\132\132\132\132\132\132\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\133\133\133\133\133\133\133\133\133\133\133\133\133\133\133\
+\\133\133\133\133\133\133\133\133\133\133\133\000\000\000\000\000\
 \\000"
 ),
  (133, 
-"\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\000\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\137\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134"
+"\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\133\133\133\133\133\133\133\133\133\133\000\000\000\000\000\000\
+\\000\133\133\133\133\133\133\133\133\133\133\133\133\133\133\133\
+\\133\133\133\133\133\133\133\133\133\133\133\000\000\000\000\133\
+\\000\133\133\133\133\133\133\133\133\133\133\133\133\133\133\133\
+\\133\133\133\133\133\133\133\133\133\133\133\000\000\000\000\000\
+\\000"
 ),
  (134, 
-"\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\136\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\135\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\134\
-\\134"
+"\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\000\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\138\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135"
 ),
  (135, 
+"\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\137\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\136\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\135\
+\\135"
+),
+ (136, 
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\134\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\134\000\000\000\
+\\000\000\135\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\135\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000"
 ),
- (138, 
+ (139, 
 "\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\141\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\140\139\000\
+\\000\142\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\141\140\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000"
 ),
- (142, 
-"\000\000\000\000\000\000\000\000\000\142\000\000\000\000\000\000\
-\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
-\\142\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+ (143, 
+"\000\000\000\000\000\000\000\000\000\143\000\000\000\000\000\000\
+\\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
+\\143\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
@@ -1130,8 +1130,8 @@
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000"
 ),
- (143, 
-"\000\000\000\000\000\000\000\000\000\000\144\000\000\000\000\000\
+ (144, 
+"\000\000\000\000\000\000\000\000\000\000\145\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
 \\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\000\
@@ -1143,12 +1143,12 @@
 ),
 (0, "")]
 fun f x = x 
-val s = map f (rev (tl (rev s))) 
+val s = List.map f (List.rev (tl (List.rev s))) 
 exception LexHackingError 
 fun look ((j,x)::r, i: int) = if i = j then x else look(r, i) 
   | look ([], i) = raise LexHackingError
 fun g {fin=x, trans=i} = {fin=x, trans=look(s,i)} 
-in Vector.fromList(map g 
+in Vector.fromList(List.map g 
 [{fin = [], trans = 0},
 {fin = [(N 2)], trans = 1},
 {fin = [(N 2)], trans = 1},
@@ -1184,7 +1184,7 @@
 {fin = [(N 12)], trans = 0},
 {fin = [(N 78)], trans = 33},
 {fin = [(N 21)], trans = 0},
-{fin = [(N 303)], trans = 0},
+{fin = [(N 306)], trans = 0},
 {fin = [(N 38)], trans = 0},
 {fin = [(N 31)], trans = 37},
 {fin = [(N 48)], trans = 0},
@@ -1193,10 +1193,10 @@
 {fin = [(N 68)], trans = 0},
 {fin = [(N 41)], trans = 42},
 {fin = [(N 45)], trans = 0},
-{fin = [(N 297)], trans = 0},
+{fin = [(N 300)], trans = 0},
 {fin = [(N 27)], trans = 45},
 {fin = [(N 36)], trans = 0},
-{fin = [(N 306)], trans = 0},
+{fin = [(N 309)], trans = 0},
 {fin = [(N 126)], trans = 48},
 {fin = [], trans = 49},
 {fin = [(N 104)], trans = 49},
@@ -1225,11 +1225,11 @@
 {fin = [(N 55)], trans = 0},
 {fin = [(N 123)], trans = 74},
 {fin = [(N 58)], trans = 75},
-{fin = [(N 294)], trans = 0},
+{fin = [(N 297)], trans = 0},
 {fin = [(N 29)], trans = 0},
-{fin = [(N 288)], trans = 78},
+{fin = [(N 291)], trans = 78},
 {fin = [(N 76)], trans = 0},
-{fin = [(N 290)], trans = 0},
+{fin = [(N 293)], trans = 0},
 {fin = [(N 82)], trans = 0},
 {fin = [(N 52)], trans = 0},
 {fin = [], trans = 83},
@@ -1280,19 +1280,20 @@
 {fin = [(N 278)], trans = 128},
 {fin = [(N 278)], trans = 129},
 {fin = [(N 209),(N 278)], trans = 102},
-{fin = [], trans = 131},
-{fin = [(N 286)], trans = 132},
-{fin = [], trans = 133},
+{fin = [(N 281)], trans = 0},
+{fin = [], trans = 132},
+{fin = [(N 289)], trans = 133},
 {fin = [], trans = 134},
 {fin = [], trans = 135},
+{fin = [], trans = 136},
 {fin = [(N 95)], trans = 0},
-{fin = [], trans = 135},
-{fin = [(N 33)], trans = 138},
-{fin = [(N 300)], trans = 0},
+{fin = [], trans = 136},
+{fin = [(N 33)], trans = 139},
+{fin = [(N 303)], trans = 0},
 {fin = [(N 64)], trans = 0},
 {fin = [(N 18)], trans = 0},
-{fin = [(N 2)], trans = 142},
-{fin = [(N 7)], trans = 143},
+{fin = [(N 2)], trans = 143},
+{fin = [(N 7)], trans = 144},
 {fin = [(N 7)], trans = 0}])
 end
 structure StartStates =
@@ -1328,7 +1329,7 @@
 	| action (i,(node::acts)::l) =
 		case node of
 		    Internal.N yyk => 
-			(let fun yymktext() = substring(!yyb,i0,i-i0)
+			(let fun yymktext() = String.substring(!yyb,i0,i-i0)
 			     val yypos = i0+ !yygone
 			open UserDeclarations Internal.StartStates
  in (yybufpos := i; case yyk of 
@@ -1369,15 +1370,16 @@
 | 27 => (col:=yypos-(!eolpos); T.COLON(!linep,!col))
 | 271 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.LOWER_WORD(yytext,!linep,!col) end
 | 278 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOLLAR_WORD(yytext,!linep,!col) end
-| 286 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOLLAR_DOLLAR_WORD(yytext,!linep,!col) end
-| 288 => (col:=yypos-(!eolpos); T.PLUS(!linep,!col))
+| 281 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOLLAR_WORD(yytext,!linep,!col) end
+| 289 => let val yytext=yymktext() in col:=yypos-(!eolpos); T.DOLLAR_DOLLAR_WORD(yytext,!linep,!col) end
 | 29 => (col:=yypos-(!eolpos); T.COMMA(!linep,!col))
-| 290 => (col:=yypos-(!eolpos); T.TIMES(!linep,!col))
-| 294 => (col:=yypos-(!eolpos); T.GENTZEN_ARROW(!linep,!col))
-| 297 => (col:=yypos-(!eolpos); T.SUBTYPE(!linep,!col))
-| 300 => (col:=yypos-(!eolpos); T.DEP_PROD(!linep,!col))
-| 303 => (col:=yypos-(!eolpos); T.DEP_SUM(!linep,!col))
-| 306 => (col:=yypos-(!eolpos); T.LET_TERM(!linep,!col))
+| 291 => (col:=yypos-(!eolpos); T.PLUS(!linep,!col))
+| 293 => (col:=yypos-(!eolpos); T.TIMES(!linep,!col))
+| 297 => (col:=yypos-(!eolpos); T.GENTZEN_ARROW(!linep,!col))
+| 300 => (col:=yypos-(!eolpos); T.SUBTYPE(!linep,!col))
+| 303 => (col:=yypos-(!eolpos); T.DEP_PROD(!linep,!col))
+| 306 => (col:=yypos-(!eolpos); T.DEP_SUM(!linep,!col))
+| 309 => (col:=yypos-(!eolpos); T.LET_TERM(!linep,!col))
 | 31 => (col:=yypos-(!eolpos); T.EQUALS(!linep,!col))
 | 33 => (col:=yypos-(!eolpos); T.EXCLAMATION(!linep,!col))
 | 36 => (col:=yypos-(!eolpos); T.LET(!linep,!col))
@@ -1414,14 +1416,14 @@
 	     if trans = #trans(Vector.sub(Internal.tab,0))
 	       then action(l,NewAcceptingLeaves
 ) else	    let val newchars= if !yydone then "" else yyinput 1024
-	    in if (size newchars)=0
+	    in if (String.size newchars)=0
 		  then (yydone := true;
 		        if (l=i0) then UserDeclarations.eof yyarg
 		                  else action(l,NewAcceptingLeaves))
 		  else (if i0=l then yyb := newchars
-		     else yyb := substring(!yyb,i0,l-i0)^newchars;
+		     else yyb := String.substring(!yyb,i0,l-i0)^newchars;
 		     yygone := !yygone+i0;
-		     yybl := size (!yyb);
+		     yybl := String.size (!yyb);
 		     scan (s,AcceptingLeaves,l-i0,0))
 	    end
 	  else let val NewChar = Char.ord(String.sub(!yyb,l))
@@ -1432,7 +1434,7 @@
 	end
 	end
 (*
-	val start= if substring(!yyb,!yybufpos-1,1)="\n"
+	val start= if String.substring(!yyb,!yybufpos-1,1)="\n"
 then !yybegin+1 else !yybegin
 *)
 	in scan(!yybegin (* start *),nil,!yybufpos,!yybufpos)
@@ -3604,7 +3606,7 @@
        fun f i =
             if i=numstates then g i
             else (Array.update(memo,i,SHIFT (STATE i)); f (i+1))
-          in f 0 handle Subscript => ()
+          in f 0 handle General.Subscript => ()
           end
 in
 val entry_to_action = fn 0 => ACCEPT | 1 => ERROR | j => Array.sub(memo,(j-2))
@@ -3614,7 +3616,7 @@
 val actionRowNumbers = string_to_list actionRowNumbers
 val actionT = let val actionRowLookUp=
 let val a=Array.fromList(actionRows) in fn i=>Array.sub(a,i) end
-in Array.fromList(map actionRowLookUp actionRowNumbers)
+in Array.fromList(List.map actionRowLookUp actionRowNumbers)
 end
 in LrTable.mkLrTable {actions=actionT,gotos=gotoT,numRules=numrules,
 numStates=numstates,initialState=STATE 0}
@@ -4851,7 +4853,7 @@
 end
 |  ( 135, ( ( _, ( MlyValue.atomic_word atomic_word, atomic_word1left,
  atomic_word1right)) :: rest671)) => let val  result = 
-MlyValue.tff_atomic_type (( Atom_type atomic_word ))
+MlyValue.tff_atomic_type (( Atom_type (atomic_word, []) ))
  in ( LrTable.NT 79, ( result, atomic_word1left, atomic_word1right), 
 rest671)
 end
@@ -5316,6 +5318,7 @@
   | "$real" => Type_Real
   | "$rat" => Type_Rat
   | "$int" => Type_Int
+  | "$_" => Type_Dummy
   | thing => raise UNRECOGNISED_SYMBOL ("defined_type", thing)
 )
 )
@@ -5590,6 +5593,7 @@
   | "$real" => TypeSymbol Type_Real
   | "$rat" => TypeSymbol Type_Rat
   | "$tType" => TypeSymbol Type_Type
+  | "$_" => TypeSymbol Type_Dummy
 
   | "$true" => Interpreted_Logic True
   | "$false" => Interpreted_Logic False
--- a/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Parser/tptp_syntax.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -66,7 +66,7 @@
     Forall | Exists | Epsilon | Iota | Lambda | Dep_Prod | Dep_Sum
 
   and tptp_base_type =
-    Type_Ind | Type_Bool | Type_Type | Type_Int | Type_Rat | Type_Real
+    Type_Ind | Type_Bool | Type_Type | Type_Int | Type_Rat | Type_Real | Type_Dummy
 
   and symbol =
       Uninterpreted of string
@@ -111,7 +111,8 @@
   and tptp_type =
       Prod_type of tptp_type * tptp_type
     | Fn_type of tptp_type * tptp_type
-    | Atom_type of string
+    | Atom_type of string * tptp_type list
+    | Var_type of string
     | Defined_type of tptp_base_type
     | Sum_type of tptp_type * tptp_type (*only THF*)
     | Fmla_type of tptp_formula
@@ -141,8 +142,6 @@
 
   val status_to_string : status_value -> string
 
-  val nameof_tff_atom_type : tptp_type -> string
-
   val pos_of_line : tptp_line -> position
 
   (*Returns the list of all files included in a directory and its
@@ -216,7 +215,7 @@
     Forall | Exists | Epsilon | Iota | Lambda | Dep_Prod | Dep_Sum
 
   and tptp_base_type =
-    Type_Ind | Type_Bool | Type_Type | Type_Int | Type_Rat | Type_Real
+    Type_Ind | Type_Bool | Type_Type | Type_Int | Type_Rat | Type_Real | Type_Dummy
 
   and symbol =
       Uninterpreted of string
@@ -261,7 +260,8 @@
   and tptp_type =
       Prod_type of tptp_type * tptp_type
     | Fn_type of tptp_type * tptp_type
-    | Atom_type of string
+    | Atom_type of string * tptp_type list
+    | Var_type of string
     | Defined_type of tptp_base_type
     | Sum_type of tptp_type * tptp_type
     | Fmla_type of tptp_formula
@@ -287,9 +287,6 @@
 
 fun debug f x = if Options.default_bool @{system_option ML_exception_trace} then (f x; ()) else ()
 
-fun nameof_tff_atom_type (Atom_type str) = str
-  | nameof_tff_atom_type _ = raise TPTP_SYNTAX "nameof_tff_atom_type called on non-atom type"
-
 fun pos_of_line tptp_line =
   case tptp_line of
       Annotated_Formula (position, _, _, _, _, _) => position
@@ -428,6 +425,7 @@
   | string_of_tptp_base_type Type_Int = "$int"
   | string_of_tptp_base_type Type_Rat = "$rat"
   | string_of_tptp_base_type Type_Real = "$real"
+  | string_of_tptp_base_type Type_Dummy = "$_"
 
 and string_of_interpreted_symbol x =
   case x of
@@ -517,7 +515,10 @@
       string_of_tptp_type tptp_type1 ^ " * " ^ string_of_tptp_type tptp_type2
   | string_of_tptp_type (Fn_type (tptp_type1, tptp_type2)) =
       string_of_tptp_type tptp_type1 ^ " > " ^ string_of_tptp_type tptp_type2
-  | string_of_tptp_type (Atom_type str) = str
+  | string_of_tptp_type (Atom_type (str, [])) = str
+  | string_of_tptp_type (Atom_type (str, tptp_types)) =
+      str ^ "(" ^ commas (map string_of_tptp_type tptp_types) ^ ")"
+  | string_of_tptp_type (Var_type str) = str
   | string_of_tptp_type (Defined_type tptp_base_type) =
       string_of_tptp_base_type tptp_base_type
   | string_of_tptp_type (Sum_type (tptp_type1, tptp_type2)) = ""
@@ -565,6 +566,7 @@
   | latex_of_tptp_base_type Type_Int = "\\\\mathsf{int} "
   | latex_of_tptp_base_type Type_Rat = "\\\\mathsf{rat} "
   | latex_of_tptp_base_type Type_Real = "\\\\mathsf{real} "
+  | latex_of_tptp_base_type Type_Dummy = "\\\\mathsf{\\\\_} "
 
 and latex_of_interpreted_symbol x =
   case x of
@@ -687,7 +689,10 @@
       latex_of_tptp_type tptp_type1 ^ " \\\\times " ^ latex_of_tptp_type tptp_type2
   | latex_of_tptp_type (Fn_type (tptp_type1, tptp_type2)) =
       latex_of_tptp_type tptp_type1 ^ " \\\\to " ^ latex_of_tptp_type tptp_type2
-  | latex_of_tptp_type (Atom_type str) = "\\\\mathrm{" ^ str ^ "}"
+  | latex_of_tptp_type (Atom_type (str, [])) = "\\\\mathrm{" ^ str ^ "}"
+  | latex_of_tptp_type (Atom_type (str, tptp_types)) =
+    "\\\\mathrm{" ^ str ^ "}(" ^ commas (map latex_of_tptp_type tptp_types) ^ ")"
+  | latex_of_tptp_type (Var_type str) = "\\\\mathrm{" ^ str ^ "}"
   | latex_of_tptp_type (Defined_type tptp_base_type) =
       latex_of_tptp_base_type tptp_base_type
   | latex_of_tptp_type (Sum_type (tptp_type1, tptp_type2)) = ""
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -2116,7 +2116,10 @@
   in
     case inference_name of
       "fo_atp_e" =>
+        HEADGOAL (Metis_Tactic.metis_tac [] ATP_Problem_Generate.combs_or_liftingN ctxt [])
+        (*NOTE To treat E as an oracle use the following line:
         HEADGOAL (etac (oracle_based_reconstruction_tac ctxt prob_name node))
+        *)
     | "copy" =>
          HEADGOAL
           (atac
@@ -2284,4 +2287,4 @@
   end
 *}
 
-end
\ No newline at end of file
+end
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -249,7 +249,7 @@
 (*given a list of tactics to be applied in sequence (i.e., they
   follow a skeleton), we build a single tactic, interleaving
   some tracing info to help with debugging.*)
-fun step_by_step_tacs verbose (thm_tacs : (thm * tactic) list) : tactic =
+fun step_by_step_tacs ctxt verbose (thm_tacs : (thm * tactic) list) : tactic =
     let
       fun interleave_tacs [] [] = all_tac
         | interleave_tacs (tac1 :: tacs1) (tac2 :: tacs2) =
@@ -258,7 +258,7 @@
       val thms_to_traceprint =
         map (fn thm => fn st =>
               (*FIXME uses makestring*)
-              print_tac (PolyML.makestring thm) st)
+              print_tac ctxt (PolyML.makestring thm) st)
 
     in
       if verbose then
@@ -272,9 +272,9 @@
 
 ML {*
 (*apply step_by_step_tacs to all problems under test*)
-val narrated_tactics =
+fun narrated_tactics ctxt =
  map (map (#3 #> the)
-      #> step_by_step_tacs false)
+      #> step_by_step_tacs ctxt false)
    the_tactics;
 
 (*produce thm*)
@@ -284,7 +284,7 @@
   map (fn (prob_name, tac) =>
          TPTP_Reconstruct.reconstruct @{context}
            (fn _ => tac) prob_name)
-    (ListPair.zip (prob_names, narrated_tactics))
+    (ListPair.zip (prob_names, narrated_tactics @{context}))
 *}
 
 
--- a/src/HOL/TPTP/atp_problem_import.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/atp_problem_import.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -13,7 +13,8 @@
   val refute_tptp_file : theory -> int -> string -> unit
   val can_tac : Proof.context -> tactic -> term -> bool
   val SOLVE_TIMEOUT :  int -> string -> tactic -> tactic
-  val atp_tac : Proof.context -> int -> (string * string) list -> int -> string -> int -> tactic
+  val atp_tac : Proof.context -> int -> (string * string) list -> int -> term list -> string ->
+    int -> tactic
   val smt_solver_tac : string -> Proof.context -> int -> tactic
   val freeze_problem_consts : theory -> term -> term
   val make_conj : term list * term list -> term list -> term
@@ -45,20 +46,18 @@
 fun read_tptp_file thy postproc file_name =
   let
     fun has_role role (_, role', _, _) = (role' = role)
-    fun get_prop (name, role, P, info) =
-      let val P' = P |> Logic.varify_global |> close_form in
-        (name, P') |> postproc
-      end
+    fun get_prop f (name, _, P, _) = P |> f |> close_form |> pair name |> postproc
+
     val path = exploded_absolute_path file_name
     val ((_, _, problem), thy) =
-      TPTP_Interpret.interpret_file true [Path.dir path, Path.explode "$TPTP"]
-                                    path [] [] thy
-    val (conjs, defs_and_nondefs) =
-      problem |> List.partition (has_role TPTP_Syntax.Role_Conjecture)
-              ||> List.partition (has_role TPTP_Syntax.Role_Definition)
+      TPTP_Interpret.interpret_file true [Path.dir path, Path.explode "$TPTP"] path [] [] thy
+    val (conjs, defs_and_nondefs) = problem
+      |> List.partition (has_role TPTP_Syntax.Role_Conjecture)
+      ||> List.partition (has_role TPTP_Syntax.Role_Definition)
   in
-    (map get_prop conjs, pairself (map get_prop) defs_and_nondefs,
-     thy |> Proof_Context.init_global)
+    (map (get_prop I) conjs,
+     pairself (map (get_prop Logic.varify_global)) defs_and_nondefs,
+     Named_Target.theory_init thy)
   end
 
 
@@ -76,12 +75,10 @@
   let
     val (conjs, (defs, nondefs), ctxt) = read_tptp_file thy snd file_name
     val thy = Proof_Context.theory_of ctxt
-    val (defs, pseudo_defs) =
-      defs |> map (ATP_Util.abs_extensionalize_term ctxt
-                   #> aptrueprop (hol_open_form I))
-           |> List.partition (is_legitimate_tptp_def
-                              o perhaps (try HOLogic.dest_Trueprop)
-                              o ATP_Util.unextensionalize_def)
+    val (defs, pseudo_defs) = defs
+      |> map (ATP_Util.abs_extensionalize_term ctxt #> aptrueprop (hol_open_form I))
+      |> List.partition (is_legitimate_tptp_def o perhaps (try HOLogic.dest_Trueprop)
+        o ATP_Util.unextensionalize_def)
     val nondefs = pseudo_defs @ nondefs
     val state = Proof.init ctxt
     val params =
@@ -104,8 +101,8 @@
     val step = 0
     val subst = []
   in
-    Nitpick.pick_nits_in_term state params Nitpick.TPTP i n step subst
-        defs nondefs (case conjs of conj :: _ => conj | [] => @{prop True});
+    Nitpick.pick_nits_in_term state params Nitpick.TPTP i n step subst defs nondefs
+      (case conjs of conj :: _ => conj | [] => @{prop True});
     ()
   end
 
@@ -127,7 +124,7 @@
        ("maxvars", "100000")]
   in
     Refute.refute_term ctxt params (defs @ nondefs)
-        (case conjs of conj :: _ => conj | [] => @{prop True})
+      (case conjs of conj :: _ => conj | [] => @{prop True})
     |> print_szs_of_outcome (not (null conjs))
   end
 
@@ -138,17 +135,16 @@
 
 fun SOLVE_TIMEOUT seconds name tac st =
   let
-    val _ = Output.urgent_message ("running " ^ name ^ " for " ^
-                                   string_of_int seconds ^ " s")
+    val _ = Output.urgent_message ("running " ^ name ^ " for " ^ string_of_int seconds ^ " s")
     val result =
-      TimeLimit.timeLimit (Time.fromSeconds seconds)
-        (fn () => SINGLE (SOLVE tac) st) ()
-      handle TimeLimit.TimeOut => NONE
-        | ERROR _ => NONE
+      TimeLimit.timeLimit (Time.fromSeconds seconds) (fn () => SINGLE (SOLVE tac) st) ()
+      handle
+        TimeLimit.TimeOut => NONE
+      | ERROR _ => NONE
   in
-    case result of
+    (case result of
       NONE => (Output.urgent_message ("FAILURE: " ^ name); Seq.empty)
-    | SOME st' => (Output.urgent_message ("SUCCESS: " ^ name); Seq.single st')
+    | SOME st' => (Output.urgent_message ("SUCCESS: " ^ name); Seq.single st'))
   end
 
 fun nitpick_finite_oracle_tac ctxt timeout i th =
@@ -157,6 +153,7 @@
       | is_safe @{typ prop} = true
       | is_safe @{typ bool} = true
       | is_safe _ = false
+
     val conj = Thm.term_of (Thm.cprem_of th i)
   in
     if exists_type (not o is_safe) conj then
@@ -179,40 +176,43 @@
         val step = 0
         val subst = []
         val (outcome, _) =
-          Nitpick.pick_nits_in_term state params Nitpick.Normal i n step subst
-                                    [] [] conj
-      in if outcome = "none" then ALLGOALS Skip_Proof.cheat_tac th else Seq.empty end
+          Nitpick.pick_nits_in_term state params Nitpick.Normal i n step subst [] [] conj
+      in
+        if outcome = "none" then ALLGOALS Skip_Proof.cheat_tac th else Seq.empty
+      end
   end
 
-fun atp_tac ctxt completeness override_params timeout prover =
+fun atp_tac ctxt completeness override_params timeout assms prover =
   let
-    val ctxt =
-      ctxt |> Config.put Sledgehammer_Prover_ATP.atp_completish (completeness > 0)
+    val thy = Proof_Context.theory_of ctxt
+    val assm_ths0 = map (Skip_Proof.make_thm thy) assms
+    val ((assm_name, _), ctxt) = ctxt
+      |> Config.put Sledgehammer_Prover_ATP.atp_completish (completeness > 0)
+      |> Local_Theory.note ((@{binding thms}, []), assm_ths0)
+
     fun ref_of th = (Facts.named (Thm.derivation_name th), [])
+    val ref_of_assms = (Facts.named assm_name, [])
   in
     Sledgehammer_Tactics.sledgehammer_as_oracle_tac ctxt
-        ([("debug", if debug then "true" else "false"),
-          ("overlord", if overlord then "true" else "false"),
-          ("provers", prover),
-          ("timeout", string_of_int timeout)] @
-         (if completeness > 0 then
-            [("type_enc",
-              if completeness = 1 then "mono_native" else "poly_tags")]
-          else
-            []) @
-         override_params)
-        {add = map ref_of [ext, @{thm someI}], del = [], only = true}
+      ([("debug", if debug then "true" else "false"),
+        ("overlord", if overlord then "true" else "false"),
+        ("provers", prover),
+        ("timeout", string_of_int timeout)] @
+       (if completeness > 0 then
+          [("type_enc", if completeness = 1 then "mono_native" else "poly_tags")]
+        else
+          []) @
+       override_params)
+      {add = ref_of_assms :: map ref_of [ext, @{thm someI}], del = [], only = true} []
   end
 
-fun sledgehammer_tac demo ctxt timeout i =
+fun sledgehammer_tac demo ctxt timeout assms i =
   let
     val frac = if demo then 16 else 12
     fun slice mult completeness prover =
       SOLVE_TIMEOUT (mult * timeout div frac)
-          (prover ^
-           (if completeness > 0 then "(" ^ string_of_int completeness ^ ")"
-            else ""))
-          (atp_tac ctxt completeness [] (mult * timeout div frac) prover i)
+        (prover ^ (if completeness > 0 then "(" ^ string_of_int completeness ^ ")" else ""))
+        (atp_tac ctxt completeness [] (mult * timeout div frac) assms prover i)
   in
     slice 2 0 ATP_Proof.spassN
     ORELSE slice 2 0 ATP_Proof.vampireN
@@ -235,15 +235,12 @@
     val ctxt = ctxt |> Context.proof_map (SMT_Config.select_solver solver)
   in SMT_Solver.smt_tac ctxt [] end
 
-fun auto_etc_tac ctxt timeout i =
-  SOLVE_TIMEOUT (timeout div 20) "nitpick"
-      (nitpick_finite_oracle_tac ctxt (timeout div 20) i)
-  ORELSE SOLVE_TIMEOUT (timeout div 10) "simp"
-      (asm_full_simp_tac ctxt i)
+fun auto_etc_tac ctxt timeout assms i =
+  SOLVE_TIMEOUT (timeout div 20) "nitpick" (nitpick_finite_oracle_tac ctxt (timeout div 20) i)
+  ORELSE SOLVE_TIMEOUT (timeout div 10) "simp" (asm_full_simp_tac ctxt i)
   ORELSE SOLVE_TIMEOUT (timeout div 10) "blast" (blast_tac ctxt i)
   ORELSE SOLVE_TIMEOUT (timeout div 5) "auto+spass"
-      (auto_tac ctxt
-       THEN ALLGOALS (atp_tac ctxt 0 [] (timeout div 5) ATP_Proof.spassN))
+    (auto_tac ctxt THEN ALLGOALS (atp_tac ctxt 0 [] (timeout div 5) assms ATP_Proof.spassN))
   ORELSE SOLVE_TIMEOUT (timeout div 10) "fast" (fast_tac ctxt i)
   ORELSE SOLVE_TIMEOUT (timeout div 20) "z3" (smt_solver_tac "z3" ctxt i)
   ORELSE SOLVE_TIMEOUT (timeout div 20) "cvc3" (smt_solver_tac "cvc3" ctxt i)
@@ -258,45 +255,44 @@
    unfold "definitions" of free variables than of constants (cf. PUZ107^5). *)
 fun freeze_problem_consts thy =
   let val is_problem_const = String.isPrefix (problem_const_prefix thy) in
-    map_aterms (fn t as Const (s, T) =>
-                   if is_problem_const s then Free (Long_Name.base_name s, T)
-                   else t
-                 | t => t)
+    map_aterms
+      (fn t as Const (s, T) => if is_problem_const s then Free (Long_Name.base_name s, T) else t
+        | t => t)
   end
 
 fun make_conj (defs, nondefs) conjs =
-  Logic.list_implies (rev defs @ rev nondefs,
-                      case conjs of conj :: _ => conj | [] => @{prop False})
+  Logic.list_implies (rev defs @ rev nondefs, case conjs of conj :: _ => conj | [] => @{prop False})
 
 fun print_szs_of_success conjs success =
   Output.urgent_message ("% SZS status " ^
-                         (if success then
-                            if null conjs then "Unsatisfiable" else "Theorem"
-                          else
-                            "Unknown"))
+    (if success then
+       if null conjs then "Unsatisfiable" else "Theorem"
+     else
+       "Unknown"))
 
 fun sledgehammer_tptp_file thy timeout file_name =
   let
-    val (conjs, assms, ctxt) =
-      read_tptp_file thy (freeze_problem_consts thy o snd) file_name
-    val conj = make_conj assms conjs
+    val (conjs, assms, ctxt) = read_tptp_file thy (freeze_problem_consts thy o snd) file_name
+    val conj = make_conj ([], []) conjs
+    val assms = op @ assms
   in
-    can_tac ctxt (sledgehammer_tac true ctxt timeout 1) conj
+    can_tac ctxt (sledgehammer_tac true ctxt timeout assms 1) conj
     |> print_szs_of_success conjs
   end
 
 fun generic_isabelle_tptp_file demo thy timeout file_name =
   let
-    val (conjs, assms, ctxt) =
-      read_tptp_file thy (freeze_problem_consts thy o snd) file_name
-    val conj = make_conj assms conjs
+    val (conjs, assms, ctxt) = read_tptp_file thy (freeze_problem_consts thy o snd) file_name
+    val conj = make_conj ([], []) conjs
+    val full_conj = make_conj assms conjs
+    val assms = op @ assms
     val (last_hope_atp, last_hope_completeness) =
       if demo then (ATP_Proof.satallaxN, 0) else (ATP_Proof.vampireN, 2)
   in
-    (can_tac ctxt (auto_etc_tac ctxt (timeout div 2) 1) conj orelse
-     can_tac ctxt (sledgehammer_tac demo ctxt (timeout div 2) 1) conj orelse
+    (can_tac ctxt (auto_etc_tac ctxt (timeout div 2) assms 1) full_conj orelse
+     can_tac ctxt (sledgehammer_tac demo ctxt (timeout div 2) assms 1) conj orelse
      can_tac ctxt (SOLVE_TIMEOUT timeout (last_hope_atp ^ "(*)")
-         (atp_tac ctxt last_hope_completeness [] timeout last_hope_atp 1)) conj)
+       (atp_tac ctxt last_hope_completeness [] timeout assms last_hope_atp 1)) full_conj)
     |> print_szs_of_success conjs
   end
 
@@ -322,7 +318,8 @@
     val uncurried_aliases = false
     val readable_names = true
     val presimp = true
-    val facts = map (apfst (rpair (Global, Non_Rec_Def))) defs @
+    val facts =
+      map (apfst (rpair (Global, Non_Rec_Def))) defs @
       map (apfst (rpair (Global, General))) nondefs
     val (atp_problem, _, _, _) =
       generate_atp_problem ctxt format Hypothesis (type_enc_of_string Strict type_enc) Translator
--- a/src/HOL/TPTP/sledgehammer_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/TPTP/sledgehammer_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -9,10 +9,10 @@
 sig
   type fact_override = Sledgehammer_Fact.fact_override
 
-  val sledgehammer_with_metis_tac :
-    Proof.context -> (string * string) list -> fact_override -> int -> tactic
-  val sledgehammer_as_oracle_tac :
-    Proof.context -> (string * string) list -> fact_override -> int -> tactic
+  val sledgehammer_with_metis_tac : Proof.context -> (string * string) list -> fact_override ->
+    thm list -> int -> tactic
+  val sledgehammer_as_oracle_tac : Proof.context -> (string * string) list -> fact_override ->
+    thm list -> int -> tactic
 end;
 
 structure Sledgehammer_Tactics : SLEDGEHAMMER_TACTICS =
@@ -26,7 +26,7 @@
 open Sledgehammer_MaSh
 open Sledgehammer_Commands
 
-fun run_prover override_params fact_override i n ctxt goal =
+fun run_prover override_params fact_override chained i n ctxt goal =
   let
     val thy = Proof_Context.theory_of ctxt
     val mode = Normal
@@ -39,37 +39,38 @@
     val reserved = reserved_isar_keyword_table ()
     val css_table = clasimpset_rule_table_of ctxt
     val facts =
-      nearly_all_facts ctxt ho_atp fact_override reserved css_table [] hyp_ts concl_t
-      |> relevant_facts ctxt params name
-             (the_default default_max_facts max_facts) fact_override hyp_ts
-             concl_t
+      nearly_all_facts ctxt ho_atp fact_override reserved css_table chained hyp_ts concl_t
+      |> relevant_facts ctxt params name (the_default default_max_facts max_facts) fact_override
+        hyp_ts concl_t
       |> hd |> snd
     val problem =
       {comment = "", state = Proof.init ctxt, goal = goal, subgoal = i, subgoal_count = n,
        factss = [("", facts)]}
   in
-    (case prover params (K (K (K ""))) problem of
+    (case prover params problem of
       {outcome = NONE, used_facts, ...} => used_facts |> map fst |> SOME
     | _ => NONE)
     handle ERROR message => (warning ("Error: " ^ message ^ "\n"); NONE)
   end
 
-fun sledgehammer_with_metis_tac ctxt override_params fact_override i th =
+fun sledgehammer_with_metis_tac ctxt override_params fact_override chained i th =
   let val override_params = override_params @ [("preplay_timeout", "0")] in
-    case run_prover override_params fact_override i i ctxt th of
+    (case run_prover override_params fact_override chained i i ctxt th of
       SOME facts =>
       Metis_Tactic.metis_tac [] ATP_Problem_Generate.combs_or_liftingN ctxt
-          (maps (thms_of_name ctxt) facts) i th
-    | NONE => Seq.empty
+        (maps (thms_of_name ctxt) facts) i th
+    | NONE => Seq.empty)
   end
 
-fun sledgehammer_as_oracle_tac ctxt override_params fact_override i th =
+fun sledgehammer_as_oracle_tac ctxt override_params fact_override chained i th =
   let
     val override_params =
       override_params @
       [("preplay_timeout", "0"),
        ("minimize", "false")]
-    val xs = run_prover override_params fact_override i i ctxt th
-  in if is_some xs then ALLGOALS Skip_Proof.cheat_tac th else Seq.empty end
+    val xs = run_prover override_params fact_override chained i i ctxt th
+  in
+    if is_some xs then ALLGOALS Skip_Proof.cheat_tac th else Seq.empty
+  end
 
 end;
--- a/src/HOL/Tools/ATP/atp_problem.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_problem.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -76,6 +76,8 @@
   val tptp_ho_exists : string
   val tptp_choice : string
   val tptp_ho_choice : string
+  val tptp_hilbert_choice : string
+  val tptp_hilbert_the : string
   val tptp_not : string
   val tptp_and : string
   val tptp_not_and : string
@@ -239,6 +241,8 @@
 val tptp_iff = "<=>"
 val tptp_not_iff = "<~>"
 val tptp_app = "@"
+val tptp_hilbert_choice = "@+"
+val tptp_hilbert_the = "@-"
 val tptp_not_infix = "!"
 val tptp_equal = "="
 val tptp_not_equal = "!="
@@ -533,11 +537,12 @@
     " : " ^ string_of_type format ty ^ ").\n"
   | tptp_string_of_line format (Formula ((ident, alt), kind, phi, source, _)) =
     tptp_string_of_format format ^ "(" ^ ident ^ ", " ^
-    tptp_string_of_role kind ^ "," ^ maybe_alt alt ^
-    "\n    (" ^ tptp_string_of_formula format phi ^ ")" ^
+    tptp_string_of_role kind ^ "," ^ "\n    (" ^
+    tptp_string_of_formula format phi ^ ")" ^
     (case source of
-       SOME tm => ", " ^ tptp_string_of_term format tm
-     | NONE => "") ^ ").\n"
+      SOME tm => ", " ^ tptp_string_of_term format tm
+    | NONE => "") ^
+    ")." ^ maybe_alt alt ^ "\n"
 
 fun tptp_lines format =
   maps (fn (_, []) => []
@@ -638,8 +643,7 @@
     fun formula pred (Formula ((ident, alt), kind, phi, _, info)) =
         if pred kind then
           let val rank = extract_isabelle_rank info in
-            "formula(" ^ dfg_string_of_formula poly gen_simp info phi ^
-            ", " ^ ident ^
+            "formula(" ^ dfg_string_of_formula poly gen_simp info phi ^ ", " ^ ident ^
             (if rank = default_rank then "" else ", " ^ string_of_int rank) ^
             ")." ^ maybe_alt alt
             |> SOME
--- a/src/HOL/Tools/ATP/atp_proof.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_proof.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -2,6 +2,7 @@
     Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
     Author:     Claire Quigley, Cambridge University Computer Laboratory
     Author:     Jasmin Blanchette, TU Muenchen
+    Author:     Mathias Fleury, ENS Rennes
 
 Abstract representation of ATP proofs and TSTP/SPASS syntax.
 *)
@@ -67,7 +68,6 @@
   val remote_prefix : string
 
   val agsyhol_core_rule : string
-  val satallax_core_rule : string
   val spass_input_rule : string
   val spass_pre_skolemize_rule : string
   val spass_skolemize_rule : string
@@ -84,10 +84,25 @@
   val scan_general_id : string list -> string * string list
   val parse_formula : string list ->
     (string, string atp_type, (string, string atp_type) atp_term, string) atp_formula * string list
-  val atp_proof_of_tstplike_proof : string -> string atp_problem -> string -> string atp_proof
   val clean_up_atp_proof_dependencies : string atp_proof -> string atp_proof
   val map_term_names_in_atp_proof : (string -> string) -> string atp_proof -> string atp_proof
   val nasty_atp_proof : string Symtab.table -> string atp_proof -> string atp_proof
+
+  val skip_term: string list -> string * string list
+  val parse_thf0_formula :string list ->
+    ('a, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, 'c) ATP_Problem.atp_formula *
+    string list
+  val dummy_atype : string ATP_Problem.atp_type
+  val role_of_tptp_string: string -> ATP_Problem.atp_formula_role
+  val parse_line: string -> ('a * string ATP_Problem.atp_problem_line list) list ->
+    string list -> ((string * string list) * ATP_Problem.atp_formula_role *
+    (string, 'b, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term,
+      'c) ATP_Problem.atp_formula
+    * string * (string * 'd list) list) list * string list
+  val core_inference : 'a -> 'b -> ('b * 'b list) * ATP_Problem.atp_formula_role *
+    ('c, 'd, (string, 'e) ATP_Problem.atp_term, 'f) ATP_Problem.atp_formula * 'a * 'g list
+  val vampire_step_name_ord : (string * 'a) * (string * 'a) -> order
+  val core_of_agsyhol_proof :  string -> string list option
 end;
 
 structure ATP_Proof : ATP_PROOF =
@@ -122,7 +137,6 @@
 val remote_prefix = "remote_"
 
 val agsyhol_core_rule = "__agsyhol_core" (* arbitrary *)
-val satallax_core_rule = "__satallax_core" (* arbitrary *)
 val spass_input_rule = "Inp"
 val spass_pre_skolemize_rule = "__Sko0" (* arbitrary *)
 val spass_skolemize_rule = "__Sko" (* arbitrary *)
@@ -277,7 +291,7 @@
   let
     fun skip _ accum [] = (accum, [])
       | skip n accum (ss as s :: ss') =
-        if s = "," andalso n = 0 then
+        if (s = "," orelse s = ".") andalso n = 0 then
           (accum, ss)
         else if member (op =) [")", "]"] s then
           if n = 0 then (accum, ss) else skip (n - 1) (s :: accum) ss'
@@ -291,11 +305,12 @@
 
 datatype source =
   File_Source of string * string option |
-  Inference_Source of string * string list
+  Inference_Source of string * string list |
+  Introduced_Source of string
 
 val dummy_phi = AAtom (ATerm (("", []), []))
 val dummy_inference = Inference_Source ("", [])
-val dummy_atype = AType(("", []), [])
+val dummy_atype = AType (("", []), [])
 
 (* "skip_term" is there to cope with Waldmeister nonsense such as "theory(equality)". *)
 fun parse_dependency x =
@@ -312,13 +327,13 @@
   (Scan.this_string "inference" |-- $$ "(" |-- scan_general_id
    --| skip_term --| $$ "," --| skip_term --| $$ "," --| $$ "["
    -- parse_dependencies --| $$ "]" --| $$ ")") x
-and skip_introduced x =
-  (Scan.this_string "introduced" |-- $$ "(" |-- skip_term
-   -- Scan.repeat ($$ "," |-- skip_term) --| $$ ")") x
+and parse_introduced_source x =
+  (Scan.this_string "introduced" |-- $$ "(" |-- scan_general_id
+   --| Scan.option ($$ "," |-- skip_term) --| $$ ")") x
 and parse_source x =
   (parse_file_source >> File_Source
    || parse_inference_source >> Inference_Source
-   || skip_introduced >> K dummy_inference (* for Vampire *)
+   || parse_introduced_source >> Introduced_Source
    || scan_general_id >> (fn s => Inference_Source ("", [s])) (* for E *)
    || skip_term >> K dummy_inference) x
 
@@ -459,22 +474,34 @@
   | role_of_tptp_string "type" = Type_Role
   | role_of_tptp_string _ = Unknown
 
-val tptp_binary_op_list = [tptp_and, tptp_not_and, tptp_or, tptp_not_or, tptp_implies, tptp_if,
-                             tptp_iff, tptp_not_iff, tptp_equal, tptp_app]
+val tptp_binary_ops =
+  [tptp_and, tptp_not_and, tptp_or, tptp_not_or, tptp_implies, tptp_if, tptp_iff, tptp_not_iff,
+   tptp_equal, tptp_app]
+
+fun parse_one_in_list xs =
+  foldl1 (op ||) (map Scan.this_string xs)
 
 fun parse_binary_op x =
-  (foldl1 (op ||) (map Scan.this_string tptp_binary_op_list)
-    >> (fn c => if c = tptp_equal then "equal" else c)) x
+  (parse_one_in_list tptp_binary_ops
+   >> (fn c => if c = tptp_equal then "equal" else c)) x
+
+val parse_fo_quantifier =
+   parse_one_in_list [tptp_forall, tptp_exists, tptp_lambda, tptp_hilbert_choice, tptp_hilbert_the]
+
+val parse_ho_quantifier =
+   parse_one_in_list [tptp_ho_forall, tptp_ho_exists, tptp_hilbert_choice, tptp_hilbert_the]
 
 fun parse_thf0_type x =
   (($$ "(" |-- parse_thf0_type --| $$ ")" || scan_general_id >> (fn t => AType ((t, []), [])))
-     -- Scan.option ($$ tptp_fun_type |-- parse_thf0_type)
+   -- Scan.option ($$ tptp_fun_type |-- parse_thf0_type)
   >> (fn (a, NONE) => a
        | (a, SOME b) => AFun (a, b))) x
 
 fun mk_ho_of_fo_quant q =
   if q = tptp_forall then tptp_ho_forall
   else if q = tptp_exists then tptp_ho_exists
+  else if q = tptp_hilbert_choice then tptp_hilbert_choice
+  else if q = tptp_hilbert_the then tptp_hilbert_the
   else raise Fail ("unrecognized quantification: " ^ q)
 
 fun remove_thf_app (ATerm ((x, ty), arg)) =
@@ -492,8 +519,7 @@
    || $$ "(" |-- parse_thf0_typed_var --| $$ ")") x
 
 fun parse_simple_thf0_term x =
-  ((Scan.this_string tptp_forall || Scan.this_string tptp_exists || Scan.this_string tptp_lambda)
-        -- ($$ "[" |-- parse_thf0_typed_var --| $$ "]" --| $$ ":") -- parse_thf0_term
+  (parse_fo_quantifier -- ($$ "[" |-- parse_thf0_typed_var --| $$ "]" --| $$ ":") -- parse_thf0_term
       >> (fn ((q, ys), t) =>
           fold_rev
             (fn (var, ty) => fn r =>
@@ -505,7 +531,7 @@
             ys t)
   || Scan.this_string tptp_not |-- parse_thf0_term >> mk_app (mk_simple_aterm tptp_not)
   || scan_general_id --| $$ ":" -- parse_thf0_type >> (fn (var, typ) => ATerm ((var, [typ]), []))
-  || (Scan.this_string tptp_ho_forall || Scan.this_string tptp_ho_exists) >> mk_simple_aterm
+  || parse_ho_quantifier >> mk_simple_aterm
   || $$ "(" |-- parse_thf0_term --| $$ ")"
   || scan_general_id >> mk_simple_aterm) x
 and parse_thf0_term x =
@@ -546,11 +572,12 @@
     |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ ","
     -- (parse_formula || skip_term >> K dummy_phi) -- parse_tstp_extra_arguments
     --| $$ ")" --| $$ "."
-   >> (fn (((num, role), phi), deps) =>
+   >> (fn (((num, role0), phi), src) =>
           let
-            val ((name, phi), rule, deps) =
+            val role = role_of_tptp_string role0
+            val ((name, phi), role', rule, deps) =
               (* Waldmeister isn't exactly helping. *)
-              (case deps of
+              (case src of
                 File_Source (_, SOME s) =>
                 (if s = waldmeister_conjecture_name then
                    (case find_formula_in_problem (mk_anot phi) problem of
@@ -563,13 +590,15 @@
                  else
                    ((num, [s |> perhaps (try (unprefix tofof_fact_prefix))]),
                     phi),
-                 "", [])
+                 role, "", [])
               | File_Source _ =>
-                (((num, map fst (find_formula_in_problem phi problem)), phi), "", [])
-              | Inference_Source (rule, deps) => (((num, []), phi), rule, deps))
-            fun mk_step () = (name, role_of_tptp_string role, phi, rule, map (rpair []) deps)
+                (((num, map fst (find_formula_in_problem phi problem)), phi), role, "", [])
+              | Inference_Source (rule, deps) => (((num, []), phi), role, rule, deps)
+              | Introduced_Source rule => (((num, []), phi), Lemma, rule, []))
+
+            fun mk_step () = (name, role', phi, rule, map (rpair []) deps)
           in
-            [(case role_of_tptp_string role of
+            [(case role' of
                Definition =>
                (case phi of
                  AAtom (ATerm (("equal", _), _)) =>
@@ -581,15 +610,13 @@
 
 (**** PARSING OF SPASS OUTPUT ****)
 
-(* SPASS returns clause references of the form "x.y". We ignore "y", whose role
-   is not clear anyway. *)
+(* SPASS returns clause references of the form "x.y". We ignore "y". *)
 val parse_dot_name = scan_general_id --| $$ "." --| scan_general_id
 
 val parse_spass_annotations =
   Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name --| Scan.option ($$ ","))) []
 
-(* It is not clear why some literals are followed by sequences of stars and/or
-   pluses. We ignore them. *)
+(* We ignore the stars and the pluses that follow literals. *)
 fun parse_decorated_atom x =
   (parse_atom --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")) x
 
@@ -598,7 +625,7 @@
   | mk_horn (neg_lits, []) = mk_anot (foldr1 (uncurry (mk_aconn AAnd)) neg_lits)
   | mk_horn (neg_lits, pos_lits) =
     mk_aconn AImplies (foldr1 (uncurry (mk_aconn AAnd)) neg_lits)
-                      (foldr1 (uncurry (mk_aconn AOr)) pos_lits)
+      (foldr1 (uncurry (mk_aconn AOr)) pos_lits)
 
 fun parse_horn_clause x =
   (Scan.repeat parse_decorated_atom --| $$ "|" --| $$ "|"
@@ -647,10 +674,6 @@
 
 fun core_inference inf fact = ((fact, [fact]), Unknown, dummy_phi, inf, [])
 
-(* Syntax: <name> *)
-fun parse_satallax_core_line x =
-  (scan_general_id --| Scan.option ($$ " ") >> (single o core_inference satallax_core_rule)) x
-
 (* Syntax: SZS core <name> ... <name> *)
 fun parse_z3_tptp_core_line x =
   (Scan.this_string "SZS core" |-- Scan.repeat ($$ " " |-- scan_general_id)
@@ -658,35 +681,17 @@
 
 fun parse_line local_name problem =
   if local_name = leo2N then parse_tstp_thf0_line problem
-  else if local_name = satallaxN then parse_satallax_core_line
   else if local_name = spassN then parse_spass_line
   else if local_name = spass_pirateN then parse_spass_pirate_line
   else if local_name = z3_tptpN then parse_z3_tptp_core_line
   else parse_tstp_line problem
 
-fun parse_proof local_name problem =
-  strip_spaces_except_between_idents
-  #> raw_explode
-  #> Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())
-       (Scan.finite Symbol.stopper (Scan.repeat1 (parse_line local_name problem) >> flat)))
-  #> fst
-
 fun core_of_agsyhol_proof s =
   (case split_lines s of
     "The transformed problem consists of the following conjectures:" :: conj ::
     _ :: proof_term :: _ => SOME (unprefix " " conj :: find_enclosed "<<" ">>" proof_term)
   | _ => NONE)
 
-fun atp_proof_of_tstplike_proof _ _ "" = []
-  | atp_proof_of_tstplike_proof local_prover problem tstp =
-    (case core_of_agsyhol_proof tstp of
-      SOME facts => facts |> map (core_inference agsyhol_core_rule)
-    | NONE =>
-      tstp ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
-      |> parse_proof local_prover problem
-      |> local_prover = vampireN ? perhaps (try (sort (vampire_step_name_ord o pairself #1)))
-      |> local_prover = zipperpositionN ? rev)
-
 fun clean_up_dependencies _ [] = []
   | clean_up_dependencies seen ((name, role, u, rule, deps) :: steps) =
     (name, role, u, rule, map_filter (fn dep => find_first (is_same_atp_step dep) seen) deps) ::
--- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -30,6 +30,9 @@
 
   val forall_of : term -> term -> term
   val exists_of : term -> term -> term
+  val simplify_bool : term -> term
+  val var_name_of_typ : typ -> string
+  val rename_bound_vars : term -> term
   val type_enc_aliases : (string * string list) list
   val unalias_type_enc : string -> string list
   val term_of_atp : Proof.context -> ATP_Problem.atp_format -> ATP_Problem_Generate.type_enc ->
@@ -50,8 +53,8 @@
     ATP_Problem_Generate.type_enc -> string Symtab.table -> (string * term) list ->
     int Symtab.table -> string atp_proof -> (term, string) atp_step list
   val repair_waldmeister_endgame : (term, 'a) atp_step list -> (term, 'a) atp_step list
-  val infer_formula_types : Proof.context -> term list -> term list
-  val introduce_spass_skolem : (term, string) atp_step list -> (term, string) atp_step list
+  val infer_formulas_types : Proof.context -> term list -> term list
+  val introduce_spass_skolems : (term, string) atp_step list -> (term, string) atp_step list
   val factify_atp_proof : (string * 'a) list -> term list -> term -> (term, string) atp_step list ->
     (term, string) atp_step list
 end;
@@ -101,6 +104,46 @@
     TFree (ww, the_default @{sort type} (Variable.def_sort ctxt (ww, ~1)))
   end
 
+fun simplify_bool ((all as Const (@{const_name All}, _)) $ Abs (s, T, t)) =
+    let val t' = simplify_bool t in
+      if loose_bvar1 (t', 0) then all $ Abs (s, T, t') else t'
+    end
+  | simplify_bool (Const (@{const_name Not}, _) $ t) = s_not (simplify_bool t)
+  | simplify_bool (Const (@{const_name conj}, _) $ t $ u) =
+    s_conj (simplify_bool t, simplify_bool u)
+  | simplify_bool (Const (@{const_name disj}, _) $ t $ u) =
+    s_disj (simplify_bool t, simplify_bool u)
+  | simplify_bool (Const (@{const_name implies}, _) $ t $ u) =
+    s_imp (simplify_bool t, simplify_bool u)
+  | simplify_bool ((t as Const (@{const_name HOL.eq}, _)) $ u $ v) =
+    (case (u, v) of
+      (Const (@{const_name True}, _), _) => v
+    | (u, Const (@{const_name True}, _)) => u
+    | (Const (@{const_name False}, _), v) => s_not v
+    | (u, Const (@{const_name False}, _)) => s_not u
+    | _ => if u aconv v then @{const True} else t $ simplify_bool u $ simplify_bool v)
+  | simplify_bool (t $ u) = simplify_bool t $ simplify_bool u
+  | simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t)
+  | simplify_bool t = t
+
+fun single_letter upper s =
+  let val s' = if String.isPrefix "o" s orelse String.isPrefix "O" s then "z" else s in
+    String.extract (Name.desymbolize (SOME upper) (Long_Name.base_name s'), 0, SOME 1)
+  end
+
+fun var_name_of_typ (Type (@{type_name fun}, [_, T])) =
+    if body_type T = HOLogic.boolT then "p" else "f"
+  | var_name_of_typ (Type (@{type_name set}, [T])) = single_letter true (var_name_of_typ T)
+  | var_name_of_typ (Type (s, Ts)) =
+    if String.isSuffix "list" s then var_name_of_typ (the_single Ts) ^ "s"
+    else single_letter false (Long_Name.base_name s)
+  | var_name_of_typ (TFree (s, _)) = single_letter false (perhaps (try (unprefix "'")) s)
+  | var_name_of_typ (TVar ((s, _), T)) = var_name_of_typ (TFree (s, T))
+
+fun rename_bound_vars (t $ u) = rename_bound_vars t $ rename_bound_vars u
+  | rename_bound_vars (Abs (_, T, t)) = Abs (var_name_of_typ T, T, rename_bound_vars t)
+  | rename_bound_vars t = t
+
 exception ATP_CLASS of string list
 exception ATP_TYPE of string atp_type list
 exception ATP_TERM of (string, string atp_type) atp_term list
@@ -116,25 +159,33 @@
 (* Type variables are given the basic sort "HOL.type". Some will later be constrained by information
    from type literals, or by type inference. *)
 fun typ_of_atp_type ctxt (ty as AType ((a, clss), tys)) =
-  let val Ts = map (typ_of_atp_type ctxt) tys in
-    (case unprefix_and_unascii type_const_prefix a of
-      SOME b => Type (invert_const b, Ts)
-    | NONE =>
-      if not (null tys) then
-        raise ATP_TYPE [ty] (* only "tconst"s have type arguments *)
-      else
-        (case unprefix_and_unascii tfree_prefix a of
-          SOME b => make_tfree ctxt b
-        | NONE =>
-          (* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018".
-             Sometimes variables from the ATP are indistinguishable from Isabelle variables, which
-             forces us to use a type parameter in all cases. *)
-          Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a,
-            (if null clss then @{sort type} else map class_of_atp_class clss))))
+    let val Ts = map (typ_of_atp_type ctxt) tys in
+      (case unprefix_and_unascii type_const_prefix a of
+        SOME b => Type (invert_const b, Ts)
+      | NONE =>
+        if not (null tys) then
+          raise ATP_TYPE [ty] (* only "tconst"s have type arguments *)
+        else
+          (case unprefix_and_unascii tfree_prefix a of
+            SOME b => make_tfree ctxt b
+          | NONE =>
+            (* The term could be an Isabelle variable or a variable from the ATP, say "X1" or "_5018".
+               Sometimes variables from the ATP are indistinguishable from Isabelle variables, which
+               forces us to use a type parameter in all cases. *)
+            Type_Infer.param 0 ("'" ^ perhaps (unprefix_and_unascii tvar_prefix) a,
+              (if null clss then @{sort type} else map class_of_atp_class clss))))
+    end
+  | typ_of_atp_type ctxt (AFun (ty1, ty2)) = typ_of_atp_type ctxt ty1 --> typ_of_atp_type ctxt ty2
+
+fun atp_type_of_atp_term (ATerm ((s, _), us)) =
+  let val tys = map atp_type_of_atp_term us in
+    if s = tptp_fun_type then
+      (case tys of
+        [ty1, ty2] => AFun (ty1, ty2)
+      | _ => raise ATP_TYPE tys)
+    else
+      AType ((s, []), tys)
   end
-  | typ_of_atp_type ctxt (AFun (a, tys)) = typ_of_atp_type ctxt a --> typ_of_atp_type ctxt tys
-
-fun atp_type_of_atp_term (ATerm ((s, _), us)) = AType ((s, []), map atp_type_of_atp_term us)
 
 fun typ_of_atp_term ctxt = typ_of_atp_type ctxt o atp_type_of_atp_term
 
@@ -150,25 +201,10 @@
   | add_type_constraint false (cl, TVar (ix, _)) = add_var (ix, cl)
   | add_type_constraint _ _ = I
 
-fun repair_var_name_raw s =
-  let
-    fun subscript_name s n = s ^ nat_subscript n
-    val s = s |> String.map Char.toLower
-  in
-    (case space_explode "_" s of
-      [_] =>
-      (case take_suffix Char.isDigit (String.explode s) of
-        (cs1 as _ :: _, cs2 as _ :: _) =>
-        subscript_name (String.implode cs1) (the (Int.fromString (String.implode cs2)))
-      | (_, _) => s)
-    | [s1, s2] => (case Int.fromString s2 of SOME n => subscript_name s1 n | NONE => s)
-    | _ => s)
-  end
-
-fun repair_var_name textual s =
+fun repair_var_name s =
   (case unprefix_and_unascii schematic_var_prefix s of
-    SOME s => s
-  | NONE => s |> textual ? repair_var_name_raw)
+    SOME s' => s'
+  | NONE => s)
 
 (* The number of type arguments of a constant, zero if it's monomorphic. For (instances of) Skolem
    pseudoconstants, this information is encoded in the constant name. *)
@@ -182,58 +218,52 @@
 
 fun slack_fastype_of t = fastype_of t handle TERM _ => Type_Infer.anyT @{sort type}
 
-(* Cope with "tt(X) = X" atoms, where "X" is existentially quantified. *)
-fun loose_aconv (Free (s, _), Free (s', _)) = s = s'
-  | loose_aconv (t, t') = t aconv t'
-
 val spass_skolem_prefix = "sk" (* "skc" or "skf" *)
 val vampire_skolem_prefix = "sK"
 
-
 fun var_index_of_textual textual = if textual then 0 else 1
 
 fun quantify_over_var textual quant_of var_s var_T t =
   let
-    val vars = ((var_s, var_index_of_textual textual), var_T) ::
-                 filter (fn ((s, _), _) => s = var_s) (Term.add_vars t [])
+    val vars =
+      ((var_s, var_index_of_textual textual), var_T) ::
+      filter (fn ((s, _), _) => s = var_s) (Term.add_vars t [])
     val normTs = vars |> AList.group (op =) |> map (apsnd hd)
     fun norm_var_types (Var (x, T)) =
         Var (x, the_default T (AList.lookup (op =) normTs x))
       | norm_var_types t = t
   in t |> map_aterms norm_var_types |> fold_rev quant_of (map Var normTs) end
 
+(* This assumes that distinct names are mapped to distinct names by "Variable.variant_frees". This
+   does not hold in general but should hold for ATP-generated Skolem function names, since these end
+   with a digit and "variant_frees" appends letters. *)
+fun fresh_up ctxt s = fst (hd (Variable.variant_frees ctxt [] [(s, ())]))
 
 (* Higher-order translation. Variables are typed (although we don't use that information). Lambdas
-   are typed.
-
-  The code is similar to term_of_atp_fo. *)
-fun term_of_atp_ho ctxt textual sym_tab =
+   are typed. The code is similar to "term_of_atp_fo". *)
+fun term_of_atp_ho ctxt sym_tab =
   let
     val thy = Proof_Context.theory_of ctxt
-    val var_index = var_index_of_textual textual
+    val var_index = var_index_of_textual true
 
     fun do_term opt_T u =
       (case u of
-        AAbs(((var, ty), term), []) =>
+        AAbs (((var, ty), term), []) =>
         let
           val typ = typ_of_atp_type ctxt ty
-          val var_name = repair_var_name textual var
+          val var_name = repair_var_name var
           val tm = do_term NONE term
-        in quantify_over_var textual lambda' var_name typ tm end
+        in quantify_over_var true lambda' var_name typ tm end
       | ATerm ((s, tys), us) =>
         if s = ""
         then error "Isar proof reconstruction failed because the ATP proof \
                      \contains unparsable material."
         else if s = tptp_equal then
-          let val ts = map (do_term NONE) us in
-            if textual andalso length ts = 2 andalso loose_aconv (hd ts, List.last ts) then
-              @{const True}
-            else
-              list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}), ts)
-          end
+          list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}),
+            map (do_term NONE) us)
         else if not (null us) then
           let
-            val args = List.map (do_term NONE) us
+            val args = map (do_term NONE) us
             val opt_T' = SOME (map slack_fastype_of args ---> the_default dummyT opt_T)
             val func = do_term opt_T' (ATerm ((s, tys), []))
           in foldl1 (op $) (func :: args) end
@@ -241,13 +271,15 @@
         else if s = tptp_and then HOLogic.conj
         else if s = tptp_implies then HOLogic.imp
         else if s = tptp_iff orelse s = tptp_equal then HOLogic.eq_const dummyT
-        else if s = tptp_not_iff orelse s = tptp_not_equal then @{term "% P Q. Q ~= P"}
-        else if s = tptp_if then @{term "% P Q. Q --> P"}
-        else if s = tptp_not_and then @{term "% P Q. ~ (P & Q)"}
-        else if s = tptp_not_or then @{term "% P Q. ~ (P | Q)"}
+        else if s = tptp_not_iff orelse s = tptp_not_equal then @{term "%P Q. Q ~= P"}
+        else if s = tptp_if then @{term "%P Q. Q --> P"}
+        else if s = tptp_not_and then @{term "%P Q. ~ (P & Q)"}
+        else if s = tptp_not_or then @{term "%P Q. ~ (P | Q)"}
         else if s = tptp_not then HOLogic.Not
         else if s = tptp_ho_forall then HOLogic.all_const dummyT
         else if s = tptp_ho_exists then HOLogic.exists_const dummyT
+        else if s = tptp_hilbert_choice then HOLogic.choice_const dummyT
+        else if s = tptp_hilbert_the then @{term "The"}
         else
           (case unprefix_and_unascii const_prefix s of
             SOME s' =>
@@ -259,7 +291,7 @@
               val Ts = map (typ_of_atp_type ctxt) tys @ map (typ_of_atp_term ctxt) type_us
               val T =
                 (if not (null Ts) andalso robust_const_num_type_args thy s' = length Ts then
-                   if textual then try (Sign.const_instance thy) (s', Ts) else NONE
+                   try (Sign.const_instance thy) (s', Ts)
                  else
                    NONE)
                 |> (fn SOME T => T
@@ -270,8 +302,6 @@
             in list_comb (t, term_ts) end
           | NONE => (* a free or schematic variable *)
             let
-              fun fresh_up s =
-                [(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst
               val ts = map (do_term NONE) us
               val T =
                 (case opt_T of
@@ -285,10 +315,8 @@
                 (case unprefix_and_unascii fixed_var_prefix s of
                   SOME s => Free (s, T)
                 | NONE =>
-                  if textual andalso not (is_tptp_variable s) then
-                    Free (s |> textual ? (repair_var_name_raw #> fresh_up), T)
-                  else
-                    Var ((repair_var_name textual s, var_index), T))
+                  if not (is_tptp_variable s) then Free (fresh_up ctxt s, T)
+                  else Var ((repair_var_name s, var_index), T))
             in list_comb (t, ts) end))
   in do_term end
 
@@ -311,12 +339,8 @@
         else if String.isPrefix native_type_prefix s then
           @{const True} (* ignore TPTP type information *)
         else if s = tptp_equal then
-          let val ts = map (do_term [] NONE) us in
-            if textual andalso length ts = 2 andalso loose_aconv (hd ts, List.last ts) then
-              @{const True}
-            else
-              list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}), ts)
-          end
+          list_comb (Const (@{const_name HOL.eq}, Type_Infer.anyT @{sort type}),
+            map (do_term [] NONE) us)
         else
           (case unprefix_and_unascii const_prefix s of
             SOME s' =>
@@ -363,16 +387,10 @@
             end
           | NONE => (* a free or schematic variable *)
             let
-              (* This assumes that distinct names are mapped to distinct names by
-                 "Variable.variant_frees". This does not hold in general but should hold for
-                 ATP-generated Skolem function names, since these end with a digit and
-                 "variant_frees" appends letters. *)
-              fun fresh_up s = [(s, ())] |> Variable.variant_frees ctxt [] |> hd |> fst
-
               val term_ts =
                 map (do_term [] NONE) us
                 (* SPASS (3.8ds) and Vampire (2.6) pass arguments to Skolem functions in reverse
-                   order, which is incompatible with the new Metis skolemizer. *)
+                   order, which is incompatible with "metis"'s new skolemizer. *)
                 |> exists (fn pre => String.isPrefix pre s)
                   [spass_skolem_prefix, vampire_skolem_prefix] ? rev
               val ts = term_ts @ extra_ts
@@ -387,19 +405,17 @@
                   SOME s => Free (s, T)
                 | NONE =>
                   if textual andalso not (is_tptp_variable s) then
-                    Free (s |> textual ? (repair_var_name_raw #> fresh_up), T)
+                    Free (s |> textual ? fresh_up ctxt, T)
                   else
-                    Var ((repair_var_name textual s, var_index), T))
+                    Var ((repair_var_name s, var_index), T))
             in list_comb (t, ts) end))
   in do_term [] end
 
 fun term_of_atp ctxt (ATP_Problem.THF _) type_enc =
-    if ATP_Problem_Generate.is_type_enc_higher_order type_enc
-    then term_of_atp_ho ctxt
+    if ATP_Problem_Generate.is_type_enc_higher_order type_enc then K (term_of_atp_ho ctxt)
     else error "Unsupported Isar reconstruction."
   | term_of_atp ctxt _ type_enc =
-    if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc)
-    then term_of_atp_fo ctxt
+    if not (ATP_Problem_Generate.is_type_enc_higher_order type_enc) then term_of_atp_fo ctxt
     else error "Unsupported Isar reconstruction."
 
 fun term_of_atom ctxt format type_enc textual sym_tab pos (u as ATerm ((s, _), _)) =
@@ -436,7 +452,7 @@
         do_formula pos (AQuant (q, xs, phi'))
         (* FIXME: TFF *)
         #>> quantify_over_var textual (case q of AForall => forall_of | AExists => exists_of)
-          (repair_var_name textual s) dummyT
+          (repair_var_name s) dummyT
       | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
       | AConn (c, [phi1, phi2]) =>
         do_formula (pos |> c = AImplies ? not) phi1
@@ -477,7 +493,7 @@
 val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
 
 fun add_fact ctxt facts ((_, ss), _, _, rule, deps) =
-  (if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule, satallax_core_rule] rule then
+  (if member (op =) [agsyhol_core_rule, leo2_extcnf_equal_neg_rule] rule then
      insert (op =) (short_thm_name ctxt ext, (Global, General))
    else
      I)
@@ -530,9 +546,9 @@
     else
       s
 
-fun set_var_index j = map_aterms (fn Var ((s, _), T) => Var ((s, j), T) | t => t)
+fun set_var_index j = map_aterms (fn Var ((s, 0), T) => Var ((s, j), T) | t => t)
 
-fun infer_formula_types ctxt =
+fun infer_formulas_types ctxt =
   map_index (uncurry (fn j => set_var_index j #> Type.constraint HOLogic.boolT))
   #> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
   #> map (set_var_index 0)
@@ -546,25 +562,26 @@
 
 fun uncombine_term thy =
   let
-    fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
-      | aux (Abs (s, T, t')) = Abs (s, T, aux t')
-      | aux (t as Const (x as (s, _))) =
+    fun uncomb (t1 $ t2) = betapply (uncomb t1, uncomb t2)
+      | uncomb (Abs (s, T, t)) = Abs (s, T, uncomb t)
+      | uncomb (t as Const (x as (s, _))) =
         (case AList.lookup (op =) combinator_table s of
           SOME thm => thm |> prop_of |> specialize_type thy x |> Logic.dest_equals |> snd
         | NONE => t)
-      | aux t = t
-  in aux end
+      | uncomb t = t
+  in uncomb end
 
-fun unlift_term lifted =
-  map_aterms (fn t as Const (s, _) =>
-                 if String.isPrefix lam_lifted_prefix s then
-                   (* FIXME: do something about the types *)
-                   (case AList.lookup (op =) lifted s of
-                     SOME t => unlift_term lifted t
-                   | NONE => t)
-                 else
-                   t
-               | t => t)
+fun unlift_aterm lifted (t as Const (s, _)) =
+    if String.isPrefix lam_lifted_prefix s then
+      (* FIXME: do something about the types *)
+      (case AList.lookup (op =) lifted s of
+        SOME t' => unlift_term lifted t'
+      | NONE => t)
+    else
+      t
+  | unlift_aterm _ t = t
+and unlift_term lifted =
+  map_aterms (unlift_aterm lifted)
 
 fun termify_line _ _ _ _ _ (_, Type_Role, _, _, _) = NONE
   | termify_line ctxt format type_enc lifted sym_tab (name, role, u, rule, deps) =
@@ -572,8 +589,9 @@
       val thy = Proof_Context.theory_of ctxt
       val t = u
         |> prop_of_atp ctxt format type_enc true sym_tab
+        |> unlift_term lifted
         |> uncombine_term thy
-        |> unlift_term lifted
+        |> simplify_bool
     in
       SOME (name, role, t, rule, deps)
     end
@@ -599,7 +617,7 @@
   nasty_atp_proof pool
   #> map_term_names_in_atp_proof repair_name
   #> map_filter (termify_line ctxt format type_enc lifted sym_tab)
-  #> map_proof_terms (infer_formula_types ctxt #> map HOLogic.mk_Trueprop)
+  #> map_proof_terms (infer_formulas_types ctxt #> map HOLogic.mk_Trueprop)
   #> local_prover = waldmeisterN ? repair_waldmeister_endgame
 
 fun take_distinct_vars seen ((t as Var _) :: ts) =
@@ -608,12 +626,12 @@
 
 fun unskolemize_term skos t =
   let
-    val is_sko = member (op =) skos
+    val is_skolem = member (op =) skos
 
     fun find_args args (t $ u) = find_args (u :: args) t #> find_args [] u
       | find_args _ (Abs (_, _, t)) = find_args [] t
       | find_args args (Free (s, _)) =
-        if is_sko s then
+        if is_skolem s then
           let val new = take_distinct_vars [] args in
             (fn [] => new | old => if length new < length old then new else old)
           end
@@ -626,7 +644,7 @@
 
     fun fix_skos args (t $ u) = fix_skos (fix_skos [] u :: args) t
       | fix_skos args (t as Free (s, T)) =
-        if is_sko s then list_comb (Free (s, funpow num_alls body_type T), drop num_alls args)
+        if is_skolem s then list_comb (Free (s, funpow num_alls body_type T), drop num_alls args)
         else list_comb (t, args)
       | fix_skos [] (Abs (s, T, t)) = Abs (s, T, fix_skos [] t)
       | fix_skos [] t = t
@@ -634,10 +652,10 @@
 
     val t' = fix_skos [] t
 
-    fun add_sko (t as Free (s, _)) = is_sko s ? insert (op aconv) t
-      | add_sko _ = I
+    fun add_skolem (t as Free (s, _)) = is_skolem s ? insert (op aconv) t
+      | add_skolem _ = I
 
-    val exs = Term.fold_aterms add_sko t' []
+    val exs = Term.fold_aterms add_skolem t' []
   in
     t'
     |> HOLogic.dest_Trueprop
@@ -646,62 +664,80 @@
     |> HOLogic.mk_Trueprop
   end
 
-fun introduce_spass_skolem [] = []
-  | introduce_spass_skolem (proof as (_, _, _, rule1, _) :: _) =
-    if rule1 = spass_input_rule then
-      let
-        fun add_sko (Free (s, _)) = String.isPrefix spass_skolem_prefix s ? insert (op =) s
-          | add_sko _ = I
+fun rename_skolem_args t =
+  let
+    fun add_skolem_args (Abs (_, _, t)) = add_skolem_args t
+      | add_skolem_args t =
+        (case strip_comb t of
+          (Free (s, _), args as _ :: _) =>
+          if String.isPrefix spass_skolem_prefix s then
+            insert (op =) (s, fst (take_prefix is_Var args))
+          else
+            fold add_skolem_args args
+        | (u, args as _ :: _) => fold add_skolem_args (u :: args)
+        | _ => I)
+
+    fun subst_of_skolem (sk, args) =
+      map_index (fn (j, Var (z, T)) => (z, Var ((sk ^ "_" ^ string_of_int j, 0), T))) args
 
-        (* union-find would be faster *)
-        fun add_cycle [] = I
-          | add_cycle ss =
-            fold (fn s => Graph.default_node (s, ())) ss
-            #> fold Graph.add_edge (ss ~~ tl ss @ [hd ss])
-
-        val (input_steps, other_steps) = List.partition (null o #5) proof
+    val subst = maps subst_of_skolem (add_skolem_args t [])
+  in
+    subst_vars ([], subst) t
+  end
 
-        val skoss = map (fn (_, _, t, _, _) => Term.fold_aterms add_sko t []) input_steps
-        val skoss_input_steps = filter_out (null o fst) (skoss ~~ input_steps)
-        val groups = Graph.strong_conn (fold add_cycle skoss Graph.empty)
+fun introduce_spass_skolems proof =
+  let
+    fun add_skolem (Free (s, _)) = String.isPrefix spass_skolem_prefix s ? insert (op =) s
+      | add_skolem _ = I
 
-        fun step_name_of_group skos = (implode skos, [])
-        fun in_group group = member (op =) group o hd
-        fun group_of sko = the (find_first (fn group => in_group group sko) groups)
+    (* union-find would be faster *)
+    fun add_cycle [] = I
+      | add_cycle ss =
+        fold (fn s => Graph.default_node (s, ())) ss
+        #> fold Graph.add_edge (ss ~~ tl ss @ [hd ss])
+
+    val (input_steps, other_steps) = List.partition (null o #5) proof
 
-        fun new_steps (skoss_steps : (string list * (term, 'a) atp_step) list) group =
-          let
-            val name = step_name_of_group group
-            val name0 = name |>> prefix "0"
-            val t =
-              skoss_steps
-              |> map (snd #> #3 #> HOLogic.dest_Trueprop)
-              |> Library.foldr1 s_conj
-              |> HOLogic.mk_Trueprop
-            val skos = fold (union (op =) o fst) skoss_steps []
-            val deps = map (snd #> #1) skoss_steps
-          in
-            [(name0, Unknown, unskolemize_term skos t, spass_pre_skolemize_rule, deps),
-             (name, Unknown, t, spass_skolemize_rule, [name0])]
-          end
+    val skoss = map (fn (_, _, t, _, _) => Term.fold_aterms add_skolem t []) input_steps
+    val skoss_input_steps = filter_out (null o fst) (skoss ~~ input_steps)
+    val groups = Graph.strong_conn (fold add_cycle skoss Graph.empty)
+
+    fun step_name_of_group skos = (implode skos, [])
+    fun in_group group = member (op =) group o hd
+    fun group_of sko = the (find_first (fn group => in_group group sko) groups)
 
-        val sko_steps =
-          maps (fn group => new_steps (filter (in_group group o fst) skoss_input_steps) group)
-            groups
+    fun new_steps (skoss_steps : (string list * (term, 'a) atp_step) list) group =
+      let
+        val name = step_name_of_group group
+        val name0 = name |>> prefix "0"
+        val t =
+          (case map (snd #> #3) skoss_steps of
+            [t] => t
+          | ts => ts
+            |> map (HOLogic.dest_Trueprop #> rename_skolem_args)
+            |> Library.foldr1 s_conj
+            |> HOLogic.mk_Trueprop)
+        val skos = fold (union (op =) o fst) skoss_steps []
+        val deps = map (snd #> #1) skoss_steps
+      in
+        [(name0, Unknown, unskolemize_term skos t, spass_pre_skolemize_rule, deps),
+         (name, Unknown, t, spass_skolemize_rule, [name0])]
+      end
 
-        val old_news =
-          map (fn (skos, (name, _, _, _, _)) => (name, [step_name_of_group (group_of skos)]))
-            skoss_input_steps
-        val repair_deps = fold replace_dependencies_in_line old_news
-      in
-        input_steps @ sko_steps @ map repair_deps other_steps
-      end
-  else
-    proof
+    val sko_steps =
+      maps (fn group => new_steps (filter (in_group group o fst) skoss_input_steps) group) groups
+
+    val old_news =
+      map (fn (skos, (name, _, _, _, _)) => (name, [step_name_of_group (group_of skos)]))
+        skoss_input_steps
+    val repair_deps = fold replace_dependencies_in_line old_news
+  in
+    input_steps @ sko_steps @ map repair_deps other_steps
+  end
 
 fun factify_atp_proof facts hyp_ts concl_t atp_proof =
   let
-    fun factify_step ((num, ss), _, t, rule, deps) =
+    fun factify_step ((num, ss), role, t, rule, deps) =
       let
         val (ss', role', t') =
           (case resolve_conjectures ss of
@@ -709,7 +745,7 @@
             if j = length hyp_ts then ([], Conjecture, concl_t) else ([], Hypothesis, nth hyp_ts j)
           | _ =>
             (case resolve_facts facts ss of
-              [] => (ss, Plain, t)
+              [] => (ss, if role = Lemma then Lemma else Plain, t)
             | facts => (map fst facts, Axiom, t)))
       in
         ((num, ss'), role', t', rule, deps)
--- a/src/HOL/Tools/ATP/atp_proof_redirect.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_proof_redirect.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -167,8 +167,9 @@
           val subseqss = map (subsequents seqs) zones
           val seqs = fold (subtract direct_sequent_eq) subseqss seqs
           val cases = map2 (fn l => fn subseqs => ([l], redirect [l] proved subseqs)) c subseqss
+          val cases' = filter_out (null o snd) cases
         in
-          s_cases cases @ redirect (succedent_of_cases cases) proved seqs
+          s_cases cases' @ redirect (succedent_of_cases cases) proved seqs
         end
   in
     redirect [] axioms seqs
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/ATP/atp_satallax.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,412 @@
+(*  Title:      HOL/Tools/ATP/atp_satallax.ML
+    Author:     Mathias Fleury, ENS Rennes
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Satallax proof parser and transformation for Sledgehammer.
+*)
+
+signature ATP_SATALLAX =
+sig
+  val atp_proof_of_tstplike_proof : string -> string ATP_Proof.atp_problem -> string ->
+    string ATP_Proof.atp_proof
+end;
+
+structure ATP_Satallax : ATP_SATALLAX =
+struct
+
+open ATP_Proof
+open ATP_Util
+open ATP_Problem
+
+(*Undocumented format:
+thf (number, plain | Axiom | ..., inference(rule, [status(thm), assumptions ([hypotheses list]),
+detailed_rule(discharge,used_hypothese_list) *], used_hypotheses_list, premises))
+(seen by tab_mat)
+
+Also seen -- but we can ignore it:
+"tab_inh (a) __11." meaning that the type a is inhabited usueful of variable eigen__11,
+*)
+fun parse_satallax_tstp_information x =
+  ((Symbol.scan_ascii_id || ($$ "$" |-- Symbol.scan_ascii_id))
+  -- Scan.option ( $$ "(" |-- Scan.option (Symbol.scan_ascii_id --| $$ ",")
+  -- ((Scan.option (($$ "[" |-- (Scan.option ((scan_general_id
+         -- Scan.repeat ($$ "," |-- scan_general_id)) >> op ::) --| $$ "]"))
+         || (skip_term >> K "") >> (fn x => SOME [x]))
+       >> (fn (SOME x) => x | NONE => NONE)) --| $$ ")"))) x
+
+fun parse_prem x =
+  ((Symbol.scan_ascii_id || scan_general_id) --| Scan.option ($$ ":" -- skip_term)) x
+
+fun parse_prems x =
+  (($$ "[" |-- parse_prem -- Scan.repeat ($$ "," |-- parse_prem) --| $$ "]")
+     >> op ::) x
+
+fun parse_tstp_satallax_extra_arguments x =
+  ($$ "," |-- scan_general_id -- (($$ "(" |-- Symbol.scan_ascii_id --| $$ "," )
+  -- ($$ "[" |-- Scan.option ((parse_satallax_tstp_information
+  -- Scan.repeat ($$ "," |-- parse_satallax_tstp_information)) >> op ::)
+  --| $$ "]") --
+  (Scan.optional ($$ "," |-- parse_prems) [] -- Scan.optional ($$ "," |-- parse_prems) []
+    >> (fn (x, []) => x | (_, x) => x))
+   --| $$ ")")) x
+
+val dummy_satallax_step = ((("~1", "tab_inh"), AAtom (ATerm(("False", []), []))), NONE)
+fun parse_tstp_thf0_satallax_line x =
+  (((Scan.this_string tptp_thf
+  -- $$ "(") |-- scan_general_id --| $$ "," -- Symbol.scan_ascii_id --| $$ ","
+  -- parse_thf0_formula -- Scan.option parse_tstp_satallax_extra_arguments --| $$ ")" --| $$ ".")
+  || (Scan.this_string "tab_inh" |-- skip_term --| $$ ".")
+     >> K dummy_satallax_step) x
+
+datatype satallax_step = Satallax_Step of {
+  id: string,
+  role: string,
+  theorem: (string, string, (string, string ATP_Problem.atp_type) ATP_Problem.atp_term, string)
+    ATP_Problem.atp_formula,
+  assumptions: string list,
+  rule: string,
+  used_assumptions: string list,
+  detailed_eigen: string,
+  generated_goal_assumptions: (string * string list) list}
+
+fun mk_satallax_step id role theorem assumptions rule used_assumptions
+    generated_goal_assumptions detailed_eigen =
+  Satallax_Step {id = id, role = role, theorem = theorem, assumptions = assumptions, rule = rule,
+    used_assumptions = used_assumptions, generated_goal_assumptions = generated_goal_assumptions,
+    detailed_eigen = detailed_eigen}
+
+fun get_assumptions (("assumptions", SOME (_ , assumptions)) :: _) =
+    the_default [] assumptions
+  | get_assumptions (_ :: l) = get_assumptions l
+  | get_assumptions [] = []
+
+fun get_detailled_eigen ((_, SOME (SOME "eigenvar" , var)) :: _) =
+    hd (the_default [""] var)
+  | get_detailled_eigen (_ :: l) = get_detailled_eigen l
+  | get_detailled_eigen [] = ""
+
+fun seperate_dependices dependencies =
+  let
+    fun sep_dep [] used_assumptions new_goals generated_assumptions _ =
+        (used_assumptions, new_goals, generated_assumptions)
+      | sep_dep (x :: l) used_assumptions new_goals generated_assumptions state =
+        if hd (raw_explode x) = "h" orelse Int.fromString x = NONE  then
+          if state = 0 then
+            sep_dep l (x :: used_assumptions) new_goals generated_assumptions state
+          else
+            sep_dep l used_assumptions new_goals (x :: generated_assumptions) 3
+        else
+          if state = 1 orelse state = 0 then
+            sep_dep l used_assumptions (x :: new_goals) generated_assumptions 1
+          else
+            raise Fail ("incorrect Satallax proof: " ^ @{make_string} l)
+  in
+    sep_dep dependencies [] [] [] 0
+  end
+
+fun create_grouped_goal_assumption rule new_goals generated_assumptions =
+  let
+    val number_of_new_goals = length new_goals
+    val number_of_new_assms = length generated_assumptions
+  in
+    if number_of_new_goals = number_of_new_assms then
+      new_goals ~~ (map single generated_assumptions)
+    else if 2 * number_of_new_goals = number_of_new_assms then
+      let
+        fun group_by_pair (new_goal :: new_goals) (assumpt1 :: assumpt2 :: generated_assumptions) =
+            (new_goal, [assumpt1, assumpt2]) :: group_by_pair new_goals generated_assumptions
+          | group_by_pair [] [] = []
+      in
+        group_by_pair new_goals generated_assumptions
+      end
+    else
+      raise Fail ("the rule " ^ rule ^" is not supported in the reconstruction.")
+  end
+fun to_satallax_step (((id, role), formula), (SOME (_,((rule, l), used_rules)))) =
+    let
+      val (used_assumptions, new_goals, generated_assumptions) = seperate_dependices used_rules
+    in
+      mk_satallax_step id role formula (get_assumptions (the_default [] l)) rule used_assumptions
+        (create_grouped_goal_assumption rule new_goals generated_assumptions)
+        (get_detailled_eigen (the_default [] l))
+    end
+  | to_satallax_step (((id, role), formula), NONE) =
+      mk_satallax_step id role formula [] "assumption" [] [] ""
+
+fun is_assumption (Satallax_Step {role, ...}) = role = "assumption" orelse role = "axiom" orelse
+  role = "negated_conjecture" orelse role = "conjecture"
+
+fun seperate_assumptions_and_steps l =
+  let
+    fun seperate_assumption [] l l' = (l, l')
+      | seperate_assumption (step :: steps) l l' =
+        if is_assumption step then
+          seperate_assumption steps (step :: l) l'
+        else
+          seperate_assumption steps l (step :: l')
+  in
+    seperate_assumption l [] []
+  end
+
+datatype satallax_proof_graph =
+  Linear_Part of {node: satallax_step, succs: satallax_proof_graph list} |
+  Tree_Part of {node: satallax_step, deps: satallax_proof_graph list}
+
+fun find_proof_step ((x as Satallax_Step {id, ...}) :: l) h =
+    if h = id then x else find_proof_step l h
+  | find_proof_step [] h = raise Fail ("not_found: " ^ @{make_string} h ^ " (probably a parsing \
+    \error)")
+
+fun remove_not_not (x as ATerm ((op1, _), [ATerm ((op2, _), [th])])) =
+    if op1 = op2 andalso op1 = tptp_not then th else x
+  | remove_not_not th = th
+
+fun tptp_term_equal (ATerm((op1, _), l1), ATerm((op2, _), l2)) = op1 = op2 andalso
+    fold2 (fn t1 => fn t2 => fn c => c andalso t1 = t2) l1 l2 true
+  | tptp_term_equal (_, _) = false
+
+val dummy_true_aterm = ATerm (("$true", [dummy_atype]), [])
+
+fun find_assumptions_to_inline ths (assm :: assms) to_inline no_inline =
+    (case List.find (curry (op =) assm o fst) no_inline of
+      SOME _ => find_assumptions_to_inline ths assms to_inline no_inline
+    | NONE =>
+      (case List.find (curry (op =) assm o fst) to_inline of
+        NONE => find_assumptions_to_inline ths assms to_inline no_inline
+      | SOME (_, th) =>
+        let
+          val simplified_ths_with_inlined_assumption =
+            (case List.partition (curry tptp_term_equal th o remove_not_not) ths of
+              ([], ths) => ATerm ((tptp_not, [dummy_atype]), [th]) :: ths
+            | (_, _) => [])
+        in
+          find_assumptions_to_inline simplified_ths_with_inlined_assumption assms to_inline no_inline
+        end))
+  | find_assumptions_to_inline ths [] _ _ = ths
+
+fun inline_if_needed_and_simplify th assms to_inline no_inline =
+    (case find_assumptions_to_inline [th] assms to_inline no_inline of
+      [] => dummy_true_aterm
+  | l => foldl1 (fn (a, b) =>
+    (case b of
+      ATerm (("$false", _), _) => a
+    | ATerm (("~", _ ), [ATerm (("$true", _), _)]) => a
+    | _ => ATerm ((tptp_or, [dummy_atype]), [a, b]))) l)
+
+fun get_conclusion (Satallax_Step {theorem = AAtom theorem, ...}) = theorem
+
+fun add_assumption new_used_assumptions ((Satallax_Step {id, role, theorem, assumptions,
+    rule, generated_goal_assumptions, used_assumptions, detailed_eigen})) =
+  mk_satallax_step id role theorem assumptions rule (new_used_assumptions @ used_assumptions)
+    generated_goal_assumptions detailed_eigen
+
+fun set_rule new_rule (Satallax_Step {id, role, theorem, assumptions,
+    generated_goal_assumptions, used_assumptions, detailed_eigen, ...}) =
+  mk_satallax_step id role theorem assumptions new_rule used_assumptions
+    generated_goal_assumptions detailed_eigen
+
+fun transform_inline_assumption hypotheses_step proof_sketch =
+  let
+    fun inline_in_step (Linear_Part {node as Satallax_Step {generated_goal_assumptions,
+          used_assumptions, rule, ...}, succs}) =
+        if generated_goal_assumptions = [] then
+          Linear_Part {node = node, succs = []}
+        else
+          let
+            (*one singel goal is created, two hypothesis can be created, for the "and" rule:
+              (A /\ B) create two hypotheses A, and B.*)
+            fun set_hypotheses_as_goal [hypothesis] succs =
+                Linear_Part {node = set_rule rule (add_assumption used_assumptions 
+                    (find_proof_step hypotheses_step hypothesis)),
+                  succs = map inline_in_step succs}
+              | set_hypotheses_as_goal (hypothesis :: new_goal_hypotheses) succs =
+                Linear_Part {node = set_rule rule (add_assumption used_assumptions
+                    (find_proof_step hypotheses_step hypothesis)),
+                  succs = [set_hypotheses_as_goal new_goal_hypotheses succs]}
+          in
+            set_hypotheses_as_goal (snd (hd generated_goal_assumptions)) succs
+          end
+      | inline_in_step (Tree_Part {node = node, deps}) =
+        Tree_Part {node = node, deps = map inline_in_step deps}
+
+    fun inline_contradictory_assumptions (Linear_Part {node as Satallax_Step{id, theorem, ...},
+       succs}) (to_inline, no_inline) =
+      let
+        val (succs, inliner) = fold_map inline_contradictory_assumptions succs 
+          (to_inline, (id, theorem) :: no_inline)
+      in
+        (Linear_Part {node = node, succs = succs}, inliner)
+      end
+    | inline_contradictory_assumptions (Tree_Part {node = Satallax_Step {id, role,
+        theorem = AAtom theorem, assumptions, rule, generated_goal_assumptions,
+        used_assumptions, detailed_eigen}, deps}) (to_inline, no_inline) =
+      let
+        val (dep', (to_inline', no_inline')) = fold_map inline_contradictory_assumptions deps 
+          (to_inline, no_inline)
+        val to_inline'' =
+          map (fn s => (s, get_conclusion (find_proof_step hypotheses_step s)))
+            (List.filter (fn s => nth_string s 0 = "h")
+              (used_assumptions @ (map snd generated_goal_assumptions |> flat)))
+          @ to_inline'
+        val node' = Satallax_Step {id = id, role = role, theorem =
+            AAtom (inline_if_needed_and_simplify theorem assumptions to_inline'' no_inline'),
+          assumptions = assumptions, rule = rule,
+          generated_goal_assumptions = generated_goal_assumptions, detailed_eigen = detailed_eigen,
+          used_assumptions =
+            List.filter (fn s => List.find (curry (op =) s o fst) to_inline = NONE)
+            used_assumptions}
+      in
+        (Tree_Part {node = node', deps = dep'}, (to_inline'', no_inline'))
+      end
+  in
+    fst (inline_contradictory_assumptions (inline_in_step proof_sketch) ([], []))
+  end
+
+fun correct_dependencies (Linear_Part {node, succs}) =
+    Linear_Part {node = node, succs = map correct_dependencies succs}
+  | correct_dependencies (Tree_Part {node, deps}) =
+    let
+      val new_used_assumptions =
+        map (fn Linear_Part {node = (Satallax_Step{id, ...}), ...} => id
+              | Tree_Part {node = (Satallax_Step{id, ...}), ...} => id) deps
+    in
+      Tree_Part {node = add_assumption new_used_assumptions node,
+        deps = map correct_dependencies deps}
+    end
+
+fun create_satallax_proof_graph (hypotheses_step, proof_sketch) =
+  let
+    fun create_step (step as Satallax_Step {generated_goal_assumptions, ...}) =
+      Linear_Part {node = step,
+        succs = map (create_step o (find_proof_step (hypotheses_step @ proof_sketch)))
+          (map fst generated_goal_assumptions)}
+    fun reverted_discharged_steps is_branched (Linear_Part {node as
+          Satallax_Step {generated_goal_assumptions, ...}, succs}) =
+        if is_branched orelse length generated_goal_assumptions > 1 then
+          Tree_Part {node = node, deps = map (reverted_discharged_steps true) succs}
+        else
+          Linear_Part {node = node, succs = map (reverted_discharged_steps is_branched) succs}
+    val proof_with_correct_sense =
+        correct_dependencies (reverted_discharged_steps false
+          (create_step (find_proof_step proof_sketch "0" )))
+  in
+    transform_inline_assumption hypotheses_step proof_with_correct_sense
+  end
+
+val satallax_known_theorems = ["eq_ind", "eq_trans2", "eq_trans3", "eq_neg_neg_id", "eq_true",
+  "eq_and_nor", "eq_or_nand", "eq_or_imp", "eq_and_imp", "eq_imp_or", "eq_iff", "eq_sym_eq",
+  "eq_forall_nexists", "eq_exists_nforall", "eq_leib1", "eq_leib2", "eq_leib3", "eq_leib4",
+  "eq_eta", "SinhE", "eq_forall_Seps", "eq_SPi_Seps", "eq_exists_Seps"]
+val is_special_satallax_rule = member (op =) satallax_known_theorems
+
+fun terms_to_upper_case var (AAbs (((var', ty), b), ts)) =
+    let
+      val bdy = terms_to_upper_case var b
+      val ts' = map (terms_to_upper_case var) ts
+    in
+      AAbs (((((var = var' ? String.implode o (map Char.toUpper) o String.explode) var'), ty),
+        bdy), ts')
+    end
+  | terms_to_upper_case var (ATerm ((var', ty), ts)) =
+    ATerm ((((var = var' ? String.implode o (map Char.toUpper) o String.explode) var'),
+      ty), map (terms_to_upper_case var) ts)
+
+fun add_quantifier_in_tree_part vars_to_add assumption_to_add (Linear_Part {node, succs}) =
+    Linear_Part {node = node, succs = map (add_quantifier_in_tree_part vars_to_add assumption_to_add) succs}
+  | add_quantifier_in_tree_part vars_to_add assumption_to_add (Tree_Part {node = Satallax_Step {rule, detailed_eigen,
+      id, role, theorem = AAtom theorem, assumptions, used_assumptions, generated_goal_assumptions},
+      deps}) =
+    let
+      val theorem' = fold (fn v => fn body => terms_to_upper_case v body) vars_to_add theorem
+      val theorem'' = theorem'
+      val node' = mk_satallax_step id role (AAtom theorem'') assumptions rule
+        (used_assumptions @ (filter (curry (op <=) (the (Int.fromString id)) o size)
+        assumption_to_add)) generated_goal_assumptions detailed_eigen
+    in
+      if detailed_eigen <> "" then
+        Tree_Part {node = node',
+          deps = map (add_quantifier_in_tree_part (detailed_eigen :: vars_to_add)
+          (used_assumptions @ assumption_to_add)) deps}
+      else
+        Tree_Part {node = node',
+                   deps = map (add_quantifier_in_tree_part vars_to_add assumption_to_add) deps}
+    end
+
+fun transform_to_standard_atp_step already_transformed proof =
+  let
+    fun create_fact_step id =
+      ((id, [id]), Axiom, AAtom (ATerm((id, []), [])), "", [])
+    fun transform_one_step already_transformed (Satallax_Step {id, theorem, used_assumptions, role,
+        rule, ...}) =
+      let
+        val role' = role_of_tptp_string role
+        val new_transformed = List.filter
+          (fn s => size s >=4 andalso not (is_special_satallax_rule s) andalso not
+          (member (op =) already_transformed s)) used_assumptions
+      in
+        (map create_fact_step new_transformed
+        @ [((id, []), if role' = Unknown then Plain else role', theorem, rule,
+           map (fn s => (s, [])) (filter_out is_special_satallax_rule used_assumptions))],
+        new_transformed @ already_transformed)
+      end
+    fun transform_steps (Linear_Part {node, succs}) (already_transformed:string list) =
+        transform_one_step already_transformed node
+        ||> (fold_map transform_steps succs)
+        ||> apfst flat
+        |> (fn (a, (b, transformed)) => (a @ b, transformed))
+      | transform_steps (Tree_Part {node, deps}) (already_transformed: string list) =
+        fold_map transform_steps deps already_transformed
+        |>> flat
+        ||> (fn transformed => transform_one_step transformed node)
+        |> (fn (a, (b, transformed)) => (a @ b, transformed))
+  in
+    fst (transform_steps proof already_transformed)
+  end
+
+fun remove_unused_dependency steps =
+  let
+    fun find_all_ids [] = []
+      | find_all_ids (((id, _), _, _, _, _) :: steps) = id :: find_all_ids steps
+    fun keep_only_used used_ids steps =
+      let
+        fun remove_unused (((id, ids), role, theorem, rule, deps) :: steps) =
+            (((id, ids), role, theorem, rule, filter (member (op =) used_ids o fst) deps) :: steps)
+          | remove_unused [] = []
+      in
+        remove_unused steps
+      end
+  in
+    keep_only_used (find_all_ids steps) steps
+  end
+
+fun parse_proof local_name problem =
+  strip_spaces_except_between_idents
+  #> raw_explode
+  #>
+    (if local_name <> satallaxN then
+      (Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())
+        (Scan.finite Symbol.stopper (Scan.repeat1 (parse_line local_name problem) >> flat)))
+         #> fst)
+    else
+      (Scan.error (!! (fn _ => raise UNRECOGNIZED_ATP_PROOF ())
+        (Scan.finite Symbol.stopper (Scan.repeat1 (parse_tstp_thf0_satallax_line))))
+        #> fst
+        #> rev
+        #> map to_satallax_step
+        #> seperate_assumptions_and_steps
+        #> create_satallax_proof_graph
+        #> add_quantifier_in_tree_part [] []
+        #> transform_to_standard_atp_step []
+        #> remove_unused_dependency))
+
+fun atp_proof_of_tstplike_proof _ _ "" = []
+  | atp_proof_of_tstplike_proof local_prover problem tstp =
+    (case core_of_agsyhol_proof tstp of
+      SOME facts => facts |> map (core_inference agsyhol_core_rule)
+    | NONE =>
+      tstp ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
+      |> parse_proof local_prover problem
+      |> local_prover = vampireN ? perhaps (try (sort (vampire_step_name_ord o pairself #1)))
+      |> local_prover = zipperpositionN ? rev)
+
+end;
--- a/src/HOL/Tools/ATP/atp_systems.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_systems.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -397,17 +397,16 @@
 
 (* LEO-II *)
 
-(* LEO-II supports definitions, but it performs significantly better on our
-   benchmarks when they are not used. *)
 val leo2_thf0 = THF (Monomorphic, THF_Without_Choice)
 
 val leo2_config : atp_config =
   {exec = K (["LEO2_HOME"], ["leo.opt", "leo"]),
-   arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ =>
-       "--foatp e --atp e=\"$E_HOME\"/eprover \
-       \--atp epclextract=\"$E_HOME\"/epclextract \
-       \--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^
-       file_name,
+   arguments = fn _ => fn full_proofs => fn _ => fn timeout => fn file_name => fn _ =>
+     "--foatp e --atp e=\"$E_HOME\"/eprover \
+     \--atp epclextract=\"$E_HOME\"/epclextract \
+     \--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^
+     (if full_proofs then "--notReplLeibnizEQ --notReplAndrewsEQ --notUseExtCnfCmbd " else "") ^
+     file_name,
    proof_delims = tstp_proof_delims,
    known_failures =
      [(TimedOut, "CPU time limit exceeded, terminating"),
@@ -431,9 +430,9 @@
 val satallax_config : atp_config =
   {exec = K (["SATALLAX_HOME"], ["satallax.opt", "satallax"]),
    arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ =>
-       "-p hocore -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name,
+       "-p tstp -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name,
    proof_delims =
-     [("% Higher-Order Unsat Core BEGIN", "% Higher-Order Unsat Core END")],
+     [("% SZS output start Proof", "% SZS output end Proof")],
    known_failures = known_szs_status_failures,
    prem_role = Hypothesis,
    best_slices =
--- a/src/HOL/Tools/ATP/atp_util.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_util.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -77,31 +77,29 @@
   | stringN_of_int k n =
     stringN_of_int (k - 1) (n div 10) ^ string_of_int (n mod 10)
 
+fun is_spaceish c = Char.isSpace c orelse c = #"\127" (* DEL -- no idea where these come from *)
+
 fun strip_spaces skip_comments is_evil =
   let
     fun strip_c_style_comment [] accum = accum
-      | strip_c_style_comment (#"*" :: #"/" :: cs) accum =
-        strip_spaces_in_list true cs accum
+      | strip_c_style_comment (#"*" :: #"/" :: cs) accum = strip_spaces_in_list true cs accum
       | strip_c_style_comment (_ :: cs) accum = strip_c_style_comment cs accum
     and strip_spaces_in_list _ [] accum = accum
       | strip_spaces_in_list true (#"%" :: cs) accum =
-        strip_spaces_in_list true (cs |> take_prefix (not_equal #"\n") |> snd)
-                             accum
-      | strip_spaces_in_list true (#"/" :: #"*" :: cs) accum =
-        strip_c_style_comment cs accum
-      | strip_spaces_in_list _ [c1] accum =
-        accum |> not (Char.isSpace c1) ? cons c1
+        strip_spaces_in_list true (cs |> take_prefix (not_equal #"\n") |> snd) accum
+      | strip_spaces_in_list true (#"/" :: #"*" :: cs) accum = strip_c_style_comment cs accum
+      | strip_spaces_in_list _ [c1] accum = accum |> not (is_spaceish c1) ? cons c1
       | strip_spaces_in_list skip_comments (cs as [_, _]) accum =
         accum |> fold (strip_spaces_in_list skip_comments o single) cs
       | strip_spaces_in_list skip_comments (c1 :: c2 :: c3 :: cs) accum =
-        if Char.isSpace c1 then
+        if is_spaceish c1 then
           strip_spaces_in_list skip_comments (c2 :: c3 :: cs) accum
-        else if Char.isSpace c2 then
-          if Char.isSpace c3 then
+        else if is_spaceish c2 then
+          if is_spaceish c3 then
             strip_spaces_in_list skip_comments (c1 :: c3 :: cs) accum
           else
             strip_spaces_in_list skip_comments (c3 :: cs)
-                (c1 :: accum |> forall is_evil [c1, c3] ? cons #" ")
+              (c1 :: accum |> forall is_evil [c1, c3] ? cons #" ")
         else
           strip_spaces_in_list skip_comments (c2 :: c3 :: cs) (cons c1 accum)
   in
--- a/src/HOL/Tools/ATP/atp_waldmeister.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/ATP/atp_waldmeister.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -181,7 +181,7 @@
     val thy = Proof_Context.theory_of ctxt
     val t = u
       |> atp_to_trm
-      |> singleton (infer_formula_types ctxt)
+      |> singleton (infer_formulas_types ctxt)
       |> HOLogic.mk_Trueprop
   in
     (name, role, t, rule, deps)
--- a/src/HOL/Tools/BNF/bnf_comp.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_comp.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -203,7 +203,7 @@
         fun tac {context = ctxt, prems = _} =
           mk_simplified_set_tac ctxt (collect_set_map_of_bnf outer);
         val set'_eq_set =
-          Goal.prove names_lthy [] [] goal tac
+          Goal.prove (*no sorry*) names_lthy [] [] goal tac
           |> Thm.close_derivation;
         val set' = fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of set'_eq_set)));
       in
--- a/src/HOL/Tools/BNF/bnf_def.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_def.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -55,10 +55,12 @@
   val in_mono_of_bnf: bnf -> thm
   val in_rel_of_bnf: bnf -> thm
   val inj_map_of_bnf: bnf -> thm
+  val inj_map_strong_of_bnf: bnf -> thm
   val map_comp0_of_bnf: bnf -> thm
   val map_comp_of_bnf: bnf -> thm
   val map_cong0_of_bnf: bnf -> thm
   val map_cong_of_bnf: bnf -> thm
+  val map_cong_simp_of_bnf: bnf -> thm
   val map_def_of_bnf: bnf -> thm
   val map_id0_of_bnf: bnf -> thm
   val map_id_of_bnf: bnf -> thm
@@ -73,6 +75,7 @@
   val rel_cong_of_bnf: bnf -> thm
   val rel_conversep_of_bnf: bnf -> thm
   val rel_mono_of_bnf: bnf -> thm
+  val rel_mono_strong0_of_bnf: bnf -> thm
   val rel_mono_strong_of_bnf: bnf -> thm
   val rel_eq_of_bnf: bnf -> thm
   val rel_flip_of_bnf: bnf -> thm
@@ -222,8 +225,10 @@
   in_mono: thm lazy,
   in_rel: thm lazy,
   inj_map: thm lazy,
+  inj_map_strong: thm lazy,
   map_comp: thm lazy,
   map_cong: thm lazy,
+  map_cong_simp: thm lazy,
   map_id: thm lazy,
   map_ident0: thm lazy,
   map_ident: thm lazy,
@@ -232,7 +237,9 @@
   rel_flip: thm lazy,
   set_map: thm lazy list,
   rel_cong: thm lazy,
+  rel_map: thm list lazy,
   rel_mono: thm lazy,
+  rel_mono_strong0: thm lazy,
   rel_mono_strong: thm lazy,
   rel_Grp: thm lazy,
   rel_conversep: thm lazy,
@@ -240,8 +247,9 @@
 };
 
 fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel
-    inj_map map_comp map_cong map_id map_ident0 map_ident map_transfer rel_eq rel_flip set_map
-    rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO = {
+    inj_map inj_map_strong map_comp map_cong map_cong_simp map_id map_ident0 map_ident
+    map_transfer rel_eq rel_flip set_map rel_cong rel_map rel_mono rel_mono_strong0 rel_mono_strong
+    rel_Grp rel_conversep rel_OO = {
   bd_Card_order = bd_Card_order,
   bd_Cinfinite = bd_Cinfinite,
   bd_Cnotzero = bd_Cnotzero,
@@ -251,8 +259,10 @@
   in_mono = in_mono,
   in_rel = in_rel,
   inj_map = inj_map,
+  inj_map_strong = inj_map_strong,
   map_comp = map_comp,
   map_cong = map_cong,
+  map_cong_simp = map_cong_simp,
   map_id = map_id,
   map_ident0 = map_ident0,
   map_ident = map_ident,
@@ -261,7 +271,9 @@
   rel_flip = rel_flip,
   set_map = set_map,
   rel_cong = rel_cong,
+  rel_map = rel_map,
   rel_mono = rel_mono,
+  rel_mono_strong0 = rel_mono_strong0,
   rel_mono_strong = rel_mono_strong,
   rel_Grp = rel_Grp,
   rel_conversep = rel_conversep,
@@ -277,8 +289,10 @@
   in_mono,
   in_rel,
   inj_map,
+  inj_map_strong,
   map_comp,
   map_cong,
+  map_cong_simp,
   map_id,
   map_ident0,
   map_ident,
@@ -287,7 +301,9 @@
   rel_flip,
   set_map,
   rel_cong,
+  rel_map,
   rel_mono,
+  rel_mono_strong0,
   rel_mono_strong,
   rel_Grp,
   rel_conversep,
@@ -301,8 +317,10 @@
     in_mono = Lazy.map f in_mono,
     in_rel = Lazy.map f in_rel,
     inj_map = Lazy.map f inj_map,
+    inj_map_strong = Lazy.map f inj_map_strong,
     map_comp = Lazy.map f map_comp,
     map_cong = Lazy.map f map_cong,
+    map_cong_simp = Lazy.map f map_cong_simp,
     map_id = Lazy.map f map_id,
     map_ident0 = Lazy.map f map_ident0,
     map_ident = Lazy.map f map_ident,
@@ -311,7 +329,9 @@
     rel_flip = Lazy.map f rel_flip,
     set_map = map (Lazy.map f) set_map,
     rel_cong = Lazy.map f rel_cong,
+    rel_map = Lazy.map (map f) rel_map,
     rel_mono = Lazy.map f rel_mono,
+    rel_mono_strong0 = Lazy.map f rel_mono_strong0,
     rel_mono_strong = Lazy.map f rel_mono_strong,
     rel_Grp = Lazy.map f rel_Grp,
     rel_conversep = Lazy.map f rel_conversep,
@@ -421,6 +441,7 @@
 val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
 val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf;
 val inj_map_of_bnf = Lazy.force o #inj_map o #facts o rep_bnf;
+val inj_map_strong_of_bnf = Lazy.force o #inj_map_strong o #facts o rep_bnf;
 val map_def_of_bnf = #map_def o #defs o rep_bnf;
 val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf;
 val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf;
@@ -430,6 +451,7 @@
 val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf;
 val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf;
 val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf;
+val map_cong_simp_of_bnf = Lazy.force o #map_cong_simp o #facts o rep_bnf;
 val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf;
 val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf;
 val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
@@ -441,6 +463,7 @@
 val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf;
 val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
 val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
+val rel_mono_strong0_of_bnf = Lazy.force o #rel_mono_strong0 o #facts o rep_bnf;
 val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf;
 val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf;
 val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf;
@@ -594,6 +617,7 @@
 val in_monoN = "in_mono";
 val in_relN = "in_rel";
 val inj_mapN = "inj_map";
+val inj_map_strongN = "inj_map_strong";
 val map_id0N = "map_id0";
 val map_idN = "map_id";
 val map_identN = "map_ident";
@@ -601,6 +625,7 @@
 val map_compN = "map_comp";
 val map_cong0N = "map_cong0";
 val map_congN = "map_cong";
+val map_cong_simpN = "map_cong_simp";
 val map_transferN = "map_transfer";
 val rel_eqN = "rel_eq";
 val rel_flipN = "rel_flip";
@@ -609,7 +634,9 @@
 val set_bdN = "set_bd";
 val rel_GrpN = "rel_Grp";
 val rel_conversepN = "rel_conversep";
+val rel_mapN = "rel_map"
 val rel_monoN = "rel_mono"
+val rel_mono_strong0N = "rel_mono_strong0"
 val rel_mono_strongN = "rel_mono_strong"
 val rel_comppN = "rel_compp";
 val rel_compp_GrpN = "rel_compp_Grp";
@@ -647,9 +674,9 @@
            (in_bdN, [Lazy.force (#in_bd facts)]),
            (in_monoN, [Lazy.force (#in_mono facts)]),
            (in_relN, [Lazy.force (#in_rel facts)]),
-           (inj_mapN, [Lazy.force (#inj_map facts)]),
            (map_comp0N, [#map_comp0 axioms]),
            (map_transferN, [Lazy.force (#map_transfer facts)]),
+           (rel_mono_strong0N, [Lazy.force (#rel_mono_strong0 facts)]),
            (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)]),
            (set_map0N, #set_map0 axioms),
            (set_bdN, #set_bd axioms)] @
@@ -664,9 +691,12 @@
     fun note_unless_dont_note (noted0, lthy0) =
       let
         val notes =
-          [(map_compN, [Lazy.force (#map_comp facts)], []),
+          [(inj_mapN, [Lazy.force (#inj_map facts)], []),
+           (inj_map_strongN, [Lazy.force (#inj_map_strong facts)], []),
+           (map_compN, [Lazy.force (#map_comp facts)], []),
            (map_cong0N, [#map_cong0 axioms], []),
            (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs),
+           (map_cong_simpN, [Lazy.force (#map_cong_simp facts)], []),
            (map_idN, [Lazy.force (#map_id facts)], []),
            (map_id0N, [#map_id0 axioms], []),
            (map_identN, [Lazy.force (#map_ident facts)], []),
@@ -676,6 +706,7 @@
            (rel_eqN, [Lazy.force (#rel_eq facts)], []),
            (rel_flipN, [Lazy.force (#rel_flip facts)], []),
            (rel_GrpN, [Lazy.force (#rel_Grp facts)], []),
+           (rel_mapN, Lazy.force (#rel_map facts), []),
            (rel_monoN, [Lazy.force (#rel_mono facts)], []),
            (set_mapN, map Lazy.force (#set_map facts), [])]
           |> filter_out (null o #2)
@@ -920,6 +951,7 @@
     val pred2RTs = map2 mk_pred2T As' Bs';
     val pred2RTsAsCs = map2 mk_pred2T As' Cs;
     val pred2RTsBsCs = map2 mk_pred2T Bs' Cs;
+    val pred2RTsCsBs = map2 mk_pred2T Cs Bs';
     val pred2RT's = map2 mk_pred2T Bs' As';
     val self_pred2RTs = map2 mk_pred2T As' As';
     val transfer_domRTs = map2 mk_pred2T As' B1Ts;
@@ -941,15 +973,19 @@
     fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel;
 
     val pre_names_lthy = lthy;
-    val (((((((((((((((fs, gs), hs), x), y), zs), ys), As),
-      As_copy), bs), Rs), Rs_copy), Ss),
+    val (((((((((((((((((((((fs, fs'), gs), hs), is), x), x'), y), zs), zs'), ys), As),
+      As_copy), bs), Rs), Rs_copy), Ss), S_AsCs), S_CsBs),
       transfer_domRs), transfer_ranRs), names_lthy) = pre_names_lthy
       |> mk_Frees "f" (map2 (curry op -->) As' Bs')
+      ||>> mk_Frees "f" (map2 (curry op -->) As' Bs')
       ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs)
       ||>> mk_Frees "h" (map2 (curry op -->) As' Ts)
+      ||>> mk_Frees "i" (map2 (curry op -->) As' Cs)
+      ||>> yield_singleton (mk_Frees "x") CA'
       ||>> yield_singleton (mk_Frees "x") CA'
       ||>> yield_singleton (mk_Frees "y") CB'
       ||>> mk_Frees "z" As'
+      ||>> mk_Frees "z" As'
       ||>> mk_Frees "y" Bs'
       ||>> mk_Frees "A" (map HOLogic.mk_setT As')
       ||>> mk_Frees "A" (map HOLogic.mk_setT As')
@@ -957,6 +993,8 @@
       ||>> mk_Frees "R" pred2RTs
       ||>> mk_Frees "R" pred2RTs
       ||>> mk_Frees "S" pred2RTsBsCs
+      ||>> mk_Frees "S" pred2RTsAsCs
+      ||>> mk_Frees "S" pred2RTsCsBs
       ||>> mk_Frees "R" transfer_domRTs
       ||>> mk_Frees "S" transfer_ranRTs;
 
@@ -981,13 +1019,13 @@
         fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
       end;
 
-    fun mk_map_cong_prem x z set f f_copy =
-      Logic.all z (Logic.mk_implies
+    fun mk_map_cong_prem mk_implies x z set f f_copy =
+      Logic.all z (mk_implies
         (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (f $ z, f_copy $ z)));
 
     val map_cong0_goal =
       let
-        val prems = map4 (mk_map_cong_prem x) zs bnf_sets_As fs fs_copy;
+        val prems = map4 (mk_map_cong_prem Logic.mk_implies x) zs bnf_sets_As fs fs_copy;
         val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
           Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
       in
@@ -1022,6 +1060,7 @@
 
     val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC';
     val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC';
+    val relCsBs = mk_bnf_rel pred2RTsCsBs CC' CB';
     val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss);
     val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss));
     val le_rel_OO_goal =
@@ -1116,20 +1155,23 @@
         val map_ident = Lazy.lazy (fn () => mk_map_ident lthy (Lazy.force map_id));
         val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms));
 
-        fun mk_map_cong () =
+        fun mk_map_cong mk_implies () =
           let
             val prem0 = mk_Trueprop_eq (x, x_copy);
-            val prems = map4 (mk_map_cong_prem x_copy) zs bnf_sets_As fs fs_copy;
+            val prems = map4 (mk_map_cong_prem mk_implies x_copy) zs bnf_sets_As fs fs_copy;
             val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
               Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy);
             val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy)
               (Logic.list_implies (prem0 :: prems, eq));
           in
-            Goal.prove_sorry lthy [] [] goal (fn _ => mk_map_cong_tac lthy (#map_cong0 axioms))
+            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
+              unfold_thms_tac lthy @{thms simp_implies_def} THEN
+              mk_map_cong_tac lthy (#map_cong0 axioms))
             |> Thm.close_derivation
           end;
 
-        val map_cong = Lazy.lazy mk_map_cong;
+        val map_cong = Lazy.lazy (mk_map_cong Logic.mk_implies);
+        val map_cong_simp = Lazy.lazy (mk_map_cong (fn (a, b) => @{term simp_implies} $ a $ b));
 
         fun mk_inj_map () =
           let
@@ -1288,7 +1330,7 @@
 
         val rel_flip = Lazy.lazy mk_rel_flip;
 
-        fun mk_rel_mono_strong () =
+        fun mk_rel_mono_strong0 () =
           let
             fun mk_prem setA setB R S a b =
               HOLogic.mk_Trueprop
@@ -1301,13 +1343,44 @@
           in
             Goal.prove_sorry lthy [] []
               (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl)))
-              (fn {context = ctxt, prems = _} => mk_rel_mono_strong_tac ctxt (Lazy.force in_rel)
+              (fn {context = ctxt, prems = _} => mk_rel_mono_strong0_tac ctxt (Lazy.force in_rel)
                 (map Lazy.force set_map))
             |> Thm.close_derivation
           end;
 
+        val rel_mono_strong0 = Lazy.lazy mk_rel_mono_strong0;
+
+        fun mk_rel_mono_strong () = Object_Logic.rulify lthy (Lazy.force rel_mono_strong0)
+
         val rel_mono_strong = Lazy.lazy mk_rel_mono_strong;
 
+        fun mk_rel_map () =
+          let
+            fun mk_goal lhs rhs =
+              fold_rev Logic.all ([x, y] @ S_CsBs @ S_AsCs @ is @ gs) (mk_Trueprop_eq (lhs, rhs));
+
+            val lhss =
+              [Term.list_comb (relCsBs, S_CsBs) $ (Term.list_comb (bnf_map_AsCs, is) $ x) $ y,
+               Term.list_comb (relAsCs, S_AsCs) $ x $ (Term.list_comb (bnf_map_BsCs, gs) $ y)];
+            val rhss =
+              [Term.list_comb (rel, map3 (fn f => fn P => fn T =>
+                 mk_vimage2p f (HOLogic.id_const T) $ P) is S_CsBs Bs') $ x $ y,
+               Term.list_comb (rel, map3 (fn f => fn P => fn T =>
+                 mk_vimage2p (HOLogic.id_const T) f $ P) gs S_AsCs As') $ x $ y];
+            val goals = map2 mk_goal lhss rhss;
+          in
+            goals
+            |> map (fn goal => Goal.prove_sorry lthy [] [] goal
+              (fn {context = ctxt, prems = _} =>
+                 mk_rel_map0_tac ctxt live (Lazy.force rel_OO) (Lazy.force rel_conversep)
+                  (Lazy.force rel_Grp) (Lazy.force map_id)))
+            |> map (unfold_thms lthy @{thms vimage2p_def[of id, unfolded id_apply]
+                 vimage2p_def[of _ id, unfolded id_apply]})
+            |> map Thm.close_derivation
+          end;
+
+        val rel_map = Lazy.lazy mk_rel_map;
+
         fun mk_map_transfer () =
           let
             val rels = map2 mk_rel_fun transfer_domRs transfer_ranRs;
@@ -1327,11 +1400,36 @@
 
         val map_transfer = Lazy.lazy mk_map_transfer;
 
+        fun mk_inj_map_strong () =
+          let
+            val assms = map5 (fn setA => fn z => fn f => fn z' => fn f' =>
+              fold_rev Logic.all [z, z']
+                (Logic.mk_implies (mk_Trueprop_mem (z, setA $ x),
+                   Logic.mk_implies (mk_Trueprop_mem (z', setA $ x'),
+                     Logic.mk_implies (mk_Trueprop_eq (f $ z, f' $ z'),
+                       mk_Trueprop_eq (z, z')))))) bnf_sets_As zs fs zs' fs';
+            val concl = Logic.mk_implies
+              (mk_Trueprop_eq
+                 (Term.list_comb (bnf_map_AsBs, fs) $ x,
+                  Term.list_comb (bnf_map_AsBs, fs') $ x'),
+               mk_Trueprop_eq (x, x'));
+            val goal = fold_rev Logic.all (x :: x' :: fs @ fs')
+              (fold_rev (curry Logic.mk_implies) assms concl);
+          in
+            Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} =>
+              mk_inj_map_strong_tac ctxt (Lazy.force rel_eq) (Lazy.force rel_map)
+                (Lazy.force rel_mono_strong))
+            |> Thm.close_derivation
+          end;
+
+        val inj_map_strong = Lazy.lazy mk_inj_map_strong;
+
         val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def;
 
         val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong
-          in_mono in_rel inj_map map_comp map_cong map_id map_ident0 map_ident map_transfer rel_eq
-          rel_flip set_map rel_cong rel_mono rel_mono_strong rel_Grp rel_conversep rel_OO;
+          in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_id map_ident0
+          map_ident map_transfer rel_eq rel_flip set_map rel_cong rel_map rel_mono rel_mono_strong0
+          rel_mono_strong rel_Grp rel_conversep rel_OO;
 
         val wits = map2 mk_witness bnf_wits wit_thms;
 
--- a/src/HOL/Tools/BNF/bnf_def_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_def_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -11,6 +11,7 @@
 sig
   val mk_collect_set_map_tac: thm list -> tactic
   val mk_in_mono_tac: int -> tactic
+  val mk_inj_map_strong_tac: Proof.context -> thm -> thm list -> thm -> tactic
   val mk_inj_map_tac: int -> thm -> thm -> thm -> thm -> tactic
   val mk_map_id: thm -> thm
   val mk_map_ident: Proof.context -> thm -> thm
@@ -23,8 +24,9 @@
   val mk_rel_OO_le_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm list -> tactic
   val mk_rel_conversep_tac: thm -> thm -> tactic
   val mk_rel_conversep_le_tac: Proof.context -> thm list -> thm -> thm -> thm -> thm list -> tactic
+  val mk_rel_map0_tac: Proof.context -> int -> thm -> thm -> thm -> thm -> tactic
   val mk_rel_mono_tac: thm list -> thm -> tactic
-  val mk_rel_mono_strong_tac: Proof.context -> thm -> thm list -> tactic
+  val mk_rel_mono_strong0_tac: Proof.context -> thm -> thm list -> tactic
 
   val mk_map_transfer_tac: Proof.context -> thm -> thm -> thm list -> thm -> thm -> tactic
 
@@ -69,6 +71,17 @@
         REPEAT_DETERM_N n o atac))
   end;
 
+fun mk_inj_map_strong_tac ctxt rel_eq rel_maps rel_mono_strong =
+  let
+    val rel_eq' = rel_eq RS @{thm predicate2_eqD};
+    val rel_maps' = map (fn thm => thm RS iffD1) rel_maps;
+  in
+    HEADGOAL (dtac (rel_eq' RS iffD2) THEN' rtac (rel_eq' RS iffD1)) THEN
+    EVERY (map (HEADGOAL o dtac) rel_maps') THEN
+    HEADGOAL (etac rel_mono_strong) THEN
+    TRYALL (Goal.assume_rule_tac ctxt)
+  end;
+
 fun mk_collect_set_map_tac set_map0s =
   (rtac (@{thm collect_comp} RS trans) THEN' rtac @{thm arg_cong[of _ _ collect]} THEN'
   EVERY' (map (fn set_map0 =>
@@ -118,6 +131,26 @@
     rtac @{thm equalityI}, rtac subset_UNIV, rtac subsetI, rtac CollectI,
     CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac map_id0])) 1;
 
+fun mk_rel_map0_tac ctxt live rel_compp rel_conversep rel_Grp map_id =
+  if live = 0 then
+    HEADGOAL (Goal.conjunction_tac) THEN
+    unfold_thms_tac ctxt @{thms id_apply} THEN
+    ALLGOALS (rtac refl)
+  else
+    let
+      val ks = 1 upto live;
+    in
+      Goal.conjunction_tac 1 THEN
+      unfold_thms_tac ctxt [rel_compp, rel_conversep, rel_Grp, @{thm vimage2p_Grp}] THEN
+      TRYALL (EVERY' [rtac iffI, rtac @{thm relcomppI}, rtac @{thm GrpI},
+        resolve_tac [map_id, refl], rtac CollectI,
+        CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks, rtac @{thm relcomppI}, atac,
+        rtac @{thm conversepI}, rtac @{thm GrpI}, resolve_tac [map_id, refl], rtac CollectI,
+        CONJ_WRAP' (K (rtac @{thm subset_UNIV})) ks,
+        REPEAT_DETERM o eresolve_tac @{thms relcomppE conversepE GrpE},
+        dtac (trans OF [sym, map_id]), hyp_subst_tac ctxt, atac])
+    end;
+
 fun mk_rel_mono_tac rel_OO_Grps in_mono =
   let
     val rel_OO_Grps_tac = if null rel_OO_Grps then K all_tac
@@ -176,7 +209,7 @@
         REPEAT_DETERM_N n o rtac @{thm fst_fstOp},
         in_tac @{thm fstOp_in},
         rtac @{thm GrpI}, rtac trans, rtac map_comp, rtac map_cong0,
-        REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, 
+        REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply,
           rtac ballE, rtac subst,
           rtac @{thm csquare_def}, rtac @{thm csquare_fstOp_sndOp}, atac, etac notE,
           etac set_mp, atac],
@@ -190,7 +223,7 @@
         in_tac @{thm sndOp_in}] 1
   end;
 
-fun mk_rel_mono_strong_tac ctxt in_rel set_maps =
+fun mk_rel_mono_strong0_tac ctxt in_rel set_maps =
   if null set_maps then atac 1
   else
     unfold_tac ctxt [in_rel] THEN
@@ -231,10 +264,10 @@
     let
       val bd'_Cinfinite = bd_Cinfinite RS @{thm Cinfinite_csum1};
       val inserts =
-        map (fn set_bd => 
+        map (fn set_bd =>
           iffD2 OF [@{thm card_of_ordLeq}, @{thm ordLeq_ordIso_trans} OF
             [set_bd, bd_Card_order RS @{thm card_of_Field_ordIso} RS @{thm ordIso_symmetric}]])
-        set_bds;        
+        set_bds;
     in
       EVERY' [rtac (Drule.rotate_prems 1 ctrans), rtac @{thm cprod_cinfinite_bound},
         rtac (ctrans OF @{thms ordLeq_csum2 ordLeq_cexp2}), rtac @{thm card_of_Card_order},
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -21,17 +21,17 @@
   val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list
 
   type lfp_sugar_thms =
-    (thm list * thm * Args.src list) * (thm list list * Args.src list)
+    (thm list * thm * Token.src list) * (thm list list * Token.src list)
 
   val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms
   val transfer_lfp_sugar_thms: Proof.context -> lfp_sugar_thms -> lfp_sugar_thms
 
   type gfp_sugar_thms =
-    ((thm list * thm) list * (Args.src list * Args.src list))
+    ((thm list * thm) list * (Token.src list * Token.src list))
     * thm list list
     * thm list list
-    * (thm list list * Args.src list)
-    * (thm list list list * Args.src list)
+    * (thm list list * Token.src list)
+    * (thm list list list * Token.src list)
 
   val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms
   val transfer_gfp_sugar_thms: Proof.context -> gfp_sugar_thms -> gfp_sugar_thms
@@ -98,11 +98,14 @@
 
 val EqN = "Eq_";
 
+val ctr_transferN = "ctr_transfer";
 val corec_codeN = "corec_code";
-val disc_map_iffN = "disc_map_iff";
-val sel_mapN = "sel_map";
-val sel_setN = "sel_set";
-val set_emptyN = "set_empty";
+val map_disc_iffN = "map_disc_iff";
+val map_selN = "map_sel";
+val set_casesN = "set_cases";
+val set_introsN = "set_intros";
+val set_inductN = "set_induct";
+val set_selN = "set_sel";
 
 structure Data = Generic_Data
 (
@@ -112,10 +115,34 @@
   fun merge data : T = Symtab.merge (K true) data;
 );
 
-fun choose_relator Rs AB = find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) Rs;
-fun build_the_rel ctxt Rs Ts A B = build_rel ctxt Ts (the o choose_relator Rs) (A, B);
+fun zipping_map f =
+  let
+    fun zmap _ [] = []
+      | zmap xs (y :: ys) = f (xs, y, ys) :: zmap (xs @ [y]) ys
+  in zmap [] end;
+
+fun choose_binary_fun fs AB =
+  find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs;
+fun build_binary_fun_app fs a b =
+  Option.map (rapp b o rapp a) (choose_binary_fun fs (fastype_of a, fastype_of b));
+
+fun build_the_rel ctxt Rs Ts A B = build_rel ctxt Ts (the o choose_binary_fun Rs) (A, B);
 fun build_rel_app ctxt Rs Ts a b = build_the_rel ctxt Rs Ts (fastype_of a) (fastype_of b) $ a $ b;
 
+val name_of_set = name_of_const "set";
+
+
+fun mk_parametricity_goals ctxt Rs fs gs =
+  let
+    val prems =
+      map (foldr1 (uncurry mk_rel_fun) o
+          uncurry (map2 (build_the_rel ctxt Rs [])) o
+          pairself (fastype_of #> strip_type #> (fn (Ts, T) => Ts @ [T])))
+        (fs ~~ gs);
+  in
+    map3 (fn prem => fn f => fn g => HOLogic.mk_Trueprop (prem $ f $ g)) prems fs gs
+  end
+
 fun fp_sugar_of ctxt =
   Symtab.lookup (Data.get (Context.Proof ctxt))
   #> Option.map (transfer_fp_sugar ctxt);
@@ -276,7 +303,7 @@
 fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p;
 
 type lfp_sugar_thms =
-  (thm list * thm * Args.src list) * (thm list list * Args.src list);
+  (thm list * thm * Token.src list) * (thm list list * Token.src list);
 
 fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) =
   ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs),
@@ -287,21 +314,21 @@
   morph_lfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
 
 type gfp_sugar_thms =
-  ((thm list * thm) list * (Args.src list * Args.src list))
+  ((thm list * thm) list * (Token.src list * Token.src list))
   * thm list list
   * thm list list
-  * (thm list list * Args.src list)
-  * (thm list list list * Args.src list);
+  * (thm list list * Token.src list)
+  * (thm list list list * Token.src list);
 
 fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs_pair),
-    corecss, disc_corecss, (disc_corec_iffss, disc_corec_iff_attrs),
-    (sel_corecsss, sel_corec_attrs)) =
+    corecss, corec_discss, (corec_disc_iffss, corec_disc_iff_attrs),
+    (corec_selsss, corec_sel_attrs)) =
   ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs,
     coinduct_attrs_pair),
    map (map (Morphism.thm phi)) corecss,
-   map (map (Morphism.thm phi)) disc_corecss,
-   (map (map (Morphism.thm phi)) disc_corec_iffss, disc_corec_iff_attrs),
-   (map (map (map (Morphism.thm phi))) sel_corecsss, sel_corec_attrs)) : gfp_sugar_thms;
+   map (map (Morphism.thm phi)) corec_discss,
+   (map (map (Morphism.thm phi)) corec_disc_iffss, corec_disc_iff_attrs),
+   (map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms;
 
 val transfer_gfp_sugar_thms =
   morph_gfp_sugar_thms o Morphism.transfer_morphism o Proof_Context.theory_of;
@@ -456,7 +483,7 @@
   #>> map (fn thm => Thm.permute_prems 0 (~1) (thm RS prop))
   ##> (fn thm => Thm.permute_prems 0 (~nn)
     (if nn = 1 then thm RS prop
-     else funpow nn (fn thm => Local_Defs.unfold lthy @{thms conj_assoc} (thm RS prop_conj)) thm));
+     else funpow nn (fn thm => unfold_thms lthy @{thms conj_assoc} (thm RS prop_conj)) thm));
 
 fun mk_induct_attrs ctrss =
   let
@@ -494,10 +521,10 @@
     val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq
       (map2 (build_the_rel lthy (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs));
 
-    val rel_induct0_thm = Goal.prove_sorry lthy [] premises goal
-      (fn {context = ctxt, prems = prems} =>
-          mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (certify ctxt) IRs) exhausts ctor_defss
-            ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs)
+    val rel_induct0_thm =
+      Goal.prove_sorry lthy [] premises goal (fn {context = ctxt, prems} =>
+        mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (certify ctxt) IRs) exhausts ctor_defss
+          ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs)
       |> singleton (Proof_Context.export names_lthy lthy)
       |> Thm.close_derivation;
   in
@@ -745,12 +772,12 @@
     val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq
       IRs (map2 (build_the_rel lthy (Rs @ IRs) []) fpA_Ts fpB_Ts)));
 
-    val rel_coinduct0_thm = Goal.prove_sorry lthy [] premises goal
-      (fn {context = ctxt, prems = prems} =>
-          mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (certify ctxt) IRs) prems
-            (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars)
-            (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects
-            rel_pre_defs abs_inverses live_nesting_rel_eqs)
+    val rel_coinduct0_thm =
+      Goal.prove_sorry lthy [] premises goal (fn {context = ctxt, prems} =>
+        mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (certify ctxt) IRs) prems
+          (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars)
+          (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects
+          rel_pre_defs abs_inverses live_nesting_rel_eqs)
       |> singleton (Proof_Context.export names_lthy lthy)
       |> Thm.close_derivation;
   in
@@ -759,6 +786,99 @@
      mk_coinduct_attributes fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss)
   end;
 
+fun derive_set_induct_thms_for_types lthy nn fpTs ctrss setss dtor_set_inducts exhausts
+    set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses =
+  let
+    fun mk_prems A Ps ctr_args t ctxt =
+      (case fastype_of t of
+        Type (type_name, innerTs) =>
+        (case bnf_of ctxt type_name of
+          NONE => ([], ctxt)
+        | SOME bnf =>
+          let
+            fun seq_assm a set ctxt =
+              let
+                val X = HOLogic.dest_setT (range_type (fastype_of set));
+                val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt;
+                val assm = mk_Trueprop_mem (x, set $ a);
+              in
+                (case build_binary_fun_app Ps x a of
+                  NONE =>
+                  mk_prems A Ps ctr_args x ctxt'
+                  |>> map (Logic.all x o Logic.mk_implies o pair assm)
+                | SOME f =>
+                  ([Logic.all x
+                      (Logic.mk_implies (assm,
+                         Logic.mk_implies (HOLogic.mk_Trueprop f,
+                           HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))],
+                   ctxt'))
+              end;
+          in
+            fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt
+            |>> flat
+          end)
+      | T =>
+        if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt)
+        else ([], ctxt));
+
+    fun mk_prems_for_ctr A Ps ctr ctxt =
+      let
+        val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt;
+      in
+        fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt'
+        |>> map (fold_rev Logic.all args) o flat
+        |>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr)))
+      end;
+
+    fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt =
+      let
+        val ((x, fp), ctxt') = ctxt
+          |> yield_singleton (mk_Frees "x") A
+          ||>> yield_singleton (mk_Frees "a") fpT;
+        val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x)
+          (the (build_binary_fun_app Ps x fp)));
+      in
+        fold_map (mk_prems_for_ctr A Ps) ctrs ctxt'
+        |>> split_list
+        |>> map_prod flat flat
+        |>> apfst (rpair concl)
+      end;
+
+    fun mk_thm ctxt fpTs ctrss sets =
+      let
+        val A = HOLogic.dest_setT (range_type (fastype_of (hd sets)));
+        val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt;
+        val (((premises, conclusion), case_names), ctxt'') =
+          (fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt')
+          |>> apfst split_list o split_list
+          |>> apfst (apfst flat)
+          |>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj))
+          |>> apsnd flat;
+
+        val thm =
+          Goal.prove_sorry lthy [] premises (HOLogic.mk_Trueprop conclusion)
+            (fn {context = ctxt, prems} =>
+               mk_set_induct0_tac ctxt (map (certify ctxt'') Ps) prems dtor_set_inducts exhausts
+                 set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses)
+          |> singleton (Proof_Context.export ctxt'' ctxt)
+          |> Thm.close_derivation;
+
+        val case_names_attr =
+          Attrib.internal (K (Rule_Cases.case_names (quasi_unambiguous_case_names case_names)));
+        val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets;
+      in
+        (thm, case_names_attr :: induct_set_attrs)
+      end
+    val consumes_attr = Attrib.internal (K (Rule_Cases.consumes 1));
+  in
+    map (fn Asets =>
+      let
+        fun massage_thm thm = rotate_prems (~1) (thm RS bspec);
+      in
+        mk_thm lthy fpTs ctrss Asets |> nn = 1 ? map_prod massage_thm (cons consumes_attr)
+      end) (transpose setss)
+  end;
+
 fun derive_coinduct_corecs_thms_for_types pre_bnfs (x, cs, cpss, ((pgss, cqgsss), _))
     dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss
     kss mss ns fp_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list)
@@ -853,7 +973,7 @@
         val rel_eqs = map rel_eq_of_bnf pre_bnfs;
         val rel_monos = map rel_mono_of_bnf pre_bnfs;
         val dtor_coinducts =
-          [dtor_coinduct, mk_strong_coinduct_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p lthy]
+          [dtor_coinduct, mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p lthy]
       in
         map2 (postproc_co_induct lthy nn mp mp_conj oo prove) dtor_coinducts goals
       end;
@@ -902,7 +1022,7 @@
         |> map (map (unfold_thms lthy @{thms case_sum_if}))
       end;
 
-    val disc_corec_iff_thmss =
+    val corec_disc_iff_thmss =
       let
         fun mk_goal c cps gcorec n k disc =
           mk_Trueprop_eq (disc $ (gcorec $ c),
@@ -915,7 +1035,7 @@
 
         val case_splitss' = map (map mk_case_split') cpss;
 
-        val tacss = map3 (map oo mk_disc_corec_iff_tac) case_splitss' corec_thmss disc_thmsss;
+        val tacss = map3 (map oo mk_corec_disc_iff_tac) case_splitss' corec_thmss disc_thmsss;
 
         fun prove goal tac =
           Goal.prove_sorry lthy [] [] goal (tac o #context)
@@ -929,11 +1049,11 @@
         map2 proves goalss tacss
       end;
 
-    fun mk_disc_corec_thms corecs discIs = map (op RS) (corecs ~~ discIs);
+    fun mk_corec_disc_thms corecs discIs = map (op RS) (corecs ~~ discIs);
 
-    val disc_corec_thmss = map2 mk_disc_corec_thms corec_thmss discIss;
+    val corec_disc_thmss = map2 mk_corec_disc_thms corec_thmss discIss;
 
-    fun mk_sel_corec_thm corec_thm sel sel_thm =
+    fun mk_corec_sel_thm corec_thm sel sel_thm =
       let
         val (domT, ranT) = dest_funT (fastype_of sel);
         val arg_cong' =
@@ -945,17 +1065,17 @@
         corec_thm RS arg_cong' RS sel_thm'
       end;
 
-    fun mk_sel_corec_thms corec_thmss =
-      map3 (map3 (map2 o mk_sel_corec_thm)) corec_thmss selsss sel_thmsss;
+    fun mk_corec_sel_thms corec_thmss =
+      map3 (map3 (map2 o mk_corec_sel_thm)) corec_thmss selsss sel_thmsss;
 
-    val sel_corec_thmsss = mk_sel_corec_thms corec_thmss;
+    val corec_sel_thmsss = mk_corec_sel_thms corec_thmss;
   in
     ((coinduct_thms_pairs,
       mk_coinduct_attributes fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss),
      corec_thmss,
-     disc_corec_thmss,
-     (disc_corec_iff_thmss, simp_attrs),
-     (sel_corec_thmsss, simp_attrs))
+     corec_disc_thmss,
+     (corec_disc_iff_thmss, simp_attrs),
+     (corec_sel_thmsss, simp_attrs))
   end;
 
 fun define_co_datatypes prepare_constraint prepare_typ prepare_term fp construct_fp
@@ -1072,7 +1192,7 @@
     val ((pre_bnfs, absT_infos), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0,
              dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors,
              ctor_dtors, ctor_injects, dtor_injects, xtor_map_thms, xtor_set_thmss, xtor_rel_thms,
-             xtor_co_rec_thms, rel_xtor_co_induct_thm, ...},
+             xtor_co_rec_thms, rel_xtor_co_induct_thm, dtor_set_induct_thms, ...},
            lthy)) =
       fp_bnf (construct_fp mixfixes map_bs rel_bs set_bss) fp_bs (map dest_TFree unsorted_As)
         (map dest_TFree killed_As) fp_eqs no_defs_lthy0
@@ -1143,8 +1263,8 @@
 
     fun massage_simple_notes base =
       filter_out (null o #2)
-      #> map (fn (thmN, thms, attrs) =>
-        ((Binding.qualify true base (Binding.name thmN), attrs), [(thms, [])]));
+      #> map (fn (thmN, thms, f_attrs) =>
+        ((Binding.qualify true base (Binding.name thmN), []), map_index (fn (i, thm) => ([thm], f_attrs i)) thms));
 
     val massage_multi_notes =
       maps (fn (thmN, thmss, attrs) =>
@@ -1168,6 +1288,7 @@
         disc_bindings), sel_bindingss), raw_sel_default_eqs) no_defs_lthy =
       let
         val fp_b_name = Binding.name_of fp_b;
+        val fpBT = B_ify fpT;
 
         val ctr_absT = domain_type (fastype_of ctor);
 
@@ -1290,55 +1411,22 @@
                          else fp_map_thm RS ctor_cong RS (ctor_dtor RS sym RS trans))))
                 |> singleton (Proof_Context.export names_lthy no_defs_lthy);
 
-              fun mk_set_thm fp_set_thm ctr_def' cxIn =
+              fun mk_set0_thm fp_set_thm ctr_def' cxIn =
                 fold_thms lthy [ctr_def']
                   (unfold_thms lthy (pre_set_defs @ fp_nesting_set_maps @ live_nesting_set_maps @
-                       (if fp = Least_FP then [] else [dtor_ctor]) @ sumprod_thms_set @
-                       abs_inverses)
+                       (if fp = Least_FP then [] else [dtor_ctor]) @ basic_sumprod_thms_set @
+                       @{thms UN_Un sup_assoc[THEN sym]} @ abs_inverses)
                      (cterm_instantiate_pos [SOME cxIn] fp_set_thm))
                 |> singleton (Proof_Context.export names_lthy no_defs_lthy);
 
-              fun mk_set_thms fp_set_thm = map2 (mk_set_thm fp_set_thm) ctr_defs' cxIns;
+              fun mk_set0_thms fp_set_thm = map2 (mk_set0_thm fp_set_thm) ctr_defs' cxIns;
 
               val map_thms = map2 mk_map_thm ctr_defs' cxIns;
-              val set_thmss = map mk_set_thms fp_set_thms;
-              val set_thms = flat set_thmss;
-
-              val sets = map (mk_set (snd (Term.dest_Type fpT))) (sets_of_bnf fp_bnf);
-
-              val set_empty_thms =
-                let
-                  val ctr_argT_namess = map ((fn Ts => fold Term.add_tfree_namesT Ts []) o
-                    binder_types o fastype_of) ctrs;
-                  val setTs = map (HOLogic.dest_setT o range_type o fastype_of) sets;
-                  val setT_names = map (fn T => the_single (Term.add_tfree_namesT T [])) setTs;
-
-                  fun mk_set_empty_goal disc set T =
-                    Logic.mk_implies (HOLogic.mk_Trueprop (disc $ u),
-                      mk_Trueprop_eq (set $ u, HOLogic.mk_set T []));
-
-                  val goals =
-                    if null discs then
-                      []
-                    else
-                      map_filter I (flat
-                        (map2 (fn names => fn disc =>
-                          map3 (fn name => fn setT => fn set =>
-                            if member (op =) names name then NONE
-                            else SOME (mk_set_empty_goal disc set setT))
-                          setT_names setTs sets)
-                        ctr_argT_namess discs));
-                in
-                  if null goals then
-                    []
-                  else
-                    Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
-                      (fn {context = ctxt, prems = _} =>
-                        mk_set_empty_tac ctxt (certify ctxt u) exhaust set_thms (flat disc_thmss))
-                      |> Conjunction.elim_balanced (length goals)
-                      |> Proof_Context.export names_lthy lthy
-                      |> map Thm.close_derivation
-                end;
+              val set0_thmss = map mk_set0_thms fp_set_thms;
+              val set0_thms = flat set0_thmss;
+              val set_thms = set0_thms
+                |> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left
+                  Un_insert_left});
 
               val rel_infos = (ctr_defs' ~~ cxIns, ctr_defs' ~~ cyIns);
 
@@ -1386,31 +1474,153 @@
                 map2 (fn th => fn 0 => th RS @{thm eq_True[THEN iffD2]} | _ => th)
                   rel_inject_thms ms;
 
-              val (disc_map_iff_thms, sel_map_thms, sel_set_thms, rel_sel_thms,
-                (rel_cases_thm, rel_cases_attrs)) =
+              val (map_disc_iff_thms, map_sel_thms, set_sel_thms, rel_sel_thms, set_intros_thms,
+                   ctr_transfer_thms, (set_cases_thms, set_cases_attrss),
+                   (rel_cases_thm, rel_cases_attrs)) =
                 let
-                  val (((Ds, As), Bs), names_lthy) = lthy
-                    |> mk_TFrees (dead_of_bnf fp_bnf)
-                    ||>> mk_TFrees (live_of_bnf fp_bnf)
-                    ||>> mk_TFrees (live_of_bnf fp_bnf);
-                  val TA as Type (_, ADs) = mk_T_of_bnf Ds As fp_bnf;
-                  val TB as Type (_, BDs) = mk_T_of_bnf Ds Bs fp_bnf;
-                  val fTs = map2 (curry op -->) As Bs;
-                  val rel = mk_rel_of_bnf Ds As Bs fp_bnf;
-                  val ((((fs, Rs), ta), tb), names_lthy) = names_lthy
+                  val live_AsBs = filter (op <>) (As ~~ Bs);
+                  val fTs = map (op -->) live_AsBs;
+                  val (((((fs, Rs), ta), tb), thesis), names_lthy) = names_lthy
                     |> mk_Frees "f" fTs
-                    ||>> mk_Frees "R" (map2 mk_pred2T As Bs)
-                    ||>> yield_singleton (mk_Frees "a") TA
-                    ||>> yield_singleton (mk_Frees "b") TB;
-                  val map_term = mk_map_of_bnf Ds As Bs fp_bnf;
-                  val discAs = map (mk_disc_or_sel ADs) discs;
-                  val selAss = map (map (mk_disc_or_sel ADs)) selss;
+                    ||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs)
+                    ||>> yield_singleton (mk_Frees "a") fpT
+                    ||>> yield_singleton (mk_Frees "b") fpBT
+                    ||>> apfst HOLogic.mk_Trueprop o
+                      yield_singleton (mk_Frees "thesis") HOLogic.boolT;
+                  val map_term = mk_map live As Bs (map_of_bnf fp_bnf);
+                  val ctrAs = map (mk_ctr As) ctrs;
+                  val ctrBs = map (mk_ctr Bs) ctrs;
+                  val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf);
+                  val setAs = map (mk_set As) (sets_of_bnf fp_bnf);
+                  val discAs = map (mk_disc_or_sel As) discs;
+                  val selAss = map (map (mk_disc_or_sel As)) selss;
                   val discAs_selAss = flat (map2 (map o pair) discAs selAss);
 
+                  val ctr_transfer_thms =
+                    let
+                      val goals = mk_parametricity_goals names_lthy Rs ctrAs ctrBs;
+                    in
+                      Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
+                        (K (mk_ctr_transfer_tac rel_intro_thms))
+                      |> Conjunction.elim_balanced (length goals)
+                      |> Proof_Context.export names_lthy lthy
+                      |> map Thm.close_derivation
+                    end;
+
+                  val (set_cases_thms, set_cases_attrss) =
+                    let
+                      fun mk_prems assms elem t ctxt =
+                        (case fastype_of t of
+                          Type (type_name, xs) =>
+                          (case bnf_of ctxt type_name of
+                            NONE => ([], ctxt)
+                          | SOME bnf =>
+                            apfst flat (fold_map (fn set => fn ctxt =>
+                              let
+                                val X = HOLogic.dest_setT (range_type (fastype_of set));
+                                val new_var = not (X = fastype_of elem);
+                                val (x, ctxt') =
+                                  if new_var then yield_singleton (mk_Frees "x") X ctxt
+                                  else (elem, ctxt);
+                              in
+                                mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt'
+                                |>> map (if new_var then Logic.all x else I)
+                              end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt))
+                        | T => rpair ctxt
+                          (if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis]
+                           else []));
+                    in
+                      split_list (map (fn set =>
+                        let
+                          val A = HOLogic.dest_setT (range_type (fastype_of set));
+                          val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy;
+                          val premss =
+                            map (fn ctr =>
+                              let
+                                val (args, names_lthy) =
+                                  mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy;
+                              in
+                                flat (zipping_map (fn (prev_args, arg, next_args) =>
+                                  let
+                                    val (args_with_elem, args_without_elem) =
+                                      if fastype_of arg = A then
+                                        (prev_args @ [elem] @ next_args, prev_args @ next_args)
+                                      else
+                                        `I (prev_args @ [arg] @ next_args);
+                                  in
+                                    mk_prems
+                                      [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))]
+                                      elem arg names_lthy
+                                    |> fst
+                                    |> map (fold_rev Logic.all args_without_elem)
+                                  end) args)
+                              end) ctrAs;
+                          val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis);
+                          val thm =
+                            Goal.prove_sorry lthy [] (flat premss) goal
+                              (fn {context = ctxt, prems} =>
+                                 mk_set_cases_tac ctxt (certify ctxt ta) prems exhaust set_thms)
+                            |> singleton (Proof_Context.export names_lthy lthy)
+                            |> Thm.close_derivation
+                            |> rotate_prems ~1;
+
+                          val consumes_attr = Attrib.internal (K (Rule_Cases.consumes 1));
+                          val cases_set_attr =
+                            Attrib.internal (K (Induct.cases_pred (name_of_set set)));
+                        in
+                          (* TODO: @{attributes [elim?]} *)
+                          (thm, [consumes_attr, cases_set_attr])
+                        end) setAs)
+                    end;
+
+                  val set_intros_thms =
+                    let
+                      fun mk_goals A setA ctr_args t ctxt =
+                        (case fastype_of t of
+                          Type (type_name, innerTs) =>
+                          (case bnf_of ctxt type_name of
+                            NONE => ([], ctxt)
+                          | SOME bnf =>
+                            apfst flat (fold_map (fn set => fn ctxt =>
+                              let
+                                val X = HOLogic.dest_setT (range_type (fastype_of set));
+                                val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt;
+                                val assm = mk_Trueprop_mem (x, set $ t);
+                              in
+                                apfst (map (Logic.mk_implies o pair assm))
+                                  (mk_goals A setA ctr_args x ctxt')
+                              end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt))
+                        | T =>
+                          (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt));
+
+                      val (goals, names_lthy) =
+                        apfst flat (fold_map (fn set => fn ctxt =>
+                          let
+                            val A = HOLogic.dest_setT (range_type (fastype_of set));
+                          in
+                            apfst flat (fold_map (fn ctr => fn ctxt =>
+                              let
+                                val (args, ctxt') =
+                                  mk_Frees "arg" (binder_types (fastype_of ctr)) ctxt;
+                                val ctr_args = Term.list_comb (ctr, args);
+                              in
+                                apfst flat (fold_map (mk_goals A set ctr_args) args ctxt')
+                              end) ctrAs ctxt)
+                          end) setAs lthy);
+                    in
+                      if null goals then []
+                      else
+                        Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
+                          (fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms)
+                        |> Conjunction.elim_balanced (length goals)
+                        |> Proof_Context.export names_lthy lthy
+                        |> map Thm.close_derivation
+                    end;
+
                   val rel_sel_thms =
                     let
-                      val discBs = map (mk_disc_or_sel BDs) discs;
-                      val selBss = map (map (mk_disc_or_sel BDs)) selss;
+                      val discBs = map (mk_disc_or_sel Bs) discs;
+                      val selBss = map (map (mk_disc_or_sel Bs)) selss;
                       val n = length discAs;
 
                       fun mk_rhs n k discA selAs discB selBs =
@@ -1433,21 +1643,18 @@
                       else
                         Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
                           (fn {context = ctxt, prems = _} =>
-                            mk_rel_sel_tac ctxt (certify ctxt ta) (certify ctxt tb) exhaust
-                              (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts
-                              rel_distinct_thms)
-                          |> Conjunction.elim_balanced (length goals)
-                          |> Proof_Context.export names_lthy lthy
+                             mk_rel_sel_tac ctxt (certify ctxt ta) (certify ctxt tb) exhaust
+                               (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts
+                               rel_distinct_thms)
+                        |> Conjunction.elim_balanced (length goals)
+                        |> Proof_Context.export names_lthy lthy
+                        |> map Thm.close_derivation
                     end;
 
                   val (rel_cases_thm, rel_cases_attrs) =
                     let
-                      val (thesis, names_lthy) = apfst HOLogic.mk_Trueprop
-                        (yield_singleton (mk_Frees "thesis") HOLogic.boolT names_lthy);
-
-                      val rel_Rs_a_b = list_comb (rel, Rs) $ ta $ tb;
-                      val ctrAs = map (mk_ctr ADs) ctrs;
-                      val ctrBs = map (mk_ctr BDs) ctrs;
+                      val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb;
+                      val ctrBs = map (mk_ctr Bs) ctrs;
 
                       fun mk_assms ctrA ctrB ctxt =
                         let
@@ -1469,13 +1676,14 @@
                         end;
 
                       val (assms, names_lthy) = fold_map2 mk_assms ctrAs ctrBs names_lthy;
-                      val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms,
-                        thesis);
-                      val thm = Goal.prove_sorry lthy [] [] goal
+                      val goal =
+                        Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis);
+                      val thm =
+                        Goal.prove_sorry lthy [] [] goal
                           (fn {context = ctxt, prems = _} =>
-                            mk_rel_cases_tac ctxt (certify ctxt ta) (certify ctxt tb) exhaust
-                              injects rel_inject_thms distincts rel_distinct_thms
-                              (map rel_eq_of_bnf live_nesting_bnfs))
+                             mk_rel_cases_tac ctxt (certify ctxt ta) (certify ctxt tb) exhaust
+                               injects rel_inject_thms distincts rel_distinct_thms
+                               (map rel_eq_of_bnf live_nesting_bnfs))
                         |> singleton (Proof_Context.export names_lthy lthy)
                         |> Thm.close_derivation;
 
@@ -1487,9 +1695,9 @@
                       (thm, [consumes_attr, case_names_attr, cases_pred_attr ""])
                     end;
 
-                  val disc_map_iff_thms =
+                  val map_disc_iff_thms =
                     let
-                      val discsB = map (mk_disc_or_sel BDs) discs;
+                      val discsB = map (mk_disc_or_sel Bs) discs;
                       val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs;
 
                       fun mk_goal (discA_t, discB) =
@@ -1506,31 +1714,34 @@
                       else
                         Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
                           (fn {context = ctxt, prems = _} =>
-                            mk_disc_map_iff_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss)
-                              map_thms)
-                          |> Conjunction.elim_balanced (length goals)
-                          |> Proof_Context.export names_lthy lthy
+                             mk_map_disc_iff_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss)
+                               map_thms)
+                        |> Conjunction.elim_balanced (length goals)
+                        |> Proof_Context.export names_lthy lthy
+                        |> map Thm.close_derivation
                     end;
 
-                  val sel_map_thms =
+                  val map_sel_thms =
                     let
                       fun mk_goal (discA, selA) =
                         let
                           val prem = Term.betapply (discA, ta);
-                          val selB = mk_disc_or_sel BDs selA;
+                          val selB = mk_disc_or_sel Bs selA;
                           val lhs = selB $ (Term.list_comb (map_term, fs) $ ta);
                           val lhsT = fastype_of lhs;
-                          val map_rhsT = map_atyps (perhaps (AList.lookup (op =) (Bs ~~ As))) lhsT;
+                          val map_rhsT =
+                            map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT;
                           val map_rhs = build_map lthy []
-                            (the o (AList.lookup (op =) ((As ~~ Bs) ~~ fs))) (map_rhsT, lhsT);
+                            (the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT);
                           val rhs = (case map_rhs of
-                            Const (@{const_name id}, _) => selA $ ta
+                              Const (@{const_name id}, _) => selA $ ta
                             | _ => map_rhs $ (selA $ ta));
                           val concl = mk_Trueprop_eq (lhs, rhs);
                         in
                           if is_refl_bool prem then concl
                           else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl)
                         end;
+
                       val goals = map mk_goal discAs_selAss;
                     in
                       if null goals then
@@ -1538,16 +1749,15 @@
                       else
                         Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
                           (fn {context = ctxt, prems = _} =>
-                            mk_sel_map_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss)
-                              map_thms (flat sel_thmss))
-                          |> Conjunction.elim_balanced (length goals)
-                          |> Proof_Context.export names_lthy lthy
+                             mk_map_sel_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss)
+                               map_thms (flat sel_thmss))
+                        |> Conjunction.elim_balanced (length goals)
+                        |> Proof_Context.export names_lthy lthy
+                        |> map Thm.close_derivation
                     end;
 
-                  val sel_set_thms =
+                  val set_sel_thms =
                     let
-                      val setsA = map (mk_set ADs) (sets_of_bnf fp_bnf);
-
                       fun mk_goal discA selA setA ctxt =
                         let
                           val prem = Term.betapply (discA, ta);
@@ -1556,7 +1766,7 @@
 
                           fun travese_nested_types t ctxt =
                             (case fastype_of t of
-                              Type (type_name, xs) =>
+                              Type (type_name, innerTs) =>
                               (case bnf_of ctxt type_name of
                                 NONE => ([], ctxt)
                               | SOME bnf =>
@@ -1571,7 +1781,7 @@
                                       |>> map (Logic.mk_implies o pair assm)
                                     end;
                                 in
-                                  fold_map (seq_assm t o mk_set xs) (sets_of_bnf bnf) ctxt
+                                  fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt
                                   |>> flat
                                 end)
                             | T =>
@@ -1596,20 +1806,22 @@
                             ([], ctxt)
                         end;
                       val (goals, names_lthy) = apfst (flat o flat) (fold_map (fn (disc, sel) =>
-                        fold_map (mk_goal disc sel) setsA) discAs_selAss names_lthy);
+                        fold_map (mk_goal disc sel) setAs) discAs_selAss names_lthy);
                     in
                       if null goals then
                         []
                       else
                         Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced goals)
                           (fn {context = ctxt, prems = _} =>
-                            mk_sel_set_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss)
-                              (flat sel_thmss) set_thms)
-                          |> Conjunction.elim_balanced (length goals)
-                          |> Proof_Context.export names_lthy lthy
+                             mk_set_sel_tac ctxt (certify ctxt ta) exhaust (flat disc_thmss)
+                               (flat sel_thmss) set0_thms)
+                        |> Conjunction.elim_balanced (length goals)
+                        |> Proof_Context.export names_lthy lthy
+                        |> map Thm.close_derivation
                     end;
                 in
-                  (disc_map_iff_thms, sel_map_thms, sel_set_thms, rel_sel_thms,
+                  (map_disc_iff_thms, map_sel_thms, set_sel_thms, rel_sel_thms, set_intros_thms,
+                    ctr_transfer_thms, (set_cases_thms, set_cases_attrss),
                     (rel_cases_thm, rel_cases_attrs))
                 end;
 
@@ -1619,17 +1831,19 @@
                 |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
 
               val notes =
-                [(disc_map_iffN, disc_map_iff_thms, simp_attrs),
-                 (mapN, map_thms, code_nitpicksimp_attrs @ simp_attrs),
-                 (rel_casesN, [rel_cases_thm], rel_cases_attrs),
-                 (rel_distinctN, rel_distinct_thms, simp_attrs),
-                 (rel_injectN, rel_inject_thms, simp_attrs),
-                 (rel_introsN, rel_intro_thms, []),
-                 (rel_selN, rel_sel_thms, []),
-                 (setN, set_thms, code_nitpicksimp_attrs @ simp_attrs),
-                 (sel_mapN, sel_map_thms, []),
-                 (sel_setN, sel_set_thms, []),
-                 (set_emptyN, set_empty_thms, [])]
+                [(ctr_transferN, ctr_transfer_thms, K []),
+                 (mapN, map_thms, K (code_nitpicksimp_attrs @ simp_attrs)),
+                 (map_disc_iffN, map_disc_iff_thms, K simp_attrs),
+                 (map_selN, map_sel_thms, K []),
+                 (rel_casesN, [rel_cases_thm], K rel_cases_attrs),
+                 (rel_distinctN, rel_distinct_thms, K simp_attrs),
+                 (rel_injectN, rel_inject_thms, K simp_attrs),
+                 (rel_introsN, rel_intro_thms, K []),
+                 (rel_selN, rel_sel_thms, K []),
+                 (setN, set_thms, K (code_nitpicksimp_attrs @ simp_attrs)),
+                 (set_casesN, set_cases_thms, nth set_cases_attrss),
+                 (set_introsN, set_intros_thms, K []),
+                 (set_selN, set_sel_thms, K [])]
                 |> massage_simple_notes fp_b_name;
 
               val (noted, lthy') =
@@ -1637,13 +1851,13 @@
                 |> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) map_thms)
                 |> fp = Least_FP
                   ? Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) rel_eq_thms)
-                |> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) set_thms)
+                |> Spec_Rules.add Spec_Rules.Equational (`(single o lhs_head_of o hd) set0_thms)
                 |> Local_Theory.notes (anonymous_notes @ notes);
 
               val subst = Morphism.thm (substitute_noted_thm noted);
             in
               (((map subst map_thms, map subst rel_inject_thms, map subst rel_distinct_thms,
-                 map (map subst) set_thmss), ctr_sugar), lthy')
+                 map (map subst) set0_thmss), ctr_sugar), lthy')
             end;
 
         fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);
@@ -1700,8 +1914,8 @@
 
         val common_notes =
           (if nn > 1 then
-             [(inductN, [induct_thm], induct_attrs),
-              (rel_inductN, common_rel_induct_thms, common_rel_induct_attrs)]
+             [(inductN, [induct_thm], K induct_attrs),
+              (rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)]
            else
              [])
           |> massage_simple_notes fp_common_name;
@@ -1729,17 +1943,17 @@
         ((((mapss, rel_injectss, rel_distinctss, setss), (_, _, ctr_defss, ctr_sugars)),
           (corecs, corec_defs)), lthy) =
       let
-        val (([(coinduct_thms, coinduct_thm), (strong_coinduct_thms, strong_coinduct_thm)],
+        val (([(coinduct_thms, coinduct_thm), (coinduct_strong_thms, coinduct_strong_thm)],
               (coinduct_attrs, common_coinduct_attrs)),
-             corec_thmss, disc_corec_thmss,
-             (disc_corec_iff_thmss, disc_corec_iff_attrs), (sel_corec_thmsss, sel_corec_attrs)) =
+             corec_thmss, corec_disc_thmss,
+             (corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) =
           derive_coinduct_corecs_thms_for_types pre_bnfs (the corecs_args_types) xtor_co_induct
             dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss
             ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs
             (Proof_Context.export lthy' no_defs_lthy) lthy;
 
         fun distinct_prems ctxt th =
-          Goal.prove ctxt [] []
+          Goal.prove (*no sorry*) ctxt [] []
             (th |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies)
             (fn _ => HEADGOAL (cut_tac th THEN' atac) THEN ALLGOALS eq_assume_tac);
 
@@ -1749,8 +1963,8 @@
                 (map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]})
                   [thm, eq_ifIN ctxt thms]));
 
-        val corec_code_thms = map (eq_ifIN lthy) corec_thmss
-        val sel_corec_thmss = map flat sel_corec_thmsss;
+        val corec_code_thms = map (eq_ifIN lthy) corec_thmss;
+        val corec_sel_thmss = map flat corec_sel_thmsss;
 
         val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type;
         val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred;
@@ -1773,40 +1987,48 @@
                (rel_coinduct_attrs, common_rel_coinduct_attrs))
             end;
 
+        val (set_induct_thms, set_induct_attrss) =
+          derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars)
+            (map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_induct_thms
+            (map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss)
+            dtor_ctors abs_inverses
+          |> split_list;
+
         val simp_thmss =
           map6 mk_simp_thms ctr_sugars
-            (map3 flat_corec_thms disc_corec_thmss disc_corec_iff_thmss sel_corec_thmss)
+            (map3 flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss)
             mapss rel_injectss rel_distinctss setss;
 
         val common_notes =
+          (set_inductN, set_induct_thms, nth set_induct_attrss) ::
           (if nn > 1 then
-            [(coinductN, [coinduct_thm], common_coinduct_attrs),
-             (rel_coinductN, common_rel_coinduct_thms, common_rel_coinduct_attrs),
-             (strong_coinductN, [strong_coinduct_thm], common_coinduct_attrs)]
-          else [])
+            [(coinductN, [coinduct_thm], K common_coinduct_attrs),
+             (coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs),
+             (rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)]
+           else [])
           |> massage_simple_notes fp_common_name;
 
         val notes =
           [(coinductN, map single coinduct_thms,
             fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]),
+           (coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs),
            (corecN, corec_thmss, K []),
            (corec_codeN, map single corec_code_thms, K code_nitpicksimp_attrs),
-           (disc_corecN, disc_corec_thmss, K []),
-           (disc_corec_iffN, disc_corec_iff_thmss, K disc_corec_iff_attrs),
+           (corec_discN, corec_disc_thmss, K []),
+           (corec_disc_iffN, corec_disc_iff_thmss, K corec_disc_iff_attrs),
+           (corec_selN, corec_sel_thmss, K corec_sel_attrs),
            (rel_coinductN, rel_coinduct_thmss, K (rel_coinduct_attrs @ [coinduct_pred_attr ""])),
-           (sel_corecN, sel_corec_thmss, K sel_corec_attrs),
-           (simpsN, simp_thmss, K []),
-           (strong_coinductN, map single strong_coinduct_thms, K coinduct_attrs)]
+           (simpsN, simp_thmss, K [])]
           |> massage_multi_notes;
       in
         lthy
         |> fold (curry (Spec_Rules.add Spec_Rules.Equational) corecs)
-          [flat sel_corec_thmss, flat corec_thmss]
+          [flat corec_sel_thmss, flat corec_thmss]
         |> Local_Theory.notes (common_notes @ notes)
         |-> register_as_fp_sugars fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos fp_nesting_bnfs
           live_nesting_bnfs fp_res ctrXs_Tsss ctr_defss ctr_sugars corecs corec_defs mapss
-          [coinduct_thm, strong_coinduct_thm] (transpose [coinduct_thms, strong_coinduct_thms])
-          corec_thmss disc_corec_thmss sel_corec_thmsss rel_injectss rel_distinctss
+          [coinduct_thm, coinduct_strong_thm] (transpose [coinduct_thms, coinduct_strong_thms])
+          corec_thmss corec_disc_thmss corec_sel_thmsss rel_injectss rel_distinctss
       end;
 
     val lthy'' = lthy'
--- a/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -10,21 +10,24 @@
 sig
   val sumprod_thms_map: thm list
   val sumprod_thms_set: thm list
+  val basic_sumprod_thms_set: thm list
   val sumprod_thms_rel: thm list
 
   val mk_coinduct_tac: Proof.context -> thm list -> int -> int list -> thm -> thm list ->
     thm list -> thm list -> thm list -> thm list -> thm list list -> thm list list list ->
     thm list list list -> tactic
   val mk_corec_tac: thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context -> tactic
+  val mk_corec_disc_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
   val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
     tactic
-  val mk_disc_corec_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
-  val mk_disc_map_iff_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic
+  val mk_ctr_transfer_tac: thm list -> tactic
   val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
   val mk_half_distinct_tac: Proof.context -> thm -> thm -> thm list -> tactic
   val mk_induct_tac: Proof.context -> int -> int list -> int list list -> int list list list ->
     thm list -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic
   val mk_inject_tac: Proof.context -> thm -> thm -> thm -> tactic
+  val mk_map_disc_iff_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic
+  val mk_map_sel_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
   val mk_rec_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> thm -> Proof.context ->
     tactic
   val mk_rel_cases_tac:  Proof.context -> cterm -> cterm -> thm -> thm list -> thm list ->
@@ -36,9 +39,11 @@
     thm list list -> thm list -> thm list -> thm list -> thm list -> tactic
   val mk_rel_sel_tac: Proof.context -> cterm -> cterm -> thm -> thm list -> thm list -> thm list ->
     thm list -> thm list -> tactic
-  val mk_sel_map_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
-  val mk_sel_set_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
-  val mk_set_empty_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> tactic
+  val mk_set_cases_tac: Proof.context -> cterm -> thm list -> thm -> thm list -> tactic
+  val mk_set_induct0_tac: Proof.context -> cterm list -> thm list -> thm list -> thm list ->
+    thm list -> thm list -> thm list -> thm list -> tactic
+  val mk_set_intros_tac: Proof.context -> thm list -> tactic
+  val mk_set_sel_tac: Proof.context -> cterm -> thm -> thm list -> thm list -> thm list -> tactic
 end;
 
 structure BNF_FP_Def_Sugar_Tactics : BNF_FP_DEF_SUGAR_TACTICS =
@@ -54,10 +59,10 @@
 val simp_thms' = @{thms simp_thms(6,7,8,11,12,15,16,22,24)};
 
 val sumprod_thms_map = @{thms id_apply map_prod_simp prod.case sum.case map_sum.simps};
-val sumprod_thms_set =
-  @{thms UN_empty UN_insert Un_empty_left Un_empty_right Un_iff UN_simps(10) UN_iff
-      Union_Un_distrib image_iff o_apply map_prod_simp
-      mem_Collect_eq prod_set_simps map_sum.simps sum_set_simps};
+val basic_sumprod_thms_set =
+  @{thms UN_empty UN_insert  UN_iff Un_empty_left Un_empty_right Un_iff Union_Un_distrib
+       o_apply map_prod_simp mem_Collect_eq prod_set_simps map_sum.simps sum_set_simps};
+val sumprod_thms_set = @{thms UN_simps(10) image_iff} @ basic_sumprod_thms_set;
 val sumprod_thms_rel = @{thms rel_sum_simps rel_prod_apply prod.inject id_apply conj_assoc};
 
 fun hhf_concl_conv cv ctxt ct =
@@ -91,6 +96,11 @@
       SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
       atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor]));
 
+fun mk_ctr_transfer_tac rel_intros =
+  HEADGOAL Goal.conjunction_tac THEN
+  ALLGOALS (REPEAT o (resolve_tac (@{thm rel_funI} :: rel_intros) THEN'
+    REPEAT_DETERM o atac));
+
 fun mk_half_distinct_tac ctxt ctor_inject abs_inject ctr_defs =
   unfold_thms_tac ctxt (ctor_inject :: abs_inject :: @{thms sum.inject} @ ctr_defs) THEN
   HEADGOAL (rtac @{thm sum.distinct(1)});
@@ -121,22 +131,13 @@
     HEADGOAL (rtac refl ORELSE' rtac (@{thm unit_eq} RS arg_cong))
   end;
 
-fun mk_disc_corec_iff_tac case_splits' corecs discs ctxt =
+fun mk_corec_disc_iff_tac case_splits' corecs discs ctxt =
   EVERY (map3 (fn case_split_tac => fn corec_thm => fn disc =>
       HEADGOAL case_split_tac THEN unfold_thms_tac ctxt [corec_thm] THEN
       HEADGOAL (asm_simp_tac (ss_only basic_simp_thms ctxt)) THEN
       (if is_refl disc then all_tac else HEADGOAL (rtac disc)))
     (map rtac case_splits' @ [K all_tac]) corecs discs);
 
-fun mk_disc_map_iff_tac ctxt ct exhaust discs maps =
-  TRYALL Goal.conjunction_tac THEN
-  ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
-    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
-  unfold_thms_tac ctxt maps THEN
-  unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI
-    handle THM _ => thm RS eqTrueI) discs) THEN
-  ALLGOALS (rtac refl ORELSE' rtac TrueI);
-
 fun solve_prem_prem_tac ctxt =
   REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
     hyp_subst_tac ctxt ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
@@ -216,6 +217,25 @@
       (1 upto nn) ns pre_rel_defs fp_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss discsss
       selsss));
 
+fun mk_map_disc_iff_tac ctxt ct exhaust discs maps =
+  TRYALL Goal.conjunction_tac THEN
+  ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
+    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
+  unfold_thms_tac ctxt maps THEN
+  unfold_thms_tac ctxt (map (fn thm => thm RS eqFalseI
+    handle THM _ => thm RS eqTrueI) discs) THEN
+  ALLGOALS (rtac refl ORELSE' rtac TrueI);
+
+fun mk_map_sel_tac ctxt ct exhaust discs maps sels =
+  TRYALL Goal.conjunction_tac THEN
+    ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
+      REPEAT_DETERM o hyp_subst_tac ctxt) THEN
+    unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @
+      @{thms not_True_eq_False not_False_eq_True}) THEN
+    TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN
+    unfold_thms_tac ctxt (maps @ sels) THEN
+    ALLGOALS (rtac refl);
+
 fun mk_rel_cases_tac ctxt ct1 ct2 exhaust injects rel_injects distincts rel_distincts rel_eqs=
   HEADGOAL (rtac (cterm_instantiate_pos [SOME ct1] exhaust) THEN_ALL_NEW
     rtac (cterm_instantiate_pos [SOME ct2] exhaust) THEN_ALL_NEW
@@ -257,12 +277,12 @@
         HEADGOAL (rtac exhaust THEN_ALL_NEW (rtac exhaust THEN_ALL_NEW
           (rtac (cterm_instantiate_pos (replicate 4 NONE @ [SOME cterm]) @{thm arg_cong2} RS iffD2)
             THEN' atac THEN' atac THEN' TRY o resolve_tac assms))) THEN
-        unfold_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @
+        unfold_thms_tac ctxt (ctor_inject :: rel_pre_list_def :: ctor_defs @ nesting_rel_eqs @
           @{thms BNF_Comp.id_bnf_comp_def vimage2p_def}) THEN
         TRYALL (hyp_subst_tac ctxt) THEN
         unfold_thms_tac ctxt (Abs_inverse :: @{thms rel_sum_simps rel_prod_apply Inl_Inr_False
           Inr_Inl_False  sum.inject prod.inject}) THEN
-        TRYALL (etac FalseE ORELSE' (REPEAT_DETERM o etac conjE) THEN' atac))
+        TRYALL (rtac refl ORELSE' etac FalseE ORELSE' (REPEAT_DETERM o etac conjE) THEN' atac))
     cterms exhausts ctor_defss ctor_injects rel_pre_list_defs Abs_inverses);
 
 fun mk_rel_sel_tac ctxt ct1 ct2 exhaust discs sels rel_injects distincts rel_distincts =
@@ -275,17 +295,7 @@
     (rel_distincts RL @{thms iffD2[OF eq_False]})) THEN
   TRYALL (resolve_tac [TrueI, refl]);
 
-fun mk_sel_map_tac ctxt ct exhaust discs maps sels =
-  TRYALL Goal.conjunction_tac THEN
-    ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
-      REPEAT_DETERM o hyp_subst_tac ctxt) THEN
-    unfold_thms_tac ctxt ((discs RL [eqTrueI, eqFalseI]) @
-      @{thms not_True_eq_False not_False_eq_True}) THEN
-    TRYALL (etac FalseE ORELSE' etac @{thm TrueE}) THEN
-    unfold_thms_tac ctxt (maps @ sels) THEN
-    ALLGOALS (rtac refl);
-
-fun mk_sel_set_tac ctxt ct exhaust discs sels sets =
+fun mk_set_sel_tac ctxt ct exhaust discs sels sets =
   TRYALL Goal.conjunction_tac THEN
     ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
       REPEAT_DETERM o hyp_subst_tac ctxt) THEN
@@ -298,12 +308,43 @@
         hyp_subst_tac ctxt) THEN'
       (rtac @{thm singletonI} ORELSE' atac));
 
-fun mk_set_empty_tac ctxt ct exhaust sets discs =
+fun mk_set_cases_tac ctxt ct assms exhaust sets =
+  HEADGOAL (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW hyp_subst_tac ctxt) THEN
+  unfold_thms_tac ctxt sets THEN
+  REPEAT_DETERM (HEADGOAL
+    (eresolve_tac @{thms FalseE emptyE singletonE UnE UN_E insertE} ORELSE'
+     hyp_subst_tac ctxt ORELSE'
+     SELECT_GOAL (SOLVE (HEADGOAL (eresolve_tac assms THEN' REPEAT_DETERM o atac)))));
+
+fun mk_set_intros_tac ctxt sets =
   TRYALL Goal.conjunction_tac THEN
-  ALLGOALS (rtac (cterm_instantiate_pos [SOME ct] exhaust) THEN_ALL_NEW
-    REPEAT_DETERM o hyp_subst_tac ctxt) THEN
-  unfold_thms_tac ctxt (sets @ map_filter (fn thm =>
-    SOME (thm RS eqFalseI) handle THM _ => NONE) discs) THEN
-  ALLGOALS (rtac refl ORELSE' etac FalseE);
+  unfold_thms_tac ctxt sets THEN
+  TRYALL (REPEAT o
+    (resolve_tac @{thms UnI1 UnI2} ORELSE'
+     eresolve_tac @{thms UN_I UN_I[rotated]}) THEN' (rtac @{thm singletonI} ORELSE' atac));
+
+fun mk_set_induct0_tac ctxt cts assms dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors
+    Abs_pre_inverses =
+  let
+    val assms_ctor_defs =
+      map (unfold_thms ctxt (@{thm BNF_Comp.id_bnf_comp_def} :: ctor_defs)) assms;
+    val exhausts' = map (fn thm => thm RS @{thm asm_rl[of "P x y" for P x y]}) exhausts
+      |> map2 (fn ct => cterm_instantiate_pos [NONE, SOME ct]) cts;
+  in
+    ALLGOALS (resolve_tac dtor_set_inducts) THEN
+    TRYALL (resolve_tac exhausts' THEN_ALL_NEW
+      (resolve_tac (map (fn ct => refl RS
+         cterm_instantiate_pos (replicate 4 NONE @ [SOME ct]) @{thm arg_cong2} RS iffD2) cts)
+        THEN' atac THEN' hyp_subst_tac ctxt)) THEN
+    unfold_thms_tac ctxt (Abs_pre_inverses @ dtor_ctors @ set_pre_defs @ ctor_defs @
+      @{thms BNF_Comp.id_bnf_comp_def o_apply sum_set_simps prod_set_simps UN_empty UN_insert
+        Un_empty_left Un_empty_right empty_iff singleton_iff}) THEN
+    REPEAT_DETERM (HEADGOAL
+      (TRY o etac UnE THEN' TRY o etac @{thm singletonE} THEN' TRY o hyp_subst_tac ctxt THEN'
+       REPEAT_DETERM o eresolve_tac @{thms UN_E UnE singletonE} THEN'
+       fold (curry (op ORELSE')) (map (fn thm =>
+         funpow (length (prems_of thm)) (fn tac => tac THEN' atac) (rtac thm)) assms_ctor_defs)
+         (etac FalseE)))
+  end;
 
 end;
--- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -240,29 +240,26 @@
     val co_recs = of_fp_res #xtor_co_recs;
     val ns = map (length o #Ts o #fp_res) fp_sugars;
 
-    fun substT rho (Type (@{type_name fun}, [T, U])) = substT rho T --> substT rho U
-      | substT rho (Type (s, Ts)) = Type (s, map (typ_subst_nonatomic rho) Ts)
-      | substT _ T = T;
-
     val typ_subst_nonatomic_sorted = fold_rev (typ_subst_nonatomic o single);
 
     fun foldT_of_recT recT =
       let
-        val ((FTXs, Xs), TX) = strip_fun_type recT |>> map_split dest_co_algT;
+        val ((FTXs, Ys), TX) = strip_fun_type recT |>> map_split dest_co_algT;
+        val Zs = union op = Xs Ys;
         fun subst (Type (C, Ts as [_, X])) =
-            if C = co_productC andalso member op = Xs X then X else Type (C, map subst Ts)
+            if C = co_productC andalso member op = Zs X then X else Type (C, map subst Ts)
           | subst (Type (C, Ts)) = Type (C, map subst Ts)
           | subst T = T;
       in
-        map2 (mk_co_algT o subst) FTXs Xs ---> TX
+        map2 (mk_co_algT o subst) FTXs Ys ---> TX
       end;
 
-    fun force_rec i TU TU_rec raw_rec =
+    fun force_rec i TU raw_rec =
       let
         val thy = Proof_Context.theory_of lthy;
 
         val approx_rec = raw_rec
-          |> force_typ names_lthy (replicate (nth ns i) dummyT ---> TU_rec);
+          |> force_typ names_lthy (replicate (nth ns i) dummyT ---> TU);
         val subst = Term.typ_subst_atomic fold_thetaAs;
 
         fun mk_fp_absT_repT fp_repT fp_absT = mk_absT thy fp_repT fp_absT ooo mk_repT;
@@ -299,9 +296,7 @@
         val i = find_index (fn T => x = T) Xs;
         val TUrec =
           (case find_first (fn f => body_fun_type (fastype_of f) = TU) recs of
-            NONE =>
-              force_rec i TU
-                (TU |> is_none b_opt ? substT (map2 mk_co_productT fpTs Xs ~~ Xs)) (nth co_recs i)
+            NONE => force_rec i TU (nth co_recs i)
           | SOME f => f);
 
         val TUs = binder_fun_types (fastype_of TUrec);
@@ -340,14 +335,21 @@
                     let
                       val (TY, (U, X)) = TU |> dest_co_algT ||> dest_co_productT;
                       val T = mk_co_algT TY U;
+                      fun mk_co_proj TU = build_map lthy [] (fn TU =>
+                        let
+                          val ((T1, T2), U) = TU |> co_swap |>> dest_co_productT
+                        in
+                          if T1 = U then co_proj1_const TU
+                          else mk_co_comp (mk_co_proj (co_swap (T1, U)),
+                            co_proj1_const (co_swap (mk_co_productT T1 T2, T1)))
+                        end) TU;
                     in
-                      (case try (force_typ names_lthy T o build_map lthy [] co_proj1_const o dest_funT) T of
-                        SOME f => mk_co_product f
-                          (fst (fst (mk_rec NONE recs lthy (mk_co_algT TY X))))
-                      | NONE => mk_map_co_product
-                          (build_map lthy [] co_proj1_const
-                            (dest_funT (mk_co_algT (dest_co_productT TY |> fst) U)))
-                          (HOLogic.id_const X))
+                      if not (can dest_co_productT TY)
+                      then mk_co_product (mk_co_proj (dest_funT T))
+                        (fst (fst (mk_rec NONE recs lthy (mk_co_algT TY X))))
+                      else mk_map_co_product
+                        (mk_co_proj (co_swap (dest_co_productT TY |> fst, U)))
+                        (HOLogic.id_const X)
                     end)
                 val smap_args = map mk_smap_arg smap_argTs;
               in
@@ -413,8 +415,8 @@
           map (fn thm => thm RS rewrite_comp_comp) @{thms map_prod.comp map_sum.comp} @
           @{thms id_apply comp_id id_comp map_prod.comp map_prod.id map_sum.comp map_sum.id};
         val rec_thms = fold_thms @ case_fp fp
-          @{thms fst_convol map_prod_o_convol convol_o}
-          @{thms case_sum_o_inj(1) case_sum_o_map_sum o_case_sum};
+          @{thms fst_convol map_prod_o_convol convol_o fst_comp_map_prod}
+          @{thms case_sum_o_inj(1) case_sum_o_map_sum o_case_sum map_sum_o_inj(1)};
 
         val eq_thm_prop_untyped = Term.aconv_untyped o pairself Thm.full_prop_of;
 
@@ -429,10 +431,11 @@
               rewrite_comp_comp)
           type_definitions bnfs);
 
-        fun mk_Rep_o_Abs thm = thm RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp;
+        fun mk_Rep_o_Abs thm = (thm RS @{thm type_copy_Rep_o_Abs})
+          |> (fn thm => [thm, thm RS rewrite_comp_comp]);
 
-        val fp_Rep_o_Abss = map mk_Rep_o_Abs fp_type_definitions;
-        val Rep_o_Abss = map mk_Rep_o_Abs type_definitions;
+        val fp_Rep_o_Abss = maps mk_Rep_o_Abs fp_type_definitions;
+        val Rep_o_Abss = maps mk_Rep_o_Abs type_definitions;
 
         fun tac {context = ctxt, prems = _} =
           unfold_thms_tac ctxt (flat [rec_thms, raw_co_rec_defs, pre_map_defs, fp_pre_map_defs,
@@ -463,7 +466,8 @@
         xtor_rel_thms = of_fp_res #xtor_rel_thms (*too general types and terms*),
         xtor_co_rec_thms = xtor_co_rec_thms,
         xtor_co_rec_o_map_thms = fp_rec_o_maps (*theorems about old constants*),
-        rel_xtor_co_induct_thm = rel_xtor_co_induct_thm}
+        rel_xtor_co_induct_thm = rel_xtor_co_induct_thm,
+        dtor_set_induct_thms = []}
        |> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy))));
   in
     (fp_res, lthy)
--- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -8,6 +8,7 @@
 signature BNF_FP_N2M_SUGAR =
 sig
   val unfold_lets_splits: term -> term
+  val unfold_splits_lets: term -> term
   val dest_map: Proof.context -> string -> term -> term * term list
 
   val mutualize_fp_sugars: BNF_Util.fp_kind -> int list -> binding list -> typ list -> term list ->
@@ -60,18 +61,22 @@
   Local_Theory.declaration {syntax = false, pervasive = false}
     (fn phi => Data.map (Typtab.update (key, morph_n2m_sugar phi n2m_sugar)));
 
-fun unfold_lets_splits (Const (@{const_name Let}, _) $ arg1 $ arg2) =
-    unfold_lets_splits (betapply (arg2, arg1))
-  | unfold_lets_splits (t as Const (@{const_name case_prod}, _) $ u) =
-    (case unfold_lets_splits u of
+fun unfold_lets_splits (Const (@{const_name Let}, _) $ t $ u) =
+    unfold_lets_splits (betapply (unfold_splits_lets u, t))
+  | unfold_lets_splits (t $ u) = betapply (unfold_lets_splits t, unfold_lets_splits u)
+  | unfold_lets_splits (Abs (s, T, t)) = Abs (s, T, unfold_lets_splits t)
+  | unfold_lets_splits t = t
+and unfold_splits_lets ((t as Const (@{const_name case_prod}, _)) $ u) =
+    (case unfold_splits_lets u of
       u' as Abs (s1, T1, Abs (s2, T2, _)) =>
       let val v = Var ((s1 ^ s2, Term.maxidx_of_term u' + 1), HOLogic.mk_prodT (T1, T2)) in
         lambda v (incr_boundvars 1 (betapplys (u', [HOLogic.mk_fst v, HOLogic.mk_snd v])))
       end
-    | _ => t)
-  | unfold_lets_splits (t $ u) = betapply (pairself unfold_lets_splits (t, u))
-  | unfold_lets_splits (Abs (s, T, t)) = Abs (s, T, unfold_lets_splits t)
-  | unfold_lets_splits t = t;
+    | _ => t $ unfold_lets_splits u)
+  | unfold_splits_lets (t as Const (@{const_name Let}, _) $ _ $ _) = unfold_lets_splits t
+  | unfold_splits_lets (t $ u) = betapply (unfold_splits_lets t, unfold_lets_splits u)
+  | unfold_splits_lets (Abs (s, T, t)) = Abs (s, T, unfold_splits_lets t)
+  | unfold_splits_lets t = unfold_lets_splits t;
 
 fun mk_map_pattern ctxt s =
   let
@@ -252,7 +257,7 @@
             fp_bs xtor_co_recs lthy
           |>> split_list;
 
-        val ((common_co_inducts, co_inductss, co_rec_thmss, disc_corec_thmss, sel_corec_thmsss),
+        val ((common_co_inducts, co_inductss, co_rec_thmss, corec_disc_thmss, corec_sel_thmsss),
              fp_sugar_thms) =
           if fp = Least_FP then
             derive_induct_recs_thms_for_types pre_bnfs recs_args_types xtor_co_induct
@@ -267,30 +272,30 @@
               dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss
               mss ns fp_abs_inverses abs_inverses (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss
               ctr_sugars co_recs co_rec_defs (Proof_Context.export lthy no_defs_lthy) lthy
-            |> `(fn ((coinduct_thms_pairs, _), corec_thmss, disc_corec_thmss, _,
-                     (sel_corec_thmsss, _)) =>
+            |> `(fn ((coinduct_thms_pairs, _), corec_thmss, corec_disc_thmss, _,
+                     (corec_sel_thmsss, _)) =>
               (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss,
-               disc_corec_thmss, sel_corec_thmsss))
+               corec_disc_thmss, corec_sel_thmsss))
             ||> (fn info => (NONE, SOME info));
 
         val phi = Proof_Context.export_morphism no_defs_lthy no_defs_lthy0;
 
         fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctr_defs ctr_sugar co_rec
-            co_rec_def maps co_inducts co_rec_thms disc_corec_thms sel_corec_thmss
+            co_rec_def maps co_inducts co_rec_thms corec_disc_thms corec_sel_thmss
             ({rel_injects, rel_distincts, ...} : fp_sugar) =
           {T = T, BT = Term.dummyT (*wrong but harmless*), X = X, fp = fp, fp_res = fp_res,
            fp_res_index = kk, pre_bnf = pre_bnf, absT_info = absT_info,
            fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs,
            ctrXs_Tss = ctrXs_Tss, ctr_defs = ctr_defs, ctr_sugar = ctr_sugar, co_rec = co_rec,
            co_rec_def = co_rec_def, maps = maps, common_co_inducts = common_co_inducts,
-           co_inducts = co_inducts, co_rec_thms = co_rec_thms, disc_co_recs = disc_corec_thms,
-           sel_co_recss = sel_corec_thmss, rel_injects = rel_injects, rel_distincts = rel_distincts}
+           co_inducts = co_inducts, co_rec_thms = co_rec_thms, disc_co_recs = corec_disc_thms,
+           sel_co_recss = corec_sel_thmss, rel_injects = rel_injects, rel_distincts = rel_distincts}
           |> morph_fp_sugar phi;
 
         val target_fp_sugars =
           map16 mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctr_defss ctr_sugars
-            co_recs co_rec_defs mapss (transpose co_inductss) co_rec_thmss disc_corec_thmss
-            sel_corec_thmsss fp_sugars0;
+            co_recs co_rec_defs mapss (transpose co_inductss) co_rec_thmss corec_disc_thmss
+            corec_sel_thmsss fp_sugars0;
 
         val n2m_sugar = (target_fp_sugars, fp_sugar_thms);
       in
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -25,7 +25,8 @@
      xtor_rel_thms: thm list,
      xtor_co_rec_thms: thm list,
      xtor_co_rec_o_map_thms: thm list,
-     rel_xtor_co_induct_thm: thm}
+     rel_xtor_co_induct_thm: thm,
+     dtor_set_induct_thms: thm list}
 
   val morph_fp_result: morphism -> fp_result -> fp_result
 
@@ -70,7 +71,10 @@
   val caseN: string
   val coN: string
   val coinductN: string
+  val coinduct_strongN: string
   val corecN: string
+  val corec_discN: string
+  val corec_disc_iffN: string
   val ctorN: string
   val ctor_dtorN: string
   val ctor_exhaustN: string
@@ -90,8 +94,6 @@
   val ctor_rel_inductN: string
   val ctor_set_inclN: string
   val ctor_set_set_inclN: string
-  val disc_corecN: string
-  val disc_corec_iffN: string
   val dtorN: string
   val dtor_coinductN: string
   val dtor_corecN: string
@@ -102,13 +104,13 @@
   val dtor_injectN: string
   val dtor_mapN: string
   val dtor_map_coinductN: string
-  val dtor_map_strong_coinductN: string
+  val dtor_map_coinduct_strongN: string
   val dtor_map_uniqueN: string
   val dtor_relN: string
   val dtor_rel_coinductN: string
   val dtor_set_inclN: string
   val dtor_set_set_inclN: string
-  val dtor_strong_coinductN: string
+  val dtor_coinduct_strongN: string
   val dtor_unfoldN: string
   val dtor_unfold_o_mapN: string
   val dtor_unfold_transferN: string
@@ -133,14 +135,13 @@
   val rel_distinctN: string
   val rel_selN: string
   val rvN: string
-  val sel_corecN: string
+  val corec_selN: string
   val set_inclN: string
   val set_set_inclN: string
   val setN: string
   val simpsN: string
   val strTN: string
   val str_initN: string
-  val strong_coinductN: string
   val sum_bdN: string
   val sum_bdTN: string
   val uniqueN: string
@@ -205,7 +206,7 @@
   val mk_xtor_un_fold_o_map_thms: BNF_Util.fp_kind -> bool -> int -> thm -> thm list -> thm list ->
     thm list -> thm list -> thm list
 
-  val mk_strong_coinduct_thm: thm -> thm list -> thm list -> (thm -> thm) -> Proof.context -> thm
+  val mk_coinduct_strong_thm: thm -> thm list -> thm list -> (thm -> thm) -> Proof.context -> thm
 
   val fp_bnf: (binding list -> (string * sort) list -> typ list * typ list list ->
       BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> 'a) ->
@@ -237,11 +238,12 @@
    xtor_rel_thms: thm list,
    xtor_co_rec_thms: thm list,
    xtor_co_rec_o_map_thms: thm list,
-   rel_xtor_co_induct_thm: thm};
+   rel_xtor_co_induct_thm: thm,
+   dtor_set_induct_thms: thm list};
 
 fun morph_fp_result phi {Ts, bnfs, ctors, dtors, xtor_co_recs, xtor_co_induct,
     dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_map_thms, xtor_set_thmss,
-    xtor_rel_thms, xtor_co_rec_thms, xtor_co_rec_o_map_thms, rel_xtor_co_induct_thm} =
+    xtor_rel_thms, xtor_co_rec_thms, xtor_co_rec_o_map_thms, rel_xtor_co_induct_thm, dtor_set_induct_thms} =
   {Ts = map (Morphism.typ phi) Ts,
    bnfs = map (morph_bnf phi) bnfs,
    ctors = map (Morphism.term phi) ctors,
@@ -257,7 +259,8 @@
    xtor_rel_thms = map (Morphism.thm phi) xtor_rel_thms,
    xtor_co_rec_thms = map (Morphism.thm phi) xtor_co_rec_thms,
    xtor_co_rec_o_map_thms = map (Morphism.thm phi) xtor_co_rec_o_map_thms,
-   rel_xtor_co_induct_thm = Morphism.thm phi rel_xtor_co_induct_thm};
+   rel_xtor_co_induct_thm = Morphism.thm phi rel_xtor_co_induct_thm,
+   dtor_set_induct_thms = map (Morphism.thm phi) dtor_set_induct_thms}; (* No idea of what this is doing... *)
 
 type fp_sugar =
   {T: typ,
@@ -390,9 +393,9 @@
 val ctor_induct2N = ctor_inductN ^ "2"
 val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN
 val dtor_coinductN = dtorN ^ "_" ^ coinductN
-val strong_coinductN = "strong_" ^ coinductN
-val dtor_map_strong_coinductN = dtor_mapN ^ "_" ^ strong_coinductN
-val dtor_strong_coinductN = dtorN ^ "_" ^ strong_coinductN
+val coinduct_strongN = coinductN ^ "_strong"
+val dtor_map_coinduct_strongN = dtor_mapN ^ "_" ^ coinduct_strongN
+val dtor_coinduct_strongN = dtorN ^ "_" ^ coinduct_strongN
 val colN = "col"
 val set_inclN = "set_incl"
 val ctor_set_inclN = ctorN ^ "_" ^ set_inclN
@@ -403,9 +406,9 @@
 
 val caseN = "case"
 val discN = "disc"
-val disc_corecN = discN ^ "_" ^ corecN
+val corec_discN = corecN ^ "_" ^ discN
 val iffN = "_iff"
-val disc_corec_iffN = discN ^ "_" ^ corecN ^ iffN
+val corec_disc_iffN = corec_discN ^ iffN
 val distinctN = "distinct"
 val rel_distinctN = relN ^ "_" ^ distinctN
 val injectN = "inject"
@@ -418,7 +421,7 @@
 val rel_inductN = relN ^ "_" ^ inductN
 val ctor_rel_inductN = ctorN ^ "_" ^ rel_inductN
 val selN = "sel"
-val sel_corecN = selN ^ "_" ^ corecN
+val corec_selN = corecN ^ "_" ^ selN
 
 fun co_prefix fp = case_fp fp "" "co";
 
@@ -631,7 +634,7 @@
     split_conj_thm (un_fold_unique OF map (case_fp fp I mk_sym) unique_prems)
   end;
 
-fun mk_strong_coinduct_thm coind rel_eqs rel_monos mk_vimage2p ctxt =
+fun mk_coinduct_strong_thm coind rel_eqs rel_monos mk_vimage2p ctxt =
   let
     val n = Thm.nprems_of coind;
     val m = Thm.nprems_of (hd rel_monos) - n;
--- a/src/HOL/Tools/BNF/bnf_gfp.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -223,7 +223,7 @@
     val rel_congs = map rel_cong_of_bnf bnfs;
     val rel_converseps = map rel_conversep_of_bnf bnfs;
     val rel_Grps = map rel_Grp_of_bnf bnfs;
-    val rel_OOs = map rel_OO_of_bnf bnfs;
+    val le_rel_OOs = map le_rel_OO_of_bnf bnfs;
     val in_rels = map in_rel_of_bnf bnfs;
 
     val timer = time (timer "Extracted terms & thms");
@@ -589,7 +589,7 @@
           HOLogic.mk_Trueprop (mk_bis Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
       in
         Goal.prove_sorry lthy [] [] (Logic.list_implies (prems, concl))
-          (K (mk_bis_O_tac lthy m bis_rel_thm rel_congs rel_OOs))
+          (K (mk_bis_O_tac lthy m bis_rel_thm rel_congs le_rel_OOs))
         |> Thm.close_derivation
         |> singleton (Proof_Context.export names_lthy lthy)
       end;
@@ -1592,7 +1592,7 @@
     val (dtor_corec_unique_thms, dtor_corec_unique_thm) =
       `split_conj_thm (split_conj_prems n
         (mor_UNIV_thm RS iffD2 RS corec_unique_mor_thm)
-        |> Local_Defs.unfold lthy (@{thms o_case_sum comp_id id_comp comp_assoc[symmetric]
+        |> unfold_thms lthy (@{thms o_case_sum comp_id id_comp comp_assoc[symmetric]
            case_sum_o_inj(1)} @ map_id0s_o_id @ sym_map_comps)
         OF replicate n @{thm arg_cong2[of _ _ _ _ case_sum, OF refl]});
 
@@ -1659,12 +1659,12 @@
 
     (*register new codatatypes as BNFs*)
     val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jset_thmss',
-      dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, lthy) =
+      dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, dtor_Jset_induct_thms, lthy) =
       if m = 0 then
         (timer, replicate n DEADID_bnf,
         map_split (`(mk_pointfree lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_ids),
         replicate n [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs,
-        mk_Jrel_DEADID_coinduct_thm (), [], lthy)
+        mk_Jrel_DEADID_coinduct_thm (), [], [], lthy)
       else let
         val fTs = map2 (curry op -->) passiveAs passiveBs;
         val gTs = map2 (curry op -->) passiveBs passiveCs;
@@ -2260,7 +2260,7 @@
             val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals);
           in
             Goal.prove_sorry lthy [] [] goal
-              (K (mk_le_rel_OO_tac Jrel_coinduct_thm dtor_Jrel_thms rel_OOs))
+              (K (mk_le_rel_OO_tac Jrel_coinduct_thm dtor_Jrel_thms le_rel_OOs))
             |> Thm.close_derivation
             |> singleton (Proof_Context.export names_lthy lthy)
           end;
@@ -2464,7 +2464,8 @@
             bs thmss)
       in
         (timer, Jbnfs, (Jmap_thms, dtor_Jmap_thms), dtor_Jset_thmss',
-          dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, lthy)
+          dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, dtor_Jset_induct_thms,
+          lthy)
       end;
 
     val dtor_unfold_o_Jmap_thms = mk_xtor_un_fold_o_map_thms Greatest_FP false m
@@ -2513,12 +2514,7 @@
           ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
         bs thmss);
 
-    (*FIXME: once the package exports all the necessary high-level characteristic theorems,
-       those should not only be concealed but rather not noted at all*)
-    val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
-
-    val (noted, lthy') =
-      lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Jbnf_notes));
+    val lthy' = lthy |> note_all ? snd o Local_Theory.notes (common_notes @ notes @ Jbnf_notes);
 
     val fp_res =
       {Ts = Ts, bnfs = Jbnfs, ctors = ctors, dtors = dtors, xtor_co_recs = corecs,
@@ -2526,8 +2522,8 @@
        ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
        xtor_map_thms = dtor_Jmap_thms, xtor_set_thmss = dtor_Jset_thmss',
        xtor_rel_thms = dtor_Jrel_thms, xtor_co_rec_thms = dtor_corec_thms,
-       xtor_co_rec_o_map_thms = dtor_corec_o_Jmap_thms, rel_xtor_co_induct_thm = Jrel_coinduct_thm}
-      |> morph_fp_result (substitute_noted_thm noted);
+       xtor_co_rec_o_map_thms = dtor_corec_o_Jmap_thms, rel_xtor_co_induct_thm = Jrel_coinduct_thm,
+       dtor_set_induct_thms = dtor_Jset_induct_thms};
   in
     timer; (fp_res, lthy')
   end;
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -69,15 +69,15 @@
    calls: corec_call list,
    discI: thm,
    sel_thms: thm list,
-   disc_excludess: thm list list,
+   distinct_discss: thm list list,
    collapse: thm,
    corec_thm: thm,
-   disc_corec: thm,
-   sel_corecs: thm list};
+   corec_disc: thm,
+   corec_sels: thm list};
 
 type corec_spec =
   {corec: term,
-   disc_exhausts: thm list,
+   exhaust_discs: thm list,
    sel_defs: thm list,
    fp_nesting_maps: thm list,
    fp_nesting_map_ident0s: thm list,
@@ -159,7 +159,7 @@
             (case fastype_of1 (bound_Ts, nth args n) of
               Type (s, Ts) =>
               (case dest_case ctxt s Ts t of
-                SOME ({sel_splits = _ :: _, ...}, conds', branches) =>
+                SOME ({split_sels = _ :: _, ...}, conds', branches) =>
                 fold_rev (uncurry fld) (map (append conds o conjuncts_s) conds' ~~ branches)
               | _ => f conds t)
             | _ => f conds t)
@@ -173,7 +173,7 @@
 
 fun case_of ctxt s =
   (case ctr_sugar_of ctxt s of
-    SOME {casex = Const (s', _), sel_splits = _ :: _, ...} => SOME s'
+    SOME {casex = Const (s', _), split_sels = _ :: _, ...} => SOME s'
   | _ => NONE);
 
 fun massage_let_if_case ctxt has_call massage_leaf =
@@ -198,7 +198,7 @@
           end
         | (c as Const (@{const_name case_prod}, _), arg :: args) =>
           massage_rec bound_Ts
-            (unfold_lets_splits (Term.list_comb (c $ Envir.eta_long bound_Ts arg, args)))
+            (unfold_splits_lets (Term.list_comb (c $ Envir.eta_long bound_Ts arg, args)))
         | (Const (c, _), args as _ :: _ :: _) =>
           (case try strip_fun_type (Sign.the_const_type thy c) of
             SOME (gen_branch_Ts, gen_body_fun_T) =>
@@ -343,8 +343,8 @@
 
 fun case_thms_of_term ctxt t =
   let val ctr_sugars = map_filter (Ctr_Sugar.ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in
-    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #disc_exhausts ctr_sugars,
-     maps #sel_splits ctr_sugars, maps #sel_split_asms ctr_sugars)
+    (maps #distincts ctr_sugars, maps #discIs ctr_sugars, maps #exhaust_discs ctr_sugars,
+     maps #split_sels ctr_sugars, maps #split_sel_asms ctr_sugars)
   end;
 
 fun basic_corec_specs_of ctxt res_T =
@@ -444,32 +444,32 @@
          else No_Corec) g_i
       | call_of _ [q_i] [g_i, g_i'] _ = Mutual_Corec (q_i, g_i, g_i');
 
-    fun mk_ctr_spec ctr disc sels p_io q_iss f_iss f_Tss discI sel_thms disc_excludess collapse
-        corec_thm disc_corec sel_corecs =
+    fun mk_ctr_spec ctr disc sels p_io q_iss f_iss f_Tss discI sel_thms distinct_discss collapse
+        corec_thm corec_disc corec_sels =
       let val nullary = not (can dest_funT (fastype_of ctr)) in
         {ctr = substA ctr, disc = substA disc, sels = map substA sels, pred = p_io,
          calls = map3 (call_of nullary) q_iss f_iss f_Tss, discI = discI, sel_thms = sel_thms,
-         disc_excludess = disc_excludess, collapse = collapse, corec_thm = corec_thm,
-         disc_corec = disc_corec, sel_corecs = sel_corecs}
+         distinct_discss = distinct_discss, collapse = collapse, corec_thm = corec_thm,
+         corec_disc = corec_disc, corec_sels = corec_sels}
       end;
 
-    fun mk_ctr_specs ({ctrs, discs, selss, discIs, sel_thmss, disc_excludesss, collapses, ...}
-        : ctr_sugar) p_is q_isss f_isss f_Tsss corec_thms disc_corecs sel_corecss =
+    fun mk_ctr_specs ({ctrs, discs, selss, discIs, sel_thmss, distinct_discsss, collapses, ...}
+        : ctr_sugar) p_is q_isss f_isss f_Tsss corec_thms corec_discs corec_selss =
       let val p_ios = map SOME p_is @ [NONE] in
         map14 mk_ctr_spec ctrs discs selss p_ios q_isss f_isss f_Tsss discIs sel_thmss
-          disc_excludesss collapses corec_thms disc_corecs sel_corecss
+          distinct_discsss collapses corec_thms corec_discs corec_selss
       end;
 
-    fun mk_spec ({T, ctr_sugar as {disc_exhausts, sel_defs, ...}, co_rec = corec,
-        co_rec_thms = corec_thms, disc_co_recs = disc_corecs,
-        sel_co_recss = sel_corecss, ...} : fp_sugar) p_is q_isss f_isss f_Tsss =
-      {corec = mk_co_rec thy Greatest_FP (substAT T) perm_Cs' corec, disc_exhausts = disc_exhausts,
+    fun mk_spec ({T, ctr_sugar as {exhaust_discs, sel_defs, ...}, co_rec = corec,
+        co_rec_thms = corec_thms, disc_co_recs = corec_discs,
+        sel_co_recss = corec_selss, ...} : fp_sugar) p_is q_isss f_isss f_Tsss =
+      {corec = mk_co_rec thy Greatest_FP (substAT T) perm_Cs' corec, exhaust_discs = exhaust_discs,
        sel_defs = sel_defs,
        fp_nesting_maps = maps (map_thms_of_typ lthy o T_of_bnf) fp_nesting_bnfs,
        fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs,
        fp_nesting_map_comps = map map_comp_of_bnf fp_nesting_bnfs,
-       ctr_specs = mk_ctr_specs ctr_sugar p_is q_isss f_isss f_Tsss corec_thms disc_corecs
-         sel_corecss};
+       ctr_specs = mk_ctr_specs ctr_sugar p_is q_isss f_isss f_Tsss corec_thms corec_discs
+         corec_selss};
   in
     ((is_some gfp_sugar_thms, map5 mk_spec fp_sugars p_iss q_issss f_issss f_Tssss, missing_res_Ts,
       co_induct_of common_coinduct_thms, strong_co_induct_of common_coinduct_thms,
@@ -694,14 +694,6 @@
         handle ListPair.UnequalLengths =>
           primcorec_error_eqn "partially applied constructor in right-hand side" rhs;
 
-(*
-val _ = tracing ("reduced\n    " ^ Syntax.string_of_term @{context} concl ^ "\nto\n    \<cdot> " ^
- (is_some eqn_data_disc_opt ? K (Syntax.string_of_term @{context} disc_concl ^ "\n    \<cdot> ")) "" ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) sel_concls) ^
- "\nfor premise(s)\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term @{context}) prems));
-*)
-
     val eqns_data_sel =
       map (dissect_coeqn_sel lthy fun_names basic_ctr_specss eqn_pos
         (SOME (abstract (List.rev fun_args) rhs)) code_rhs_opt eqn0 (SOME ctr)) sel_concls;
@@ -890,11 +882,6 @@
       | currys Ts t = t $ mk_tuple1_balanced (List.rev Ts) (map Bound (length Ts - 1 downto 0))
           |> fold_rev (Term.abs o pair Name.uu) Ts;
 
-(*
-val _ = tracing ("corecursor arguments:\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (map (Syntax.string_of_term lthy) corec_args));
-*)
-
     val excludess' =
       disc_eqnss
       |> map (map (fn x => (#fun_args x, #ctr_no x, #prems x, #auto_gen x))
@@ -1006,12 +993,8 @@
       |> map2 (curry (op |>)) (map (map (fn {ctr, sels, ...} =>
         (ctr, map (K []) sels))) basic_ctr_specss);
 
-(*
-val _ = tracing ("callssss = " ^ @{make_string} callssss);
-*)
-
-    val ((n2m, corec_specs', _, coinduct_thm, strong_coinduct_thm, coinduct_thms,
-          strong_coinduct_thms), lthy') =
+    val ((n2m, corec_specs', _, coinduct_thm, coinduct_strong_thm, coinduct_thms,
+          coinduct_strong_thms), lthy') =
       corec_specs_of bs arg_Ts res_Ts frees callssss lthy;
     val corec_specs = take actual_nn corec_specs';
     val ctr_specss = map #ctr_specs corec_specs;
@@ -1051,11 +1034,6 @@
       else
         tac_opt;
 
-(*
-val _ = tracing ("exclusiveness properties:\n    \<cdot> " ^
- space_implode "\n    \<cdot> " (maps (map (Syntax.string_of_term lthy o snd)) excludess'));
-*)
-
     val excludess'' = map3 (fn tac_opt => fn sequential => map (fn (j, goal) =>
         (j, (Option.map (Goal.prove_sorry lthy [] [] goal #> Thm.close_derivation)
            (exclude_tac tac_opt sequential j), goal))))
@@ -1083,15 +1061,15 @@
         de_facto_exhaustives disc_eqnss
       |> list_all_fun_args []
     val nchotomy_taut_thmss =
-      map5 (fn tac_opt => fn {disc_exhausts = res_disc_exhausts, ...} =>
+      map5 (fn tac_opt => fn {exhaust_discs = res_exhaust_discs, ...} =>
           fn {code_rhs_opt, ...} :: _ => fn [] => K []
             | [goal] => fn true =>
               let
-                val (_, _, arg_disc_exhausts, _, _) =
+                val (_, _, arg_exhaust_discs, _, _) =
                   case_thms_of_term lthy (the_default Term.dummy code_rhs_opt);
               in
                 [Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
-                   mk_primcorec_nchotomy_tac ctxt (res_disc_exhausts @ arg_disc_exhausts))
+                   mk_primcorec_nchotomy_tac ctxt (res_exhaust_discs @ arg_exhaust_discs))
                  |> Thm.close_derivation]
               end
             | false =>
@@ -1155,7 +1133,7 @@
             []
           else
             let
-              val {disc, disc_corec, ...} = nth ctr_specs ctr_no;
+              val {disc, corec_disc, ...} = nth ctr_specs ctr_no;
               val k = 1 + ctr_no;
               val m = length prems;
               val goal =
@@ -1168,7 +1146,7 @@
               if prems = [@{term False}] then
                 []
               else
-                mk_primcorec_disc_tac lthy def_thms disc_corec k m excludesss
+                mk_primcorec_disc_tac lthy def_thms corec_disc k m excludesss
                 |> K |> Goal.prove_sorry lthy [] [] goal
                 |> Thm.close_derivation
                 |> pair (#disc (nth ctr_specs ctr_no))
@@ -1185,8 +1163,8 @@
             val ctr_no = find_index (curry (op =) ctr o #ctr) ctr_specs;
             val prems = the_default (maps (s_not_conj o #prems) disc_eqns)
               (find_first (curry (op =) ctr_no o #ctr_no) disc_eqns |> Option.map #prems);
-            val sel_corec = find_index (curry (op =) sel) (#sels ctr_spec)
-              |> nth (#sel_corecs ctr_spec);
+            val corec_sel = find_index (curry (op =) sel) (#sels ctr_spec)
+              |> nth (#corec_sels ctr_spec);
             val k = 1 + ctr_no;
             val m = length prems;
             val goal =
@@ -1196,10 +1174,10 @@
               |> HOLogic.mk_Trueprop o HOLogic.mk_eq
               |> curry Logic.list_implies (map HOLogic.mk_Trueprop prems)
               |> curry Logic.list_all (map dest_Free fun_args);
-            val (distincts, _, _, sel_splits, sel_split_asms) = case_thms_of_term lthy rhs_term;
+            val (distincts, _, _, split_sels, split_sel_asms) = case_thms_of_term lthy rhs_term;
           in
-            mk_primcorec_sel_tac lthy def_thms distincts sel_splits sel_split_asms fp_nesting_maps
-              fp_nesting_map_ident0s fp_nesting_map_comps sel_corec k m excludesss
+            mk_primcorec_sel_tac lthy def_thms distincts split_sels split_sel_asms fp_nesting_maps
+              fp_nesting_map_ident0s fp_nesting_map_comps corec_sel k m excludesss
             |> K |> Goal.prove_sorry lthy [] [] goal
             |> Thm.close_derivation
             |> `(is_some code_rhs_opt ? fold_thms lthy sel_defs) (*mildly too aggressive*)
@@ -1328,16 +1306,16 @@
                     val (raw_goal, goal) = (raw_rhs, rhs)
                       |> pairself (curry mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args)
                         #> curry Logic.list_all (map dest_Free fun_args));
-                    val (distincts, discIs, _, sel_splits, sel_split_asms) =
+                    val (distincts, discIs, _, split_sels, split_sel_asms) =
                       case_thms_of_term lthy raw_rhs;
 
-                    val raw_code_thm = mk_primcorec_raw_code_tac lthy distincts discIs sel_splits
-                        sel_split_asms ms ctr_thms
+                    val raw_code_thm = mk_primcorec_raw_code_tac lthy distincts discIs split_sels
+                        split_sel_asms ms ctr_thms
                         (if exhaustive_code then try the_single nchotomys else NONE)
                       |> K |> Goal.prove_sorry lthy [] [] raw_goal
                       |> Thm.close_derivation;
                   in
-                    mk_primcorec_code_tac lthy distincts sel_splits raw_code_thm
+                    mk_primcorec_code_tac lthy distincts split_sels raw_code_thm
                     |> K |> Goal.prove_sorry lthy [] [] goal
                     |> Thm.close_derivation
                     |> single
@@ -1359,14 +1337,14 @@
             []
           else
             let
-              val {disc, disc_excludess, ...} = nth ctr_specs ctr_no;
+              val {disc, distinct_discss, ...} = nth ctr_specs ctr_no;
               val goal =
                 mk_Trueprop_eq (applied_fun_of fun_name fun_T fun_args |> curry betapply disc,
                   mk_conjs prems)
                 |> curry Logic.list_all (map dest_Free fun_args);
             in
               mk_primcorec_disc_iff_tac lthy (map (fst o dest_Free) exhaust_fun_args)
-                (the_single exhaust_thms) disc_thms disc_thmss' (flat disc_excludess)
+                (the_single exhaust_thms) disc_thms disc_thmss' (flat distinct_discss)
               |> K |> Goal.prove_sorry lthy [] [] goal
               |> Thm.close_derivation
               |> fold (fn rule => perhaps (try (fn thm => Meson.first_order_resolve thm rule)))
@@ -1407,7 +1385,7 @@
            (exhaustN, nontriv_exhaust_thmss, []),
            (selN, sel_thmss, simp_attrs),
            (simpsN, simp_thmss, []),
-           (strong_coinductN, map (if n2m then single else K []) strong_coinduct_thms, [])]
+           (coinduct_strongN, map (if n2m then single else K []) coinduct_strong_thms, [])]
           |> maps (fn (thmN, thmss, attrs) =>
             map2 (fn fun_name => fn thms =>
                 ((Binding.qualify true fun_name (Binding.name thmN), attrs), [(thms, [])]))
@@ -1416,7 +1394,7 @@
 
         val common_notes =
           [(coinductN, if n2m then [coinduct_thm] else [], []),
-           (strong_coinductN, if n2m then [strong_coinduct_thm] else [], [])]
+           (coinduct_strongN, if n2m then [coinduct_strong_thm] else [], [])]
           |> filter_out (null o #2)
           |> map (fn (thmN, thms, attrs) =>
             ((Binding.qualify true common_name (Binding.name thmN), attrs), [(thms, [])]));
--- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -54,8 +54,8 @@
 fun distinct_in_prems_tac distincts =
   eresolve_tac (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' atac;
 
-fun mk_primcorec_nchotomy_tac ctxt disc_exhausts =
-  HEADGOAL (Method.insert_tac disc_exhausts THEN' clean_blast_tac ctxt);
+fun mk_primcorec_nchotomy_tac ctxt exhaust_discs =
+  HEADGOAL (Method.insert_tac exhaust_discs THEN' clean_blast_tac ctxt);
 
 fun mk_primcorec_exhaust_tac ctxt frees n nchotomy =
   let val ks = 1 upto n in
@@ -105,11 +105,11 @@
 fun prelude_tac ctxt defs thm =
   unfold_thms_tac ctxt defs THEN HEADGOAL (rtac thm) THEN unfold_thms_tac ctxt unfold_lets;
 
-fun mk_primcorec_disc_tac ctxt defs disc_corec k m excludesss =
-  prelude_tac ctxt defs disc_corec THEN cases_tac ctxt k m excludesss;
+fun mk_primcorec_disc_tac ctxt defs corec_disc k m excludesss =
+  prelude_tac ctxt defs corec_disc THEN cases_tac ctxt k m excludesss;
 
 fun mk_primcorec_disc_iff_tac ctxt fun_exhaust_frees fun_exhaust fun_discs fun_discss
-    disc_excludes =
+    distinct_discs =
   HEADGOAL (rtac iffI THEN'
     rtac fun_exhaust THEN'
     K (exhaust_inst_as_projs_tac ctxt fun_exhaust_frees) THEN'
@@ -121,7 +121,7 @@
             rtac FalseE THEN'
             (rotate_tac 1 THEN' dtac fun_disc' THEN' REPEAT_DETERM o atac ORELSE'
              cut_tac fun_disc') THEN'
-            dresolve_tac disc_excludes THEN' etac notE THEN' atac)
+            dresolve_tac distinct_discs THEN' etac notE THEN' atac)
       fun_discss) THEN'
     (etac FalseE ORELSE'
      resolve_tac (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' atac));
--- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -103,7 +103,7 @@
 val ord_eq_le_trans = @{thm ord_eq_le_trans};
 val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans};
 val snd_convol_fun_cong_sym = @{thm snd_convol[unfolded convol_def]} RS fun_cong RS sym;
-val sum_weak_case_cong = @{thm sum.weak_case_cong};
+val sum_case_cong_weak = @{thm sum.case_cong_weak};
 val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
 val Collect_splitD_set_mp = @{thm Collect_splitD[OF set_mp]};
 val rev_bspec = Drule.rotate_prems 1 bspec;
@@ -277,21 +277,21 @@
         REPEAT_DETERM o etac allE,
         rtac @{thm conversepI}, etac mp, etac @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1;
 
-fun mk_bis_O_tac ctxt m bis_rel rel_congs rel_OOs =
+fun mk_bis_O_tac ctxt m bis_rel rel_congs le_rel_OOs =
   EVERY' [rtac (bis_rel RS iffD2), REPEAT_DETERM o dtac (bis_rel RS iffD1),
     REPEAT_DETERM o etac conjE, rtac conjI,
     CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) rel_congs,
-    CONJ_WRAP' (fn (rel_cong, rel_OO) =>
+    CONJ_WRAP' (fn (rel_cong, le_rel_OO) =>
       EVERY' [rtac allI, rtac allI, rtac impI,
         rtac (rel_cong RS @{thm eq_refl} RS @{thm predicate2D}),
         REPEAT_DETERM_N m o rtac @{thm eq_OO},
         REPEAT_DETERM_N (length rel_congs) o rtac @{thm relcompp_in_rel},
-        rtac (rel_OO RS sym RS @{thm eq_refl} RS @{thm predicate2D}),
+        rtac (le_rel_OO RS @{thm predicate2D}),
         etac @{thm relcompE},
         REPEAT_DETERM o dtac prod_injectD,
         etac conjE, hyp_subst_tac ctxt,
         REPEAT_DETERM o etac allE, rtac @{thm relcomppI},
-        etac mp, atac, etac mp, atac]) (rel_congs ~~ rel_OOs)] 1;
+        etac mp, atac, etac mp, atac]) (rel_congs ~~ le_rel_OOs)] 1;
 
 fun mk_bis_Gr_tac ctxt bis_rel rel_Grps mor_images morEs coalg_ins =
   unfold_thms_tac ctxt (bis_rel :: @{thm eq_alt} :: @{thm in_rel_Gr} :: rel_Grps) THEN
@@ -430,7 +430,7 @@
         CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
           CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI,
             rtac (@{thm append_Cons} RS arg_cong RS trans), rtac (rv_Cons RS trans),
-            if n = 1 then K all_tac else etac (sum_weak_case_cong RS trans),
+            if n = 1 then K all_tac else etac (sum_case_cong_weak RS trans),
             rtac (mk_sum_caseN n i RS trans), atac])
           ks])
         rv_Conss)
@@ -565,7 +565,7 @@
               rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS iffD2),
               rtac exI, rtac conjI,
               if n = 1 then rtac refl
-              else etac (sum_weak_case_cong RS trans) THEN' rtac (mk_sum_caseN n i),
+              else etac (sum_case_cong_weak RS trans) THEN' rtac (mk_sum_caseN n i),
               CONJ_WRAP' (fn (set_map0, (set_Lev, set_image_Lev)) =>
                 EVERY' [rtac (set_map0 RS trans), rtac equalityI, rtac @{thm image_subsetI},
                   rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
@@ -603,7 +603,7 @@
         CONVERSION (Conv.top_conv
           (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
         if n = 1 then K all_tac
-        else (rtac (sum_weak_case_cong RS trans) THEN'
+        else (rtac (sum_case_cong_weak RS trans) THEN'
           rtac (mk_sum_caseN n i) THEN' rtac (mk_sum_caseN n i RS trans)),
         rtac (map_comp_id RS trans), rtac (map_cong0 OF replicate m refl),
         EVERY' (map3 (fn i' => fn to_sbd_inj => fn from_to_sbd =>
@@ -628,7 +628,7 @@
             CONVERSION (Conv.top_conv
               (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
             if n = 1 then K all_tac
-            else rtac sum_weak_case_cong THEN' rtac (mk_sum_caseN n i' RS trans),
+            else rtac sum_case_cong_weak THEN' rtac (mk_sum_caseN n i' RS trans),
             SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl])
         ks to_sbd_injs from_to_sbds)];
   in
@@ -898,14 +898,14 @@
   EVERY' (map rtac [@{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
     @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite]) 1;
 
-fun mk_le_rel_OO_tac coinduct rel_Jrels rel_OOs =
-  EVERY' (rtac coinduct :: map2 (fn rel_Jrel => fn rel_OO =>
+fun mk_le_rel_OO_tac coinduct rel_Jrels le_rel_OOs =
+  EVERY' (rtac coinduct :: map2 (fn rel_Jrel => fn le_rel_OO =>
     let val Jrel_imp_rel = rel_Jrel RS iffD1;
     in
-      EVERY' [rtac (rel_OO RS sym RS @{thm eq_refl} RS @{thm predicate2D}), etac @{thm relcomppE},
+      EVERY' [rtac (le_rel_OO RS @{thm predicate2D}), etac @{thm relcomppE},
       rtac @{thm relcomppI}, etac Jrel_imp_rel, etac Jrel_imp_rel]
     end)
-  rel_Jrels rel_OOs) 1;
+  rel_Jrels le_rel_OOs) 1;
 
 fun mk_wit_tac ctxt n dtor_ctors dtor_set wit coind_wits =
   ALLGOALS (TRY o (eresolve_tac coind_wits THEN' rtac refl)) THEN
@@ -973,8 +973,6 @@
                 hyp_subst_tac ctxt,
                 dtac (in_Jrel RS iffD1),
                 dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
-                TRY o
-                  EVERY' [dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac, hyp_subst_tac ctxt],
                 REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
             (rev (active_set_map0s ~~ in_Jrels))])
         (dtor_sets ~~ passive_set_map0s),
--- a/src/HOL/Tools/BNF/bnf_lfp.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -167,8 +167,8 @@
     val map_id0s = map map_id0_of_bnf bnfs;
     val map_ids = map map_id_of_bnf bnfs;
     val set_mapss = map set_map_of_bnf bnfs;
-    val rel_mono_strongs = map rel_mono_strong_of_bnf bnfs;
-    val rel_OOs = map rel_OO_of_bnf bnfs;
+    val rel_mono_strong0s = map rel_mono_strong0_of_bnf bnfs;
+    val le_rel_OOs = map le_rel_OO_of_bnf bnfs;
 
     val timer = time (timer "Extracted terms & thms");
 
@@ -1197,7 +1197,7 @@
     val (ctor_rec_unique_thms, ctor_rec_unique_thm) =
       `split_conj_thm (split_conj_prems n
         (mor_UNIV_thm RS iffD2 RS rec_unique_mor_thm)
-        |> Local_Defs.unfold lthy (@{thms convol_o comp_id id_comp comp_assoc fst_convol} @
+        |> unfold_thms lthy (@{thms convol_o comp_id id_comp comp_assoc fst_convol} @
            map_id0s @ sym_map_comps) OF replicate n @{thm arg_cong2[of _ _ _ _ convol, OF refl]});
 
     val timer = time (timer "rec definitions & thms");
@@ -1691,7 +1691,7 @@
           in
             Goal.prove_sorry lthy [] [] goal
               (fn {context = ctxt, prems = _} => mk_le_rel_OO_tac ctxt m induct ctor_nchotomy_thms
-                ctor_Irel_thms rel_mono_strongs rel_OOs)
+                ctor_Irel_thms rel_mono_strong0s le_rel_OOs)
             |> Thm.close_derivation
             |> singleton (Proof_Context.export names_lthy lthy)
           end;
@@ -1762,7 +1762,7 @@
     val Irel_induct_thm =
       mk_rel_xtor_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's
         (fn {context = ctxt, prems = IHs} => mk_rel_induct_tac ctxt IHs m ctor_induct2_thm ks
-           ctor_Irel_thms rel_mono_strongs) lthy;
+           ctor_Irel_thms rel_mono_strong0s) lthy;
 
     val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs;
     val ctor_fold_transfer_thms =
@@ -1801,12 +1801,7 @@
           ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]))
         bs thmss);
 
-    (*FIXME: once the package exports all the necessary high-level characteristic theorems,
-       those should not only be concealed but rather not noted at all*)
-    val maybe_conceal_notes = note_all = false ? map (apfst (apfst Binding.conceal));
-
-    val (noted, lthy') =
-      lthy |> Local_Theory.notes (maybe_conceal_notes (common_notes @ notes @ Ibnf_notes));
+    val lthy' = lthy |> note_all ? snd o Local_Theory.notes (common_notes @ notes @ Ibnf_notes);
 
     val fp_res =
       {Ts = Ts, bnfs = Ibnfs, ctors = ctors, dtors = dtors, xtor_co_recs = recs,
@@ -1814,8 +1809,8 @@
        ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms,
        xtor_map_thms = ctor_Imap_thms, xtor_set_thmss = ctor_Iset_thmss',
        xtor_rel_thms = ctor_Irel_thms, xtor_co_rec_thms = ctor_rec_thms,
-       xtor_co_rec_o_map_thms = ctor_rec_o_Imap_thms, rel_xtor_co_induct_thm = Irel_induct_thm}
-      |> morph_fp_result (substitute_noted_thm noted);
+       xtor_co_rec_o_map_thms = ctor_rec_o_Imap_thms, rel_xtor_co_induct_thm = Irel_induct_thm,
+       dtor_set_induct_thms = []};
   in
     timer; (fp_res, lthy')
   end;
--- a/src/HOL/Tools/BNF/bnf_lfp_compat.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_compat.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -80,10 +80,27 @@
     val (orig_descr' :: nested_descrs, _) =
       Datatype_Aux.unfold_datatypes lthy orig_descr all_infos orig_descr nn_fp;
 
+    fun cliquify_descr [] = []
+      | cliquify_descr [entry] = [[entry]]
+      | cliquify_descr (full_descr as (_, (T_name1, _, _)) :: _) =
+        let
+          val nn =
+            if member (op =) fpT_names T_name1 then
+              nn_fp
+            else
+              (case Symtab.lookup all_infos T_name1 of
+                SOME {descr, ...} =>
+                length (filter_out (exists Datatype_Aux.is_rec_type o #2 o snd) descr)
+              | NONE => raise Fail "unknown old-style datatype");
+        in
+          chop nn full_descr ||> cliquify_descr |> op ::
+        end;
+
     (* put nested types before the types that nest them, as needed for N2M *)
     val descrs = burrow reindex_desc (orig_descr' :: rev nested_descrs);
     val (cliques, descr) =
-      split_list (flat (map_index (fn (i, descr) => map (pair i) descr) descrs));
+      split_list (flat (map_index (fn (i, descr) => map (pair i) descr)
+        (maps cliquify_descr descrs)));
 
     val dest_dtyp = Datatype_Aux.typ_of_dtyp descr;
 
@@ -121,13 +138,13 @@
     val inducts = map (the_single o #co_inducts) fp_sugars;
 
     fun mk_info (kk, {T = Type (T_name0, _), ctr_sugar = {casex, exhaust, nchotomy, injects,
-        distincts, case_thms, case_cong, weak_case_cong, split, split_asm, ...}, ...} : fp_sugar) =
+        distincts, case_thms, case_cong, case_cong_weak, split, split_asm, ...}, ...} : fp_sugar) =
       (T_name0,
        {index = kk, descr = descr, inject = injects, distinct = distincts, induct = induct,
-       inducts = inducts, exhaust = exhaust, nchotomy = nchotomy, rec_names = recs,
-       rec_rewrites = rec_thms, case_name = fst (dest_Const casex), case_rewrites = case_thms,
-       case_cong = case_cong, weak_case_cong = weak_case_cong, split = split,
-       split_asm = split_asm});
+        inducts = inducts, exhaust = exhaust, nchotomy = nchotomy, rec_names = recs,
+        rec_rewrites = rec_thms, case_name = fst (dest_Const casex), case_rewrites = case_thms,
+        case_cong = case_cong, case_cong_weak = case_cong_weak, split = split,
+        split_asm = split_asm});
 
     val infos = map_index mk_info (take nn_fp fp_sugars);
 
--- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar_more.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -16,7 +16,7 @@
 open BNF_FP_N2M_Sugar
 open BNF_LFP_Rec_Sugar
 
-val nested_simps = @{thms id_def split comp_def fst_conv snd_conv};
+val nested_simps = @{thms o_def[abs_def] id_def split fst_conv snd_conv};
 
 fun is_new_datatype ctxt s =
   (case fp_sugar_of ctxt s of SOME {fp = Least_FP, ...} => true | _ => false);
--- a/src/HOL/Tools/BNF/bnf_lfp_size.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -64,15 +64,17 @@
   funpow n (fn thm => thm RS @{thm fun_cong_unused_0} handle THM _ => thm RS fun_cong);
 
 val rec_o_map_simp_thms =
-  @{thms o_def id_def case_prod_app case_sum_map_sum case_prod_map_prod BNF_Comp.id_bnf_comp_def};
+  @{thms o_def[abs_def] id_def case_prod_app case_sum_map_sum case_prod_map_prod
+      BNF_Comp.id_bnf_comp_def};
 
 fun mk_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s abs_inverses
     ctor_rec_o_map =
   unfold_thms_tac ctxt [rec_def] THEN
-  HEADGOAL (rtac (ctor_rec_o_map RS trans)) THEN
-  PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion) THEN
-  HEADGOAL (asm_simp_tac (ss_only (pre_map_defs @
-    distinct Thm.eq_thm_prop (live_nesting_map_ident0s @ abs_inverses) @ rec_o_map_simp_thms) ctxt));
+  HEADGOAL (rtac (ctor_rec_o_map RS trans) THEN'
+    CONVERSION Thm.eta_long_conversion THEN'
+    asm_simp_tac (ss_only (pre_map_defs @
+        distinct Thm.eq_thm_prop (live_nesting_map_ident0s @ abs_inverses) @ rec_o_map_simp_thms)
+      ctxt));
 
 val size_o_map_simp_thms = @{thms prod_inj_map inj_on_id snd_comp_apfst[unfolded apfst_def]};
 
@@ -303,7 +305,7 @@
               curry HOLogic.mk_eq) rec_o_map_lhss rec_o_map_rhss;
           val rec_o_map_thms =
             map3 (fn goal => fn rec_def => fn ctor_rec_o_map =>
-                Goal.prove lthy2 [] [] goal (fn {context = ctxt, ...} =>
+                Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} =>
                   mk_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s abs_inverses
                     ctor_rec_o_map)
                 |> Thm.close_derivation)
@@ -333,7 +335,7 @@
              in terms of "map", which is not the end of the world. *)
           val size_o_map_thmss =
             map3 (fn goal => fn size_def => the_list o try (fn rec_o_map =>
-                Goal.prove lthy2 [] [] goal (fn {context = ctxt, ...} =>
+                Goal.prove (*no sorry*) lthy2 [] [] goal (fn {context = ctxt, ...} =>
                   mk_size_o_map_tac ctxt size_def rec_o_map all_inj_maps nested_size_maps)
                 |> Thm.close_derivation))
               size_o_map_goals size_defs rec_o_map_thms
--- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -582,17 +582,17 @@
     (induct_tac THEN' EVERY' (map3 mk_map_cong0 ctor_maps map_cong0s set_setsss)) 1
   end;
 
-fun mk_le_rel_OO_tac ctxt m induct ctor_nchotomys ctor_Irels rel_mono_strongs rel_OOs =
+fun mk_le_rel_OO_tac ctxt m induct ctor_nchotomys ctor_Irels rel_mono_strong0s le_rel_OOs =
   EVERY' (rtac induct ::
-  map4 (fn nchotomy => fn Irel => fn rel_mono => fn rel_OO =>
+  map4 (fn nchotomy => fn Irel => fn rel_mono => fn le_rel_OO =>
     EVERY' [rtac impI, etac (nchotomy RS @{thm nchotomy_relcomppE}),
       REPEAT_DETERM_N 2 o dtac (Irel RS iffD1), rtac (Irel RS iffD2),
-      rtac rel_mono, rtac (rel_OO RS @{thm predicate2_eqD} RS iffD2),
+      rtac rel_mono, rtac (le_rel_OO RS @{thm predicate2D}),
       rtac @{thm relcomppI}, atac, atac,
       REPEAT_DETERM_N m o EVERY' [rtac ballI, rtac ballI, rtac impI, atac],
-      REPEAT_DETERM_N (length rel_OOs) o
+      REPEAT_DETERM_N (length le_rel_OOs) o
         EVERY' [rtac ballI, rtac ballI, Goal.assume_rule_tac ctxt]])
-  ctor_nchotomys ctor_Irels rel_mono_strongs rel_OOs) 1;
+  ctor_nchotomys ctor_Irels rel_mono_strong0s le_rel_OOs) 1;
 
 (* BNF tactics *)
 
@@ -682,8 +682,6 @@
                   @{thms case_prodE iffD1[OF prod.inject, elim_format]}),
                 hyp_subst_tac ctxt,
                 dtac (in_Irel RS iffD1), dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
-                TRY o
-                  EVERY' [dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac, hyp_subst_tac ctxt],
                 REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
             (rev (active_set_map0s ~~ in_Irels))])
         (ctor_sets ~~ passive_set_map0s),
@@ -703,19 +701,19 @@
     EVERY' [rtac iffI, if_tac, only_if_tac] 1
   end;
 
-fun mk_rel_induct_tac ctxt IHs m ctor_induct2 ks ctor_rels rel_mono_strongs =
+fun mk_rel_induct_tac ctxt IHs m ctor_induct2 ks ctor_rels rel_mono_strong0s =
   let val n = length ks;
   in
     unfold_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN
     EVERY' [REPEAT_DETERM o rtac allI, rtac ctor_induct2,
-      EVERY' (map3 (fn IH => fn ctor_rel => fn rel_mono_strong =>
+      EVERY' (map3 (fn IH => fn ctor_rel => fn rel_mono_strong0 =>
         EVERY' [rtac impI, dtac (ctor_rel RS iffD1), rtac (IH RS @{thm spec2} RS mp),
-          etac rel_mono_strong,
+          etac rel_mono_strong0,
           REPEAT_DETERM_N m o rtac @{thm ballI[OF ballI[OF imp_refl]]},
           EVERY' (map (fn j =>
             EVERY' [select_prem_tac n (dtac asm_rl) j, rtac @{thm ballI[OF ballI]},
               Goal.assume_rule_tac ctxt]) ks)])
-      IHs ctor_rels rel_mono_strongs)] 1
+      IHs ctor_rels rel_mono_strong0s)] 1
   end;
 
 fun mk_fold_transfer_tac ctxt m ctor_rel_induct map_transfers folds =
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -18,7 +18,7 @@
      distincts: thm list,
      case_thms: thm list,
      case_cong: thm,
-     weak_case_cong: thm,
+     case_cong_weak: thm,
      split: thm,
      split_asm: thm,
      disc_defs: thm list,
@@ -26,13 +26,13 @@
      discIs: thm list,
      sel_defs: thm list,
      sel_thmss: thm list list,
-     disc_excludesss: thm list list list,
-     disc_exhausts: thm list,
-     sel_exhausts: thm list,
+     distinct_discsss: thm list list list,
+     exhaust_discs: thm list,
+     exhaust_sels: thm list,
      collapses: thm list,
      expands: thm list,
-     sel_splits: thm list,
-     sel_split_asms: thm list,
+     split_sels: thm list,
+     split_sel_asms: thm list,
      case_eq_ifs: thm list};
 
   val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
@@ -90,7 +90,7 @@
    distincts: thm list,
    case_thms: thm list,
    case_cong: thm,
-   weak_case_cong: thm,
+   case_cong_weak: thm,
    split: thm,
    split_asm: thm,
    disc_defs: thm list,
@@ -98,19 +98,19 @@
    discIs: thm list,
    sel_defs: thm list,
    sel_thmss: thm list list,
-   disc_excludesss: thm list list list,
-   disc_exhausts: thm list,
-   sel_exhausts: thm list,
+   distinct_discsss: thm list list list,
+   exhaust_discs: thm list,
+   exhaust_sels: thm list,
    collapses: thm list,
    expands: thm list,
-   sel_splits: thm list,
-   sel_split_asms: thm list,
+   split_sels: thm list,
+   split_sel_asms: thm list,
    case_eq_ifs: thm list};
 
 fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
-    case_thms, case_cong, weak_case_cong, split, split_asm, disc_defs, disc_thmss, discIs, sel_defs,
-    sel_thmss, disc_excludesss, disc_exhausts, sel_exhausts, collapses, expands, sel_splits,
-    sel_split_asms, case_eq_ifs} =
+    case_thms, case_cong, case_cong_weak, split, split_asm, disc_defs, disc_thmss, discIs, sel_defs,
+    sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels,
+    split_sel_asms, case_eq_ifs} =
   {ctrs = map (Morphism.term phi) ctrs,
    casex = Morphism.term phi casex,
    discs = map (Morphism.term phi) discs,
@@ -121,7 +121,7 @@
    distincts = map (Morphism.thm phi) distincts,
    case_thms = map (Morphism.thm phi) case_thms,
    case_cong = Morphism.thm phi case_cong,
-   weak_case_cong = Morphism.thm phi weak_case_cong,
+   case_cong_weak = Morphism.thm phi case_cong_weak,
    split = Morphism.thm phi split,
    split_asm = Morphism.thm phi split_asm,
    disc_defs = map (Morphism.thm phi) disc_defs,
@@ -129,13 +129,13 @@
    discIs = map (Morphism.thm phi) discIs,
    sel_defs = map (Morphism.thm phi) sel_defs,
    sel_thmss = map (map (Morphism.thm phi)) sel_thmss,
-   disc_excludesss = map (map (map (Morphism.thm phi))) disc_excludesss,
-   disc_exhausts = map (Morphism.thm phi) disc_exhausts,
-   sel_exhausts = map (Morphism.thm phi) sel_exhausts,
+   distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss,
+   exhaust_discs = map (Morphism.thm phi) exhaust_discs,
+   exhaust_sels = map (Morphism.thm phi) exhaust_sels,
    collapses = map (Morphism.thm phi) collapses,
    expands = map (Morphism.thm phi) expands,
-   sel_splits = map (Morphism.thm phi) sel_splits,
-   sel_split_asms = map (Morphism.thm phi) sel_split_asms,
+   split_sels = map (Morphism.thm phi) split_sels,
+   split_sel_asms = map (Morphism.thm phi) split_sel_asms,
    case_eq_ifs = map (Morphism.thm phi) case_eq_ifs};
 
 val transfer_ctr_sugar =
@@ -200,23 +200,24 @@
 val case_congN = "case_cong";
 val case_eq_ifN = "case_eq_if";
 val collapseN = "collapse";
-val disc_excludeN = "disc_exclude";
-val disc_exhaustN = "disc_exhaust";
 val discN = "disc";
 val discIN = "discI";
 val distinctN = "distinct";
+val distinct_discN = "distinct_disc";
 val exhaustN = "exhaust";
+val exhaust_discN = "exhaust_disc";
 val expandN = "expand";
 val injectN = "inject";
 val nchotomyN = "nchotomy";
 val selN = "sel";
-val sel_exhaustN = "sel_exhaust";
-val sel_splitN = "sel_split";
-val sel_split_asmN = "sel_split_asm";
+val exhaust_selN = "exhaust_sel";
 val splitN = "split";
+val split_asmN = "split_asm";
+val split_selN = "split_sel";
+val split_sel_asmN = "split_sel_asm";
 val splitsN = "splits";
-val split_asmN = "split_asm";
-val weak_case_cong_thmsN = "weak_case_cong";
+val split_selsN = "split_sels";
+val case_cong_weak_thmsN = "case_cong_weak";
 
 val cong_attrs = @{attributes [cong]};
 val dest_attrs = @{attributes [dest]};
@@ -256,12 +257,6 @@
 fun mk_disc_or_sel Ts t =
   subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
 
-fun name_of_const what t =
-  (case head_of t of
-    Const (s, _) => s
-  | Free (s, _) => s
-  | _ => error ("Cannot extract name of " ^ what));
-
 val name_of_ctr = name_of_const "constructor";
 
 fun name_of_disc t =
@@ -683,7 +678,7 @@
               ks goals inject_thmss distinct_thmsss
           end;
 
-        val (case_cong_thm, weak_case_cong_thm) =
+        val (case_cong_thm, case_cong_weak_thm) =
           let
             fun mk_prem xctr xs xf xg =
               fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
@@ -739,9 +734,9 @@
           end;
 
         val (sel_defs, all_sel_thms, sel_thmss, disc_defs, disc_thmss, nontriv_disc_thmss,
-             discI_thms, nontriv_discI_thms, disc_exclude_thms, disc_exclude_thmsss,
-             disc_exhaust_thms, sel_exhaust_thms, all_collapse_thms, safe_collapse_thms,
-             expand_thms, sel_split_thms, sel_split_asm_thms, case_eq_if_thms) =
+             discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss,
+             exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms,
+             expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms) =
           if no_discs_sels then
             ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [])
           else
@@ -833,7 +828,7 @@
                 flat (map2 (fn b => if is_discI_boring b then K [] else single) disc_bindings
                   discI_thms);
 
-              val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_thmsss)) =
+              val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) =
                 let
                   fun mk_goal [] = []
                     | mk_goal [((_, udisc), (_, udisc'))] =
@@ -849,25 +844,25 @@
                   val half_goalss = map mk_goal half_pairss;
                   val half_thmss =
                     map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
-                        fn disc_thm => [prove (mk_half_disc_exclude_tac lthy m discD disc_thm) goal])
+                        fn disc_thm => [prove (mk_half_distinct_disc_tac lthy m discD disc_thm) goal])
                       half_goalss half_pairss (flat disc_thmss');
 
                   val other_half_goalss = map (mk_goal o map swap) half_pairss;
                   val other_half_thmss =
-                    map2 (map2 (prove o mk_other_half_disc_exclude_tac)) half_thmss
+                    map2 (map2 (prove o mk_other_half_distinct_disc_tac)) half_thmss
                       other_half_goalss;
                 in
                   join_halves n half_thmss other_half_thmss ||> `transpose
                   |>> has_alternate_disc_def ? K []
                 end;
 
-              val disc_exhaust_thm =
+              val exhaust_disc_thm =
                 let
                   fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc];
                   val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs));
                 in
                   Goal.prove_sorry lthy [] [] goal (fn _ =>
-                    mk_disc_exhaust_tac n exhaust_thm discI_thms)
+                    mk_exhaust_disc_tac n exhaust_thm discI_thms)
                   |> Thm.close_derivation
                 end;
 
@@ -896,13 +891,13 @@
               val swapped_all_collapse_thms =
                 map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms;
 
-              val sel_exhaust_thm =
+              val exhaust_sel_thm =
                 let
                   fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)];
                   val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs));
                 in
                   Goal.prove_sorry lthy [] [] goal (fn _ =>
-                    mk_sel_exhaust_tac n disc_exhaust_thm swapped_all_collapse_thms)
+                    mk_exhaust_sel_tac n exhaust_disc_thm swapped_all_collapse_thms)
                   |> Thm.close_derivation
                 end;
 
@@ -925,14 +920,14 @@
                     map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss;
                 in
                   Goal.prove_sorry lthy [] [] goal (fn _ =>
-                    mk_expand_tac lthy n ms (inst_thm u disc_exhaust_thm)
-                      (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
-                      disc_exclude_thmsss')
+                    mk_expand_tac lthy n ms (inst_thm u exhaust_disc_thm)
+                      (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss
+                      distinct_disc_thmsss')
                   |> singleton (Proof_Context.export names_lthy lthy)
                   |> Thm.close_derivation
                 end;
 
-              val (sel_split_thm, sel_split_asm_thm) =
+              val (split_sel_thm, split_sel_asm_thm) =
                 let
                   val zss = map (K []) xss;
                   val goal = mk_split_goal usel_ctrs zss usel_fs;
@@ -955,9 +950,9 @@
                 end;
             in
               (sel_defs, all_sel_thms, sel_thmss, disc_defs, disc_thmss, nontriv_disc_thmss,
-               discI_thms, nontriv_discI_thms, disc_exclude_thms, disc_exclude_thmsss,
-               [disc_exhaust_thm], [sel_exhaust_thm], all_collapse_thms, safe_collapse_thms,
-               [expand_thm], [sel_split_thm], [sel_split_asm_thm], [case_eq_if_thm])
+               discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss,
+               [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms,
+               [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm])
             end;
 
         val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
@@ -972,30 +967,29 @@
            (flat nontriv_disc_eq_thmss, code_nitpicksimp_attrs)]
           |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
 
-        (* "exhaust" is deliberately put first to avoid apparent circular dependencies in the proof
-           objects, which would confuse MaSh. *)
         val notes =
-          [(exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
-           (caseN, case_thms, code_nitpicksimp_simp_attrs),
+          [(caseN, case_thms, code_nitpicksimp_simp_attrs),
            (case_congN, [case_cong_thm], []),
+           (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs),
            (case_eq_ifN, case_eq_if_thms, []),
            (collapseN, safe_collapse_thms, simp_attrs),
            (discN, flat nontriv_disc_thmss, simp_attrs),
            (discIN, nontriv_discI_thms, []),
-           (disc_excludeN, disc_exclude_thms, dest_attrs),
-           (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
            (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs),
+           (distinct_discN, distinct_disc_thms, dest_attrs),
+           (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
+           (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]),
+           (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]),
            (expandN, expand_thms, []),
            (injectN, inject_thms, iff_attrs @ inductsimp_attrs),
            (nchotomyN, [nchotomy_thm], []),
            (selN, all_sel_thms, code_nitpicksimp_simp_attrs),
-           (sel_exhaustN, sel_exhaust_thms, [exhaust_case_names_attr]),
-           (sel_splitN, sel_split_thms, []),
-           (sel_split_asmN, sel_split_asm_thms, []),
            (splitN, [split_thm], []),
            (split_asmN, [split_asm_thm], []),
-           (splitsN, [split_thm, split_asm_thm], []),
-           (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
+           (split_selN, split_sel_thms, []),
+           (split_sel_asmN, split_sel_asm_thms, []),
+           (split_selsN, split_sel_thms @ split_sel_asm_thms, []),
+           (splitsN, [split_thm, split_asm_thm], [])]
           |> filter_out (null o #2)
           |> map (fn (thmN, thms, attrs) =>
             ((qualify true (Binding.name thmN), attrs), [(thms, [])]));
@@ -1025,12 +1019,12 @@
         val ctr_sugar =
           {ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm,
            nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms,
-           case_thms = case_thms, case_cong = case_cong_thm, weak_case_cong = weak_case_cong_thm,
+           case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm,
            split = split_thm, split_asm = split_asm_thm, disc_defs = disc_defs,
            disc_thmss = disc_thmss, discIs = discI_thms, sel_defs = sel_defs, sel_thmss = sel_thmss,
-           disc_excludesss = disc_exclude_thmsss, disc_exhausts = disc_exhaust_thms,
-           sel_exhausts = sel_exhaust_thms, collapses = all_collapse_thms, expands = expand_thms,
-           sel_splits = sel_split_thms, sel_split_asms = sel_split_asm_thms,
+           distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms,
+           exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms,
+           split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms,
            case_eq_ifs = case_eq_if_thms}
           |> morph_ctr_sugar (substitute_noted_thm noted);
       in
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_tactics.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -22,13 +22,13 @@
   val mk_case_eq_if_tac: Proof.context -> int -> thm -> thm list -> thm list list ->
     thm list list -> tactic
   val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
-  val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
+  val mk_exhaust_disc_tac: int -> thm -> thm list -> tactic
+  val mk_exhaust_sel_tac: int -> thm -> thm list -> tactic
   val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
     thm list list list -> thm list list list -> tactic
-  val mk_half_disc_exclude_tac: Proof.context -> int -> thm -> thm -> tactic
+  val mk_half_distinct_disc_tac: Proof.context -> int -> thm -> thm -> tactic
   val mk_nchotomy_tac: int -> thm -> tactic
-  val mk_other_half_disc_exclude_tac: thm -> tactic
-  val mk_sel_exhaust_tac: int -> thm -> thm list -> tactic
+  val mk_other_half_distinct_disc_tac: thm -> tactic
   val mk_split_tac: Proof.context -> thm -> thm list -> thm list list -> thm list list ->
     thm list list list -> tactic
   val mk_split_asm_tac: Proof.context -> thm -> tactic
@@ -67,21 +67,21 @@
     (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac])
      |> k = 1 ? swap |> op @)));
 
-fun mk_half_disc_exclude_tac ctxt m discD disc' =
+fun mk_half_distinct_disc_tac ctxt m discD disc' =
   HEADGOAL (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
     rtac disc');
 
-fun mk_other_half_disc_exclude_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half);
+fun mk_other_half_distinct_disc_tac half = HEADGOAL (etac @{thm contrapos_pn} THEN' etac half);
 
-fun mk_disc_or_sel_exhaust_tac n exhaust destIs =
+fun mk_exhaust_disc_or_sel_tac n exhaust destIs =
   HEADGOAL (rtac exhaust THEN'
     EVERY' (map2 (fn k => fn destI => dtac destI THEN'
       select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) destIs));
 
-val mk_disc_exhaust_tac = mk_disc_or_sel_exhaust_tac;
+val mk_exhaust_disc_tac = mk_exhaust_disc_or_sel_tac;
 
-fun mk_sel_exhaust_tac n disc_exhaust collapses =
-  mk_disc_or_sel_exhaust_tac n disc_exhaust collapses ORELSE
+fun mk_exhaust_sel_tac n exhaust_disc collapses =
+  mk_exhaust_disc_or_sel_tac n exhaust_disc collapses ORELSE
   HEADGOAL (etac meta_mp THEN' resolve_tac collapses);
 
 fun mk_collapse_tac ctxt m discD sels =
@@ -92,17 +92,17 @@
        REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac ctxt THEN'
        SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl));
 
-fun mk_expand_tac ctxt n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss
-    disc_excludesss' =
+fun mk_expand_tac ctxt n ms uexhaust_disc vexhaust_disc uncollapses distinct_discsss
+    distinct_discsss' =
   if ms = [0] then
     HEADGOAL (rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses))) THEN'
-      TRY o EVERY' [rtac udisc_exhaust, atac, rtac vdisc_exhaust, atac])
+      TRY o EVERY' [rtac uexhaust_disc, atac, rtac vexhaust_disc, atac])
   else
     let val ks = 1 upto n in
-      HEADGOAL (rtac udisc_exhaust THEN'
-        EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' => fn uuncollapse =>
-          EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac, rtac sym, rtac vdisc_exhaust,
-            EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse =>
+      HEADGOAL (rtac uexhaust_disc THEN'
+        EVERY' (map5 (fn k => fn m => fn distinct_discss => fn distinct_discss' => fn uuncollapse =>
+          EVERY' [rtac (uuncollapse RS trans) THEN' TRY o atac, rtac sym, rtac vexhaust_disc,
+            EVERY' (map4 (fn k' => fn distinct_discs => fn distinct_discs' => fn vuncollapse =>
               EVERY'
                 (if k' = k then
                    [rtac (vuncollapse RS trans), TRY o atac] @
@@ -113,11 +113,11 @@
                        REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE,
                        asm_simp_tac (ss_only [] ctxt)])
                  else
-                   [dtac (the_single (if k = n then disc_excludes else disc_excludes')),
+                   [dtac (the_single (if k = n then distinct_discs else distinct_discs')),
                     etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
                     atac, atac]))
-              ks disc_excludess disc_excludess' uncollapses)])
-          ks ms disc_excludesss disc_excludesss' uncollapses))
+              ks distinct_discss distinct_discss' uncollapses)])
+          ks ms distinct_discsss distinct_discsss' uncollapses))
     end;
 
 fun mk_case_same_ctr_tac ctxt injects =
--- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_util.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -39,6 +39,8 @@
     (string * sort) list * Proof.context
   val variant_tfrees: string list -> Proof.context -> typ list * Proof.context
 
+  val name_of_const: string -> term -> string
+
   val typ_subst_nonatomic: (typ * typ) list -> typ -> typ
   val subst_nonatomic_types: (typ * typ) list -> term -> term
 
@@ -177,6 +179,12 @@
   apfst (map TFree) o
     variant_types (map (ensure_prefix "'") ss) (replicate (length ss) @{sort type});
 
+fun name_of_const what t =
+  (case head_of t of
+    Const (s, _) => s
+  | Free (s, _) => s
+  | _ => error ("Cannot extract name of " ^ what));
+
 (*Replace each Ti by Ui (starting from the leaves); inst = [(T1, U1), ..., (Tn, Un)].*)
 fun typ_subst_nonatomic [] = I
   | typ_subst_nonatomic inst =
--- a/src/HOL/Tools/Datatype/datatype_aux.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Datatype/datatype_aux.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -27,7 +27,7 @@
     case_name : string,
     case_rewrites : thm list,
     case_cong : thm,
-    weak_case_cong : thm,
+    case_cong_weak : thm,
     split : thm,
     split_asm: thm}
 end
@@ -195,7 +195,7 @@
    case_name : string,
    case_rewrites : thm list,
    case_cong : thm,
-   weak_case_cong : thm,
+   case_cong_weak : thm,
    split : thm,
    split_asm: thm};
 
@@ -311,10 +311,9 @@
               else is_nonempty_dt (i :: is) i
           | arg_nonempty _ = true;
       in exists (forall (arg_nonempty o strip_dtyp) o snd) constrs end
-  in
-    assert_all (fn (i, _) => is_nonempty_dt [i] i) (hd descr)
-      (fn (_, (s, _, _)) => raise Datatype_Empty s)
-  end;
+    val _ = hd descr |> forall (fn (i, (s, _, _)) =>
+      is_nonempty_dt [i] i orelse raise Datatype_Empty s)
+  in () end;
 
 (* unfold a list of mutually recursive datatype specifications *)
 (* all types of the form DtType (dt_name, [..., DtRec _, ...]) *)
--- a/src/HOL/Tools/Datatype/datatype_data.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Datatype/datatype_data.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -95,7 +95,7 @@
   head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of case_rewrite))));
 
 fun ctr_sugar_of_info ({exhaust, nchotomy, inject, distinct, case_rewrites, case_cong,
-    weak_case_cong, split, split_asm, ...} : Datatype_Aux.info) =
+    case_cong_weak, split, split_asm, ...} : Datatype_Aux.info) =
   {ctrs = ctrs_of_exhaust exhaust,
    casex = case_of_case_rewrite (hd case_rewrites),
    discs = [],
@@ -106,7 +106,7 @@
    distincts = distinct,
    case_thms = case_rewrites,
    case_cong = case_cong,
-   weak_case_cong = weak_case_cong,
+   case_cong_weak = case_cong_weak,
    split = split,
    split_asm = split_asm,
    disc_defs = [],
@@ -114,13 +114,13 @@
    discIs = [],
    sel_defs = [],
    sel_thmss = [],
-   disc_excludesss = [],
-   disc_exhausts = [],
-   sel_exhausts = [],
+   distinct_discsss = [],
+   exhaust_discs = [],
+   exhaust_sels = [],
    collapses = [],
    expands = [],
-   sel_splits = [],
-   sel_split_asms = [],
+   split_sels = [],
+   split_sel_asms = [],
    case_eq_ifs = []};
 
 fun register dt_infos =
--- a/src/HOL/Tools/Datatype/datatype_prop.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -18,7 +18,7 @@
   val make_cases : string list -> descr list -> theory -> term list list
   val make_splits : string list -> descr list -> theory -> (term * term) list
   val make_case_combs : string list -> descr list -> theory -> string -> term list
-  val make_weak_case_congs : string list -> descr list -> theory -> term list
+  val make_case_cong_weaks : string list -> descr list -> theory -> term list
   val make_case_congs : string list -> descr list -> theory -> term list
   val make_nchotomys : descr list -> term list
 end;
@@ -330,7 +330,7 @@
 
 (************************* additional rules for TFL ***************************)
 
-fun make_weak_case_congs case_names descr thy =
+fun make_case_cong_weaks case_names descr thy =
   let
     val case_combs = make_case_combs case_names descr thy "f";
 
--- a/src/HOL/Tools/Datatype/rep_datatype.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Datatype/rep_datatype.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -258,7 +258,8 @@
     thy2
     |> Sign.add_path (space_implode "_" new_type_names)
     |> Global_Theory.note_thmss ""
-      [((Binding.name "rec", [Nitpick_Simps.add]), [(rec_thms, [])])]
+      [((Binding.name "rec", [Named_Theorems.add @{named_theorems nitpick_simp}]),
+          [(rec_thms, [])])]
     ||> Sign.parent_path
     |-> (fn thms => pair (reccomb_names, maps #2 thms))
   end;
@@ -347,7 +348,8 @@
     val case_names = map case_name_of case_thms;
   in
     thy2
-    |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms)
+    |> Context.theory_map
+        ((fold o fold) (Named_Theorems.add_thm @{named_theorems nitpick_simp}) case_thms)
     |> Sign.parent_path
     |> Datatype_Aux.store_thmss "case" new_type_names case_thms
     |-> (fn thmss => pair (thmss, case_names))
@@ -393,16 +395,16 @@
     |-> (fn (thms1, thms2) => pair (thms1 ~~ thms2))
   end;
 
-fun prove_weak_case_congs new_type_names case_names descr thy =
+fun prove_case_cong_weaks new_type_names case_names descr thy =
   let
-    fun prove_weak_case_cong t =
+    fun prove_case_cong_weak t =
      Goal.prove_sorry_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
        (fn {prems, ...} => EVERY [rtac (hd prems RS arg_cong) 1]);
 
-    val weak_case_congs =
-      map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs case_names descr thy);
+    val case_cong_weaks =
+      map prove_case_cong_weak (Datatype_Prop.make_case_cong_weaks case_names descr thy);
 
-  in thy |> Datatype_Aux.store_thms "weak_case_cong" new_type_names weak_case_congs end;
+  in thy |> Datatype_Aux.store_thms "case_cong_weak" new_type_names case_cong_weaks end;
 
 
 (* additional theorems for TFL *)
@@ -468,7 +470,7 @@
 
 fun make_dt_info descr induct inducts rec_names rec_rewrites
     (index, (((((((((((_, (tname, _, _))), inject), distinct),
-      exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong),
+      exhaust), nchotomy), case_name), case_rewrites), case_cong), case_cong_weak),
         (split, split_asm))) =
   (tname,
    {index = index,
@@ -484,7 +486,7 @@
     case_name = case_name,
     case_rewrites = case_rewrites,
     case_cong = case_cong,
-    weak_case_cong = weak_case_cong,
+    case_cong_weak = case_cong_weak,
     split = split,
     split_asm = split_asm});
 
@@ -511,8 +513,8 @@
       |> prove_case_thms config new_type_names descr rec_names rec_rewrites;
     val (case_congs, thy7) = thy6
       |> prove_case_congs new_type_names case_names descr nchotomys case_rewrites;
-    val (weak_case_congs, thy8) = thy7
-      |> prove_weak_case_congs new_type_names case_names descr;
+    val (case_cong_weaks, thy8) = thy7
+      |> prove_case_cong_weaks new_type_names case_names descr;
     val (splits, thy9) = thy8
       |> prove_split_thms config new_type_names case_names descr
         inject distinct exhaust case_rewrites;
@@ -522,7 +524,7 @@
       map_index
         (make_dt_info flat_descr induct inducts rec_names rec_rewrites)
         (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~
-          case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits);
+          case_names ~~ case_rewrites ~~ case_congs ~~ case_cong_weaks ~~ splits);
     val dt_names = map fst dt_infos;
     val prfx = Binding.qualify true (space_implode "_" new_type_names);
     val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites;
@@ -551,7 +553,7 @@
         ((Binding.empty, [iff_add]), [(flat inject, [])]),
         ((Binding.empty, [Classical.safe_elim NONE]),
           [(map (fn th => th RS notE) (flat distinct), [])]),
-        ((Binding.empty, [Simplifier.cong_add]), [(weak_case_congs, [])]),
+        ((Binding.empty, [Simplifier.cong_add]), [(case_cong_weaks, [])]),
         ((Binding.empty, [Induct.induct_simp_add]), [(flat (distinct @ inject), [])])] @
           named_rules @ unnamed_rules)
     |> snd
--- a/src/HOL/Tools/Function/function.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/function.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -279,9 +279,7 @@
 
 (* setup *)
 
-val setup =
-  setup_case_cong
-  #> Function_Common.Termination_Simps.setup
+val setup = setup_case_cong
 
 val get_congs = Function_Ctx_Tree.get_function_congs
 
--- a/src/HOL/Tools/Function/function_common.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/function_common.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -12,7 +12,7 @@
   defname : string,
     (* contains no logical entities: invariant under morphisms: *)
   add_simps : (binding -> binding) -> string -> (binding -> binding) ->
-    Attrib.src list -> thm list -> local_theory -> thm list * local_theory,
+    Token.src list -> thm list -> local_theory -> thm list * local_theory,
   fnames : string list,
   case_names : string list,
   fs : term list,
@@ -37,7 +37,7 @@
   defname : string,
     (* contains no logical entities: invariant under morphisms: *)
   add_simps : (binding -> binding) -> string -> (binding -> binding) ->
-    Attrib.src list -> thm list -> local_theory -> thm list * local_theory,
+    Token.src list -> thm list -> local_theory -> thm list * local_theory,
   fnames : string list,
   case_names : string list,
   fs : term list,
@@ -81,7 +81,6 @@
   val import_function_data : term -> Proof.context -> info option
   val import_last_function : Proof.context -> info option
   val add_function_data : info -> Context.generic -> Context.generic
-  structure Termination_Simps: NAMED_THMS
   val set_termination_prover : (Proof.context -> tactic) -> Context.generic -> Context.generic
   val get_termination_prover : Proof.context -> tactic
   val store_termination_rule : thm -> Context.generic -> Context.generic
@@ -224,15 +223,6 @@
   #> store_termination_rule termination
 
 
-(* Simp rules for termination proofs *)
-
-structure Termination_Simps = Named_Thms
-(
-  val name = @{binding termination_simp}
-  val description = "simplification rules for termination proofs"
-)
-
-
 (* Default Termination Prover *)
 
 (* FIXME just one data slot (record) per program unit *)
--- a/src/HOL/Tools/Function/lexicographic_order.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/lexicographic_order.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -182,7 +182,7 @@
     val (domT, _) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of rel))
 
     val measure_funs = (* 1: generate measures *)
-      MeasureFunctions.get_measure_functions ctxt domT
+      Measure_Functions.get_measure_functions ctxt domT
 
     val table = (* 2: create table *)
       map (fn t => map (mk_cell thy solve_tac (dest_term t)) measure_funs) tl
@@ -212,7 +212,7 @@
 fun lexicographic_order_tac quiet ctxt =
   TRY (Function_Common.apply_termination_rule ctxt 1) THEN
   lex_order_tac quiet ctxt
-    (auto_tac (ctxt addsimps Function_Common.Termination_Simps.get ctxt))
+    (auto_tac (ctxt addsimps (Named_Theorems.get ctxt @{named_theorems termination_simp})))
 
 val setup =
   Context.theory_map (Function_Common.set_termination_prover (lexicographic_order_tac false))
--- a/src/HOL/Tools/Function/measure_functions.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/measure_functions.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -7,25 +7,18 @@
 signature MEASURE_FUNCTIONS =
 sig
   val get_measure_functions : Proof.context -> typ -> term list
-  val setup : theory -> theory
 end
 
-structure MeasureFunctions : MEASURE_FUNCTIONS =
+structure Measure_Functions : MEASURE_FUNCTIONS =
 struct
 
 (** User-declared size functions **)
-structure Measure_Heuristic_Rules = Named_Thms
-(
-  val name = @{binding measure_function}
-  val description =
-    "rules that guide the heuristic generation of measure functions"
-);
 
 fun mk_is_measure t =
   Const (@{const_name is_measure}, fastype_of t --> HOLogic.boolT) $ t
 
 fun find_measures ctxt T =
-  DEPTH_SOLVE (resolve_tac (Measure_Heuristic_Rules.get ctxt) 1)
+  DEPTH_SOLVE (resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems measure_function})) 1)
     (HOLogic.mk_Trueprop (mk_is_measure (Var (("f",0), T --> HOLogic.natT)))
      |> cterm_of (Proof_Context.theory_of ctxt) |> Goal.init)
   |> Seq.map (prop_of #> (fn _ $ (_ $ (_ $ f)) => f))
@@ -53,6 +46,4 @@
 
 val get_measure_functions = mk_all_measure_funs
 
-val setup = Measure_Heuristic_Rules.setup
-
 end
--- a/src/HOL/Tools/Function/partial_function.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/partial_function.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,7 +6,6 @@
 
 signature PARTIAL_FUNCTION =
 sig
-  val setup: theory -> theory
   val init: string -> term -> term -> thm -> thm -> thm option -> declaration
 
   val mono_tac: Proof.context -> int -> tactic
@@ -54,13 +53,6 @@
 val lookup_mode = Symtab.lookup o Modes.get o Context.Proof;
 
 
-structure Mono_Rules = Named_Thms
-(
-  val name = @{binding partial_function_mono};
-  val description = "monotonicity rules for partial function definitions";
-);
-
-
 (*** Automated monotonicity proofs ***)
 
 fun strip_cases ctac = ctac #> Seq.map snd;
@@ -109,7 +101,7 @@
 fun mono_tac ctxt =
   K (Local_Defs.unfold_tac ctxt [@{thm curry_def}])
   THEN' (TRY o REPEAT_ALL_NEW
-   (resolve_tac (Mono_Rules.get ctxt)
+   (resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems partial_function_mono}))
      ORELSE' split_cases_tac ctxt));
 
 
@@ -321,7 +313,4 @@
     ((mode -- (Parse.fixes -- (Parse.where_ |-- Parse_Spec.spec)))
       >> (fn (mode, (fixes, spec)) => add_partial_function_cmd mode fixes spec));
 
-
-val setup = Mono_Rules.setup;
-
 end
--- a/src/HOL/Tools/Function/scnp_reconstruct.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/scnp_reconstruct.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -331,7 +331,7 @@
 
 fun decomp_scnp_tac orders ctxt =
   let
-    val extra_simps = Function_Common.Termination_Simps.get ctxt
+    val extra_simps = Named_Theorems.get ctxt @{named_theorems termination_simp}
     val autom_tac = auto_tac (ctxt addsimps extra_simps)
   in
      gen_sizechange_tac orders autom_tac ctxt
--- a/src/HOL/Tools/Function/size.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/size.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -200,7 +200,7 @@
         val ([(_, size_thms)], thy'') = thy'
           |> Global_Theory.note_thmss ""
             [((Binding.name "size",
-                [Simplifier.simp_add, Nitpick_Simps.add,
+                [Simplifier.simp_add, Named_Theorems.add @{named_theorems nitpick_simp},
                  Thm.declaration_attribute (fn thm =>
                    Context.mapping (Code.add_default_eqn thm) I)]),
               [(size_eqns, [])])];
--- a/src/HOL/Tools/Function/termination.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Function/termination.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -199,7 +199,7 @@
     val thy = Proof_Context.theory_of ctxt
     val sk = mk_sum_skel rel
     val Ts = node_types sk T
-    val M = Inttab.make (map_index (apsnd (MeasureFunctions.get_measure_functions ctxt)) Ts)
+    val M = Inttab.make (map_index (apsnd (Measure_Functions.get_measure_functions ctxt)) Ts)
     val chain_cache = Cache.create Term2tab.empty Term2tab.lookup Term2tab.update
       (prove_chain thy chain_tac)
     val descent_cache = Cache.create Term3tab.empty Term3tab.lookup Term3tab.update
--- a/src/HOL/Tools/Lifting/lifting_bnf.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Lifting/lifting_bnf.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -64,8 +64,9 @@
     val T_rel = list_comb (mk_rel_of_bnf Ds As Bs bnf, nth argss' 3)
     val concl = mk_Quotient [R_rel, Abs_map, Rep_map, T_rel] |> HOLogic.mk_Trueprop
     val goal = Logic.list_implies (assms, concl)
-    val thm = Goal.prove ctxt [] [] goal 
+    val thm = Goal.prove_sorry ctxt [] [] goal 
       (fn {context = ctxt, prems = _} => Quotient_tac bnf ctxt 1)
+      |> Thm.close_derivation
   in
     Drule.zero_var_indexes (singleton (Variable.export ctxt old_ctxt) thm)
   end
--- a/src/HOL/Tools/Lifting/lifting_def.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Lifting/lifting_def.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -13,7 +13,8 @@
     (binding * mixfix) -> typ -> term -> thm -> thm list -> local_theory -> local_theory
 
   val lift_def_cmd:
-    (binding * string option * mixfix) * string * (Facts.ref * Args.src list) list -> local_theory -> Proof.state
+    (binding * string option * mixfix) * string * (Facts.ref * Token.src list) list ->
+    local_theory -> Proof.state
 
   val can_generate_code_cert: thm -> bool
 end
@@ -699,7 +700,7 @@
 val liftdef_parser =
   (((Parse.binding -- (@{keyword "::"} |-- (Parse.typ >> SOME) -- Parse.opt_mixfix')) >> Parse.triple2)
     --| @{keyword "is"} -- Parse.term -- 
-      Scan.optional (@{keyword "parametric"} |-- Parse.!!! Parse_Spec.xthms1) []) >> Parse.triple1
+      Scan.optional (@{keyword "parametric"} |-- Parse.!!! Parse.xthms1) []) >> Parse.triple1
 val _ =
   Outer_Syntax.local_theory_to_proof @{command_spec "lift_definition"}
     "definition for constants over the quotient type"
--- a/src/HOL/Tools/Lifting/lifting_info.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Lifting/lifting_info.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -277,13 +277,8 @@
 
 (* theorems that a relator of an eq_onp is an eq_onp of the corresponding predicate *)
 
-structure Relator_Eq_Onp = Named_Thms
-(
-  val name = @{binding relator_eq_onp}
-  val description = "theorems that a relator of an eq_onp is an eq_onp of the corresponding predicate"
-)
-
-fun get_relator_eq_onp_rules ctxt = map safe_mk_meta_eq (Relator_Eq_Onp.get ctxt)
+fun get_relator_eq_onp_rules ctxt =
+  map safe_mk_meta_eq (rev (Named_Theorems.get ctxt @{named_theorems relator_eq_onp}))
 
 (* info about reflexivity rules *)
 
@@ -525,12 +520,13 @@
 val setup =
   quot_map_attribute_setup
   #> quot_del_attribute_setup
-  #> Relator_Eq_Onp.setup
   #> relator_distr_attribute_setup
 
 (* setup fixed eq_onp rules *)
 
-val _ = Context.>> (fold (Relator_Eq_Onp.add_thm o Transfer.prep_transfer_domain_thm @{context}) 
+val _ = Context.>>
+  (fold (Named_Theorems.add_thm @{named_theorems relator_eq_onp} o
+    Transfer.prep_transfer_domain_thm @{context})
   @{thms composed_equiv_rel_eq_onp composed_equiv_rel_eq_eq_onp})
 
 (* setup fixed reflexivity rules *)
--- a/src/HOL/Tools/Lifting/lifting_setup.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Lifting/lifting_setup.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -221,9 +221,9 @@
 
     val thy = Proof_Context.theory_of lthy
     val dummy_thm = Thm.transfer thy Drule.dummy_thm
-    val pointer = Outer_Syntax.scan Position.none bundle_name
+    val pointer = Outer_Syntax.scan (Keyword.get_lexicons ()) Position.none bundle_name
     val restore_lifting_att = 
-      ([dummy_thm], [Args.src ("Lifting.lifting_restore_internal", Position.none) pointer])
+      ([dummy_thm], [Token.src ("Lifting.lifting_restore_internal", Position.none) pointer])
   in
     lthy 
       |> Local_Theory.declaration {syntax = false, pervasive = true}
@@ -796,8 +796,8 @@
 val _ = 
   Outer_Syntax.local_theory @{command_spec "setup_lifting"}
     "setup lifting infrastructure" 
-      (opt_gen_code -- Parse_Spec.xthm -- Scan.option Parse_Spec.xthm 
-      -- Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse_Spec.xthm) >> 
+      (opt_gen_code -- Parse.xthm -- Scan.option Parse.xthm 
+      -- Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse.xthm) >> 
         (fn (((gen_code, xthm), opt_reflp_xthm), opt_par_xthm) => 
           setup_lifting_cmd gen_code xthm opt_reflp_xthm opt_par_xthm))
 
@@ -969,7 +969,7 @@
     case bundle of
       [(_, [arg_src])] => 
         let
-          val (name, _) = Args.syntax (Scan.lift Args.name) arg_src ctxt
+          val (name, _) = Token.syntax (Scan.lift Args.name) arg_src ctxt
             handle ERROR _ => error "The provided bundle is not a lifting bundle."
         in name end
       | _ => error "The provided bundle is not a lifting bundle."
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -416,17 +416,6 @@
                    \in the problem.")
       else
         ()
-    val _ =
-      if mode = Normal andalso
-         exists (fn Type (@{type_name Datatype.node}, _) => true | _ => false)
-                all_Ts then
-        print_nt (K ("Warning: The problem involves directly or indirectly the \
-                     \internal type " ^ quote @{type_name Datatype.node} ^
-                     ". This type is very Nitpick-unfriendly, and its presence \
-                     \usually indicates either a failure of abstraction or a \
-                     \quirk in Nitpick."))
-      else
-        ()
     val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts
     val _ =
       if verbose andalso not unique_scope then
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -974,7 +974,7 @@
 fun zero_const T = Const (@{const_name zero_class.zero}, T)
 fun suc_const T = Const (@{const_name Suc}, T --> T)
 
-fun uncached_data_type_constrs ({thy, ctxt, ...} : hol_context) (T as Type (s, _)) =
+fun uncached_data_type_constrs ({ctxt, ...} : hol_context) (T as Type (s, _)) =
     if is_interpreted_type s then
       []
     else
@@ -1927,12 +1927,13 @@
     Const (s, _) => (s, t)
   | t' => raise TERM ("Nitpick_HOL.pair_for_prop", [t, t'])
 
-fun def_table_for get ctxt subst =
-  ctxt |> get |> map (pair_for_prop o subst_atomic subst)
+fun def_table_for ts subst =
+  ts |> map (pair_for_prop o subst_atomic subst)
        |> AList.group (op =) |> Symtab.make
 
 fun const_def_tables ctxt subst ts =
-  (def_table_for (map prop_of o Nitpick_Unfolds.get) ctxt subst,
+  (def_table_for
+    (map prop_of (rev (Named_Theorems.get ctxt @{named_theorems nitpick_unfold}))) subst,
    fold (fn (s, t) => Symtab.map_default (s, []) (cons t))
         (map pair_for_prop ts) Symtab.empty)
 
@@ -1943,14 +1944,15 @@
 
 fun const_simp_table ctxt =
   def_table_for (map_filter (equationalize_term ctxt "nitpick_simp" o prop_of)
-                 o Nitpick_Simps.get) ctxt
+    (rev (Named_Theorems.get ctxt @{named_theorems nitpick_simp})))
 
 fun const_psimp_table ctxt =
   def_table_for (map_filter (equationalize_term ctxt "nitpick_psimp" o prop_of)
-                 o Nitpick_Psimps.get) ctxt
+    (rev (Named_Theorems.get ctxt @{named_theorems nitpick_psimp})))
 
 fun const_choice_spec_table ctxt subst =
-  map (subst_atomic subst o prop_of) (Nitpick_Choice_Specs.get ctxt)
+  map (subst_atomic subst o prop_of)
+    (rev (Named_Theorems.get ctxt @{named_theorems nitpick_choice_spec}))
   |> const_nondef_table
 
 fun inductive_intro_table ctxt subst def_tables =
@@ -1958,9 +1960,8 @@
     def_table_for
         (maps (map (unfold_mutually_inductive_preds thy def_tables o prop_of)
                o snd o snd)
-         o filter (fn (cat, _) => cat = Spec_Rules.Inductive orelse
-                                  cat = Spec_Rules.Co_Inductive)
-         o Spec_Rules.get) ctxt subst
+         (filter (fn (cat, _) => cat = Spec_Rules.Inductive orelse
+                                 cat = Spec_Rules.Co_Inductive) (Spec_Rules.get ctxt))) subst
   end
 
 fun ground_theorem_table thy =
--- a/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Predicate_Compile/code_prolog.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -1047,7 +1047,8 @@
     val thy = Proof_Context.theory_of ctxt
     val ((((full_constname, constT), vs'), intro), thy1) =
       Predicate_Compile_Aux.define_quickcheck_predicate t' thy
-    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
+    val thy2 =
+      Context.theory_map (Named_Theorems.add_thm @{named_theorems code_pred_def} intro) thy1
     val thy3 = Predicate_Compile.preprocess preprocess_options (Const (full_constname, constT)) thy2
     val ctxt' = Proof_Context.init_global thy3
     val _ = tracing "Generating prolog program..."
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -1042,7 +1042,7 @@
   let
     val ctxt = Proof_Context.init_global thy  (* FIXME proper context!? *)
     val process =
-      rewrite_rule ctxt (Predicate_Compile_Simps.get ctxt)
+      rewrite_rule ctxt (Named_Theorems.get ctxt @{named_theorems code_pred_simp})
     fun process_False intro_t =
       if member (op =) (Logic.strip_imp_prems intro_t) @{prop "False"}
       then NONE else SOME intro_t
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -161,8 +161,8 @@
 fun inline_equations thy th =
   let
     val ctxt = Proof_Context.init_global thy
-    val inline_defs = Predicate_Compile_Inline_Defs.get ctxt
-    val th' = (Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps inline_defs)) th
+    val inline_defs = Named_Theorems.get ctxt @{named_theorems code_pred_inline}
+    val th' = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps inline_defs) th
     (*val _ = print_step options
       ("Inlining " ^ (Syntax.string_of_term_global thy (prop_of th))
        ^ "with " ^ (commas (map ((Syntax.string_of_term_global thy) o prop_of) inline_defs))
@@ -208,12 +208,12 @@
           NONE
     fun filter_defs ths = map_filter filtering (map (normalize thy o Thm.transfer thy) ths)
     val spec =
-      (case filter_defs (Predicate_Compile_Alternative_Defs.get ctxt) of
+      (case filter_defs (Named_Theorems.get ctxt @{named_theorems code_pred_def}) of
         [] =>
           (case Spec_Rules.retrieve ctxt t of
             [] => error ("No specification for " ^ Syntax.string_of_term_global thy t)
           | ((_, (_, ths)) :: _) => filter_defs ths)
-      | ths => rev ths)
+      | ths => ths)
     val _ =
       if show_intermediate_results options then
         tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -231,7 +231,8 @@
     val thy = Proof_Context.theory_of ctxt
     val ((((full_constname, constT), vs'), intro), thy1) =
       Predicate_Compile_Aux.define_quickcheck_predicate t' thy
-    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
+    val thy2 =
+      Context.theory_map (Named_Theorems.add_thm @{named_theorems code_pred_def} intro) thy1
     val (thy3, _) = cpu_time "predicate preprocessing"
         (fn () => Predicate_Compile.preprocess options (Const (full_constname, constT)) thy2)
     val (thy4, _) = cpu_time "random_dseq core compilation"
--- a/src/HOL/Tools/Qelim/cooper.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Qelim/cooper.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -862,9 +862,8 @@
   let
     val simpset_ctxt =
       put_simpset (fst (get ctxt)) ctxt delsimps del_ths addsimps add_ths
-    val aprems = Arith_Data.get_arith_facts ctxt
   in
-    Method.insert_tac aprems
+    Method.insert_tac (rev (Named_Theorems.get ctxt @{named_theorems arith}))
     THEN_ALL_NEW Object_Logic.full_atomize_tac ctxt
     THEN_ALL_NEW CONVERSION Thm.eta_long_conversion
     THEN_ALL_NEW simp_tac simpset_ctxt
--- a/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/exhaustive_generators.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -198,7 +198,7 @@
   
 fun contains_recursive_type_under_function_types xs =
   exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT =>
-    (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs
+    (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype_Aux.DtRec _) => true | _ => false)))) xs
     
 fun instantiate_datatype (name, constprfx, sort, mk_equations, mk_T, argnames)
     config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
@@ -254,7 +254,7 @@
   let
     val cnstrs = flat (maps
       (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
-      (Symtab.dest (Datatype.get_all (Proof_Context.theory_of ctxt))))
+      (Symtab.dest (Datatype_Data.get_all (Proof_Context.theory_of ctxt))))
     fun is_constrt (Const (s, T), ts) = (case (AList.lookup (op =) cnstrs s, body_type T) of
         (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname'
       | _ => false)
@@ -537,19 +537,19 @@
   Quickcheck_Common.datatype_interpretation (@{sort exhaustive}, instantiate_exhaustive_datatype)
 
 val setup_bounded_forall_datatype_interpretation =
-  Datatype.interpretation (Quickcheck_Common.ensure_sort
+  Datatype_Data.interpretation (Quickcheck_Common.ensure_sort
     (((@{sort type}, @{sort type}), @{sort bounded_forall}),
-    (Datatype.the_descr, instantiate_bounded_forall_datatype)))
+    (Datatype_Data.the_descr, instantiate_bounded_forall_datatype)))
 
 val active = Attrib.setup_config_bool @{binding quickcheck_exhaustive_active} (K true);
 
 val setup =
   Quickcheck_Common.datatype_interpretation (@{sort full_exhaustive}, instantiate_full_exhaustive_datatype)
-(* #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype
+(* #> Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
       (((@{sort typerep}, @{sort term_of}), @{sort exhaustive}), instantiate_exhaustive_datatype))
-  #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype
+  #> Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
       (((@{sort typerep}, @{sort term_of}), @{sort fast_exhaustive}), instantiate_fast_exhaustive_datatype))
-  #> Datatype.interpretation (Quickcheck_Common.ensure_sort_datatype
+  #> Datatype_Data.interpretation (Quickcheck_Common.ensure_sort_datatype
       (((@{sort type}, @{sort type}), @{sort bounded_forall}), instantiate_bounded_forall_datatype))*)
   #> Context.theory_map (Quickcheck.add_tester ("exhaustive", (active, test_goals)))
   #> Context.theory_map (Quickcheck.add_batch_generator ("exhaustive", compile_generator_exprs))
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -175,7 +175,7 @@
     
 fun contains_recursive_type_under_function_types xs =
   exists (fn (_, (_, _, cs)) => cs |> exists (snd #> exists (fn dT =>
-    (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype.DtRec _) => true | _ => false)))) xs
+    (case Datatype_Aux.strip_dtyp dT of (_ :: _, Datatype_Aux.DtRec _) => true | _ => false)))) xs
 
 fun instantiate_narrowing_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy =
   let
--- a/src/HOL/Tools/Quickcheck/quickcheck_common.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quickcheck/quickcheck_common.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -23,15 +23,15 @@
   val generator_test_goal_terms :
     generator -> Proof.context -> bool -> (string * typ) list
       -> (term * term list) list -> Quickcheck.result list
-  type instantiation = Datatype.config -> Datatype.descr -> (string * sort) list
+  type instantiation = Datatype_Aux.config -> Datatype_Aux.descr -> (string * sort) list
      -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
   val ensure_sort :
     (((sort * sort) * sort) *
       ((theory -> string list -> Datatype_Aux.descr * (string * sort) list * string list
         * string * (string list * string list) * (typ list * typ list)) * instantiation))
-    -> Datatype.config -> string list -> theory -> theory
+    -> Datatype_Aux.config -> string list -> theory -> theory
   val ensure_common_sort_datatype :
-    (sort * instantiation) -> Datatype.config -> string list -> theory -> theory
+    (sort * instantiation) -> Datatype_Aux.config -> string list -> theory -> theory
   val datatype_interpretation : (sort * instantiation) -> theory -> theory
   val gen_mk_parametric_generator_expr :
    (((Proof.context -> term * term list -> term) * term) * typ)
@@ -387,7 +387,7 @@
 
 (** ensuring sort constraints **)
 
-type instantiation = Datatype.config -> Datatype.descr -> (string * sort) list
+type instantiation = Datatype_Aux.config -> Datatype_Aux.descr -> (string * sort) list
   -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory
 
 fun perhaps_constrain thy insts raw_vs =
@@ -419,9 +419,9 @@
   end;
 
 fun ensure_common_sort_datatype (sort, instantiate) =
-  ensure_sort (((@{sort typerep}, @{sort term_of}), sort), (Datatype.the_descr, instantiate))
+  ensure_sort (((@{sort typerep}, @{sort term_of}), sort), (Datatype_Data.the_descr, instantiate))
 
-val datatype_interpretation = Datatype.interpretation o ensure_common_sort_datatype
+val datatype_interpretation = Datatype_Data.interpretation o ensure_common_sort_datatype
   
 (** generic parametric compilation **)
 
--- a/src/HOL/Tools/Quotient/quotient_def.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quotient/quotient_def.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -84,8 +84,7 @@
               Quotient_Info.update_quotconsts c qcinfo
           | _ => I))
       |> (snd oo Local_Theory.note) 
-        ((rsp_thm_name, [Attrib.internal (K Quotient_Info.rsp_rules_add)]),
-        [rsp_thm])
+        ((rsp_thm_name, @{attributes [quot_respect]}), [rsp_thm])
   in
     (qconst_data, lthy'')
   end
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -33,17 +33,6 @@
   val dest_quotconsts_global: theory -> quotconsts list
   val dest_quotconsts: Proof.context -> quotconsts list
   val print_quotconsts: Proof.context -> unit
-
-  val equiv_rules: Proof.context -> thm list
-  val equiv_rules_add: attribute
-  val rsp_rules: Proof.context -> thm list
-  val rsp_rules_add: attribute
-  val prs_rules: Proof.context -> thm list
-  val prs_rules_add: attribute
-  val id_simps: Proof.context -> thm list
-  val quotient_rules: Proof.context -> thm list
-  val quotient_rules_add: attribute
-  val setup: theory -> theory
 end;
 
 structure Quotient_Info: QUOTIENT_INFO =
@@ -69,16 +58,17 @@
 
 (* FIXME export proper internal update operation!? *)
 
-val quotmaps_attribute_setup =
-  Attrib.setup @{binding mapQ3}
-    ((Args.type_name {proper = true, strict = true} --| Scan.lift @{keyword "="}) --
-      (Scan.lift @{keyword "("} |--
-        Args.const {proper = true, strict = true} --| Scan.lift @{keyword ","} --
-        Attrib.thm --| Scan.lift @{keyword ")"}) >>
-      (fn (tyname, (relname, qthm)) =>
-        let val minfo = {relmap = relname, quot_thm = qthm}
-        in Thm.declaration_attribute (fn _ => Quotmaps.map (Symtab.update (tyname, minfo))) end))
-    "declaration of map information"
+val _ =
+  Theory.setup
+   (Attrib.setup @{binding mapQ3}
+      ((Args.type_name {proper = true, strict = true} --| Scan.lift @{keyword "="}) --
+        (Scan.lift @{keyword "("} |--
+          Args.const {proper = true, strict = true} --| Scan.lift @{keyword ","} --
+          Attrib.thm --| Scan.lift @{keyword ")"}) >>
+        (fn (tyname, (relname, qthm)) =>
+          let val minfo = {relmap = relname, quot_thm = qthm}
+          in Thm.declaration_attribute (fn _ => Quotmaps.map (Symtab.update (tyname, minfo))) end))
+      "declaration of map information")
 
 fun print_quotmaps ctxt =
   let
@@ -235,66 +225,6 @@
     |> Pretty.writeln
   end
 
-(* equivalence relation theorems *)
-structure Equiv_Rules = Named_Thms
-(
-  val name = @{binding quot_equiv}
-  val description = "equivalence relation theorems"
-)
-
-val equiv_rules = Equiv_Rules.get
-val equiv_rules_add = Equiv_Rules.add
-
-(* respectfulness theorems *)
-structure Rsp_Rules = Named_Thms
-(
-  val name = @{binding quot_respect}
-  val description = "respectfulness theorems"
-)
-
-val rsp_rules = Rsp_Rules.get
-val rsp_rules_add = Rsp_Rules.add
-
-(* preservation theorems *)
-structure Prs_Rules = Named_Thms
-(
-  val name = @{binding quot_preserve}
-  val description = "preservation theorems"
-)
-
-val prs_rules = Prs_Rules.get
-val prs_rules_add = Prs_Rules.add
-
-(* id simplification theorems *)
-structure Id_Simps = Named_Thms
-(
-  val name = @{binding id_simps}
-  val description = "identity simp rules for maps"
-)
-
-val id_simps = Id_Simps.get
-
-(* quotient theorems *)
-structure Quotient_Rules = Named_Thms
-(
-  val name = @{binding quot_thm}
-  val description = "quotient theorems"
-)
-
-val quotient_rules = Quotient_Rules.get
-val quotient_rules_add = Quotient_Rules.add
-
-
-(* theory setup *)
-
-val setup =
-  quotmaps_attribute_setup #>
-  Equiv_Rules.setup #>
-  Rsp_Rules.setup #>
-  Prs_Rules.setup #>
-  Id_Simps.setup #>
-  Quotient_Rules.setup
-
 
 (* outer syntax commands *)
 
--- a/src/HOL/Tools/Quotient/quotient_tacs.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quotient/quotient_tacs.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -55,14 +55,14 @@
 (** solvers for equivp and quotient assumptions **)
 
 fun equiv_tac ctxt =
-  REPEAT_ALL_NEW (resolve_tac (Quotient_Info.equiv_rules ctxt))
+  REPEAT_ALL_NEW (resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv})))
 
 val equiv_solver = mk_solver "Equivalence goal solver" equiv_tac
 
 fun quotient_tac ctxt =
   (REPEAT_ALL_NEW (FIRST'
     [rtac @{thm identity_quotient3},
-     resolve_tac (Quotient_Info.quotient_rules ctxt)]))
+     resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems quot_thm}))]))
 
 val quotient_solver = mk_solver "Quotient goal solver" quotient_tac
 
@@ -144,11 +144,12 @@
 
 fun reflp_get ctxt =
   map_filter (fn th => if prems_of th = [] then SOME (OF1 @{thm equivp_reflp} th) else NONE
-    handle THM _ => NONE) (Quotient_Info.equiv_rules ctxt)
+    handle THM _ => NONE) (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv}))
 
 val eq_imp_rel = @{lemma "equivp R ==> a = b --> R a b" by (simp add: equivp_reflp)}
 
-fun eq_imp_rel_get ctxt = map (OF1 eq_imp_rel) (Quotient_Info.equiv_rules ctxt)
+fun eq_imp_rel_get ctxt =
+  map (OF1 eq_imp_rel) (rev (Named_Theorems.get ctxt @{named_theorems quot_equiv}))
 
 fun regularize_tac ctxt =
   let
@@ -370,7 +371,8 @@
 
      (* respectfulness of constants; in particular of a simple relation *)
   | _ $ (Const _) $ (Const _)  (* rel_fun, list_rel, etc but not equality *)
-      => resolve_tac (Quotient_Info.rsp_rules ctxt) THEN_ALL_NEW quotient_tac ctxt
+      => resolve_tac (rev (Named_Theorems.get ctxt @{named_theorems quot_respect}))
+          THEN_ALL_NEW quotient_tac ctxt
 
       (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
       (* observe map_fun *)
@@ -516,20 +518,20 @@
 
   4. test for refl
 *)
-fun clean_tac lthy =
+fun clean_tac ctxt =
   let
-    val thy =  Proof_Context.theory_of lthy
+    val thy =  Proof_Context.theory_of ctxt
     val defs = map (Thm.symmetric o #def) (Quotient_Info.dest_quotconsts_global thy)
-    val prs = Quotient_Info.prs_rules lthy
-    val ids = Quotient_Info.id_simps lthy
+    val prs = rev (Named_Theorems.get ctxt @{named_theorems quot_preserve})
+    val ids = rev (Named_Theorems.get ctxt @{named_theorems id_simps})
     val thms =
       @{thms Quotient3_abs_rep Quotient3_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs
 
-    val simpset = (mk_minimal_simpset lthy) addsimps thms addSolver quotient_solver
+    val simpset = (mk_minimal_simpset ctxt) addsimps thms addSolver quotient_solver
   in
     EVERY' [
-      map_fun_tac lthy,
-      lambda_prs_tac lthy,
+      map_fun_tac ctxt,
+      lambda_prs_tac ctxt,
       simp_tac simpset,
       TRY o rtac refl]
   end
--- a/src/HOL/Tools/Quotient/quotient_type.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Quotient/quotient_type.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -13,7 +13,7 @@
     ((binding * binding) option * bool * thm option) -> Proof.context -> Proof.state
 
   val quotient_type_cmd: ((((((bool * string list) * binding) * mixfix) * string) * (bool * string)) *
-    (binding * binding) option) * (Facts.ref * Args.src list) option -> Proof.context -> Proof.state
+    (binding * binding) option) * (Facts.ref * Token.src list) option -> Proof.context -> Proof.state
 end;
 
 structure Quotient_Type: QUOTIENT_TYPE =
@@ -206,11 +206,10 @@
       |> init_quotient_infr gen_code quotient_thm equiv_thm opt_par_thm
       |> (snd oo Local_Theory.note)
         ((equiv_thm_name,
-          if partial then [] else [Attrib.internal (K Quotient_Info.equiv_rules_add)]),
+          if partial then [] else @{attributes [quot_equiv]}),
           [equiv_thm])
       |> (snd oo Local_Theory.note)
-        ((quotient_thm_name, [Attrib.internal (K Quotient_Info.quotient_rules_add)]),
-          [quotient_thm])
+        ((quotient_thm_name, @{attributes [quot_thm]}), [quotient_thm])
   in
     (quotients, lthy4)
   end
@@ -344,7 +343,7 @@
       Parse.opt_mixfix -- (@{keyword "="} |-- Parse.typ) --
         (@{keyword "/"} |-- (partial -- Parse.term))  --
         Scan.option (@{keyword "morphisms"} |-- Parse.!!! (Parse.binding -- Parse.binding))
-          -- Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse_Spec.xthm)
+          -- Scan.option (@{keyword "parametric"} |-- Parse.!!! Parse.xthm)
 
 val _ =
   Outer_Syntax.local_theory_to_proof @{command_spec "quotient_type"}
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -158,7 +158,7 @@
     | @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct)
     | _ => raise CTERM ("no equation", [ct]))
 
-  fun get_constrs thy (Type (n, _)) = these (Datatype.get_constrs thy n)
+  fun get_constrs thy (Type (n, _)) = these (Datatype_Data.get_constrs thy n)
     | get_constrs _ _ = []
 
   fun is_constr thy (n, T) =
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -7,9 +7,7 @@
 signature Z3_PROOF_RECONSTRUCTION =
 sig
   val add_z3_rule: thm -> Context.generic -> Context.generic
-  val reconstruct: Proof.context -> SMT_Translate.recon -> string list ->
-    int list * thm
-  val setup: theory -> theory
+  val reconstruct: Proof.context -> SMT_Translate.recon -> string list -> int list * thm
 end
 
 structure Z3_Proof_Reconstruction: Z3_PROOF_RECONSTRUCTION =
@@ -23,8 +21,6 @@
 
 (* net of schematic rules *)
 
-val z3_ruleN = "z3_rule"
-
 local
   val description = "declaration of Z3 proof rules"
 
@@ -55,9 +51,9 @@
 fun by_schematic_rule ctxt ct =
   the (Z3_Proof_Tools.net_instance (Z3_Rules.get (Context.Proof ctxt)) ct)
 
-val z3_rules_setup =
-  Attrib.setup (Binding.name z3_ruleN) (Attrib.add_del add del) description #>
-  Global_Theory.add_thms_dynamic (Binding.name z3_ruleN, Net.content o Z3_Rules.get)
+val _ = Theory.setup
+ (Attrib.setup @{binding z3_rule} (Attrib.add_del add del) description #>
+  Global_Theory.add_thms_dynamic (@{binding z3_rule}, Net.content o Z3_Rules.get))
 
 end
 
@@ -84,8 +80,7 @@
       Pretty.big_list ("Z3 found a proof," ^
         " but proof reconstruction failed at the following subgoal:")
         (pretty_goal ctxt thms (Thm.term_of ct)),
-      Pretty.str ("Adding a rule to the lemma group " ^ quote z3_ruleN ^
-        " might solve this problem.")])
+      Pretty.str ("Declaring a rule as [z3_rule] might solve this problem.")])
 
     fun apply [] ct = error (try_apply_err ct)
       | apply (prover :: provers) ct =
@@ -671,12 +666,6 @@
      * normal forms for polynoms (integer/real arithmetic)
      * quantifier elimination over linear arithmetic
      * ... ? **)
-structure Z3_Simps = Named_Thms
-(
-  val name = @{binding z3_simp}
-  val description = "simplification rules for Z3 proof reconstruction"
-)
-
 local
   fun spec_meta_eq_of thm =
     (case try (fn th => th RS @{thm spec}) thm of
@@ -880,7 +869,8 @@
     val (asserted, steps, vars, ctxt1) =
       Z3_Proof_Parser.parse ctxt typs terms output
 
-    val simpset = Z3_Proof_Tools.make_simpset ctxt1 (Z3_Simps.get ctxt1)
+    val simpset =
+      Z3_Proof_Tools.make_simpset ctxt1 (Named_Theorems.get ctxt1 @{named_theorems z3_simp})
 
     val ((is, rules), cxp as (ctxt2, _)) =
       add_asserted outer_ctxt rewrite_rules assms asserted ctxt1
@@ -895,6 +885,4 @@
 
 end
 
-val setup = z3_rules_setup #> Z3_Simps.setup
-
 end
--- a/src/HOL/Tools/SMT2/smt2_normalize.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/smt2_normalize.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -144,7 +144,7 @@
     | @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct)
     | _ => raise CTERM ("no equation", [ct]))
 
-  fun get_constrs thy (Type (n, _)) = these (Datatype.get_constrs thy n)
+  fun get_constrs thy (Type (n, _)) = these (Datatype_Data.get_constrs thy n)
     | get_constrs _ _ = []
 
   fun is_constr thy (n, T) =
--- a/src/HOL/Tools/SMT2/smt2_real.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/smt2_real.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -36,7 +36,7 @@
   fold (SMT2_Builtin.add_builtin_fun' SMTLIB2_Interface.smtlib2C) [
     (@{const less (real)}, "<"),
     (@{const less_eq (real)}, "<="),
-    (@{const uminus (real)}, "~"),
+    (@{const uminus (real)}, "-"),
     (@{const plus (real)}, "+"),
     (@{const minus (real)}, "-") ] #>
   SMT2_Builtin.add_builtin_fun SMTLIB2_Interface.smtlib2C
--- a/src/HOL/Tools/SMT2/smt2_systems.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/smt2_systems.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -37,6 +37,8 @@
     val (l, ls) = split_first (snd (take_prefix (curry (op =) "") lines))
   in (test_outcome solver_name l, ls) end
 
+fun on_first_non_unsupported_line test_outcome solver_name lines =
+  on_first_line test_outcome solver_name (filter (curry (op <>) "unsupported") lines)
 
 (* CVC3 *)
 
@@ -85,6 +87,27 @@
 
 end
 
+(* veriT *)
+
+val veriT: SMT2_Solver.solver_config = {
+  name = "veriT",
+  class = K SMTLIB2_Interface.smtlib2C,
+  avail = make_avail "VERIT",
+  command = make_command "VERIT",
+  options = (fn ctxt => [
+    "--proof-version=1",
+    "--proof=-",
+    "--proof-prune",
+    "--proof-merge",
+    "--disable-print-success",
+    "--disable-banner",
+    "--max-time=" ^ string_of_int (Real.ceil (Config.get ctxt SMT2_Config.timeout))]),
+  smt_options = [],
+  default_max_relevant = 120 (* FUDGE *),
+  outcome = on_first_non_unsupported_line (outcome_of "unsat" "sat"
+    "warning : proof_done: status is still open"),
+  parse_proof = SOME VeriT_Proof_Parse.parse_proof,
+  replay = NONE }
 
 (* Z3 *)
 
@@ -160,6 +183,7 @@
 val _ = Theory.setup (
   SMT2_Solver.add_solver cvc3 #>
   SMT2_Solver.add_solver cvc4 #>
+  SMT2_Solver.add_solver veriT #>
   SMT2_Solver.add_solver z3)
 
 end;
--- a/src/HOL/Tools/SMT2/smtlib2.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/smtlib2.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -118,7 +118,7 @@
       read l cs None ((S (rev ts1) :: ts2) :: tss)
   | read l ("#" :: "x" :: cs) None (ts :: tss) =
       token read_hex l cs ts tss
-  | read l ("#" :: cs) None (ts :: tss) =
+  | read l ("#" :: "b" :: cs) None (ts :: tss) =
       token read_bin l cs ts tss
   | read l (":" :: cs) None (ts :: tss) =
       token (read_sym Key) l cs ts tss
--- a/src/HOL/Tools/SMT2/smtlib2_interface.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/smtlib2_interface.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -11,6 +11,7 @@
   val add_logic: int * (term list -> string option) -> Context.generic -> Context.generic
   val translate_config: Proof.context -> SMT2_Translate.config
   val assert_index_of_name: string -> int
+  val assert_prefix : string
 end;
 
 structure SMTLIB2_Interface: SMTLIB2_INTERFACE =
@@ -48,7 +49,7 @@
     (@{const If ('a)}, "ite"),
     (@{const less (int)}, "<"),
     (@{const less_eq (int)}, "<="),
-    (@{const uminus (int)}, "~"),
+    (@{const uminus (int)}, "-"),
     (@{const plus (int)}, "+"),
     (@{const minus (int)}, "-")] #>
   SMT2_Builtin.add_builtin_fun smtlib2C
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smtlib2_isar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,71 @@
+(*  Title:      HOL/Tools/SMT2/smtlib2_isar.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+    Author:     Mathias Fleury, ENS Rennes
+
+General tools for Isar proof reconstruction.
+*)
+
+signature SMTLIB2_ISAR =
+sig
+  val unlift_term: term list -> term -> term
+  val postprocess_step_conclusion: theory -> thm list -> term list -> term -> term
+  val normalizing_prems : Proof.context -> term -> (string * string list) list
+  val distinguish_conjecture_and_hypothesis : ''a list -> ''b -> ''b -> ''b list ->
+    (''a * 'c) list -> 'c list -> 'c -> (ATP_Problem.atp_formula_role * 'c) option
+  val unskolemize_names: term -> term
+end;
+
+structure SMTLIB2_Isar: SMTLIB2_ISAR =
+struct
+
+open ATP_Util
+open ATP_Problem
+open ATP_Proof_Reconstruct
+
+fun unlift_term ll_defs =
+  let
+    val lifted = map (ATP_Util.extract_lambda_def dest_Free o ATP_Util.hol_open_form I) ll_defs
+
+    fun un_free (t as Free (s, _)) =
+       (case AList.lookup (op =) lifted s of
+         SOME t => un_term t
+       | NONE => t)
+     | un_free t = t
+    and un_term t = map_aterms un_free t
+  in un_term end
+
+(* It is not entirely clear why this should be necessary, especially for abstractions variables. *)
+val unskolemize_names =
+  Term.map_abs_vars (perhaps (try Name.dest_skolem))
+  #> Term.map_aterms (perhaps (try (fn Free (s, T) => Free (Name.dest_skolem s, T))))
+
+fun postprocess_step_conclusion thy rewrite_rules ll_defs =
+  Raw_Simplifier.rewrite_term thy rewrite_rules []
+  #> Object_Logic.atomize_term thy
+  #> not (null ll_defs) ? unlift_term ll_defs
+  #> simplify_bool
+  #> unskolemize_names
+  #> HOLogic.mk_Trueprop
+
+fun normalizing_prems ctxt concl0 =
+  SMT2_Normalize.case_bool_entry :: SMT2_Normalize.special_quant_table @
+  SMT2_Normalize.abs_min_max_table
+  |> map_filter (fn (c, th) =>
+    if exists_Const (curry (op =) c o fst) concl0 then
+      let val s = short_thm_name ctxt th in SOME (s, [s]) end
+    else
+      NONE)
+
+fun distinguish_conjecture_and_hypothesis ss id conjecture_id prem_ids fact_helper_ts hyp_ts
+    concl_t =
+  (case ss of
+    [s] => SOME (Axiom, the (AList.lookup (op =) fact_helper_ts s))
+  | _ =>
+    if id = conjecture_id then
+      SOME (Conjecture, concl_t)
+    else
+     (case find_index (curry (op =) id) prem_ids of
+       ~1 => NONE (* lambda-lifting definition *)
+     | i => SOME (Hypothesis, nth hyp_ts i)))
+
+end;
--- a/src/HOL/Tools/SMT2/smtlib2_proof.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/smtlib2_proof.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -21,8 +21,7 @@
     ('a, 'b) context -> 'c * ('d, 'b) context
   val next_id: ('a, 'b) context -> int * ('a, 'b) context
   val with_fresh_names: (('a list, 'b) context ->
-    term * ((string * (string * typ)) list, 'b) context) -> ('c, 'b) context ->
-    (term * string list) * ('c, 'b) context
+    term * ((string * (string * typ)) list, 'b) context) -> ('c, 'b) context -> (term * string list)
 
   (*type and term parsers*)
   type type_parser = SMTLIB2.tree * typ list -> typ option
@@ -56,7 +55,7 @@
   extra: 'a}
 
 fun mk_context ctxt id syms typs funs extra: ('a, 'b) context =
-  {ctxt=ctxt, id=id, syms=syms, typs=typs, funs=funs, extra=extra}
+  {ctxt = ctxt, id = id, syms = syms, typs = typs, funs = funs, extra = extra}
 
 fun empty_context ctxt typs funs = mk_context ctxt 1 Symtab.empty typs funs []
 
@@ -82,7 +81,7 @@
 fun next_id ({ctxt, id, syms, typs, funs, extra}: ('a, 'b) context) =
   (id, mk_context ctxt (id + 1) syms typs funs extra)
 
-fun with_fresh_names f ({ctxt, id, syms, typs, funs, extra}: ('a, 'b) context) =
+fun with_fresh_names f ({ctxt, id, syms, typs, funs, ...}: ('a, 'b) context) =
   let
     fun bind (_, v as (_, T)) t = Logic.all_const T $ Term.absfree v t
 
@@ -93,11 +92,11 @@
       singleton (Proof_Context.standard_term_check_finish ctxt)
     fun infer ctxt t = if needs_infer t then infer_types ctxt t else t
 
-    val (t, {ctxt=ctxt', extra=names, ...}: ((string * (string * typ)) list, 'b) context) =
+    val (t, {ctxt = ctxt', extra = names, ...}: ((string * (string * typ)) list, 'b) context) =
       f (mk_context ctxt id syms typs funs [])
     val t' = infer ctxt' (fold_rev bind names (HOLogic.mk_Trueprop t))
   in
-    ((t', map fst names), mk_context ctxt id syms typs funs extra)
+    (t', map fst names)
   end
 
 fun lookup_typ ({typs, ...}: ('a, 'b) context) = Symtab.lookup typs
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/verit_isar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,59 @@
+(*  Title:      HOL/Tools/SMT2/verit_isar.ML
+    Author:     Mathias Fleury, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+
+VeriT proofs as generic ATP proofs for Isar proof reconstruction.
+*)
+
+signature VERIT_ISAR =
+sig
+  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
+  val atp_proof_of_veriT_proof: Proof.context -> term list -> thm list -> term list -> term ->
+    (string * term) list -> int list -> int -> (int * string) list -> VeriT_Proof.veriT_step list ->
+    (term, string) ATP_Proof.atp_step list
+end;
+
+structure VeriT_Isar: VERIT_ISAR =
+struct
+
+open ATP_Util
+open ATP_Problem
+open ATP_Proof
+open ATP_Proof_Reconstruct
+open SMTLIB2_Isar
+open VeriT_Proof
+
+fun atp_proof_of_veriT_proof ctxt ll_defs rewrite_rules hyp_ts concl_t fact_helper_ts prem_ids
+    conjecture_id fact_helper_ids proof =
+  let
+    val thy = Proof_Context.theory_of ctxt
+    fun steps_of (VeriT_Proof.VeriT_Step {id, rule, prems, concl, ...}) =
+      let
+        val concl' = postprocess_step_conclusion thy rewrite_rules ll_defs concl
+        fun standard_step role = ((id, []), role, concl', rule, map (fn id => (id, [])) prems)
+      in
+        if rule = veriT_input_rule then
+          let val ss = the_list (AList.lookup (op =) fact_helper_ids (the (Int.fromString id))) in
+            (case distinguish_conjecture_and_hypothesis ss (the (Int.fromString id))
+                conjecture_id prem_ids fact_helper_ts hyp_ts concl_t of
+              NONE => []
+            | SOME (role0, concl00) =>
+              let
+                val name0 = (id ^ "a", ss)
+                val concl0 = unskolemize_names concl00
+              in
+                [(name0, role0, concl0, rule, []),
+                 ((id, []), Plain, concl', veriT_rewrite_rule,
+                  name0 :: normalizing_prems ctxt concl0)]
+              end)
+          end
+        else if rule = veriT_tmp_ite_elim_rule then
+          [standard_step Lemma]
+        else
+          [standard_step Plain]
+      end
+  in
+    maps steps_of proof
+  end
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/verit_proof.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,334 @@
+(*  Title:      HOL/Tools/SMT2/verit_proof.ML
+    Author:     Mathias Fleury, ENS Rennes
+    Author:     Sascha Boehme, TU Muenchen
+
+VeriT proofs: parsing and abstract syntax tree.
+*)
+
+signature VERIT_PROOF =
+sig
+  (*proofs*)
+  datatype veriT_step = VeriT_Step of {
+    id: string,
+    rule: string,
+    prems: string list,
+    concl: term,
+    fixes: string list}
+
+  (*proof parser*)
+  val parse: typ Symtab.table -> term Symtab.table -> string list ->
+    Proof.context -> veriT_step list * Proof.context
+
+  val veriT_step_prefix : string
+  val veriT_input_rule: string
+  val veriT_la_generic_rule : string
+  val veriT_rewrite_rule : string
+  val veriT_simp_arith_rule : string
+  val veriT_tmp_ite_elim_rule : string
+  val veriT_tmp_skolemize_rule : string
+end;
+
+structure VeriT_Proof: VERIT_PROOF =
+struct
+
+open SMTLIB2_Proof
+
+datatype veriT_node = VeriT_Node of {
+  id: string,
+  rule: string,
+  prems: string list,
+  concl: term,
+  bounds: string list}
+
+fun mk_node id rule prems concl bounds =
+  VeriT_Node {id = id, rule = rule, prems = prems, concl = concl, bounds = bounds}
+
+datatype veriT_step = VeriT_Step of {
+  id: string,
+  rule: string,
+  prems: string list,
+  concl: term,
+  fixes: string list}
+
+fun mk_step id rule prems concl fixes =
+  VeriT_Step {id = id, rule = rule, prems = prems, concl = concl, fixes = fixes}
+
+val veriT_step_prefix = ".c"
+val veriT_alpha_conv_rule = "tmp_alphaconv"
+val veriT_input_rule = "input"
+val veriT_la_generic_rule = "la_generic"
+val veriT_rewrite_rule = "__rewrite" (* arbitrary *)
+val veriT_simp_arith_rule = "simp_arith"
+val veriT_tmp_ite_elim_rule = "tmp_ite_elim"
+val veriT_tmp_skolemize_rule = "tmp_skolemize"
+
+(* proof parser *)
+
+fun node_of p cx =
+  ([], cx)
+  ||>> `(with_fresh_names (term_of p))
+  |>> snd
+
+(*in order to get Z3-style quantification*)
+fun repair_quantification (SMTLIB2.S (SMTLIB2.Sym "forall" :: l)) =
+    let val (quantified_vars, t) = split_last (map repair_quantification l)
+    in
+      SMTLIB2.S (SMTLIB2.Sym "forall" :: SMTLIB2.S quantified_vars :: t :: [])
+    end
+  | repair_quantification (SMTLIB2.S (SMTLIB2.Sym "exists" :: l)) =
+    let val (quantified_vars, t) = split_last (map repair_quantification l)
+    in
+      SMTLIB2.S (SMTLIB2.Sym "exists" :: SMTLIB2.S quantified_vars :: t :: [])
+    end
+  | repair_quantification (SMTLIB2.S l) = SMTLIB2.S (map repair_quantification l)
+  | repair_quantification x = x
+
+fun replace_bound_var_by_free_var (q $ Abs (var, ty, u)) free_var =
+    (case List.find (fn v => String.isPrefix v var) free_var of
+      NONE => q $ Abs (var, ty, replace_bound_var_by_free_var u free_var)
+    | SOME _ => replace_bound_var_by_free_var (Term.subst_bound (Free (var, ty), u)) free_var)
+  | replace_bound_var_by_free_var (u $ v) free_vars = replace_bound_var_by_free_var u free_vars $
+     replace_bound_var_by_free_var v free_vars
+  | replace_bound_var_by_free_var u _ = u
+
+fun find_type_in_formula (Abs(v, ty, u)) var_name =
+    if String.isPrefix var_name v then SOME ty else find_type_in_formula u var_name
+  | find_type_in_formula (u $ v) var_name =
+    (case find_type_in_formula u var_name of
+      NONE => find_type_in_formula v var_name
+    | a => a)
+  | find_type_in_formula _ _ = NONE
+
+fun add_bound_variables_to_ctxt cx bounds concl =
+    fold (fn a => fn b => update_binding a b)
+      (map (fn s => ((s, Term (Free (s, the_default dummyT (find_type_in_formula concl s))))))
+       bounds) cx
+
+fun update_step_and_cx (st as VeriT_Node {id, rule, prems, concl, bounds}) cx =
+  if rule = veriT_tmp_ite_elim_rule then
+    (mk_node id rule prems concl bounds, add_bound_variables_to_ctxt cx bounds concl)
+  else if rule = veriT_tmp_skolemize_rule then
+    let
+      val concl' = replace_bound_var_by_free_var concl bounds
+    in
+      (mk_node id rule prems concl' [], add_bound_variables_to_ctxt cx bounds concl)
+    end
+  else
+    (st, cx)
+
+(*FIXME: using a reference would be better to know th numbers of the steps to add*)
+fun fix_subproof_steps ((((id_of_father_step, rule), prems), subproof), ((step_concl, bounds),
+    cx)) =
+  let
+    fun mk_prop_of_term concl = (fastype_of concl = @{typ "bool"} ?
+      curry (op $) @{term "Trueprop"}) concl
+    fun inline_assumption assumption assumption_id (st as VeriT_Node {id, rule, prems, concl,
+        bounds}) =
+      if List.find (curry (op =) assumption_id) prems <> NONE then
+        let
+          val prems' = filter_out (curry (op =) assumption_id) prems
+        in
+          mk_node id rule (filter_out (curry (op =) assumption_id) prems')
+            (Const (@{const_name "Pure.imp"}, @{typ "prop"} --> @{typ "prop"} --> @{typ "prop"})
+            $ mk_prop_of_term assumption $ mk_prop_of_term concl) bounds
+        end
+      else
+        st
+    fun find_input_steps_and_inline [] last_step = ([], last_step)
+      | find_input_steps_and_inline (VeriT_Node {id = id', rule, prems, concl, bounds} :: steps)
+          last_step =
+        if rule = veriT_input_rule then
+          find_input_steps_and_inline (map (inline_assumption concl id') steps) last_step
+        else
+          apfst (cons (mk_node (id_of_father_step ^ id') rule prems concl bounds))
+            (find_input_steps_and_inline steps (id_of_father_step ^ id'))
+    val (subproof', last_step_id) = find_input_steps_and_inline subproof ""
+    val prems' =
+      if last_step_id = "" then prems
+      else
+        (case prems of
+          NONE => SOME [last_step_id]
+        | SOME l => SOME (last_step_id :: l))
+  in
+    (subproof', (((((id_of_father_step, rule), prems'), step_concl), bounds), cx))
+  end
+
+(*
+(set id rule :clauses(...) :args(..) :conclusion (...)).
+or
+(set id subproof (set ...) :conclusion (...)).
+*)
+
+fun parse_proof_step cx =
+  let
+    fun rotate_pair (a, (b, c)) = ((a, b), c)
+    fun get_id (SMTLIB2.S [SMTLIB2.Sym "set", SMTLIB2.Sym id, SMTLIB2.S l]) = (id, l)
+      | get_id t = raise Fail ("unrecognized VeriT Proof" ^ PolyML.makestring t)
+    fun parse_rule (SMTLIB2.Sym rule :: l) = (rule, l)
+    fun parse_source (SMTLIB2.Key "clauses" :: SMTLIB2.S source ::l) =
+        (SOME (map (fn (SMTLIB2.Sym id) => id) source), l)
+      | parse_source l = (NONE, l)
+    fun parse_subproof cx id_of_father_step ((subproof_step as SMTLIB2.S (SMTLIB2.Sym "set" :: _)) :: l) =
+        let val (subproof_steps, cx') = parse_proof_step cx subproof_step in
+          apfst (apfst (curry (op @) subproof_steps)) (parse_subproof cx' id_of_father_step l)
+        end
+      | parse_subproof cx _ l = (([], cx), l)
+    fun skip_args (SMTLIB2.Key "args" :: SMTLIB2.S _ :: l) = l
+      | skip_args l = l
+    fun parse_conclusion (SMTLIB2.Key "conclusion" :: SMTLIB2.S concl :: []) = concl
+    fun make_or_from_clausification l =
+      foldl1 (fn ((concl1, bounds1), (concl2, bounds2)) =>
+        (HOLogic.mk_disj (perhaps (try HOLogic.dest_Trueprop) concl1,
+        perhaps (try HOLogic.dest_Trueprop) concl2), bounds1 @ bounds2)) l
+    fun to_node (((((id, rule), prems), concl), bounds), cx) = (mk_node id rule
+      (the_default [] prems) concl bounds, cx)
+  in
+    get_id
+    ##> parse_rule
+    #> rotate_pair
+    ##> parse_source
+    #> rotate_pair
+    ##> skip_args
+    #> (fn (((id, rule), prems), sub) => (((id, rule), prems), parse_subproof cx id sub))
+    #> rotate_pair
+    ##> parse_conclusion
+    ##> map repair_quantification
+    #> (fn ((((id, rule), prems), (subproof, cx)), terms) =>
+         (((((id, rule), prems), subproof), fold_map (fn t => fn cx => node_of t cx) terms cx)))
+    ##> apfst (fn [] => (@{const False}, []) | concls => make_or_from_clausification concls)
+    #> fix_subproof_steps
+    ##> to_node
+    #> (fn (subproof, (step, cx)) => (subproof @ [step], cx))
+    #-> fold_map update_step_and_cx
+  end
+
+(*subproofs are written on multiple lines: SMTLIB can not parse then, because parentheses are
+unbalanced on each line*)
+fun seperate_into_steps lines =
+  let
+    fun count ("(" :: l) n = count l (n+1)
+      | count (")" :: l) n = count l (n-1)
+      | count (_ :: l) n = count l n
+      | count [] n = n
+    fun seperate (line :: l) actual_lines m =
+        let val n = count (raw_explode line) 0 in
+          if m + n = 0 then
+            [actual_lines ^ line] :: seperate l "" 0
+          else seperate l (actual_lines ^ line) (m + n)
+        end
+      | seperate [] _ 0 = []
+  in
+    seperate lines "" 0
+  end
+
+ (* VeriT adds @ before every variable. *)
+fun remove_all_at (SMTLIB2.Sym v :: l) = SMTLIB2.Sym (perhaps (try (unprefix "@")) v) :: remove_all_at l
+  | remove_all_at (SMTLIB2.S l :: l') = SMTLIB2.S (remove_all_at l) :: remove_all_at l'
+  | remove_all_at (SMTLIB2.Key v :: l) = SMTLIB2.Key v :: remove_all_at l
+  | remove_all_at (v :: l) = v :: remove_all_at l
+  | remove_all_at [] = []
+
+fun find_in_which_step_defined var (VeriT_Node {id, bounds, ...} :: l) =
+    (case List.find (fn v => String.isPrefix v var) bounds of
+      NONE => find_in_which_step_defined var l
+    | SOME _ => id)
+  | find_in_which_step_defined var _ = raise Fail ("undefined " ^ var)
+
+(*Yes every case is possible: the introduced var is not on a special size of the equality sign.*)
+fun find_ite_var_in_term (Const ("HOL.If", _) $ _ $
+      (Const (@{const_name "HOL.eq"}, _) $ Free (var1, _) $ Free (var2, _) ) $
+      (Const (@{const_name "HOL.eq"}, _) $ Free (var3, _) $ Free (var4, _) )) =
+    let
+      fun get_number_of_ite_transformed_var var =
+        perhaps (try (unprefix "ite")) var
+        |> Int.fromString
+      fun is_equal_and_has_correct_substring var var' var'' =
+        if var = var' andalso String.isPrefix "ite" var then SOME var'
+        else if var = var'' andalso String.isPrefix "ite" var then SOME var'' else NONE
+      val var1_introduced_var = is_equal_and_has_correct_substring var1 var3 var4
+      val var2_introduced_var = is_equal_and_has_correct_substring var3 var1 var2
+    in
+      (case (var1_introduced_var, var2_introduced_var) of
+        (SOME a, SOME b) =>
+          (*ill-generated case, might be possible when applying the rule to max a a. Only if the
+          variable have been introduced before. Probably an impossible edge case*)
+          (case (get_number_of_ite_transformed_var a, get_number_of_ite_transformed_var b) of
+            (SOME a, SOME b) => if a < b then var2_introduced_var else var1_introduced_var
+            (*Otherwise, it is a name clase between a parameter name and the introduced variable.
+             Or the name convention has been changed.*)
+          | (NONE, SOME _) => var2_introduced_var
+          | (SOME _, NONE) => var2_introduced_var)
+      | (_, SOME _) => var2_introduced_var
+      | (SOME _, _) => var1_introduced_var)
+    end
+  | find_ite_var_in_term (Const (@{const_name "If"}, _) $ _ $
+      (Const (@{const_name "HOL.eq"}, _) $ Free (var, _) $ _ ) $
+      (Const (@{const_name "HOL.eq"}, _) $ Free (var', _) $ _ )) =
+    if var = var' then SOME var else NONE
+  | find_ite_var_in_term (Const (@{const_name "If"}, _) $ _ $
+      (Const (@{const_name "HOL.eq"}, _) $ _ $ Free (var, _)) $
+      (Const (@{const_name "HOL.eq"}, _) $ _ $ Free (var', _))) =
+    if var = var' then SOME var else NONE
+  | find_ite_var_in_term (p $ q) =
+    (case find_ite_var_in_term p of
+      NONE => find_ite_var_in_term q
+    | x => x)
+  | find_ite_var_in_term (Abs (_, _, body)) = find_ite_var_in_term body
+  | find_ite_var_in_term _ = NONE
+
+fun correct_veriT_step steps (st as VeriT_Node {id, rule, prems, concl, bounds}) =
+  if rule = veriT_tmp_ite_elim_rule then
+    if bounds = [] then
+      (*if the introduced var has already been defined, adding the definition as a dependency*)
+      let
+        val new_prems =
+          (case find_ite_var_in_term concl of
+            NONE => prems
+          | SOME var => find_in_which_step_defined var steps :: prems)
+      in
+        VeriT_Node {id = id, rule = rule, prems = new_prems, concl = concl, bounds = bounds}
+      end
+    else
+      (*some new variables are created*)
+      let
+        val concl' = replace_bound_var_by_free_var concl bounds
+      in
+        mk_node id rule prems concl' []
+      end
+  else
+    st
+
+fun remove_alpha_conversion _ [] = []
+  | remove_alpha_conversion replace_table (VeriT_Node {id, rule, prems, concl, bounds} :: steps) =
+    let
+      fun correct_dependency prems =
+        map (fn x => perhaps (Symtab.lookup replace_table) x) prems
+      fun find_predecessor prem = perhaps (Symtab.lookup replace_table) prem
+    in
+      if rule = veriT_alpha_conv_rule then
+        remove_alpha_conversion (Symtab.update (id, find_predecessor (hd prems))
+          replace_table) steps
+      else
+        VeriT_Node {id = id, rule = rule, prems = correct_dependency prems,
+          concl = concl, bounds = bounds} :: remove_alpha_conversion replace_table steps
+    end
+
+fun correct_veriT_steps steps =
+  steps
+  |> map (correct_veriT_step steps)
+  |> remove_alpha_conversion Symtab.empty
+
+fun parse typs funs lines ctxt =
+  let
+    val smtlib2_lines_without_at =
+      remove_all_at (map SMTLIB2.parse (seperate_into_steps lines))
+    val (u, env) = apfst flat (fold_map (fn l => fn cx => parse_proof_step cx l)
+      smtlib2_lines_without_at (empty_context ctxt typs funs))
+    val t = correct_veriT_steps u
+    fun node_to_step (VeriT_Node {id, rule, prems, concl, bounds, ...}) =
+      mk_step id rule prems concl bounds
+   in
+    (map node_to_step t, ctxt_of env)
+  end
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/verit_proof_parse.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,104 @@
+(*  Title:      HOL/Tools/SMT2/verit_proof_parse.ML
+    Author:     Mathias Fleury, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+
+VeriT proof parsing.
+*)
+
+signature VERIT_PROOF_PARSE =
+sig
+  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
+  val parse_proof: Proof.context -> SMT2_Translate.replay_data ->
+    ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
+    SMT2_Solver.parsed_proof
+end;
+
+structure VeriT_Proof_Parse: VERIT_PROOF_PARSE =
+struct
+
+open ATP_Util
+open ATP_Problem
+open ATP_Proof
+open ATP_Proof_Reconstruct
+open VeriT_Isar
+open VeriT_Proof
+
+fun find_and_add_missing_dependances steps assms ll_offset =
+  let
+    fun prems_to_theorem_number [] id repl = (([], []), (id, repl))
+      | prems_to_theorem_number (x :: ths) id replaced =
+        (case Int.fromString (perhaps (try (unprefix SMTLIB2_Interface.assert_prefix)) x) of
+          NONE =>
+          let
+            val ((prems, iidths), (id', replaced')) = prems_to_theorem_number ths id replaced
+          in
+            ((x :: prems, iidths), (id', replaced'))
+          end
+        | SOME th =>
+          (case Option.map snd (List.find (fst #> curry (op =) x) replaced) of
+            NONE =>
+            let
+              val id' = if th = ll_offset then 0 else id - ll_offset (* 0: for the conjecture*)
+              val ((prems, iidths), (id'', replaced')) =
+                prems_to_theorem_number ths (if th <> ll_offset then id + 1 else id)
+                  ((x, string_of_int id') :: replaced)
+            in
+              ((string_of_int id' :: prems, (th, (id', th - ll_offset)) :: iidths),
+               (id'', replaced'))
+            end
+          | SOME x =>
+            let
+              val ((prems, iidths), (id', replaced')) = prems_to_theorem_number ths id replaced
+            in ((x :: prems, iidths), (id', replaced')) end))
+    fun update_step (VeriT_Proof.VeriT_Step {prems, id = id0, rule = rule0,
+        concl = concl0, fixes = fixes0}) (id, replaced) =
+      let val ((prems', iidths), (id', replaced)) = prems_to_theorem_number prems id replaced
+      in
+        ((VeriT_Proof.VeriT_Step {id = id0, rule = rule0, prems = prems', concl = concl0,
+           fixes = fixes0}, iidths), (id', replaced))
+      end
+  in
+    fold_map update_step steps (1, [])
+    |> fst
+    |> `(map snd)
+    ||> (map fst)
+    |>> flat
+    |>> map (fn (_, (id, tm_id)) => let val (i, tm) = nth assms tm_id in (i, (id, tm)) end)
+  end
+
+fun add_missing_steps iidths =
+  let
+    fun add_single_step (_, (id, th)) = VeriT_Proof.VeriT_Step {id = string_of_int id,
+      rule = veriT_input_rule, prems = [], concl = prop_of th, fixes = []}
+  in map add_single_step iidths end
+
+fun parse_proof _
+    ({context = ctxt, typs, terms, ll_defs, rewrite_rules, assms} : SMT2_Translate.replay_data)
+    xfacts prems concl output =
+  let
+    val (steps, _) = VeriT_Proof.parse typs terms output ctxt
+    val (iidths, steps'') = find_and_add_missing_dependances steps assms (length ll_defs)
+    val steps' = add_missing_steps iidths @ steps''
+    fun id_of_index i = the_default ~1 (Option.map fst (AList.lookup (op =) iidths i))
+
+    val prems_i = 1
+    val facts_i = prems_i + length prems
+    val conjecture_i = 0
+    val ll_offset = id_of_index conjecture_i
+    val prem_ids = map id_of_index (prems_i upto facts_i - 1)
+    val helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths
+
+    val fact_ids = map_filter (fn (i, (id, _)) =>
+      (try (apsnd (nth xfacts)) (id, i - facts_i))) iidths
+    val fact_helper_ts =
+      map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, prop_of th)) helper_ids @
+      map (fn (_, ((s, _), th)) => (s, prop_of th)) fact_ids
+    val fact_helper_ids =
+      map (apsnd (ATP_Util.short_thm_name ctxt)) helper_ids @ map (apsnd (fst o fst)) fact_ids
+  in
+    {outcome = NONE, fact_ids = fact_ids,
+     atp_proof = fn () => atp_proof_of_veriT_proof ctxt ll_defs rewrite_rules prems concl
+       fact_helper_ts prem_ids ll_offset fact_helper_ids steps'}
+  end
+
+end;
--- a/src/HOL/Tools/SMT2/z3_new_isar.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/z3_new_isar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -18,6 +18,7 @@
 open ATP_Problem
 open ATP_Proof
 open ATP_Proof_Reconstruct
+open SMTLIB2_Isar
 
 val z3_apply_def_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Apply_Def
 val z3_hypothesis_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Hypothesis
@@ -62,104 +63,49 @@
           end
       end
 
-fun simplify_bool ((all as Const (@{const_name All}, _)) $ Abs (s, T, t)) =
-    let val t' = simplify_bool t in
-      if loose_bvar1 (t', 0) then all $ Abs (s, T, t') else t'
+fun dest_alls (Const (@{const_name Pure.all}, _) $ Abs (abs as (_, T, _))) =
+    let val (s', t') = Term.dest_abs abs in
+      dest_alls t' |>> cons (s', T)
     end
-  | simplify_bool (@{const Not} $ t) = s_not (simplify_bool t)
-  | simplify_bool (@{const conj} $ t $ u) = s_conj (simplify_bool t, simplify_bool u)
-  | simplify_bool (@{const disj} $ t $ u) = s_disj (simplify_bool t, simplify_bool u)
-  | simplify_bool (@{const implies} $ t $ u) = s_imp (simplify_bool t, simplify_bool u)
-  | simplify_bool (@{const HOL.eq (bool)} $ t $ u) = s_iff (simplify_bool t, simplify_bool u)
-  | simplify_bool (t as Const (@{const_name HOL.eq}, _) $ u $ v) =
-    if u aconv v then @{const True} else t
-  | simplify_bool (t $ u) = simplify_bool t $ simplify_bool u
-  | simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t)
-  | simplify_bool t = t
-
-(* It is not entirely clear why this should be necessary, especially for abstractions variables. *)
-val unskolemize_names =
-  Term.map_abs_vars (perhaps (try Name.dest_skolem))
-  #> Term.map_aterms (perhaps (try (fn Free (s, T) => Free (Name.dest_skolem s, T))))
+  | dest_alls t = ([], t)
 
-fun strip_alls (Const (@{const_name All}, _) $ Abs (s, T, body)) = strip_alls body |>> cons (s, T)
-  | strip_alls t = ([], t)
-
-fun push_skolem_all_under_iff t =
-  (case strip_alls t of
-    (qs as _ :: _,
-     (t0 as Const (@{const_name HOL.eq}, _)) $ (t1 as Const (@{const_name Ex}, _) $ _) $ t2) =>
-    t0 $ HOLogic.list_all (qs, t1) $ HOLogic.list_all (qs, t2)
-  | _ => t)
-
-fun unlift_term ll_defs =
-  let
-    val lifted = map (ATP_Util.extract_lambda_def dest_Free o ATP_Util.hol_open_form I) ll_defs
-
-    fun un_free (t as Free (s, _)) =
-       (case AList.lookup (op =) lifted s of
-         SOME t => un_term t
-       | NONE => t)
-     | un_free t = t
-    and un_term t = map_aterms un_free t
-  in un_term end
+val reorder_foralls =
+  dest_alls
+  #>> sort_wrt fst
+  #-> fold_rev (Logic.all o Free);
 
 fun atp_proof_of_z3_proof ctxt ll_defs rewrite_rules hyp_ts concl_t fact_helper_ts prem_ids
     conjecture_id fact_helper_ids proof =
   let
     val thy = Proof_Context.theory_of ctxt
 
-    val unlift_term = unlift_term ll_defs
-
     fun steps_of (Z3_New_Proof.Z3_Step {id, rule, prems, concl, ...}) =
       let
         val sid = string_of_int id
 
-        val concl' =
-          concl
-          |> Raw_Simplifier.rewrite_term thy rewrite_rules []
-          |> Object_Logic.atomize_term thy
-          |> simplify_bool
-          |> not (null ll_defs) ? unlift_term
-          |> unskolemize_names
-          |> push_skolem_all_under_iff
-          |> HOLogic.mk_Trueprop
-
+        val concl' = concl
+          |> reorder_foralls (* crucial for skolemization steps *)
+          |> postprocess_step_conclusion thy rewrite_rules ll_defs
         fun standard_step role =
           ((sid, []), role, concl', Z3_New_Proof.string_of_rule rule,
            map (fn id => (string_of_int id, [])) prems)
       in
         (case rule of
           Z3_New_Proof.Asserted =>
-          let
-            val ss = the_list (AList.lookup (op =) fact_helper_ids id)
-            val name0 = (sid ^ "a", ss)
-
-            val (role0, concl0) =
-              (case ss of
-                [s] => (Axiom, the (AList.lookup (op =) fact_helper_ts s))
-              | _ =>
-                if id = conjecture_id then
-                  (Conjecture, concl_t)
-                else
-                  (Hypothesis,
-                   (case find_index (curry (op =) id) prem_ids of
-                     ~1 => concl
-                   | i => nth hyp_ts i)))
-
-            val normalize_prems =
-              SMT2_Normalize.case_bool_entry :: SMT2_Normalize.special_quant_table @
-              SMT2_Normalize.abs_min_max_table
-              |> map_filter (fn (c, th) =>
-                if exists_Const (curry (op =) c o fst) concl0 then
-                  let val s = short_thm_name ctxt th in SOME (s, [s]) end
-                else
-                  NONE)
-          in
-            (if role0 = Axiom then []
-             else [(name0, role0, concl0, Z3_New_Proof.string_of_rule rule, [])]) @
-            [((sid, []), Plain, concl', Z3_New_Proof.string_of_rule Z3_New_Proof.Rewrite,
-              name0 :: normalize_prems)]
+          let val ss = the_list (AList.lookup (op =) fact_helper_ids id) in
+            (case distinguish_conjecture_and_hypothesis ss id conjecture_id prem_ids fact_helper_ts
+                hyp_ts concl_t of
+              NONE => []
+            | SOME (role0, concl00) =>
+              let
+                val name0 = (sid ^ "a", ss)
+                val concl0 = unskolemize_names concl00
+              in
+                (if role0 = Axiom then []
+                 else [(name0, role0, concl0, Z3_New_Proof.string_of_rule rule, [])]) @
+                [((sid, []), Plain, concl', Z3_New_Proof.string_of_rule Z3_New_Proof.Rewrite,
+                  name0 :: normalizing_prems ctxt concl0)]
+              end)
           end
         | Z3_New_Proof.Rewrite => [standard_step Lemma]
         | Z3_New_Proof.Rewrite_Star => [standard_step Lemma]
--- a/src/HOL/Tools/SMT2/z3_new_proof.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/SMT2/z3_new_proof.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -7,14 +7,14 @@
 signature Z3_NEW_PROOF =
 sig
   (*proof rules*)
-  datatype z3_rule = True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity |
-    Symmetry | Transitivity | Transitivity_Star | Monotonicity | Quant_Intro |
-    Distributivity | And_Elim | Not_Or_Elim | Rewrite | Rewrite_Star |
-    Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars |
-    Dest_Eq_Res | Quant_Inst | Hypothesis | Lemma | Unit_Resolution |
-    Iff_True | Iff_False | Commutativity | Def_Axiom | Intro_Def | Apply_Def |
-    Iff_Oeq | Nnf_Pos | Nnf_Neg | Nnf_Star | Cnf_Star | Skolemize |
-    Modus_Ponens_Oeq | Th_Lemma of string
+  datatype z3_rule =
+    True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity | Symmetry | Transitivity |
+    Transitivity_Star | Monotonicity | Quant_Intro | Distributivity | And_Elim | Not_Or_Elim |
+    Rewrite | Rewrite_Star | Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars |
+    Dest_Eq_Res | Quant_Inst | Hypothesis | Lemma | Unit_Resolution | Iff_True | Iff_False |
+    Commutativity | Def_Axiom | Intro_Def | Apply_Def | Iff_Oeq | Nnf_Pos | Nnf_Neg | Nnf_Star |
+    Cnf_Star | Skolemize | Modus_Ponens_Oeq | Th_Lemma of string
+
   val is_assumption: z3_rule -> bool
   val string_of_rule: z3_rule -> string
 
@@ -40,16 +40,14 @@
 
 (* proof rules *)
 
-datatype z3_rule = True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity |
-  Symmetry | Transitivity | Transitivity_Star | Monotonicity | Quant_Intro |
-  Distributivity | And_Elim | Not_Or_Elim | Rewrite | Rewrite_Star |
-  Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars | Dest_Eq_Res |
-  Quant_Inst | Hypothesis | Lemma | Unit_Resolution | Iff_True | Iff_False |
-  Commutativity | Def_Axiom | Intro_Def | Apply_Def | Iff_Oeq | Nnf_Pos |
-  Nnf_Neg | Nnf_Star | Cnf_Star | Skolemize | Modus_Ponens_Oeq |
-  Th_Lemma of string
-  (* TODO: some proof rules come with further information
-     that is currently dropped by the parser *)
+datatype z3_rule =
+  True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity | Symmetry | Transitivity |
+  Transitivity_Star | Monotonicity | Quant_Intro | Distributivity | And_Elim | Not_Or_Elim |
+  Rewrite | Rewrite_Star | Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars |
+  Dest_Eq_Res | Quant_Inst | Hypothesis | Lemma | Unit_Resolution | Iff_True | Iff_False |
+  Commutativity | Def_Axiom | Intro_Def | Apply_Def | Iff_Oeq | Nnf_Pos | Nnf_Neg | Nnf_Star |
+  Cnf_Star | Skolemize | Modus_Ponens_Oeq | Th_Lemma of string
+  (* some proof rules include further information that is currently dropped by the parser *)
 
 val rule_names = Symtab.make [
   ("true-axiom", True_Axiom),
@@ -102,7 +100,7 @@
     SOME rule => rule
   | NONE => error ("unknown Z3 proof rule " ^ quote name))
 
-fun string_of_rule (Th_Lemma kind) = "th-lemma " ^ kind
+fun string_of_rule (Th_Lemma kind) = "th-lemma" ^ (if kind = "" then "" else " " ^ kind)
   | string_of_rule r =
       let fun eq_rule (s, r') = if r = r' then SOME s else NONE
       in the (Symtab.get_first eq_rule rule_names) end
@@ -118,7 +116,7 @@
   bounds: string list}
 
 fun mk_node id rule prems concl bounds =
-  Z3_Node {id=id, rule=rule, prems=prems, concl=concl, bounds=bounds}
+  Z3_Node {id = id, rule = rule, prems = prems, concl = concl, bounds = bounds}
 
 datatype z3_step = Z3_Step of {
   id: int,
@@ -129,8 +127,8 @@
   is_fix_step: bool}
 
 fun mk_step id rule prems concl fixes is_fix_step =
-  Z3_Step {id=id, rule=rule, prems=prems, concl=concl, fixes=fixes,
-    is_fix_step=is_fix_step}
+  Z3_Step {id = id, rule = rule, prems = prems, concl = concl, fixes = fixes,
+    is_fix_step = is_fix_step}
 
 
 (* proof parser *)
@@ -161,7 +159,7 @@
       in
         cx
         |> fold_map node_of ps
-        ||>> with_fresh_names (term_of p)
+        ||>> `(with_fresh_names (term_of p))
         ||>> next_id
         |>> (fn ((prems, (t, ns)), id) => mk_node id r prems t ns)
       end
@@ -188,10 +186,8 @@
     val ts = dest_seq (SMTLIB2.parse lines)
     val (node, cx) = parse' ts (empty_context ctxt typs funs)
   in (node, ctxt_of cx) end
-  handle SMTLIB2.PARSE (l, msg) =>
-           error ("parsing error at line " ^ string_of_int l ^ ": " ^ msg)
-       | SMTLIB2_PARSE (msg, t) =>
-           error (msg ^ ": " ^ SMTLIB2.str_of t)
+  handle SMTLIB2.PARSE (l, msg) => error ("parsing error at line " ^ string_of_int l ^ ": " ^ msg)
+       | SMTLIB2_PARSE (msg, t) => error (msg ^ ": " ^ SMTLIB2.str_of t)
 
 
 (* handling of bound variables *)
@@ -253,7 +249,7 @@
 
 fun dest_alls t = dest_all (Term.maxidx_of_term t + 1) t
 
-fun match_rule ctxt env (Z3_Node {bounds=bs', concl=t', ...}) bs t =
+fun match_rule ctxt env (Z3_Node {bounds = bs', concl = t', ...}) bs t =
   let
     val t'' = singleton (Variable.polymorphic ctxt) t'
     val (i, obj) = dest_alls (subst_types ctxt env bs t)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -10,7 +10,8 @@
 sig
   type fact = Sledgehammer_Fact.fact
   type fact_override = Sledgehammer_Fact.fact_override
-  type minimize_command = Sledgehammer_Proof_Methods.minimize_command
+  type proof_method = Sledgehammer_Proof_Methods.proof_method
+  type play_outcome = Sledgehammer_Proof_Methods.play_outcome
   type mode = Sledgehammer_Prover.mode
   type params = Sledgehammer_Prover.params
 
@@ -19,10 +20,11 @@
   val timeoutN : string
   val unknownN : string
 
+  val play_one_line_proof : bool -> Time.time -> (string * 'a) list -> Proof.state -> int ->
+    proof_method * proof_method list list -> (string * 'a) list * (proof_method * play_outcome)
   val string_of_factss : (string * fact list) list -> string
   val run_sledgehammer : params -> mode -> (string -> unit) option -> int -> fact_override ->
-    ((string * string list) list -> string -> minimize_command) -> Proof.state ->
-    bool * (string * Proof.state)
+    Proof.state -> bool * (string * Proof.state)
 end;
 
 structure Sledgehammer : SLEDGEHAMMER =
@@ -34,6 +36,9 @@
 open Sledgehammer_Util
 open Sledgehammer_Fact
 open Sledgehammer_Proof_Methods
+open Sledgehammer_Isar_Proof
+open Sledgehammer_Isar_Preplay
+open Sledgehammer_Isar_Minimize
 open Sledgehammer_Prover
 open Sledgehammer_Prover_ATP
 open Sledgehammer_Prover_Minimize
@@ -49,17 +54,61 @@
 fun max_outcome_code codes =
   NONE
   |> fold (fn candidate =>
-              fn accum as SOME _ => accum
-               | NONE => if member (op =) codes candidate then SOME candidate else NONE)
-       ordered_outcome_codes
+      fn accum as SOME _ => accum
+       | NONE => if member (op =) codes candidate then SOME candidate else NONE)
+    ordered_outcome_codes
   |> the_default unknownN
 
 fun prover_description verbose name num_facts =
   (quote name,
    if verbose then " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts else "")
 
-fun launch_prover (params as {debug, verbose, spy, blocking, max_facts, timeout, expect, ...}) mode
-    output_result minimize_command only learn
+fun is_metis_method (Metis_Method _) = true
+  | is_metis_method _ = false
+
+fun play_one_line_proof minimize timeout used_facts state i (preferred_meth, methss) =
+  if timeout = Time.zeroTime then
+    (used_facts, (preferred_meth, Play_Timed_Out Time.zeroTime))
+  else
+    let
+      val fact_names = map fst used_facts
+
+      val {context = ctxt, facts = chained, goal} = Proof.goal state
+      val (_, hyp_ts, concl_t) = strip_subgoal goal i ctxt
+      val goal_t = Logic.list_implies (map prop_of chained @ hyp_ts, concl_t)
+
+      fun try_methss ress [] =
+          (used_facts,
+           (case AList.lookup (op =) ress preferred_meth of
+             SOME play => (preferred_meth, play)
+           | NONE => hd (sort (play_outcome_ord o pairself snd) (rev ress))))
+        | try_methss ress (meths :: methss) =
+          let
+            fun mk_step fact_names meths =
+              Prove ([], [], ("", 0), goal_t, [], ([], fact_names), meths, "")
+          in
+            (case preplay_isar_step ctxt timeout [] (mk_step fact_names meths) of
+              (res as (meth, Played time)) :: _ =>
+              (* if a fact is needed by an ATP, it will be needed by "metis" *)
+              if not minimize orelse is_metis_method meth then
+                (used_facts, res)
+              else
+                let
+                  val (time', used_names') =
+                    minimized_isar_step ctxt time (mk_step fact_names [meth])
+                    ||> (facts_of_isar_step #> snd)
+                  val used_facts' = filter (member (op =) used_names' o fst) used_facts
+                in
+                  (used_facts', (meth, Played time'))
+                end
+            | ress' => try_methss (ress' @ ress) methss)
+          end
+    in
+      try_methss [] methss
+    end
+
+fun launch_prover (params as {debug, verbose, spy, blocking, max_facts, minimize, timeout,
+      preplay_timeout, expect, ...}) mode output_result only learn
     {comment, state, goal, subgoal, subgoal_count, factss as (_, facts) :: _} name =
   let
     val ctxt = Proof.context_of state
@@ -75,10 +124,9 @@
       {comment = comment, state = state, goal = goal, subgoal = subgoal,
        subgoal_count = subgoal_count,
        factss = factss
-         |> map (apsnd ((not (is_ho_atp ctxt name)
-                         ? filter_out (fn ((_, (_, Induction)), _) => true
-                                        | _ => false))
-                        #> take num_facts))}
+       |> map (apsnd ((not (is_ho_atp ctxt name)
+           ? filter_out (fn ((_, (_, Induction)), _) => true | _ => false))
+         #> take num_facts))}
 
     fun print_used_facts used_facts used_from =
       tag_list 1 used_from
@@ -115,8 +163,8 @@
             |> AList.group (op =)
             |> map (fn (indices, fact_filters) => commas fact_filters ^ ": " ^ indices)
         in
-          "Success: Found proof with " ^ string_of_int num_used_facts ^
-          " of " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
+          "Success: Found proof with " ^ string_of_int num_used_facts ^ " of " ^
+          string_of_int num_facts ^ " fact" ^ plural_s num_facts ^
           (if num_used_facts = 0 then "" else ": " ^ commas filter_infos)
         end
       | spying_str_of_res {outcome = SOME failure, ...} =
@@ -124,16 +172,17 @@
 
     fun really_go () =
       problem
-      |> get_minimizing_prover ctxt mode learn name params minimize_command
-      |> verbose
-         ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} =>
-                   print_used_facts used_facts used_from
-                 | _ => ())
+      |> get_minimizing_prover ctxt mode learn name params
+      |> verbose ? tap (fn {outcome = NONE, used_facts as _ :: _, used_from, ...} =>
+          print_used_facts used_facts used_from
+        | _ => ())
       |> spy ? tap (fn res => spying spy (fn () => (state, subgoal, name, spying_str_of_res res)))
-      |> (fn {outcome, preplay, message, message_tail, ...} =>
-             (if outcome = SOME ATP_Proof.TimedOut then timeoutN
-              else if is_some outcome then noneN
-              else someN, fn () => message (Lazy.force preplay) ^ message_tail))
+      |> (fn {outcome, used_facts, preferred_methss, message, ...} =>
+        (if outcome = SOME ATP_Proof.TimedOut then timeoutN
+         else if is_some outcome then noneN
+         else someN,
+         fn () => message (fn () => play_one_line_proof minimize preplay_timeout used_facts state
+           subgoal preferred_methss)))
 
     fun go () =
       let
@@ -203,7 +252,7 @@
       (if filter = "" then "" else quote filter ^ ": ") ^ string_of_facts facts) factss)
 
 fun run_sledgehammer (params as {verbose, spy, blocking, provers, max_facts, ...}) mode
-    output_result i (fact_override as {only, ...}) minimize_command state =
+    output_result i (fact_override as {only, ...}) state =
   if null provers then
     error "No prover is set."
   else
@@ -260,7 +309,7 @@
               {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
                factss = factss}
             val learn = mash_learn_proof ctxt params (prop_of goal) all_facts
-            val launch = launch_prover params mode output_result minimize_command only learn
+            val launch = launch_prover params mode output_result only learn
           in
             if mode = Auto_Try then
               (unknownN, state)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -31,7 +31,6 @@
 val z3N = "z3"
 
 val runN = "run"
-val minN = "min"
 val messagesN = "messages"
 val supported_proversN = "supported_provers"
 val kill_allN = "kill_all"
@@ -92,17 +91,17 @@
    ("max_mono_iters", "smart"),
    ("max_new_mono_instances", "smart"),
    ("isar_proofs", "smart"),
-   ("compress", "10"),
+   ("compress", "smart"),
    ("try0", "true"),
    ("smt_proofs", "smart"),
    ("slice", "true"),
-   ("minimize", "smart"),
-   ("preplay_timeout", "2")]
+   ("minimize", "true"),
+   ("preplay_timeout", "1")]
 
 val alias_params =
   [("prover", ("provers", [])), (* undocumented *)
    ("dont_preplay", ("preplay_timeout", ["0"])),
-   ("dont_compress", ("compress", ["0"])),
+   ("dont_compress", ("compress", ["1"])),
    ("isar_proof", ("isar_proofs", [])) (* legacy *)]
 val negated_alias_params =
   [("no_debug", "debug"),
@@ -119,9 +118,6 @@
    ("dont_try0", "try0"),
    ("no_smt_proofs", "smt_proofs")]
 
-val params_not_for_minimize =
-  ["provers", "blocking", "fact_filter", "max_facts", "fact_thresholds", "slice", "minimize"];
-
 val property_dependent_params = ["provers", "timeout"]
 
 fun is_known_raw_param s =
@@ -295,11 +291,11 @@
     val max_new_mono_instances =
       lookup_option lookup_int "max_new_mono_instances"
     val isar_proofs = lookup_option lookup_bool "isar_proofs"
-    val compress = Real.max (1.0, lookup_real "compress")
+    val compress = Option.map (curry Real.max 1.0) (lookup_option lookup_real "compress")
     val try0 = lookup_bool "try0"
     val smt_proofs = lookup_option lookup_bool "smt_proofs"
     val slice = mode <> Auto_Try andalso lookup_bool "slice"
-    val minimize = if mode = Auto_Try then NONE else lookup_option lookup_bool "minimize"
+    val minimize = mode <> Auto_Try andalso lookup_bool "minimize"
     val timeout = lookup_time "timeout"
     val preplay_timeout = if mode = Auto_Try then Time.zeroTime else lookup_time "preplay_timeout"
     val expect = lookup_string "expect"
@@ -318,33 +314,6 @@
 
 (* Sledgehammer the given subgoal *)
 
-val default_minimize_prover = metisN
-
-fun is_raw_param_relevant_for_minimize (name, _) = not (member (op =) params_not_for_minimize name)
-
-fun string_of_raw_param (key, values) =
-  key ^ (case implode_param values of "" => "" | value => " = " ^ value)
-
-fun nice_string_of_raw_param (p as (key, ["false"])) =
-    (case AList.find (op =) negated_alias_params key of
-       [neg] => neg
-     | _ => string_of_raw_param p)
-  | nice_string_of_raw_param p = string_of_raw_param p
-
-fun minimize_command override_params i more_override_params prover_name fact_names =
-  let
-    val params =
-      (override_params |> filter_out (AList.defined (op =) more_override_params o fst)) @
-      more_override_params
-      |> filter is_raw_param_relevant_for_minimize
-      |> map nice_string_of_raw_param
-      |> (if prover_name = default_minimize_prover then I else cons prover_name)
-      |> space_implode ", "
-  in
-    sledgehammerN ^ " " ^ minN ^ (if params = "" then "" else enclose " [" "]" params) ^
-    " (" ^ space_implode " " fact_names ^ ")" ^ (if i = 1 then "" else " " ^ string_of_int i)
-  end
-
 val default_learn_prover_timeout = 2.0
 
 fun hammer_away override_params subcommand opt_i fact_override state =
@@ -357,20 +326,9 @@
   in
     if subcommand = runN then
       let val i = the_default 1 opt_i in
-        run_sledgehammer (get_params Normal thy override_params) Normal NONE i fact_override
-          (minimize_command override_params i) state
-        |> K ()
+        ignore (run_sledgehammer (get_params Normal thy override_params) Normal NONE i fact_override
+          state)
       end
-    else if subcommand = minN then
-      let
-        val ctxt = ctxt |> Config.put instantiate_inducts false
-        val i = the_default 1 opt_i
-        val params =
-          get_params Minimize thy (("provers", [default_minimize_prover]) :: override_params)
-        val goal = prop_of (#goal (Proof.goal state))
-        val facts = nearly_all_facts ctxt false fact_override Symtab.empty Termtab.empty [] [] goal
-        val learn = mash_learn_proof ctxt params goal facts
-      in run_minimize params learn i (#add fact_override) state end
     else if subcommand = messagesN then
       messages opt_i
     else if subcommand = supported_proversN then
@@ -391,8 +349,7 @@
                 ([("timeout", [string_of_real default_learn_prover_timeout]),
                   ("slice", ["false"])] @
                  override_params @
-                 [("minimize", ["true"]),
-                  ("preplay_timeout", ["0"])]))
+                 [("preplay_timeout", ["0"])]))
            fact_override (#facts (Proof.goal state))
            (subcommand = learn_proverN orelse subcommand = relearn_proverN))
     else if subcommand = running_learnersN then
@@ -407,7 +364,8 @@
   Toplevel.unknown_proof o
   Toplevel.keep (hammer_away params subcommand opt_i fact_override o Toplevel.proof_of)
 
-fun string_of_raw_param (name, value) = name ^ " = " ^ implode_param value
+fun string_of_raw_param (key, values) =
+  key ^ (case implode_param values of "" => "" | value => " = " ^ value)
 
 fun sledgehammer_params_trans params =
   Toplevel.theory (fold set_default_raw_param params #> tap (fn thy =>
@@ -421,7 +379,7 @@
 val parse_value = Scan.repeat1 (Parse.xname || Parse.float_number || parse_query_bang)
 val parse_param = parse_key -- Scan.optional (@{keyword "="} |-- parse_value) []
 val parse_params = Scan.optional (Args.bracks (Parse.list parse_param)) []
-val parse_fact_refs = Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse_Spec.xthm)
+val parse_fact_refs = Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) Parse.xthm)
 val parse_fact_override_chunk =
   (Args.add |-- Args.colon |-- parse_fact_refs >> add_fact_override)
   || (Args.del |-- Args.colon |-- parse_fact_refs >> del_fact_override)
@@ -446,8 +404,7 @@
     val mode = if auto then Auto_Try else Try
     val i = 1
   in
-    run_sledgehammer (get_params mode thy []) mode NONE i no_fact_override (minimize_command [] i)
-      state
+    run_sledgehammer (get_params mode thy []) mode NONE i no_fact_override state
   end
 
 val _ = Try.tool_setup (sledgehammerN, (40, @{system_option auto_sledgehammer}, try_sledgehammer))
@@ -456,24 +413,23 @@
   Query_Operation.register sledgehammerN (fn {state = st, args, output_result} =>
     (case try Toplevel.proof_of st of
       SOME state =>
-        let
-          val thy = Proof.theory_of state
-          val ctxt = Proof.context_of state
-          val [provers_arg, isar_proofs_arg] = args
+      let
+        val thy = Proof.theory_of state
+        val ctxt = Proof.context_of state
+        val [provers_arg, isar_proofs_arg] = args
 
-          val override_params =
-            ((if provers_arg = "" then [] else [("provers", space_explode " " provers_arg)]) @
-              [("isar_proofs", [isar_proofs_arg]),
-               ("blocking", ["true"]),
-               ("minimize", ["true"]),
-               ("debug", ["false"]),
-               ("verbose", ["false"]),
-               ("overlord", ["false"])])
-            |> map (normalize_raw_param ctxt)
-
-          val _ = run_sledgehammer (get_params Normal thy override_params) Normal
-            (SOME output_result) 1 no_fact_override (minimize_command override_params 1) state
-        in () end
+        val override_params =
+          ((if provers_arg = "" then [] else [("provers", space_explode " " provers_arg)]) @
+            [("isar_proofs", [isar_proofs_arg]),
+             ("blocking", ["true"]),
+             ("debug", ["false"]),
+             ("verbose", ["false"]),
+             ("overlord", ["false"])])
+          |> map (normalize_raw_param ctxt)
+      in
+        ignore (run_sledgehammer (get_params Normal thy override_params) Normal
+          (SOME output_result) 1 no_fact_override state)
+      end
     | NONE => error "Unknown proof context"))
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -14,15 +14,15 @@
   type fact = (string * stature) * thm
 
   type fact_override =
-    {add : (Facts.ref * Attrib.src list) list,
-     del : (Facts.ref * Attrib.src list) list,
+    {add : (Facts.ref * Token.src list) list,
+     del : (Facts.ref * Token.src list) list,
      only : bool}
 
   val instantiate_inducts : bool Config.T
   val no_fact_override : fact_override
 
   val fact_of_ref : Proof.context -> unit Symtab.table -> thm list -> status Termtab.table ->
-    Facts.ref * Attrib.src list -> ((string * stature) * thm) list
+    Facts.ref * Token.src list -> ((string * stature) * thm) list
   val backquote_thm : Proof.context -> thm -> string
   val is_blacklisted_or_something : Proof.context -> bool -> string -> bool
   val clasimpset_rule_table_of : Proof.context -> status Termtab.table
@@ -50,8 +50,8 @@
 type fact = (string * stature) * thm
 
 type fact_override =
-  {add : (Facts.ref * Attrib.src list) list,
-   del : (Facts.ref * Attrib.src list) list,
+  {add : (Facts.ref * Token.src list) list,
+   del : (Facts.ref * Token.src list) list,
    only : bool}
 
 (* gracefully handle huge background theories *)
@@ -77,8 +77,7 @@
   | explode_interval max (Facts.From i) = i upto i + max - 1
   | explode_interval _ (Facts.Single i) = [i]
 
-val backquote =
-  raw_explode #> map (fn "`" => "\\`" | s => s) #> implode #> enclose "`" "`"
+val backquote = raw_explode #> map (fn "`" => "\\`" | s => s) #> implode #> enclose "`" "`"
 
 (* unfolding these can yield really huge terms *)
 val risky_defs = @{thms Bit0_def Bit1_def}
@@ -101,9 +100,8 @@
   else Local
 
 val may_be_induction =
-  exists_subterm (fn Var (_, Type (@{type_name fun}, [_, T])) =>
-                     body_type T = @{typ bool}
-                   | _ => false)
+  exists_subterm (fn Var (_, Type (@{type_name fun}, [_, T])) => body_type T = @{typ bool}
+    | _ => false)
 
 (* TODO: get rid of *)
 fun normalize_vars t =
@@ -118,14 +116,12 @@
 
     fun norm (t $ u) = norm t ##>> norm u #>> op $
       | norm (Const (s, T)) = normT T #>> curry Const s
-      | norm (Var (z as (_, T))) =
-        normT T
+      | norm (Var (z as (_, T))) = normT T
         #> (fn (T, (accumT, (known, n))) =>
-               (case find_index (equal z) known of
-                 ~1 => (Var ((Name.uu, n), T), (accumT, (z :: known, n + 1)))
-               | j => (Var ((Name.uu, n - j - 1), T), (accumT, (known, n)))))
-      | norm (Abs (_, T, t)) =
-        norm t ##>> normT T #>> (fn (t, T) => Abs (Name.uu, T, t))
+          (case find_index (equal z) known of
+            ~1 => (Var ((Name.uu, n), T), (accumT, (z :: known, n + 1)))
+          | j => (Var ((Name.uu, n - j - 1), T), (accumT, (known, n)))))
+      | norm (Abs (_, T, t)) = norm t ##>> normT T #>> (fn (t, T) => Abs (Name.uu, T, t))
       | norm (Bound j) = pair (Bound j)
       | norm (Free (s, T)) = normT T #>> curry Free s
   in fst (norm t (([], 0), ([], 0))) end
@@ -163,50 +159,43 @@
 fun fact_of_ref ctxt reserved chained css (xthm as (xref, args)) =
   let
     val ths = Attrib.eval_thms ctxt [xthm]
-    val bracket =
-      map (enclose "[" "]" o Pretty.str_of o Args.pretty_src ctxt) args
-      |> implode
+    val bracket = implode (map (enclose "[" "]" o Pretty.str_of o Token.pretty_src ctxt) args)
 
     fun nth_name j =
       (case xref of
-        Facts.Fact s => backquote s ^ bracket
+        Facts.Fact s => backquote (simplify_spaces (unyxml s)) ^ bracket
       | Facts.Named (("", _), _) => "[" ^ bracket ^ "]"
-      | Facts.Named ((name, _), NONE) =>
-        make_name reserved (length ths > 1) (j + 1) name ^ bracket
+      | Facts.Named ((name, _), NONE) => make_name reserved (length ths > 1) (j + 1) name ^ bracket
       | Facts.Named ((name, _), SOME intervals) =>
         make_name reserved true
           (nth (maps (explode_interval (length ths)) intervals) j) name ^ bracket)
 
     fun add_nth th (j, rest) =
       let val name = nth_name j in
-        (j + 1, ((name, stature_of_thm false [] chained css name th), th)
-                :: rest)
+        (j + 1, ((name, stature_of_thm false [] chained css name th), th) :: rest)
       end
   in
     (0, []) |> fold add_nth ths |> snd
   end
 
-(* Reject theorems with names like "List.filter.filter_list_def" or
-  "Accessible_Part.acc.defs", as these are definitions arising from packages. *)
+(* Reject theorems with names like "List.filter.filter_list_def" or "Accessible_Part.acc.defs", as
+   these are definitions arising from packages. *)
 fun is_package_def s =
   let val ss = Long_Name.explode s in
     length ss > 2 andalso not (hd ss = "local") andalso
     exists (fn suf => String.isSuffix suf s)
-           ["_case_def", "_rec_def", "_size_def", "_size_overloaded_def"]
+      ["_case_def", "_rec_def", "_size_def", "_size_overloaded_def"]
   end
 
 (* FIXME: put other record thms here, or declare as "no_atp" *)
 fun multi_base_blacklist ctxt ho_atp =
-  ["defs", "select_defs", "update_defs", "split", "splits", "split_asm",
-   "ext_cases", "eq.simps", "eq.refl", "nchotomy", "case_cong",
-   "weak_case_cong", "nat_of_char_simps", "nibble.simps",
+  ["defs", "select_defs", "update_defs", "split", "splits", "split_asm", "ext_cases", "eq.simps",
+   "eq.refl", "nchotomy", "case_cong", "case_cong_weak", "nat_of_char_simps", "nibble.simps",
    "nibble.distinct"]
-  |> not (ho_atp orelse Config.get ctxt instantiate_inducts) ?
-        append ["induct", "inducts"]
+  |> not (ho_atp orelse Config.get ctxt instantiate_inducts) ? append ["induct", "inducts"]
   |> map (prefix Long_Name.separator)
 
-(* The maximum apply depth of any "metis" call in "Metis_Examples" (on
-   2007-10-31) was 11. *)
+(* The maximum apply depth of any "metis" call in "Metis_Examples" (back in 2007) was 11. *)
 val max_apply_depth = 18
 
 fun apply_depth (f $ t) = Int.max (apply_depth f, apply_depth t + 1)
@@ -217,23 +206,20 @@
 
 (* FIXME: Ad hoc list *)
 val technical_prefixes =
-  ["ATP", "Code_Evaluation", "Datatype", "Enum", "Lazy_Sequence",
-   "Limited_Sequence", "Meson", "Metis", "Nitpick",
-   "Quickcheck_Random", "Quickcheck_Exhaustive", "Quickcheck_Narrowing",
+  ["ATP", "Code_Evaluation", "Datatype", "Enum", "Lazy_Sequence", "Limited_Sequence", "Meson",
+   "Metis", "Nitpick", "Quickcheck_Random", "Quickcheck_Exhaustive", "Quickcheck_Narrowing",
    "Random_Sequence", "Sledgehammer", "SMT"]
   |> map (suffix Long_Name.separator)
 
-fun is_technical_const (s, _) =
-  exists (fn pref => String.isPrefix pref s) technical_prefixes
+fun is_technical_const s = exists (fn pref => String.isPrefix pref s) technical_prefixes
 
 (* FIXME: make more reliable *)
 val sep_class_sep = Long_Name.separator ^ "class" ^ Long_Name.separator
 
-fun is_low_level_class_const (s, _) =
+fun is_low_level_class_const s =
   s = @{const_name equal_class.equal} orelse String.isSubstring sep_class_sep s
 
 val sep_that = Long_Name.separator ^ Obtain.thatN
-
 val skolem_thesis = Name.skolem Auto_Bind.thesisN
 
 fun is_that_fact th =
@@ -259,11 +245,11 @@
       | is_interesting_subterm _ = false
 
     fun interest_of_bool t =
-      if exists_Const (is_technical_const orf is_low_level_class_const orf
-                       type_has_top_sort o snd) t then
+      if exists_Const ((is_technical_const o fst) orf (is_low_level_class_const o fst) orf
+          type_has_top_sort o snd) t then
         Deal_Breaker
       else if exists_type (exists_subtype (curry (op =) @{typ prop})) t orelse
-              not (exists_subterm is_interesting_subterm t) then
+          not (exists_subterm is_interesting_subterm t) then
         Boring
       else
         Interesting
@@ -281,8 +267,7 @@
 
     val t = prop_of th
   in
-    (interest_of_prop [] t <> Interesting andalso
-     not (Thm.eq_thm_prop (@{thm ext}, th))) orelse
+    (interest_of_prop [] t <> Interesting andalso not (Thm.eq_thm_prop (@{thm ext}, th))) orelse
     is_that_fact th
   end
 
@@ -297,19 +282,18 @@
    prefixes of all free variables. In the worse case scenario, where the fact
    won't be resolved correctly, the user can fix it manually, e.g., by giving a
    name to the offending fact. *)
-fun all_prefixes_of s =
-  map (fn i => String.extract (s, 0, SOME i)) (1 upto size s - 1)
+fun all_prefixes_of s = map (fn i => String.extract (s, 0, SOME i)) (1 upto size s - 1)
 
 fun close_form t =
   (t, [] |> Term.add_free_names t |> maps all_prefixes_of)
   |> fold (fn ((s, i), T) => fn (t', taken) =>
-              let val s' = singleton (Name.variant_list taken) s in
-                ((if fastype_of t' = HOLogic.boolT then HOLogic.all_const
-                  else Logic.all_const) T
-                 $ Abs (s', T, abstract_over (Var ((s, i), T), t')),
-                 s' :: taken)
-              end)
-          (Term.add_vars t [] |> sort_wrt (fst o fst))
+      let val s' = singleton (Name.variant_list taken) s in
+        ((if fastype_of t' = HOLogic.boolT then HOLogic.all_const
+          else Logic.all_const) T
+         $ Abs (s', T, abstract_over (Var ((s, i), T), t')),
+         s' :: taken)
+      end)
+    (Term.add_vars t [] |> sort_wrt (fst o fst))
   |> fst
 
 fun backquote_term ctxt = close_form #> hackish_string_of_term ctxt #> backquote
@@ -322,26 +306,23 @@
       Termtab.empty
     else
       let
-        fun add stature th =
-          Termtab.update (normalize_vars (prop_of th), stature)
+        fun add stature th = Termtab.update (normalize_vars (prop_of th), stature)
 
-        val {safeIs, (* safeEs, *) hazIs, (* hazEs, *) ...} =
-          ctxt |> claset_of |> Classical.rep_cs
+        val {safeIs, (* safeEs, *) hazIs, (* hazEs, *) ...} = ctxt |> claset_of |> Classical.rep_cs
         val intros = Item_Net.content safeIs @ Item_Net.content hazIs
 (* Add once it is used:
-        val elims =
-          Item_Net.content safeEs @ Item_Net.content hazEs
+        val elims = Item_Net.content safeEs @ Item_Net.content hazEs
           |> map Classical.classical_rule
 *)
-        val specs = ctxt |> Spec_Rules.get
-        val (rec_defs, nonrec_defs) =
-          specs |> filter (curry (op =) Spec_Rules.Equational o fst)
-                |> maps (snd o snd)
-                |> filter_out (member Thm.eq_thm_prop risky_defs)
-                |> List.partition (is_rec_def o prop_of)
-        val spec_intros =
-          specs |> filter (member (op =) [Spec_Rules.Inductive, Spec_Rules.Co_Inductive] o fst)
-                |> maps (snd o snd)
+        val specs = Spec_Rules.get ctxt
+        val (rec_defs, nonrec_defs) = specs
+          |> filter (curry (op =) Spec_Rules.Equational o fst)
+          |> maps (snd o snd)
+          |> filter_out (member Thm.eq_thm_prop risky_defs)
+          |> List.partition (is_rec_def o prop_of)
+        val spec_intros = specs
+          |> filter (member (op =) [Spec_Rules.Inductive, Spec_Rules.Co_Inductive] o fst)
+          |> maps (snd o snd)
       in
         Termtab.empty
         |> fold (add Simp o snd) simps
@@ -358,7 +339,7 @@
 fun normalize_eq (@{const Trueprop} $ (t as (t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
     if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else t0 $ t2 $ t1
   | normalize_eq (@{const Trueprop} $ (t as @{const Not}
-        $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2))) =
+      $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2))) =
     if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t else HOLogic.mk_not (t0 $ t2 $ t1)
   | normalize_eq (Const (@{const_name Pure.eq}, Type (_, [T, _])) $ t1 $ t2) =
     (if Term_Ord.fast_term_ord (t1, t2) <> GREATER then (t1, t2) else (t2, t1))
@@ -398,8 +379,7 @@
 
 fun struct_induct_rule_on th =
   (case Logic.strip_horn (prop_of th) of
-    (prems, @{const Trueprop}
-            $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) =>
+    (prems, @{const Trueprop} $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) =>
     if not (is_TVar ind_T) andalso length prems > 1 andalso
        exists (exists_subterm (curry (op aconv) p)) prems andalso
        not (exists (exists_subterm (curry (op aconv) a)) prems) then
@@ -416,13 +396,14 @@
         if (s, T) = ind_x orelse can dest_funT T then t else Var ((s, 0), T)
       | varify_noninducts t = t
 
-    val p_inst =
-      concl_prop |> map_aterms varify_noninducts |> close_form
-                 |> lambda (Free ind_x)
-                 |> hackish_string_of_term ctxt
+    val p_inst = concl_prop
+      |> map_aterms varify_noninducts
+      |> close_form
+      |> lambda (Free ind_x)
+      |> hackish_string_of_term ctxt
   in
-    ((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]",
-      stature), th |> Rule_Insts.read_instantiate ctxt [((p_name, 0), p_inst)] [])
+    ((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]", stature),
+     th |> Rule_Insts.read_instantiate ctxt [((p_name, 0), p_inst)] [])
   end
 
 fun type_match thy (T1, T2) =
@@ -436,7 +417,7 @@
       stmt_xs
       |> filter (fn (_, T) => type_match thy (T, ind_T))
       |> map_filter (try (TimeLimit.timeLimit instantiate_induct_timeout
-           (instantiate_induct_rule ctxt stmt p_name ax)))
+        (instantiate_induct_rule ctxt stmt p_name ax)))
     end
   | NONE => [ax])
 
@@ -451,7 +432,9 @@
         (hyp_ts |> filter_out (null o external_frees), concl_t)
         |> Logic.list_implies |> Object_Logic.atomize_term thy
       val ind_stmt_xs = external_frees ind_stmt
-    in maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs) end
+    in
+      maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs)
+    end
   else
     I
 
@@ -464,10 +447,8 @@
     val thy = Proof_Context.theory_of ctxt
     val global_facts = Global_Theory.facts_of thy
     val is_too_complex =
-      if generous orelse fact_count global_facts >= max_facts_for_complex_check then
-        K false
-      else
-        is_too_complex
+      if generous orelse fact_count global_facts >= max_facts_for_complex_check then K false
+      else is_too_complex
     val local_facts = Proof_Context.facts_of ctxt
     val named_locals = local_facts |> Facts.dest_static true [global_facts]
     val assms = Assumption.all_assms_of ctxt
@@ -479,8 +460,7 @@
     val unnamed_locals =
       union Thm.eq_thm_prop (Facts.props local_facts) chained
       |> filter is_good_unnamed_local |> map (pair "" o single)
-    val full_space =
-      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
+    val full_space = Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
     val is_blacklisted_or_something = is_blacklisted_or_something ctxt ho_atp
 
     fun add_facts global foldx facts =
@@ -532,12 +512,12 @@
                  end)) ths (n, accum))
           end)
   in
-    (* The single-theorem names go before the multiple-theorem ones (e.g.,
-       "xxx" vs. "xxx(3)"), so that single names are preferred when both are
-       available. *)
-    `I [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
-          |> add_facts true Facts.fold_static global_facts global_facts
-          |> op @
+    (* The single-theorem names go before the multiple-theorem ones (e.g., "xxx" vs. "xxx(3)"), so
+       that single names are preferred when both are available. *)
+    `I []
+    |> add_facts false fold local_facts (unnamed_locals @ named_locals)
+    |> add_facts true Facts.fold_static global_facts global_facts
+    |> op @
   end
 
 fun nearly_all_facts ctxt ho_atp {add, del, only} reserved css chained hyp_ts concl_t =
@@ -545,20 +525,19 @@
     []
   else
     let
-      val chained =
-        chained
-        |> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
+      val chained = chained |> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
     in
       (if only then
          maps (map (fn ((name, stature), th) => ((K name, stature), th))
-               o fact_of_ref ctxt reserved chained css) add
+           o fact_of_ref ctxt reserved chained css) add
        else
          let
            val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del)
            val facts =
              all_facts ctxt false ho_atp reserved add chained css
              |> filter_out ((member Thm.eq_thm_prop del orf
-               (No_ATPs.member ctxt andf not o member Thm.eq_thm_prop add)) o snd)
+               (Named_Theorems.member ctxt @{named_theorems no_atp} andf
+                 not o member Thm.eq_thm_prop add)) o snd)
          in
            facts
          end)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -16,7 +16,7 @@
   val trace : bool Config.T
 
   type isar_params =
-    bool * (string option * string option) * Time.time * real * bool * bool
+    bool * (string option * string option) * Time.time * real option * bool * bool
     * (term, string) atp_step list * thm
 
   val proof_text : Proof.context -> bool -> bool option -> bool option -> (unit -> isar_params) ->
@@ -46,18 +46,29 @@
 
 val trace = Attrib.setup_config_bool @{binding sledgehammer_isar_trace} (K false)
 
+val e_definition_rule = "definition"
 val e_skolemize_rule = "skolemize"
+val leo2_extcnf_forall_neg_rule = "extcnf_forall_neg"
+val satallax_skolemize_rule = "tab_ex"
 val spass_pirate_datatype_rule = "DT"
 val vampire_skolemisation_rule = "skolemisation"
-(* TODO: Use "Z3_Proof.string_of_rule" once it is moved to Isabelle *)
-val z3_skolemize_rule = "sk"
-val z3_th_lemma_rule = "th-lemma"
+val veriT_la_generic_rule = "la_generic"
+val veriT_simp_arith_rule = "simp_arith"
+val veriT_tmp_ite_elim_rule = "tmp_ite_elim"
+val veriT_tmp_skolemize_rule = "tmp_skolemize"
+val z3_skolemize_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Skolemize
+val z3_th_lemma_rule_prefix = Z3_New_Proof.string_of_rule (Z3_New_Proof.Th_Lemma "")
+val zipperposition_cnf_rule = "cnf"
 
 val skolemize_rules =
-  [e_skolemize_rule, spass_skolemize_rule, vampire_skolemisation_rule, z3_skolemize_rule]
+  [e_definition_rule, e_skolemize_rule, leo2_extcnf_forall_neg_rule, satallax_skolemize_rule,
+   spass_skolemize_rule, vampire_skolemisation_rule, veriT_tmp_ite_elim_rule,
+   veriT_tmp_skolemize_rule, z3_skolemize_rule, zipperposition_cnf_rule]
 
 val is_skolemize_rule = member (op =) skolemize_rules
-val is_arith_rule = String.isPrefix z3_th_lemma_rule
+fun is_arith_rule rule =
+  String.isPrefix z3_th_lemma_rule_prefix rule orelse rule = veriT_simp_arith_rule orelse
+  rule = veriT_la_generic_rule
 val is_datatype_rule = String.isPrefix spass_pirate_datatype_rule
 
 fun raw_label_of_num num = (num, 0)
@@ -71,59 +82,74 @@
 fun add_line_pass1 (line as (name, role, t, rule, [])) lines =
     (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or
        definitions. *)
-    if role = Conjecture orelse role = Negated_Conjecture then line :: lines
-    else if t aconv @{prop True} then map (replace_dependencies_in_line (name, [])) lines
-    else if role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then line :: lines
-    else if role = Axiom then lines (* axioms (facts) need no proof lines *)
-    else map (replace_dependencies_in_line (name, [])) lines
+    if role = Conjecture orelse role = Negated_Conjecture then
+      line :: lines
+    else if t aconv @{prop True} then
+      map (replace_dependencies_in_line (name, [])) lines
+    else if role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then
+      line :: lines
+    else if role = Axiom then
+      lines (* axioms (facts) need no proof lines *)
+    else
+      map (replace_dependencies_in_line (name, [])) lines
   | add_line_pass1 line lines = line :: lines
 
-fun add_lines_pass2 res _ [] = rev res
-  | add_lines_pass2 res prev_t ((line as (name, role, t, rule, deps)) :: lines) =
+fun add_lines_pass2 res [] = rev res
+  | add_lines_pass2 res ((line as (name, role, t, rule, deps)) :: lines) =
     let
-      fun looks_boring () =
-        t aconv @{prop True} orelse t aconv @{prop False} orelse t aconv prev_t orelse
-        length deps < 2
+      fun normalize role =
+        role = Conjecture ? (HOLogic.dest_Trueprop #> s_not #> HOLogic.mk_Trueprop)
+
+      val norm_t = normalize role t
+      val is_duplicate =
+        exists (fn (prev_name, prev_role, prev_t, _, _) =>
+            member (op =) deps prev_name andalso
+            Term.aconv_untyped (normalize prev_role prev_t, norm_t))
+          res
+
+      fun looks_boring () = t aconv @{prop False} orelse length deps < 2
 
       fun is_skolemizing_line (_, _, _, rule', deps') =
         is_skolemize_rule rule' andalso member (op =) deps' name
+
       fun is_before_skolemize_rule () = exists is_skolemizing_line lines
     in
-      if role <> Plain orelse is_skolemize_rule rule orelse is_arith_rule rule orelse
-         is_datatype_rule rule orelse null lines orelse not (looks_boring ()) orelse
-         is_before_skolemize_rule () then
-        add_lines_pass2 (line :: res) t lines
+      if is_duplicate orelse
+          (role = Plain andalso not (is_skolemize_rule rule) andalso
+           not (is_arith_rule rule) andalso not (is_datatype_rule rule) andalso
+           not (null lines) andalso looks_boring () andalso not (is_before_skolemize_rule ())) then
+        add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines)
       else
-        add_lines_pass2 res t (map (replace_dependencies_in_line (name, deps)) lines)
+        add_lines_pass2 (line :: res) lines
     end
 
 type isar_params =
-  bool * (string option * string option) * Time.time * real * bool * bool
+  bool * (string option * string option) * Time.time * real option * bool * bool
   * (term, string) atp_step list * thm
 
 val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method, SATx_Method]
-val simp_based_methods = [Auto_Method, Simp_Method, Fastforce_Method, Force_Method]
+val basic_simp_based_methods = [Auto_Method, Simp_Method, Force_Method]
 val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method]
 
-val arith_methods = basic_arith_methods @ simp_based_methods @ basic_systematic_methods
+val arith_methods = basic_arith_methods @ basic_simp_based_methods @ basic_systematic_methods
 val datatype_methods = [Simp_Method, Simp_Size_Method]
-val systematic_methods0 = basic_systematic_methods @ basic_arith_methods @ simp_based_methods @
-  [Metis_Method (SOME no_typesN, NONE)]
-val rewrite_methods = simp_based_methods @ basic_systematic_methods @ basic_arith_methods
-val skolem_methods = basic_systematic_methods
+val systematic_methods =
+  basic_systematic_methods @ basic_arith_methods @ basic_simp_based_methods @
+  [Metis_Method (SOME full_typesN, NONE), Metis_Method (SOME no_typesN, NONE)]
+val rewrite_methods = basic_simp_based_methods @ basic_systematic_methods @ basic_arith_methods
+val skolem_methods = basic_systematic_methods @ [Auto_Choice_Method]
 
 fun isar_proof_text ctxt debug isar_proofs smt_proofs isar_params
-    (one_line_params as ((_, one_line_play), banner, used_facts, minimize_command, subgoal,
-       subgoal_count)) =
+    (one_line_params as ((used_facts, (_, one_line_play)), banner, subgoal, subgoal_count)) =
   let
     val _ = if debug then Output.urgent_message "Constructing Isar proof..." else ()
 
     fun generate_proof_text () =
       let
-        val (verbose, alt_metis_args, preplay_timeout, compress, try0, minimize, atp_proof, goal) =
+        val (verbose, alt_metis_args, preplay_timeout, compress, try0, minimize, atp_proof0, goal) =
           isar_params ()
 
-        val systematic_methods = insert (op =) (Metis_Method alt_metis_args) systematic_methods0
+        val systematic_methods' = insert (op =) (Metis_Method alt_metis_args) systematic_methods
 
         fun massage_methods (meths as meth :: _) =
           if not try0 then [meth]
@@ -135,7 +161,10 @@
         val ctxt = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd
 
         val do_preplay = preplay_timeout <> Time.zeroTime
-        val compress = if isar_proofs = NONE andalso do_preplay then 1000.0 else compress
+        val compress =
+          (case compress of
+            NONE => if isar_proofs = NONE andalso do_preplay then 1000.0 else 10.0
+          | SOME n => n)
 
         fun is_fixed ctxt = Variable.is_declared ctxt orf Name.is_skolem
         fun skolems_of ctxt t = Term.add_frees t [] |> filter_out (is_fixed ctxt o fst) |> rev
@@ -144,9 +173,8 @@
           if keep_role role then SOME ((raw_label_of_num num, t), rule) else NONE
 
         val atp_proof =
-          atp_proof
-          |> rpair [] |-> fold_rev add_line_pass1
-          |> add_lines_pass2 [] Term.dummy
+          fold_rev add_line_pass1 atp_proof0 []
+          |> add_lines_pass2 []
 
         val conjs =
           map_filter (fn (name, role, _, _, _) =>
@@ -169,7 +197,7 @@
         val (lems, _) =
           fold_map add_lemma (map_filter (get_role (curry (op =) Lemma)) atp_proof) ctxt
 
-        val bot = atp_proof |> List.last |> #1
+        val bot = #1 (List.last atp_proof)
 
         val refute_graph =
           atp_proof
@@ -196,18 +224,20 @@
 
         val rule_of_clause_id = fst o the o Symtab.lookup steps o fst
 
-        fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> close_form
+        val finish_off = close_form #> rename_bound_vars
+
+        fun prop_of_clause [(num, _)] = Symtab.lookup steps num |> the |> snd |> finish_off
           | prop_of_clause names =
             let
-              val lits = map (HOLogic.dest_Trueprop o snd)
-                (map_filter (Symtab.lookup steps o fst) names)
+              val lits =
+                map (HOLogic.dest_Trueprop o snd) (map_filter (Symtab.lookup steps o fst) names)
             in
               (case List.partition (can HOLogic.dest_not) lits of
                 (negs as _ :: _, pos as _ :: _) =>
                 s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos)
               | _ => fold (curry s_disj) lits @{term False})
             end
-            |> HOLogic.mk_Trueprop |> close_form
+            |> HOLogic.mk_Trueprop |> finish_off
 
         fun maybe_show outer c = (outer andalso eq_set (op =) (c, conjs)) ? cons Show
 
@@ -215,7 +245,7 @@
             accum
             |> (if tainted = [] then
                   cons (Prove (if outer then [Show] else [], [], no_label, concl_t, [],
-                    (the_list predecessor, []), massage_methods systematic_methods, ""))
+                    sort_facts (the_list predecessor, []), massage_methods systematic_methods', ""))
                 else
                   I)
             |> rev
@@ -226,12 +256,14 @@
               val rule = rule_of_clause_id id
               val skolem = is_skolemize_rule rule
 
-              val deps = fold add_fact_of_dependencies gamma ([], [])
+              val deps = ([], [])
+                |> fold add_fact_of_dependencies gamma
+                |> sort_facts
               val meths =
                 (if skolem then skolem_methods
                  else if is_arith_rule rule then arith_methods
                  else if is_datatype_rule rule then datatype_methods
-                 else systematic_methods)
+                 else systematic_methods')
                 |> massage_methods
 
               fun prove sub facts = Prove (maybe_show outer c [], [], l, t, sub, facts, meths, "")
@@ -241,16 +273,23 @@
                 (case gamma of
                   [g] =>
                   if skolem andalso is_clause_tainted g then
-                    let val subproof = Proof (skolems_of ctxt (prop_of_clause g), [], rev accum) in
-                      isar_steps outer (SOME l) [prove [subproof] ([], [])] infs
-                    end
+                    (case skolems_of ctxt (prop_of_clause g) of
+                      [] => steps_of_rest (prove [] deps)
+                    | skos =>
+                      let val subproof = Proof (skos, [], rev accum) in
+                        isar_steps outer (SOME l) [prove [subproof] ([], [])] infs
+                      end)
                   else
                     steps_of_rest (prove [] deps)
                 | _ => steps_of_rest (prove [] deps))
               else
                 steps_of_rest
-                  (if skolem then Prove ([], skolems_of ctxt t, l, t, [], deps, meths, "")
-                   else prove [] deps)
+                  (if skolem then
+                     (case skolems_of ctxt t of
+                       [] => prove [] deps
+                     | skos => Prove ([], skos, l, t, [], deps, meths, ""))
+                   else
+                     prove [] deps)
             end
           | isar_steps outer predecessor accum (Cases cases :: infs) =
             let
@@ -262,7 +301,7 @@
               val step =
                 Prove (maybe_show outer c [], [], l, t,
                   map isar_case (filter_out (null o snd) cases),
-                  (the_list predecessor, []), massage_methods systematic_methods, "")
+                  sort_facts (the_list predecessor, []), massage_methods systematic_methods', "")
             in
               isar_steps outer (SOME l) (step :: accum) infs
             end
@@ -318,18 +357,12 @@
                (keep_fastest_method_of_isar_step (!preplay_data)
                 #> minimize ? minimize_isar_step_dependencies ctxt preplay_data)
           |> tap (trace_isar_proof "Minimized")
-          (* It's not clear whether this is worth the trouble (and if so, "compress" has an
-             unnatural semantics): *)
-(*
-          |> minimize
-               ? (compress_isar_proof ctxt compress preplay_timeout preplay_data
-                  #> tap (trace_isar_proof "Compressed again"))
-*)
           |> `(preplay_outcome_of_isar_proof (!preplay_data))
           ||> (comment_isar_proof comment_of
                #> chain_isar_proof
                #> kill_useless_labels_in_isar_proof
-               #> relabel_isar_proof_nicely)
+               #> relabel_isar_proof_nicely
+               #> rationalize_obtains_in_isar_proofs ctxt)
       in
         (case add_isar_steps (steps_of_isar_proof isar_proof) 0 of
           1 =>
@@ -338,12 +371,11 @@
                (case isar_proof of
                  Proof (_, _, [Prove (_, _, _, _, _, (_, gfs), meth :: _, _)]) =>
                  let val used_facts' = filter (member (op =) gfs o fst) used_facts in
-                   ((meth, play_outcome), banner, used_facts', minimize_command, subgoal,
-                    subgoal_count)
+                   ((used_facts', (meth, play_outcome)), banner, subgoal, subgoal_count)
                  end)
              else
                one_line_params) ^
-          (if isar_proofs = SOME true then "\n(No structured proof available -- proof too simple.)"
+          (if isar_proofs = SOME true then "\n(No Isar proof available.)"
            else "")
         | num_steps =>
           let
@@ -352,7 +384,7 @@
               (if do_preplay then [string_of_play_outcome play_outcome] else [])
           in
             one_line_proof_text ctxt 0 one_line_params ^
-            "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
+            "\n\nIsar proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
             Active.sendback_markup [Markup.padding_command]
               (string_of_isar_proof ctxt subgoal subgoal_count isar_proof)
           end)
@@ -375,7 +407,7 @@
   | Play_Failed => true)
 
 fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained
-    (one_line_params as (preplay, _, _, _, _, _)) =
+    (one_line_params as ((_, preplay), _, _, _)) =
   (if isar_proofs = SOME true orelse
       (isar_proofs = NONE andalso isar_proof_would_be_a_good_idea smt_proofs preplay) then
      isar_proof_text ctxt debug isar_proofs smt_proofs isar_params
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -53,10 +53,10 @@
            updates as Prove (qs', xs', l', t', subproofs', facts', meths', comment') :: updates') =
         (if l = l' then
            update_subproofs subproofs' updates'
-           |>> (fn subproofs' => Prove (qs', xs', l', t', subproofs', facts', meths', comment'))
+           |>> (fn subproofs'' => Prove (qs', xs', l', t', subproofs'', facts', meths', comment'))
          else
            update_subproofs subproofs updates
-           |>> (fn subproofs => Prove (qs, xs, l, t, subproofs, facts, meths, comment)))
+           |>> (fn subproofs' => Prove (qs, xs, l, t, subproofs', facts, meths, comment)))
         |>> (fn step => step :: steps)
       | update_step step (steps, updates) = (step :: steps, updates)
     and update_subproofs [] updates = ([], updates)
@@ -64,9 +64,9 @@
       | update_subproofs (proof :: subproofs) updates =
         update_proof proof (update_subproofs subproofs updates)
     and update_proof proof (proofs, []) = (proof :: proofs, [])
-      | update_proof (Proof (fix, assms, steps)) (proofs, updates) =
-        let val (steps, updates) = update_steps steps updates in
-          (Proof (fix, assms, steps) :: proofs, updates)
+      | update_proof (Proof (xs, assms, steps)) (proofs, updates) =
+        let val (steps', updates') = update_steps steps updates in
+          (Proof (xs, assms, steps') :: proofs, updates')
         end
   in
     fst (update_steps steps (rev updates))
@@ -87,18 +87,19 @@
     (hopeful @ hopeless, hopeless)
   end
 
-fun merge_steps preplay_data (Prove ([], fix1, l1, _, subproofs1, (lfs1, gfs1), meths1, comment1))
-      (Prove (qs2, fix2, l2, t, subproofs2, (lfs2, gfs2), meths2, comment2)) =
+fun merge_steps preplay_data (Prove ([], xs1, l1, _, subproofs1, (lfs1, gfs1), meths1, comment1))
+      (Prove (qs2, xs2, l2, t, subproofs2, (lfs2, gfs2), meths2, comment2)) =
   let
     val (meths, hopeless) = merge_methods preplay_data (l1, meths1) (l2, meths2)
     val lfs = union (op =) lfs1 (remove (op =) l1 lfs2)
     val gfs = union (op =) gfs1 gfs2
   in
-    (Prove (qs2, if member (op =) qs2 Show then [] else union (op =) fix1 fix2, l2, t,
-       subproofs1 @ subproofs2, (lfs, gfs), meths, comment1 ^ comment2), hopeless)
+    (Prove (qs2, inter (op =) (Term.add_frees t []) (xs1 @ xs2), l2, t,
+       subproofs1 @ subproofs2, sort_facts (lfs, gfs), meths, comment1 ^ comment2),
+     hopeless)
   end
 
-val merge_slack_time = seconds 0.005
+val merge_slack_time = seconds 0.01
 val merge_slack_factor = 1.5
 
 fun adjust_merge_timeout max time =
@@ -117,7 +118,7 @@
       val (compress_further, decrement_step_count) =
         let
           val number_of_steps = add_isar_steps (steps_of_isar_proof proof) 0
-          val target_number_of_steps = Real.round (Real.fromInt number_of_steps / compress)
+          val target_number_of_steps = Real.ceil (Real.fromInt number_of_steps / compress)
           val delta = Unsynchronized.ref (number_of_steps - target_number_of_steps)
         in
           (fn () => !delta > 0, fn () => delta := !delta - 1)
@@ -153,10 +154,10 @@
         | _ => preplay_timeout)
 
       (* elimination of trivial, one-step subproofs *)
-      fun elim_one_subproof time (step as Prove (qs, fix, l, t, _, (lfs, gfs), meths, comment)) subs
+      fun elim_one_subproof time (step as Prove (qs, xs, l, t, _, (lfs, gfs), meths, comment)) subs
           nontriv_subs =
         if null subs orelse not (compress_further ()) then
-          Prove (qs, fix, l, t, List.revAppend (nontriv_subs, subs), (lfs, gfs), meths, comment)
+          Prove (qs, xs, l, t, List.revAppend (nontriv_subs, subs), (lfs, gfs), meths, comment)
         else
           (case subs of
             (sub as Proof (_, assms, [Prove (_, _, l', _, [], (lfs', gfs'), meths', _)])) :: subs =>
@@ -167,14 +168,14 @@
               val gfs'' = union (op =) gfs' gfs
               val (meths'' as _ :: _, hopeless) =
                 merge_methods (!preplay_data) (l', meths') (l, meths)
-              val step'' = Prove (qs, fix, l, t, subs'', (lfs'', gfs''), meths'', comment)
+              val step'' = Prove (qs, xs, l, t, subs'', (lfs'', gfs''), meths'', comment)
 
               (* check if the modified step can be preplayed fast enough *)
               val timeout = adjust_merge_timeout preplay_timeout (Time.+ (time, reference_time l'))
             in
               (case preplay_isar_step ctxt timeout hopeless step'' of
                 meths_outcomes as (_, Played time'') :: _ =>
-                (* l' successfully eliminated *)
+                (* "l'" successfully eliminated *)
                 (decrement_step_count ();
                  set_preplay_outcomes_of_isar_step ctxt time'' preplay_data step'' meths_outcomes;
                  map (replace_successor l' [l]) lfs';
@@ -192,8 +193,9 @@
 
       fun compress_top_level steps =
         let
-          fun cand_key (l, t_size) = (length (get_successors l), t_size)
-          val cand_ord = prod_ord int_ord (int_ord o swap) o pairself cand_key
+          val cand_key = apfst (length o get_successors)
+          val cand_ord =
+            prod_ord int_ord (prod_ord (int_ord o swap) (int_ord o swap)) o pairself cand_key
 
           fun pop_next_candidate [] = (NONE, [])
             | pop_next_candidate (cands as (cand :: cands')) =
@@ -215,13 +217,12 @@
                   val time_slice = time_mult (1.0 / Real.fromInt (length labels)) (reference_time l)
                   val timeouts =
                     map (adjust_merge_timeout preplay_timeout o curry Time.+ time_slice) times0
-                  val meths_outcomess =
-                    map3 (preplay_isar_step ctxt) timeouts hopelesses succs'
+                  val meths_outcomess = map3 (preplay_isar_step ctxt) timeouts hopelesses succs'
                 in
                   (case try (map (fn (_, Played time) :: _ => time)) meths_outcomess of
                     NONE => steps
                   | SOME times =>
-                    (* candidate successfully eliminated *)
+                    (* "l" successfully eliminated *)
                     (decrement_step_count ();
                      map3 (fn time => set_preplay_outcomes_of_isar_step ctxt time preplay_data)
                        times succs' meths_outcomess;
@@ -236,12 +237,17 @@
             else
               (case pop_next_candidate candidates of
                 (NONE, _) => steps (* no more candidates for elimination *)
-              | (SOME (l, _), candidates') =>
+              | (SOME (l, (num_xs, _)), candidates') =>
                 (case find_index (curry (op =) (SOME l) o label_of_isar_step) steps of
                   ~1 => steps
                 | i =>
-                  let val successors = get_successors l in
-                    if length successors > compress_degree then
+                  let
+                    val successors = get_successors l
+                    val num_successors = length successors
+                  in
+                    (* Careful with "obtain", so we don't "obtain" twice the same variable after a
+                       merge. *)
+                    if num_successors > (if num_xs > 0 then 1 else compress_degree) then
                       steps
                     else
                       steps
@@ -249,7 +255,7 @@
                       |> compression_loop candidates'
                   end))
 
-          fun add_cand (Prove (_, _, l, t, _, _, _, _)) = cons (l, size_of_term t)
+          fun add_cand (Prove (_, xs, l, t, _, _, _, _)) = cons (l, (length xs, size_of_term t))
             | add_cand _ = I
 
           (* the very last step is not a candidate *)
@@ -264,12 +270,13 @@
          can be eliminated. In the best case, this once again leads to a proof whose proof steps do
          not have subproofs. Applying this approach recursively will result in a flat proof in the
          best cast. *)
-      fun compress_proof (proof as (Proof (fix, assms, steps))) =
-        if compress_further () then Proof (fix, assms, compress_steps steps) else proof
+      fun compress_proof (proof as (Proof (xs, assms, steps))) =
+        if compress_further () then Proof (xs, assms, compress_steps steps) else proof
       and compress_steps steps =
         (* bottom-up: compress innermost proofs first *)
-        steps |> map (fn step => step |> compress_further () ? compress_sub_levels)
-              |> compress_further () ? compress_top_level
+        steps
+        |> map (fn step => step |> compress_further () ? compress_sub_levels)
+        |> compress_further () ? compress_top_level
       and compress_sub_levels (Prove (qs, xs, l, t, subproofs, facts, meths, comment)) =
           (* compress subproofs *)
           Prove (qs, xs, l, t, map compress_proof subproofs, facts, meths, comment)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_minimize.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_minimize.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -12,6 +12,7 @@
   type isar_preplay_data = Sledgehammer_Isar_Preplay.isar_preplay_data
 
   val keep_fastest_method_of_isar_step : isar_preplay_data -> isar_step -> isar_step
+  val minimized_isar_step : Proof.context -> Time.time -> isar_step -> Time.time * isar_step
   val minimize_isar_step_dependencies : Proof.context -> isar_preplay_data Unsynchronized.ref ->
     isar_step -> isar_step
   val postprocess_isar_proof_remove_unreferenced_steps : (isar_step -> isar_step) -> isar_proof ->
@@ -35,28 +36,33 @@
 
 val slack = seconds 0.025
 
+fun minimized_isar_step ctxt time
+    (Prove (qs, xs, l, t, subproofs, (lfs0, gfs00), meths as meth :: _, comment)) =
+  let
+    val gfs0 = filter (thms_influence_proof_method ctxt meth o thms_of_name ctxt) gfs00
+
+    fun mk_step_lfs_gfs lfs gfs =
+      Prove (qs, xs, l, t, subproofs, sort_facts (lfs, gfs), meths, comment)
+
+    fun minimize_half _ min_facts [] time = (min_facts, time)
+      | minimize_half mk_step min_facts (fact :: facts) time =
+        (case preplay_isar_step_for_method ctxt (Time.+ (time, slack)) meth
+            (mk_step (min_facts @ facts)) of
+          Played time' => minimize_half mk_step min_facts facts time'
+        | _ => minimize_half mk_step (fact :: min_facts) facts time)
+
+    val (min_lfs, time') = minimize_half (fn lfs => mk_step_lfs_gfs lfs gfs0) [] lfs0 time
+    val (min_gfs, time'') = minimize_half (mk_step_lfs_gfs min_lfs) [] gfs0 time'
+  in
+    (time'', mk_step_lfs_gfs min_lfs min_gfs)
+  end
+
 fun minimize_isar_step_dependencies ctxt preplay_data
-      (step as Prove (qs, xs, l, t, subproofs, (lfs0, gfs00), meths as meth :: _, comment)) =
+      (step as Prove (_, _, l, _, _, _, meth :: _, _)) =
     (case Lazy.force (preplay_outcome_of_isar_step_for_method (!preplay_data) l meth) of
       Played time =>
-      let
-        val gfs0 = filter (influence_proof_method ctxt meth o thms_of_name ctxt) gfs00
-
-        fun mk_step_lfs_gfs lfs gfs = Prove (qs, xs, l, t, subproofs, (lfs, gfs), meths, comment)
-
-        fun minimize_facts _ min_facts [] time = (min_facts, time)
-          | minimize_facts mk_step min_facts (fact :: facts) time =
-            (case preplay_isar_step_for_method ctxt (Time.+ (time, slack)) meth
-                (mk_step (min_facts @ facts)) of
-              Played time' => minimize_facts mk_step min_facts facts time'
-            | _ => minimize_facts mk_step (fact :: min_facts) facts time)
-
-        val (min_lfs, time') = minimize_facts (fn lfs => mk_step_lfs_gfs lfs gfs0) [] lfs0 time
-        val (min_gfs, time'') = minimize_facts (mk_step_lfs_gfs min_lfs) [] gfs0 time'
-
-        val step' = mk_step_lfs_gfs min_lfs min_gfs
-      in
-        set_preplay_outcomes_of_isar_step ctxt time'' preplay_data step' [(meth, Played time'')];
+      let val (time', step') = minimized_isar_step ctxt time step in
+        set_preplay_outcomes_of_isar_step ctxt time' preplay_data step' [(meth, Played time')];
         step'
       end
     | _ => step (* don't touch steps that time out or fail *))
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -41,6 +41,27 @@
 
 val trace = Attrib.setup_config_bool @{binding sledgehammer_preplay_trace} (K false)
 
+fun par_list_map_filter_with_timeout _ _ _ _ [] = []
+  | par_list_map_filter_with_timeout get_time min_timeout timeout0 f xs =
+    let
+      val next_timeout = Unsynchronized.ref timeout0
+
+      fun apply_f x =
+        let val timeout = !next_timeout in
+          if Time.<= (timeout, min_timeout) then
+            NONE
+          else
+            let val y = f timeout x in
+              (case get_time y of
+                SOME time => next_timeout := time
+              | _ => ());
+              SOME y
+            end
+        end
+    in
+      map_filter I (Par_List.map apply_f xs)
+    end
+
 fun enrich_context_with_local_facts proof ctxt =
   let
     val thy = Proof_Context.theory_of ctxt
@@ -140,16 +161,18 @@
 fun preplay_isar_step_for_method ctxt timeout meth =
   try (raw_preplay_step_for_method ctxt timeout meth) #> the_default Play_Failed
 
-fun preplay_isar_step ctxt timeout hopeless step =
+val min_preplay_timeout = seconds 0.002
+
+fun preplay_isar_step ctxt timeout0 hopeless step =
   let
-    fun try_method meth =
-      (case preplay_isar_step_for_method ctxt timeout meth step of
-        outcome as Played _ => SOME (meth, outcome)
-      | _ => NONE)
+    fun preplay timeout meth = (meth, preplay_isar_step_for_method ctxt timeout meth step)
+    fun get_time (_, Played time) = SOME time
+      | get_time _ = NONE
 
     val meths = proof_methods_of_isar_step step |> subtract (op =) hopeless
   in
-    the_list (Par_List.get_some try_method meths)
+    par_list_map_filter_with_timeout get_time min_preplay_timeout timeout0 preplay meths
+    |> sort (play_outcome_ord o pairself snd)
   end
 
 type isar_preplay_data = (proof_method * play_outcome Lazy.lazy) list Canonical_Label_Tab.table
@@ -181,7 +204,7 @@
   if Lazy.is_finished outcome then Lazy.force outcome else Play_Timed_Out Time.zeroTime
 
 fun get_best_method_outcome force =
-  tap (List.app (K () o Lazy.future Future.default_params o snd)) (* optional parallelism *)
+  tap (List.app (K () o Lazy.future Future.default_params o snd)) (* could be parallelized *)
   #> map (apsnd force)
   #> sort (play_outcome_ord o pairself snd)
   #> hd
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -26,9 +26,9 @@
 
   val label_ord : label * label -> order
   val string_of_label : label -> string
+  val sort_facts : facts -> facts
 
   val steps_of_isar_proof : isar_proof -> isar_step list
-
   val label_of_isar_step : isar_step -> label option
   val facts_of_isar_step : isar_step -> facts
   val proof_methods_of_isar_step : isar_step -> proof_method list
@@ -46,6 +46,7 @@
   val kill_useless_labels_in_isar_proof : isar_proof -> isar_proof
   val relabel_isar_proof_canonically : isar_proof -> isar_proof
   val relabel_isar_proof_nicely : isar_proof -> isar_proof
+  val rationalize_obtains_in_isar_proofs : Proof.context -> isar_proof -> isar_proof
 
   val string_of_isar_proof : Proof.context -> int -> int -> isar_proof -> string
 end;
@@ -75,10 +76,24 @@
 
 val no_label = ("", ~1)
 
-val label_ord = pairself swap #> prod_ord int_ord fast_string_ord
+(* cf. "label_ord" below *)
+val assume_prefix = "a"
+val have_prefix = "f"
+
+fun label_ord ((s1, i1), (s2, i2)) =
+  (case int_ord (pairself String.size (s1, s2)) of
+    EQUAL =>
+    (case string_ord (s1, s2) of
+      EQUAL => int_ord (i1, i2)
+    | ord => ord (* "assume" before "have" *))
+  | ord => ord)
 
 fun string_of_label (s, num) = s ^ string_of_int num
 
+(* Put the nearest local label first, since it's the most likely to be replaced by a "hence".
+   (Some preplaying proof methods, e.g. "blast", react poorly to fact reorderings.) *)
+fun sort_facts (lfs, gfs) = (sort (label_ord o swap) lfs, sort string_ord gfs)
+
 fun steps_of_isar_proof (Proof (_, _, steps)) = steps
 
 fun label_of_isar_step (Prove (_, _, l, _, _, _, _, _)) = SOME l
@@ -174,9 +189,6 @@
     fst (relabel_proof proof (0, []))
   end
 
-val assume_prefix = "a"
-val have_prefix = "f"
-
 val relabel_isar_proof_nicely =
   let
     fun next_label depth prefix l (accum as (next, subst)) =
@@ -206,6 +218,35 @@
     relabel_proof [] 0
   end
 
+fun stutter_single_letter s = String.extract (s, 0, SOME 1) ^ s
+
+fun rationalize_obtains_in_isar_proofs ctxt =
+  let
+    fun rename_obtains xs (subst, ctxt) =
+      let
+        val Ts = map snd xs
+        val new_names0 = map (stutter_single_letter o var_name_of_typ o body_type) Ts
+        val (new_names, ctxt') = Variable.variant_fixes new_names0 ctxt
+        val ys = map2 pair new_names Ts
+      in
+        (ys, ((map Free xs ~~ map Free ys) @ subst, ctxt'))
+      end
+
+    fun rationalize_step (Prove (qs, xs, l, t, subs, facts, meths, comment)) subst_ctxt =
+        let
+          val (xs', subst_ctxt' as (subst', _)) = rename_obtains xs subst_ctxt
+          val t' = subst_atomic subst' t
+          val subs' = map (rationalize_proof subst_ctxt') subs
+        in
+          (Prove (qs, xs', l, t', subs', facts, meths, comment), subst_ctxt')
+        end
+    and rationalize_proof (subst_ctxt as (subst, _)) (Proof (fix, assms, steps)) =
+      Proof (fix, map (apsnd (subst_atomic subst)) assms,
+        fst (fold_map rationalize_step steps subst_ctxt))
+  in
+    rationalize_proof ([], ctxt)
+  end
+
 val indent_size = 2
 
 fun string_of_isar_proof ctxt0 i n proof =
@@ -290,18 +331,17 @@
     and add_step_pre ind qs subs (s, ctxt) =
       (s ^ of_subproofs ind ctxt qs subs ^ of_indent ind, ctxt)
     and add_step ind (Let (t1, t2)) =
-        add_str (of_indent ind ^ "let ")
-        #> add_term t1 #> add_str " = " #> add_term t2 #> add_str "\n"
+        add_str (of_indent ind ^ "let ") #> add_term t1 #> add_str " = " #> add_term t2
+        #> add_str "\n"
       | add_step ind (Prove (qs, xs, l, t, subs, (ls, ss), meth :: _, comment)) =
         add_step_pre ind qs subs
         #> (case xs of
-             [] => add_str (of_have qs (length subs) ^ " ")
-           | _ => add_str (of_obtain qs (length subs) ^ " ") #> add_frees xs #> add_str " where ")
+            [] => add_str (of_have qs (length subs) ^ " ")
+          | _ => add_str (of_obtain qs (length subs) ^ " ") #> add_frees xs #> add_str " where ")
         #> add_str (of_label l)
         #> add_term t
-        #> add_str (" " ^
-             of_method (sort_distinct label_ord ls) (sort_distinct string_ord ss) meth ^
-             (if comment = "" then "" else " (* " ^ comment ^ " *)") ^ "\n")
+        #> add_str (" " ^ of_method ls ss meth ^
+          (if comment = "" then "" else " (* " ^ comment ^ " *)") ^ "\n")
     and add_steps ind = fold (add_step ind)
     and of_proof ind ctxt (Proof (xs, assms, steps)) =
       ("", ctxt)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -814,7 +814,7 @@
       {comment = "Goal: " ^ goal_name, state = Proof.init ctxt, goal = goal, subgoal = 1,
        subgoal_count = 1, factss = [("", facts)]}
   in
-    get_minimizing_prover ctxt MaSh (K ()) prover params (K (K (K ""))) problem
+    get_minimizing_prover ctxt MaSh (K ()) prover params problem
   end
 
 val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
@@ -974,8 +974,8 @@
   thms_in_proof max_dependencies (SOME name_tabs) th
   |> Option.map (fn deps =>
     if deps = [typedef_dep] orelse deps = [class_some_dep] orelse
-       exists (member (op =) fundef_ths) deps orelse exists (member (op =) typedef_ths) deps orelse
-       is_size_def deps th then
+        exists (member (op =) fundef_ths) deps orelse exists (member (op =) typedef_ths) deps orelse
+        is_size_def deps th then
       []
     else
       deps)
@@ -1378,7 +1378,7 @@
             let
               val name = nickname_of_thm th
               val feats = features_of ctxt (theory_of_thm th) stature [prop_of th]
-              val deps = deps_of status th |> these
+              val deps = these (deps_of status th)
               val num_nontrivial = num_nontrivial |> not (null deps) ? Integer.add 1
               val learns = (name, parents, feats, deps) :: learns
               val (learns, next_commit) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -18,7 +18,7 @@
     Simp_Method |
     Simp_Size_Method |
     Auto_Method |
-    Fastforce_Method |
+    Auto_Choice_Method |
     Force_Method |
     Linarith_Method |
     Presburger_Method |
@@ -29,14 +29,13 @@
     Play_Timed_Out of Time.time |
     Play_Failed
 
-  type minimize_command = string list -> string
   type one_line_params =
-    (proof_method * play_outcome) * string * (string * stature) list * minimize_command * int * int
+    ((string * stature) list * (proof_method * play_outcome)) * string * int * int
 
   val is_proof_method_direct : proof_method -> bool
   val string_of_proof_method : Proof.context -> string list -> proof_method -> string
   val tac_of_proof_method : Proof.context -> thm list * thm list -> proof_method -> int -> tactic
-  val influence_proof_method : Proof.context -> proof_method -> thm list -> bool
+  val thms_influence_proof_method : Proof.context -> proof_method -> thm list -> bool
   val string_of_play_outcome : play_outcome -> string
   val play_outcome_ord : play_outcome * play_outcome -> order
   val one_line_proof_text : Proof.context -> int -> one_line_params -> string
@@ -58,7 +57,7 @@
   Simp_Method |
   Simp_Size_Method |
   Auto_Method |
-  Fastforce_Method |
+  Auto_Choice_Method |
   Force_Method |
   Linarith_Method |
   Presburger_Method |
@@ -69,9 +68,8 @@
   Play_Timed_Out of Time.time |
   Play_Failed
 
-type minimize_command = string list -> string
 type one_line_params =
-  (proof_method * play_outcome) * string * (string * stature) list * minimize_command * int * int
+  ((string * stature) list * (proof_method * play_outcome)) * string * int * int
 
 fun is_proof_method_direct (Metis_Method _) = true
   | is_proof_method_direct Meson_Method = true
@@ -96,7 +94,7 @@
       | Simp_Method => if null ss then "simp" else "simp add:"
       | Simp_Size_Method => "simp add: " ^ short_thm_name ctxt @{thm size_ne_size_imp_ne}
       | Auto_Method => "auto"
-      | Fastforce_Method => "fastforce"
+      | Auto_Choice_Method => "atomize_elim, auto intro!: " ^ short_thm_name ctxt @{thm choice}
       | Force_Method => "force"
       | Linarith_Method => "linarith"
       | Presburger_Method => "presburger"
@@ -129,18 +127,20 @@
     (case meth of
       SATx_Method => SAT.satx_tac ctxt
     | Blast_Method => blast_tac ctxt
-    | Auto_Method => K (Clasimp.auto_tac (silence_methods ctxt))
-    | Fastforce_Method => Clasimp.fast_force_tac (silence_methods ctxt)
-    | Force_Method => Clasimp.force_tac (silence_methods ctxt)
+    | Auto_Method => SELECT_GOAL (Clasimp.auto_tac (silence_methods ctxt))
+    | Auto_Choice_Method =>
+      Atomize_Elim.atomize_elim_tac ctxt THEN'
+      SELECT_GOAL (Clasimp.auto_tac (silence_methods ctxt addSIs @{thms choice}))
+    | Force_Method => SELECT_GOAL (Clasimp.auto_tac (silence_methods ctxt))
     | Linarith_Method =>
       let val ctxt = Config.put Lin_Arith.verbose false ctxt in Lin_Arith.tac ctxt end
     | Presburger_Method => Cooper.tac true [] [] ctxt
     | Algebra_Method => Groebner.algebra_tac [] [] ctxt))
 
 val simp_based_methods =
-  [Simp_Method, Simp_Size_Method, Auto_Method, Fastforce_Method, Force_Method]
+  [Simp_Method, Simp_Size_Method, Auto_Method, Auto_Choice_Method, Force_Method]
 
-fun influence_proof_method ctxt meth ths =
+fun thms_influence_proof_method ctxt meth ths =
   not (member (op =) simp_based_methods meth) orelse
   let val ctxt' = silence_methods ctxt in
     (* unfortunate pointer comparison -- but it's always safe to consider a theorem useful *)
@@ -179,29 +179,13 @@
     (s |> s <> "" ? enclose " (" ")") ^ "."
   end
 
-fun minimize_line _ [] = ""
-  | minimize_line minimize_command ss =
-    (case minimize_command ss of
-      "" => ""
-    | command => "\nTo minimize: " ^ Active.sendback_markup [Markup.padding_command] command ^ ".")
-
-fun split_used_facts facts =
-  facts
-  |> List.partition (fn (_, (sc, _)) => sc = Chained)
-  |> pairself (sort_distinct (string_ord o pairself fst))
-
 fun one_line_proof_text ctxt num_chained
-    ((meth, play), banner, used_facts, minimize_command, subgoal, subgoal_count) =
-  let
-    val (chained, extra) = split_used_facts used_facts
-
-    val try_line =
-      map fst extra
-      |> proof_method_command ctxt meth subgoal subgoal_count (map fst chained) num_chained
-      |> (if play = Play_Failed then enclose "One-line proof reconstruction failed: " "."
-          else try_command_line banner play)
-  in
-    try_line ^ minimize_line minimize_command (map fst (extra @ chained))
+    ((used_facts, (meth, play)), banner, subgoal, subgoal_count) =
+  let val (chained, extra) = List.partition (fn (_, (sc, _)) => sc = Chained) used_facts in
+    map fst extra
+    |> proof_method_command ctxt meth subgoal subgoal_count (map fst chained) num_chained
+    |> (if play = Play_Failed then enclose "One-line proof reconstruction failed: " "."
+        else try_command_line banner play)
   end
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -14,7 +14,6 @@
   type fact = Sledgehammer_Fact.fact
   type proof_method = Sledgehammer_Proof_Methods.proof_method
   type play_outcome = Sledgehammer_Proof_Methods.play_outcome
-  type minimize_command = Sledgehammer_Proof_Methods.minimize_command
 
   datatype mode = Auto_Try | Try | Normal | MaSh | Auto_Minimize | Minimize
 
@@ -36,11 +35,11 @@
      max_mono_iters : int option,
      max_new_mono_instances : int option,
      isar_proofs : bool option,
-     compress : real,
+     compress : real option,
      try0 : bool,
      smt_proofs : bool option,
      slice : bool,
-     minimize : bool option,
+     minimize : bool,
      timeout : Time.time,
      preplay_timeout : Time.time,
      expect : string}
@@ -57,32 +56,21 @@
     {outcome : atp_failure option,
      used_facts : (string * stature) list,
      used_from : fact list,
+     preferred_methss : proof_method * proof_method list list,
      run_time : Time.time,
-     preplay : (proof_method * play_outcome) Lazy.lazy,
-     message : proof_method * play_outcome -> string,
-     message_tail : string}
+     message : (unit -> (string * stature) list * (proof_method * play_outcome)) -> string}
 
-  type prover =
-    params -> ((string * string list) list -> string -> minimize_command) -> prover_problem ->
-    prover_result
+  type prover = params -> prover_problem -> prover_result
 
   val SledgehammerN : string
   val str_of_mode : mode -> string
-  val smtN : string
   val overlord_file_location_of_prover : string -> string * string
   val proof_banner : mode -> string -> string
-  val extract_proof_method : params -> proof_method -> string * (string * string list) list
-  val is_proof_method : string -> bool
   val is_atp : theory -> string -> bool
-  val bunch_of_proof_methods : bool -> bool -> string -> proof_method list
+  val bunches_of_proof_methods : bool -> bool -> bool -> string -> proof_method list list
   val is_fact_chained : (('a * stature) * 'b) -> bool
   val filter_used_facts : bool -> (''a * stature) list -> ((''a * stature) * 'b) list ->
     ((''a * stature) * 'b) list
-  val play_one_line_proof : mode -> bool -> Time.time -> ((string * 'a) * thm) list ->
-    Proof.state -> int -> proof_method -> proof_method list -> proof_method * play_outcome
-  val isar_supported_prover_of : theory -> string -> string
-  val choose_minimize_command : theory -> params -> ((string * string list) list -> string -> 'a) ->
-    string -> proof_method * play_outcome -> 'a
   val repair_monomorph_context : int option -> int -> int option -> int -> Proof.context ->
     Proof.context
 
@@ -118,11 +106,6 @@
   | str_of_mode Auto_Minimize = "Auto_Minimize"
   | str_of_mode Minimize = "Minimize"
 
-val smtN = "smt"
-
-val proof_method_names = [metisN, smtN]
-val is_proof_method = member (op =) proof_method_names
-
 val is_atp = member (op =) o supported_atps
 
 type params =
@@ -143,11 +126,11 @@
    max_mono_iters : int option,
    max_new_mono_instances : int option,
    isar_proofs : bool option,
-   compress : real,
+   compress : real option,
    try0 : bool,
    smt_proofs : bool option,
    slice : bool,
-   minimize : bool option,
+   minimize : bool,
    timeout : Time.time,
    preplay_timeout : Time.time,
    expect : string}
@@ -164,14 +147,11 @@
   {outcome : atp_failure option,
    used_facts : (string * stature) list,
    used_from : fact list,
+   preferred_methss : proof_method * proof_method list list,
    run_time : Time.time,
-   preplay : (proof_method * play_outcome) Lazy.lazy,
-   message : proof_method * play_outcome -> string,
-   message_tail : string}
+   message : (unit -> (string * stature) list * (proof_method * play_outcome)) -> string}
 
-type prover =
-  params -> ((string * string list) list -> string -> minimize_command)
-  -> prover_problem -> prover_result
+type prover = params -> prover_problem -> prover_result
 
 fun overlord_file_location_of_prover prover = (getenv "ISABELLE_HOME_USER", "prob_" ^ prover)
 
@@ -181,101 +161,29 @@
   | Try => "Sledgehammer (" ^ quote name ^ ") found a proof"
   | _ => "Try this")
 
-fun bunch_of_proof_methods smt_proofs needs_full_types desperate_lam_trans =
-  (if needs_full_types then
-     [Metis_Method (SOME full_typesN, NONE),
-      Metis_Method (SOME really_full_type_enc, NONE),
-      Metis_Method (SOME full_typesN, SOME desperate_lam_trans),
-      Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)]
+fun bunches_of_proof_methods try0 smt_proofs needs_full_types desperate_lam_trans =
+  (if try0 then
+    [[Simp_Method, Auto_Method, Blast_Method, Linarith_Method],
+     [Meson_Method, Force_Method, Presburger_Method]]
    else
-     [Metis_Method (NONE, NONE),
-      Metis_Method (SOME full_typesN, NONE),
-      Metis_Method (SOME no_typesN, SOME desperate_lam_trans),
-      Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)]) @
-  (if smt_proofs then [SMT2_Method] else [])
-
-fun extract_proof_method ({type_enc, lam_trans, ...} : params)
-      (Metis_Method (type_enc', lam_trans')) =
-    let
-      val override_params =
-        (if is_none type_enc' andalso is_none type_enc then []
-         else [("type_enc", [hd (unalias_type_enc (type_enc' |> the_default partial_typesN))])]) @
-        (if is_none lam_trans' andalso is_none lam_trans then []
-         else [("lam_trans", [lam_trans' |> the_default default_metis_lam_trans])])
-    in (metisN, override_params) end
-  | extract_proof_method _ SMT2_Method = (smtN, [])
-
-(* based on "Mirabelle.can_apply" and generalized *)
-fun timed_apply timeout tac state i =
-  let
-    val {context = ctxt, facts, goal} = Proof.goal state
-    val full_tac = Method.insert_tac facts i THEN tac ctxt i
-  in
-    TimeLimit.timeLimit timeout (try (Seq.pull o full_tac)) goal
-  end
-
-fun timed_proof_method timeout ths meth =
-  timed_apply timeout (fn ctxt => tac_of_proof_method ctxt ([], ths) meth)
+     []) @
+  [[Metis_Method (if needs_full_types then SOME full_typesN else NONE, NONE)],
+   (if needs_full_types then
+      [Metis_Method (NONE, NONE),
+       Metis_Method (SOME really_full_type_enc, NONE),
+       Metis_Method (SOME full_typesN, SOME desperate_lam_trans),
+       Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)]
+    else
+      [Metis_Method (SOME full_typesN, NONE),
+       Metis_Method (SOME no_typesN, SOME desperate_lam_trans),
+       Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)])] @
+  (if smt_proofs then [[SMT2_Method]] else [])
 
 fun is_fact_chained ((_, (sc, _)), _) = sc = Chained
 
 fun filter_used_facts keep_chained used =
   filter ((member (op =) used o fst) orf (if keep_chained then is_fact_chained else K false))
 
-fun play_one_line_proof mode verbose timeout pairs state i preferred (meths as meth :: _) =
-  let
-    val ctxt = Proof.context_of state
-
-    fun get_preferred meths = if member (op =) meths preferred then preferred else meth
-  in
-    if timeout = Time.zeroTime then
-      (get_preferred meths, Play_Timed_Out Time.zeroTime)
-    else
-      let
-        val _ = if mode = Minimize then Output.urgent_message "Preplaying proof..." else ()
-        val ths = pairs |> sort_wrt (fst o fst) |> map snd
-        fun play [] [] = (get_preferred meths, Play_Failed)
-          | play timed_outs [] = (get_preferred timed_outs, Play_Timed_Out timeout)
-          | play timed_out (meth :: meths) =
-            let
-              val _ =
-                if verbose then
-                  Output.urgent_message ("Trying \"" ^ string_of_proof_method ctxt [] meth ^
-                    "\" for " ^ string_of_time timeout ^ "...")
-                else
-                  ()
-              val timer = Timer.startRealTimer ()
-            in
-              (case timed_proof_method timeout ths meth state i of
-                SOME (SOME _) => (meth, Played (Timer.checkRealTimer timer))
-              | _ => play timed_out meths)
-            end
-            handle TimeLimit.TimeOut => play (meth :: timed_out) meths
-      in
-        play [] meths
-      end
-  end
-
-val canonical_isar_supported_prover = eN
-val z3N = "z3"
-
-fun isar_supported_prover_of thy name =
-  if is_atp thy name orelse name = z3N then name
-  else if is_atp_installed thy canonical_isar_supported_prover then canonical_isar_supported_prover
-  else name
-
-(* FIXME: See the analogous logic in the function "maybe_minimize" in
-   "sledgehammer_prover_minimize.ML". *)
-fun choose_minimize_command thy (params as {isar_proofs, ...}) minimize_command name preplay =
-  let
-    val maybe_isar_name = name |> isar_proofs = SOME true ? isar_supported_prover_of thy
-    val (min_name, override_params) =
-      (case preplay of
-        (meth as Metis_Method _, Played _) =>
-        if isar_proofs = SOME true then (maybe_isar_name, []) else extract_proof_method params meth
-      | _ => (maybe_isar_name, []))
-  in minimize_command override_params min_name end
-
 val max_fact_instances = 10 (* FUDGE *)
 
 fun repair_monomorph_context max_iters best_max_iters max_new_instances best_max_new_instances =
@@ -288,7 +196,6 @@
   let
     val thy = Proof_Context.theory_of ctxt
     val (remote_provers, local_provers) =
-      proof_method_names @
       sort_strings (supported_atps thy) @
       sort_strings (SMT2_Config.available_solvers_of ctxt)
       |> List.partition (String.isPrefix remote_prefix)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_atp.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -30,6 +30,7 @@
 open ATP_Problem_Generate
 open ATP_Proof_Reconstruct
 open ATP_Waldmeister
+open ATP_Satallax
 open ATP_Systems
 open Sledgehammer_Util
 open Sledgehammer_Proof_Methods
@@ -126,15 +127,13 @@
    the only ATP that does not honor its time limit. *)
 val atp_timeout_slack = seconds 1.0
 
-(* Important messages are important but not so important that users want to see
-   them each time. *)
+(* Important messages are important but not so important that users want to see them each time. *)
 val atp_important_message_keep_quotient = 25
 
 fun run_atp mode name
-    (params as {debug, verbose, overlord, type_enc, strict, lam_trans, uncurried_aliases,
-       fact_filter, max_facts, max_mono_iters, max_new_mono_instances, isar_proofs, compress, try0,
-       smt_proofs, slice, minimize, timeout, preplay_timeout, ...})
-    minimize_command
+    ({debug, verbose, overlord, type_enc, strict, lam_trans, uncurried_aliases, fact_filter,
+     max_facts, max_mono_iters, max_new_mono_instances, isar_proofs, compress, try0, smt_proofs,
+     slice, minimize, timeout, preplay_timeout, ...} : params)
     ({comment, state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
   let
     val thy = Proof.theory_of state
@@ -295,19 +294,21 @@
               |>> (if overlord then prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") else I)
               |> fst |> split_time
               |> (fn accum as (output, _) =>
-                     (accum,
-                      extract_tstplike_proof_and_outcome verbose proof_delims known_failures output
-                      |>> atp_proof_of_tstplike_proof (perhaps (try (unprefix remote_prefix)) name) atp_problem
-                      handle UNRECOGNIZED_ATP_PROOF () => ([], SOME ProofUnparsable)))
+                (accum,
+                 extract_tstplike_proof_and_outcome verbose proof_delims known_failures output
+                 |>> atp_proof_of_tstplike_proof (perhaps (try (unprefix remote_prefix)) name)
+                   atp_problem
+                 handle UNRECOGNIZED_ATP_PROOF () => ([], SOME ProofUnparsable)))
               handle TimeLimit.TimeOut => (("", slice_timeout), ([], SOME TimedOut))
 
             val outcome =
               (case outcome of
                 NONE =>
-                (case used_facts_in_unsound_atp_proof ctxt (map fst facts) atp_proof
-                      |> Option.map (sort string_ord) of
+                (case used_facts_in_unsound_atp_proof ctxt (map fst facts) atp_proof of
                   SOME facts =>
-                  let val failure = UnsoundProof (is_type_enc_sound type_enc, facts) in
+                  let
+                    val failure = UnsoundProof (is_type_enc_sound type_enc, sort string_ord facts)
+                  in
                     if debug then (warning (string_of_atp_failure failure); NONE) else SOME failure
                   end
                 | NONE => NONE)
@@ -354,22 +355,18 @@
       else
         ""
 
-    val (used_facts, preplay, message, message_tail) =
+    val (used_facts, preferred_methss, message) =
       (case outcome of
         NONE =>
         let
-          val used_facts = used_facts_in_atp_proof ctxt (map fst used_from) atp_proof
+          val used_facts = sort_wrt fst (used_facts_in_atp_proof ctxt (map fst used_from) atp_proof)
           val needs_full_types = is_typed_helper_used_in_atp_proof atp_proof
-          val meths =
-            bunch_of_proof_methods (smt_proofs <> SOME false) needs_full_types
-              (if atp_proof_prefers_lifting atp_proof then liftingN else hide_lamsN)
+          val preferred_methss =
+            (Metis_Method (NONE, NONE),
+             bunches_of_proof_methods try0 (smt_proofs <> SOME false) needs_full_types
+               (if atp_proof_prefers_lifting atp_proof then liftingN else hide_lamsN))
         in
-          (used_facts,
-           Lazy.lazy (fn () =>
-             let val used_pairs = used_from |> filter_used_facts false used_facts in
-               play_one_line_proof mode verbose preplay_timeout used_pairs state subgoal (hd meths)
-                 meths
-             end),
+          (used_facts, preferred_methss,
            fn preplay =>
               let
                 val _ = if verbose then Output.urgent_message "Generating proof text..." else ()
@@ -383,33 +380,29 @@
                     val atp_proof =
                       atp_proof
                       |> termify_atp_proof ctxt name format type_enc pool lifted sym_tab
-                      |> spass ? introduce_spass_skolem
+                      |> spass ? introduce_spass_skolems
                       |> factify_atp_proof (map fst used_from) hyp_ts concl_t
                   in
                     (verbose, (metis_type_enc, metis_lam_trans), preplay_timeout, compress, try0,
-                     minimize <> SOME false, atp_proof, goal)
+                     minimize, atp_proof, goal)
                   end
 
-                val one_line_params =
-                  (preplay, proof_banner mode name, used_facts,
-                   choose_minimize_command thy params minimize_command name preplay, subgoal,
-                   subgoal_count)
+                val one_line_params = (preplay (), proof_banner mode name, subgoal, subgoal_count)
                 val num_chained = length (#facts (Proof.goal state))
               in
-                proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained one_line_params
-              end,
-           (if verbose then "\nATP real CPU time: " ^ string_of_time run_time ^ "." else "") ^
-           (if important_message <> "" then
-              "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message
-            else
-              ""))
+                proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained
+                  one_line_params ^
+                (if important_message <> "" then
+                   "\n\nImportant message from Dr. Geoff Sutcliffe:\n" ^ important_message
+                 else
+                   "")
+              end)
         end
       | SOME failure =>
-        ([], Lazy.value (Metis_Method (NONE, NONE), Play_Failed),
-         fn _ => string_of_atp_failure failure, ""))
+        ([], (Auto_Method (* dummy *), []), fn _ => string_of_atp_failure failure))
   in
-    {outcome = outcome, used_facts = used_facts, used_from = used_from, run_time = run_time,
-     preplay = preplay, message = message, message_tail = message_tail}
+    {outcome = outcome, used_facts = used_facts, used_from = used_from,
+     preferred_methss = preferred_methss, run_time = run_time, message = message}
   end
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_minimize.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_minimize.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -20,18 +20,11 @@
   val get_prover : Proof.context -> mode -> string -> prover
 
   val binary_min_facts : int Config.T
-  val auto_minimize_min_facts : int Config.T
-  val auto_minimize_max_time : real Config.T
   val minimize_facts : (thm list -> unit) -> string -> params -> bool -> int -> int ->
-    Proof.state -> thm -> (proof_method * play_outcome) Lazy.lazy option ->
-    ((string * stature) * thm list) list ->
+    Proof.state -> thm -> ((string * stature) * thm list) list ->
     ((string * stature) * thm list) list option
-      * ((proof_method * play_outcome) Lazy.lazy * ((proof_method * play_outcome) -> string)
-         * string)
+    * ((unit -> (string * stature) list * (proof_method * play_outcome)) -> string)
   val get_minimizing_prover : Proof.context -> mode -> (thm list -> unit) -> string -> prover
-
-  val run_minimize : params -> (thm list -> unit) -> int -> (Facts.ref * Attrib.src list) list ->
-    Proof.state -> unit
 end;
 
 structure Sledgehammer_Prover_Minimize : SLEDGEHAMMER_PROVER_MINIMIZE =
@@ -50,59 +43,17 @@
 open Sledgehammer_Prover_ATP
 open Sledgehammer_Prover_SMT2
 
-fun run_proof_method mode name (params as {verbose, timeout, type_enc, lam_trans, ...})
-    minimize_command
-    ({state, subgoal, subgoal_count, factss = (_, facts) :: _, ...} : prover_problem) =
-  let
-    val meth =
-      if name = metisN then Metis_Method (type_enc, lam_trans)
-      else if name = smtN then SMT2_Method
-      else raise Fail ("unknown proof_method: " ^ quote name)
-    val used_facts = facts |> map fst
-  in
-    (case play_one_line_proof (if mode = Minimize then Normal else mode) verbose timeout facts state
-        subgoal meth [meth] of
-      play as (_, Played time) =>
-      {outcome = NONE, used_facts = used_facts, used_from = facts, run_time = time,
-       preplay = Lazy.value play,
-       message = fn play =>
-          let
-            val ctxt = Proof.context_of state
-            val (_, override_params) = extract_proof_method params meth
-            val one_line_params =
-              (play, proof_banner mode name, used_facts, minimize_command override_params name,
-               subgoal, subgoal_count)
-            val num_chained = length (#facts (Proof.goal state))
-          in
-            one_line_proof_text ctxt num_chained one_line_params
-          end,
-       message_tail = ""}
-    | play =>
-      let
-        val failure = (case play of (_, Play_Failed) => GaveUp | _ => TimedOut)
-      in
-        {outcome = SOME failure, used_facts = [], used_from = [],
-         run_time = Time.zeroTime, preplay = Lazy.value play,
-         message = fn _ => string_of_atp_failure failure, message_tail = ""}
-      end)
-  end
-
 fun is_prover_supported ctxt =
   let val thy = Proof_Context.theory_of ctxt in
-    is_proof_method orf is_atp thy orf is_smt2_prover ctxt
+    is_atp thy orf is_smt2_prover ctxt
   end
 
 fun is_prover_installed ctxt =
-  is_proof_method orf is_smt2_prover ctxt orf
-  is_atp_installed (Proof_Context.theory_of ctxt)
-
-val proof_method_default_max_facts = 20
+  is_smt2_prover ctxt orf is_atp_installed (Proof_Context.theory_of ctxt)
 
 fun default_max_facts_of_prover ctxt name =
   let val thy = Proof_Context.theory_of ctxt in
-    if is_proof_method name then
-      proof_method_default_max_facts
-    else if is_atp thy name then
+    if is_atp thy name then
       fold (Integer.max o fst o #1 o fst o snd) (#best_slices (get_atp thy name ()) ctxt) 0
     else if is_smt2_prover ctxt name then
       SMT2_Solver.default_max_relevant ctxt name
@@ -112,8 +63,7 @@
 
 fun get_prover ctxt mode name =
   let val thy = Proof_Context.theory_of ctxt in
-    if is_proof_method name then run_proof_method mode name
-    else if is_atp thy name then run_atp mode name
+    if is_atp thy name then run_atp mode name
     else if is_smt2_prover ctxt name then run_smt2_solver mode name
     else error ("No such prover: " ^ name ^ ".")
   end
@@ -149,8 +99,7 @@
     val problem =
       {comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
        factss = [("", facts)]}
-    val result as {outcome, used_facts, run_time, ...} =
-      prover params (K (K (K ""))) problem
+    val result as {outcome, used_facts, run_time, ...} = prover params problem
   in
     print silent (fn () =>
       (case outcome of
@@ -180,11 +129,6 @@
    actually needed, we heuristically set the threshold to 10 facts. *)
 val binary_min_facts =
   Attrib.setup_config_int @{binding sledgehammer_minimize_binary_min_facts} (K 20)
-val auto_minimize_min_facts =
-  Attrib.setup_config_int @{binding sledgehammer_auto_minimize_min_facts}
-      (fn generic => Config.get_generic generic binary_min_facts)
-val auto_minimize_max_time =
-  Attrib.setup_config_real @{binding sledgehammer_auto_minimize_max_time} (K 5.0)
 
 fun linear_minimize test timeout result xs =
   let
@@ -216,7 +160,7 @@
           (case test timeout (sup @ l0) of
             result as {outcome = NONE, used_facts, ...} =>
             min depth result (filter_used_facts true used_facts sup)
-                      (filter_used_facts true used_facts l0)
+              (filter_used_facts true used_facts l0)
           | _ =>
             (case test timeout (sup @ r0) of
               result as {outcome = NONE, used_facts, ...} =>
@@ -244,7 +188,7 @@
   end
 
 fun minimize_facts do_learn prover_name (params as {learn, timeout, ...}) silent i n state goal
-    preplay0 facts =
+    facts =
   let
     val ctxt = Proof.context_of state
     val prover = get_prover ctxt (if silent then Auto_Minimize else Minimize) prover_name
@@ -262,135 +206,50 @@
          val min =
            if length facts >= Config.get ctxt binary_min_facts then binary_minimize
            else linear_minimize
-         val (min_facts, {preplay, message, message_tail, ...}) =
-           min test (new_timeout timeout run_time) result facts
+         val (min_facts, {message, ...}) = min test (new_timeout timeout run_time) result facts
        in
          print silent (fn () => cat_lines
-             ["Minimized to " ^ n_facts (map fst min_facts)] ^
-              (case min_facts |> filter is_fact_chained |> length of
-                 0 => ""
-               | n => "\n(including " ^ string_of_int n ^ " chained)") ^ ".");
+           ["Minimized to " ^ n_facts (map fst min_facts)] ^
+            (case min_facts |> filter is_fact_chained |> length of
+              0 => ""
+            | n => "\n(including " ^ string_of_int n ^ " chained)") ^ ".");
          (if learn then do_learn (maps snd min_facts) else ());
-         (SOME min_facts,
-          (if is_some preplay0 andalso length min_facts = length facts then the preplay0
-           else preplay,
-           message, message_tail))
+         (SOME min_facts, message)
        end
-     | {outcome = SOME TimedOut, preplay, ...} =>
-       (NONE, (preplay, fn _ =>
+     | {outcome = SOME TimedOut, ...} =>
+       (NONE, fn _ =>
           "Timeout: You can increase the time limit using the \"timeout\" option (e.g., \
-          \timeout = " ^ string_of_int (10 + Time.toMilliseconds timeout div 1000) ^ "\").", ""))
-     | {preplay, message, ...} => (NONE, (preplay, prefix "Prover error: " o message, ""))))
-    handle ERROR msg =>
-      (NONE, (Lazy.value (Metis_Method (NONE, NONE), Play_Failed), fn _ => "Error: " ^ msg, ""))
+          \timeout = " ^ string_of_int (10 + Time.toMilliseconds timeout div 1000) ^ "\").")
+     | {message, ...} => (NONE, (prefix "Prover error: " o message))))
+    handle ERROR msg => (NONE, fn _ => "Error: " ^ msg)
   end
 
-fun adjust_proof_method_params override_params
-    ({debug, verbose, overlord, spy, blocking, provers, type_enc, strict, lam_trans,
-      uncurried_aliases, learn, fact_filter, max_facts, fact_thresholds, max_mono_iters,
-      max_new_mono_instances, isar_proofs, compress, try0, smt_proofs, slice, minimize, timeout,
-      preplay_timeout, expect} : params) =
-  let
-    fun lookup_override name default_value =
-      (case AList.lookup (op =) override_params name of
-        SOME [s] => SOME s
-      | _ => default_value)
-    (* Only those options that proof_methods are interested in are considered here. *)
-    val type_enc = lookup_override "type_enc" type_enc
-    val lam_trans = lookup_override "lam_trans" lam_trans
-  in
-    {debug = debug, verbose = verbose, overlord = overlord, spy = spy, blocking = blocking,
-     provers = provers, type_enc = type_enc, strict = strict, lam_trans = lam_trans,
-     uncurried_aliases = uncurried_aliases, learn = learn, fact_filter = fact_filter,
-     max_facts = max_facts, fact_thresholds = fact_thresholds, max_mono_iters = max_mono_iters,
-     max_new_mono_instances = max_new_mono_instances, isar_proofs = isar_proofs,
-     compress = compress, try0 = try0, smt_proofs = smt_proofs, slice = slice, minimize = minimize,
-     timeout = timeout, preplay_timeout = preplay_timeout, expect = expect}
-  end
-
-fun maybe_minimize ctxt mode do_learn name (params as {verbose, isar_proofs, minimize, ...})
+fun maybe_minimize mode do_learn name (params as {verbose, minimize, ...})
     ({state, goal, subgoal, subgoal_count, ...} : prover_problem)
-    (result as {outcome, used_facts, used_from, run_time, preplay, message, message_tail} :
-     prover_result) =
+    (result as {outcome, used_facts, used_from, preferred_methss, run_time, message}
+     : prover_result) =
   if is_some outcome orelse null used_facts then
     result
   else
     let
-      val thy = Proof_Context.theory_of ctxt
-      val num_facts = length used_facts
-
-      val ((perhaps_minimize, (minimize_name, params)), preplay) =
-        if mode = Normal then
-          if num_facts >= Config.get ctxt auto_minimize_min_facts then
-            ((true, (name, params)), preplay)
-          else
-            let
-              fun can_min_fast_enough time =
-                0.001
-                * Real.fromInt ((num_facts + 1) * Time.toMilliseconds time)
-                <= Config.get ctxt auto_minimize_max_time
-              fun prover_fast_enough () = can_min_fast_enough run_time
-            in
-              (case Lazy.force preplay of
-                 (meth as Metis_Method _, Played timeout) =>
-                 if isar_proofs = SOME true then
-                   (* Cheat: Assume the selected ATP is as fast as "metis" for the goal it proved
-                      itself. *)
-                   (can_min_fast_enough timeout, (isar_supported_prover_of thy name, params))
-                 else if can_min_fast_enough timeout then
-                   (true, extract_proof_method params meth
-                          ||> (fn override_params =>
-                                  adjust_proof_method_params override_params params))
-                 else
-                   (prover_fast_enough (), (name, params))
-               | (SMT2_Method, Played timeout) =>
-                 (* Cheat: Assume the original prover is as fast as "smt" for the goal it proved
-                    itself. *)
-                 (can_min_fast_enough timeout, (name, params))
-               | _ => (prover_fast_enough (), (name, params)),
-               preplay)
-            end
-        else
-          ((false, (name, params)), preplay)
-      val minimize = minimize |> the_default perhaps_minimize
-      val (used_facts, (preplay, message, _)) =
+      val (used_facts, message) =
         if minimize then
-          minimize_facts do_learn minimize_name params
+          minimize_facts do_learn name params
             (not verbose orelse (mode <> Normal andalso mode <> MaSh)) subgoal subgoal_count state
-            goal (SOME preplay) (filter_used_facts true used_facts (map (apsnd single) used_from))
+            goal (filter_used_facts true used_facts (map (apsnd single) used_from))
           |>> Option.map (map fst)
         else
-          (SOME used_facts, (preplay, message, ""))
+          (SOME used_facts, message)
     in
       (case used_facts of
         SOME used_facts =>
-        {outcome = NONE, used_facts = used_facts, used_from = used_from, run_time = run_time,
-         preplay = preplay, message = message, message_tail = message_tail}
+        {outcome = NONE, used_facts = sort_wrt fst used_facts, used_from = used_from,
+         preferred_methss = preferred_methss, run_time = run_time, message = message}
       | NONE => result)
     end
 
-fun get_minimizing_prover ctxt mode do_learn name params minimize_command problem =
-  get_prover ctxt mode name params minimize_command problem
-  |> maybe_minimize ctxt mode do_learn name params problem
-
-fun run_minimize (params as {provers, ...}) do_learn i refs state =
-  let
-    val ctxt = Proof.context_of state
-    val {goal, facts = chained_ths, ...} = Proof.goal state
-    val reserved = reserved_isar_keyword_table ()
-    val css = clasimpset_rule_table_of ctxt
-    val facts = refs |> maps (map (apsnd single) o fact_of_ref ctxt reserved chained_ths css)
-  in
-    (case subgoal_count state of
-      0 => Output.urgent_message "No subgoal!"
-    | n =>
-      (case provers of
-        [] => error "No prover is set."
-      | prover :: _ =>
-        (kill_provers ();
-         minimize_facts do_learn prover params false i n state goal NONE facts
-         |> (fn (_, (preplay, message, message_tail)) =>
-                Output.urgent_message (message (Lazy.force preplay) ^ message_tail)))))
-  end
+fun get_minimizing_prover ctxt mode do_learn name params problem =
+  get_prover ctxt mode name params problem
+  |> maybe_minimize mode do_learn name params problem
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt2.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt2.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -179,7 +179,7 @@
 
 fun run_smt2_solver mode name (params as {debug, verbose, isar_proofs, compress, try0, smt_proofs,
       minimize, preplay_timeout, ...})
-    minimize_command ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
+    ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
   let
     val thy = Proof.theory_of state
     val ctxt = Proof.context_of state
@@ -189,36 +189,36 @@
     val {outcome, filter_result = {fact_ids, atp_proof, ...}, used_from, run_time} =
       smt2_filter_loop name params state goal subgoal factss
     val used_named_facts = map snd fact_ids
-    val used_facts = map fst used_named_facts
+    val used_facts = sort_wrt fst (map fst used_named_facts)
     val outcome = Option.map failure_of_smt2_failure outcome
 
-    val (preplay, message, message_tail) =
+    val (preferred_methss, message) =
       (case outcome of
         NONE =>
-        (Lazy.lazy (fn () =>
-           play_one_line_proof mode verbose preplay_timeout used_named_facts state subgoal
-             SMT2_Method (bunch_of_proof_methods (smt_proofs <> SOME false) false liftingN)),
-         fn preplay =>
-            let
-              fun isar_params () =
-                (verbose, (NONE, NONE), preplay_timeout, compress, try0, minimize <> SOME false,
-                 atp_proof (), goal)
+        let
+          val preferred_methss =
+            (SMT2_Method, bunches_of_proof_methods try0 (smt_proofs <> SOME false) false liftingN)
+        in
+          (preferred_methss,
+           fn preplay =>
+             let
+               val _ = if verbose then Output.urgent_message "Generating proof text..." else ()
 
-              val one_line_params =
-                (preplay, proof_banner mode name, used_facts,
-                 choose_minimize_command thy params minimize_command name preplay, subgoal,
-                 subgoal_count)
-              val num_chained = length (#facts (Proof.goal state))
-            in
-              proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained one_line_params
-            end,
-         if verbose then "\nSMT solver real CPU time: " ^ string_of_time run_time ^ "." else "")
-      | SOME failure =>
-        (Lazy.value (Metis_Method (NONE, NONE), Play_Failed),
-         fn _ => string_of_atp_failure failure, ""))
+               fun isar_params () =
+                 (verbose, (NONE, NONE), preplay_timeout, compress, try0, minimize, atp_proof (),
+                  goal)
+
+               val one_line_params = (preplay (), proof_banner mode name, subgoal, subgoal_count)
+               val num_chained = length (#facts (Proof.goal state))
+             in
+               proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained
+                 one_line_params
+             end)
+        end
+      | SOME failure => ((Auto_Method (* dummy *), []), fn _ => string_of_atp_failure failure))
   in
-    {outcome = outcome, used_facts = used_facts, used_from = used_from, run_time = run_time,
-     preplay = preplay, message = message, message_tail = message_tail}
+    {outcome = outcome, used_facts = used_facts, used_from = used_from,
+     preferred_methss = preferred_methss, run_time = run_time, message = message}
   end
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_util.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -101,7 +101,10 @@
           else (false, false)
       in
         if anonymous then
-          app_body (if enter_class then map_inclass_name else map_name) (Future.join body) accum
+          (case Future.peek body of
+            SOME (Exn.Res the_body) =>
+            app_body (if enter_class then map_inclass_name else map_name) the_body accum
+          | NONE => accum)
         else
           (case map_name name of
             SOME name' =>
@@ -133,7 +136,7 @@
     |> Symbol.source
     |> Token.source {do_recover = SOME false} lex Position.start
     |> Token.source_proper
-    |> Source.source Token.stopper (Parse_Spec.xthms1 >> get) NONE
+    |> Source.source Token.stopper (Parse.xthms1 >> get) NONE
     |> Source.exhaust
   end
 
@@ -144,9 +147,9 @@
   Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) (print_mode_value ())) f x
 
 fun hackish_string_of_term ctxt =
-  with_vanilla_print_mode (Syntax.string_of_term ctxt) #> simplify_spaces
+  with_vanilla_print_mode (Syntax.string_of_term ctxt) #> unyxml #> simplify_spaces
 
-val spying_version = "c"
+val spying_version = "d"
 
 fun spying false _ = ()
   | spying true f =
--- a/src/HOL/Tools/TFL/tfl.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/TFL/tfl.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -434,7 +434,7 @@
      val case_simpset =
        put_simpset HOL_basic_ss ctxt
           addsimps case_rewrites
-          |> fold (Simplifier.add_cong o #weak_case_cong o snd)
+          |> fold (Simplifier.add_cong o #case_cong_weak o snd)
               (Symtab.dest (Datatype.get_all theory))
      val corollaries' = map (Simplifier.simplify case_simpset) corollaries
      val extract = Rules.CONTEXT_REWRITE_RULE
--- a/src/HOL/Tools/Transfer/transfer_bnf.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/Transfer/transfer_bnf.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -112,8 +112,9 @@
   let
     val old_ctxt = ctxt
     val (goal, ctxt) = generate_relation_constraint_goal ctxt bnf constraint_def
-    val thm = Goal.prove ctxt [] [] goal
+    val thm = Goal.prove_sorry ctxt [] [] goal
       (fn {context = ctxt, prems = _} => side_constraint_tac bnf [constraint_def] ctxt 1)
+      |> Thm.close_derivation
   in
     Drule.zero_var_indexes (singleton (Variable.export ctxt old_ctxt) thm)
   end
@@ -122,8 +123,9 @@
   let
     val old_ctxt = ctxt
     val (goal, ctxt) = generate_relation_constraint_goal ctxt bnf constraint_def
-    val thm = Goal.prove ctxt [] [] goal
+    val thm = Goal.prove_sorry ctxt [] [] goal
       (fn {context = ctxt, prems = _} => bi_constraint_tac constraint_def side_constraints ctxt 1)
+      |> Thm.close_derivation
   in
     Drule.zero_var_indexes (singleton (Variable.export ctxt old_ctxt) thm)
   end
@@ -219,8 +221,9 @@
     val lhs = mk_Domainp (list_comb (relator, args))
     val rhs = mk_pred pred_def (map mk_Domainp args) T
     val goal = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop
-    val thm = Goal.prove ctxt [] [] goal
+    val thm = Goal.prove_sorry ctxt [] [] goal
       (fn {context = ctxt, prems = _} => Domainp_tac bnf pred_def ctxt 1)
+      |> Thm.close_derivation
   in
     Drule.zero_var_indexes (singleton (Variable.export ctxt old_ctxt) thm)
   end
@@ -244,8 +247,9 @@
     val lhs = list_comb (relator, map mk_eq_onp args)
     val rhs = mk_eq_onp (mk_pred pred_def args T)
     val goal = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop
-    val thm = Goal.prove ctxt [] [] goal
+    val thm = Goal.prove_sorry ctxt [] [] goal
       (fn {context = ctxt, prems = _} => pred_eq_onp_tac bnf pred_def ctxt 1)
+      |> Thm.close_derivation
   in
     Drule.zero_var_indexes (singleton (Variable.export ctxt old_ctxt) thm)
   end
--- a/src/HOL/Tools/arith_data.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/arith_data.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -9,7 +9,6 @@
   val arith_tac: Proof.context -> int -> tactic
   val verbose_arith_tac: Proof.context -> int -> tactic
   val add_tactic: string -> (bool -> Proof.context -> int -> tactic) -> theory -> theory
-  val get_arith_facts: Proof.context -> thm list
 
   val mk_number: typ -> int -> term
   val mk_sum: typ -> term list -> term
@@ -28,15 +27,7 @@
 structure Arith_Data: ARITH_DATA =
 struct
 
-(* slots for plugging in arithmetic facts and tactics *)
-
-structure Arith_Facts = Named_Thms
-(
-  val name = @{binding arith}
-  val description = "arith facts -- only ground formulas"
-);
-
-val get_arith_facts = Arith_Facts.get;
+(* slot for plugging in tactics *)
 
 structure Arith_Tactics = Theory_Data
 (
@@ -58,11 +49,12 @@
 val verbose_arith_tac = gen_arith_tac true;
 
 val setup =
-  Arith_Facts.setup #>
   Method.setup @{binding arith}
     (Scan.succeed (fn ctxt =>
-      METHOD (fn facts => HEADGOAL (Method.insert_tac (get_arith_facts ctxt @ facts)
-        THEN' verbose_arith_tac ctxt))))
+      METHOD (fn facts =>
+        HEADGOAL
+        (Method.insert_tac (rev (Named_Theorems.get ctxt @{named_theorems arith}) @ facts)
+          THEN' verbose_arith_tac ctxt))))
     "various arithmetic decision procedures";
 
 
--- a/src/HOL/Tools/coinduction.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/coinduction.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -149,9 +149,8 @@
 val setup =
   Method.setup @{binding coinduction}
     (arbitrary -- Scan.option coinduct_rule >>
-      (fn (arbitrary, opt_rule) => fn ctxt =>
-        RAW_METHOD_CASES (fn facts =>
-          Seq.DETERM (coinduction_tac ctxt arbitrary opt_rule facts))))
+      (fn (arbitrary, opt_rule) => fn ctxt => fn facts =>
+        Seq.DETERM (coinduction_tac ctxt arbitrary opt_rule facts)))
     "coinduction on types or predicates/sets";
 
 end;
--- a/src/HOL/Tools/groebner.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/groebner.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -924,7 +924,7 @@
 
 fun presimplify ctxt add_thms del_thms =
   asm_full_simp_tac (put_simpset HOL_basic_ss ctxt
-    addsimps (Algebra_Simplification.get ctxt)
+    addsimps (Named_Theorems.get ctxt @{named_theorems algebra})
     delsimps del_thms addsimps add_thms);
 
 fun ring_tac add_ths del_ths ctxt =
--- a/src/HOL/Tools/inductive.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/inductive.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -53,7 +53,7 @@
     (binding * string option * mixfix) list ->
     (binding * string option * mixfix) list ->
     (Attrib.binding * string) list ->
-    (Facts.ref * Attrib.src list) list ->
+    (Facts.ref * Token.src list) list ->
     local_theory -> inductive_result * local_theory
   val add_inductive_global: inductive_flags ->
     ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list ->
@@ -76,7 +76,7 @@
     term list -> (binding * mixfix) list ->
     local_theory -> inductive_result * local_theory
   val declare_rules: binding -> bool -> bool -> string list -> term list ->
-    thm list -> binding list -> Attrib.src list list -> (thm * string list * int) list ->
+    thm list -> binding list -> Token.src list list -> (thm * string list * int) list ->
     thm list -> thm -> local_theory -> thm list * thm list * thm list * thm * thm list * local_theory
   val add_ind_def: add_ind_def
   val gen_add_inductive_i: add_ind_def -> inductive_flags ->
@@ -85,7 +85,7 @@
   val gen_add_inductive: add_ind_def -> bool -> bool ->
     (binding * string option * mixfix) list ->
     (binding * string option * mixfix) list ->
-    (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
+    (Attrib.binding * string) list -> (Facts.ref * Token.src list) list ->
     local_theory -> inductive_result * local_theory
   val gen_ind_decl: add_ind_def -> bool -> (local_theory -> local_theory) parser
 end;
@@ -1157,7 +1157,7 @@
 fun gen_ind_decl mk_def coind =
   Parse.fixes -- Parse.for_fixes --
   Scan.optional Parse_Spec.where_alt_specs [] --
-  Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse_Spec.xthms1) []
+  Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse.xthms1) []
   >> (fn (((preds, params), specs), monos) =>
       (snd o gen_add_inductive mk_def true coind preds params specs monos));
 
--- a/src/HOL/Tools/inductive_set.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/inductive_set.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -20,7 +20,7 @@
   val add_inductive: bool -> bool ->
     (binding * string option * mixfix) list ->
     (binding * string option * mixfix) list ->
-    (Attrib.binding * string) list -> (Facts.ref * Attrib.src list) list ->
+    (Attrib.binding * string) list -> (Facts.ref * Token.src list) list ->
     local_theory -> Inductive.inductive_result * local_theory
   val mono_add: attribute
   val mono_del: attribute
--- a/src/HOL/Tools/lin_arith.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/lin_arith.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -796,7 +796,7 @@
 
 (* FIXME !?? *)
 fun add_arith_facts ctxt =
-  Simplifier.add_prems (Arith_Data.get_arith_facts ctxt) ctxt;
+  Simplifier.add_prems (rev (Named_Theorems.get ctxt @{named_theorems arith})) ctxt;
 
 val simproc = add_arith_facts #> Fast_Arith.lin_arith_simproc;
 
@@ -889,7 +889,7 @@
   Method.setup @{binding linarith}
     (Scan.succeed (fn ctxt =>
       METHOD (fn facts =>
-        HEADGOAL (Method.insert_tac (Arith_Data.get_arith_facts ctxt @ facts)
+        HEADGOAL (Method.insert_tac (rev (Named_Theorems.get ctxt @{named_theorems arith}) @ facts)
           THEN' tac ctxt)))) "linear arithmetic" #>
   Arith_Data.add_tactic "linear arithmetic" gen_tac;
 
--- a/src/HOL/Tools/recdef.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/recdef.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -15,17 +15,17 @@
   val cong_del: attribute
   val wf_add: attribute
   val wf_del: attribute
-  val add_recdef: bool -> xstring -> string -> ((binding * string) * Attrib.src list) list ->
-    Attrib.src option -> theory -> theory
+  val add_recdef: bool -> xstring -> string -> ((binding * string) * Token.src list) list ->
+    Token.src option -> theory -> theory
       * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
   val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list ->
     theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}
-  val defer_recdef: xstring -> string list -> (Facts.ref * Attrib.src list) list
+  val defer_recdef: xstring -> string list -> (Facts.ref * Token.src list) list
     -> theory -> theory * {induct_rules: thm}
   val defer_recdef_i: xstring -> term list -> thm list -> theory -> theory * {induct_rules: thm}
-  val recdef_tc: bstring * Attrib.src list -> xstring -> int option -> bool ->
+  val recdef_tc: bstring * Token.src list -> xstring -> int option -> bool ->
     local_theory -> Proof.state
-  val recdef_tc_i: bstring * Attrib.src list -> string -> int option -> bool ->
+  val recdef_tc_i: bstring * Token.src list -> string -> int option -> bool ->
     local_theory -> Proof.state
   val setup: theory -> theory
 end;
@@ -143,15 +143,15 @@
 val recdef_wfN = "recdef_wf";
 
 val recdef_modifiers =
- [Args.$$$ recdef_simpN -- Args.colon >> K ((I, simp_add): Method.modifier),
-  Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (I, simp_add),
-  Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (I, simp_del),
-  Args.$$$ recdef_congN -- Args.colon >> K (I, cong_add),
-  Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (I, cong_add),
-  Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (I, cong_del),
-  Args.$$$ recdef_wfN -- Args.colon >> K (I, wf_add),
-  Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (I, wf_add),
-  Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (I, wf_del)] @
+ [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
+  Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
+  Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
+  Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add @{here}),
+  Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
+  Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here}),
+  Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add @{here}),
+  Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add @{here}),
+  Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del @{here})] @
   Clasimp.clasimp_modifiers;
 
 
@@ -164,7 +164,7 @@
     val ctxt =
       (case opt_src of
         NONE => ctxt0
-      | SOME src => #2 (Args.syntax (Method.sections recdef_modifiers) src ctxt0));
+      | SOME src => #2 (Token.syntax (Method.sections recdef_modifiers) src ctxt0));
     val {simps, congs, wfs} = get_hints ctxt;
     val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong};
   in (ctxt', rev (map snd congs), wfs) end;
@@ -202,7 +202,8 @@
       tfl_fn not_permissive ctxt congs wfs name R eqs thy;
     val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx);
     val simp_att =
-      if null tcs then [Simplifier.simp_add, Nitpick_Simps.add, Code.add_default_eqn_attribute]
+      if null tcs then [Simplifier.simp_add,
+        Named_Theorems.add @{named_theorems nitpick_simp}, Code.add_default_eqn_attribute]
       else [];
     val ((simps' :: rules', [induct']), thy) =
       thy
@@ -291,7 +292,7 @@
 val hints =
   @{keyword "("} |--
     Parse.!!! (Parse.position @{keyword "hints"} -- Parse.args --| @{keyword ")"})
-  >> uncurry Args.src;
+  >> uncurry Token.src;
 
 val recdef_decl =
   Scan.optional
@@ -308,7 +309,7 @@
 val defer_recdef_decl =
   Parse.name -- Scan.repeat1 Parse.prop --
   Scan.optional
-    (@{keyword "("} |-- @{keyword "congs"} |-- Parse.!!! (Parse_Spec.xthms1 --| @{keyword ")"})) []
+    (@{keyword "("} |-- @{keyword "congs"} |-- Parse.!!! (Parse.xthms1 --| @{keyword ")"})) []
   >> (fn ((f, eqs), congs) => #1 o defer_recdef f eqs congs);
 
 val _ =
--- a/src/HOL/Tools/record.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/record.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -1785,10 +1785,10 @@
   Ctr_Sugar.default_register_ctr_sugar_global T_name
     {ctrs = [ctr], casex = Term.dummy, discs = [], selss = [sels], exhaust = exhaust,
      nchotomy = Drule.dummy_thm, injects = [inject], distincts = [], case_thms = [],
-     case_cong = Drule.dummy_thm, weak_case_cong = Drule.dummy_thm, split = Drule.dummy_thm,
+     case_cong = Drule.dummy_thm, case_cong_weak = Drule.dummy_thm, split = Drule.dummy_thm,
      split_asm = Drule.dummy_thm, disc_defs = [], disc_thmss = [], discIs = [], sel_defs = sel_defs,
-     sel_thmss = [sel_thms], disc_excludesss = [], disc_exhausts = [], sel_exhausts = [],
-     collapses = [], expands = [], sel_splits = [], sel_split_asms = [], case_eq_ifs = []};
+     sel_thmss = [sel_thms], distinct_discsss = [], exhaust_discs = [], exhaust_sels = [],
+     collapses = [], expands = [], split_sels = [], split_sel_asms = [], case_eq_ifs = []};
 
 
 (* definition *)
--- a/src/HOL/Tools/try0.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Tools/try0.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -48,18 +48,18 @@
       NONE
   end;
 
-val parse_method =
-  enclose "(" ")"
-  #> Outer_Syntax.scan Position.start
-  #> filter Token.is_proper
-  #> Scan.read Token.stopper Method.parse
-  #> (fn SOME (Method.Source src, _) => src | _ => raise Fail "expected Source");
+fun parse_method s =
+  enclose "(" ")" s
+  |> Outer_Syntax.scan (Keyword.get_lexicons ()) Position.start
+  |> filter Token.is_proper
+  |> Scan.read Token.stopper Method.parse
+  |> (fn SOME (Method.Source src, _) => src | _ => raise Fail "expected Source");
 
 fun apply_named_method_on_first_goal ctxt =
   parse_method
   #> Method.method_cmd ctxt
   #> Method.Basic
-  #> (fn m => Method.Select_Goals (Method.no_combinator_info, 1, m))
+  #> (fn m => Method.Combinator (Method.no_combinator_info, Method.Select_Goals 1, [m]))
   #> Proof.refine;
 
 fun add_attr_text (NONE, _) s = s
@@ -173,10 +173,10 @@
 
 fun string_of_xthm (xref, args) =
   Facts.string_of_ref xref ^
-  implode (map (enclose "[" "]" o Pretty.str_of o Args.pretty_src @{context}) args);
+  implode (map (enclose "[" "]" o Pretty.str_of o Token.pretty_src @{context}) args);
 
 val parse_fact_refs =
-  Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) (Parse_Spec.xthm >> string_of_xthm));
+  Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) (Parse.xthm >> string_of_xthm));
 
 val parse_attr =
   Args.$$$ "simp" |-- Args.colon |-- parse_fact_refs >> (fn ss => (ss, [], [], []))
--- a/src/HOL/Topological_Spaces.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/HOL/Topological_Spaces.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -9,17 +9,8 @@
 imports Main Conditionally_Complete_Lattices
 begin
 
-ML {*
-
-structure Continuous_Intros = Named_Thms
-(
-  val name = @{binding continuous_intros}
-  val description = "Structural introduction rules for continuity"
-)
-
-*}
-
-setup Continuous_Intros.setup
+named_theorems continuous_intros "structural introduction rules for continuity"
+
 
 subsection {* Topological space *}
 
@@ -1100,20 +1091,12 @@
 lemma tendsto_eq_rhs: "(f ---> x) F \<Longrightarrow> x = y \<Longrightarrow> (f ---> y) F"
   by simp
 
-ML {*
-
-structure Tendsto_Intros = Named_Thms
-(
-  val name = @{binding tendsto_intros}
-  val description = "introduction rules for tendsto"
-)
-
-*}
-
+named_theorems tendsto_intros "introduction rules for tendsto"
 setup {*
-  Tendsto_Intros.setup #>
   Global_Theory.add_thms_dynamic (@{binding tendsto_eq_intros},
-    map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm])) o Tendsto_Intros.get o Context.proof_of);
+    fn context =>
+      Named_Theorems.get (Context.proof_of context) @{named_theorems tendsto_intros}
+      |> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm])))
 *}
 
 lemma (in topological_space) tendsto_def:
--- a/src/Provers/clasimp.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Provers/clasimp.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -146,7 +146,7 @@
       ORELSE'
       (CHANGED o nodup_depth_tac (addss ctxt) n);  (* slower but more general *)
   in
-    PARALLEL_GOALS (ALLGOALS (Simplifier.asm_full_simp_tac ctxt)) THEN
+    PARALLEL_ALLGOALS (Simplifier.asm_full_simp_tac ctxt) THEN
     TRY (Classical.safe_tac ctxt) THEN
     REPEAT_DETERM (FIRSTGOAL main_tac) THEN
     TRY (Classical.safe_tac (addSss ctxt)) THEN
@@ -191,9 +191,9 @@
 val iffN = "iff";
 
 val iff_modifiers =
- [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K ((I, iff_add): Method.modifier),
-  Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (I, iff_add'),
-  Args.$$$ iffN -- Args.del -- Args.colon >> K (I, iff_del)];
+ [Args.$$$ iffN -- Scan.option Args.add -- Args.colon >> K (Method.modifier iff_add @{here}),
+  Args.$$$ iffN -- Scan.option Args.add -- Args.query_colon >> K (Method.modifier iff_add' @{here}),
+  Args.$$$ iffN -- Args.del -- Args.colon >> K (Method.modifier iff_del @{here})];
 
 val clasimp_modifiers =
   Simplifier.simp_modifiers @ Splitter.split_modifiers @
--- a/src/Provers/classical.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Provers/classical.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -928,13 +928,13 @@
 (* automatic methods *)
 
 val cla_modifiers =
- [Args.$$$ destN -- Args.bang_colon >> K ((I, safe_dest NONE): Method.modifier),
-  Args.$$$ destN -- Args.colon >> K (I, haz_dest NONE),
-  Args.$$$ elimN -- Args.bang_colon >> K (I, safe_elim NONE),
-  Args.$$$ elimN -- Args.colon >> K (I, haz_elim NONE),
-  Args.$$$ introN -- Args.bang_colon >> K (I, safe_intro NONE),
-  Args.$$$ introN -- Args.colon >> K (I, haz_intro NONE),
-  Args.del -- Args.colon >> K (I, rule_del)];
+ [Args.$$$ destN -- Args.bang_colon >> K (Method.modifier (safe_dest NONE) @{here}),
+  Args.$$$ destN -- Args.colon >> K (Method.modifier (haz_dest NONE) @{here}),
+  Args.$$$ elimN -- Args.bang_colon >> K (Method.modifier (safe_elim NONE) @{here}),
+  Args.$$$ elimN -- Args.colon >> K (Method.modifier (haz_elim NONE) @{here}),
+  Args.$$$ introN -- Args.bang_colon >> K (Method.modifier (safe_intro NONE) @{here}),
+  Args.$$$ introN -- Args.colon >> K (Method.modifier (haz_intro NONE) @{here}),
+  Args.del -- Args.colon >> K (Method.modifier rule_del @{here})];
 
 fun cla_method tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD o tac);
 fun cla_method' tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD' o tac);
--- a/src/Provers/splitter.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Provers/splitter.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -457,9 +457,9 @@
 (* methods *)
 
 val split_modifiers =
- [Args.$$$ splitN -- Args.colon >> K ((I, split_add): Method.modifier),
-  Args.$$$ splitN -- Args.add -- Args.colon >> K (I, split_add),
-  Args.$$$ splitN -- Args.del -- Args.colon >> K (I, split_del)];
+ [Args.$$$ splitN -- Args.colon >> K (Method.modifier split_add @{here}),
+  Args.$$$ splitN -- Args.add -- Args.colon >> K (Method.modifier split_add @{here}),
+  Args.$$$ splitN -- Args.del -- Args.colon >> K (Method.modifier split_del @{here})];
 
 
 (* theory setup *)
--- a/src/Pure/Concurrent/future.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Concurrent/future.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -37,7 +37,7 @@
   def join: A
   def map[B](f: A => B): Future[B] = Future.fork { f(join) }
 
-  override def toString =
+  override def toString: String =
     peek match {
       case None => "<future>"
       case Some(Exn.Exn(_)) => "<failed>"
--- a/src/Pure/GUI/gui.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/GUI/gui.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -7,9 +7,10 @@
 
 package isabelle
 
-
 import java.lang.{ClassLoader, ClassNotFoundException, NoSuchMethodException}
-import java.awt.{Image, Component, Container, Toolkit, Window, Font, KeyboardFocusManager}
+import java.io.{FileInputStream, BufferedInputStream}
+import java.awt.{GraphicsEnvironment, Image, Component, Container, Toolkit, Window, Font,
+  KeyboardFocusManager}
 import java.awt.font.{TextAttribute, TransformAttribute, FontRenderContext, LineMetrics}
 import java.awt.geom.AffineTransform
 import javax.swing.{ImageIcon, JOptionPane, UIManager, JLayeredPane, JFrame, JWindow, JDialog,
@@ -232,5 +233,15 @@
     import scala.collection.JavaConversions._
     font.deriveFont(Map(TextAttribute.TRANSFORM -> new TransformAttribute(transform)))
   }
+
+  def font(family: String = "IsabelleText", size: Int = 1, bold: Boolean = false): Font =
+    new Font(family, if (bold) Font.BOLD else Font.PLAIN, size)
+
+  def install_fonts()
+  {
+    val ge = GraphicsEnvironment.getLocalGraphicsEnvironment()
+    for (font <- Path.split(Isabelle_System.getenv_strict("ISABELLE_FONTS")))
+      ge.registerFont(Font.createFont(Font.TRUETYPE_FONT, font.file))
+  }
 }
 
--- a/src/Pure/GUI/html5_panel.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/GUI/html5_panel.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -1,5 +1,5 @@
 /*  Title:      Pure/GUI/html5_panel.scala
-    Module:     PIDE-GUI
+    Module:     PIDE-JFX
     Author:     Makarius
 
 HTML5 panel based on Java FX WebView.
@@ -54,7 +54,7 @@
 class HTML5_Panel extends javafx.embed.swing.JFXPanel
 {
   private val future =
-    JFX_Thread.future {
+    JFX_GUI.Thread.future {
       val pane = new Web_View_Workaround
 
       val web_view = pane.web_view
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/GUI/jfx_gui.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,55 @@
+/*  Title:      Pure/GUI/jfx_thread.scala
+    Module:     PIDE-JFX
+    Author:     Makarius
+
+Basic GUI tools (for Java FX).
+*/
+
+package isabelle
+
+
+import java.io.{FileInputStream, BufferedInputStream}
+
+import javafx.application.{Platform => JFX_Platform}
+import javafx.scene.text.{Font => JFX_Font}
+
+
+object JFX_GUI
+{
+  /* evaluation within the Java FX application thread */
+
+  object Thread
+  {
+    def assert() = Predef.assert(JFX_Platform.isFxApplicationThread())
+    def require() = Predef.require(JFX_Platform.isFxApplicationThread())
+
+    def later(body: => Unit)
+    {
+      if (JFX_Platform.isFxApplicationThread()) body
+      else JFX_Platform.runLater(new Runnable { def run = body })
+    }
+
+    def future[A](body: => A): Future[A] =
+    {
+      if (JFX_Platform.isFxApplicationThread()) Future.value(body)
+      else {
+        val promise = Future.promise[A]
+        later { promise.fulfill_result(Exn.capture(body)) }
+        promise
+      }
+    }
+  }
+
+
+  /* Isabelle fonts */
+
+  def install_fonts()
+  {
+    for (font <- Path.split(Isabelle_System.getenv_strict("ISABELLE_FONTS"))) {
+      val stream = new BufferedInputStream(new FileInputStream(font.file))
+      try { JFX_Font.loadFont(stream, 1.0) }
+      finally { stream.close }
+    }
+  }
+
+}
--- a/src/Pure/GUI/jfx_thread.scala	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-/*  Title:      Pure/GUI/jfx_thread.scala
-    Module:     PIDE-GUI
-    Author:     Makarius
-
-Evaluation within the Java FX application thread.
-*/
-
-package isabelle
-
-import javafx.application.{Platform => JFX_Platform}
-
-
-object JFX_Thread
-{
-  /* checks */
-
-  def assert() = Predef.assert(JFX_Platform.isFxApplicationThread())
-  def require() = Predef.require(JFX_Platform.isFxApplicationThread())
-
-
-  /* asynchronous context switch */
-
-  def later(body: => Unit)
-  {
-    if (JFX_Platform.isFxApplicationThread()) body
-    else JFX_Platform.runLater(new Runnable { def run = body })
-  }
-
-  def future[A](body: => A): Future[A] =
-  {
-    if (JFX_Platform.isFxApplicationThread()) Future.value(body)
-    else {
-      val promise = Future.promise[A]
-      later { promise.fulfill_result(Exn.capture(body)) }
-      promise
-    }
-  }
-}
--- a/src/Pure/General/binding.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/binding.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -31,6 +31,7 @@
   val conceal: binding -> binding
   val pretty: binding -> Pretty.T
   val print: binding -> string
+  val pp: binding -> Pretty.T
   val bad: binding -> string
   val check: binding -> unit
 end;
@@ -133,6 +134,8 @@
 
 val print = Pretty.str_of o pretty;
 
+val pp = pretty o set_pos Position.none;
+
 
 (* check *)
 
--- a/src/Pure/General/name_space.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/name_space.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -14,7 +14,7 @@
   val kind_of: T -> string
   val defined_entry: T -> string -> bool
   val the_entry: T -> string ->
-    {concealed: bool, group: serial option, theory_name: string, pos: Position.T, id: serial}
+    {concealed: bool, group: serial option, theory_name: string, pos: Position.T, serial: serial}
   val entry_ord: T -> string * string -> order
   val markup: T -> string -> Markup.T
   val is_concealed: T -> string -> bool
@@ -92,10 +92,10 @@
   group: serial option,
   theory_name: string,
   pos: Position.T,
-  id: serial};
+  serial: serial};
 
-fun entry_markup def kind (name, {pos, id, ...}: entry) =
-  Markup.properties (Position.entity_properties_of def id pos) (Markup.entity kind name);
+fun entry_markup def kind (name, {pos, serial, ...}: entry) =
+  Markup.properties (Position.entity_properties_of def serial pos) (Markup.entity kind name);
 
 fun print_entry_ref kind (name, entry) =
   quote (Markup.markup (entry_markup false kind (name, entry)) name);
@@ -152,7 +152,7 @@
     NONE => error (undefined kind name)
   | SOME (_, entry) => entry);
 
-fun entry_ord space = int_ord o pairself (#id o the_entry space);
+fun entry_ord space = int_ord o pairself (#serial o the_entry space);
 
 fun markup (Name_Space {kind, entries, ...}) name =
   (case Change_Table.lookup entries name of
@@ -275,7 +275,7 @@
         else (Library.merge (op =) (names1, names2), Library.merge (op =) (names1', names2'))));
     val entries' = (entries1, entries2) |> Change_Table.join
       (fn name => fn ((_, entry1), (_, entry2)) =>
-        if #id entry1 = #id entry2 then raise Change_Table.SAME
+        if #serial entry1 = #serial entry2 then raise Change_Table.SAME
         else err_dup kind' (name, entry1) (name, entry2) Position.none);
   in make_name_space (kind', internals', entries') end;
 
@@ -448,7 +448,7 @@
       group = group,
       theory_name = theory_name,
       pos = pos,
-      id = serial ()};
+      serial = serial ()};
     val space' =
       space |> map_name_space (fn (kind, internals, entries) =>
         let
--- a/src/Pure/General/position.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/position.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -153,9 +153,9 @@
 
 val def_properties_of = properties_of #> map (fn (x, y) => ("def_" ^ x, y));
 
-fun entity_properties_of def id pos =
-  if def then (Markup.defN, Markup.print_int id) :: properties_of pos
-  else (Markup.refN, Markup.print_int id) :: def_properties_of pos;
+fun entity_properties_of def serial pos =
+  if def then (Markup.defN, Markup.print_int serial) :: properties_of pos
+  else (Markup.refN, Markup.print_int serial) :: def_properties_of pos;
 
 fun default_properties default props =
   if exists (member (op =) Markup.position_properties o #1) props then props
--- a/src/Pure/General/position.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/position.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -54,7 +54,7 @@
 
   object Range
   {
-    def apply(range: Symbol.Range): T = Offset(range.start) ::: Offset(range.stop)
+    def apply(range: Symbol.Range): T = Offset(range.start) ::: End_Offset(range.stop)
     def unapply(pos: T): Option[Symbol.Range] =
       (pos, pos) match {
         case (Offset(start), End_Offset(stop)) if start <= stop => Some(Text.Range(start, stop))
@@ -63,35 +63,35 @@
       }
   }
 
-  object Id_Offset
+  object Id_Offset0
   {
     def unapply(pos: T): Option[(Long, Symbol.Offset)] =
-      (pos, pos) match {
-        case (Id(id), Offset(offset)) => Some((id, offset))
+      pos match {
+        case Id(id) => Some((id, Offset.unapply(pos) getOrElse 0))
         case _ => None
       }
   }
 
-  object Def_Id_Offset
+  object Def_Id_Offset0
   {
     def unapply(pos: T): Option[(Long, Symbol.Offset)] =
-      (pos, pos) match {
-        case (Def_Id(id), Def_Offset(offset)) => Some((id, offset))
+      pos match {
+        case Def_Id(id) => Some((id, Def_Offset.unapply(pos) getOrElse 0))
         case _ => None
       }
   }
 
-  object Reported
+  object Identified
   {
-    def unapply(pos: T): Option[(Long, Symbol.Text_Chunk.Name, Symbol.Range)] =
-      (pos, pos) match {
-        case (Id(id), Range(range)) =>
+    def unapply(pos: T): Option[(Long, Symbol.Text_Chunk.Name)] =
+      pos match {
+        case Id(id) =>
           val chunk_name =
             pos match {
               case File(name) => Symbol.Text_Chunk.File(name)
               case _ => Symbol.Text_Chunk.Default
             }
-          Some((id, chunk_name, range))
+          Some((id, chunk_name))
         case _ => None
       }
   }
--- a/src/Pure/General/properties.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/properties.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -27,7 +27,7 @@
 
     object Int
     {
-      def apply(x: scala.Int): java.lang.String = x.toString
+      def apply(x: scala.Int): java.lang.String = Library.signed_string_of_int(x)
       def unapply(s: java.lang.String): Option[scala.Int] =
         try { Some(Integer.parseInt(s)) }
         catch { case _: NumberFormatException => None }
@@ -35,7 +35,7 @@
 
     object Long
     {
-      def apply(x: scala.Long): java.lang.String = x.toString
+      def apply(x: scala.Long): java.lang.String = Library.signed_string_of_long(x)
       def unapply(s: java.lang.String): Option[scala.Long] =
         try { Some(java.lang.Long.parseLong(s)) }
         catch { case _: NumberFormatException => None }
--- a/src/Pure/General/symbol_pos.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/symbol_pos.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -41,6 +41,7 @@
   val implode_range: Position.T -> Position.T -> T list -> text * Position.range
   val explode: text * Position.T -> T list
   type source = {delimited: bool, text: text, pos: Position.T}
+  val source_explode: source -> T list
   val source_content: source -> string * Position.T
   val scan_ident: T list -> T list * T list
   val is_identifier: string -> bool
@@ -261,6 +262,8 @@
 
 type source = {delimited: bool, text: text, pos: Position.T};
 
+fun source_explode {delimited = _, text, pos} = explode (text, pos);
+
 fun source_content {delimited = _, text, pos} =
   let val syms = explode (text, pos) in (content syms, pos) end;
 
--- a/src/Pure/General/time.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/time.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -40,7 +40,7 @@
   def is_zero: Boolean = ms == 0
   def is_relevant: Boolean = ms >= 1
 
-  override def toString = Time.print_seconds(seconds)
+  override def toString: String = Time.print_seconds(seconds)
 
   def message: String = toString + "s"
 }
--- a/src/Pure/General/timing.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/General/timing.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -38,6 +38,6 @@
   def message: String =
     elapsed.message + " elapsed time, " + cpu.message + " cpu time, " + gc.message + " GC time"
 
-  override def toString = message
+  override def toString: String = message
 }
 
--- a/src/Pure/Isar/args.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/args.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -1,22 +1,12 @@
 (*  Title:      Pure/Isar/args.ML
     Author:     Markus Wenzel, TU Muenchen
 
-Parsing with implicit value assignment.  Concrete argument syntax of
+Quasi-inner syntax based on outer tokens: concrete argument syntax of
 attributes, methods etc.
 *)
 
 signature ARGS =
 sig
-  type src
-  val src: xstring * Position.T -> Token.T list -> src
-  val name_of_src: src -> string * Position.T
-  val range_of_src: src -> Position.T
-  val unparse_src: src -> string list
-  val pretty_src: Proof.context -> src -> Pretty.T
-  val check_src: Proof.context -> 'a Name_Space.table -> src -> src * 'a
-  val transform_values: morphism -> src -> src
-  val init_assignable: src -> src
-  val closure: src -> src
   val context: Proof.context context_parser
   val theory: theory context_parser
   val $$$ : string -> string parser
@@ -43,17 +33,20 @@
   val symbol: string parser
   val liberal_name: string parser
   val var: indexname parser
-  val internal_text: string parser
+  val internal_source: Token.src parser
+  val internal_name: (string * morphism) parser
   val internal_typ: typ parser
   val internal_term: term parser
   val internal_fact: thm list parser
   val internal_attribute: (morphism -> attribute) parser
-  val named_text: (string -> string) -> string parser
+  val internal_declaration: declaration parser
+  val named_source: (Token.T -> Token.src) -> Token.src parser
   val named_typ: (string -> typ) -> typ parser
   val named_term: (string -> term) -> term parser
-  val named_fact: (string -> thm list) -> thm list parser
-  val named_attribute:
-    (string * Position.T -> morphism -> attribute) -> (morphism -> attribute) parser
+  val named_fact: (string -> string option * thm list) -> thm list parser
+  val named_attribute: (string * Position.T -> morphism -> attribute) ->
+    (morphism -> attribute) parser
+  val text_declaration: (Symbol_Pos.source -> declaration) -> declaration parser
   val typ_abbrev: typ context_parser
   val typ: typ context_parser
   val term: term context_parser
@@ -63,82 +56,13 @@
   val type_name: {proper: bool, strict: bool} -> string context_parser
   val const: {proper: bool, strict: bool} -> string context_parser
   val goal_spec: ((int -> tactic) -> tactic) context_parser
-  val attribs: (xstring * Position.T -> string) -> src list parser
-  val opt_attribs: (xstring * Position.T -> string) -> src list parser
-  val syntax_generic: 'a context_parser -> src -> Context.generic -> 'a * Context.generic
-  val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context
+  val attribs: (xstring * Position.T -> string) -> Token.src list parser
+  val opt_attribs: (xstring * Position.T -> string) -> Token.src list parser
 end;
 
 structure Args: ARGS =
 struct
 
-(** datatype src **)
-
-datatype src =
-  Src of
-   {name: string * Position.T,
-    args: Token.T list,
-    output_info: (string * Markup.T) option};
-
-fun src name args = Src {name = name, args = args, output_info = NONE};
-
-fun name_of_src (Src {name, ...}) = name;
-
-fun range_of_src (Src {name = (_, pos), args, ...}) =
-  if null args then pos
-  else Position.set_range (pos, #2 (Token.range_of args));
-
-fun unparse_src (Src {args, ...}) = map Token.unparse args;
-
-fun pretty_src ctxt src =
-  let
-    val Src {name = (name, _), args, output_info} = src;
-    val prt_name =
-      (case output_info of
-        NONE => Pretty.str name
-      | SOME (_, markup) => Pretty.mark_str (markup, name));
-    val prt_thm = Pretty.backquote o Display.pretty_thm ctxt;
-    fun prt_arg arg =
-      (case Token.get_value arg of
-        SOME (Token.Literal markup) =>
-          let val x = Token.content_of arg
-          in Pretty.mark_str (Token.keyword_markup markup x, x) end
-      | SOME (Token.Text s) => Pretty.str (quote s)
-      | SOME (Token.Typ T) => Syntax.pretty_typ ctxt T
-      | SOME (Token.Term t) => Syntax.pretty_term ctxt t
-      | SOME (Token.Fact ths) => Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths))
-      | _ => Pretty.mark_str (Token.markup arg, Token.unparse arg));
-  in Pretty.block (Pretty.breaks (prt_name :: map prt_arg args)) end;
-
-
-(* check *)
-
-fun check_src ctxt table (Src {name = (xname, pos), args, output_info = _}) =
-  let
-    val (name, x) = Name_Space.check (Context.Proof ctxt) table (xname, pos);
-    val space = Name_Space.space_of_table table;
-    val kind = Name_Space.kind_of space;
-    val markup = Name_Space.markup space name;
-  in (Src {name = (name, pos), args = args, output_info = SOME (kind, markup)}, x) end;
-
-
-(* values *)
-
-fun map_args f (Src {name, args, output_info}) =
-  Src {name = name, args = map f args, output_info = output_info};
-
-fun transform_values phi = map_args (Token.map_value
-  (fn Token.Typ T => Token.Typ (Morphism.typ phi T)
-    | Token.Term t => Token.Term (Morphism.term phi t)
-    | Token.Fact ths => Token.Fact (Morphism.fact phi ths)
-    | Token.Attribute att => Token.Attribute (Morphism.transform phi att)
-    | tok => tok));
-
-val init_assignable = map_args Token.init_assignable;
-val closure = map_args Token.closure;
-
-
-
 (** argument scanners **)
 
 (* context *)
@@ -209,23 +133,32 @@
 fun evaluate mk eval arg =
   let val x = eval arg in (Token.assign (SOME (mk x)) arg; x) end;
 
-val internal_text = value (fn Token.Text s => s);
+val internal_source = value (fn Token.Source src => src);
+val internal_name = value (fn Token.Name a => a);
 val internal_typ = value (fn Token.Typ T => T);
 val internal_term = value (fn Token.Term t => t);
-val internal_fact = value (fn Token.Fact ths => ths);
+val internal_fact = value (fn Token.Fact (_, ths) => ths);
 val internal_attribute = value (fn Token.Attribute att => att);
+val internal_declaration = value (fn Token.Declaration decl => decl);
 
-fun named_text intern = internal_text || named >> evaluate Token.Text (intern o Token.content_of);
-fun named_typ readT = internal_typ || named >> evaluate Token.Typ (readT o Token.inner_syntax_of);
+fun named_source read = internal_source || named >> evaluate Token.Source read;
+
+fun named_typ read = internal_typ || named >> evaluate Token.Typ (read o Token.inner_syntax_of);
 fun named_term read = internal_term || named >> evaluate Token.Term (read o Token.inner_syntax_of);
 
-fun named_fact get = internal_fact || named >> evaluate Token.Fact (get o Token.content_of) ||
-  alt_string >> evaluate Token.Fact (get o Token.inner_syntax_of);
+fun named_fact get =
+  internal_fact ||
+  named >> evaluate Token.Fact (get o Token.content_of) >> #2 ||
+  alt_string >> evaluate Token.Fact (get o Token.inner_syntax_of) >> #2;
 
 fun named_attribute att =
   internal_attribute ||
   named >> evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.pos_of tok));
 
+fun text_declaration read =
+  internal_declaration ||
+  text_token >> evaluate Token.Declaration (read o Token.source_position_of);
+
 
 (* terms and types *)
 
@@ -265,43 +198,10 @@
 fun attribs check =
   let
     fun intern tok = check (Token.content_of tok, Token.pos_of tok);
-    val attrib_name = internal_text || (symbolic || named) >> evaluate Token.Text intern;
-    val attrib = Parse.position attrib_name -- Parse.!!! Parse.args >> uncurry src;
+    val attrib_name = internal_name >> #1 || (symbolic || named) >> evaluate Token.name0 intern;
+    val attrib = Parse.position attrib_name -- Parse.!!! Parse.args >> uncurry Token.src;
   in $$$ "[" |-- Parse.!!! (Parse.list attrib --| $$$ "]") end;
 
 fun opt_attribs check = Scan.optional (attribs check) [];
 
-
-
-(** syntax wrapper **)
-
-fun syntax_generic scan (Src {name = (name, pos), args = args0, output_info}) context =
-  let
-    val args1 = map Token.init_assignable args0;
-    fun reported_text () =
-      if Context_Position.is_visible_generic context then
-        ((pos, Markup.operator) :: maps (Token.reports_of_value o Token.closure) args1)
-        |> map (fn (p, m) => Position.reported_text p m "")
-      else [];
-  in
-    (case Scan.error (Scan.finite' Token.stopper (Scan.option scan)) (context, args1) of
-      (SOME x, (context', [])) =>
-        let val _ = Output.report (reported_text ())
-        in (x, context') end
-    | (_, (_, args2)) =>
-        let
-          val print_name =
-            (case output_info of
-              NONE => quote name
-            | SOME (kind, markup) => plain_words kind ^ " " ^ quote (Markup.markup markup name));
-          val print_args =
-            if null args2 then "" else ":\n  " ^ space_implode " " (map Token.print args2);
-        in
-          error ("Bad arguments for " ^ print_name ^ Position.here pos ^ print_args ^
-            Markup.markup_report (implode (reported_text ())))
-        end)
-  end;
-
-fun syntax scan src = apsnd Context.the_proof o syntax_generic scan src o Context.Proof;
-
 end;
--- a/src/Pure/Isar/attrib.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/attrib.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,19 +6,22 @@
 
 signature ATTRIB =
 sig
-  type src = Args.src
-  type binding = binding * src list
+  type binding = binding * Token.src list
   val empty_binding: binding
   val is_empty_binding: binding -> bool
   val print_attributes: Proof.context -> unit
+  val define_global: Binding.binding -> (Token.src -> attribute) ->
+    string -> theory -> string * theory
+  val define: Binding.binding -> (Token.src -> attribute) ->
+    string -> local_theory -> string * local_theory
   val check_name_generic: Context.generic -> xstring * Position.T -> string
   val check_name: Proof.context -> xstring * Position.T -> string
-  val check_src: Proof.context -> src -> src
-  val pretty_attribs: Proof.context -> src list -> Pretty.T list
-  val attribute: Proof.context -> src -> attribute
-  val attribute_global: theory -> src -> attribute
-  val attribute_cmd: Proof.context -> src -> attribute
-  val attribute_cmd_global: theory -> src -> attribute
+  val check_src: Proof.context -> Token.src -> Token.src
+  val pretty_attribs: Proof.context -> Token.src list -> Pretty.T list
+  val attribute: Proof.context -> Token.src -> attribute
+  val attribute_global: theory -> Token.src -> attribute
+  val attribute_cmd: Proof.context -> Token.src -> attribute
+  val attribute_cmd_global: theory -> Token.src -> attribute
   val map_specs: ('a list -> 'att list) ->
     (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list
   val map_facts: ('a list -> 'att list) ->
@@ -27,24 +30,32 @@
   val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) ->
     (('c * 'a list) * ('b * 'a list) list) list ->
     (('c * 'att list) * ('fact * 'att list) list) list
-  val global_notes: string -> (binding * (thm list * src list) list) list ->
+  val global_notes: string -> (binding * (thm list * Token.src list) list) list ->
     theory -> (string * thm list) list * theory
-  val local_notes: string -> (binding * (thm list * src list) list) list ->
+  val local_notes: string -> (binding * (thm list * Token.src list) list) list ->
     Proof.context -> (string * thm list) list * Proof.context
-  val generic_notes: string -> (binding * (thm list * src list) list) list ->
+  val generic_notes: string -> (binding * (thm list * Token.src list) list) list ->
     Context.generic -> (string * thm list) list * Context.generic
-  val eval_thms: Proof.context -> (Facts.ref * src list) list -> thm list
+  val eval_thms: Proof.context -> (Facts.ref * Token.src list) list -> thm list
+  val attribute_syntax: attribute context_parser -> Token.src -> attribute
   val setup: Binding.binding -> attribute context_parser -> string -> theory -> theory
-  val attribute_setup: bstring * Position.T -> Symbol_Pos.source -> string -> theory -> theory
-  val internal: (morphism -> attribute) -> src
+  val local_setup: Binding.binding -> attribute context_parser -> string ->
+    local_theory -> local_theory
+  val attribute_setup: bstring * Position.T -> Symbol_Pos.source -> string ->
+    local_theory -> local_theory
+  val internal: (morphism -> attribute) -> Token.src
   val add_del: attribute -> attribute -> attribute context_parser
-  val thm_sel: Facts.interval list parser
   val thm: thm context_parser
   val thms: thm list context_parser
   val multi_thm: thm list context_parser
+  val check_attribs: Proof.context -> Token.src list -> Token.src list
+  val read_attribs: Proof.context -> Symbol_Pos.source -> Token.src list
+  val transform_facts: morphism ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list
   val partial_evaluation: Proof.context ->
-    (binding * (thm list * Args.src list) list) list ->
-    (binding * (thm list * Args.src list) list) list
+    (binding * (thm list * Token.src list) list) list ->
+    (binding * (thm list * Token.src list) list) list
   val print_options: Proof.context -> unit
   val config_bool: Binding.binding ->
     (Context.generic -> bool) -> bool Config.T * (theory -> theory)
@@ -73,9 +84,7 @@
 
 (* source and bindings *)
 
-type src = Args.src;
-
-type binding = binding * src list;
+type binding = binding * Token.src list;
 
 val empty_binding: binding = (Binding.empty, []);
 fun is_empty_binding ((b, srcs): binding) = Binding.is_empty b andalso null srcs;
@@ -86,19 +95,25 @@
 
 (* theory data *)
 
-structure Attributes = Theory_Data
+structure Attributes = Generic_Data
 (
-  type T = ((src -> attribute) * string) Name_Space.table;
+  type T = ((Token.src -> attribute) * string) Name_Space.table;
   val empty : T = Name_Space.empty_table "attribute";
   val extend = I;
   fun merge data : T = Name_Space.merge_tables data;
 );
 
-val get_attributes = Attributes.get o Context.theory_of;
+val get_attributes = Attributes.get o Context.Proof;
+
+fun transfer_attributes ctxt =
+  let
+    val attribs0 = Attributes.get (Context.Theory (Proof_Context.theory_of ctxt));
+    val attribs' = Name_Space.merge_tables (attribs0, get_attributes ctxt);
+  in Context.proof_map (Attributes.put attribs') ctxt end;
 
 fun print_attributes ctxt =
   let
-    val attribs = get_attributes (Context.Proof ctxt);
+    val attribs = get_attributes ctxt;
     fun prt_attr (name, (_, "")) = Pretty.mark_str name
       | prt_attr (name, (_, comment)) =
           Pretty.block
@@ -108,33 +123,47 @@
     |> Pretty.writeln_chunks
   end;
 
-val attribute_space = Name_Space.space_of_table o get_attributes o Context.Proof;
+val attribute_space = Name_Space.space_of_table o get_attributes;
+
+
+(* define *)
 
-fun add_attribute name att comment thy = thy
-  |> Attributes.map (Name_Space.define (Context.Theory thy) true (name, (att, comment)) #> snd);
+fun define_global binding att comment thy =
+  let
+    val context = Context.Theory thy;
+    val (name, attribs') =
+      Name_Space.define context true (binding, (att, comment)) (Attributes.get context);
+  in (name, Context.the_theory (Attributes.put attribs' context)) end;
+
+fun define binding att comment =
+  Local_Theory.background_theory_result (define_global binding att comment)
+  #-> (fn name =>
+    Local_Theory.map_contexts (K transfer_attributes)
+    #> Local_Theory.generic_alias Attributes.map binding name
+    #> pair name);
 
 
 (* check *)
 
-fun check_name_generic context = #1 o Name_Space.check context (get_attributes context);
+fun check_name_generic context = #1 o Name_Space.check context (Attributes.get context);
 val check_name = check_name_generic o Context.Proof;
 
 fun check_src ctxt src =
- (Context_Position.report ctxt (Args.range_of_src src) Markup.language_attribute;
-  #1 (Args.check_src ctxt (get_attributes (Context.Proof ctxt)) src));
+ (Context_Position.report ctxt (Token.range_of_src src) Markup.language_attribute;
+  #1 (Token.check_src ctxt (get_attributes ctxt) src));
 
 
 (* pretty printing *)
 
 fun pretty_attribs _ [] = []
-  | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Args.pretty_src ctxt) srcs)];
+  | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Token.pretty_src ctxt) srcs)];
 
 
 (* get attributes *)
 
 fun attribute_generic context =
-  let val table = get_attributes context
-  in fn src => #1 (Name_Space.get table (#1 (Args.name_of_src src))) src end;
+  let val table = Attributes.get context
+  in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end;
 
 val attribute = attribute_generic o Context.Proof;
 val attribute_global = attribute_generic o Context.Theory;
@@ -171,23 +200,27 @@
 
 (* attribute setup *)
 
-fun setup name scan =
-  add_attribute name
-    (fn src => fn (ctxt, th) =>
-      let val (a, ctxt') = Args.syntax_generic scan src ctxt in a (ctxt', th) end);
+fun attribute_syntax scan src (context, th) =
+  let val (a, context') = Token.syntax_generic scan src context in a (context', th) end;
+
+fun setup binding scan comment = define_global binding (attribute_syntax scan) comment #> snd;
+fun local_setup binding scan comment = define binding (attribute_syntax scan) comment #> snd;
 
 fun attribute_setup name source cmt =
-  Context.theory_map (ML_Context.expression (#pos source)
+  (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
+    ML_Lex.read_source false source @
+    ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")"))
+  |> ML_Context.expression (#pos source)
     "val (name, scan, comment): binding * attribute context_parser * string"
-    "Context.map_theory (Attrib.setup name scan comment)"
-    (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
-      ML_Lex.read_source false source @
-      ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")")));
+    "Context.map_proof (Attrib.local_setup name scan comment)"
+  |> Context.proof_map;
 
 
 (* internal attribute *)
 
-fun internal att = Args.src ("Pure.attribute", Position.none) [Token.mk_attribute att];
+fun internal att =
+  Token.src ("Pure.attribute", Position.none)
+    [Token.make_value "<attribute>" (Token.Attribute att)];
 
 val _ = Theory.setup
   (setup (Binding.make ("attribute", @{here}))
@@ -203,11 +236,6 @@
 
 (** parsing attributed theorems **)
 
-val thm_sel = Parse.$$$ "(" |-- Parse.list1
- (Parse.nat --| Parse.minus -- Parse.nat >> Facts.FromTo ||
-  Parse.nat --| Parse.minus >> Facts.From ||
-  Parse.nat >> Facts.Single) --| Parse.$$$ ")";
-
 local
 
 val fact_name = Args.internal_fact >> K "<fact>" || Args.name;
@@ -216,7 +244,9 @@
   let
     val get = Proof_Context.get_fact_generic context;
     val get_fact = get o Facts.Fact;
-    fun get_named pos name = get (Facts.Named ((name, pos), NONE));
+    fun get_named is_sel pos name =
+      let val (a, ths) = get (Facts.Named ((name, pos), NONE))
+      in (if is_sel then NONE else a, ths) end;
   in
     Parse.$$$ "[" |-- Args.attribs (check_name_generic context) --| Parse.$$$ "]" >> (fn srcs =>
       let
@@ -226,9 +256,10 @@
     ||
     (Scan.ahead Args.alt_name -- Args.named_fact get_fact
       >> (fn (s, fact) => ("", Facts.Fact s, fact)) ||
-     Scan.ahead (Parse.position fact_name) :|-- (fn (name, pos) =>
-      Args.named_fact (get_named pos) -- Scan.option thm_sel
-        >> (fn (fact, sel) => (name, Facts.Named ((name, pos), sel), fact))))
+     Scan.ahead (Parse.position fact_name -- Scan.option Parse.thm_sel) :|--
+      (fn ((name, pos), sel) =>
+        Args.named_fact (get_named (is_some sel) pos) --| Scan.option Parse.thm_sel
+          >> (fn fact => (name, Facts.Named ((name, pos), sel), fact))))
     -- Args.opt_attribs (check_name_generic context) >> (fn ((name, thmref, fact), srcs) =>
       let
         val ths = Facts.select thmref fact;
@@ -247,6 +278,32 @@
 end;
 
 
+(* read attribute source *)
+
+fun check_attribs ctxt raw_srcs =
+  let
+    val srcs = map (check_src ctxt) raw_srcs;
+    val _ = map (attribute ctxt) srcs;
+  in srcs end;
+
+fun read_attribs ctxt source =
+  let
+    val syms = Symbol_Pos.source_explode source;
+    val lex = #1 (Keyword.get_lexicons ());
+  in
+    (case Token.read lex Parse.attribs (syms, #pos source) of
+      [raw_srcs] => check_attribs ctxt raw_srcs
+    | _ => error ("Malformed attributes" ^ Position.here (#pos source)))
+  end;
+
+
+(* transform fact expressions *)
+
+fun transform_facts phi = map (fn ((a, atts), bs) =>
+  ((Morphism.binding phi a, map (Token.transform_src phi) atts),
+    bs |> map (fn (ths, btts) => (Morphism.fact phi ths, map (Token.transform_src phi) btts))));
+
+
 
 (** partial evaluation -- observing rule/declaration/mixed attributes **)
 
@@ -256,13 +313,13 @@
 
 fun apply_att src (context, th) =
   let
-    val src1 = Args.init_assignable src;
+    val src1 = Token.init_assignable_src src;
     val result = attribute_generic context src1 (context, th);
-    val src2 = Args.closure src1;
+    val src2 = Token.closure_src src1;
   in (src2, result) end;
 
 fun err msg src =
-  let val (name, pos) = Args.name_of_src src
+  let val (name, pos) = Token.name_of_src src
   in error (msg ^ " " ^ quote name ^ Position.here pos) end;
 
 fun eval src ((th, dyn), (decls, context)) =
--- a/src/Pure/Isar/bundle.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/bundle.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,13 +6,13 @@
 
 signature BUNDLE =
 sig
-  type bundle = (thm list * Args.src list) list
+  type bundle = (thm list * Token.src list) list
   val check: Proof.context -> xstring * Position.T -> string
   val get_bundle: Proof.context -> string -> bundle
   val get_bundle_cmd: Proof.context -> xstring * Position.T -> bundle
-  val bundle: binding * (thm list * Args.src list) list ->
+  val bundle: binding * (thm list * Token.src list) list ->
     (binding * typ option * mixfix) list -> local_theory -> local_theory
-  val bundle_cmd: binding * (Facts.ref * Args.src list) list ->
+  val bundle_cmd: binding * (Facts.ref * Token.src list) list ->
     (binding * string option * mixfix) list -> local_theory -> local_theory
   val includes: string list -> Proof.context -> Proof.context
   val includes_cmd: (xstring * Position.T) list -> Proof.context -> Proof.context
@@ -31,10 +31,10 @@
 
 (* maintain bundles *)
 
-type bundle = (thm list * Args.src list) list;
+type bundle = (thm list * Token.src list) list;
 
 fun transform_bundle phi : bundle -> bundle =
-  map (fn (fact, atts) => (Morphism.fact phi fact, map (Args.transform_values phi) atts));
+  map (fn (fact, atts) => (Morphism.fact phi fact, map (Token.transform_src phi) atts));
 
 structure Data = Generic_Data
 (
--- a/src/Pure/Isar/calculation.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/calculation.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -14,10 +14,10 @@
   val sym_del: attribute
   val symmetric: attribute
   val also: thm list option -> bool -> Proof.state -> Proof.state Seq.result Seq.seq
-  val also_cmd: (Facts.ref * Attrib.src list) list option ->
+  val also_cmd: (Facts.ref * Token.src list) list option ->
     bool -> Proof.state -> Proof.state Seq.result Seq.seq
   val finally: thm list option -> bool -> Proof.state -> Proof.state Seq.result Seq.seq
-  val finally_cmd: (Facts.ref * Attrib.src list) list option -> bool ->
+  val finally_cmd: (Facts.ref * Token.src list) list option -> bool ->
     Proof.state -> Proof.state Seq.result Seq.seq
   val moreover: bool -> Proof.state -> Proof.state
   val ultimately: bool -> Proof.state -> Proof.state
@@ -201,7 +201,7 @@
 (* outer syntax *)
 
 val calc_args =
-  Scan.option (@{keyword "("} |-- Parse.!!! ((Parse_Spec.xthms1 --| @{keyword ")"})));
+  Scan.option (@{keyword "("} |-- Parse.!!! ((Parse.xthms1 --| @{keyword ")"})));
 
 val _ =
   Outer_Syntax.command @{command_spec "also"} "combine calculation and current facts"
--- a/src/Pure/Isar/code.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/code.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -49,13 +49,13 @@
   val add_nbe_eqn: thm -> theory -> theory
   val add_abs_eqn: thm -> theory -> theory
   val add_abs_eqn_attribute: attribute
-  val add_abs_eqn_attrib: Attrib.src
+  val add_abs_eqn_attrib: Token.src
   val add_default_eqn: thm -> theory -> theory
   val add_default_eqn_attribute: attribute
-  val add_default_eqn_attrib: Attrib.src
+  val add_default_eqn_attrib: Token.src
   val add_nbe_default_eqn: thm -> theory -> theory
   val add_nbe_default_eqn_attribute: attribute
-  val add_nbe_default_eqn_attrib: Attrib.src
+  val add_nbe_default_eqn_attrib: Token.src
   val del_eqn: thm -> theory -> theory
   val del_eqns: string -> theory -> theory
   val del_exception: string -> theory -> theory
--- a/src/Pure/Isar/element.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/element.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -17,18 +17,15 @@
     Constrains of (string * 'typ) list |
     Assumes of (Attrib.binding * ('term * 'term list) list) list |
     Defines of (Attrib.binding * ('term * 'term list)) list |
-    Notes of string * (Attrib.binding * ('fact * Attrib.src list) list) list
+    Notes of string * (Attrib.binding * ('fact * Token.src list) list) list
   type context = (string, string, Facts.ref) ctxt
   type context_i = (typ, term, thm list) ctxt
   val map_ctxt: {binding: binding -> binding, typ: 'typ -> 'a, term: 'term -> 'b,
-    pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Attrib.src -> Attrib.src} ->
+    pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Token.src -> Token.src} ->
     ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt
-  val map_ctxt_attrib: (Attrib.src -> Attrib.src) ->
+  val map_ctxt_attrib: (Token.src -> Token.src) ->
     ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt
   val transform_ctxt: morphism -> context_i -> context_i
-  val transform_facts: morphism ->
-    (Attrib.binding * (thm list * Args.src list) list) list ->
-    (Attrib.binding * (thm list * Args.src list) list) list
   val pretty_stmt: Proof.context -> statement_i -> Pretty.T list
   val pretty_ctxt: Proof.context -> context_i -> Pretty.T list
   val pretty_statement: Proof.context -> string -> thm -> Pretty.T
@@ -78,7 +75,7 @@
   Constrains of (string * 'typ) list |
   Assumes of (Attrib.binding * ('term * 'term list) list) list |
   Defines of (Attrib.binding * ('term * 'term list)) list |
-  Notes of string * (Attrib.binding * ('fact * Attrib.src list) list) list;
+  Notes of string * (Attrib.binding * ('fact * Token.src list) list) list;
 
 type context = (string, string, Facts.ref) ctxt;
 type context_i = (typ, term, thm list) ctxt;
@@ -103,10 +100,7 @@
   term = Morphism.term phi,
   pattern = Morphism.term phi,
   fact = Morphism.fact phi,
-  attrib = Args.transform_values phi};
-
-fun transform_facts phi facts =
-  Notes ("", facts) |> transform_ctxt phi |> (fn Notes (_, facts') => facts');
+  attrib = Token.transform_src phi};
 
 
 
@@ -270,8 +264,8 @@
 local
 
 val refine_witness =
-  Proof.refine (Method.Basic (K (RAW_METHOD
-    (K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (rtac Drule.protectI))))))))));
+  Proof.refine (Method.Basic (K (NO_CASES o
+    K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (rtac Drule.protectI)))))))));
 
 fun gen_witness_proof proof after_qed wit_propss eq_props =
   let
@@ -509,14 +503,14 @@
 fun activate_i elem ctxt =
   let
     val elem' =
-      (case map_ctxt_attrib Args.init_assignable elem of
+      (case map_ctxt_attrib Token.init_assignable_src elem of
         Defines defs =>
           Defines (defs |> map (fn ((a, atts), (t, ps)) =>
             ((Thm.def_binding_optional (Binding.name (#1 (#1 (Local_Defs.cert_def ctxt t)))) a, atts),
               (t, ps))))
       | e => e);
     val ctxt' = Context.proof_map (init elem') ctxt;
-  in (map_ctxt_attrib Args.closure elem', ctxt') end;
+  in (map_ctxt_attrib Token.closure_src elem', ctxt') end;
 
 fun activate raw_elem ctxt =
   let val elem = raw_elem |> map_ctxt
--- a/src/Pure/Isar/expression.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/expression.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -924,7 +924,7 @@
 fun subscribe_or_activate lthy =
   if Named_Target.is_theory lthy
   then Local_Theory.subscription
-  else Local_Theory.activate;
+  else Locale.activate_fragment;
 
 fun subscribe_locale_only lthy =
   let
@@ -964,7 +964,7 @@
     (K Local_Theory.subscription) expression raw_eqns;
 
 fun ephemeral_interpretation x =
-  gen_local_theory_interpretation cert_interpretation (K Local_Theory.activate) x;
+  gen_local_theory_interpretation cert_interpretation (K Locale.activate_fragment) x;
 
 fun interpretation x =
   gen_local_theory_interpretation cert_interpretation subscribe_or_activate x;
--- a/src/Pure/Isar/generic_target.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/generic_target.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -23,9 +23,9 @@
     bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory
   val notes:
-    (string -> (Attrib.binding * (thm list * Args.src list) list) list ->
-      (Attrib.binding * (thm list * Args.src list) list) list -> local_theory -> local_theory) ->
-    string -> (Attrib.binding * (thm list * Args.src list) list) list -> local_theory ->
+    (string -> (Attrib.binding * (thm list * Token.src list) list) list ->
+      (Attrib.binding * (thm list * Token.src list) list) list -> local_theory -> local_theory) ->
+    string -> (Attrib.binding * (thm list * Token.src list) list) list -> local_theory ->
     (string * thm list) list * local_theory
   val abbrev: (string * bool -> binding * mixfix -> term ->
       term list * term list -> local_theory -> local_theory) ->
@@ -35,8 +35,8 @@
   val theory_foundation: ((binding * typ) * mixfix) * (binding * term) ->
     term list * term list -> local_theory -> (term * thm) * local_theory
   val theory_notes: string ->
-    (Attrib.binding * (thm list * Args.src list) list) list ->
-    (Attrib.binding * (thm list * Args.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
     local_theory -> local_theory
   val theory_declaration: declaration -> local_theory -> local_theory
   val theory_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list ->
@@ -46,8 +46,8 @@
 
   (* locale operations *)
   val locale_notes: string -> string ->
-    (Attrib.binding * (thm list * Args.src list) list) list ->
-    (Attrib.binding * (thm list * Args.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
     local_theory -> local_theory
   val locale_target_declaration: string -> bool -> declaration -> local_theory -> local_theory
   val locale_declaration: string -> {syntax: bool, pervasive: bool} -> declaration ->
@@ -66,7 +66,7 @@
 (** notes **)
 
 fun standard_facts lthy ctxt =
-  Element.transform_facts (Local_Theory.standard_morphism lthy ctxt);
+  Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt);
 
 fun standard_notes pred kind facts lthy =
   Local_Theory.map_contexts (fn level => fn ctxt =>
@@ -358,6 +358,6 @@
 
 fun locale_dependency locale dep_morph mixin export =
   (Local_Theory.raw_theory ooo Locale.add_dependency locale) dep_morph mixin export
-  #> Local_Theory.activate_nonbrittle dep_morph mixin export;
+  #> Locale.activate_fragment_nonbrittle dep_morph mixin export;
 
 end;
--- a/src/Pure/Isar/isar_cmd.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/isar_cmd.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -16,7 +16,7 @@
   val translations: (xstring * string) Syntax.trrule list -> theory -> theory
   val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory
   val oracle: bstring * Position.T -> Symbol_Pos.source -> theory -> theory
-  val add_defs: (bool * bool) * ((binding * string) * Attrib.src list) list -> theory -> theory
+  val add_defs: (bool * bool) * ((binding * string) * Token.src list) list -> theory -> theory
   val declaration: {syntax: bool, pervasive: bool} ->
     Symbol_Pos.source -> local_theory -> local_theory
   val simproc_setup: string * Position.T -> string list -> Symbol_Pos.source ->
@@ -39,14 +39,11 @@
   val thy_deps: Toplevel.transition -> Toplevel.transition
   val locale_deps: Toplevel.transition -> Toplevel.transition
   val class_deps: Toplevel.transition -> Toplevel.transition
-  val thm_deps: (Facts.ref * Attrib.src list) list -> Toplevel.transition -> Toplevel.transition
-  val unused_thms: (string list * string list option) option ->
-    Toplevel.transition -> Toplevel.transition
-  val print_stmts: string list * (Facts.ref * Attrib.src list) list
+  val print_stmts: string list * (Facts.ref * Token.src list) list
     -> Toplevel.transition -> Toplevel.transition
-  val print_thms: string list * (Facts.ref * Attrib.src list) list
+  val print_thms: string list * (Facts.ref * Token.src list) list
     -> Toplevel.transition -> Toplevel.transition
-  val print_prfs: bool -> string list * (Facts.ref * Attrib.src list) list option
+  val print_prfs: bool -> string list * (Facts.ref * Token.src list) list option
     -> Toplevel.transition -> Toplevel.transition
   val print_prop: (string list * string) -> Toplevel.transition -> Toplevel.transition
   val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition
@@ -314,28 +311,6 @@
       |> sort (int_ord o pairself #1) |> map #2;
   in Graph_Display.display_graph gr end);
 
-fun thm_deps args = Toplevel.unknown_theory o Toplevel.keep (fn state =>
-  Thm_Deps.thm_deps (Toplevel.theory_of state)
-    (Attrib.eval_thms (Toplevel.context_of state) args));
-
-
-(* find unused theorems *)
-
-fun unused_thms opt_range = Toplevel.keep (fn state =>
-  let
-    val thy = Toplevel.theory_of state;
-    val ctxt = Toplevel.context_of state;
-    fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]);
-    val get_theory = Context.get_theory thy;
-  in
-    Thm_Deps.unused_thms
-      (case opt_range of
-        NONE => (Theory.parents_of thy, [thy])
-      | SOME (xs, NONE) => (map get_theory xs, [thy])
-      | SOME (xs, SOME ys) => (map get_theory xs, map get_theory ys))
-    |> map pretty_thm |> Pretty.writeln_chunks
-  end);
-
 
 (* print theorems, terms, types etc. *)
 
--- a/src/Pure/Isar/isar_syn.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/isar_syn.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -228,7 +228,7 @@
 
 val _ =
   Outer_Syntax.local_theory' @{command_spec "declare"} "declare theorems"
-    (Parse.and_list1 Parse_Spec.xthms1 -- Parse.for_fixes
+    (Parse.and_list1 Parse.xthms1 -- Parse.for_fixes
       >> (fn (facts, fixes) =>
           #2 oo Specification.theorems_cmd "" [(Attrib.empty_binding, flat facts)] fixes));
 
@@ -329,16 +329,16 @@
     (Parse.ML_source >> Isar_Cmd.local_setup);
 
 val _ =
-  Outer_Syntax.command @{command_spec "attribute_setup"} "define attribute in ML"
+  Outer_Syntax.local_theory @{command_spec "attribute_setup"} "define attribute in ML"
     (Parse.position Parse.name --
         Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
-      >> (fn (name, (txt, cmt)) => Toplevel.theory (Attrib.attribute_setup name txt cmt)));
+      >> (fn (name, (txt, cmt)) => Attrib.attribute_setup name txt cmt));
 
 val _ =
-  Outer_Syntax.command @{command_spec "method_setup"} "define proof method in ML"
+  Outer_Syntax.local_theory @{command_spec "method_setup"} "define proof method in ML"
     (Parse.position Parse.name --
         Parse.!!! (@{keyword "="} |-- Parse.ML_source -- Scan.optional Parse.text "")
-      >> (fn (name, (txt, cmt)) => Toplevel.theory (Method.method_setup name txt cmt)));
+      >> (fn (name, (txt, cmt)) => Method.method_setup name txt cmt));
 
 val _ =
   Outer_Syntax.local_theory @{command_spec "declaration"} "generic ML declaration"
@@ -398,7 +398,7 @@
 
 val _ =
   Outer_Syntax.local_theory @{command_spec "bundle"} "define bundle of declarations"
-    ((Parse.binding --| @{keyword "="}) -- Parse_Spec.xthms1 -- Parse.for_fixes
+    ((Parse.binding --| @{keyword "="}) -- Parse.xthms1 -- Parse.for_fixes
       >> (uncurry Bundle.bundle_cmd));
 
 val _ =
@@ -567,7 +567,7 @@
 
 (* facts *)
 
-val facts = Parse.and_list1 Parse_Spec.xthms1;
+val facts = Parse.and_list1 Parse.xthms1;
 
 val _ =
   Outer_Syntax.command @{command_spec "then"} "forward chaining"
@@ -640,7 +640,7 @@
     ((@{keyword "("} |--
       Parse.!!! (Parse.position Parse.xname -- Scan.repeat (Parse.maybe Parse.binding)
         --| @{keyword ")"}) ||
-      Parse.position Parse.xname >> rpair []) -- Parse_Spec.opt_attribs >> (fn ((c, xs), atts) =>
+      Parse.position Parse.xname >> rpair []) -- Parse.opt_attribs >> (fn ((c, xs), atts) =>
         Toplevel.proof (Proof.invoke_case_cmd (c, xs, atts))));
 
 
@@ -744,7 +744,7 @@
 val _ =
   Outer_Syntax.improper_command @{command_spec "Isabelle.command"} "evaluate nested Isabelle command"
     (props_text :|-- (fn (pos, str) =>
-      (case Outer_Syntax.parse pos str of
+      (case Outer_Syntax.parse (Outer_Syntax.get_syntax ()) pos str of
         [tr] => Scan.succeed (K tr)
       | _ => Scan.fail_with (K (fn () => "exactly one command expected")))
       handle ERROR msg => Scan.fail_with (K (fn () => msg))));
@@ -894,10 +894,6 @@
     (Scan.succeed Isar_Cmd.class_deps);
 
 val _ =
-  Outer_Syntax.improper_command @{command_spec "thm_deps"} "visualize theorem dependencies"
-    (Parse_Spec.xthms1 >> Isar_Cmd.thm_deps);
-
-val _ =
   Outer_Syntax.improper_command @{command_spec "print_binds"}
     "print term bindings of proof context -- Proof General legacy"
     (Scan.succeed (Toplevel.unknown_context o
@@ -924,19 +920,19 @@
 val _ =
   Outer_Syntax.improper_command @{command_spec "print_statement"}
     "print theorems as long statements"
-    (opt_modes -- Parse_Spec.xthms1 >> Isar_Cmd.print_stmts);
+    (opt_modes -- Parse.xthms1 >> Isar_Cmd.print_stmts);
 
 val _ =
   Outer_Syntax.improper_command @{command_spec "thm"} "print theorems"
-    (opt_modes -- Parse_Spec.xthms1 >> Isar_Cmd.print_thms);
+    (opt_modes -- Parse.xthms1 >> Isar_Cmd.print_thms);
 
 val _ =
   Outer_Syntax.improper_command @{command_spec "prf"} "print proof terms of theorems"
-    (opt_modes -- Scan.option Parse_Spec.xthms1 >> Isar_Cmd.print_prfs false);
+    (opt_modes -- Scan.option Parse.xthms1 >> Isar_Cmd.print_prfs false);
 
 val _ =
   Outer_Syntax.improper_command @{command_spec "full_prf"} "print full proof terms of theorems"
-    (opt_modes -- Scan.option Parse_Spec.xthms1 >> Isar_Cmd.print_prfs true);
+    (opt_modes -- Scan.option Parse.xthms1 >> Isar_Cmd.print_prfs true);
 
 val _ =
   Outer_Syntax.improper_command @{command_spec "prop"} "read and print proposition"
@@ -956,11 +952,6 @@
     (Scan.succeed (Toplevel.unknown_theory o
       Toplevel.keep (Code.print_codesetup o Toplevel.theory_of)));
 
-val _ =
-  Outer_Syntax.improper_command @{command_spec "unused_thms"} "find unused theorems"
-    (Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --
-       Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >> Isar_Cmd.unused_thms);
-
 
 
 (** system commands (for interactive mode only) **)
--- a/src/Pure/Isar/local_theory.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/local_theory.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -7,6 +7,11 @@
 type local_theory = Proof.context;
 type generic_theory = Context.generic;
 
+structure Attrib =
+struct
+  type binding = binding * Token.src list;
+end;
+
 signature LOCAL_THEORY =
 sig
   type operations
@@ -14,6 +19,7 @@
   val restore: local_theory -> local_theory
   val level: Proof.context -> int
   val assert_bottom: bool -> local_theory -> local_theory
+  val mark_brittle: local_theory -> local_theory
   val assert_nonbrittle: local_theory -> local_theory
   val open_target: Name_Space.naming -> operations -> (local_theory -> local_theory) ->
     local_theory -> local_theory
@@ -42,9 +48,9 @@
   val define_internal: (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory
   val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory
-  val notes: (Attrib.binding * (thm list * Attrib.src list) list) list ->
+  val notes: (Attrib.binding * (thm list * Token.src list) list) list ->
     local_theory -> (string * thm list) list * local_theory
-  val notes_kind: string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
+  val notes_kind: string -> (Attrib.binding * (thm list * Token.src list) list) list ->
     local_theory -> (string * thm list) list * local_theory
   val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory
@@ -52,16 +58,16 @@
   val subscription: string * morphism -> (morphism * bool) option -> morphism ->
     local_theory -> local_theory
   val pretty: local_theory -> Pretty.T list
+  val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory
   val set_defsort: sort -> local_theory -> local_theory
   val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
+  val generic_alias:
+    (('a Name_Space.table -> 'a Name_Space.table) -> Context.generic -> Context.generic) ->
+    binding -> string -> local_theory -> local_theory
   val class_alias: binding -> class -> local_theory -> local_theory
   val type_alias: binding -> string -> local_theory -> local_theory
   val const_alias: binding -> string -> local_theory -> local_theory
-  val activate: string * morphism -> (morphism * bool) option -> morphism ->
-    local_theory -> local_theory
-  val activate_nonbrittle: string * morphism -> (morphism * bool) option -> morphism ->
-    local_theory -> local_theory
   val init: Name_Space.naming -> operations -> Proof.context -> local_theory
   val exit: local_theory -> Proof.context
   val exit_global: local_theory -> theory
@@ -80,7 +86,7 @@
  {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory ->
     (term * (string * thm)) * local_theory,
   notes: string ->
-    (Attrib.binding * (thm list * Attrib.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
     local_theory -> (string * thm list) list * local_theory,
   abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory ->
     (term * term) * local_theory,
@@ -98,7 +104,8 @@
   target: Proof.context};
 
 fun make_lthy (naming, operations, after_close, brittle, target) : lthy =
-  {naming = naming, operations = operations, after_close = after_close, brittle = brittle, target = target};
+  {naming = naming, operations = operations, after_close = after_close,
+    brittle = brittle, target = target};
 
 
 (* context data *)
@@ -115,7 +122,7 @@
 val bottom_of = List.last o Data.get o assert;
 val top_of = hd o Data.get o assert;
 
-fun map_bottom f =
+fun map_top f =
   assert #>
   Data.map (fn {naming, operations, after_close, brittle, target} :: parents =>
     make_lthy (f (naming, operations, after_close, brittle, target)) :: parents);
@@ -162,14 +169,13 @@
 (* brittle context -- implicit for nested structures *)
 
 fun mark_brittle lthy =
-  if level lthy = 1
-  then map_bottom (fn (naming, operations, after_close, brittle, target) =>
-    (naming, operations, after_close, true, target)) lthy
+  if level lthy = 1 then
+    map_top (fn (naming, operations, after_close, _, target) =>
+      (naming, operations, after_close, true, target)) lthy
   else lthy;
 
 fun assert_nonbrittle lthy =
-  if #brittle (top_of lthy)
-  then error "Brittle local theory context"
+  if #brittle (top_of lthy) then error "Brittle local theory context"
   else lthy;
 
 
@@ -179,7 +185,7 @@
 val full_name = Name_Space.full_name o naming_of;
 
 fun map_naming f =
-  map_bottom (fn (naming, operations, after_close, brittle, target) =>
+  map_top (fn (naming, operations, after_close, brittle, target) =>
     (f naming, operations, after_close, brittle, target));
 
 val conceal = map_naming Name_Space.conceal;
@@ -251,12 +257,12 @@
 
 val operations_of = #operations o top_of;
 
+fun operation f lthy = f (operations_of lthy) lthy;
+fun operation2 f x y = operation (fn ops => f ops x y);
+
 
 (* primitives *)
 
-fun operation f lthy = f (operations_of lthy) lthy;
-fun operation2 f x y = operation (fn ops => f ops x y);
-
 val pretty = operation #pretty;
 val abbrev = operation2 #abbrev;
 val define = operation2 #define false;
@@ -267,11 +273,28 @@
   assert_bottom true #> operation (fn ops => #subscription ops dep_morph mixin export);
 
 
-(* basic derived operations *)
+(* theorems *)
 
 val notes = notes_kind "";
 fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single;
 
+fun add_thms_dynamic (binding, f) lthy =
+  lthy
+  |> background_theory_result (fn thy => thy
+      |> Global_Theory.add_thms_dynamic' (Sign.inherit_naming thy lthy) (binding, f))
+  |-> (fn name =>
+    map_contexts (fn _ => fn ctxt =>
+      Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #>
+    declaration {syntax = false, pervasive = false} (fn phi =>
+      let val binding' = Morphism.binding phi binding in
+        Context.mapping
+          (Global_Theory.alias_fact binding' name)
+          (Proof_Context.fact_alias binding' name)
+      end));
+
+
+(* default sort *)
+
 fun set_defsort S =
   declaration {syntax = true, pervasive = false}
     (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S)));
@@ -298,25 +321,21 @@
 
 (* name space aliases *)
 
-fun alias global_alias local_alias b name =
+fun generic_alias app b name =
+  declaration {syntax = false, pervasive = false} (fn phi => fn context =>
+    let
+      val naming = Name_Space.naming_of context;
+      val b' = Morphism.binding phi b;
+    in app (Name_Space.alias_table naming b' name) context end);
+
+fun syntax_alias global_alias local_alias b name =
   declaration {syntax = true, pervasive = false} (fn phi =>
     let val b' = Morphism.binding phi b
     in Context.mapping (global_alias b' name) (local_alias b' name) end);
 
-val class_alias = alias Sign.class_alias Proof_Context.class_alias;
-val type_alias = alias Sign.type_alias Proof_Context.type_alias;
-val const_alias = alias Sign.const_alias Proof_Context.const_alias;
-
-
-(* activation of locale fragments *)
-
-fun activate_nonbrittle dep_morph mixin export =
-  map_bottom (fn (naming, operations, after_close, brittle, target) =>
-    (naming, operations, after_close, brittle,
-      (Context.proof_map ooo Locale.add_registration) dep_morph mixin export target));
-
-fun activate dep_morph mixin export =
-  mark_brittle #> activate_nonbrittle dep_morph mixin export;
+val class_alias = syntax_alias Sign.class_alias Proof_Context.class_alias;
+val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias;
+val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias;
 
 
 
--- a/src/Pure/Isar/locale.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/locale.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -34,7 +34,7 @@
     term option * term list ->
     thm option * thm option -> thm list ->
     declaration list ->
-    (string * (Attrib.binding * (thm list * Attrib.src list) list) list) list ->
+    (string * (Attrib.binding * (thm list * Token.src list) list) list) list ->
     (string * morphism) list -> theory -> theory
   val intern: theory -> xstring -> string
   val check: theory -> xstring * Position.T -> string
@@ -49,7 +49,7 @@
   val specification_of: theory -> string -> term option * term list
 
   (* Storing results *)
-  val add_thmss: string -> string -> (Attrib.binding * (thm list * Attrib.src list) list) list ->
+  val add_thmss: string -> string -> (Attrib.binding * (thm list * Token.src list) list) list ->
     Proof.context -> Proof.context
   val add_declaration: string -> bool -> declaration -> Proof.context -> Proof.context
 
@@ -70,6 +70,10 @@
   (* Registrations and dependencies *)
   val add_registration: string * morphism -> (morphism * bool) option ->
     morphism -> Context.generic -> Context.generic
+  val activate_fragment: string * morphism -> (morphism * bool) option -> morphism ->
+    local_theory -> local_theory
+  val activate_fragment_nonbrittle: string * morphism -> (morphism * bool) option -> morphism ->
+    local_theory -> local_theory
   val amend_registration: string * morphism -> morphism * bool ->
     morphism -> Context.generic -> Context.generic
   val registrations_of: Context.generic -> string -> (string * morphism) list
@@ -123,7 +127,7 @@
   (** dynamic part **)
   syntax_decls: (declaration * serial) list,
     (* syntax declarations *)
-  notes: ((string * (Attrib.binding * (thm list * Attrib.src list) list) list) * serial) list,
+  notes: ((string * (Attrib.binding * (thm list * Token.src list) list) list) * serial) list,
     (* theorem declarations *)
   dependencies: ((string * (morphism * morphism)) * serial) list
     (* locale dependencies (sublocale relation) in reverse order *),
@@ -514,6 +518,19 @@
   end;
 
 
+(* locale fragments within local theory *)
+
+fun activate_fragment_nonbrittle dep_morph mixin export lthy =
+  let val n = Local_Theory.level lthy in
+    lthy |> Local_Theory.map_contexts (fn level =>
+      level = n - 1 ? Context.proof_map (add_registration dep_morph mixin export))
+  end;
+
+fun activate_fragment dep_morph mixin export =
+  Local_Theory.mark_brittle #> activate_fragment_nonbrittle dep_morph mixin export;
+
+
+
 (*** Dependencies ***)
 
 (* FIXME dead code!?
@@ -561,7 +578,7 @@
           (* Registrations *)
           (fn thy =>
             fold_rev (fn (_, morph) =>
-              snd o Attrib.global_notes kind (Element.transform_facts morph facts))
+              snd o Attrib.global_notes kind (Attrib.transform_facts morph facts))
             (registrations_of (Context.Theory thy) loc) thy));
 
 
--- a/src/Pure/Isar/method.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/method.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,10 +6,7 @@
 
 signature METHOD =
 sig
-  type method
-  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
+  type method = thm list -> cases_tactic
   val METHOD_CASES: (thm list -> cases_tactic) -> method
   val METHOD: (thm list -> tactic) -> method
   val fail: method
@@ -41,20 +38,15 @@
   val erule: Proof.context -> int -> thm list -> method
   val drule: Proof.context -> int -> thm list -> method
   val frule: Proof.context -> int -> thm list -> method
-  val set_tactic: (thm list -> tactic) -> Proof.context -> Proof.context
-  val tactic: Symbol_Pos.source -> Proof.context -> method
-  val raw_tactic: Symbol_Pos.source -> Proof.context -> method
-  type src = Args.src
+  val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic
   type combinator_info
   val no_combinator_info: combinator_info
+  datatype combinator = Then | Orelse | Try | Repeat1 | Select_Goals of int
   datatype text =
-    Source of src |
+    Source of Token.src |
     Basic of Proof.context -> method |
-    Then of combinator_info * text list |
-    Orelse of combinator_info * text list |
-    Try of combinator_info * text |
-    Repeat1 of combinator_info * text |
-    Select_Goals of combinator_info * int * text
+    Combinator of combinator_info * combinator * text list
+  val map_source: (Token.src -> Token.src) -> text -> text
   val primitive_text: (Proof.context -> thm -> thm) -> text
   val succeed_text: text
   val default_text: text
@@ -64,13 +56,21 @@
   val finish_text: text option * bool -> text
   val print_methods: Proof.context -> unit
   val check_name: Proof.context -> xstring * Position.T -> string
-  val method: Proof.context -> src -> Proof.context -> method
-  val method_cmd: Proof.context -> src -> Proof.context -> method
+  val method_syntax: (Proof.context -> method) context_parser ->
+    Token.src -> Proof.context -> method
   val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory
-  val method_setup: bstring * Position.T -> Symbol_Pos.source -> string -> theory -> theory
-  type modifier = (Proof.context -> Proof.context) * attribute
-  val section: modifier parser list -> thm list context_parser
-  val sections: modifier parser list -> thm list list context_parser
+  val local_setup: binding -> (Proof.context -> method) context_parser -> string ->
+    local_theory -> local_theory
+  val method_setup: bstring * Position.T -> Symbol_Pos.source -> string ->
+    local_theory -> local_theory
+  val method: Proof.context -> Token.src -> Proof.context -> method
+  val method_closure: Proof.context -> Token.src -> Token.src
+  val method_cmd: Proof.context -> Token.src -> Proof.context -> method
+  val evaluate: text -> Proof.context -> method
+  type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}
+  val modifier: attribute -> Position.T -> modifier
+  val section: modifier parser list -> declaration context_parser
+  val sections: modifier parser list -> declaration list context_parser
   type text_range = text * Position.range
   val text: text_range option -> text option
   val position: text_range option -> Position.T
@@ -84,20 +84,12 @@
 
 (** proof methods **)
 
-(* datatype method *)
-
-datatype method = Meth of thm list -> cases_tactic;
-
-fun apply meth ctxt = let val Meth m = meth ctxt in m end;
+(* method *)
 
-val RAW_METHOD_CASES = Meth;
-
-fun RAW_METHOD tac = RAW_METHOD_CASES (NO_CASES o tac);
+type method = thm list -> cases_tactic;
 
-fun METHOD_CASES tac = RAW_METHOD_CASES (fn facts =>
-  Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts));
-
-fun METHOD tac = RAW_METHOD (fn facts => ALLGOALS Goal.conjunction_tac THEN tac facts);
+fun METHOD_CASES tac : method = fn facts => Seq.THEN (ALLGOALS Goal.conjunction_tac, tac facts);
+fun METHOD tac : method = fn facts => NO_CASES (ALLGOALS Goal.conjunction_tac THEN tac facts);
 
 val fail = METHOD (K no_tac);
 val succeed = METHOD (K all_tac);
@@ -150,7 +142,7 @@
 fun atomize false ctxt =
       SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac ctxt)
   | atomize true ctxt =
-      RAW_METHOD (K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt)));
+      NO_CASES o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt));
 
 
 (* this -- resolve facts directly *)
@@ -254,74 +246,90 @@
 val try_intros_tac = gen_intros_tac TRYALL;
 
 
-(* ML tactics *)
-
-structure ML_Tactic = Proof_Data
-(
-  type T = thm list -> tactic;
-  fun init _ = undefined;
-);
-
-val set_tactic = ML_Tactic.put;
-
-fun ml_tactic source ctxt =
-  let
-    val ctxt' = ctxt |> Context.proof_map
-      (ML_Context.expression (#pos source)
-        "fun tactic (facts: thm list) : tactic"
-        "Context.map_proof (Method.set_tactic tactic)" (ML_Lex.read_source false source));
-  in Context.setmp_thread_data (SOME (Context.Proof ctxt)) (ML_Tactic.get ctxt') end;
-
-fun tactic source ctxt = METHOD (ml_tactic source ctxt);
-fun raw_tactic source ctxt = RAW_METHOD (ml_tactic source ctxt);
-
-
 
 (** method syntax **)
 
-(* method text *)
+(* context data *)
+
+structure Data = Generic_Data
+(
+  type T =
+    ((Token.src -> Proof.context -> method) * string) Name_Space.table *  (*methods*)
+    (morphism -> thm list -> tactic) option;  (*ML tactic*)
+  val empty : T = (Name_Space.empty_table "method", NONE);
+  val extend = I;
+  fun merge ((tab, tac), (tab', tac')) : T =
+    (Name_Space.merge_tables (tab, tab'), merge_options (tac, tac'));
+);
+
+val get_methods = fst o Data.get;
+val map_methods = Data.map o apfst;
+
+
+(* ML tactic *)
+
+val set_tactic = Data.map o apsnd o K o SOME;
 
-type src = Args.src;
+fun the_tactic context =
+  (case snd (Data.get context) of
+    SOME tac => tac
+  | NONE => raise Fail "Undefined ML tactic");
+
+val parse_tactic =
+  Scan.state :|-- (fn context =>
+    Scan.lift (Args.text_declaration (fn source =>
+      let
+        val context' = context |>
+          ML_Context.expression (#pos source)
+            "fun tactic (morphism: Morphism.morphism) (facts: thm list) : tactic"
+            "Method.set_tactic tactic" (ML_Lex.read_source false source);
+        val tac = the_tactic context';
+      in
+        fn phi =>
+          set_tactic (fn _ => Context.setmp_thread_data (SOME context) (tac phi))
+      end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context))));
+
+
+(* method text *)
 
 datatype combinator_info = Combinator_Info of {keywords: Position.T list};
 fun combinator_info keywords = Combinator_Info {keywords = keywords};
 val no_combinator_info = combinator_info [];
 
+datatype combinator = Then | Orelse | Try | Repeat1 | Select_Goals of int;
+
 datatype text =
-  Source of src |
+  Source of Token.src |
   Basic of Proof.context -> method |
-  Then of combinator_info * text list |
-  Orelse of combinator_info * text list |
-  Try of combinator_info * text |
-  Repeat1 of combinator_info * text |
-  Select_Goals of combinator_info * int * text;
+  Combinator of combinator_info * combinator * text list;
+
+fun map_source f (Source src) = Source (f src)
+  | map_source _ (Basic meth) = Basic meth
+  | map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts);
 
 fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r);
 val succeed_text = Basic (K succeed);
-val default_text = Source (Args.src ("default", Position.none) []);
+val default_text = Source (Token.src ("default", Position.none) []);
 val this_text = Basic (K this);
 val done_text = Basic (K (SIMPLE_METHOD all_tac));
 fun sorry_text int = Basic (fn ctxt => cheating ctxt int);
 
 fun finish_text (NONE, immed) = Basic (finish immed)
-  | finish_text (SOME txt, immed) = Then (no_combinator_info, [txt, Basic (finish immed)]);
+  | finish_text (SOME txt, immed) =
+      Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]);
 
 
 (* method definitions *)
 
-structure Methods = Theory_Data
-(
-  type T = ((src -> Proof.context -> method) * string) Name_Space.table;
-  val empty : T = Name_Space.empty_table "method";
-  val extend = I;
-  fun merge data : T = Name_Space.merge_tables data;
-);
-
-val get_methods = Methods.get o Proof_Context.theory_of;
+fun transfer_methods ctxt =
+  let
+    val meths0 = get_methods (Context.Theory (Proof_Context.theory_of ctxt));
+    val meths' = Name_Space.merge_tables (meths0, get_methods (Context.Proof ctxt));
+  in Context.proof_map (map_methods (K meths')) ctxt end;
 
 fun print_methods ctxt =
   let
-    val meths = get_methods ctxt;
+    val meths = get_methods (Context.Proof ctxt);
     fun prt_meth (name, (_, "")) = Pretty.mark_str name
       | prt_meth (name, (_, comment)) =
           Pretty.block
@@ -331,66 +339,166 @@
     |> Pretty.writeln_chunks
   end;
 
-fun add_method name meth comment thy = thy
-  |> Methods.map (Name_Space.define (Context.Theory thy) true (name, (meth, comment)) #> snd);
+
+(* define *)
+
+fun define_global binding meth comment thy =
+  let
+    val context = Context.Theory thy;
+    val (name, meths') =
+      Name_Space.define context true (binding, (meth, comment)) (get_methods context);
+  in (name, Context.the_theory (map_methods (K meths') context)) end;
+
+fun define binding meth comment =
+  Local_Theory.background_theory_result (define_global binding meth comment)
+  #-> (fn name =>
+    Local_Theory.map_contexts (K transfer_methods)
+    #> Local_Theory.generic_alias map_methods binding name
+    #> pair name);
 
 
 (* check *)
 
-fun check_name ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_methods ctxt);
-fun check_src ctxt src = Args.check_src ctxt (get_methods ctxt) src;
+fun check_name ctxt =
+  let val context = Context.Proof ctxt
+  in #1 o Name_Space.check context (get_methods context) end;
+
+fun check_src ctxt src =
+  Token.check_src ctxt (get_methods (Context.Proof ctxt)) src;
 
 
-(* get methods *)
+(* method setup *)
+
+fun method_syntax scan src ctxt : method =
+  let val (m, ctxt') = Token.syntax scan src ctxt in m ctxt' end;
+
+fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd;
+fun local_setup binding scan comment = define binding (method_syntax scan) comment #> snd;
+
+fun method_setup name source cmt =
+  (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
+    ML_Lex.read_source false source @
+    ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")"))
+  |> ML_Context.expression (#pos source)
+    "val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string"
+    "Context.map_proof (Method.local_setup name scan comment)"
+  |> Context.proof_map;
+
+
+(* prepare methods *)
 
 fun method ctxt =
-  let val table = get_methods ctxt
-  in fn src => #1 (Name_Space.get table (#1 (Args.name_of_src src))) src end;
+  let val table = get_methods (Context.Proof ctxt)
+  in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end;
 
 fun method_closure ctxt0 src0 =
   let
-    val (src1, meth) = check_src ctxt0 src0;
-    val src2 = Args.init_assignable src1;
+    val (src1, _) = check_src ctxt0 src0;
+    val src2 = Token.init_assignable_src src1;
     val ctxt = Context_Position.not_really ctxt0;
-    val _ = Seq.pull (apply (method ctxt src2) ctxt [] (Goal.protect 0 Drule.dummy_thm));
-  in Args.closure src2 end;
+    val _ = Seq.pull (method ctxt src2 ctxt [] (Goal.protect 0 Drule.dummy_thm));
+  in Token.closure_src src2 end;
 
 fun method_cmd ctxt = method ctxt o method_closure ctxt;
 
 
-(* method setup *)
+(* evaluate method text *)
+
+local
+
+fun APPEND_CASES (meth: cases_tactic) (cases, st) =
+  meth st |> Seq.map (fn (cases', st') => (cases @ cases', st'));
+
+fun BYPASS_CASES (tac: tactic) (cases, st) =
+  tac st |> Seq.map (pair cases);
 
-fun setup name scan =
-  add_method name
-    (fn src => fn ctxt => let val (m, ctxt') = Args.syntax scan src ctxt in m ctxt' end);
+val op THEN = Seq.THEN;
+
+fun SELECT_GOALS n method =
+  BYPASS_CASES
+    (ALLGOALS Goal.conjunction_tac THEN PRIMITIVE (Goal.restrict 1 n) THEN Goal.conjunction_tac 1)
+  THEN method
+  THEN BYPASS_CASES (PRIMITIVE (Goal.unrestrict 1));
+
+fun COMBINATOR1 comb [meth] = comb meth
+  | COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument";
 
-fun method_setup name source cmt =
-  Context.theory_map (ML_Context.expression (#pos source)
-    "val (name, scan, comment): binding * (Proof.context -> Proof.method) context_parser * string"
-    "Context.map_theory (Method.setup name scan comment)"
-    (ML_Lex.read Position.none ("(" ^ ML_Syntax.make_binding name ^ ", ") @
-      ML_Lex.read_source false source @
-      ML_Lex.read Position.none (", " ^ ML_Syntax.print_string cmt ^ ")")));
+fun combinator Then = Seq.EVERY
+  | combinator Orelse = Seq.FIRST
+  | combinator Try = COMBINATOR1 Seq.TRY
+  | combinator Repeat1 = COMBINATOR1 Seq.REPEAT1
+  | combinator (Select_Goals n) = COMBINATOR1 (SELECT_GOALS n);
+
+in
+
+fun evaluate text ctxt =
+  let
+    fun eval (Basic meth) = APPEND_CASES o meth ctxt
+      | eval (Source src) = APPEND_CASES o method_cmd ctxt src ctxt
+      | eval (Combinator (_, c, txts)) =
+          let
+            val comb = combinator c;
+            val meths = map eval txts;
+          in fn facts => comb (map (fn meth => meth facts) meths) end;
+    val meth = eval text;
+  in fn facts => fn st => meth facts ([], st) end;
+
+end;
 
 
 
 (** concrete syntax **)
 
-(* sections *)
+(* type modifier *)
+
+type modifier =
+  {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T};
 
-type modifier = (Proof.context -> Proof.context) * attribute;
+fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos};
+
+
+(* sections *)
 
 local
 
-fun thms ss = Scan.repeat (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm) >> flat;
-fun app (f, att) ths context = fold_map (Thm.apply_attribute att) ths (Context.map_proof f context);
+fun sect (modifier : modifier parser) = Scan.depend (fn context =>
+  let
+    val ctxt = Context.proof_of context;
+    fun prep_att src =
+      let
+        val src' = Attrib.check_src ctxt src;
+        val _ = List.app (Token.assign NONE) (Token.args_of_src src');
+      in src' end;
+    val parse_thm =
+      Parse.xthm >> (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs));
+  in
+    Scan.ahead Parse.not_eof -- modifier -- Scan.repeat (Scan.unless modifier parse_thm) >>
+      (fn ((tok, {init, attribute, pos}), thms) =>
+        let
+          val decl =
+            (case Token.get_value tok of
+              SOME (Token.Declaration decl) => decl
+            | _ =>
+                let
+                  val facts =
+                    Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)]
+                    |> map (fn (_, bs) => ((Binding.empty, [Attrib.internal (K attribute)]), bs));
+                  val _ =
+                    Context_Position.report ctxt (Token.pos_of tok)
+                      (Markup.entity Markup.method_modifierN ""
+                        |> Markup.properties (Position.def_properties_of pos));
+                  fun decl phi =
+                    Context.mapping I init #>
+                    Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd;
+                  val _ = Token.assign (SOME (Token.Declaration decl)) tok;
+                in decl end);
+        in (Morphism.form decl context, decl) end)
+  end);
 
 in
 
-fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|--
-  (fn (m, ths) => Scan.succeed (swap (app m ths context))));
-
-fun sections ss = Scan.repeat (section ss);
+val section = sect o Scan.first;
+val sections = Scan.repeat o section;
 
 end;
 
@@ -419,16 +527,8 @@
 
 fun keyword_positions (Source _) = []
   | keyword_positions (Basic _) = []
-  | keyword_positions (Then (Combinator_Info {keywords}, texts)) =
-      keywords @ maps keyword_positions texts
-  | keyword_positions (Orelse (Combinator_Info {keywords}, texts)) =
-      keywords @ maps keyword_positions texts
-  | keyword_positions (Try (Combinator_Info {keywords}, text)) =
-      keywords @ keyword_positions text
-  | keyword_positions (Repeat1 (Combinator_Info {keywords}, text)) =
-      keywords @ keyword_positions text
-  | keyword_positions (Select_Goals (Combinator_Info {keywords}, _, text)) =
-      keywords @ keyword_positions text;
+  | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) =
+      keywords @ maps keyword_positions texts;
 
 in
 
@@ -450,29 +550,29 @@
 local
 
 fun meth4 x =
- (Parse.position Parse.xname >> (fn name => Source (Args.src name [])) ||
+ (Parse.position Parse.xname >> (fn name => Source (Token.src name [])) ||
   Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok =>
-    Source (Args.src ("cartouche", Token.pos_of tok) [tok])) ||
+    Source (Token.src ("cartouche", Token.pos_of tok) [tok])) ||
   Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x
 and meth3 x =
  (meth4 -- Parse.position (Parse.$$$ "?")
-    >> (fn (m, (_, pos)) => Try (combinator_info [pos], m)) ||
+    >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) ||
   meth4 -- Parse.position (Parse.$$$ "+")
-    >> (fn (m, (_, pos)) => Repeat1 (combinator_info [pos], m)) ||
+    >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) ||
   meth4 --
     (Parse.position (Parse.$$$ "[") -- Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]"))
     >> (fn (m, (((_, pos1), n), (_, pos2))) =>
-        Select_Goals (combinator_info [pos1, pos2], n, m)) ||
+        Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) ||
   meth4) x
 and meth2 x =
- (Parse.position Parse.xname -- Parse.args1 is_symid_meth >> (Source o uncurry Args.src) ||
+ (Parse.position Parse.xname -- Parse.args1 is_symid_meth >> (Source o uncurry Token.src) ||
   meth3) x
 and meth1 x =
   (Parse.enum1_positions "," meth2
-    >> (fn ([m], _) => m | (ms, ps) => Then (combinator_info ps, ms))) x
+    >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x
 and meth0 x =
   (Parse.enum1_positions "|" meth1
-    >> (fn ([m], _) => m | (ms, ps) => Orelse (combinator_info ps, ms))) x;
+    >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x;
 
 in
 
@@ -514,9 +614,9 @@
   setup @{binding rotate_tac} (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >>
     (fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i))))
       "rotate assumptions of goal" #>
-  setup @{binding tactic} (Scan.lift Args.text_source_position >> tactic)
+  setup @{binding tactic} (parse_tactic >> (K o METHOD))
     "ML tactic as proof method" #>
-  setup @{binding raw_tactic} (Scan.lift Args.text_source_position >> raw_tactic)
+  setup @{binding raw_tactic} (parse_tactic >> (fn tac => fn _ => NO_CASES o tac))
     "ML tactic as raw proof method");
 
 
@@ -526,8 +626,6 @@
 
 end;
 
-val RAW_METHOD_CASES = Method.RAW_METHOD_CASES;
-val RAW_METHOD = Method.RAW_METHOD;
 val METHOD_CASES = Method.METHOD_CASES;
 val METHOD = Method.METHOD;
 val SIMPLE_METHOD = Method.SIMPLE_METHOD;
--- a/src/Pure/Isar/outer_syntax.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/outer_syntax.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -31,8 +31,10 @@
     (local_theory -> Proof.state) parser -> unit
   val help_outer_syntax: string list -> unit
   val print_outer_syntax: unit -> unit
-  val scan: Position.T -> string -> Token.T list
-  val parse: Position.T -> string -> Toplevel.transition list
+  val scan: Scan.lexicon * Scan.lexicon -> Position.T -> string -> Token.T list
+  val parse: (Scan.lexicon * Scan.lexicon) * outer_syntax ->
+    Position.T -> string -> Toplevel.transition list
+  val parse_spans: Token.T list -> Command_Span.span list
   type isar
   val isar: TextIO.instream -> bool -> isar
   val side_comments: Token.T list -> Token.T list
@@ -256,18 +258,49 @@
 
 (* off-line scanning/parsing *)
 
-fun scan pos str =
+fun scan lexs pos =
+  Source.of_string #>
+  Symbol.source #>
+  Token.source {do_recover = SOME false} (K lexs) pos #>
+  Source.exhaust;
+
+fun parse (lexs, syntax) pos str =
   Source.of_string str
   |> Symbol.source
-  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
+  |> Token.source {do_recover = SOME false} (K lexs) pos
+  |> toplevel_source false NONE (K (lookup_commands syntax))
   |> Source.exhaust;
 
-fun parse pos str =
-  Source.of_string str
-  |> Symbol.source
-  |> Token.source {do_recover = SOME false} Keyword.get_lexicons pos
-  |> toplevel_source false NONE lookup_commands_dynamic
-  |> Source.exhaust;
+
+(* parse spans *)
+
+local
+
+fun ship span =
+  let
+    val kind =
+      if not (null span) andalso Token.is_command (hd span) andalso not (exists Token.is_error span)
+      then Command_Span.Command_Span (Token.content_of (hd span), Token.pos_of (hd span))
+      else if forall Token.is_improper span then Command_Span.Ignored_Span
+      else Command_Span.Malformed_Span;
+  in cons (Command_Span.Span (kind, span)) end;
+
+fun flush (result, content, improper) =
+  result
+  |> not (null content) ? ship (rev content)
+  |> not (null improper) ? ship (rev improper);
+
+fun parse tok (result, content, improper) =
+  if Token.is_command tok then (flush (result, content, improper), [tok], [])
+  else if Token.is_improper tok then (result, content, tok :: improper)
+  else (result, tok :: (improper @ content), []);
+
+in
+
+fun parse_spans toks =
+  fold parse toks ([], [], []) |> flush |> rev;
+
+end;
 
 
 (* interactive source of toplevel transformers *)
--- a/src/Pure/Isar/outer_syntax.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/outer_syntax.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -57,8 +57,8 @@
   def keyword_kind_files(name: String): Option[(String, List[String])] = keywords.get(name)
   def keyword_kind(name: String): Option[String] = keyword_kind_files(name).map(_._1)
 
-  def load(span: List[Token]): Option[List[String]] =
-    keywords.get(Command.name(span)) match {
+  def load_command(name: String): Option[List[String]] =
+    keywords.get(name) match {
       case Some((Keyword.THY_LOAD, exts)) => Some(exts)
       case _ => None
     }
@@ -124,18 +124,16 @@
 
   /* token language */
 
-  def scan(input: Reader[Char]): List[Token] =
+  def scan(input: CharSequence): List[Token] =
   {
+    var in: Reader[Char] = new CharSequenceReader(input)
     Token.Parsers.parseAll(
-        Token.Parsers.rep(Token.Parsers.token(lexicon, is_command)), input) match {
+        Token.Parsers.rep(Token.Parsers.token(lexicon, is_command)), in) match {
       case Token.Parsers.Success(tokens, _) => tokens
-      case _ => error("Unexpected failure of tokenizing input:\n" + input.source.toString)
+      case _ => error("Unexpected failure of tokenizing input:\n" + input.toString)
     }
   }
 
-  def scan(input: CharSequence): List[Token] =
-    scan(new CharSequenceReader(input))
-
   def scan_line(input: CharSequence, context: Scan.Line_Context): (List[Token], Scan.Line_Context) =
   {
     var in: Reader[Char] = new CharSequenceReader(input)
@@ -152,6 +150,47 @@
   }
 
 
+  /* parse_spans */
+
+  def parse_spans(toks: List[Token]): List[Command_Span.Span] =
+  {
+    val result = new mutable.ListBuffer[Command_Span.Span]
+    val content = new mutable.ListBuffer[Token]
+    val improper = new mutable.ListBuffer[Token]
+
+    def ship(span: List[Token])
+    {
+      val kind =
+        if (!span.isEmpty && span.head.is_command && !span.exists(_.is_error)) {
+          val name = span.head.source
+          val pos = Position.Range(Text.Range(0, Symbol.iterator(name).length) + 1)
+          Command_Span.Command_Span(name, pos)
+        }
+        else if (span.forall(_.is_improper)) Command_Span.Ignored_Span
+        else Command_Span.Malformed_Span
+      result += Command_Span.Span(kind, span)
+    }
+
+    def flush()
+    {
+      if (!content.isEmpty) { ship(content.toList); content.clear }
+      if (!improper.isEmpty) { ship(improper.toList); improper.clear }
+    }
+
+    for (tok <- toks) {
+      if (tok.is_command) { flush(); content += tok }
+      else if (tok.is_improper) improper += tok
+      else { content ++= improper; improper.clear; content += tok }
+    }
+    flush()
+
+    result.toList
+  }
+
+  def parse_spans(input: CharSequence): List[Command_Span.Span] =
+    parse_spans(scan(input))
+
+
   /* language context */
 
   def set_language_context(context: Completion.Language_Context): Outer_Syntax =
--- a/src/Pure/Isar/parse.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/parse.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,8 +6,6 @@
 
 signature PARSE =
 sig
-  type 'a parser = Token.T list -> 'a * Token.T list
-  type 'a context_parser = Context.generic * Token.T list -> 'a * (Context.generic * Token.T list)
   val group: (unit -> string) -> (Token.T list -> 'a) -> Token.T list -> 'a
   val !!! : (Token.T list -> 'a) -> Token.T list -> 'a
   val !!!! : (Token.T list -> 'a) -> Token.T list -> 'a
@@ -109,15 +107,16 @@
   val opt_target: (xstring * Position.T) option parser
   val args: Token.T list parser
   val args1: (string -> bool) -> Token.T list parser
+  val attribs: Token.src list parser
+  val opt_attribs: Token.src list parser
+  val thm_sel: Facts.interval list parser
+  val xthm: (Facts.ref * Token.src list) parser
+  val xthms1: (Facts.ref * Token.src list) list parser
 end;
 
 structure Parse: PARSE =
 struct
 
-type 'a parser = Token.T list -> 'a * Token.T list;
-type 'a context_parser = Context.generic * Token.T list -> 'a * (Context.generic * Token.T list);
-
-
 (** error handling **)
 
 (* group atomic parsers (no cuts!) *)
@@ -439,8 +438,27 @@
 
 end;
 
+
+(* attributes *)
+
+val attrib = position liberal_name -- !!! args >> uncurry Token.src;
+val attribs = $$$ "[" |-- list attrib --| $$$ "]";
+val opt_attribs = Scan.optional attribs [];
+
+
+(* theorem references *)
+
+val thm_sel = $$$ "(" |-- list1
+ (nat --| minus -- nat >> Facts.FromTo ||
+  nat --| minus >> Facts.From ||
+  nat >> Facts.Single) --| $$$ ")";
+
+val xthm =
+  $$$ "[" |-- attribs --| $$$ "]" >> pair (Facts.named "") ||
+  (literal_fact >> Facts.Fact ||
+    position xname -- Scan.option thm_sel >> Facts.Named) -- opt_attribs;
+
+val xthms1 = Scan.repeat1 xthm;
+
 end;
 
-type 'a parser = 'a Parse.parser;
-type 'a context_parser = 'a Parse.context_parser;
-
--- a/src/Pure/Isar/parse_spec.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/parse_spec.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,17 +6,13 @@
 
 signature PARSE_SPEC =
 sig
-  val attribs: Attrib.src list parser
-  val opt_attribs: Attrib.src list parser
   val thm_name: string -> Attrib.binding parser
   val opt_thm_name: string -> Attrib.binding parser
   val spec: (Attrib.binding * string) parser
   val specs: (Attrib.binding * string list) parser
   val alt_specs: (Attrib.binding * string) list parser
   val where_alt_specs: (Attrib.binding * string) list parser
-  val xthm: (Facts.ref * Attrib.src list) parser
-  val xthms1: (Facts.ref * Attrib.src list) list parser
-  val name_facts: (Attrib.binding * (Facts.ref * Attrib.src list) list) list parser
+  val name_facts: (Attrib.binding * (Facts.ref * Token.src list) list) list parser
   val constdecl: (binding * string option * mixfix) parser
   val constdef: ((binding * string option * mixfix) option * (Attrib.binding * string)) parser
   val includes: (xstring * Position.T) list parser
@@ -37,14 +33,11 @@
 
 (* theorem specifications *)
 
-val attrib = Parse.position Parse.liberal_name -- Parse.!!! Parse.args >> uncurry Args.src;
-val attribs = Parse.$$$ "[" |-- Parse.list attrib --| Parse.$$$ "]";
-val opt_attribs = Scan.optional attribs [];
-
-fun thm_name s = Parse.binding -- opt_attribs --| Parse.$$$ s;
+fun thm_name s = Parse.binding -- Parse.opt_attribs --| Parse.$$$ s;
 
 fun opt_thm_name s =
-  Scan.optional ((Parse.binding -- opt_attribs || attribs >> pair Binding.empty) --| Parse.$$$ s)
+  Scan.optional
+    ((Parse.binding -- Parse.opt_attribs || Parse.attribs >> pair Binding.empty) --| Parse.$$$ s)
     Attrib.empty_binding;
 
 val spec = opt_thm_name ":" -- Parse.prop;
@@ -56,14 +49,7 @@
 
 val where_alt_specs = Parse.where_ |-- Parse.!!! alt_specs;
 
-val xthm =
-  Parse.$$$ "[" |-- attribs --| Parse.$$$ "]" >> pair (Facts.named "") ||
-  (Parse.literal_fact >> Facts.Fact ||
-    Parse.position Parse.xname -- Scan.option Attrib.thm_sel >> Facts.Named) -- opt_attribs;
-
-val xthms1 = Scan.repeat1 xthm;
-
-val name_facts = Parse.and_list1 (opt_thm_name "=" -- xthms1);
+val name_facts = Parse.and_list1 (opt_thm_name "=" -- Parse.xthms1);
 
 
 (* basic constant specifications *)
@@ -103,7 +89,7 @@
     >> Element.Assumes ||
   Parse.$$$ "defines" |-- Parse.!!! (Parse.and_list1 (opt_thm_name ":" -- Parse.propp))
     >> Element.Defines ||
-  Parse.$$$ "notes" |-- Parse.!!! (Parse.and_list1 (opt_thm_name "=" -- xthms1))
+  Parse.$$$ "notes" |-- Parse.!!! (Parse.and_list1 (opt_thm_name "=" -- Parse.xthms1))
     >> (curry Element.Notes "");
 
 fun plus1_unless test scan =
--- a/src/Pure/Isar/proof.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/proof.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -61,18 +61,18 @@
   val chain: state -> state
   val chain_facts: thm list -> state -> state
   val note_thmss: (Thm.binding * (thm list * attribute list) list) list -> state -> state
-  val note_thmss_cmd: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
+  val note_thmss_cmd: (Attrib.binding * (Facts.ref * Token.src list) list) list -> state -> state
   val from_thmss: ((thm list * attribute list) list) list -> state -> state
-  val from_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val from_thmss_cmd: ((Facts.ref * Token.src list) list) list -> state -> state
   val with_thmss: ((thm list * attribute list) list) list -> state -> state
-  val with_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val with_thmss_cmd: ((Facts.ref * Token.src list) list) list -> state -> state
   val using: ((thm list * attribute list) list) list -> state -> state
-  val using_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val using_cmd: ((Facts.ref * Token.src list) list) list -> state -> state
   val unfolding: ((thm list * attribute list) list) list -> state -> state
-  val unfolding_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
+  val unfolding_cmd: ((Facts.ref * Token.src list) list) list -> state -> state
   val invoke_case: (string * Position.T) * binding option list * attribute list ->
     state -> state
-  val invoke_case_cmd: (string * Position.T) * binding option list * Attrib.src list ->
+  val invoke_case_cmd: (string * Position.T) * binding option list * Token.src list ->
     state -> state
   val begin_block: state -> state
   val next_block: state -> state
@@ -395,48 +395,24 @@
   Rule_Cases.make_common
     (Thm.theory_of_thm st, Thm.prop_of st) (map (rpair [] o rpair []) (goals st));
 
-fun apply_method current_context method 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 method ctxt using goal |> Seq.map (fn (meth_cases, goal') =>
+fun apply_method text ctxt state =
+  #2 (#2 (find_goal state))
+  |> (fn {statement, messages = _, using, goal, before_qed, after_qed} =>
+    Method.evaluate text ctxt using goal
+    |> Seq.map (fn (meth_cases, goal') =>
       state
       |> map_goal
           (Proof_Context.update_cases false (no_goal_cases goal @ goal_cases goal') #>
            Proof_Context.update_cases true meth_cases)
-          (K (statement, [], using, goal', before_qed, after_qed)) I)
-  end;
-
-fun select_goals n meth state =
-  ALLGOALS Goal.conjunction_tac (#2 (#2 (get_goal state)))
-  |> Seq.maps (fn goal =>
-    state
-    |> Seq.lift provide_goal ((PRIMITIVE (Goal.restrict 1 n) THEN Goal.conjunction_tac 1) goal)
-    |> Seq.maps meth
-    |> Seq.maps (fn state' => state'
-      |> Seq.lift provide_goal (PRIMITIVE (Goal.unrestrict 1) (#2 (#2 (get_goal state')))))
-    |> Seq.maps (apply_method true (K Method.succeed)));
-
-fun apply_text current_context text state =
-  let
-    val ctxt = context_of state;
-
-    fun eval (Method.Basic m) = apply_method current_context m
-      | eval (Method.Source src) = apply_method current_context (Method.method_cmd ctxt 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)
-      | eval (Method.Repeat1 (_, txt)) = Seq.REPEAT1 (eval txt)
-      | eval (Method.Select_Goals (_, n, txt)) = select_goals n (eval txt);
-  in eval text state end;
+          (K (statement, [], using, goal', before_qed, after_qed)) I));
 
 in
 
-val refine = apply_text true;
-val refine_end = apply_text false;
-fun refine_insert ths = Seq.hd o refine (Method.Basic (K (Method.insert ths)));
+fun refine text state = apply_method text (context_of state) state;
+fun refine_end text state = apply_method text (#1 (find_goal state)) state;
+
+fun refine_insert ths =
+  Seq.hd o refine (Method.Basic (K (Method.insert ths)));
 
 end;
 
@@ -892,9 +868,9 @@
   in map (Logic.mk_term o Var) vars end;
 
 fun refine_terms n =
-  refine (Method.Basic (K (RAW_METHOD
-    (K (HEADGOAL (PRECISE_CONJUNCTS n
-      (HEADGOAL (CONJUNCTS (ALLGOALS (rtac Drule.termI))))))))))
+  refine (Method.Basic (K (NO_CASES o
+    K (HEADGOAL (PRECISE_CONJUNCTS n
+      (HEADGOAL (CONJUNCTS (ALLGOALS (rtac Drule.termI)))))))))
   #> Seq.hd;
 
 in
--- a/src/Pure/Isar/proof_context.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/proof_context.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -50,6 +50,7 @@
   val markup_const: Proof.context -> string -> string
   val pretty_const: Proof.context -> string -> Pretty.T
   val transfer: theory -> Proof.context -> Proof.context
+  val transfer_facts: theory -> Proof.context -> Proof.context
   val background_theory: (theory -> theory) -> Proof.context -> Proof.context
   val background_theory_result: (theory -> 'a * theory) -> Proof.context -> 'a * Proof.context
   val facts_of: Proof.context -> Facts.T
@@ -120,7 +121,7 @@
     (term list list * (Proof.context -> Proof.context)) * Proof.context
   val fact_tac: Proof.context -> thm list -> int -> tactic
   val some_fact_tac: Proof.context -> int -> tactic
-  val get_fact_generic: Context.generic -> Facts.ref -> thm list
+  val get_fact_generic: Context.generic -> Facts.ref -> string option * thm list
   val get_fact: Proof.context -> Facts.ref -> thm list
   val get_fact_single: Proof.context -> Facts.ref -> thm
   val get_thms: Proof.context -> xstring -> thm list
@@ -154,6 +155,7 @@
   val class_alias: binding -> class -> Proof.context -> Proof.context
   val type_alias: binding -> string -> Proof.context -> Proof.context
   val const_alias: binding -> string -> Proof.context -> Proof.context
+  val fact_alias: binding -> string -> Proof.context -> Proof.context
   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
   val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context
   val revert_abbrev: string -> string -> Proof.context -> Proof.context
@@ -322,16 +324,19 @@
   map_tsig (fn tsig as (local_tsig, global_tsig) =>
     let val thy_tsig = Sign.tsig_of thy in
       if Type.eq_tsig (thy_tsig, global_tsig) then tsig
-      else (Type.merge_tsig (Context.pretty ctxt) (local_tsig, thy_tsig), thy_tsig)
+      else (Type.merge_tsig (Context.pretty ctxt) (local_tsig, thy_tsig), thy_tsig)  (*historic merge order*)
     end) |>
   map_consts (fn consts as (local_consts, global_consts) =>
     let val thy_consts = Sign.consts_of thy in
       if Consts.eq_consts (thy_consts, global_consts) then consts
-      else (Consts.merge (local_consts, thy_consts), thy_consts)
+      else (Consts.merge (local_consts, thy_consts), thy_consts)  (*historic merge order*)
     end);
 
 fun transfer thy = Context.raw_transfer thy #> transfer_syntax thy;
 
+fun transfer_facts thy =
+  map_facts (fn local_facts => Facts.merge (Global_Theory.facts_of thy, local_facts));
+
 fun background_theory f ctxt = transfer (f (theory_of ctxt)) ctxt;
 
 fun background_theory_result f ctxt =
@@ -943,22 +948,27 @@
             [thm] => thm
           | [] => err "Failed to retrieve literal fact"
           | _ => err "Ambiguous specification of literal fact");
-      in pick ("", Position.none) [thm] end
-  | retrieve pick context (Facts.Named ((xname, pos), ivs)) =
+      in pick true ("", Position.none) [thm] end
+  | retrieve pick context (Facts.Named ((xname, pos), sel)) =
       let
         val thy = Context.theory_of context;
-        val (name, thms) =
+        fun immediate thm = {name = xname, static = true, thms = [Thm.transfer thy thm]};
+        val {name, static, thms} =
           (case xname of
-            "" => (xname, [Thm.transfer thy Drule.dummy_thm])
-          | "_" => (xname, [Thm.transfer thy Drule.asm_rl])
+            "" => immediate Drule.dummy_thm
+          | "_" => immediate Drule.asm_rl
           | _ => retrieve_generic context (xname, pos));
-      in pick (name, pos) (Facts.select (Facts.Named ((name, pos), ivs)) thms) end;
+        val thms' = Facts.select (Facts.Named ((name, pos), sel)) thms;
+      in pick (static orelse is_some sel) (name, pos) thms' end;
 
 in
 
-val get_fact_generic = retrieve (K I);
-val get_fact = retrieve (K I) o Context.Proof;
-val get_fact_single = retrieve Facts.the_single o Context.Proof;
+val get_fact_generic =
+  retrieve (fn static => fn (name, _) => fn thms =>
+    (if static then NONE else SOME name, thms));
+
+val get_fact = retrieve (K (K I)) o Context.Proof;
+val get_fact_single = retrieve (K Facts.the_single) o Context.Proof;
 
 fun get_thms ctxt = get_fact ctxt o Facts.named;
 fun get_thm ctxt = get_fact_single ctxt o Facts.named;
@@ -1084,6 +1094,7 @@
 fun class_alias b c ctxt = (map_tsig o apfst) (Type.class_alias (naming_of ctxt) b c) ctxt;
 fun type_alias b c ctxt = (map_tsig o apfst) (Type.type_alias (naming_of ctxt) b c) ctxt;
 fun const_alias b c ctxt = (map_consts o apfst) (Consts.alias (naming_of ctxt) b c) ctxt;
+fun fact_alias b c ctxt = map_facts (Facts.alias (naming_of ctxt) b c) ctxt;
 
 
 (* local constants *)
@@ -1173,11 +1184,8 @@
       (Name_Space.del_table name cases, index)
   | update_case context is_proper (name, SOME c) (cases, index) =
       let
-        val binding =
-          Binding.make (name, Position.none)
-          |> not is_proper ? Binding.conceal;
-        val (_, cases') = cases
-          |> Name_Space.define context false (binding, ((c, is_proper), index));
+        val binding = Binding.name name |> not is_proper ? Binding.conceal;
+        val (_, cases') = Name_Space.define context false (binding, ((c, is_proper), index)) cases;
         val index' = index + 1;
       in (cases', index') end;
 
--- a/src/Pure/Isar/rule_cases.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/rule_cases.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -330,7 +330,9 @@
 
 fun get_case_concl name (a, b) =
   if a = case_concl_tagN then
-    (case explode_args b of c :: cs => if c = name then SOME cs else NONE)
+    (case explode_args b of
+      c :: cs => if c = name then SOME cs else NONE
+    | [] => NONE)
   else NONE;
 
 fun add_case_concl (name, cs) = Thm.map_tags (fn tags =>
--- a/src/Pure/Isar/specification.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/specification.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -51,11 +51,11 @@
   val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory
   val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory
   val theorems: string ->
-    (Attrib.binding * (thm list * Attrib.src list) list) list ->
+    (Attrib.binding * (thm list * Token.src list) list) list ->
     (binding * typ option * mixfix) list ->
     bool -> local_theory -> (string * thm list) list * local_theory
   val theorems_cmd: string ->
-    (Attrib.binding * (Facts.ref * Attrib.src list) list) list ->
+    (Attrib.binding * (Facts.ref * Token.src list) list) list ->
     (binding * string option * mixfix) list ->
     bool -> local_theory -> (string * thm list) list * local_theory
   val theorem: string -> Method.text option ->
@@ -301,7 +301,7 @@
 
     val facts' = facts
       |> Attrib.partial_evaluation ctxt'
-      |> Element.transform_facts (Proof_Context.export_morphism ctxt' lthy);
+      |> Attrib.transform_facts (Proof_Context.export_morphism ctxt' lthy);
     val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts';
     val _ = Proof_Display.print_results int (Position.thread_data ()) lthy' ((kind, ""), res);
   in (res, lthy') end;
--- a/src/Pure/Isar/token.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Isar/token.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -10,14 +10,28 @@
     Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
     Nat | Float | String | AltString | Verbatim | Cartouche | Space | Comment | InternalValue |
     Error of string | Sync | EOF
+  val str_of_kind: kind -> string
   type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T}
+  type T
+  type src
   datatype value =
-    Literal of bool * Markup.T | Text of string | Typ of typ | Term of term | Fact of thm list |
-    Attribute of morphism -> attribute | Files of file Exn.result list
-  type T
-  val str_of_kind: kind -> string
+    Source of src |
+    Literal of bool * Markup.T |
+    Name of string * morphism |
+    Typ of typ |
+    Term of term |
+    Fact of string option * thm list |
+    Attribute of morphism -> attribute |
+    Declaration of declaration |
+    Files of file Exn.result list
+  val name0: string -> value
   val pos_of: T -> Position.T
   val range_of: T list -> Position.range
+  val src: xstring * Position.T -> T list -> src
+  val name_of_src: src -> string * Position.T
+  val args_of_src: src -> T list
+  val range_of_src: src -> Position.T
+  val check_src: Proof.context -> 'a Name_Space.table -> src -> src * 'a
   val eof: T
   val is_eof: T -> bool
   val not_eof: T -> bool
@@ -47,21 +61,24 @@
   val report: T -> Position.report_text
   val markup: T -> Markup.T
   val unparse: T -> string
+  val unparse_src: src -> string list
   val print: T -> string
   val text_of: T -> string * string
   val get_files: T -> file Exn.result list
   val put_files: file Exn.result list -> T -> T
+  val make_value: string -> value -> T
   val get_value: T -> value option
   val map_value: (value -> value) -> T -> T
   val reports_of_value: T -> Position.report list
-  val mk_text: string -> T
-  val mk_typ: typ -> T
-  val mk_term: term -> T
-  val mk_fact: thm list -> T
-  val mk_attribute: (morphism -> attribute) -> T
+  val transform: morphism -> T -> T
+  val transform_src: morphism -> src -> src
   val init_assignable: T -> T
+  val init_assignable_src: src -> src
   val assign: value option -> T -> unit
   val closure: T -> T
+  val closure_src: src -> src
+  val pretty_value: Proof.context -> T -> Pretty.T
+  val pretty_src: Proof.context -> src -> Pretty.T
   val ident_or_symbolic: string -> bool
   val source_proper: (T, 'a) Source.source -> (T, (T, 'a) Source.source) Source.source
   val source': {do_recover: bool option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
@@ -69,7 +86,12 @@
   val source: {do_recover: bool option} -> (unit -> Scan.lexicon * Scan.lexicon) ->
     Position.T -> (Symbol.symbol, 'a) Source.source -> (T,
       (Symbol_Pos.T, Position.T * (Symbol.symbol, 'a) Source.source) Source.source) Source.source
-  val read_antiq: Scan.lexicon -> (T list -> 'a * T list) -> Symbol_Pos.T list * Position.T -> 'a
+  type 'a parser = T list -> 'a * T list
+  type 'a context_parser = Context.generic * T list -> 'a * (Context.generic * T list)
+  val read: Scan.lexicon -> 'a parser -> Symbol_Pos.T list * Position.T -> 'a list
+  val read_antiq: Scan.lexicon -> 'a parser -> Symbol_Pos.T list * Position.T -> 'a
+  val syntax_generic: 'a context_parser -> src -> Context.generic -> 'a * Context.generic
+  val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context
 end;
 
 structure Token: TOKEN =
@@ -77,39 +99,13 @@
 
 (** tokens **)
 
-(* token values *)
-
-(*The value slot assigns an (optional) internal value to a token,
-  usually as a side-effect of special scanner setup (see also
-  args.ML).  Note that an assignable ref designates an intermediate
-  state of internalization -- it is NOT meant to persist.*)
-
-type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T};
-
-datatype value =
-  Literal of bool * Markup.T |
-  Text of string |
-  Typ of typ |
-  Term of term |
-  Fact of thm list |
-  Attribute of morphism -> attribute |
-  Files of file Exn.result list;
-
-datatype slot =
-  Slot |
-  Value of value option |
-  Assignable of value option Unsynchronized.ref;
-
-
-(* datatype token *)
+(* token kind *)
 
 datatype kind =
   Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
   Nat | Float | String | AltString | Verbatim | Cartouche | Space | Comment | InternalValue |
   Error of string | Sync | EOF;
 
-datatype T = Token of (Symbol_Pos.text * Position.range) * (kind * string) * slot;
-
 val str_of_kind =
  fn Command => "command"
   | Keyword => "keyword"
@@ -135,6 +131,42 @@
 val delimited_kind = member (op =) [String, AltString, Verbatim, Cartouche, Comment];
 
 
+(* datatype token *)
+
+(*The value slot assigns an (optional) internal value to a token,
+  usually as a side-effect of special scanner setup (see also
+  args.ML).  Note that an assignable ref designates an intermediate
+  state of internalization -- it is NOT meant to persist.*)
+
+type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T};
+
+datatype T = Token of (Symbol_Pos.text * Position.range) * (kind * string) * slot
+
+and src =
+  Src of
+   {name: string * Position.T,
+    args: T list,
+    output_info: (string * Markup.T) option}
+
+and slot =
+  Slot |
+  Value of value option |
+  Assignable of value option Unsynchronized.ref
+
+and value =
+  Source of src |
+  Literal of bool * Markup.T |
+  Name of string * morphism |
+  Typ of typ |
+  Term of term |
+  Fact of string option * thm list |  (*optional name for dynamic fact, i.e. fact "variable"*)
+  Attribute of morphism -> attribute |
+  Declaration of declaration |
+  Files of file Exn.result list;
+
+fun name0 a = Name (a, Morphism.identity);
+
+
 (* position *)
 
 fun pos_of (Token ((_, (pos, _)), _, _)) = pos;
@@ -146,6 +178,29 @@
   | range_of [] = Position.no_range;
 
 
+(* src *)
+
+fun src name args = Src {name = name, args = args, output_info = NONE};
+
+fun map_args f (Src {name, args, output_info}) =
+  Src {name = name, args = map f args, output_info = output_info};
+
+fun name_of_src (Src {name, ...}) = name;
+fun args_of_src (Src {args, ...}) = args;
+
+fun range_of_src (Src {name = (_, pos), args, ...}) =
+  if null args then pos
+  else Position.set_range (pos, #2 (range_of args));
+
+fun check_src ctxt table (Src {name = (xname, pos), args, output_info = _}) =
+  let
+    val (name, x) = Name_Space.check (Context.Proof ctxt) table (xname, pos);
+    val space = Name_Space.space_of_table table;
+    val kind = Name_Space.kind_of space;
+    val markup = Name_Space.markup space name;
+  in (Src {name = (name, pos), args = args, output_info = SOME (kind, markup)}, x) end;
+
+
 (* control tokens *)
 
 fun mk_eof pos = Token (("", (pos, Position.none)), (EOF, ""), Slot);
@@ -288,6 +343,8 @@
   | EOF => ""
   | _ => x);
 
+fun unparse_src (Src {args, ...}) = map unparse args;
+
 fun print tok = Markup.markup (markup tok) (unparse tok);
 
 fun text_of tok =
@@ -320,6 +377,9 @@
 
 (* access values *)
 
+fun make_value name v =
+  Token ((name, Position.no_range), (InternalValue, name), Value (SOME v));
+
 fun get_value (Token (_, _, Value v)) = v
   | get_value _ = NONE;
 
@@ -340,15 +400,21 @@
   | _ => []);
 
 
-(* make values *)
-
-fun mk_value k v = Token ((k, Position.no_range), (InternalValue, k), Value (SOME v));
+(* transform *)
 
-val mk_text = mk_value "<text>" o Text;
-val mk_typ = mk_value "<typ>" o Typ;
-val mk_term = mk_value "<term>" o Term;
-val mk_fact = mk_value "<fact>" o Fact;
-val mk_attribute = mk_value "<attribute>" o Attribute;
+fun transform phi =
+  map_value (fn v =>
+    (case v of
+      Source src => Source (transform_src phi src)
+    | Literal _ => v
+    | Name (a, psi) => Name (a, psi $> phi)
+    | Typ T => Typ (Morphism.typ phi T)
+    | Term t => Term (Morphism.term phi t)
+    | Fact (a, ths) => Fact (a, Morphism.fact phi ths)
+    | Attribute att => Attribute (Morphism.transform phi att)
+    | Declaration decl => Declaration (Morphism.transform phi decl)
+    | Files _ => v))
+and transform_src phi = map_args (transform phi);
 
 
 (* static binding *)
@@ -358,6 +424,8 @@
   | init_assignable (tok as Token (_, _, Assignable r)) = (r := NONE; tok)
   | init_assignable tok = tok;
 
+val init_assignable_src = map_args init_assignable;
+
 (*2nd stage: assign values as side-effect of scanning*)
 fun assign v (Token (_, _, Assignable r)) = r := v
   | assign _ _ = ();
@@ -366,6 +434,32 @@
 fun closure (Token (x, y, Assignable (Unsynchronized.ref v))) = Token (x, y, Value v)
   | closure tok = tok;
 
+val closure_src = map_args closure;
+
+
+(* pretty *)
+
+fun pretty_value ctxt tok =
+  (case get_value tok of
+    SOME (Literal markup) =>
+      let val x = content_of tok
+      in Pretty.mark_str (keyword_markup markup x, x) end
+  | SOME (Name (a, _)) => Pretty.str (quote a)
+  | SOME (Typ T) => Syntax.pretty_typ ctxt T
+  | SOME (Term t) => Syntax.pretty_term ctxt t
+  | SOME (Fact (_, ths)) =>
+      Pretty.enclose "(" ")" (Pretty.breaks (map (Pretty.backquote o Display.pretty_thm ctxt) ths))
+  | _ => Pretty.mark_str (markup tok, unparse tok));
+
+fun pretty_src ctxt src =
+  let
+    val Src {name = (name, _), args, output_info} = src;
+    val prt_name =
+      (case output_info of
+        NONE => Pretty.str name
+      | SOME (_, markup) => Pretty.mark_str (markup, name));
+    val prt_arg = pretty_value ctxt;
+  in Pretty.block (Pretty.breaks (prt_name :: map prt_arg args)) end;
 
 
 (** scanners **)
@@ -492,19 +586,64 @@
 end;
 
 
-(* read_antiq *)
+
+(** parsers **)
+
+type 'a parser = T list -> 'a * T list;
+type 'a context_parser = Context.generic * T list -> 'a * (Context.generic * T list);
+
+
+(* read source *)
+
+fun read lex scan (syms, pos) =
+  Source.of_list syms
+  |> source' {do_recover = NONE} (K (lex, Scan.empty_lexicon))
+  |> source_proper
+  |> Source.source stopper (Scan.error (Scan.bulk scan)) NONE
+  |> Source.exhaust;
 
 fun read_antiq lex scan (syms, pos) =
   let
-    fun err msg = cat_error msg ("Malformed antiquotation" ^ Position.here pos ^ ":\n" ^
-      "@{" ^ Symbol_Pos.content syms ^ "}");
+    fun err msg =
+      cat_error msg ("Malformed antiquotation" ^ Position.here pos ^ ":\n" ^
+        "@{" ^ Symbol_Pos.content syms ^ "}");
+    val res = read lex scan (syms, pos) handle ERROR msg => err msg;
+  in (case res of [x] => x | _ => err "") end;
+
+
+(* wrapped syntax *)
 
-    val res =
-      Source.of_list syms
-      |> source' {do_recover = NONE} (K (lex, Scan.empty_lexicon))
-      |> source_proper
-      |> Source.source stopper (Scan.error (Scan.bulk scan)) NONE
-      |> Source.exhaust;
-  in (case res of [x] => x | _ => err "") handle ERROR msg => err msg end;
+fun syntax_generic scan (Src {name = (name, pos), args = args0, output_info}) context =
+  let
+    val args1 = map init_assignable args0;
+    fun reported_text () =
+      if Context_Position.is_visible_generic context then
+        ((pos, Markup.operator) :: maps (reports_of_value o closure) args1)
+        |> map (fn (p, m) => Position.reported_text p m "")
+      else [];
+  in
+    (case Scan.error (Scan.finite' stopper (Scan.option scan)) (context, args1) of
+      (SOME x, (context', [])) =>
+        let val _ = Output.report (reported_text ())
+        in (x, context') end
+    | (_, (_, args2)) =>
+        let
+          val print_name =
+            (case output_info of
+              NONE => quote name
+            | SOME (kind, markup) => plain_words kind ^ " " ^ quote (Markup.markup markup name));
+          val print_args =
+            if null args2 then "" else ":\n  " ^ space_implode " " (map print args2);
+        in
+          error ("Bad arguments for " ^ print_name ^ Position.here pos ^ print_args ^
+            Markup.markup_report (implode (reported_text ())))
+        end)
+  end;
+
+fun syntax scan src = apsnd Context.the_proof o syntax_generic scan src o Context.Proof;
 
 end;
+
+type 'a parser = 'a Token.parser;
+type 'a context_parser = 'a Token.context_parser;
+
--- a/src/Pure/ML/ml_antiquotation.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/ML/ml_antiquotation.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -8,7 +8,7 @@
 sig
   val variant: string -> Proof.context -> string * Proof.context
   val declaration: binding -> 'a context_parser ->
-    (Args.src -> 'a -> Proof.context -> ML_Context.decl * Proof.context) ->
+    (Token.src -> 'a -> Proof.context -> ML_Context.decl * Proof.context) ->
     theory -> theory
   val inline: binding -> string context_parser -> theory -> theory
   val value: binding -> string context_parser -> theory -> theory
@@ -40,7 +40,7 @@
 fun declaration name scan body =
   ML_Context.add_antiquotation name
     (fn src => fn orig_ctxt =>
-      let val (x, _) = Args.syntax scan src orig_ctxt
+      let val (x, _) = Token.syntax scan src orig_ctxt
       in body src x orig_ctxt end);
 
 fun inline name scan =
@@ -62,7 +62,7 @@
     (fn src => fn () => fn ctxt =>
       let
         val (a, ctxt') = variant "position" ctxt;
-        val (_, pos) = Args.name_of_src src;
+        val (_, pos) = Token.name_of_src src;
         val env = "val " ^ a ^ " = " ^ ML_Syntax.print_position pos ^ ";\n";
         val body = "Isabelle." ^ a;
       in (K (env, body), ctxt') end) #>
--- a/src/Pure/ML/ml_antiquotations.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/ML/ml_antiquotations.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -52,6 +52,16 @@
           ML_Syntax.atomic (ML_Syntax.print_term t))));
 
 
+(* embedded source *)
+
+val _ = Theory.setup
+ (ML_Antiquotation.value @{binding source}
+    (Scan.lift Args.text_source_position >> (fn {delimited, text, pos} =>
+      "{delimited = " ^ Bool.toString delimited ^
+      ", text = " ^ ML_Syntax.print_string text ^
+      ", pos = " ^ ML_Syntax.print_position pos ^ "}")));
+
+
 (* ML support *)
 
 val _ = Theory.setup
@@ -64,7 +74,7 @@
     (Scan.lift (Scan.optional Args.name "Output.writeln"))
       (fn src => fn output => fn ctxt =>
         let
-          val (_, pos) = Args.name_of_src src;
+          val (_, pos) = Token.name_of_src src;
           val (a, ctxt') = ML_Antiquotation.variant "output" ctxt;
           val env =
             "val " ^ a ^ ": string -> unit =\n\
--- a/src/Pure/ML/ml_context.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/ML/ml_context.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -14,7 +14,7 @@
   val exec: (unit -> unit) -> Context.generic -> Context.generic
   val check_antiquotation: Proof.context -> xstring * Position.T -> string
   type decl = Proof.context -> string * string
-  val add_antiquotation: binding -> (Args.src -> Proof.context -> decl * Proof.context) ->
+  val add_antiquotation: binding -> (Token.src -> Proof.context -> decl * Proof.context) ->
     theory -> theory
   val print_antiquotations: Proof.context -> unit
   val eval_antiquotes: ML_Lex.token Antiquote.antiquote list * Position.T ->
@@ -55,7 +55,7 @@
 type decl = Proof.context -> string * string;  (*final context -> ML env, ML body*)
 structure Antiquotations = Theory_Data
 (
-  type T = (Args.src -> Proof.context -> decl * Proof.context) Name_Space.table;
+  type T = (Token.src -> Proof.context -> decl * Proof.context) Name_Space.table;
   val empty : T = Name_Space.empty_table Markup.ML_antiquotationN;
   val extend = I;
   fun merge data : T = Name_Space.merge_tables data;
@@ -75,7 +75,7 @@
   |> Pretty.writeln;
 
 fun apply_antiquotation src ctxt =
-  let val (src', f) = Args.check_src ctxt (get_antiquotations ctxt) src
+  let val (src', f) = Token.check_src ctxt (get_antiquotations ctxt) src
   in f src' ctxt end;
 
 
@@ -85,7 +85,7 @@
 
 val antiq =
   Parse.!!! (Parse.position Parse.xname -- Parse.args --| Scan.ahead Parse.eof)
-  >> uncurry Args.src;
+  >> uncurry Token.src;
 
 val begin_env0 = ML_Lex.tokenize "structure Isabelle =\nstruct\n";
 
--- a/src/Pure/ML/ml_thms.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/ML/ml_thms.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,7 +6,7 @@
 
 signature ML_THMS =
 sig
-  val the_attributes: Proof.context -> int -> Args.src list
+  val the_attributes: Proof.context -> int -> Token.src list
   val the_thmss: Proof.context -> thm list list
   val get_stored_thms: unit -> thm list
   val get_stored_thm: unit -> thm
@@ -25,7 +25,7 @@
 
 structure Data = Proof_Data
 (
-  type T = Args.src list Inttab.table * thms list;
+  type T = Token.src list Inttab.table * thms list;
   fun init _ = (Inttab.empty, []);
 );
 
@@ -37,18 +37,17 @@
 val cons_thms = Data.map o apsnd o cons;
 
 
-
 (* attribute source *)
 
 val _ = Theory.setup
-  (ML_Antiquotation.declaration @{binding attributes} (Scan.lift Parse_Spec.attribs)
+  (ML_Antiquotation.declaration @{binding attributes} (Scan.lift Parse.attribs)
     (fn _ => fn raw_srcs => fn ctxt =>
       let
         val i = serial ();
-        val srcs = map (Attrib.check_src ctxt) raw_srcs;
-        val _ = map (Attrib.attribute ctxt) srcs;
+
         val (a, ctxt') = ctxt
-          |> ML_Antiquotation.variant "attributes" ||> put_attributes (i, srcs);
+          |> ML_Antiquotation.variant "attributes"
+          ||> put_attributes (i, Attrib.check_attribs ctxt raw_srcs);
         val ml =
           ("val " ^ a ^ " = ML_Thms.the_attributes ML_context " ^
             string_of_int i ^ ";\n", "Isabelle." ^ a);
--- a/src/Pure/PIDE/command.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/command.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -121,9 +121,9 @@
   in {src_path = src_path, lines = lines, digest = digest, pos = file_pos} end
 
 fun resolve_files master_dir blobs toks =
-  (case Thy_Syntax.parse_spans toks of
+  (case Outer_Syntax.parse_spans toks of
     [span] => span
-      |> Thy_Syntax.resolve_files (fn cmd => fn (path, pos) =>
+      |> Command_Span.resolve_files (fn cmd => fn (path, pos) =>
         let
           fun make_file src_path (Exn.Res (file_node, NONE)) =
                 Exn.interruptible_capture (fn () =>
@@ -140,7 +140,7 @@
             map2 make_file src_paths blobs
           else error ("Misalignment of inlined files" ^ Position.here pos)
         end)
-      |> Thy_Syntax.span_content
+      |> Command_Span.content
   | _ => toks);
 
 in
--- a/src/Pure/PIDE/command.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/command.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -189,8 +189,7 @@
               def bad(): Unit = Output.warning("Ignored report message: " + msg)
 
               msg match {
-                case XML.Elem(
-                    Markup(name, atts @ Position.Reported(id, chunk_name, symbol_range)), args) =>
+                case XML.Elem(Markup(name, atts @ Position.Identified(id, chunk_name)), args) =>
 
                   val target =
                     if (self_id(id) && command.chunks.isDefinedAt(chunk_name))
@@ -198,8 +197,8 @@
                     else if (chunk_name == Symbol.Text_Chunk.Default) other_id(id)
                     else None
 
-                  target match {
-                    case Some((target_name, target_chunk)) =>
+                  (target, atts) match {
+                    case (Some((target_name, target_chunk)), Position.Range(symbol_range)) =>
                       target_chunk.incorporate(symbol_range) match {
                         case Some(range) =>
                           val props = Position.purge(atts)
@@ -207,7 +206,7 @@
                           state.add_markup(false, target_name, info)
                         case None => bad(); state
                       }
-                    case None =>
+                    case _ =>
                       // silently ignore excessive reports
                       state
                   }
@@ -232,7 +231,8 @@
               if (Protocol.is_inlined(message)) {
                 for {
                   (chunk_name, chunk) <- command.chunks.iterator
-                  range <- Protocol.message_positions(self_id, chunk_name, chunk, message)
+                  range <- Protocol.message_positions(
+                    self_id, command.position, chunk_name, chunk, message)
                 } st = st.add_markup(false, chunk_name, Text.Info(range, message2))
               }
               st
@@ -250,39 +250,18 @@
 
   /* make commands */
 
-  def name(span: List[Token]): String =
-    span.find(_.is_command) match { case Some(tok) => tok.source case _ => "" }
-
-  private def source_span(span: List[Token]): (String, List[Token]) =
-  {
-    val source: String =
-      span match {
-        case List(tok) => tok.source
-        case _ => span.map(_.source).mkString
-      }
-
-    val span1 = new mutable.ListBuffer[Token]
-    var i = 0
-    for (Token(kind, s) <- span) {
-      val n = s.length
-      val s1 = source.substring(i, i + n)
-      span1 += Token(kind, s1)
-      i += n
-    }
-    (source, span1.toList)
-  }
-
   def apply(
     id: Document_ID.Command,
     node_name: Document.Node.Name,
     blobs: List[Blob],
-    span: List[Token]): Command =
+    span: Command_Span.Span): Command =
   {
-    val (source, span1) = source_span(span)
+    val (source, span1) = span.compact_source
     new Command(id, node_name, blobs, span1, source, Results.empty, Markup_Tree.empty)
   }
 
-  val empty: Command = Command(Document_ID.none, Document.Node.Name.empty, Nil, Nil)
+  val empty: Command =
+    Command(Document_ID.none, Document.Node.Name.empty, Nil, Command_Span.empty)
 
   def unparsed(
     id: Document_ID.Command,
@@ -290,8 +269,8 @@
     results: Results,
     markup: Markup_Tree): Command =
   {
-    val (source1, span) = source_span(List(Token(Token.Kind.UNPARSED, source)))
-    new Command(id, Document.Node.Name.empty, Nil, span, source1, results, markup)
+    val (source1, span1) = Command_Span.unparsed(source).compact_source
+    new Command(id, Document.Node.Name.empty, Nil, span1, source1, results, markup)
   }
 
   def unparsed(source: String): Command =
@@ -333,25 +312,30 @@
     val id: Document_ID.Command,
     val node_name: Document.Node.Name,
     val blobs: List[Command.Blob],
-    val span: List[Token],
+    val span: Command_Span.Span,
     val source: String,
     val init_results: Command.Results,
     val init_markup: Markup_Tree)
 {
+  /* name */
+
+  def name: String =
+    span.kind match { case Command_Span.Command_Span(name, _) => name case _ => "" }
+
+  def position: Position.T =
+    span.kind match { case Command_Span.Command_Span(_, pos) => pos case _ => Position.none }
+
+  override def toString: String = id + "/" + span.kind.toString
+
+
   /* classification */
 
+  def is_proper: Boolean = span.kind.isInstanceOf[Command_Span.Command_Span]
+  def is_ignored: Boolean = span.kind == Command_Span.Ignored_Span
+
   def is_undefined: Boolean = id == Document_ID.none
-  val is_unparsed: Boolean = span.exists(_.is_unparsed)
-  val is_unfinished: Boolean = span.exists(_.is_unfinished)
-
-  val is_ignored: Boolean = !span.exists(_.is_proper)
-  val is_malformed: Boolean = !is_ignored && (!span.head.is_command || span.exists(_.is_error))
-  def is_command: Boolean = !is_ignored && !is_malformed
-
-  def name: String = Command.name(span)
-
-  override def toString =
-    id + "/" + (if (is_command) name else if (is_ignored) "IGNORED" else "MALFORMED")
+  val is_unparsed: Boolean = span.content.exists(_.is_unparsed)
+  val is_unfinished: Boolean = span.content.exists(_.is_unfinished)
 
 
   /* blobs */
@@ -379,7 +363,8 @@
   def range: Text.Range = chunk.range
 
   val proper_range: Text.Range =
-    Text.Range(0, (length /: span.reverse.iterator.takeWhile(_.is_improper))(_ - _.source.length))
+    Text.Range(0,
+      (length /: span.content.reverse.iterator.takeWhile(_.is_improper))(_ - _.source.length))
 
   def source(range: Text.Range): String = source.substring(range.start, range.stop)
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/PIDE/command_span.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,70 @@
+(*  Title:      Pure/PIDE/command_span.ML
+    Author:     Makarius
+
+Syntactic representation of command spans.
+*)
+
+signature COMMAND_SPAN =
+sig
+  datatype kind = Command_Span of string * Position.T | Ignored_Span | Malformed_Span
+  datatype span = Span of kind * Token.T list
+  val kind: span -> kind
+  val content: span -> Token.T list
+  val resolve_files: (string -> Path.T * Position.T -> Token.file Exn.result list) -> span -> span
+end;
+
+structure Command_Span: COMMAND_SPAN =
+struct
+
+datatype kind = Command_Span of string * Position.T | Ignored_Span | Malformed_Span;
+datatype span = Span of kind * Token.T list;
+
+fun kind (Span (k, _)) = k;
+fun content (Span (_, toks)) = toks;
+
+
+(* resolve inlined files *)
+
+local
+
+fun clean ((i1, t1) :: (i2, t2) :: toks) =
+      if Token.keyword_with (fn s => s = "%" orelse s = "--") t1 then clean toks
+      else (i1, t1) :: clean ((i2, t2) :: toks)
+  | clean toks = toks;
+
+fun clean_tokens toks =
+  ((0 upto length toks - 1) ~~ toks)
+  |> filter (fn (_, tok) => Token.is_proper tok)
+  |> clean;
+
+fun find_file ((_, tok) :: toks) =
+      if Token.is_command tok then
+        toks |> get_first (fn (i, tok) =>
+          if Token.is_name tok then
+            SOME (i, (Path.explode (Token.content_of tok), Token.pos_of tok))
+              handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok))
+          else NONE)
+      else NONE
+  | find_file [] = NONE;
+
+in
+
+fun resolve_files read_files span =
+  (case span of
+    Span (Command_Span (cmd, pos), toks) =>
+      if Keyword.is_theory_load cmd then
+        (case find_file (clean_tokens toks) of
+          NONE => error ("Bad file argument of command " ^ quote cmd ^ Position.here pos)
+        | SOME (i, path) =>
+            let
+              val toks' = toks |> map_index (fn (j, tok) =>
+                if i = j then Token.put_files (read_files cmd path) tok
+                else tok);
+            in Span (Command_Span (cmd, pos), toks') end)
+      else span
+  | _ => span);
+
+end;
+
+end;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/PIDE/command_span.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,104 @@
+/*  Title:      Pure/PIDE/command_span.scala
+    Author:     Makarius
+
+Syntactic representation of command spans.
+*/
+
+package isabelle
+
+
+import scala.collection.mutable
+
+
+object Command_Span
+{
+  sealed abstract class Kind {
+    override def toString: String =
+      this match {
+        case Command_Span(name, _) => if (name != "") name else "<command>"
+        case Ignored_Span => "<ignored>"
+        case Malformed_Span => "<malformed>"
+      }
+  }
+  case class Command_Span(name: String, pos: Position.T) extends Kind
+  case object Ignored_Span extends Kind
+  case object Malformed_Span extends Kind
+
+  sealed case class Span(kind: Kind, content: List[Token])
+  {
+    def compact_source: (String, Span) =
+    {
+      val source: String =
+        content match {
+          case List(tok) => tok.source
+          case toks => toks.map(_.source).mkString
+        }
+
+      val content1 = new mutable.ListBuffer[Token]
+      var i = 0
+      for (Token(kind, s) <- content) {
+        val n = s.length
+        val s1 = source.substring(i, i + n)
+        content1 += Token(kind, s1)
+        i += n
+      }
+      (source, Span(kind, content1.toList))
+    }
+  }
+
+  val empty: Span = Span(Ignored_Span, Nil)
+
+  def unparsed(source: String): Span =
+    Span(Malformed_Span, List(Token(Token.Kind.UNPARSED, source)))
+
+
+  /* resolve inlined files */
+
+  private def find_file(tokens: List[Token]): Option[String] =
+  {
+    def clean(toks: List[Token]): List[Token] =
+      toks match {
+        case t :: _ :: ts if t.is_keyword && (t.source == "%" || t.source == "--") => clean(ts)
+        case t :: ts => t :: clean(ts)
+        case Nil => Nil
+      }
+    clean(tokens.filter(_.is_proper)) match {
+      case tok :: toks if tok.is_command => toks.find(_.is_name).map(_.content)
+      case _ => None
+    }
+  }
+
+  def span_files(syntax: Prover.Syntax, span: Span): List[String] =
+    span.kind match {
+      case Command_Span(name, _) =>
+        syntax.load_command(name) match {
+          case Some(exts) =>
+            find_file(span.content) match {
+              case Some(file) =>
+                if (exts.isEmpty) List(file)
+                else exts.map(ext => file + "." + ext)
+              case None => Nil
+            }
+          case None => Nil
+        }
+      case _ => Nil
+    }
+
+  def resolve_files(
+      resources: Resources,
+      syntax: Prover.Syntax,
+      node_name: Document.Node.Name,
+      span: Span,
+      get_blob: Document.Node.Name => Option[Document.Blob])
+    : List[Command.Blob] =
+  {
+    span_files(syntax, span).map(file_name =>
+      Exn.capture {
+        val name =
+          Document.Node.Name(resources.append(node_name.master_dir, Path.explode(file_name)))
+        val blob = get_blob(name).map(blob => ((blob.bytes.sha1_digest, blob.chunk)))
+        (name, blob)
+      })
+  }
+}
+
--- a/src/Pure/PIDE/document.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/document.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -318,7 +318,7 @@
       val span =
         Lazy.lazy (fn () =>
           Position.setmp_thread_data (Position.id_only id)
-            (fn () => Thy_Syntax.parse_tokens (Keyword.get_lexicons ()) (Position.id id) text) ());
+            (fn () => Outer_Syntax.scan (Keyword.get_lexicons ()) (Position.id id) text) ());
       val _ =
         Position.setmp_thread_data (Position.id_only id)
           (fn () => Output.status (Markup.markup_only Markup.accepted)) ();
--- a/src/Pure/PIDE/markup.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/markup.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -72,6 +72,7 @@
   val fixedN: string val fixed: string -> T
   val caseN: string val case_: string -> T
   val dynamic_factN: string val dynamic_fact: string -> T
+  val method_modifierN: string
   val tfreeN: string val tfree: T
   val tvarN: string val tvar: T
   val freeN: string val free: T
@@ -365,7 +366,7 @@
 val (hiddenN, hidden) = markup_elem "hidden";
 
 
-(* formal entities *)
+(* misc entities *)
 
 val system_optionN = "system_option";
 
@@ -378,6 +379,8 @@
 val (caseN, case_) = markup_string "case" nameN;
 val (dynamic_factN, dynamic_fact) = markup_string "dynamic_fact" nameN;
 
+val method_modifierN = "method_modifier";
+
 
 (* inner syntax *)
 
--- a/src/Pure/PIDE/markup.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/markup.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -169,7 +169,7 @@
   val HIDDEN = "hidden"
 
 
-  /* logical entities */
+  /* misc entities */
 
   val CLASS = "class"
   val TYPE_NAME = "type_name"
--- a/src/Pure/PIDE/markup_tree.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/markup_tree.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -261,7 +261,7 @@
    make_body(root_range, Nil, overlapping(root_range))
   }
 
-  override def toString =
+  override def toString: String =
     branches.toList.map(_._2) match {
       case Nil => "Empty"
       case list => list.mkString("Tree(", ",", ")")
--- a/src/Pure/PIDE/protocol.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/protocol.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -312,17 +312,27 @@
 
   def message_positions(
     self_id: Document_ID.Generic => Boolean,
+    command_position: Position.T,
     chunk_name: Symbol.Text_Chunk.Name,
     chunk: Symbol.Text_Chunk,
     message: XML.Elem): Set[Text.Range] =
   {
     def elem_positions(props: Properties.T, set: Set[Text.Range]): Set[Text.Range] =
       props match {
-        case Position.Reported(id, name, symbol_range)
-        if self_id(id) && name == chunk_name =>
-          chunk.incorporate(symbol_range) match {
-            case Some(range) => set + range
-            case _ => set
+        case Position.Identified(id, name) if self_id(id) && name == chunk_name =>
+          val opt_range =
+            Position.Range.unapply(props) orElse {
+              if (name == Symbol.Text_Chunk.Default)
+                Position.Range.unapply(command_position)
+              else None
+            }
+          opt_range match {
+            case Some(symbol_range) =>
+              chunk.incorporate(symbol_range) match {
+                case Some(range) => set + range
+                case _ => set
+              }
+            case None => set
           }
         case _ => set
       }
@@ -343,8 +353,25 @@
 }
 
 
-trait Protocol extends Prover
+trait Protocol
 {
+  /* text */
+
+  def encode(s: String): String
+  def decode(s: String): String
+
+  object Encode
+  {
+    val string: XML.Encode.T[String] = (s => XML.Encode.string(encode(s)))
+  }
+
+
+  /* protocol commands */
+
+  def protocol_command_bytes(name: String, args: Bytes*): Unit
+  def protocol_command(name: String, args: String*): Unit
+
+
   /* options */
 
   def options(opts: Options): Unit =
--- a/src/Pure/PIDE/prover.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/prover.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -1,12 +1,16 @@
 /*  Title:      Pure/PIDE/prover.scala
     Author:     Makarius
+    Options:    :folding=explicit:
 
-General prover operations.
+Prover process wrapping.
 */
 
 package isabelle
 
 
+import java.io.{InputStream, OutputStream, BufferedReader, BufferedOutputStream, IOException}
+
+
 object Prover
 {
   /* syntax */
@@ -14,12 +18,23 @@
   trait Syntax
   {
     def add_keywords(keywords: Thy_Header.Keywords): Syntax
-    def scan(input: CharSequence): List[Token]
-    def load(span: List[Token]): Option[List[String]]
+    def parse_spans(input: CharSequence): List[Command_Span.Span]
+    def load_command(name: String): Option[List[String]]
     def load_commands_in(text: String): Boolean
   }
 
 
+  /* underlying system process */
+
+  trait System_Process
+  {
+    def stdout: BufferedReader
+    def stderr: BufferedReader
+    def terminate: Unit
+    def join: Int
+  }
+
+
   /* messages */
 
   sealed abstract class Message
@@ -68,44 +83,287 @@
 }
 
 
-trait Prover
+abstract class Prover(
+  receiver: Prover.Message => Unit,
+  system_channel: System_Channel,
+  system_process: Prover.System_Process) extends Protocol
 {
-  /* text and tree data */
+  /** receiver output **/
+
+  val xml_cache: XML.Cache = new XML.Cache()
+
+  private def system_output(text: String)
+  {
+    receiver(new Prover.Output(XML.Elem(Markup(Markup.SYSTEM, Nil), List(XML.Text(text)))))
+  }
+
+  private def protocol_output(props: Properties.T, bytes: Bytes)
+  {
+    receiver(new Prover.Protocol_Output(props, bytes))
+  }
+
+  private def output(kind: String, props: Properties.T, body: XML.Body)
+  {
+    if (kind == Markup.INIT) system_channel.accepted()
 
-  def encode(s: String): String
-  def decode(s: String): String
+    val main = XML.Elem(Markup(kind, props), Protocol.clean_message(body))
+    val reports = Protocol.message_reports(props, body)
+    for (msg <- main :: reports) receiver(new Prover.Output(xml_cache.elem(msg)))
+  }
+
+  private def exit_message(rc: Int)
+  {
+    output(Markup.EXIT, Markup.Return_Code(rc), List(XML.Text("Return code: " + rc.toString)))
+  }
+
+
 
-  object Encode
+  /** process manager **/
+
+  private val (_, process_result) =
+    Simple_Thread.future("process_result") { system_process.join }
+
+  private def terminate_process()
   {
-    val string: XML.Encode.T[String] = (s => XML.Encode.string(encode(s)))
+    try { system_process.terminate }
+    catch {
+      case exn @ ERROR(_) => system_output("Failed to terminate prover process: " + exn.getMessage)
+    }
   }
 
-  def xml_cache: XML.Cache
+  private val process_manager = Simple_Thread.fork("process_manager")
+  {
+    val (startup_failed, startup_errors) =
+    {
+      var finished: Option[Boolean] = None
+      val result = new StringBuilder(100)
+      while (finished.isEmpty && (system_process.stderr.ready || !process_result.is_finished)) {
+        while (finished.isEmpty && system_process.stderr.ready) {
+          try {
+            val c = system_process.stderr.read
+            if (c == 2) finished = Some(true)
+            else result += c.toChar
+          }
+          catch { case _: IOException => finished = Some(false) }
+        }
+        Thread.sleep(10)
+      }
+      (finished.isEmpty || !finished.get, result.toString.trim)
+    }
+    if (startup_errors != "") system_output(startup_errors)
+
+    if (startup_failed) {
+      terminate_process()
+      process_result.join
+      exit_message(127)
+    }
+    else {
+      val (command_stream, message_stream) = system_channel.rendezvous()
+
+      command_input_init(command_stream)
+      val stdout = physical_output(false)
+      val stderr = physical_output(true)
+      val message = message_output(message_stream)
+
+      val rc = process_result.join
+      system_output("process terminated")
+      command_input_close()
+      for (thread <- List(stdout, stderr, message)) thread.join
+      system_output("process_manager terminated")
+      exit_message(rc)
+    }
+    system_channel.accepted()
+  }
+
+
+  /* management methods */
+
+  def join() { process_manager.join() }
+
+  def terminate()
+  {
+    command_input_close()
+    system_output("Terminating prover process")
+    terminate_process()
+  }
+
+
+
+  /** process streams **/
+
+  /* command input */
+
+  private var command_input: Option[Consumer_Thread[List[Bytes]]] = None
+
+  private def command_input_close(): Unit = command_input.foreach(_.shutdown)
+
+  private def command_input_init(raw_stream: OutputStream)
+  {
+    val name = "command_input"
+    val stream = new BufferedOutputStream(raw_stream)
+    command_input =
+      Some(
+        Consumer_Thread.fork(name)(
+          consume =
+            {
+              case chunks =>
+                try {
+                  Bytes(chunks.map(_.length).mkString("", ",", "\n")).write(stream)
+                  chunks.foreach(_.write(stream))
+                  stream.flush
+                  true
+                }
+                catch { case e: IOException => system_output(name + ": " + e.getMessage); false }
+            },
+          finish = { case () => stream.close; system_output(name + " terminated") }
+        )
+      )
+  }
 
 
-  /* process management */
+  /* physical output */
+
+  private def physical_output(err: Boolean): Thread =
+  {
+    val (name, reader, markup) =
+      if (err) ("standard_error", system_process.stderr, Markup.STDERR)
+      else ("standard_output", system_process.stdout, Markup.STDOUT)
 
-  def join(): Unit
-  def terminate(): Unit
-
-  def protocol_command_bytes(name: String, args: Bytes*): Unit
-  def protocol_command(name: String, args: String*): Unit
+    Simple_Thread.fork(name) {
+      try {
+        var result = new StringBuilder(100)
+        var finished = false
+        while (!finished) {
+          //{{{
+          var c = -1
+          var done = false
+          while (!done && (result.length == 0 || reader.ready)) {
+            c = reader.read
+            if (c >= 0) result.append(c.asInstanceOf[Char])
+            else done = true
+          }
+          if (result.length > 0) {
+            output(markup, Nil, List(XML.Text(decode(result.toString))))
+            result.length = 0
+          }
+          else {
+            reader.close
+            finished = true
+          }
+          //}}}
+        }
+      }
+      catch { case e: IOException => system_output(name + ": " + e.getMessage) }
+      system_output(name + " terminated")
+    }
+  }
 
 
-  /* PIDE protocol commands */
+  /* message output */
+
+  private def message_output(stream: InputStream): Thread =
+  {
+    class EOF extends Exception
+    class Protocol_Error(msg: String) extends Exception(msg)
+
+    val name = "message_output"
+    Simple_Thread.fork(name) {
+      val default_buffer = new Array[Byte](65536)
+      var c = -1
 
-  def options(opts: Options): Unit
+      def read_int(): Int =
+      //{{{
+      {
+        var n = 0
+        c = stream.read
+        if (c == -1) throw new EOF
+        while (48 <= c && c <= 57) {
+          n = 10 * n + (c - 48)
+          c = stream.read
+        }
+        if (c != 10)
+          throw new Protocol_Error("malformed header: expected integer followed by newline")
+        else n
+      }
+      //}}}
 
-  def define_blob(digest: SHA1.Digest, bytes: Bytes): Unit
-  def define_command(command: Command): Unit
+      def read_chunk_bytes(): (Array[Byte], Int) =
+      //{{{
+      {
+        val n = read_int()
+        val buf =
+          if (n <= default_buffer.size) default_buffer
+          else new Array[Byte](n)
+
+        var i = 0
+        var m = 0
+        do {
+          m = stream.read(buf, i, n - i)
+          if (m != -1) i += m
+        }
+        while (m != -1 && n > i)
+
+        if (i != n)
+          throw new Protocol_Error("bad chunk (unexpected EOF after " + i + " of " + n + " bytes)")
+
+        (buf, n)
+      }
+      //}}}
 
-  def discontinue_execution(): Unit
-  def cancel_exec(id: Document_ID.Exec): Unit
+      def read_chunk(): XML.Body =
+      {
+        val (buf, n) = read_chunk_bytes()
+        YXML.parse_body_failsafe(UTF8.decode_chars(decode, buf, 0, n))
+      }
 
-  def update(old_id: Document_ID.Version, new_id: Document_ID.Version,
-    edits: List[Document.Edit_Command]): Unit
-  def remove_versions(versions: List[Document.Version]): Unit
+      try {
+        do {
+          try {
+            val header = read_chunk()
+            header match {
+              case List(XML.Elem(Markup(name, props), Nil)) =>
+                val kind = name.intern
+                if (kind == Markup.PROTOCOL) {
+                  val (buf, n) = read_chunk_bytes()
+                  protocol_output(props, Bytes(buf, 0, n))
+                }
+                else {
+                  val body = read_chunk()
+                  output(kind, props, body)
+                }
+              case _ =>
+                read_chunk()
+                throw new Protocol_Error("bad header: " + header.toString)
+            }
+          }
+          catch { case _: EOF => }
+        }
+        while (c != -1)
+      }
+      catch {
+        case e: IOException => system_output("Cannot read message:\n" + e.getMessage)
+        case e: Protocol_Error => system_output("Malformed message:\n" + e.getMessage)
+      }
+      stream.close
 
-  def dialog_result(serial: Long, result: String): Unit
+      system_output(name + " terminated")
+    }
+  }
+
+
+
+  /** protocol commands **/
+
+  def protocol_command_bytes(name: String, args: Bytes*): Unit =
+    command_input match {
+      case Some(thread) => thread.send(Bytes(name) :: args.toList)
+      case None => error("Uninitialized command input thread")
+    }
+
+  def protocol_command(name: String, args: String*)
+  {
+    receiver(new Prover.Input(name, args.toList))
+    protocol_command_bytes(name, args.map(Bytes(_)): _*)
+  }
 }
 
--- a/src/Pure/PIDE/resources.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/resources.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -132,7 +132,7 @@
 fun excursion master_dir last_timing init elements =
   let
     fun prepare_span span =
-      Thy_Syntax.span_content span
+      Command_Span.content span
       |> Command.read init master_dir []
       |> (fn tr => Toplevel.put_timing (last_timing tr) tr);
 
@@ -157,8 +157,8 @@
     val _ = Thy_Header.define_keywords header;
 
     val lexs = Keyword.get_lexicons ();
-    val toks = Thy_Syntax.parse_tokens lexs text_pos text;
-    val spans = Thy_Syntax.parse_spans toks;
+    val toks = Outer_Syntax.scan lexs text_pos text;
+    val spans = Outer_Syntax.parse_spans toks;
     val elements = Thy_Syntax.parse_elements spans;
 
     fun init () =
--- a/src/Pure/PIDE/resources.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/resources.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -56,8 +56,8 @@
 
   def loaded_files(syntax: Prover.Syntax, text: String): List[String] =
     if (syntax.load_commands_in(text)) {
-      val spans = Thy_Syntax.parse_spans(syntax.scan(text))
-      spans.iterator.map(Thy_Syntax.span_files(syntax, _)).flatten.toList
+      val spans = syntax.parse_spans(text)
+      spans.iterator.map(Command_Span.span_files(syntax, _)).flatten.toList
     }
     else Nil
 
@@ -126,6 +126,6 @@
   /* prover process */
 
   def start_prover(receiver: Prover.Message => Unit, name: String, args: List[String]): Prover =
-    new Isabelle_Process(receiver, args) with Protocol
+    Isabelle_Process(receiver, args)
 }
 
--- a/src/Pure/PIDE/session.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/session.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -1,6 +1,6 @@
 /*  Title:      Pure/PIDE/session.scala
     Author:     Makarius
-    Options:    :folding=explicit:collapseFolds=1:
+    Options:    :folding=explicit:
 
 PIDE editor session, potentially with running prover process.
 */
--- a/src/Pure/PIDE/text.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/text.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -40,7 +40,7 @@
     if (start > stop)
       error("Bad range: [" + start.toString + ":" + stop.toString + "]")
 
-    override def toString = "[" + start.toString + ":" + stop.toString + "]"
+    override def toString: String = "[" + start.toString + ":" + stop.toString + "]"
 
     def length: Int = stop - start
 
@@ -116,7 +116,7 @@
         case other: Perspective => ranges == other.ranges
         case _ => false
       }
-    override def toString = ranges.toString
+    override def toString: String = ranges.toString
   }
 
 
@@ -141,7 +141,7 @@
 
   final class Edit private(val is_insert: Boolean, val start: Offset, val text: String)
   {
-    override def toString =
+    override def toString: String =
       (if (is_insert) "Insert(" else "Remove(") + (start, text).toString + ")"
 
 
--- a/src/Pure/PIDE/xml.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/PIDE/xml.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -21,7 +21,7 @@
 
   type Attributes = Properties.T
 
-  sealed abstract class Tree { override def toString = string_of_tree(this) }
+  sealed abstract class Tree { override def toString: String = string_of_tree(this) }
   case class Elem(markup: Markup, body: List[Tree]) extends Tree
   {
     def name: String = markup.name
@@ -150,12 +150,17 @@
     private def trim_bytes(s: String): String = new String(s.toCharArray)
 
     private def cache_string(x: String): String =
-      lookup(x) match {
-        case Some(y) => y
-        case None =>
-          val z = trim_bytes(x)
-          if (z.length > max_string) z else store(z)
-      }
+      if (x == "true") "true"
+      else if (x == "false") "false"
+      else if (x == "0.0") "0.0"
+      else if (Library.is_small_int(x)) Library.signed_string_of_int(Integer.parseInt(x))
+      else
+        lookup(x) match {
+          case Some(y) => y
+          case None =>
+            val z = trim_bytes(x)
+            if (z.length > max_string) z else store(z)
+        }
     private def cache_props(x: Properties.T): Properties.T =
       if (x.isEmpty) x
       else
@@ -214,9 +219,9 @@
 
     /* atomic values */
 
-    def long_atom(i: Long): String = i.toString
+    def long_atom(i: Long): String = Library.signed_string_of_long(i)
 
-    def int_atom(i: Int): String = i.toString
+    def int_atom(i: Int): String = Library.signed_string_of_int(i)
 
     def bool_atom(b: Boolean): String = if (b) "1" else "0"
 
--- a/src/Pure/Pure.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Pure.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -99,6 +99,7 @@
   and "realizability" :: thy_decl == ""
   and "extract_type" "extract" :: thy_decl
   and "find_theorems" "find_consts" :: diag
+  and "named_theorems" :: thy_decl
   and "ProofGeneral.process_pgip" "ProofGeneral.pr" "ProofGeneral.undo"
     "ProofGeneral.restart" "ProofGeneral.kill_proof" "ProofGeneral.inform_file_processed"
     "ProofGeneral.inform_file_retracted" :: control
@@ -111,10 +112,12 @@
 ML_file "Isar/calculation.ML"
 ML_file "Tools/rail.ML"
 ML_file "Tools/rule_insts.ML";
+ML_file "Tools/thm_deps.ML";
 ML_file "Tools/find_theorems.ML"
 ML_file "Tools/find_consts.ML"
 ML_file "Tools/proof_general_pure.ML"
 ML_file "Tools/simplifier_trace.ML"
+ML_file "Tools/named_theorems.ML"
 
 
 section {* Basic attributes *}
--- a/src/Pure/ROOT	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/ROOT	Wed Aug 27 15:52:58 2014 +0200
@@ -160,6 +160,7 @@
     "ML/ml_syntax.ML"
     "PIDE/active.ML"
     "PIDE/command.ML"
+    "PIDE/command_span.ML"
     "PIDE/document.ML"
     "PIDE/document_id.ML"
     "PIDE/execution.ML"
@@ -201,7 +202,6 @@
     "Thy/latex.ML"
     "Thy/present.ML"
     "Thy/term_style.ML"
-    "Thy/thm_deps.ML"
     "Thy/thy_header.ML"
     "Thy/thy_info.ML"
     "Thy/thy_output.ML"
@@ -233,6 +233,7 @@
     "morphism.ML"
     "name.ML"
     "net.ML"
+    "par_tactical.ML"
     "pattern.ML"
     "primitive_defs.ML"
     "proofterm.ML"
--- a/src/Pure/ROOT.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/ROOT.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -234,8 +234,10 @@
 use "Isar/parse.ML";
 use "Isar/args.ML";
 
-(*theory sources*)
+(*theory specifications*)
+use "Isar/local_theory.ML";
 use "Thy/thy_header.ML";
+use "PIDE/command_span.ML";
 use "Thy/thy_syntax.ML";
 use "Thy/html.ML";
 use "Thy/latex.ML";
@@ -253,6 +255,7 @@
 (** bootstrap phase 2: towards Pure.thy and final ML toplevel setup *)
 
 (*basic proof engine*)
+use "par_tactical.ML";
 use "Isar/proof_display.ML";
 use "Isar/attrib.ML";
 use "Isar/context_rules.ML";
@@ -263,7 +266,6 @@
 
 (*local theories and targets*)
 use "Isar/locale.ML";
-use "Isar/local_theory.ML";
 use "Isar/generic_target.ML";
 use "Isar/overloading.ML";
 use "axclass.ML";
@@ -309,7 +311,6 @@
 use "PIDE/document.ML";
 
 (*theory and proof operations*)
-use "Thy/thm_deps.ML";
 use "Isar/isar_cmd.ML";
 
 use "subgoal.ML";
@@ -342,7 +343,7 @@
 toplevel_pp ["Task_Queue", "task"] "Pretty.str o Task_Queue.str_of_task";
 toplevel_pp ["Task_Queue", "group"] "Pretty.str o Task_Queue.str_of_group";
 toplevel_pp ["Position", "T"] "Pretty.position";
-toplevel_pp ["Binding", "binding"] "Pretty.str o Binding.print";
+toplevel_pp ["Binding", "binding"] "Binding.pp";
 toplevel_pp ["Thm", "thm"] "Proof_Display.pp_thm";
 toplevel_pp ["Thm", "cterm"] "Proof_Display.pp_cterm";
 toplevel_pp ["Thm", "ctyp"] "Proof_Display.pp_ctyp";
--- a/src/Pure/Syntax/syntax.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Syntax/syntax.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -185,10 +185,10 @@
   let
     fun parse_tree tree =
       let
-        val {delimited, text, pos} = token_source tree;
-        val syms = Symbol_Pos.explode (text, pos);
-        val _ = Context_Position.report ctxt pos (markup delimited);
-      in parse (syms, pos) end;
+        val source = token_source tree;
+        val syms = Symbol_Pos.source_explode source;
+        val _ = Context_Position.report ctxt (#pos source) (markup (#delimited source));
+      in parse (syms, #pos source) end;
   in
     (case YXML.parse_body str handle Fail msg => error msg of
       body as [tree as XML.Elem ((name, _), _)] =>
--- a/src/Pure/Syntax/syntax_phases.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Syntax/syntax_phases.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -461,11 +461,11 @@
           else decode_appl ps asts
       | decode ps (Ast.Appl asts) = decode_appl ps asts;
 
-    val {text, pos, ...} = Syntax.read_token str;
-    val syms = Symbol_Pos.explode (text, pos);
+    val source = Syntax.read_token str;
+    val syms = Symbol_Pos.source_explode source;
     val ast =
-      parse_asts ctxt true root (syms, pos)
-      |> uncurry (report_result ctxt pos)
+      parse_asts ctxt true root (syms, #pos source)
+      |> uncurry (report_result ctxt (#pos source))
       |> decode [];
     val _ = Context_Position.reports_text ctxt (! reports);
   in ast end;
--- a/src/Pure/System/isabelle_font.scala	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,36 +0,0 @@
-/*  Title:      Pure/System/isabelle_font.scala
-    Author:     Makarius
-
-Isabelle font support.
-*/
-
-package isabelle
-
-
-import java.awt.{GraphicsEnvironment, Font}
-import java.io.{FileInputStream, BufferedInputStream}
-import javafx.scene.text.{Font => JFX_Font}
-
-
-object Isabelle_Font
-{
-  def apply(family: String = "IsabelleText", size: Int = 1, bold: Boolean = false): Font =
-    new Font(family, if (bold) Font.BOLD else Font.PLAIN, size)
-
-  def install_fonts()
-  {
-    val ge = GraphicsEnvironment.getLocalGraphicsEnvironment()
-    for (font <- Path.split(Isabelle_System.getenv_strict("ISABELLE_FONTS")))
-      ge.registerFont(Font.createFont(Font.TRUETYPE_FONT, font.file))
-  }
-
-  def install_fonts_jfx()
-  {
-    for (font <- Path.split(Isabelle_System.getenv_strict("ISABELLE_FONTS"))) {
-      val stream = new BufferedInputStream(new FileInputStream(font.file))
-      try { JFX_Font.loadFont(stream, 1.0) }
-      finally { stream.close }
-    }
-  }
-}
-
--- a/src/Pure/System/isabelle_process.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/System/isabelle_process.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -1,8 +1,7 @@
 (*  Title:      Pure/System/isabelle_process.ML
     Author:     Makarius
 
-Isabelle process wrapper, based on private fifos for maximum
-robustness and performance, or local socket for maximum portability.
+Isabelle process wrapper.
 *)
 
 signature ISABELLE_PROCESS =
@@ -108,8 +107,12 @@
     fun standard_message props name body =
       if forall (fn s => s = "") body then ()
       else
-        message name
-          (fold_rev Properties.put props (Position.properties_of (Position.thread_data ()))) body;
+        let
+          val props' =
+            (case (Properties.defined props Markup.idN, Position.get_id (Position.thread_data ())) of
+              (false, SOME id') => props @ [(Markup.idN, id')]
+            | _ => props);
+        in message name props' body end;
   in
     Output.status_fn := standard_message [] Markup.statusN;
     Output.report_fn := standard_message [] Markup.reportN;
--- a/src/Pure/System/isabelle_process.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/System/isabelle_process.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -1,318 +1,43 @@
 /*  Title:      Pure/System/isabelle_process.scala
     Author:     Makarius
-    Options:    :folding=explicit:collapseFolds=1:
 
-Isabelle process management -- always reactive due to multi-threaded I/O.
+Isabelle process wrapper.
 */
 
 package isabelle
 
 
-import java.io.{InputStream, OutputStream, BufferedOutputStream, IOException}
-
-
-class Isabelle_Process(
-  receiver: Prover.Message => Unit = Console.println(_),
-  prover_args: List[String] = Nil)
+object Isabelle_Process
 {
-  /* text and tree data */
-
-  def encode(s: String): String = Symbol.encode(s)
-  def decode(s: String): String = Symbol.decode(s)
-
-  val xml_cache = new XML.Cache()
-
-
-  /* output */
-
-  private def system_output(text: String)
-  {
-    receiver(new Prover.Output(XML.Elem(Markup(Markup.SYSTEM, Nil), List(XML.Text(text)))))
-  }
-
-  private def protocol_output(props: Properties.T, bytes: Bytes)
-  {
-    receiver(new Prover.Protocol_Output(props, bytes))
-  }
-
-  private def output(kind: String, props: Properties.T, body: XML.Body)
-  {
-    if (kind == Markup.INIT) system_channel.accepted()
-
-    val main = XML.Elem(Markup(kind, props), Protocol.clean_message(body))
-    val reports = Protocol.message_reports(props, body)
-    for (msg <- main :: reports) receiver(new Prover.Output(xml_cache.elem(msg)))
-  }
-
-  private def exit_message(rc: Int)
-  {
-    output(Markup.EXIT, Markup.Return_Code(rc), List(XML.Text("Return code: " + rc.toString)))
-  }
-
-
-
-  /** process manager **/
-
-  def command_line(channel: System_Channel, args: List[String]): List[String] =
-    Isabelle_System.getenv_strict("ISABELLE_PROCESS") :: (channel.isabelle_args ::: args)
-
-  private val system_channel = System_Channel()
-
-  private val process =
-    try {
-      val cmdline = command_line(system_channel, prover_args)
-      new Isabelle_System.Managed_Process(null, null, false, cmdline: _*)
-    }
-    catch { case e: IOException => system_channel.accepted(); throw(e) }
-
-  private val (_, process_result) =
-    Simple_Thread.future("process_result") { process.join }
-
-  private def terminate_process()
+  def apply(
+    receiver: Prover.Message => Unit = Console.println(_),
+    prover_args: List[String] = Nil): Isabelle_Process =
   {
-    try { process.terminate }
-    catch { case e: IOException => system_output("Failed to terminate Isabelle: " + e.getMessage) }
-  }
-
-  private val process_manager = Simple_Thread.fork("process_manager")
-  {
-    val (startup_failed, startup_errors) =
-    {
-      var finished: Option[Boolean] = None
-      val result = new StringBuilder(100)
-      while (finished.isEmpty && (process.stderr.ready || !process_result.is_finished)) {
-        while (finished.isEmpty && process.stderr.ready) {
-          try {
-            val c = process.stderr.read
-            if (c == 2) finished = Some(true)
-            else result += c.toChar
-          }
-          catch { case _: IOException => finished = Some(false) }
-        }
-        Thread.sleep(10)
+    val system_channel = System_Channel()
+    val system_process =
+      try {
+        val cmdline =
+          Isabelle_System.getenv_strict("ISABELLE_PROCESS") ::
+            (system_channel.prover_args ::: prover_args)
+        val process =
+          new Isabelle_System.Managed_Process(null, null, false, cmdline: _*) with
+            Prover.System_Process
+        process.stdin.close
+        process
       }
-      (finished.isEmpty || !finished.get, result.toString.trim)
-    }
-    if (startup_errors != "") system_output(startup_errors)
+      catch { case exn @ ERROR(_) => system_channel.accepted(); throw(exn) }
 
-    process.stdin.close
-    if (startup_failed) {
-      terminate_process()
-      process_result.join
-      exit_message(127)
-    }
-    else {
-      val (command_stream, message_stream) = system_channel.rendezvous()
-
-      command_input_init(command_stream)
-      val stdout = physical_output(false)
-      val stderr = physical_output(true)
-      val message = message_output(message_stream)
+    new Isabelle_Process(receiver, system_channel, system_process)
+  }
+}
 
-      val rc = process_result.join
-      system_output("process terminated")
-      command_input_close()
-      for (thread <- List(stdout, stderr, message)) thread.join
-      system_output("process_manager terminated")
-      exit_message(rc)
-    }
-    system_channel.accepted()
-  }
-
-
-  /* management methods */
-
-  def join() { process_manager.join() }
-
-  def interrupt()
+class Isabelle_Process private(
+    receiver: Prover.Message => Unit,
+    system_channel: System_Channel,
+    system_process: Prover.System_Process)
+  extends Prover(receiver, system_channel, system_process)
   {
-    try { process.interrupt }
-    catch { case e: IOException => system_output("Failed to interrupt Isabelle: " + e.getMessage) }
-  }
-
-  def terminate()
-  {
-    command_input_close()
-    system_output("Terminating Isabelle process")
-    terminate_process()
+    def encode(s: String): String = Symbol.encode(s)
+    def decode(s: String): String = Symbol.decode(s)
   }
 
-
-
-  /** process streams **/
-
-  /* command input */
-
-  private var command_input: Option[Consumer_Thread[List[Bytes]]] = None
-
-  private def command_input_close(): Unit = command_input.foreach(_.shutdown)
-
-  private def command_input_init(raw_stream: OutputStream)
-  {
-    val name = "command_input"
-    val stream = new BufferedOutputStream(raw_stream)
-    command_input =
-      Some(
-        Consumer_Thread.fork(name)(
-          consume =
-            {
-              case chunks =>
-                try {
-                  Bytes(chunks.map(_.length).mkString("", ",", "\n")).write(stream)
-                  chunks.foreach(_.write(stream))
-                  stream.flush
-                  true
-                }
-                catch { case e: IOException => system_output(name + ": " + e.getMessage); false }
-            },
-          finish = { case () => stream.close; system_output(name + " terminated") }
-        )
-      )
-  }
-
-
-  /* physical output */
-
-  private def physical_output(err: Boolean): Thread =
-  {
-    val (name, reader, markup) =
-      if (err) ("standard_error", process.stderr, Markup.STDERR)
-      else ("standard_output", process.stdout, Markup.STDOUT)
-
-    Simple_Thread.fork(name) {
-      try {
-        var result = new StringBuilder(100)
-        var finished = false
-        while (!finished) {
-          //{{{
-          var c = -1
-          var done = false
-          while (!done && (result.length == 0 || reader.ready)) {
-            c = reader.read
-            if (c >= 0) result.append(c.asInstanceOf[Char])
-            else done = true
-          }
-          if (result.length > 0) {
-            output(markup, Nil, List(XML.Text(decode(result.toString))))
-            result.length = 0
-          }
-          else {
-            reader.close
-            finished = true
-          }
-          //}}}
-        }
-      }
-      catch { case e: IOException => system_output(name + ": " + e.getMessage) }
-      system_output(name + " terminated")
-    }
-  }
-
-
-  /* message output */
-
-  private def message_output(stream: InputStream): Thread =
-  {
-    class EOF extends Exception
-    class Protocol_Error(msg: String) extends Exception(msg)
-
-    val name = "message_output"
-    Simple_Thread.fork(name) {
-      val default_buffer = new Array[Byte](65536)
-      var c = -1
-
-      def read_int(): Int =
-      //{{{
-      {
-        var n = 0
-        c = stream.read
-        if (c == -1) throw new EOF
-        while (48 <= c && c <= 57) {
-          n = 10 * n + (c - 48)
-          c = stream.read
-        }
-        if (c != 10)
-          throw new Protocol_Error("malformed header: expected integer followed by newline")
-        else n
-      }
-      //}}}
-
-      def read_chunk_bytes(): (Array[Byte], Int) =
-      //{{{
-      {
-        val n = read_int()
-        val buf =
-          if (n <= default_buffer.size) default_buffer
-          else new Array[Byte](n)
-
-        var i = 0
-        var m = 0
-        do {
-          m = stream.read(buf, i, n - i)
-          if (m != -1) i += m
-        }
-        while (m != -1 && n > i)
-
-        if (i != n)
-          throw new Protocol_Error("bad chunk (unexpected EOF after " + i + " of " + n + " bytes)")
-
-        (buf, n)
-      }
-      //}}}
-
-      def read_chunk(): XML.Body =
-      {
-        val (buf, n) = read_chunk_bytes()
-        YXML.parse_body_failsafe(UTF8.decode_chars(decode, buf, 0, n))
-      }
-
-      try {
-        do {
-          try {
-            val header = read_chunk()
-            header match {
-              case List(XML.Elem(Markup(name, props), Nil)) =>
-                val kind = name.intern
-                if (kind == Markup.PROTOCOL) {
-                  val (buf, n) = read_chunk_bytes()
-                  protocol_output(props, Bytes(buf, 0, n))
-                }
-                else {
-                  val body = read_chunk()
-                  output(kind, props, body)
-                }
-              case _ =>
-                read_chunk()
-                throw new Protocol_Error("bad header: " + header.toString)
-            }
-          }
-          catch { case _: EOF => }
-        }
-        while (c != -1)
-      }
-      catch {
-        case e: IOException => system_output("Cannot read message:\n" + e.getMessage)
-        case e: Protocol_Error => system_output("Malformed message:\n" + e.getMessage)
-      }
-      stream.close
-
-      system_output(name + " terminated")
-    }
-  }
-
-
-
-  /** protocol commands **/
-
-  def protocol_command_bytes(name: String, args: Bytes*): Unit =
-    command_input match {
-      case Some(thread) => thread.send(Bytes(name) :: args.toList)
-      case None => error("Uninitialized command input thread")
-    }
-
-  def protocol_command(name: String, args: String*)
-  {
-    receiver(new Prover.Input(name, args.toList))
-    protocol_command_bytes(name, args.map(Bytes(_)): _*)
-  }
-}
--- a/src/Pure/System/system_channel.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/System/system_channel.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -22,7 +22,7 @@
 abstract class System_Channel
 {
   def params: List[String]
-  def isabelle_args: List[String]
+  def prover_args: List[String]
   def rendezvous(): (OutputStream, InputStream)
   def accepted(): Unit
 }
@@ -60,7 +60,7 @@
 
   def params: List[String] = List(fifo1, fifo2)
 
-  val isabelle_args: List[String] = List ("-W", fifo1 + ":" + fifo2)
+  val prover_args: List[String] = List ("-W", fifo1 + ":" + fifo2)
 
   def rendezvous(): (OutputStream, InputStream) =
   {
@@ -81,7 +81,7 @@
 
   def params: List[String] = List("127.0.0.1", server.getLocalPort.toString)
 
-  def isabelle_args: List[String] = List("-T", "127.0.0.1:" + server.getLocalPort)
+  def prover_args: List[String] = List("-T", "127.0.0.1:" + server.getLocalPort)
 
   def rendezvous(): (OutputStream, InputStream) =
   {
--- a/src/Pure/Thy/present.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Thy/present.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -330,9 +330,7 @@
         val _ =
           Isabelle_System.isabelle_tool "latex"
             ("-o sty " ^ File.shell_path (Path.append doc_dir (Path.basic "root.tex")));
-        val _ =
-          if null doc_files then Isabelle_System.copy_dir document_path doc_dir
-          else List.app (fn file => Isabelle_System.copy_file_base file doc_dir) doc_files;
+        val _ = List.app (fn file => Isabelle_System.copy_file_base file doc_dir) doc_files;
         val _ =
           (case opt_graphs of
             NONE => ()
@@ -360,12 +358,6 @@
         NONE => []
       | SOME path => map (document_job path false) documents);
 
-    val _ =
-      if not (null jobs) andalso null doc_files then
-        Output.physical_stderr ("### Legacy feature! Document preparation for session " ^ quote name ^
-          " without 'document_files'\n")
-      else ();
-
     val _ = jobs |> Par_List.map (fn job => job ()) |> List.app (op |>);
   in
     browser_info := empty_browser_info;
--- a/src/Pure/Thy/term_style.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Thy/term_style.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -35,8 +35,8 @@
 fun parse_single ctxt =
   Parse.position Parse.xname -- Parse.args >> (fn (name, args) =>
     let
-      val (src, parse) = Args.check_src ctxt (get_data ctxt) (Args.src name args);
-      val (f, _) = Args.syntax (Scan.lift parse) src ctxt;
+      val (src, parse) = Token.check_src ctxt (get_data ctxt) (Token.src name args);
+      val (f, _) = Token.syntax (Scan.lift parse) src ctxt;
     in f ctxt end);
 
 val parse = Args.context :|-- (fn ctxt => Scan.lift
--- a/src/Pure/Thy/thm_deps.ML	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,99 +0,0 @@
-(*  Title:      Pure/Thy/thm_deps.ML
-    Author:     Stefan Berghofer, TU Muenchen
-
-Visualize dependencies of theorems.
-*)
-
-signature THM_DEPS =
-sig
-  val thm_deps: theory -> thm list -> unit
-  val unused_thms: theory list * theory list -> (string * thm) list
-end;
-
-structure Thm_Deps: THM_DEPS =
-struct
-
-(* thm_deps *)
-
-fun thm_deps thy thms =
-  let
-    fun add_dep ("", _, _) = I
-      | add_dep (name, _, PBody {thms = thms', ...}) =
-          let
-            val prefix = #1 (split_last (Long_Name.explode name));
-            val session =
-              (case prefix of
-                a :: _ =>
-                  (case try (Context.get_theory thy) a of
-                    SOME thy =>
-                      (case Present.session_name thy of
-                        "" => []
-                      | session => [session])
-                  | NONE => [])
-               | _ => ["global"]);
-            val parents = filter_out (fn s => s = "") (map (#1 o #2) thms');
-            val entry =
-              {name = Long_Name.base_name name,
-               ID = name,
-               dir = space_implode "/" (session @ prefix),
-               unfold = false,
-               path = "",
-               parents = parents,
-               content = []};
-          in cons entry end;
-    val deps = Proofterm.fold_body_thms add_dep (Thm.proof_bodies_of thms) [];
-  in Graph_Display.display_graph (sort_wrt #ID deps) end;
-
-
-(* unused_thms *)
-
-fun unused_thms (base_thys, thys) =
-  let
-    fun add_fact space (name, ths) =
-      if exists (fn thy => Global_Theory.defined_fact thy name) base_thys then I
-      else
-        let val {concealed, group, ...} = Name_Space.the_entry space name in
-          fold_rev (fn th =>
-            (case Thm.derivation_name th of
-              "" => I
-            | a => cons (a, (th, concealed, group)))) ths
-        end;
-    fun add_facts facts = Facts.fold_static (add_fact (Facts.space_of facts)) facts;
-
-    val new_thms =
-      fold (add_facts o Global_Theory.facts_of) thys []
-      |> sort_distinct (string_ord o pairself #1);
-
-    val used =
-      Proofterm.fold_body_thms
-        (fn (a, _, _) => a <> "" ? Symtab.update (a, ()))
-        (map Proofterm.strip_thm (Thm.proof_bodies_of (map (#1 o #2) new_thms)))
-        Symtab.empty;
-
-    fun is_unused a = not (Symtab.defined used a);
-
-    (* groups containing at least one used theorem *)
-    val used_groups = fold (fn (a, (_, _, group)) =>
-      if is_unused a then I
-      else
-        (case group of
-          NONE => I
-        | SOME grp => Inttab.update (grp, ()))) new_thms Inttab.empty;
-
-    val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, seen_groups) =>
-      if not concealed andalso
-        (* FIXME replace by robust treatment of thm groups *)
-        member (op =) [Thm.theoremK, Thm.lemmaK, Thm.corollaryK] (Thm.legacy_get_kind th) andalso
-        is_unused a
-      then
-        (case group of
-           NONE => ((a, th) :: thms, seen_groups)
-         | SOME grp =>
-             if Inttab.defined used_groups grp orelse
-               Inttab.defined seen_groups grp then q
-             else ((a, th) :: thms, Inttab.update (grp, ()) seen_groups))
-      else q) new_thms ([], Inttab.empty);
-  in rev thms' end;
-
-end;
-
--- a/src/Pure/Thy/thy_info.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Thy/thy_info.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -380,7 +380,7 @@
 
 fun script_thy pos txt =
   let
-    val trs = Outer_Syntax.parse pos txt;
+    val trs = Outer_Syntax.parse (Outer_Syntax.get_syntax ()) pos txt;
     val end_pos = if null trs then pos else Toplevel.pos_of (List.last trs);
     val end_state = fold (Toplevel.command_exception true) trs Toplevel.toplevel;
   in Toplevel.end_theory end_pos end_state end;
--- a/src/Pure/Thy/thy_output.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Thy/thy_output.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -18,7 +18,7 @@
   val check_option: Proof.context -> xstring * Position.T -> string
   val print_antiquotations: Proof.context -> unit
   val antiquotation: binding -> 'a context_parser ->
-    ({source: Args.src, state: Toplevel.state, context: Proof.context} -> 'a -> string) ->
+    ({source: Token.src, state: Toplevel.state, context: Proof.context} -> 'a -> string) ->
       theory -> theory
   val boolean: string -> bool
   val integer: string -> int
@@ -30,9 +30,9 @@
   val pretty_text: Proof.context -> string -> Pretty.T
   val pretty_term: Proof.context -> term -> Pretty.T
   val pretty_thm: Proof.context -> thm -> Pretty.T
-  val str_of_source: Args.src -> string
+  val str_of_source: Token.src -> string
   val maybe_pretty_source: (Proof.context -> 'a -> Pretty.T) -> Proof.context ->
-    Args.src -> 'a list -> Pretty.T list
+    Token.src -> 'a list -> Pretty.T list
   val output: Proof.context -> Pretty.T list -> string
   val verb_text: string -> string
 end;
@@ -67,7 +67,7 @@
 structure Antiquotations = Theory_Data
 (
   type T =
-    (Args.src -> Toplevel.state -> Proof.context -> string) Name_Space.table *
+    (Token.src -> Toplevel.state -> Proof.context -> string) Name_Space.table *
       (string -> Proof.context -> Proof.context) Name_Space.table;
   val empty : T =
     (Name_Space.empty_table Markup.document_antiquotationN,
@@ -91,7 +91,7 @@
 fun check_option ctxt = #1 o Name_Space.check (Context.Proof ctxt) (#2 (get_antiquotations ctxt));
 
 fun command src state ctxt =
-  let val (src', f) = Args.check_src ctxt (#1 (get_antiquotations ctxt)) src
+  let val (src', f) = Token.check_src ctxt (#1 (get_antiquotations ctxt)) src
   in f src' state ctxt end;
 
 fun option ((xname, pos), s) ctxt =
@@ -114,7 +114,7 @@
 fun antiquotation name scan body =
   add_command name
     (fn src => fn state => fn ctxt =>
-      let val (x, ctxt') = Args.syntax scan src ctxt
+      let val (x, ctxt') = Token.syntax scan src ctxt
       in body {source = src, state = state, context = ctxt'} x end);
 
 
@@ -151,7 +151,7 @@
 val antiq =
   Parse.!!!
     (Parse.position Parse.liberal_name -- properties -- Parse.args --| Scan.ahead Parse.eof)
-  >> (fn ((name, props), args) => (props, Args.src name args));
+  >> (fn ((name, props), args) => (props, Token.src name args));
 
 end;
 
@@ -534,7 +534,7 @@
 
 (* default output *)
 
-val str_of_source = space_implode " " o Args.unparse_src;
+val str_of_source = space_implode " " o Token.unparse_src;
 
 fun maybe_pretty_source pretty ctxt src xs =
   map (pretty ctxt) xs  (*always pretty in order to exhibit errors!*)
@@ -624,7 +624,7 @@
       Scan.lift (Parse.position (Args.$$$ "by") -- Method.parse -- Scan.option Method.parse))
     (fn {source, context = ctxt, ...} => fn ((prop_token, prop), (((_, by_pos), m1), m2)) =>
       let
-        val prop_src = Args.src (Args.name_of_src source) [prop_token];
+        val prop_src = Token.src (Token.name_of_src source) [prop_token];
 
         val reports = (by_pos, Markup.keyword1) :: maps Method.reports_of (m1 :: the_list m2);
         val _ = Context_Position.reports ctxt reports;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Thy/thy_structure.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,71 @@
+/*  Title:      Pure/Thy/thy_structure.scala
+    Author:     Makarius
+
+Nested structure of theory source.
+*/
+
+package isabelle
+
+
+import scala.collection.mutable
+import scala.annotation.tailrec
+
+
+object Thy_Structure
+{
+  sealed abstract class Entry { def length: Int }
+  case class Block(val name: String, val body: List[Entry]) extends Entry
+  {
+    val length: Int = (0 /: body)(_ + _.length)
+  }
+  case class Atom(val command: Command) extends Entry
+  {
+    def length: Int = command.length
+  }
+
+  def parse(syntax: Outer_Syntax, node_name: Document.Node.Name, text: CharSequence): Entry =
+  {
+    /* stack operations */
+
+    def buffer(): mutable.ListBuffer[Entry] = new mutable.ListBuffer[Entry]
+    var stack: List[(Int, String, mutable.ListBuffer[Entry])] =
+      List((0, node_name.toString, buffer()))
+
+    @tailrec def close(level: Int => Boolean)
+    {
+      stack match {
+        case (lev, name, body) :: (_, _, body2) :: rest if level(lev) =>
+          body2 += Block(name, body.toList)
+          stack = stack.tail
+          close(level)
+        case _ =>
+      }
+    }
+
+    def result(): Entry =
+    {
+      close(_ => true)
+      val (_, name, body) = stack.head
+      Block(name, body.toList)
+    }
+
+    def add(command: Command)
+    {
+      syntax.heading_level(command) match {
+        case Some(i) =>
+          close(_ > i)
+          stack = (i + 1, command.source, buffer()) :: stack
+        case None =>
+      }
+      stack.head._3 += Atom(command)
+    }
+
+
+    /* result structure */
+
+    val spans = syntax.parse_spans(text)
+    spans.foreach(span => add(Command(Document_ID.none, node_name, Nil, span)))
+    result()
+  }
+}
+
--- a/src/Pure/Thy/thy_syntax.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Thy/thy_syntax.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,39 +6,21 @@
 
 signature THY_SYNTAX =
 sig
-  val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> Token.T list
   val reports_of_tokens: Token.T list -> bool * Position.report_text list
   val present_token: Token.T -> Output.output
-  datatype span_kind = Command of string * Position.T | Ignored | Malformed
-  datatype span = Span of span_kind * Token.T list
-  val span_kind: span -> span_kind
-  val span_content: span -> Token.T list
-  val present_span: span -> Output.output
-  val parse_spans: Token.T list -> span list
-  val resolve_files: (string -> Path.T * Position.T -> Token.file Exn.result list) -> span -> span
+  val present_span: Command_Span.span -> Output.output
   datatype 'a element = Element of 'a * ('a element list * 'a) option
   val atom: 'a -> 'a element
   val map_element: ('a -> 'b) -> 'a element -> 'b element
   val flat_element: 'a element -> 'a list
   val last_element: 'a element -> 'a
-  val parse_elements: span list -> span element list
+  val parse_elements: Command_Span.span list -> Command_Span.span element list
 end;
 
 structure Thy_Syntax: THY_SYNTAX =
 struct
 
-(** tokens **)
-
-(* parse *)
-
-fun parse_tokens lexs pos =
-  Source.of_string #>
-  Symbol.source #>
-  Token.source {do_recover = SOME false} (K lexs) pos #>
-  Source.exhaust;
-
-
-(* present *)
+(** presentation **)
 
 local
 
@@ -60,97 +42,12 @@
   let val results = map reports_of_token toks
   in (exists fst results, maps snd results) end;
 
+end;
+
 fun present_token tok =
   Markup.enclose (Token.markup tok) (Output.output (Token.unparse tok));
 
-end;
-
-
-
-(** spans **)
-
-(* type span *)
-
-datatype span_kind = Command of string * Position.T | Ignored | Malformed;
-datatype span = Span of span_kind * Token.T list;
-
-fun span_kind (Span (k, _)) = k;
-fun span_content (Span (_, toks)) = toks;
-
-val present_span = implode o map present_token o span_content;
-
-
-(* parse *)
-
-local
-
-fun make_span toks =
-  if not (null toks) andalso Token.is_command (hd toks) then
-    Span (Command (Token.content_of (hd toks), Token.pos_of (hd toks)), toks)
-  else if forall Token.is_improper toks then Span (Ignored, toks)
-  else Span (Malformed, toks);
-
-fun flush (result, span, improper) =
-  result
-  |> not (null span) ? cons (rev span)
-  |> not (null improper) ? cons (rev improper);
-
-fun parse tok (result, span, improper) =
-  if Token.is_command tok then (flush (result, span, improper), [tok], [])
-  else if Token.is_improper tok then (result, span, tok :: improper)
-  else (result, tok :: (improper @ span), []);
-
-in
-
-fun parse_spans toks =
-  fold parse toks ([], [], [])
-  |> flush |> rev |> map make_span;
-
-end;
-
-
-(* inlined files *)
-
-local
-
-fun clean ((i1, t1) :: (i2, t2) :: toks) =
-      if Token.keyword_with (fn s => s = "%" orelse s = "--") t1 then clean toks
-      else (i1, t1) :: clean ((i2, t2) :: toks)
-  | clean toks = toks;
-
-fun clean_tokens toks =
-  ((0 upto length toks - 1) ~~ toks)
-  |> filter (fn (_, tok) => Token.is_proper tok)
-  |> clean;
-
-fun find_file ((_, tok) :: toks) =
-      if Token.is_command tok then
-        toks |> get_first (fn (i, tok) =>
-          if Token.is_name tok then
-            SOME (i, (Path.explode (Token.content_of tok), Token.pos_of tok))
-              handle ERROR msg => error (msg ^ Position.here (Token.pos_of tok))
-          else NONE)
-      else NONE
-  | find_file [] = NONE;
-
-in
-
-fun resolve_files read_files span =
-  (case span of
-    Span (Command (cmd, pos), toks) =>
-      if Keyword.is_theory_load cmd then
-        (case find_file (clean_tokens toks) of
-          NONE => error ("Bad file argument of command " ^ quote cmd ^ Position.here pos)
-        | SOME (i, path) =>
-            let
-              val toks' = toks |> map_index (fn (j, tok) =>
-                if i = j then Token.put_files (read_files cmd path) tok
-                else tok);
-            in Span (Command (cmd, pos), toks') end)
-      else span
-  | _ => span);
-
-end;
+val present_span = implode o map present_token o Command_Span.content;
 
 
 
@@ -174,9 +71,9 @@
 
 (* scanning spans *)
 
-val eof = Span (Command ("", Position.none), []);
+val eof = Command_Span.Span (Command_Span.Command_Span ("", Position.none), []);
 
-fun is_eof (Span (Command ("", _), _)) = true
+fun is_eof (Command_Span.Span (Command_Span.Command_Span ("", _), _)) = true
   | is_eof _ = false;
 
 val not_eof = not o is_eof;
@@ -189,10 +86,13 @@
 local
 
 fun command_with pred =
-  Scan.one (fn (Span (Command (name, _), _)) => pred name | _ => false);
+  Scan.one
+    (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) => pred name | _ => false);
 
 val proof_atom =
-  Scan.one (fn (Span (Command (name, _), _)) => Keyword.is_proof_body name | _ => true) >> atom;
+  Scan.one
+    (fn (Command_Span.Span (Command_Span.Command_Span (name, _), _)) => Keyword.is_proof_body name
+      | _ => true) >> atom;
 
 fun proof_element x = (command_with Keyword.is_proof_goal -- proof_rest >> element || proof_atom) x
 and proof_rest x = (Scan.repeat proof_element -- command_with Keyword.is_qed) x;
--- a/src/Pure/Thy/thy_syntax.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Thy/thy_syntax.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -13,93 +13,6 @@
 
 object Thy_Syntax
 {
-  /** nested structure **/
-
-  object Structure
-  {
-    sealed abstract class Entry { def length: Int }
-    case class Block(val name: String, val body: List[Entry]) extends Entry
-    {
-      val length: Int = (0 /: body)(_ + _.length)
-    }
-    case class Atom(val command: Command) extends Entry
-    {
-      def length: Int = command.length
-    }
-
-    def parse(syntax: Outer_Syntax, node_name: Document.Node.Name, text: CharSequence): Entry =
-    {
-      /* stack operations */
-
-      def buffer(): mutable.ListBuffer[Entry] = new mutable.ListBuffer[Entry]
-      var stack: List[(Int, String, mutable.ListBuffer[Entry])] =
-        List((0, node_name.toString, buffer()))
-
-      @tailrec def close(level: Int => Boolean)
-      {
-        stack match {
-          case (lev, name, body) :: (_, _, body2) :: rest if level(lev) =>
-            body2 += Block(name, body.toList)
-            stack = stack.tail
-            close(level)
-          case _ =>
-        }
-      }
-
-      def result(): Entry =
-      {
-        close(_ => true)
-        val (_, name, body) = stack.head
-        Block(name, body.toList)
-      }
-
-      def add(command: Command)
-      {
-        syntax.heading_level(command) match {
-          case Some(i) =>
-            close(_ > i)
-            stack = (i + 1, command.source, buffer()) :: stack
-          case None =>
-        }
-        stack.head._3 += Atom(command)
-      }
-
-
-      /* result structure */
-
-      val spans = parse_spans(syntax.scan(text))
-      spans.foreach(span => add(Command(Document_ID.none, node_name, Nil, span)))
-      result()
-    }
-  }
-
-
-
-  /** parse spans **/
-
-  def parse_spans(toks: List[Token]): List[List[Token]] =
-  {
-    val result = new mutable.ListBuffer[List[Token]]
-    val span = new mutable.ListBuffer[Token]
-    val improper = new mutable.ListBuffer[Token]
-
-    def flush()
-    {
-      if (!span.isEmpty) { result += span.toList; span.clear }
-      if (!improper.isEmpty) { result += improper.toList; improper.clear }
-    }
-    for (tok <- toks) {
-      if (tok.is_command) { flush(); span += tok }
-      else if (tok.is_improper) improper += tok
-      else { span ++= improper; improper.clear; span += tok }
-    }
-    flush()
-
-    result.toList
-  }
-
-
-
   /** perspective **/
 
   def command_perspective(
@@ -231,58 +144,12 @@
   }
 
 
-  /* inlined files */
-
-  private def find_file(tokens: List[Token]): Option[String] =
-  {
-    def clean(toks: List[Token]): List[Token] =
-      toks match {
-        case t :: _ :: ts if t.is_keyword && (t.source == "%" || t.source == "--") => clean(ts)
-        case t :: ts => t :: clean(ts)
-        case Nil => Nil
-      }
-    clean(tokens.filter(_.is_proper)) match {
-      case tok :: toks if tok.is_command => toks.find(_.is_name).map(_.content)
-      case _ => None
-    }
-  }
-
-  def span_files(syntax: Prover.Syntax, span: List[Token]): List[String] =
-    syntax.load(span) match {
-      case Some(exts) =>
-        find_file(span) match {
-          case Some(file) =>
-            if (exts.isEmpty) List(file)
-            else exts.map(ext => file + "." + ext)
-          case None => Nil
-        }
-      case None => Nil
-    }
-
-  def resolve_files(
-      resources: Resources,
-      syntax: Prover.Syntax,
-      node_name: Document.Node.Name,
-      span: List[Token],
-      get_blob: Document.Node.Name => Option[Document.Blob])
-    : List[Command.Blob] =
-  {
-    span_files(syntax, span).map(file_name =>
-      Exn.capture {
-        val name =
-          Document.Node.Name(resources.append(node_name.master_dir, Path.explode(file_name)))
-        val blob = get_blob(name).map(blob => ((blob.bytes.sha1_digest, blob.chunk)))
-        (name, blob)
-      })
-  }
-
-
   /* reparse range of command spans */
 
   @tailrec private def chop_common(
       cmds: List[Command],
-      blobs_spans: List[(List[Command.Blob], List[Token])])
-    : (List[Command], List[(List[Command.Blob], List[Token])]) =
+      blobs_spans: List[(List[Command.Blob], Command_Span.Span)])
+    : (List[Command], List[(List[Command.Blob], Command_Span.Span)]) =
   {
     (cmds, blobs_spans) match {
       case (cmd :: cmds, (blobs, span) :: rest) if cmd.blobs == blobs && cmd.span == span =>
@@ -301,8 +168,8 @@
   {
     val cmds0 = commands.iterator(first, last).toList
     val blobs_spans0 =
-      parse_spans(syntax.scan(cmds0.iterator.map(_.source).mkString)).
-        map(span => (resolve_files(resources, syntax, name, span, get_blob), span))
+      syntax.parse_spans(cmds0.iterator.map(_.source).mkString).
+        map(span => (Command_Span.resolve_files(resources, syntax, name, span, get_blob), span))
 
     val (cmds1, blobs_spans1) = chop_common(cmds0, blobs_spans0)
 
@@ -337,8 +204,8 @@
     val visible = perspective.commands.toSet
 
     def next_invisible_command(cmds: Linear_Set[Command], from: Command): Command =
-      cmds.iterator(from).dropWhile(cmd => !cmd.is_command || visible(cmd))
-        .find(_.is_command) getOrElse cmds.last
+      cmds.iterator(from).dropWhile(cmd => !cmd.is_proper || visible(cmd))
+        .find(_.is_proper) getOrElse cmds.last
 
     @tailrec def recover(cmds: Linear_Set[Command]): Linear_Set[Command] =
       cmds.find(_.is_unparsed) match {
--- a/src/Pure/Tools/build.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Tools/build.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -1,6 +1,6 @@
 /*  Title:      Pure/Tools/build.scala
     Author:     Makarius
-    Options:    :folding=explicit:collapseFolds=1:
+    Options:    :folding=explicit:
 
 Build and manage Isabelle sessions.
 */
--- a/src/Pure/Tools/find_consts.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Tools/find_consts.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -143,7 +143,7 @@
 in
 
 fun read_query pos str =
-  Outer_Syntax.scan pos str
+  Outer_Syntax.scan (Keyword.get_lexicons ()) pos str
   |> filter Token.is_proper
   |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query --| Scan.ahead Parse.eof)))
   |> #1;
--- a/src/Pure/Tools/find_theorems.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Tools/find_theorems.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -528,7 +528,7 @@
 in
 
 fun read_query pos str =
-  Outer_Syntax.scan pos str
+  Outer_Syntax.scan (Keyword.get_lexicons ()) pos str
   |> filter Token.is_proper
   |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query --| Scan.ahead Parse.eof)))
   |> #1;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Tools/named_theorems.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,91 @@
+(*  Title:      Pure/Tools/named_theorems.ML
+    Author:     Makarius
+
+Named collections of theorems in canonical order.
+*)
+
+signature NAMED_THEOREMS =
+sig
+  val member: Proof.context -> string -> thm -> bool
+  val get: Proof.context -> string -> thm list
+  val add_thm: string -> thm -> Context.generic -> Context.generic
+  val del_thm: string -> thm -> Context.generic -> Context.generic
+  val add: string -> attribute
+  val del: string -> attribute
+  val declare: binding -> string -> local_theory -> string * local_theory
+end;
+
+structure Named_Theorems: NAMED_THEOREMS =
+struct
+
+(* context data *)
+
+structure Data = Generic_Data
+(
+  type T = thm Item_Net.T Symtab.table;
+  val empty: T = Symtab.empty;
+  val extend = I;
+  val merge : T * T -> T = Symtab.join (K Item_Net.merge);
+);
+
+fun new_entry name =
+  Data.map (fn data =>
+    if Symtab.defined data name
+    then error ("Duplicate declaration of named theorems: " ^ quote name)
+    else Symtab.update (name, Thm.full_rules) data);
+
+fun the_entry context name =
+  (case Symtab.lookup (Data.get context) name of
+    NONE => error ("Undeclared named theorems " ^ quote name)
+  | SOME entry => entry);
+
+fun map_entry name f context =
+  (the_entry context name; Data.map (Symtab.map_entry name f) context);
+
+
+(* maintain content *)
+
+fun member ctxt = Item_Net.member o the_entry (Context.Proof ctxt);
+
+fun content context = rev o Item_Net.content o the_entry context;
+val get = content o Context.Proof;
+
+fun add_thm name = map_entry name o Item_Net.update;
+fun del_thm name = map_entry name o Item_Net.remove;
+
+val add = Thm.declaration_attribute o add_thm;
+val del = Thm.declaration_attribute o del_thm;
+
+
+(* declaration *)
+
+fun declare binding descr lthy =
+  let
+    val name = Name_Space.full_name (Local_Theory.naming_of lthy) binding;
+    val description =
+      "declaration of " ^ (if descr = "" then Binding.name_of binding ^ " rules" else descr);
+    val lthy' = lthy
+      |> Local_Theory.background_theory (Context.theory_map (new_entry name))
+      |> Local_Theory.map_contexts (K (Context.proof_map (new_entry name)))
+      |> Local_Theory.add_thms_dynamic (binding, fn context => content context name)
+      |> Attrib.local_setup binding (Attrib.add_del (add name) (del name)) description
+  in (name, lthy') end;
+
+val _ =
+  Outer_Syntax.local_theory @{command_spec "named_theorems"}
+    "declare named collection of theorems"
+    (Parse.binding -- Scan.optional Parse.text "" >> (fn (b, descr) => snd o declare b descr));
+
+
+(* ML antiquotation *)
+
+val _ = Theory.setup
+  (ML_Antiquotation.inline @{binding named_theorems}
+    (Args.context -- Scan.lift (Parse.position Args.name) >> (fn (ctxt, (xname, pos)) =>
+      let
+        val thy = Proof_Context.theory_of ctxt;
+        val name = Global_Theory.check_fact thy (xname, pos);
+        val _ = get ctxt name handle ERROR msg => cat_error msg (Position.here pos);
+      in ML_Syntax.print_string name end)));
+
+end;
--- a/src/Pure/Tools/rule_insts.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/Tools/rule_insts.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -64,12 +64,6 @@
     val t' = f t;
   in if t aconv t' then NONE else SOME (t, t') end;
 
-val add_used =
-  (Thm.fold_terms o fold_types o fold_atyps)
-    (fn TFree (a, _) => insert (op =) a
-      | TVar ((a, _), _) => insert (op =) a
-      | _ => I);
-
 in
 
 fun read_termTs ctxt ss Ts =
@@ -131,8 +125,7 @@
   let
     val ctxt' = ctxt
       |> Proof_Context.read_vars fixes |-> Proof_Context.add_fixes |> #2
-      |> Variable.declare_thm thm
-      |> fold (fn a => Variable.declare_names (Logic.mk_type (TFree (a, dummyS)))) (add_used thm []);  (* FIXME !? *)
+      |> Variable.declare_thm thm;
     val tvars = Thm.fold_terms Term.add_tvars thm [];
     val vars = Thm.fold_terms Term.add_vars thm [];
     val insts = read_insts ctxt' mixed_insts (tvars, vars);
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Tools/thm_deps.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,124 @@
+(*  Title:      Pure/Tools/thm_deps.ML
+    Author:     Stefan Berghofer, TU Muenchen
+
+Visualize dependencies of theorems.
+*)
+
+signature THM_DEPS =
+sig
+  val thm_deps: theory -> thm list -> unit
+  val unused_thms: theory list * theory list -> (string * thm) list
+end;
+
+structure Thm_Deps: THM_DEPS =
+struct
+
+(* thm_deps *)
+
+fun thm_deps thy thms =
+  let
+    fun add_dep ("", _, _) = I
+      | add_dep (name, _, PBody {thms = thms', ...}) =
+          let
+            val prefix = #1 (split_last (Long_Name.explode name));
+            val session =
+              (case prefix of
+                a :: _ =>
+                  (case try (Context.get_theory thy) a of
+                    SOME thy =>
+                      (case Present.session_name thy of
+                        "" => []
+                      | session => [session])
+                  | NONE => [])
+               | _ => ["global"]);
+            val parents = filter_out (fn s => s = "") (map (#1 o #2) thms');
+            val entry =
+              {name = Long_Name.base_name name,
+               ID = name,
+               dir = space_implode "/" (session @ prefix),
+               unfold = false,
+               path = "",
+               parents = parents,
+               content = []};
+          in cons entry end;
+    val deps = Proofterm.fold_body_thms add_dep (Thm.proof_bodies_of thms) [];
+  in Graph_Display.display_graph (sort_wrt #ID deps) end;
+
+val _ =
+  Outer_Syntax.improper_command @{command_spec "thm_deps"} "visualize theorem dependencies"
+    (Parse.xthms1 >> (fn args =>
+      Toplevel.unknown_theory o Toplevel.keep (fn state =>
+        thm_deps (Toplevel.theory_of state) (Attrib.eval_thms (Toplevel.context_of state) args))));
+
+
+(* unused_thms *)
+
+fun unused_thms (base_thys, thys) =
+  let
+    fun add_fact space (name, ths) =
+      if exists (fn thy => Global_Theory.defined_fact thy name) base_thys then I
+      else
+        let val {concealed, group, ...} = Name_Space.the_entry space name in
+          fold_rev (fn th =>
+            (case Thm.derivation_name th of
+              "" => I
+            | a => cons (a, (th, concealed, group)))) ths
+        end;
+    fun add_facts facts = Facts.fold_static (add_fact (Facts.space_of facts)) facts;
+
+    val new_thms =
+      fold (add_facts o Global_Theory.facts_of) thys []
+      |> sort_distinct (string_ord o pairself #1);
+
+    val used =
+      Proofterm.fold_body_thms
+        (fn (a, _, _) => a <> "" ? Symtab.update (a, ()))
+        (map Proofterm.strip_thm (Thm.proof_bodies_of (map (#1 o #2) new_thms)))
+        Symtab.empty;
+
+    fun is_unused a = not (Symtab.defined used a);
+
+    (*groups containing at least one used theorem*)
+    val used_groups = fold (fn (a, (_, _, group)) =>
+      if is_unused a then I
+      else
+        (case group of
+          NONE => I
+        | SOME grp => Inttab.update (grp, ()))) new_thms Inttab.empty;
+
+    val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, seen_groups) =>
+      if not concealed andalso
+        (* FIXME replace by robust treatment of thm groups *)
+        member (op =) [Thm.theoremK, Thm.lemmaK, Thm.corollaryK] (Thm.legacy_get_kind th) andalso
+        is_unused a
+      then
+        (case group of
+           NONE => ((a, th) :: thms, seen_groups)
+         | SOME grp =>
+             if Inttab.defined used_groups grp orelse
+               Inttab.defined seen_groups grp then q
+             else ((a, th) :: thms, Inttab.update (grp, ()) seen_groups))
+      else q) new_thms ([], Inttab.empty);
+  in rev thms' end;
+
+val _ =
+  Outer_Syntax.improper_command @{command_spec "unused_thms"} "find unused theorems"
+    (Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --
+       Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >> (fn opt_range =>
+        Toplevel.keep (fn state =>
+          let
+            val thy = Toplevel.theory_of state;
+            val ctxt = Toplevel.context_of state;
+            fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]);
+            val get_theory = Context.get_theory thy;
+          in
+            unused_thms
+              (case opt_range of
+                NONE => (Theory.parents_of thy, [thy])
+              | SOME (xs, NONE) => (map get_theory xs, [thy])
+              | SOME (xs, SOME ys) => (map get_theory xs, map get_theory ys))
+            |> map pretty_thm |> Pretty.writeln_chunks
+          end)));
+
+end;
+
--- a/src/Pure/build-jars	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/build-jars	Wed Aug 27 15:52:58 2014 +0200
@@ -20,7 +20,7 @@
   GUI/gui.scala
   GUI/gui_thread.scala
   GUI/html5_panel.scala
-  GUI/jfx_thread.scala
+  GUI/jfx_gui.scala
   GUI/popup.scala
   GUI/system_dialog.scala
   GUI/wrap_panel.scala
@@ -54,6 +54,7 @@
   Isar/token.scala
   ML/ml_lex.scala
   PIDE/command.scala
+  PIDE/command_span.scala
   PIDE/document.scala
   PIDE/document_id.scala
   PIDE/editor.scala
@@ -71,7 +72,6 @@
   System/command_line.scala
   System/invoke_scala.scala
   System/isabelle_charset.scala
-  System/isabelle_font.scala
   System/isabelle_process.scala
   System/isabelle_system.scala
   System/options.scala
@@ -83,6 +83,7 @@
   Thy/present.scala
   Thy/thy_header.scala
   Thy/thy_info.scala
+  Thy/thy_structure.scala
   Thy/thy_syntax.scala
   Tools/build.scala
   Tools/build_console.scala
--- a/src/Pure/facts.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/facts.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -22,13 +22,15 @@
   type T
   val empty: T
   val space_of: T -> Name_Space.T
+  val alias: Name_Space.naming -> binding -> string -> T -> T
   val is_concealed: T -> string -> bool
   val check: Context.generic -> T -> xstring * Position.T -> string
   val intern: T -> xstring -> string
   val extern: Proof.context -> T -> string -> xstring
   val markup_extern: Proof.context -> T -> string -> Markup.T * xstring
   val lookup: Context.generic -> T -> string -> (bool * thm list) option
-  val retrieve: Context.generic -> T -> xstring * Position.T -> string * thm list
+  val retrieve: Context.generic -> T -> xstring * Position.T ->
+    {name: string, static: bool, thms: thm list}
   val defined: T -> string -> bool
   val fold_static: (string * thm list -> 'a -> 'a) -> T -> 'a -> 'a
   val dest_static: bool -> T list -> T -> (string * thm list) list
@@ -143,6 +145,9 @@
 
 val space_of = Name_Space.space_of_table o facts_of;
 
+fun alias naming binding name (Facts {facts, props}) =
+  make_facts (Name_Space.alias_table naming binding name facts) props;
+
 val is_concealed = Name_Space.is_concealed o space_of;
 
 fun check context facts (xname, pos) =
@@ -156,7 +161,7 @@
 
 val intern = Name_Space.intern o space_of;
 fun extern ctxt = Name_Space.extern ctxt o space_of;
-fun markup_extern ctxt = Name_Space.markup_extern ctxt o space_of
+fun markup_extern ctxt = Name_Space.markup_extern ctxt o space_of;
 
 
 (* retrieve *)
@@ -172,14 +177,18 @@
 fun retrieve context facts (xname, pos) =
   let
     val name = check context facts (xname, pos);
-    val thms =
+    val (static, thms) =
       (case lookup context facts name of
         SOME (static, thms) =>
           (if static then ()
            else Context_Position.report_generic context pos (Markup.dynamic_fact name);
-           thms)
+           (static, thms))
       | NONE => error ("Unknown fact " ^ quote name ^ Position.here pos));
-  in (name, map (Thm.transfer (Context.theory_of context)) thms) end;
+  in
+   {name = name,
+    static = static,
+    thms = map (Thm.transfer (Context.theory_of context)) thms}
+  end;
 
 
 (* static content *)
--- a/src/Pure/global_theory.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/global_theory.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -10,6 +10,7 @@
   val check_fact: theory -> xstring * Position.T -> string
   val intern_fact: theory -> xstring -> string
   val defined_fact: theory -> string -> bool
+  val alias_fact: binding -> string -> theory -> theory
   val hide_fact: bool -> string -> theory -> theory
   val get_thms: theory -> xstring -> thm list
   val get_thm: theory -> xstring -> thm
@@ -28,6 +29,8 @@
   val add_thms: ((binding * thm) * attribute list) list -> theory -> thm list * theory
   val add_thm: (binding * thm) * attribute list -> theory -> thm * theory
   val add_thmss: ((binding * thm list) * attribute list) list -> theory -> thm list list * theory
+  val add_thms_dynamic': Context.generic -> binding * (Context.generic -> thm list) ->
+    theory -> string * theory
   val add_thms_dynamic: binding * (Context.generic -> thm list) -> theory -> theory
   val note_thmss: string -> (Thm.binding * (thm list * attribute list) list) list
     -> theory -> (string * thm list) list * theory
@@ -60,13 +63,16 @@
 val intern_fact = Facts.intern o facts_of;
 val defined_fact = Facts.defined o facts_of;
 
+fun alias_fact binding name thy =
+  Data.map (Facts.alias (Sign.naming_of thy) binding name) thy;
+
 fun hide_fact fully name = Data.map (Facts.hide fully name);
 
 
 (* retrieve theorems *)
 
 fun get_thms thy xname =
-  #2 (Facts.retrieve (Context.Theory thy) (facts_of thy) (xname, Position.none));
+  #thms (Facts.retrieve (Context.Theory thy) (facts_of thy) (xname, Position.none));
 
 fun get_thm thy xname =
   Facts.the_single (xname, Position.none) (get_thms thy xname);
@@ -153,10 +159,14 @@
 val add_thm = yield_singleton add_thms;
 
 
-(* add_thms_dynamic *)
+(* dynamic theorems *)
 
-fun add_thms_dynamic (b, f) thy = thy
-  |> Data.map (Facts.add_dynamic (Context.Theory thy) (b, f) #> snd);
+fun add_thms_dynamic' context arg thy =
+  let val (name, facts') = Facts.add_dynamic context arg (Data.get thy)
+  in (name, Data.put facts' thy) end;
+
+fun add_thms_dynamic arg thy =
+  add_thms_dynamic' (Context.Theory thy) arg thy |> snd;
 
 
 (* note_thmss *)
--- a/src/Pure/goal.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/goal.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -11,8 +11,6 @@
   val PREFER_GOAL: tactic -> int -> tactic
   val CONJUNCTS: tactic -> int -> tactic
   val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic
-  val PARALLEL_CHOICE: tactic list -> tactic
-  val PARALLEL_GOALS: tactic -> tactic
 end;
 
 signature GOAL =
@@ -340,50 +338,6 @@
       etac (Raw_Simplifier.norm_hhf ctxt (Thm.trivial R)) THEN_ALL_NEW assume_tac);
   in fold_rev (curry op APPEND') tacs (K no_tac) i end);
 
-
-
-(** parallel tacticals **)
-
-(* parallel choice of single results *)
-
-fun PARALLEL_CHOICE tacs st =
-  (case Par_List.get_some (fn tac => SINGLE tac st) tacs of
-    NONE => Seq.empty
-  | SOME st' => Seq.single st');
-
-
-(* parallel refinement of non-schematic goal by single results *)
-
-local
-
-exception FAILED of unit;
-
-fun retrofit st' =
-  rotate_prems ~1 #>
-  Thm.bicompose {flatten = false, match = false, incremented = false}
-      (false, conclude st', Thm.nprems_of st') 1;
-
-in
-
-fun PARALLEL_GOALS tac =
-  Thm.adjust_maxidx_thm ~1 #>
-  (fn st =>
-    if not (Multithreading.enabled ()) orelse Thm.maxidx_of st >= 0 orelse Thm.nprems_of st <= 1
-    then DETERM tac st
-    else
-      let
-        fun try_tac g =
-          (case SINGLE tac g of
-            NONE => raise FAILED ()
-          | SOME g' => g');
-
-        val goals = Drule.strip_imp_prems (Thm.cprop_of st);
-        val results = Par_List.map (try_tac o init) goals;
-      in EVERY (map retrofit (rev results)) st end
-      handle FAILED () => Seq.empty);
-
-end;
-
 end;
 
 structure Basic_Goal: BASIC_GOAL = Goal;
--- a/src/Pure/library.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/library.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -32,7 +32,6 @@
   exception ERROR of string
   val error: string -> 'a
   val cat_error: string -> string -> 'a
-  val assert_all: ('a -> bool) -> 'a list -> ('a -> string) -> unit
 
   (*pairs*)
   val pair: 'a -> 'b -> 'a * 'b
@@ -269,12 +268,6 @@
   | cat_error msg "" = error msg
   | cat_error msg1 msg2 = error (msg1 ^ "\n" ^ msg2);
 
-fun assert_all pred list msg =
-  let
-    fun ass [] = ()
-      | ass (x :: xs) = if pred x then ass xs else error (msg x);
-  in ass list end;
-
 
 (* pairs *)
 
@@ -649,14 +642,14 @@
 
 local
   val zero = ord "0";
-  val small = 10000: int;
-  val small_table = Vector.tabulate (small, Int.toString);
+  val small_int = 10000: int;
+  val small_int_table = Vector.tabulate (small_int, Int.toString);
 in
 
 fun string_of_int i =
   if i < 0 then Int.toString i
   else if i < 10 then chr (zero + i)
-  else if i < small then Vector.sub (small_table, i)
+  else if i < small_int then Vector.sub (small_int_table, i)
   else Int.toString i;
 
 end;
--- a/src/Pure/library.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/library.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -32,6 +32,33 @@
     error(cat_message(msg1, msg2))
 
 
+  /* integers */
+
+  private val small_int = 10000
+  private lazy val small_int_table =
+  {
+    val array = new Array[String](small_int)
+    for (i <- 0 until small_int) array(i) = i.toString
+    array
+  }
+
+  def is_small_int(s: String): Boolean =
+  {
+    val len = s.length
+    1 <= len && len <= 4 &&
+    s.forall(c => '0' <= c && c <= '9') &&
+    (len == 1 || s(0) != '0')
+  }
+
+  def signed_string_of_long(i: Long): String =
+    if (0 <= i && i < small_int) small_int_table(i.toInt)
+    else i.toString
+
+  def signed_string_of_int(i: Int): String =
+    if (0 <= i && i < small_int) small_int_table(i)
+    else i.toString
+
+
   /* separated chunks */
 
   def separate[A](s: A, list: List[A]): List[A] =
--- a/src/Pure/more_thm.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/more_thm.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -234,7 +234,7 @@
 fun is_dummy thm =
   (case try Logic.dest_term (Thm.concl_of thm) of
     NONE => false
-  | SOME t => Term.is_dummy_pattern t);
+  | SOME t => Term.is_dummy_pattern (Term.head_of t));
 
 fun plain_prop_of raw_thm =
   let
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/par_tactical.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -0,0 +1,68 @@
+(*  Title:      Pure/par_tactical.ML
+    Author:     Makarius
+
+Parallel tacticals.
+*)
+
+signature BASIC_PAR_TACTICAL =
+sig
+  val PARALLEL_CHOICE: tactic list -> tactic
+  val PARALLEL_GOALS: tactic -> tactic
+  val PARALLEL_ALLGOALS: (int -> tactic) -> tactic
+end;
+
+signature PAR_TACTICAL =
+sig
+  include BASIC_PAR_TACTICAL
+end;
+
+structure Par_Tactical: PAR_TACTICAL =
+struct
+
+(* parallel choice of single results *)
+
+fun PARALLEL_CHOICE tacs st =
+  (case Par_List.get_some (fn tac => SINGLE tac st) tacs of
+    NONE => Seq.empty
+  | SOME st' => Seq.single st');
+
+
+(* parallel refinement of non-schematic goal by single results *)
+
+local
+
+exception FAILED of unit;
+
+fun retrofit st' =
+  rotate_prems ~1 #>
+  Thm.bicompose {flatten = false, match = false, incremented = false}
+      (false, Goal.conclude st', Thm.nprems_of st') 1;
+
+in
+
+fun PARALLEL_GOALS tac =
+  Thm.adjust_maxidx_thm ~1 #>
+  (fn st =>
+    if not (Multithreading.enabled ()) orelse Thm.maxidx_of st >= 0 orelse Thm.nprems_of st <= 1
+    then DETERM tac st
+    else
+      let
+        fun try_tac g =
+          (case SINGLE tac g of
+            NONE => raise FAILED ()
+          | SOME g' => g');
+
+        val goals = Drule.strip_imp_prems (Thm.cprop_of st);
+        val results = Par_List.map (try_tac o Goal.init) goals;
+      in EVERY (map retrofit (rev results)) st end
+      handle FAILED () => Seq.empty);
+
+end;
+
+val PARALLEL_ALLGOALS = PARALLEL_GOALS o ALLGOALS;
+
+end;
+
+structure Basic_Par_Tactical: BASIC_PAR_TACTICAL = Par_Tactical;
+open Basic_Par_Tactical;
+
--- a/src/Pure/simplifier.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Pure/simplifier.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -332,21 +332,23 @@
 (** method syntax **)
 
 val cong_modifiers =
- [Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier),
-  Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add),
-  Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)];
+ [Args.$$$ congN -- Args.colon >> K (Method.modifier cong_add @{here}),
+  Args.$$$ congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
+  Args.$$$ congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here})];
 
 val simp_modifiers =
- [Args.$$$ simpN -- Args.colon >> K (I, simp_add),
-  Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add),
-  Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del),
-  Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >> K (Raw_Simplifier.clear_simpset, simp_add)]
+ [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
+  Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
+  Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
+  Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >>
+    K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = @{here}}]
    @ cong_modifiers;
 
 val simp_modifiers' =
- [Args.add -- Args.colon >> K (I, simp_add),
-  Args.del -- Args.colon >> K (I, simp_del),
-  Args.$$$ onlyN -- Args.colon >> K (Raw_Simplifier.clear_simpset, simp_add)]
+ [Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
+  Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
+  Args.$$$ onlyN -- Args.colon >>
+    K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = @{here}}]
    @ cong_modifiers;
 
 val simp_options =
@@ -374,7 +376,7 @@
   Method.setup @{binding simp_all}
     (simp_method more_mods (fn ctxt => fn tac => fn facts =>
       ALLGOALS (Method.insert_tac facts) THEN
-        (CHANGED_PROP o PARALLEL_GOALS o ALLGOALS o tac) ctxt))
+        (CHANGED_PROP o PARALLEL_ALLGOALS o tac) ctxt))
     "simplification (all goals)";
 
 fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn ctxt0 =>
--- a/src/Sequents/prover.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Sequents/prover.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -106,8 +106,8 @@
 
 fun method tac =
   Method.sections
-   [Args.$$$ "add" -- Args.bang_colon >> K (I, safe_add),
-    Args.$$$ "add" -- Args.colon >> K (I, unsafe_add)]
+   [Args.$$$ "add" -- Args.bang_colon >> K (Method.modifier safe_add @{here}),
+    Args.$$$ "add" -- Args.colon >> K (Method.modifier unsafe_add @{here})]
   >> K (SIMPLE_METHOD' o tac);
 
 
--- a/src/Tools/Code/code_target.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/Code/code_target.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -695,7 +695,7 @@
   let
     val ctxt = Proof_Context.init_global (Thy_Info.get_theory thyname);
     val parse = Scan.read Token.stopper (Parse.!!! code_exprP) o
-      (filter Token.is_proper o Outer_Syntax.scan Position.none);
+      (filter Token.is_proper o Outer_Syntax.scan (Keyword.get_lexicons ()) Position.none);
   in case parse cmd_expr
    of SOME f => (writeln "Now generating code..."; f ctxt)
     | NONE => error ("Bad directive " ^ quote cmd_expr)
--- a/src/Tools/atomize_elim.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/atomize_elim.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -6,28 +6,22 @@
 
 signature ATOMIZE_ELIM =
 sig
-  val declare_atomize_elim : attribute  
-
   val atomize_elim_conv : Proof.context -> conv
   val full_atomize_elim_conv : Proof.context -> conv
-
   val atomize_elim_tac : Proof.context -> int -> tactic
-
-  val setup : theory -> theory
 end
 
-structure AtomizeElim : ATOMIZE_ELIM =
+structure Atomize_Elim : ATOMIZE_ELIM =
 struct
 
-(* theory data *)
+(* atomize_elim rules *)
 
-structure AtomizeElimData = Named_Thms
-(
-  val name = @{binding atomize_elim};
-  val description = "atomize_elim rewrite rule";
-);
+val named_theorems =
+  Context.>>> (Context.map_theory_result
+   (Named_Target.theory_init #>
+    Named_Theorems.declare @{binding atomize_elim} "atomize_elim rewrite rule" ##>
+    Local_Theory.exit_global));
 
-val declare_atomize_elim = AtomizeElimData.add
 
 (* More conversions: *)
 local open Conv in
@@ -50,7 +44,7 @@
     in Thm.equal_intr fwd back end
 
 
-(* convert !!thesis. (!!x y z. ... => thesis) ==> ... 
+(* convert !!thesis. (!!x y z. ... => thesis) ==> ...
                      ==> (!!a b c. ... => thesis)
                      ==> thesis
 
@@ -60,7 +54,7 @@
 *)
 fun atomize_elim_conv ctxt ct =
     (forall_conv (prems_conv ~1 o Object_Logic.atomize_prems o snd) ctxt
-    then_conv Raw_Simplifier.rewrite ctxt true (AtomizeElimData.get ctxt)
+    then_conv Raw_Simplifier.rewrite ctxt true (Named_Theorems.get ctxt named_theorems)
     then_conv (fn ct' => if can Object_Logic.dest_judgment ct'
                          then all_conv ct'
                          else raise CTERM ("atomize_elim", [ct', ct])))
@@ -71,7 +65,7 @@
 fun thesis_miniscope_conv inner_cv ctxt =
     let
       (* check if the outermost quantifier binds the conclusion *)
-      fun is_forall_thesis t = 
+      fun is_forall_thesis t =
           case Logic.strip_assums_concl t of
             (_ $ Bound i) => i = length (Logic.strip_params t) - 1
           | _ => false
@@ -93,11 +87,11 @@
           end
 
       (* move current quantifier to the right *)
-      fun moveq_conv ctxt = 
+      fun moveq_conv ctxt =
           (rewr_conv @{thm swap_params} then_conv (forall_conv (moveq_conv o snd) ctxt))
           else_conv (movea_conv ctxt)
 
-      fun ms_conv ctxt ct = 
+      fun ms_conv ctxt ct =
           if is_forall_thesis (term_of ct)
           then moveq_conv ctxt ct
           else (implies_concl_conv (ms_conv ctxt)
@@ -105,7 +99,7 @@
             (forall_conv (ms_conv o snd) ctxt))
             ct
     in
-      ms_conv ctxt 
+      ms_conv ctxt
     end
 
 val full_atomize_elim_conv = thesis_miniscope_conv atomize_elim_conv
@@ -117,13 +111,13 @@
     let
       val thy = Proof_Context.theory_of ctxt
       val _ $ thesis = Logic.strip_assums_concl subg
-                       
+
       (* Introduce a quantifier first if the thesis is not bound *)
-      val quantify_thesis = 
+      val quantify_thesis =
           if is_Bound thesis then all_tac
           else let
               val cthesis = cterm_of thy thesis
-              val rule = instantiate' [SOME (ctyp_of_term cthesis)] [NONE, SOME cthesis] 
+              val rule = instantiate' [SOME (ctyp_of_term cthesis)] [NONE, SOME cthesis]
                          @{thm meta_spec}
             in
               compose_tac (false, rule, 1) i
@@ -133,9 +127,9 @@
       THEN (CONVERSION (full_atomize_elim_conv ctxt) i)
     end)
 
-val setup =
-  Method.setup @{binding atomize_elim} (Scan.succeed (SIMPLE_METHOD' o atomize_elim_tac))
-    "convert obtains statement to atomic object logic goal"
-  #> AtomizeElimData.setup
+val _ =
+  Theory.setup
+    (Method.setup @{binding atomize_elim} (Scan.succeed (SIMPLE_METHOD' o atomize_elim_tac))
+      "convert obtains statement to atomic object logic goal")
 
 end
--- a/src/Tools/induct.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/induct.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -932,15 +932,13 @@
           (cases_tac ctxt (not no_simp) insts opt_rule facts)))))
     "case analysis on types or predicates/sets";
 
-fun gen_induct_setup binding itac =
+fun gen_induct_setup binding tac =
   Method.setup binding
     (Scan.lift (Args.mode no_simpN) --
       (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) --
         (arbitrary -- taking -- Scan.option induct_rule)) >>
-      (fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn ctxt =>
-        RAW_METHOD_CASES (fn facts =>
-          Seq.DETERM
-            (HEADGOAL (itac ctxt (not no_simp) insts arbitrary taking opt_rule facts)))))
+      (fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn ctxt => fn facts =>
+        Seq.DETERM (HEADGOAL (tac ctxt (not no_simp) insts arbitrary taking opt_rule facts))))
     "induction on types or predicates/sets";
 
 val induct_setup = gen_induct_setup @{binding induct} induct_tac;
@@ -948,9 +946,8 @@
 val coinduct_setup =
   Method.setup @{binding coinduct}
     (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >>
-      (fn ((insts, taking), opt_rule) => fn ctxt =>
-        RAW_METHOD_CASES (fn facts =>
-          Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts)))))
+      (fn ((insts, taking), opt_rule) => fn ctxt => fn facts =>
+        Seq.DETERM (HEADGOAL (coinduct_tac ctxt insts taking opt_rule facts))))
     "coinduction on types or predicates/sets";
 
 end;
--- a/src/Tools/intuitionistic.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/intuitionistic.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -73,18 +73,18 @@
 val elimN = "elim";
 val destN = "dest";
 
-fun modifier name kind kind' att =
+fun modifier name kind kind' att pos =
   Args.$$$ name |-- (kind >> K NONE || kind' |-- Parse.nat --| Args.colon >> SOME)
-    >> (pair (I: Proof.context -> Proof.context) o att);
+    >> (fn arg => Method.modifier (att arg) pos);
 
 val modifiers =
- [modifier destN Args.bang_colon Args.bang Context_Rules.dest_bang,
-  modifier destN Args.colon (Scan.succeed ()) Context_Rules.dest,
-  modifier elimN Args.bang_colon Args.bang Context_Rules.elim_bang,
-  modifier elimN Args.colon (Scan.succeed ()) Context_Rules.elim,
-  modifier introN Args.bang_colon Args.bang Context_Rules.intro_bang,
-  modifier introN Args.colon (Scan.succeed ()) Context_Rules.intro,
-  Args.del -- Args.colon >> K (I, Context_Rules.rule_del)];
+ [modifier destN Args.bang_colon Args.bang Context_Rules.dest_bang @{here},
+  modifier destN Args.colon (Scan.succeed ()) Context_Rules.dest @{here},
+  modifier elimN Args.bang_colon Args.bang Context_Rules.elim_bang @{here},
+  modifier elimN Args.colon (Scan.succeed ()) Context_Rules.elim @{here},
+  modifier introN Args.bang_colon Args.bang Context_Rules.intro_bang @{here},
+  modifier introN Args.colon (Scan.succeed ()) Context_Rules.intro @{here},
+  Args.del -- Args.colon >> K (Method.modifier Context_Rules.rule_del @{here})];
 
 in
 
--- a/src/Tools/jEdit/src/document_model.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/document_model.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -104,13 +104,10 @@
       val snapshot = this.snapshot()
 
       val document_view_ranges =
-        if (is_theory) {
-          for {
-            doc_view <- PIDE.document_views(buffer)
-            range <- doc_view.perspective(snapshot).ranges
-          } yield range
-        }
-        else Nil
+        for {
+          doc_view <- PIDE.document_views(buffer)
+          range <- doc_view.perspective(snapshot).ranges
+        } yield range
 
       val load_ranges =
         for {
--- a/src/Tools/jEdit/src/documentation_dockable.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/documentation_dockable.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -10,7 +10,7 @@
 import isabelle._
 
 import java.awt.{Dimension, GridLayout}
-import java.awt.event.{MouseEvent, MouseAdapter}
+import java.awt.event.{KeyEvent, KeyAdapter, MouseEvent, MouseAdapter}
 import javax.swing.{JTree, JScrollPane, JComponent}
 import javax.swing.tree.{DefaultMutableTreeNode, TreeSelectionModel}
 import javax.swing.event.{TreeSelectionEvent, TreeSelectionListener}
@@ -24,13 +24,13 @@
 
   private case class Documentation(name: String, title: String, path: Path)
   {
-    override def toString =
+    override def toString: String =
       "<html><b>" + HTML.encode(name) + "</b>:  " + HTML.encode(title) + "</html>"
   }
 
   private case class Text_File(name: String, path: Path)
   {
-    override def toString = name
+    override def toString: String = name
   }
 
   private val root = new DefaultMutableTreeNode
@@ -47,38 +47,57 @@
 
   private val tree = new JTree(root)
 
-  for (cond <-
-    List(JComponent.WHEN_FOCUSED,
-      JComponent.WHEN_ANCESTOR_OF_FOCUSED_COMPONENT,
-      JComponent.WHEN_IN_FOCUSED_WINDOW)) tree.setInputMap(cond, null)
+  override def focusOnDefaultComponent { tree.requestFocusInWindow }
 
   if (!OperatingSystem.isMacOSLF)
     tree.putClientProperty("JTree.lineStyle", "Angled")
   tree.getSelectionModel.setSelectionMode(TreeSelectionModel.SINGLE_TREE_SELECTION)
+
+  private def action(node: DefaultMutableTreeNode)
+  {
+    node.getUserObject match {
+      case Text_File(_, path) =>
+        PIDE.editor.goto_file(view, Isabelle_System.platform_path(path))
+      case Documentation(_, _, path) =>
+        if (path.is_file)
+          PIDE.editor.goto_file(view, Isabelle_System.platform_path(path))
+        else {
+          Future.fork {
+            try { Doc.view(path) }
+            catch {
+              case exn: Throwable =>
+                GUI.error_dialog(view,
+                  "Documentation error", GUI.scrollable_text(Exn.message(exn)))
+            }
+          }
+        }
+      case _ =>
+    }
+  }
+
+  tree.addKeyListener(new KeyAdapter {
+    override def keyPressed(e: KeyEvent)
+    {
+      if (e.getKeyCode == KeyEvent.VK_ENTER) {
+        e.consume
+        val path = tree.getSelectionPath
+        if (path != null) {
+          path.getLastPathComponent match {
+            case node: DefaultMutableTreeNode => action(node)
+            case _ =>
+          }
+        }
+      }
+    }
+  })
   tree.addMouseListener(new MouseAdapter {
-    override def mouseClicked(e: MouseEvent) {
+    override def mouseClicked(e: MouseEvent)
+    {
       val click = tree.getPathForLocation(e.getX, e.getY)
       if (click != null && e.getClickCount == 1) {
         (click.getLastPathComponent, tree.getLastSelectedPathComponent) match {
           case (node: DefaultMutableTreeNode, node1: DefaultMutableTreeNode) if node == node1 =>
-            node.getUserObject match {
-              case Text_File(_, path) =>
-                PIDE.editor.goto_file(view, Isabelle_System.platform_path(path))
-              case Documentation(_, _, path) =>
-                if (path.is_file)
-                  PIDE.editor.goto_file(view, Isabelle_System.platform_path(path))
-                else {
-                  Future.fork {
-                    try { Doc.view(path) }
-                    catch {
-                      case exn: Throwable =>
-                        GUI.error_dialog(view,
-                          "Documentation error", GUI.scrollable_text(Exn.message(exn)))
-                    }
-                  }
-                }
-              case _ =>
-            }
+            action(node)
           case _ =>
         }
       }
--- a/src/Tools/jEdit/src/isabelle_logic.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/isabelle_logic.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -24,7 +24,7 @@
 
   private class Logic_Entry(val name: String, val description: String)
   {
-    override def toString = description
+    override def toString: String = description
   }
 
   def logic_selector(autosave: Boolean): Option_Component =
--- a/src/Tools/jEdit/src/isabelle_sidekick.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/isabelle_sidekick.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -21,7 +21,7 @@
 object Isabelle_Sidekick
 {
   def int_to_pos(offset: Text.Offset): Position =
-    new Position { def getOffset = offset; override def toString = offset.toString }
+    new Position { def getOffset = offset; override def toString: String = offset.toString }
 
   class Asset(name: String, start: Text.Offset, end: Text.Offset) extends IAsset
   {
@@ -39,7 +39,7 @@
     override def setStart(start: Position) = _start = start
     override def getEnd: Position = _end
     override def setEnd(end: Position) = _end = end
-    override def toString = _name
+    override def toString: String = _name
   }
 
   def swing_markup_tree(tree: Markup_Tree, parent: DefaultMutableTreeNode,
@@ -95,13 +95,11 @@
     node_name: Buffer => Option[Document.Node.Name])
   extends Isabelle_Sidekick(name)
 {
-  import Thy_Syntax.Structure
-
   override def parser(buffer: Buffer, syntax: Outer_Syntax, data: SideKickParsedData): Boolean =
   {
-    def make_tree(offset: Text.Offset, entry: Structure.Entry): List[DefaultMutableTreeNode] =
+    def make_tree(offset: Text.Offset, entry: Thy_Structure.Entry): List[DefaultMutableTreeNode] =
       entry match {
-        case Structure.Block(name, body) =>
+        case Thy_Structure.Block(name, body) =>
           val node =
             new DefaultMutableTreeNode(
               new Isabelle_Sidekick.Asset(Library.first_line(name), offset, offset + entry.length))
@@ -111,8 +109,8 @@
               i + e.length
             })
           List(node)
-        case Structure.Atom(command)
-        if command.is_command && syntax.heading_level(command).isEmpty =>
+        case Thy_Structure.Atom(command)
+        if command.is_proper && syntax.heading_level(command).isEmpty =>
           val node =
             new DefaultMutableTreeNode(
               new Isabelle_Sidekick.Asset(command.name, offset, offset + entry.length))
@@ -123,7 +121,7 @@
     node_name(buffer) match {
       case Some(name) =>
         val text = JEdit_Lib.buffer_text(buffer)
-        val structure = Structure.parse(syntax, name, text)
+        val structure = Thy_Structure.parse(syntax, name, text)
         make_tree(0, structure) foreach (node => data.root.add(node))
         true
       case None => false
@@ -175,7 +173,7 @@
                   new Isabelle_Sidekick.Asset(command.toString, range.start, range.stop) {
                     override def getShortString: String = content
                     override def getLongString: String = info_text
-                    override def toString = quote(content) + " " + range.toString
+                    override def toString: String = quote(content) + " " + range.toString
                   })
               })
         }
--- a/src/Tools/jEdit/src/plugin.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/plugin.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -373,7 +373,7 @@
 
       PIDE.plugin = this
       Isabelle_System.init()
-      Isabelle_Font.install_fonts()
+      GUI.install_fonts()
 
       PIDE.options.update(Options.init())
       PIDE.completion_history.load()
--- a/src/Tools/jEdit/src/rendering.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/rendering.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -376,7 +376,7 @@
                 case Position.Def_Line_File(line, name) =>
                   val offset = Position.Def_Offset.unapply(props) getOrElse 0
                   PIDE.editor.hyperlink_source_file(name, line, offset)
-                case Position.Def_Id_Offset(id, offset) =>
+                case Position.Def_Id_Offset0(id, offset) =>
                   PIDE.editor.hyperlink_command_id(snapshot, id, offset)
                 case _ => None
               }
@@ -388,7 +388,7 @@
                 case Position.Line_File(line, name) =>
                   val offset = Position.Offset.unapply(props) getOrElse 0
                   PIDE.editor.hyperlink_source_file(name, line, offset)
-                case Position.Id_Offset(id, offset) =>
+                case Position.Id_Offset0(id, offset) =>
                   PIDE.editor.hyperlink_command_id(snapshot, id, offset)
                 case _ => None
               }
@@ -480,7 +480,9 @@
             Some(Text.Info(r, (t1 + t2, info)))
           case (prev, Text.Info(r, XML.Elem(Markup.Entity(kind, name), _))) =>
             val kind1 = Word.implode(Word.explode('_', kind))
-            val txt1 = kind1 + " " + quote(name)
+            val txt1 =
+              if (name == "") kind1
+              else kind1 + " " + quote(name)
             val t = prev.info._1
             val txt2 =
               if (kind == Markup.COMMAND && t.elapsed.seconds >= timing_threshold)
--- a/src/Tools/jEdit/src/symbols_dockable.scala	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/Tools/jEdit/src/symbols_dockable.scala	Wed Aug 27 15:52:58 2014 +0200
@@ -24,8 +24,8 @@
 
     font =
       Symbol.fonts.get(symbol) match {
-        case None => Isabelle_Font(size = font_size)
-        case Some(font_family) => Isabelle_Font(family = font_family, size = font_size)
+        case None => GUI.font(size = font_size)
+        case Some(font_family) => GUI.font(family = font_family, size = font_size)
       }
     action =
       Action(decoded) {
--- a/src/ZF/Bin.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/Bin.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -594,4 +594,92 @@
      "(integ_of(w) = x $* y) \<longleftrightarrow> (x $* y = integ_of(w))"
 by auto
 
+(** To simplify inequalities involving integer negation and literals,
+    such as -x = #3
+**)
+
+lemmas [simp] =
+  zminus_equation [where y = "integ_of(w)"]
+  equation_zminus [where x = "integ_of(w)"]
+  for w
+
+lemmas [iff] =
+  zminus_zless [where y = "integ_of(w)"]
+  zless_zminus [where x = "integ_of(w)"]
+  for w
+
+lemmas [iff] =
+  zminus_zle [where y = "integ_of(w)"]
+  zle_zminus [where x = "integ_of(w)"]
+  for w
+
+lemmas [simp] =
+  Let_def [where s = "integ_of(w)"] for w
+
+
+(*** Simprocs for numeric literals ***)
+
+(** Combining of literal coefficients in sums of products **)
+
+lemma zless_iff_zdiff_zless_0: "(x $< y) \<longleftrightarrow> (x$-y $< #0)"
+  by (simp add: zcompare_rls)
+
+lemma eq_iff_zdiff_eq_0: "[| x \<in> int; y \<in> int |] ==> (x = y) \<longleftrightarrow> (x$-y = #0)"
+  by (simp add: zcompare_rls)
+
+lemma zle_iff_zdiff_zle_0: "(x $<= y) \<longleftrightarrow> (x$-y $<= #0)"
+  by (simp add: zcompare_rls)
+
+
+(** For combine_numerals **)
+
+lemma left_zadd_zmult_distrib: "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k"
+  by (simp add: zadd_zmult_distrib zadd_ac)
+
+
+(** For cancel_numerals **)
+
+lemmas rel_iff_rel_0_rls =
+  zless_iff_zdiff_zless_0 [where y = "u $+ v"]
+  eq_iff_zdiff_eq_0 [where y = "u $+ v"]
+  zle_iff_zdiff_zle_0 [where y = "u $+ v"]
+  zless_iff_zdiff_zless_0 [where y = n]
+  eq_iff_zdiff_eq_0 [where y = n]
+  zle_iff_zdiff_zle_0 [where y = n]
+  for u v (* FIXME n (!?) *)
+
+lemma eq_add_iff1: "(i$*u $+ m = j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m = intify(n))"
+  apply (simp add: zdiff_def zadd_zmult_distrib)
+  apply (simp add: zcompare_rls)
+  apply (simp add: zadd_ac)
+  done
+
+lemma eq_add_iff2: "(i$*u $+ m = j$*u $+ n) \<longleftrightarrow> (intify(m) = (j$-i)$*u $+ n)"
+  apply (simp add: zdiff_def zadd_zmult_distrib)
+  apply (simp add: zcompare_rls)
+  apply (simp add: zadd_ac)
+  done
+
+lemma less_add_iff1: "(i$*u $+ m $< j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m $< n)"
+  apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls)
+  done
+
+lemma less_add_iff2: "(i$*u $+ m $< j$*u $+ n) \<longleftrightarrow> (m $< (j$-i)$*u $+ n)"
+  apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls)
+  done
+
+lemma le_add_iff1: "(i$*u $+ m $<= j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m $<= n)"
+  apply (simp add: zdiff_def zadd_zmult_distrib)
+  apply (simp add: zcompare_rls)
+  apply (simp add: zadd_ac)
+  done
+
+lemma le_add_iff2: "(i$*u $+ m $<= j$*u $+ n) \<longleftrightarrow> (m $<= (j$-i)$*u $+ n)"
+  apply (simp add: zdiff_def zadd_zmult_distrib)
+  apply (simp add: zcompare_rls)
+  apply (simp add: zadd_ac)
+  done
+
+ML_file "int_arith.ML"
+
 end
--- a/src/ZF/IntArith.thy	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,94 +0,0 @@
-
-theory IntArith imports Bin
-begin
-
-
-(** To simplify inequalities involving integer negation and literals,
-    such as -x = #3
-**)
-
-lemmas [simp] =
-  zminus_equation [where y = "integ_of(w)"]
-  equation_zminus [where x = "integ_of(w)"]
-  for w
-
-lemmas [iff] =
-  zminus_zless [where y = "integ_of(w)"]
-  zless_zminus [where x = "integ_of(w)"]
-  for w
-
-lemmas [iff] =
-  zminus_zle [where y = "integ_of(w)"]
-  zle_zminus [where x = "integ_of(w)"]
-  for w
-
-lemmas [simp] =
-  Let_def [where s = "integ_of(w)"] for w
-
-
-(*** Simprocs for numeric literals ***)
-
-(** Combining of literal coefficients in sums of products **)
-
-lemma zless_iff_zdiff_zless_0: "(x $< y) \<longleftrightarrow> (x$-y $< #0)"
-  by (simp add: zcompare_rls)
-
-lemma eq_iff_zdiff_eq_0: "[| x \<in> int; y \<in> int |] ==> (x = y) \<longleftrightarrow> (x$-y = #0)"
-  by (simp add: zcompare_rls)
-
-lemma zle_iff_zdiff_zle_0: "(x $<= y) \<longleftrightarrow> (x$-y $<= #0)"
-  by (simp add: zcompare_rls)
-
-
-(** For combine_numerals **)
-
-lemma left_zadd_zmult_distrib: "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k"
-  by (simp add: zadd_zmult_distrib zadd_ac)
-
-
-(** For cancel_numerals **)
-
-lemmas rel_iff_rel_0_rls =
-  zless_iff_zdiff_zless_0 [where y = "u $+ v"]
-  eq_iff_zdiff_eq_0 [where y = "u $+ v"]
-  zle_iff_zdiff_zle_0 [where y = "u $+ v"]
-  zless_iff_zdiff_zless_0 [where y = n]
-  eq_iff_zdiff_eq_0 [where y = n]
-  zle_iff_zdiff_zle_0 [where y = n]
-  for u v (* FIXME n (!?) *)
-
-lemma eq_add_iff1: "(i$*u $+ m = j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m = intify(n))"
-  apply (simp add: zdiff_def zadd_zmult_distrib)
-  apply (simp add: zcompare_rls)
-  apply (simp add: zadd_ac)
-  done
-
-lemma eq_add_iff2: "(i$*u $+ m = j$*u $+ n) \<longleftrightarrow> (intify(m) = (j$-i)$*u $+ n)"
-  apply (simp add: zdiff_def zadd_zmult_distrib)
-  apply (simp add: zcompare_rls)
-  apply (simp add: zadd_ac)
-  done
-
-lemma less_add_iff1: "(i$*u $+ m $< j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m $< n)"
-  apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls)
-  done
-
-lemma less_add_iff2: "(i$*u $+ m $< j$*u $+ n) \<longleftrightarrow> (m $< (j$-i)$*u $+ n)"
-  apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls)
-  done
-
-lemma le_add_iff1: "(i$*u $+ m $<= j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m $<= n)"
-  apply (simp add: zdiff_def zadd_zmult_distrib)
-  apply (simp add: zcompare_rls)
-  apply (simp add: zadd_ac)
-  done
-
-lemma le_add_iff2: "(i$*u $+ m $<= j$*u $+ n) \<longleftrightarrow> (m $<= (j$-i)$*u $+ n)"
-  apply (simp add: zdiff_def zadd_zmult_distrib)
-  apply (simp add: zcompare_rls)
-  apply (simp add: zadd_ac)
-  done
-
-ML_file "int_arith.ML"
-
-end
--- a/src/ZF/IntDiv_ZF.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/IntDiv_ZF.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -29,7 +29,9 @@
 
 header{*The Division Operators Div and Mod*}
 
-theory IntDiv_ZF imports IntArith OrderArith begin
+theory IntDiv_ZF
+imports Bin OrderArith
+begin
 
 definition
   quorem :: "[i,i] => o"  where
--- a/src/ZF/Tools/datatype_package.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/Tools/datatype_package.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -32,8 +32,8 @@
   val add_datatype_i: term * term list -> Ind_Syntax.constructor_spec list list ->
     thm list * thm list * thm list -> theory -> theory * inductive_result * datatype_result
   val add_datatype: string * string list -> (string * string list * mixfix) list list ->
-    (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list *
-    (Facts.ref * Attrib.src list) list -> theory -> theory * inductive_result * datatype_result
+    (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list *
+    (Facts.ref * Token.src list) list -> theory -> theory * inductive_result * datatype_result
 end;
 
 functor Add_datatype_def_Fun
@@ -437,9 +437,9 @@
     ("define " ^ coind_prefix ^ "datatype")
     ((Scan.optional ((@{keyword "\<subseteq>"} || @{keyword "<="}) |-- Parse.!!! Parse.term) "") --
       Parse.and_list1 (Parse.term -- (@{keyword "="} |-- Parse.enum1 "|" con_decl)) --
-      Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse_Spec.xthms1) [] --
-      Scan.optional (@{keyword "type_intros"} |-- Parse.!!! Parse_Spec.xthms1) [] --
-      Scan.optional (@{keyword "type_elims"} |-- Parse.!!! Parse_Spec.xthms1) []
+      Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse.xthms1) [] --
+      Scan.optional (@{keyword "type_intros"} |-- Parse.!!! Parse.xthms1) [] --
+      Scan.optional (@{keyword "type_elims"} |-- Parse.!!! Parse.xthms1) []
       >> (Toplevel.theory o mk_datatype));
 
 end;
--- a/src/ZF/Tools/induct_tacs.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/Tools/induct_tacs.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -12,8 +12,8 @@
   val induct_tac: Proof.context -> string -> int -> tactic
   val exhaust_tac: Proof.context -> string -> int -> tactic
   val rep_datatype_i: thm -> thm -> thm list -> thm list -> theory -> theory
-  val rep_datatype: Facts.ref * Attrib.src list -> Facts.ref * Attrib.src list ->
-    (Facts.ref * Attrib.src list) list -> (Facts.ref * Attrib.src list) list -> theory -> theory
+  val rep_datatype: Facts.ref * Token.src list -> Facts.ref * Token.src list ->
+    (Facts.ref * Token.src list) list -> (Facts.ref * Token.src list) list -> theory -> theory
   val setup: theory -> theory
 end;
 
@@ -190,10 +190,10 @@
 
 val _ =
   Outer_Syntax.command @{command_spec "rep_datatype"} "represent existing set inductively"
-    ((@{keyword "elimination"} |-- Parse.!!! Parse_Spec.xthm) --
-     (@{keyword "induction"} |-- Parse.!!! Parse_Spec.xthm) --
-     (@{keyword "case_eqns"} |-- Parse.!!! Parse_Spec.xthms1) --
-     Scan.optional (@{keyword "recursor_eqns"} |-- Parse.!!! Parse_Spec.xthms1) []
+    ((@{keyword "elimination"} |-- Parse.!!! Parse.xthm) --
+     (@{keyword "induction"} |-- Parse.!!! Parse.xthm) --
+     (@{keyword "case_eqns"} |-- Parse.!!! Parse.xthms1) --
+     Scan.optional (@{keyword "recursor_eqns"} |-- Parse.!!! Parse.xthms1) []
      >> (fn (((x, y), z), w) => Toplevel.theory (rep_datatype x y z w)));
 
 end;
--- a/src/ZF/Tools/inductive_package.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/Tools/inductive_package.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -29,9 +29,9 @@
     ((binding * term) * attribute list) list ->
     thm list * thm list * thm list * thm list -> theory -> theory * inductive_result
   val add_inductive: string list * string ->
-    ((binding * string) * Attrib.src list) list ->
-    (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list *
-    (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list ->
+    ((binding * string) * Token.src list) list ->
+    (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list *
+    (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list ->
     theory -> theory * inductive_result
 end;
 
@@ -586,10 +586,10 @@
       ((@{keyword "\<subseteq>"} || @{keyword "<="}) |-- Parse.term))) --
   (@{keyword "intros"} |--
     Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) --
-  Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse_Spec.xthms1) [] --
-  Scan.optional (@{keyword "con_defs"} |-- Parse.!!! Parse_Spec.xthms1) [] --
-  Scan.optional (@{keyword "type_intros"} |-- Parse.!!! Parse_Spec.xthms1) [] --
-  Scan.optional (@{keyword "type_elims"} |-- Parse.!!! Parse_Spec.xthms1) []
+  Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse.xthms1) [] --
+  Scan.optional (@{keyword "con_defs"} |-- Parse.!!! Parse.xthms1) [] --
+  Scan.optional (@{keyword "type_intros"} |-- Parse.!!! Parse.xthms1) [] --
+  Scan.optional (@{keyword "type_elims"} |-- Parse.!!! Parse.xthms1) []
   >> (Toplevel.theory o mk_ind);
 
 val _ =
--- a/src/ZF/Tools/primrec_package.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/Tools/primrec_package.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -8,7 +8,7 @@
 
 signature PRIMREC_PACKAGE =
 sig
-  val add_primrec: ((binding * string) * Attrib.src list) list -> theory -> theory * thm list
+  val add_primrec: ((binding * string) * Token.src list) list -> theory -> theory * thm list
   val add_primrec_i: ((binding * term) * attribute list) list -> theory -> theory * thm list
 end;
 
--- a/src/ZF/Tools/twos_compl.ML	Wed Aug 27 11:33:00 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,124 +0,0 @@
-(*  Title:      ZF/Tools/twos_compl.ML
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1993  University of Cambridge
-
-ML code for Arithmetic on binary integers; the model for theory Bin
-
-   The sign Pls stands for an infinite string of leading 0s.
-   The sign Min stands for an infinite string of leading 1s.
-
-See int_of_binary for the numerical interpretation.  A number can have
-multiple representations, namely leading 0s with sign Pls and leading 1s with
-sign Min.  A number is in NORMAL FORM if it has no such extra bits.
-
-The representation expects that (m mod 2) is 0 or 1, even if m is negative;
-For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1
-
-Still needs division!
-*)
-
-infix 5 $$ $$$
-
-(*Recursive datatype of binary integers*)
-datatype bin = Pls | Min | $$ of bin * int;
-
-(** Conversions between bin and int **)
-
-fun int_of_binary Pls = 0
-  | int_of_binary Min = ~1
-  | int_of_binary (w$$b) = 2 * int_of_binary w + b;
-
-fun binary_of_int 0 = Pls
-  | binary_of_int ~1 = Min
-  | binary_of_int n = binary_of_int (n div 2) $$ (n mod 2);
-
-(*** Addition ***)
-
-(*Attach a bit while preserving the normal form.  Cases left as default
-  are Pls$$$1 and Min$$$0. *)
-fun  Pls $$$ 0 = Pls
-  | Min $$$ 1 = Min
-  |     v $$$ x = v$$x;
-
-(*Successor of an integer, assumed to be in normal form.
-  If w$$1 is normal then w is not Min, so bin_succ(w) $$ 0 is normal.
-  But Min$$0 is normal while Min$$1 is not.*)
-fun bin_succ Pls = Pls$$1
-  | bin_succ Min = Pls
-  | bin_succ (w$$1) = bin_succ(w) $$ 0
-  | bin_succ (w$$0) = w $$$ 1;
-
-(*Predecessor of an integer, assumed to be in normal form.
-  If w$$0 is normal then w is not Pls, so bin_pred(w) $$ 1 is normal.
-  But Pls$$1 is normal while Pls$$0 is not.*)
-fun bin_pred Pls = Min
-  | bin_pred Min = Min$$0
-  | bin_pred (w$$1) = w $$$ 0
-  | bin_pred (w$$0) = bin_pred(w) $$ 1;
-
-(*Sum of two binary integers in normal form.  
-  Ensure last $$ preserves normal form! *)
-fun bin_add (Pls, w) = w
-  | bin_add (Min, w) = bin_pred w
-  | bin_add (v$$x, Pls) = v$$x
-  | bin_add (v$$x, Min) = bin_pred (v$$x)
-  | bin_add (v$$x, w$$y) = 
-      bin_add(v, if x+y=2 then bin_succ w else w) $$$ ((x+y) mod 2);
-
-(*** Subtraction ***)
-
-(*Unary minus*)
-fun bin_minus Pls = Pls
-  | bin_minus Min = Pls$$1
-  | bin_minus (w$$1) = bin_pred (bin_minus(w) $$$ 0)
-  | bin_minus (w$$0) = bin_minus(w) $$ 0;
-
-(*** Multiplication ***)
-
-(*product of two bins; a factor of 0 might cause leading 0s in result*)
-fun bin_mult (Pls, _) = Pls
-  | bin_mult (Min, v) = bin_minus v
-  | bin_mult (w$$1, v) = bin_add(bin_mult(w,v) $$$ 0,  v)
-  | bin_mult (w$$0, v) = bin_mult(w,v) $$$ 0;
-
-(*** Testing ***)
-
-(*tests addition*)
-fun checksum m n =
-    let val wm = binary_of_int m
-        and wn = binary_of_int n
-        val wsum = bin_add(wm,wn)
-    in  if m+n = int_of_binary wsum then (wm, wn, wsum, m+n)
-        else raise Match
-    end;
-
-fun bfact n = if n=0 then  Pls$$1  
-              else  bin_mult(binary_of_int n, bfact(n-1));
-
-(*Examples...
-bfact 5;
-int_of_binary it;
-bfact 69;
-int_of_binary it;
-
-(*For {HOL,ZF}/ex/BinEx.ML*)
-bin_add(binary_of_int 13, binary_of_int 19);
-bin_add(binary_of_int 1234, binary_of_int 5678);
-bin_add(binary_of_int 1359, binary_of_int ~2468);
-bin_add(binary_of_int 93746, binary_of_int ~46375);
-bin_minus(binary_of_int 65745);
-bin_minus(binary_of_int ~54321);
-bin_mult(binary_of_int 13, binary_of_int 19);
-bin_mult(binary_of_int ~84, binary_of_int 51);
-bin_mult(binary_of_int 255, binary_of_int 255);
-bin_mult(binary_of_int 1359, binary_of_int ~2468);
-
-
-(*leading zeros??*)
-bin_add(binary_of_int 1234, binary_of_int ~1234);
-bin_mult(binary_of_int 1234, Pls);
-
-(*leading ones??*)
-bin_add(binary_of_int 1, binary_of_int ~2);
-bin_add(binary_of_int 1234, binary_of_int ~1235);
-*)
--- a/src/ZF/Tools/typechk.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/Tools/typechk.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -115,8 +115,8 @@
 val typecheck_setup =
   Method.setup @{binding typecheck}
     (Method.sections
-      [Args.add -- Args.colon >> K (I, TC_add),
-       Args.del -- Args.colon >> K (I, TC_del)]
+      [Args.add -- Args.colon >> K (Method.modifier TC_add @{here}),
+       Args.del -- Args.colon >> K (Method.modifier TC_del @{here})]
       >> K (fn ctxt => SIMPLE_METHOD (CHANGED (typecheck_tac ctxt))))
     "ZF type-checking";
 
--- a/src/ZF/UNITY/Constrains.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/UNITY/Constrains.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -453,6 +453,9 @@
   used by Always_Int_I) *)
 lemmas Always_thin = thin_rl [of "F \<in> Always(A)"]
 
+(*To allow expansion of the program's definition when appropriate*)
+named_theorems program "program definitions"
+
 ML
 {*
 (*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
@@ -461,13 +464,6 @@
 (*Combines a list of invariance THEOREMS into one.*)
 val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS @{thm Always_Int_I});
 
-(*To allow expansion of the program's definition when appropriate*)
-structure Program_Defs = Named_Thms
-(
-  val name = @{binding program}
-  val description = "program definitions"
-);
-
 (*proves "co" properties when the program is specified*)
 
 fun constrains_tac ctxt =
@@ -481,7 +477,7 @@
               (* Three subgoals *)
               rewrite_goal_tac ctxt [@{thm st_set_def}] 3,
               REPEAT (force_tac ctxt 2),
-              full_simp_tac (ctxt addsimps (Program_Defs.get ctxt)) 1,
+              full_simp_tac (ctxt addsimps (Named_Theorems.get ctxt @{named_theorems program})) 1,
               ALLGOALS (clarify_tac ctxt),
               REPEAT (FIRSTGOAL (etac @{thm disjE})),
               ALLGOALS (clarify_tac ctxt),
@@ -495,8 +491,6 @@
     rtac @{thm AlwaysI} i THEN force_tac ctxt i THEN constrains_tac ctxt i;
 *}
 
-setup Program_Defs.setup
-
 method_setup safety = {*
   Scan.succeed (SIMPLE_METHOD' o constrains_tac) *}
   "for proving safety properties"
--- a/src/ZF/UNITY/SubstAx.thy	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/UNITY/SubstAx.thy	Wed Aug 27 15:52:58 2014 +0200
@@ -358,7 +358,7 @@
                 REPEAT (ares_tac [@{thm LeadsTo_Basis}, @{thm leadsTo_Basis},
                                   @{thm EnsuresI}, @{thm ensuresI}] 1),
             (*now there are two subgoals: co & transient*)
-            simp_tac (ctxt addsimps (Program_Defs.get ctxt)) 2,
+            simp_tac (ctxt addsimps (Named_Theorems.get ctxt @{named_theorems program})) 2,
             res_inst_tac ctxt [(("act", 0), sact)] @{thm transientI} 2,
                (*simplify the command's domain*)
             simp_tac (ctxt addsimps [@{thm domain_def}]) 3,
--- a/src/ZF/int_arith.ML	Wed Aug 27 11:33:00 2014 +0200
+++ b/src/ZF/int_arith.ML	Wed Aug 27 15:52:58 2014 +0200
@@ -151,7 +151,7 @@
 
 structure CancelNumeralsCommon =
   struct
-  val mk_sum            = (fn T:typ => mk_sum)
+  val mk_sum            = (fn _ : typ => mk_sum)
   val dest_sum          = dest_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff 1
@@ -236,7 +236,7 @@
   type coeff            = int
   val iszero            = (fn x => x = 0)
   val add               = op + 
-  val mk_sum            = (fn T:typ => long_mk_sum) (*to work for #2*x $+ #3*x *)
+  val mk_sum            = (fn _ : typ => long_mk_sum) (*to work for #2*x $+ #3*x *)
   val dest_sum          = dest_sum
   val mk_coeff          = mk_coeff
   val dest_coeff        = dest_coeff 1
@@ -284,7 +284,7 @@
   type coeff            = int
   val iszero            = (fn x => x = 0)
   val add               = op *
-  val mk_sum            = (fn T:typ => mk_prod)
+  val mk_sum            = (fn _ : typ => mk_prod)
   val dest_sum          = dest_prod
   fun mk_coeff(k,t) = if t=one then mk_numeral k
                       else raise TERM("mk_coeff", [])