merged
authorhuffman
Tue, 23 Feb 2010 14:44:43 -0800
changeset 35349 f9801fdeb789
parent 35348 c6331256b087 (current diff)
parent 35329 cac5a37fb638 (diff)
child 35350 0df9c8a37f64
merged
src/HOL/Algebras.thy
src/HOL/Hoare/HoareAbort.thy
src/HOL/W0/README.html
src/HOL/W0/ROOT.ML
src/HOL/W0/W0.thy
src/HOL/W0/document/root.tex
--- a/NEWS	Tue Feb 23 14:44:24 2010 -0800
+++ b/NEWS	Tue Feb 23 14:44:43 2010 -0800
@@ -122,9 +122,6 @@
 
 INCOMPATIBILITY.
 
-* New theory Algebras contains generic algebraic structures and
-generic algebraic operations.
-
 * HOLogic.strip_psplit: types are returned in syntactic order, similar
 to other strip and tuple operations.  INCOMPATIBILITY.
 
--- a/doc-src/Nitpick/nitpick.tex	Tue Feb 23 14:44:24 2010 -0800
+++ b/doc-src/Nitpick/nitpick.tex	Tue Feb 23 14:44:43 2010 -0800
@@ -150,11 +150,11 @@
 \postw
 
 The results presented here were obtained using the JNI version of MiniSat and
-with multithreading disabled to reduce nondeterminism. This was done by adding
-the line
+with multithreading disabled to reduce nondeterminism and a time limit of
+15~seconds (instead of 30~seconds). This was done by adding the line
 
 \prew
-\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSat\_JNI}, \,\textit{max\_threads}~= 1]
+\textbf{nitpick\_params} [\textit{sat\_solver}~= \textit{MiniSat\_JNI}, \,\textit{max\_threads}~= 1, \,\textit{timeout} = 15$\,s$]
 \postw
 
 after the \textbf{begin} keyword. The JNI version of MiniSat is bundled with
@@ -499,7 +499,7 @@
 
 \prew
 \textbf{lemma} ``$P~\textit{Suc}$'' \\
-\textbf{nitpick} [\textit{card} = 1--6] \\[2\smallskipamount]
+\textbf{nitpick} \\[2\smallskipamount]
 \slshape
 Nitpick found no counterexample.
 \postw
@@ -1617,7 +1617,7 @@
 ``$w \in A \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = b] + 1$'' \\
 ``$w \in B \longleftrightarrow \textit{length}~[x \mathbin{\leftarrow} w.\; x = b] = \textit{length}~[x \mathbin{\leftarrow} w.\; x = a] + 1$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
-\slshape Nitpick found no counterexample.
+\slshape Nitpick ran out of time after checking 7 of 8 scopes.
 \postw
 
 \subsection{AA Trees}
@@ -1726,7 +1726,7 @@
 ``$\textit{dataset}~(\textit{skew}~t) = \textit{dataset}~t$'' \\
 ``$\textit{dataset}~(\textit{split}~t) = \textit{dataset}~t$'' \\
 \textbf{nitpick} \\[2\smallskipamount]
-{\slshape Nitpick found no counterexample.}
+{\slshape Nitpick ran out of time after checking 7 of 8 scopes.}
 \postw
 
 Furthermore, applying \textit{skew} or \textit{split} to a well-formed tree
--- a/src/HOL/Algebra/abstract/Ideal2.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Algebra/abstract/Ideal2.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,6 +1,5 @@
 (*
     Ideals for commutative rings
-    $Id$
     Author: Clemens Ballarin, started 24 September 1999
 *)
 
@@ -23,9 +22,8 @@
 
 text {* Principle ideal domains *}
 
-axclass pid < "domain"
-  pid_ax: "is_ideal I ==> is_pideal I"
-
+class pid =
+  assumes pid_ax: "is_ideal (I :: 'a::domain \<Rightarrow> _) \<Longrightarrow> is_pideal I"
 
 (* is_ideal *)
 
--- a/src/HOL/Algebras.thy	Tue Feb 23 14:44:24 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,55 +0,0 @@
-(*  Title:      HOL/Algebras.thy
-    Author:     Florian Haftmann, TU Muenchen
-*)
-
-header {* Generic algebraic structures and operations *}
-
-theory Algebras
-imports HOL
-begin
-
-text {*
-  These locales provide basic structures for interpretation into
-  bigger structures;  extensions require careful thinking, otherwise
-  undesired effects may occur due to interpretation.
-*}
-
-ML {*
-structure Ac_Simps = Named_Thms(
-  val name = "ac_simps"
-  val description = "associativity and commutativity simplification rules"
-)
-*}
-
-setup Ac_Simps.setup
-
-locale semigroup =
-  fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)
-  assumes assoc [ac_simps]: "a * b * c = a * (b * c)"
-
-locale abel_semigroup = semigroup +
-  assumes commute [ac_simps]: "a * b = b * a"
-begin
-
-lemma left_commute [ac_simps]:
-  "b * (a * c) = a * (b * c)"
-proof -
-  have "(b * a) * c = (a * b) * c"
-    by (simp only: commute)
-  then show ?thesis
-    by (simp only: assoc)
-qed
-
-end
-
-locale semilattice = abel_semigroup +
-  assumes idem [simp]: "a * a = a"
-begin
-
-lemma left_idem [simp]:
-  "a * (a * b) = a * b"
-  by (simp add: assoc [symmetric])
-
-end
-
-end
\ No newline at end of file
--- a/src/HOL/Bali/Decl.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Bali/Decl.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -230,18 +230,22 @@
 
 datatype memberid = fid vname | mid sig
 
-axclass has_memberid < "type"
-consts
- memberid :: "'a::has_memberid \<Rightarrow> memberid"
+class has_memberid =
+  fixes memberid :: "'a \<Rightarrow> memberid"
 
-instance memberdecl::has_memberid ..
+instantiation memberdecl :: has_memberid
+begin
 
-defs (overloaded)
+definition
 memberdecl_memberid_def:
   "memberid m \<equiv> (case m of
                     fdecl (vn,f)  \<Rightarrow> fid vn
                   | mdecl (sig,m) \<Rightarrow> mid sig)"
 
+instance ..
+
+end
+
 lemma memberid_fdecl_simp[simp]: "memberid (fdecl (vn,f)) = fid vn"
 by (simp add: memberdecl_memberid_def)
 
@@ -254,12 +258,17 @@
 lemma memberid_mdecl_simp1: "memberid (mdecl m) = mid (fst m)"
 by (cases m) (simp add: memberdecl_memberid_def)
 
-instance * :: (type, has_memberid) has_memberid ..
+instantiation * :: (type, has_memberid) has_memberid
+begin
 
-defs (overloaded)
+definition
 pair_memberid_def:
   "memberid p \<equiv> memberid (snd p)"
 
+instance ..
+
+end
+
 lemma memberid_pair_simp[simp]: "memberid (c,m) = memberid m"
 by (simp add: pair_memberid_def)
 
--- a/src/HOL/Bali/DeclConcepts.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Bali/DeclConcepts.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -79,41 +79,60 @@
 text {* overloaded selector @{text accmodi} to select the access modifier 
 out of various HOL types *}
 
-axclass has_accmodi < "type"
-consts accmodi:: "'a::has_accmodi \<Rightarrow> acc_modi"
+class has_accmodi =
+  fixes accmodi:: "'a \<Rightarrow> acc_modi"
+
+instantiation acc_modi :: has_accmodi
+begin
 
-instance acc_modi::has_accmodi ..
+definition
+  acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
 
-defs (overloaded)
-acc_modi_accmodi_def: "accmodi (a::acc_modi) \<equiv> a"
+instance ..
+
+end
 
 lemma acc_modi_accmodi_simp[simp]: "accmodi (a::acc_modi) = a"
 by (simp add: acc_modi_accmodi_def)
 
-instance decl_ext_type:: ("type") has_accmodi ..
+instantiation decl_ext_type:: (type) has_accmodi
+begin
 
-defs (overloaded)
-decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
+definition
+  decl_acc_modi_def: "accmodi (d::('a:: type) decl_scheme) \<equiv> access d"
 
+instance ..
+
+end
 
 lemma decl_acc_modi_simp[simp]: "accmodi (d::('a::type) decl_scheme) = access d"
 by (simp add: decl_acc_modi_def)
 
-instance * :: ("type",has_accmodi) has_accmodi ..
+instantiation * :: (type, has_accmodi) has_accmodi
+begin
 
-defs (overloaded)
-pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
+definition
+  pair_acc_modi_def: "accmodi p \<equiv> (accmodi (snd p))"
+
+instance ..
+
+end
 
 lemma pair_acc_modi_simp[simp]: "accmodi (x,a) = (accmodi a)"
 by (simp add: pair_acc_modi_def)
 
-instance memberdecl :: has_accmodi ..
+instantiation memberdecl :: has_accmodi
+begin
 
-defs (overloaded)
-memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
+definition
+  memberdecl_acc_modi_def: "accmodi m \<equiv> (case m of
                                           fdecl f \<Rightarrow> accmodi f
                                         | mdecl m \<Rightarrow> accmodi m)"
 
+instance ..
+
+end
+
 lemma memberdecl_fdecl_acc_modi_simp[simp]:
  "accmodi (fdecl m) = accmodi m"
 by (simp add: memberdecl_acc_modi_def)
@@ -125,21 +144,35 @@
 text {* overloaded selector @{text declclass} to select the declaring class 
 out of various HOL types *}
 
-axclass has_declclass < "type"
-consts declclass:: "'a::has_declclass \<Rightarrow> qtname"
+class has_declclass =
+  fixes declclass:: "'a \<Rightarrow> qtname"
+
+instantiation qtname_ext_type :: (type) has_declclass
+begin
 
-instance qtname_ext_type::("type") has_declclass ..
+definition
+  "declclass q \<equiv> \<lparr> pid = pid q, tid = tid q \<rparr>"
+
+instance ..
 
-defs (overloaded)
-qtname_declclass_def: "declclass (q::qtname) \<equiv> q"
+end
+
+lemma qtname_declclass_def:
+  "declclass q \<equiv> (q::qtname)"
+  by (induct q) (simp add: declclass_qtname_ext_type_def)
 
 lemma qtname_declclass_simp[simp]: "declclass (q::qtname) = q"
 by (simp add: qtname_declclass_def)
 
-instance * :: ("has_declclass","type") has_declclass ..
+instantiation * :: (has_declclass, type) has_declclass
+begin
 
-defs (overloaded)
-pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
+definition
+  pair_declclass_def: "declclass p \<equiv> declclass (fst p)"
+
+instance ..
+
+end
 
 lemma pair_declclass_simp[simp]: "declclass (c,x) = declclass c" 
 by (simp add: pair_declclass_def)
@@ -147,25 +180,38 @@
 text {* overloaded selector @{text is_static} to select the static modifier 
 out of various HOL types *}
 
+class has_static =
+  fixes is_static :: "'a \<Rightarrow> bool"
 
-axclass has_static < "type"
-consts is_static :: "'a::has_static \<Rightarrow> bool"
+instantiation decl_ext_type :: (has_static) has_static
+begin
 
-instance decl_ext_type :: ("has_static") has_static ..
+instance ..
+
+end
 
-instance member_ext_type :: ("type") has_static ..
+instantiation member_ext_type :: (type) has_static
+begin
+
+instance ..
 
-defs (overloaded)
-static_field_type_is_static_def: 
- "is_static (m::('b member_scheme)) \<equiv> static m"
+end
+
+axiomatization where
+  static_field_type_is_static_def: "is_static (m::('a member_scheme)) \<equiv> static m"
 
 lemma member_is_static_simp: "is_static (m::'a member_scheme) = static m"
 by (simp add: static_field_type_is_static_def)
 
-instance * :: ("type","has_static") has_static ..
+instantiation * :: (type, has_static) has_static
+begin
 
-defs (overloaded)
-pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
+definition
+  pair_is_static_def: "is_static p \<equiv> is_static (snd p)"
+
+instance ..
+
+end
 
 lemma pair_is_static_simp [simp]: "is_static (x,s) = is_static s"
 by (simp add: pair_is_static_def)
@@ -173,14 +219,19 @@
 lemma pair_is_static_simp1: "is_static p = is_static (snd p)"
 by (simp add: pair_is_static_def)
 
-instance memberdecl:: has_static ..
+instantiation memberdecl :: has_static
+begin
 
-defs (overloaded)
+definition
 memberdecl_is_static_def: 
  "is_static m \<equiv> (case m of
                     fdecl f \<Rightarrow> is_static f
                   | mdecl m \<Rightarrow> is_static m)"
 
+instance ..
+
+end
+
 lemma memberdecl_is_static_fdecl_simp[simp]:
  "is_static (fdecl f) = is_static f"
 by (simp add: memberdecl_is_static_def)
@@ -389,18 +440,32 @@
 text {* overloaded selector @{text resTy} to select the result type 
 out of various HOL types *}
 
-axclass has_resTy < "type"
-consts resTy:: "'a::has_resTy \<Rightarrow> ty"
+class has_resTy =
+  fixes resTy:: "'a \<Rightarrow> ty"
+
+instantiation decl_ext_type :: (has_resTy) has_resTy
+begin
 
-instance decl_ext_type :: ("has_resTy") has_resTy ..
+instance ..
+
+end
+
+instantiation member_ext_type :: (has_resTy) has_resTy
+begin
 
-instance member_ext_type :: ("has_resTy") has_resTy ..
+instance ..
 
-instance mhead_ext_type :: ("type") has_resTy ..
+end
+
+instantiation mhead_ext_type :: (type) has_resTy
+begin
 
-defs (overloaded)
-mhead_ext_type_resTy_def: 
- "resTy (m::('b mhead_scheme)) \<equiv> resT m"
+instance ..
+
+end
+
+axiomatization where
+  mhead_ext_type_resTy_def: "resTy (m::('b mhead_scheme)) \<equiv> resT m"
 
 lemma mhead_resTy_simp: "resTy (m::'a mhead_scheme) = resT m"
 by (simp add: mhead_ext_type_resTy_def)
@@ -408,10 +473,15 @@
 lemma resTy_mhead [simp]:"resTy (mhead m) = resTy m"
 by (simp add: mhead_def mhead_resTy_simp)
 
-instance * :: ("type","has_resTy") has_resTy ..
+instantiation * :: ("type","has_resTy") has_resTy
+begin
 
-defs (overloaded)
-pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
+definition
+  pair_resTy_def: "resTy p \<equiv> resTy (snd p)"
+
+instance ..
+
+end
 
 lemma pair_resTy_simp[simp]: "resTy (x,m) = resTy m"
 by (simp add: pair_resTy_def)
--- a/src/HOL/Bali/Name.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Bali/Name.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -48,29 +48,34 @@
           pid :: pname  
           tid :: tname
 
-axclass has_pname < "type"
-consts pname::"'a::has_pname \<Rightarrow> pname"
+class has_pname =
+  fixes pname :: "'a \<Rightarrow> pname"
 
-instance pname::has_pname ..
+instantiation pname :: has_pname
+begin
 
-defs (overloaded)
-pname_pname_def: "pname (p::pname) \<equiv> p"
+definition
+  pname_pname_def: "pname (p::pname) \<equiv> p"
 
-axclass has_tname < "type"
-consts tname::"'a::has_tname \<Rightarrow> tname"
+instance ..
+
+end
 
-instance tname::has_tname ..
+class has_tname =
+  fixes tname :: "'a \<Rightarrow> tname"
 
-defs (overloaded)
-tname_tname_def: "tname (t::tname) \<equiv> t"
+instantiation tname :: has_tname
+begin
 
-axclass has_qtname < type
-consts qtname:: "'a::has_qtname \<Rightarrow> qtname"
+definition
+  tname_tname_def: "tname (t::tname) \<equiv> t"
+
+instance ..
 
-instance qtname_ext_type :: (type) has_qtname ..
+end
 
-defs (overloaded)
-qtname_qtname_def: "qtname (q::qtname) \<equiv> q"
+definition
+  qtname_qtname_def: "qtname (q::'a qtname_ext_type) \<equiv> q"
 
 translations
   "mname"  <= "Name.mname"
--- a/src/HOL/Bali/Term.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Bali/Term.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -295,8 +295,8 @@
 following.
 *}
 
-axclass inj_term < "type"
-consts inj_term:: "'a::inj_term \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
+class inj_term =
+  fixes inj_term:: "'a \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
 
 text {* How this overloaded injections work can be seen in the theory 
 @{text DefiniteAssignment}. Other big inductive relations on
@@ -308,10 +308,15 @@
 as bridge between the different conventions.  
 *}
 
-instance stmt::inj_term ..
+instantiation stmt :: inj_term
+begin
 
-defs (overloaded)
-stmt_inj_term_def: "\<langle>c::stmt\<rangle> \<equiv> In1r c"
+definition
+  stmt_inj_term_def: "\<langle>c::stmt\<rangle> \<equiv> In1r c"
+
+instance ..
+
+end
 
 lemma stmt_inj_term_simp: "\<langle>c::stmt\<rangle> = In1r c"
 by (simp add: stmt_inj_term_def)
@@ -319,10 +324,15 @@
 lemma  stmt_inj_term [iff]: "\<langle>x::stmt\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
   by (simp add: stmt_inj_term_simp)
 
-instance expr::inj_term ..
+instantiation expr :: inj_term
+begin
 
-defs (overloaded)
-expr_inj_term_def: "\<langle>e::expr\<rangle> \<equiv> In1l e"
+definition
+  expr_inj_term_def: "\<langle>e::expr\<rangle> \<equiv> In1l e"
+
+instance ..
+
+end
 
 lemma expr_inj_term_simp: "\<langle>e::expr\<rangle> = In1l e"
 by (simp add: expr_inj_term_def)
@@ -330,10 +340,15 @@
 lemma expr_inj_term [iff]: "\<langle>x::expr\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
   by (simp add: expr_inj_term_simp)
 
-instance var::inj_term ..
+instantiation var :: inj_term
+begin
 
-defs (overloaded)
-var_inj_term_def: "\<langle>v::var\<rangle> \<equiv> In2 v"
+definition
+  var_inj_term_def: "\<langle>v::var\<rangle> \<equiv> In2 v"
+
+instance ..
+
+end
 
 lemma var_inj_term_simp: "\<langle>v::var\<rangle> = In2 v"
 by (simp add: var_inj_term_def)
@@ -341,10 +356,32 @@
 lemma var_inj_term [iff]: "\<langle>x::var\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
   by (simp add: var_inj_term_simp)
 
-instance "list":: (type) inj_term ..
+class expr_of =
+  fixes expr_of :: "'a \<Rightarrow> expr"
+
+instantiation expr :: expr_of
+begin
+
+definition
+  "expr_of = (\<lambda>(e::expr). e)"
+
+instance ..
+
+end 
 
-defs (overloaded)
-expr_list_inj_term_def: "\<langle>es::expr list\<rangle> \<equiv> In3 es"
+instantiation list :: (expr_of) inj_term
+begin
+
+definition
+  "\<langle>es::'a list\<rangle> \<equiv> In3 (map expr_of es)"
+
+instance ..
+
+end
+
+lemma expr_list_inj_term_def:
+  "\<langle>es::expr list\<rangle> \<equiv> In3 es"
+  by (simp add: inj_term_list_def expr_of_expr_def)
 
 lemma expr_list_inj_term_simp: "\<langle>es::expr list\<rangle> = In3 es"
 by (simp add: expr_list_inj_term_def)
--- a/src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -56,10 +56,12 @@
 
 text {* Let's find out which assertions of @{text max} are hard to prove: *}
 
-boogie_status (narrow timeout: 4) max
+boogie_status (narrow step_timeout: 4 final_timeout: 15) max
   (simp add: labels, (fast | metis)?)
-  -- {* The argument @{text timeout} is optional, restricting the runtime of
-        each narrowing step by the given number of seconds. *}
+  -- {* The argument @{text step_timeout} is optional, restricting the runtime
+        of each intermediate narrowing step by the given number of seconds. *}
+  -- {* The argument @{text final_timeout} is optional, restricting the
+        runtime of the final narrowing steps by the given number of seconds. *}
   -- {* The last argument should be a method to be applied at each narrowing
         step. Note that different methods may be composed to one method by
         a language similar to regular expressions. *}
--- a/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -156,26 +156,27 @@
     |> Seq.hd
     |> Proof.global_done_proof
 in
-fun boogie_narrow_vc timeout vc_name meth thy =
+fun boogie_narrow_vc (quick, timeout) vc_name meth thy =
   let
-    val tp = TimeLimit.timeLimit (Time.fromSeconds timeout) (prove thy meth)
+    fun tp t = TimeLimit.timeLimit (Time.fromSeconds t) (prove thy meth)
 
-    fun try_vc (tag, split_tag) split vc = (trying tag;
-      (case try tp vc of
+    fun try_vc t (tag, split_tag) split vc = (trying tag;
+      (case try (tp t) vc of
         SOME _ => (success_on tag; [])
       | NONE => (failure_on tag split_tag; split vc)))
 
     fun some_asserts vc =
-      try_vc (string_of_asserts vc,
-        if Boogie_VCs.size_of vc = 1 then "." else ", further splitting ...")
-        (divide some_asserts) vc
+      let
+        val (t, sep) = if Boogie_VCs.size_of vc = 1 then (timeout, ".")
+          else (quick, ", further splitting ...")
+      in try_vc t (string_of_asserts vc, sep) (divide some_asserts) vc end
 
     fun single_path p =
-      try_vc (string_of_path p, ", splitting into assertions ...")
+      try_vc quick (string_of_path p, ", splitting into assertions ...")
         (divide some_asserts)
 
     val complete_vc =
-      try_vc ("full goal", ", splitting into paths ...")
+      try_vc quick ("full goal", ", splitting into paths ...")
         (par_map (uncurry single_path) o itemize_paths o Boogie_VCs.paths_of)
 
     val unsolved = complete_vc (get_vc thy vc_name)
@@ -262,15 +263,18 @@
       (Toplevel.print o Toplevel.local_theory_to_proof NONE (boogie_vc args))))
 
 
-val default_timeout = 5
+val quick_timeout = 5
+val default_timeout = 20
+
+fun timeout name = Scan.optional (scan_val name OuterParse.nat)
 
 val status_test =
   scan_arg (
-    (Args.$$$ "scan" >> K boogie_scan_vc ||
-     Args.$$$ "narrow" >> K boogie_narrow_vc) --
-    Scan.optional (scan_val "timeout" OuterParse.nat) default_timeout) --
+    Args.$$$ "scan" |-- timeout "timeout" quick_timeout >> boogie_scan_vc ||
+    Args.$$$ "narrow" |-- timeout "step_timeout" quick_timeout --
+      timeout "final_timeout" default_timeout >> boogie_narrow_vc) --
   vc_name -- Method.parse >>
-  (fn (((f, timeout), vc_name), meth) => f timeout vc_name meth)
+  (fn ((f, vc_name), meth) => f vc_name meth)
 
 val status_vc =
   (scan_arg
--- a/src/HOL/Code_Evaluation.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Code_Evaluation.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -76,7 +76,8 @@
         andalso can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort typerep};
     in if need_inst then add_term_of tyco raw_vs thy else thy end;
 in
-  Code.type_interpretation ensure_term_of
+  Code.datatype_interpretation ensure_term_of
+  #> Code.abstype_interpretation ensure_term_of
 end
 *}
 
@@ -114,7 +115,7 @@
       val has_inst = can (Sorts.mg_domain (Sign.classes_of thy) tyco) @{sort term_of};
     in if has_inst then add_term_of_code tyco raw_vs cs thy else thy end;
 in
-  Code.type_interpretation ensure_term_of_code
+  Code.datatype_interpretation ensure_term_of_code
 end
 *}
 
--- a/src/HOL/Groups.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Groups.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -9,6 +9,65 @@
 uses ("~~/src/Provers/Arith/abel_cancel.ML")
 begin
 
+subsection {* Fact collections *}
+
+ML {*
+structure Algebra_Simps = Named_Thms(
+  val name = "algebra_simps"
+  val description = "algebra simplification rules"
+)
+*}
+
+setup Algebra_Simps.setup
+
+text{* The rewrites accumulated in @{text algebra_simps} deal with the
+classical algebraic structures of groups, rings and family. They simplify
+terms by multiplying everything out (in case of a ring) and bringing sums and
+products into a canonical form (by ordered rewriting). As a result it decides
+group and ring equalities but also helps with inequalities.
+
+Of course it also works for fields, but it knows nothing about multiplicative
+inverses or division. This is catered for by @{text field_simps}. *}
+
+
+ML {*
+structure Ac_Simps = Named_Thms(
+  val name = "ac_simps"
+  val description = "associativity and commutativity simplification rules"
+)
+*}
+
+setup Ac_Simps.setup
+
+
+subsection {* Abstract structures *}
+
+text {*
+  These locales provide basic structures for interpretation into
+  bigger structures;  extensions require careful thinking, otherwise
+  undesired effects may occur due to interpretation.
+*}
+
+locale semigroup =
+  fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "*" 70)
+  assumes assoc [ac_simps]: "a * b * c = a * (b * c)"
+
+locale abel_semigroup = semigroup +
+  assumes commute [ac_simps]: "a * b = b * a"
+begin
+
+lemma left_commute [ac_simps]:
+  "b * (a * c) = a * (b * c)"
+proof -
+  have "(b * a) * c = (a * b) * c"
+    by (simp only: commute)
+  then show ?thesis
+    by (simp only: assoc)
+qed
+
+end
+
+
 subsection {* Generic operations *}
 
 class zero = 
@@ -64,37 +123,6 @@
 
 use "~~/src/Provers/Arith/abel_cancel.ML"
 
-text {*
-  The theory of partially ordered groups is taken from the books:
-  \begin{itemize}
-  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
-  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
-  \end{itemize}
-  Most of the used notions can also be looked up in 
-  \begin{itemize}
-  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
-  \item \emph{Algebra I} by van der Waerden, Springer.
-  \end{itemize}
-*}
-
-ML {*
-structure Algebra_Simps = Named_Thms(
-  val name = "algebra_simps"
-  val description = "algebra simplification rules"
-)
-*}
-
-setup Algebra_Simps.setup
-
-text{* The rewrites accumulated in @{text algebra_simps} deal with the
-classical algebraic structures of groups, rings and family. They simplify
-terms by multiplying everything out (in case of a ring) and bringing sums and
-products into a canonical form (by ordered rewriting). As a result it decides
-group and ring equalities but also helps with inequalities.
-
-Of course it also works for fields, but it knows nothing about multiplicative
-inverses or division. This is catered for by @{text field_simps}. *}
-
 
 subsection {* Semigroups and Monoids *}
 
@@ -144,19 +172,6 @@
 
 theorems mult_ac = mult_assoc mult_commute mult_left_commute
 
-class ab_semigroup_idem_mult = ab_semigroup_mult +
-  assumes mult_idem: "x * x = x"
-
-sublocale ab_semigroup_idem_mult < times!: semilattice times proof
-qed (fact mult_idem)
-
-context ab_semigroup_idem_mult
-begin
-
-lemmas mult_left_idem = times.left_idem
-
-end
-
 class monoid_add = zero + semigroup_add +
   assumes add_0_left [simp]: "0 + a = a"
     and add_0_right [simp]: "a + 0 = a"
@@ -411,6 +426,19 @@
 
 subsection {* (Partially) Ordered Groups *} 
 
+text {*
+  The theory of partially ordered groups is taken from the books:
+  \begin{itemize}
+  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
+  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
+  \end{itemize}
+  Most of the used notions can also be looked up in 
+  \begin{itemize}
+  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
+  \item \emph{Algebra I} by van der Waerden, Springer.
+  \end{itemize}
+*}
+
 class ordered_ab_semigroup_add = order + ab_semigroup_add +
   assumes add_left_mono: "a \<le> b \<Longrightarrow> c + a \<le> c + b"
 begin
--- a/src/HOL/Hoare/Examples.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/Examples.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Hoare/Examples.thy
-    ID:         $Id$
     Author:     Norbert Galm
     Copyright   1998 TUM
 
 Various examples.
 *)
 
-theory Examples imports Hoare Arith2 begin
+theory Examples imports Hoare_Logic Arith2 begin
 
 (*** ARITHMETIC ***)
 
--- a/src/HOL/Hoare/ExamplesAbort.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/ExamplesAbort.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,12 +1,11 @@
 (*  Title:      HOL/Hoare/ExamplesAbort.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   1998 TUM
 
 Some small examples for programs that may abort.
 *)
 
-theory ExamplesAbort imports HoareAbort begin
+theory ExamplesAbort imports Hoare_Logic_Abort begin
 
 lemma "VARS x y z::nat
  {y = z & z \<noteq> 0} z \<noteq> 0 \<rightarrow> x := y div z {x = 1}"
--- a/src/HOL/Hoare/HeapSyntax.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/HeapSyntax.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -3,7 +3,7 @@
     Copyright   2002 TUM
 *)
 
-theory HeapSyntax imports Hoare Heap begin
+theory HeapSyntax imports Hoare_Logic Heap begin
 
 subsection "Field access and update"
 
--- a/src/HOL/Hoare/HeapSyntaxAbort.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/HeapSyntaxAbort.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -3,7 +3,7 @@
     Copyright   2002 TUM
 *)
 
-theory HeapSyntaxAbort imports HoareAbort Heap begin
+theory HeapSyntaxAbort imports Hoare_Logic_Abort Heap begin
 
 subsection "Field access and update"
 
--- a/src/HOL/Hoare/Hoare.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/Hoare.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,245 +1,9 @@
-(*  Title:      HOL/Hoare/Hoare.thy
-    Author:     Leonor Prensa Nieto & Tobias Nipkow
-    Copyright   1998 TUM
-
-Sugared semantic embedding of Hoare logic.
-Strictly speaking a shallow embedding (as implemented by Norbert Galm
-following Mike Gordon) would suffice. Maybe the datatype com comes in useful
-later.
+(*  Author:     Tobias Nipkow
+    Copyright   1998-2003 TUM
 *)
 
 theory Hoare
-imports Main
-uses ("hoare_tac.ML")
+imports Examples ExamplesAbort Pointers0 Pointer_Examples Pointer_ExamplesAbort SchorrWaite Separation
 begin
 
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
-
-datatype
- 'a com = Basic "'a \<Rightarrow> 'a"
-   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
-   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
-   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
-
-abbreviation annskip ("SKIP") where "SKIP == Basic id"
-
-types 'a sem = "'a => 'a => bool"
-
-consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
-primrec
-"iter 0 b S = (%s s'. s ~: b & (s=s'))"
-"iter (Suc n) b S = (%s s'. s : b & (? s''. S s s'' & iter n b S s'' s'))"
-
-consts Sem :: "'a com => 'a sem"
-primrec
-"Sem(Basic f) s s' = (s' = f s)"
-"Sem(c1;c2) s s' = (? s''. Sem c1 s s'' & Sem c2 s'' s')"
-"Sem(IF b THEN c1 ELSE c2 FI) s s' = ((s  : b --> Sem c1 s s') &
-                                      (s ~: b --> Sem c2 s s'))"
-"Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')"
-
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
-  "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
-
-
-
-(** parse translations **)
-
-syntax
-  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
-
-syntax
- "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
-                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
-syntax ("" output)
- "_hoare"      :: "['a assn,'a com,'a assn] => bool"
-                 ("{_} // _ // {_}" [0,55,0] 50)
-ML {*
-
-local
-
-fun abs((a,T),body) =
-  let val a = absfree(a, dummyT, body)
-  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
-in
-
-fun mk_abstuple [x] body = abs (x, body)
-  | mk_abstuple (x::xs) body =
-      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
-
-fun mk_fbody a e [x as (b,_)] = if a=b then e else Syntax.free b
-  | mk_fbody a e ((b,_)::xs) =
-      Syntax.const @{const_syntax Pair} $ (if a=b then e else Syntax.free b) $ mk_fbody a e xs;
-
-fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
 end
-*}
-
-(* bexp_tr & assn_tr *)
-(*all meta-variables for bexp except for TRUE are translated as if they
-  were boolean expressions*)
-ML{*
-fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
-  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
-
-fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
-*}
-(* com_tr *)
-ML{*
-fun com_tr (Const(@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
-      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
-  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
-  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
-      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
-  | com_tr t _ = t (* if t is just a Free/Var *)
-*}
-
-(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
-ML{*
-local
-
-fun var_tr(Free(a,_)) = (a,Bound 0) (* Bound 0 = dummy term *)
-  | var_tr(Const (@{syntax_const "_constrain"}, _) $ (Free (a,_)) $ T) = (a,T);
-
-fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
-  | vars_tr t = [var_tr t]
-
-in
-fun hoare_vars_tr [vars, pre, prg, post] =
-      let val xs = vars_tr vars
-      in Syntax.const @{const_syntax Valid} $
-         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
-      end
-  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
-end
-*}
-
-parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
-
-
-(*****************************************************************************)
-
-(*** print translations ***)
-ML{*
-fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
-                            subst_bound (Syntax.free v, dest_abstuple body)
-  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
-  | dest_abstuple trm = trm;
-
-fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
-  | abs2list (Abs(x,T,t)) = [Free (x, T)]
-  | abs2list _ = [];
-
-fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
-  | mk_ts (Abs(x,_,t)) = mk_ts t
-  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
-  | mk_ts t = [t];
-
-fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
-           ((Syntax.free x)::(abs2list t), mk_ts t)
-  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
-  | mk_vts t = raise Match;
-
-fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
-  | find_ch ((v,t)::vts) i xs =
-      if t = Bound i then find_ch vts (i-1) xs
-      else (true, (v, subst_bounds (xs, t)));
-
-fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
-  | is_f (Abs(x,_,t)) = true
-  | is_f t = false;
-*}
-
-(* assn_tr' & bexp_tr'*)
-ML{*
-fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | assn_tr' (Const (@{const_syntax inter}, _) $
-        (Const (@{const_syntax Collect},_) $ T1) $ (Const (@{const_syntax Collect},_) $ T2)) =
-      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
-  | assn_tr' t = t;
-
-fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | bexp_tr' t = t;
-*}
-
-(*com_tr' *)
-ML{*
-fun mk_assign f =
-  let val (vs, ts) = mk_vts f;
-      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
-  in
-    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
-    else Syntax.const @{const_syntax annskip}
-  end;
-
-fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
-      if is_f f then mk_assign f
-      else Syntax.const @{const_syntax Basic} $ f
-  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
-      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
-      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
-      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
-  | com_tr' t = t;
-
-fun spec_tr' [p, c, q] =
-  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
-*}
-
-print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
-
-lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
-by (auto simp:Valid_def)
-
-lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
-by (auto simp:Valid_def)
-
-lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
-by (auto simp:Valid_def)
-
-lemma CondRule:
- "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
-  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
-by (auto simp:Valid_def)
-
-lemma iter_aux: "! s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
-       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)";
-apply(induct n)
- apply clarsimp
-apply(simp (no_asm_use))
-apply blast
-done
-
-lemma WhileRule:
- "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
-apply (clarsimp simp:Valid_def)
-apply(drule iter_aux)
-  prefer 2 apply assumption
- apply blast
-apply blast
-done
-
-
-lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
-  by blast
-
-lemmas AbortRule = SkipRule  -- "dummy version"
-use "hoare_tac.ML"
-
-method_setup vcg = {*
-  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
-  "verification condition generator"
-
-method_setup vcg_simp = {*
-  Scan.succeed (fn ctxt =>
-    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
-  "verification condition generator plus simplification"
-
-end
--- a/src/HOL/Hoare/HoareAbort.thy	Tue Feb 23 14:44:24 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,269 +0,0 @@
-(*  Title:      HOL/Hoare/HoareAbort.thy
-    Author:     Leonor Prensa Nieto & Tobias Nipkow
-    Copyright   2003 TUM
-
-Like Hoare.thy, but with an Abort statement for modelling run time errors.
-*)
-
-theory HoareAbort
-imports Main
-uses ("hoare_tac.ML")
-begin
-
-types
-    'a bexp = "'a set"
-    'a assn = "'a set"
-
-datatype
- 'a com = Basic "'a \<Rightarrow> 'a"
-   | Abort
-   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
-   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
-   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
-
-abbreviation annskip ("SKIP") where "SKIP == Basic id"
-
-types 'a sem = "'a option => 'a option => bool"
-
-consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
-primrec
-"iter 0 b S = (\<lambda>s s'. s \<notin> Some ` b \<and> s=s')"
-"iter (Suc n) b S =
-  (\<lambda>s s'. s \<in> Some ` b \<and> (\<exists>s''. S s s'' \<and> iter n b S s'' s'))"
-
-consts Sem :: "'a com => 'a sem"
-primrec
-"Sem(Basic f) s s' = (case s of None \<Rightarrow> s' = None | Some t \<Rightarrow> s' = Some(f t))"
-"Sem Abort s s' = (s' = None)"
-"Sem(c1;c2) s s' = (\<exists>s''. Sem c1 s s'' \<and> Sem c2 s'' s')"
-"Sem(IF b THEN c1 ELSE c2 FI) s s' =
- (case s of None \<Rightarrow> s' = None
-  | Some t \<Rightarrow> ((t \<in> b \<longrightarrow> Sem c1 s s') \<and> (t \<notin> b \<longrightarrow> Sem c2 s s')))"
-"Sem(While b x c) s s' =
- (if s = None then s' = None else \<exists>n. iter n b (Sem c) s s')"
-
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
-  "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
-
-
-
-(** parse translations **)
-
-syntax
-  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
-
-syntax
-  "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
-                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
-syntax ("" output)
-  "_hoare"      :: "['a assn,'a com,'a assn] => bool"
-                 ("{_} // _ // {_}" [0,55,0] 50)
-ML {*
-
-local
-fun free a = Free(a,dummyT)
-fun abs((a,T),body) =
-  let val a = absfree(a, dummyT, body)
-  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
-in
-
-fun mk_abstuple [x] body = abs (x, body)
-  | mk_abstuple (x::xs) body =
-      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
-
-fun mk_fbody a e [x as (b,_)] = if a=b then e else free b
-  | mk_fbody a e ((b,_)::xs) =
-      Syntax.const @{const_syntax Pair} $ (if a=b then e else free b) $ mk_fbody a e xs;
-
-fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
-end
-*}
-
-(* bexp_tr & assn_tr *)
-(*all meta-variables for bexp except for TRUE are translated as if they
-  were boolean expressions*)
-ML{*
-fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
-  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
-
-fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
-*}
-(* com_tr *)
-ML{*
-fun com_tr (Const (@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
-      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
-  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
-  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
-      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
-  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
-      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
-  | com_tr t _ = t (* if t is just a Free/Var *)
-*}
-
-(* triple_tr *)  (* FIXME does not handle "_idtdummy" *)
-ML{*
-local
-
-fun var_tr (Free (a, _)) = (a, Bound 0) (* Bound 0 = dummy term *)
-  | var_tr (Const (@{syntax_const "_constrain"}, _) $ Free (a, _) $ T) = (a, T);
-
-fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
-  | vars_tr t = [var_tr t]
-
-in
-fun hoare_vars_tr [vars, pre, prg, post] =
-      let val xs = vars_tr vars
-      in Syntax.const @{const_syntax Valid} $
-         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
-      end
-  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
-end
-*}
-
-parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
-
-
-(*****************************************************************************)
-
-(*** print translations ***)
-ML{*
-fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
-      subst_bound (Syntax.free v, dest_abstuple body)
-  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
-  | dest_abstuple trm = trm;
-
-fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
-  | abs2list (Abs(x,T,t)) = [Free (x, T)]
-  | abs2list _ = [];
-
-fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
-  | mk_ts (Abs(x,_,t)) = mk_ts t
-  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
-  | mk_ts t = [t];
-
-fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
-           ((Syntax.free x)::(abs2list t), mk_ts t)
-  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
-  | mk_vts t = raise Match;
-
-fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
-  | find_ch ((v,t)::vts) i xs =
-      if t = Bound i then find_ch vts (i-1) xs
-      else (true, (v, subst_bounds (xs,t)));
-
-fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
-  | is_f (Abs(x,_,t)) = true
-  | is_f t = false;
-*}
-
-(* assn_tr' & bexp_tr'*)
-ML{*
-fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | assn_tr' (Const (@{const_syntax inter},_) $ (Const (@{const_syntax Collect},_) $ T1) $
-        (Const (@{const_syntax Collect},_) $ T2)) =
-      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
-  | assn_tr' t = t;
-
-fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
-  | bexp_tr' t = t;
-*}
-
-(*com_tr' *)
-ML{*
-fun mk_assign f =
-  let val (vs, ts) = mk_vts f;
-      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
-  in
-    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
-    else Syntax.const @{const_syntax annskip}
-  end;
-
-fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
-      if is_f f then mk_assign f else Syntax.const @{const_syntax Basic} $ f
-  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
-      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
-      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
-  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
-      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
-  | com_tr' t = t;
-
-fun spec_tr' [p, c, q] =
-  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
-*}
-
-print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
-
-(*** The proof rules ***)
-
-lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
-by (auto simp:Valid_def)
-
-lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
-by (auto simp:Valid_def)
-
-lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
-by (auto simp:Valid_def)
-
-lemma CondRule:
- "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
-  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
-by (fastsimp simp:Valid_def image_def)
-
-lemma iter_aux:
- "! s s'. Sem c s s' \<longrightarrow> s \<in> Some ` (I \<inter> b) \<longrightarrow> s' \<in> Some ` I \<Longrightarrow>
-  (\<And>s s'. s \<in> Some ` I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' \<in> Some ` (I \<inter> -b))";
-apply(unfold image_def)
-apply(induct n)
- apply clarsimp
-apply(simp (no_asm_use))
-apply blast
-done
-
-lemma WhileRule:
- "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
-apply(simp add:Valid_def)
-apply(simp (no_asm) add:image_def)
-apply clarify
-apply(drule iter_aux)
-  prefer 2 apply assumption
- apply blast
-apply blast
-done
-
-lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
-by(auto simp:Valid_def)
-
-
-subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
-
-lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
-  by blast
-
-use "hoare_tac.ML"
-
-method_setup vcg = {*
-  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
-  "verification condition generator"
-
-method_setup vcg_simp = {*
-  Scan.succeed (fn ctxt =>
-    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
-  "verification condition generator plus simplification"
-
-(* Special syntax for guarded statements and guarded array updates: *)
-
-syntax
-  guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
-  array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
-translations
-  "P \<rightarrow> c" == "IF P THEN c ELSE CONST Abort FI"
-  "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
-  (* reverse translation not possible because of duplicate "a" *)
-
-text{* Note: there is no special syntax for guarded array access. Thus
-you must write @{text"j < length a \<rightarrow> a[i] := a!j"}. *}
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare/Hoare_Logic.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -0,0 +1,245 @@
+(*  Title:      HOL/Hoare/Hoare.thy
+    Author:     Leonor Prensa Nieto & Tobias Nipkow
+    Copyright   1998 TUM
+
+Sugared semantic embedding of Hoare logic.
+Strictly speaking a shallow embedding (as implemented by Norbert Galm
+following Mike Gordon) would suffice. Maybe the datatype com comes in useful
+later.
+*)
+
+theory Hoare_Logic
+imports Main
+uses ("hoare_tac.ML")
+begin
+
+types
+    'a bexp = "'a set"
+    'a assn = "'a set"
+
+datatype
+ 'a com = Basic "'a \<Rightarrow> 'a"
+   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
+   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
+   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
+
+abbreviation annskip ("SKIP") where "SKIP == Basic id"
+
+types 'a sem = "'a => 'a => bool"
+
+consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
+primrec
+"iter 0 b S = (%s s'. s ~: b & (s=s'))"
+"iter (Suc n) b S = (%s s'. s : b & (? s''. S s s'' & iter n b S s'' s'))"
+
+consts Sem :: "'a com => 'a sem"
+primrec
+"Sem(Basic f) s s' = (s' = f s)"
+"Sem(c1;c2) s s' = (? s''. Sem c1 s s'' & Sem c2 s'' s')"
+"Sem(IF b THEN c1 ELSE c2 FI) s s' = ((s  : b --> Sem c1 s s') &
+                                      (s ~: b --> Sem c2 s s'))"
+"Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')"
+
+constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+  "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
+
+
+
+(** parse translations **)
+
+syntax
+  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
+
+syntax
+ "_hoare_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
+                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
+syntax ("" output)
+ "_hoare"      :: "['a assn,'a com,'a assn] => bool"
+                 ("{_} // _ // {_}" [0,55,0] 50)
+ML {*
+
+local
+
+fun abs((a,T),body) =
+  let val a = absfree(a, dummyT, body)
+  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
+in
+
+fun mk_abstuple [x] body = abs (x, body)
+  | mk_abstuple (x::xs) body =
+      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
+
+fun mk_fbody a e [x as (b,_)] = if a=b then e else Syntax.free b
+  | mk_fbody a e ((b,_)::xs) =
+      Syntax.const @{const_syntax Pair} $ (if a=b then e else Syntax.free b) $ mk_fbody a e xs;
+
+fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
+end
+*}
+
+(* bexp_tr & assn_tr *)
+(*all meta-variables for bexp except for TRUE are translated as if they
+  were boolean expressions*)
+ML{*
+fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
+  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
+
+fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
+*}
+(* com_tr *)
+ML{*
+fun com_tr (Const(@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
+      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
+  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
+  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
+      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
+  | com_tr t _ = t (* if t is just a Free/Var *)
+*}
+
+(* triple_tr *)    (* FIXME does not handle "_idtdummy" *)
+ML{*
+local
+
+fun var_tr(Free(a,_)) = (a,Bound 0) (* Bound 0 = dummy term *)
+  | var_tr(Const (@{syntax_const "_constrain"}, _) $ (Free (a,_)) $ T) = (a,T);
+
+fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
+  | vars_tr t = [var_tr t]
+
+in
+fun hoare_vars_tr [vars, pre, prg, post] =
+      let val xs = vars_tr vars
+      in Syntax.const @{const_syntax Valid} $
+         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
+      end
+  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
+end
+*}
+
+parse_translation {* [(@{syntax_const "_hoare_vars"}, hoare_vars_tr)] *}
+
+
+(*****************************************************************************)
+
+(*** print translations ***)
+ML{*
+fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
+                            subst_bound (Syntax.free v, dest_abstuple body)
+  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
+  | dest_abstuple trm = trm;
+
+fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
+  | abs2list (Abs(x,T,t)) = [Free (x, T)]
+  | abs2list _ = [];
+
+fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
+  | mk_ts (Abs(x,_,t)) = mk_ts t
+  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
+  | mk_ts t = [t];
+
+fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
+           ((Syntax.free x)::(abs2list t), mk_ts t)
+  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
+  | mk_vts t = raise Match;
+
+fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
+  | find_ch ((v,t)::vts) i xs =
+      if t = Bound i then find_ch vts (i-1) xs
+      else (true, (v, subst_bounds (xs, t)));
+
+fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
+  | is_f (Abs(x,_,t)) = true
+  | is_f t = false;
+*}
+
+(* assn_tr' & bexp_tr'*)
+ML{*
+fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | assn_tr' (Const (@{const_syntax inter}, _) $
+        (Const (@{const_syntax Collect},_) $ T1) $ (Const (@{const_syntax Collect},_) $ T2)) =
+      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
+  | assn_tr' t = t;
+
+fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | bexp_tr' t = t;
+*}
+
+(*com_tr' *)
+ML{*
+fun mk_assign f =
+  let val (vs, ts) = mk_vts f;
+      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
+  in
+    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
+    else Syntax.const @{const_syntax annskip}
+  end;
+
+fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
+      if is_f f then mk_assign f
+      else Syntax.const @{const_syntax Basic} $ f
+  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
+      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
+      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
+      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
+  | com_tr' t = t;
+
+fun spec_tr' [p, c, q] =
+  Syntax.const @{syntax_const "_hoare"} $ assn_tr' p $ com_tr' c $ assn_tr' q
+*}
+
+print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
+
+lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
+by (auto simp:Valid_def)
+
+lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
+by (auto simp:Valid_def)
+
+lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
+by (auto simp:Valid_def)
+
+lemma CondRule:
+ "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
+  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
+by (auto simp:Valid_def)
+
+lemma iter_aux: "! s s'. Sem c s s' --> s : I & s : b --> s' : I ==>
+       (\<And>s s'. s : I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' : I & s' ~: b)";
+apply(induct n)
+ apply clarsimp
+apply(simp (no_asm_use))
+apply blast
+done
+
+lemma WhileRule:
+ "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
+apply (clarsimp simp:Valid_def)
+apply(drule iter_aux)
+  prefer 2 apply assumption
+ apply blast
+apply blast
+done
+
+
+lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
+  by blast
+
+lemmas AbortRule = SkipRule  -- "dummy version"
+use "hoare_tac.ML"
+
+method_setup vcg = {*
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
+  "verification condition generator"
+
+method_setup vcg_simp = {*
+  Scan.succeed (fn ctxt =>
+    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
+  "verification condition generator plus simplification"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -0,0 +1,269 @@
+(*  Title:      HOL/Hoare/HoareAbort.thy
+    Author:     Leonor Prensa Nieto & Tobias Nipkow
+    Copyright   2003 TUM
+
+Like Hoare.thy, but with an Abort statement for modelling run time errors.
+*)
+
+theory Hoare_Logic_Abort
+imports Main
+uses ("hoare_tac.ML")
+begin
+
+types
+    'a bexp = "'a set"
+    'a assn = "'a set"
+
+datatype
+ 'a com = Basic "'a \<Rightarrow> 'a"
+   | Abort
+   | Seq "'a com" "'a com"               ("(_;/ _)"      [61,60] 60)
+   | Cond "'a bexp" "'a com" "'a com"    ("(1IF _/ THEN _ / ELSE _/ FI)"  [0,0,0] 61)
+   | While "'a bexp" "'a assn" "'a com"  ("(1WHILE _/ INV {_} //DO _ /OD)"  [0,0,0] 61)
+
+abbreviation annskip ("SKIP") where "SKIP == Basic id"
+
+types 'a sem = "'a option => 'a option => bool"
+
+consts iter :: "nat => 'a bexp => 'a sem => 'a sem"
+primrec
+"iter 0 b S = (\<lambda>s s'. s \<notin> Some ` b \<and> s=s')"
+"iter (Suc n) b S =
+  (\<lambda>s s'. s \<in> Some ` b \<and> (\<exists>s''. S s s'' \<and> iter n b S s'' s'))"
+
+consts Sem :: "'a com => 'a sem"
+primrec
+"Sem(Basic f) s s' = (case s of None \<Rightarrow> s' = None | Some t \<Rightarrow> s' = Some(f t))"
+"Sem Abort s s' = (s' = None)"
+"Sem(c1;c2) s s' = (\<exists>s''. Sem c1 s s'' \<and> Sem c2 s'' s')"
+"Sem(IF b THEN c1 ELSE c2 FI) s s' =
+ (case s of None \<Rightarrow> s' = None
+  | Some t \<Rightarrow> ((t \<in> b \<longrightarrow> Sem c1 s s') \<and> (t \<notin> b \<longrightarrow> Sem c2 s s')))"
+"Sem(While b x c) s s' =
+ (if s = None then s' = None else \<exists>n. iter n b (Sem c) s s')"
+
+constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+  "Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
+
+
+
+(** parse translations **)
+
+syntax
+  "_assign"  :: "id => 'b => 'a com"        ("(2_ :=/ _)" [70,65] 61)
+
+syntax
+  "_hoare_abort_vars" :: "[idts, 'a assn,'a com,'a assn] => bool"
+                 ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
+syntax ("" output)
+  "_hoare_abort"      :: "['a assn,'a com,'a assn] => bool"
+                 ("{_} // _ // {_}" [0,55,0] 50)
+ML {*
+
+local
+fun free a = Free(a,dummyT)
+fun abs((a,T),body) =
+  let val a = absfree(a, dummyT, body)
+  in if T = Bound 0 then a else Const(Syntax.constrainAbsC,dummyT) $ a $ T end
+in
+
+fun mk_abstuple [x] body = abs (x, body)
+  | mk_abstuple (x::xs) body =
+      Syntax.const @{const_syntax split} $ abs (x, mk_abstuple xs body);
+
+fun mk_fbody a e [x as (b,_)] = if a=b then e else free b
+  | mk_fbody a e ((b,_)::xs) =
+      Syntax.const @{const_syntax Pair} $ (if a=b then e else free b) $ mk_fbody a e xs;
+
+fun mk_fexp a e xs = mk_abstuple xs (mk_fbody a e xs)
+end
+*}
+
+(* bexp_tr & assn_tr *)
+(*all meta-variables for bexp except for TRUE are translated as if they
+  were boolean expressions*)
+ML{*
+fun bexp_tr (Const ("TRUE", _)) xs = Syntax.const "TRUE"   (* FIXME !? *)
+  | bexp_tr b xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs b;
+
+fun assn_tr r xs = Syntax.const @{const_syntax Collect} $ mk_abstuple xs r;
+*}
+(* com_tr *)
+ML{*
+fun com_tr (Const (@{syntax_const "_assign"},_) $ Free (a,_) $ e) xs =
+      Syntax.const @{const_syntax Basic} $ mk_fexp a e xs
+  | com_tr (Const (@{const_syntax Basic},_) $ f) xs = Syntax.const @{const_syntax Basic} $ f
+  | com_tr (Const (@{const_syntax Seq},_) $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Seq} $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) xs =
+      Syntax.const @{const_syntax Cond} $ bexp_tr b xs $ com_tr c1 xs $ com_tr c2 xs
+  | com_tr (Const (@{const_syntax While},_) $ b $ I $ c) xs =
+      Syntax.const @{const_syntax While} $ bexp_tr b xs $ assn_tr I xs $ com_tr c xs
+  | com_tr t _ = t (* if t is just a Free/Var *)
+*}
+
+(* triple_tr *)  (* FIXME does not handle "_idtdummy" *)
+ML{*
+local
+
+fun var_tr (Free (a, _)) = (a, Bound 0) (* Bound 0 = dummy term *)
+  | var_tr (Const (@{syntax_const "_constrain"}, _) $ Free (a, _) $ T) = (a, T);
+
+fun vars_tr (Const (@{syntax_const "_idts"}, _) $ idt $ vars) = var_tr idt :: vars_tr vars
+  | vars_tr t = [var_tr t]
+
+in
+fun hoare_vars_tr [vars, pre, prg, post] =
+      let val xs = vars_tr vars
+      in Syntax.const @{const_syntax Valid} $
+         assn_tr pre xs $ com_tr prg xs $ assn_tr post xs
+      end
+  | hoare_vars_tr ts = raise TERM ("hoare_vars_tr", ts);
+end
+*}
+
+parse_translation {* [(@{syntax_const "_hoare_abort_vars"}, hoare_vars_tr)] *}
+
+
+(*****************************************************************************)
+
+(*** print translations ***)
+ML{*
+fun dest_abstuple (Const (@{const_syntax split},_) $ (Abs(v,_, body))) =
+      subst_bound (Syntax.free v, dest_abstuple body)
+  | dest_abstuple (Abs(v,_, body)) = subst_bound (Syntax.free v, body)
+  | dest_abstuple trm = trm;
+
+fun abs2list (Const (@{const_syntax split},_) $ (Abs(x,T,t))) = Free (x, T)::abs2list t
+  | abs2list (Abs(x,T,t)) = [Free (x, T)]
+  | abs2list _ = [];
+
+fun mk_ts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = mk_ts t
+  | mk_ts (Abs(x,_,t)) = mk_ts t
+  | mk_ts (Const (@{const_syntax Pair},_) $ a $ b) = a::(mk_ts b)
+  | mk_ts t = [t];
+
+fun mk_vts (Const (@{const_syntax split},_) $ (Abs(x,_,t))) =
+           ((Syntax.free x)::(abs2list t), mk_ts t)
+  | mk_vts (Abs(x,_,t)) = ([Syntax.free x], [t])
+  | mk_vts t = raise Match;
+
+fun find_ch [] i xs = (false, (Syntax.free "not_ch", Syntax.free "not_ch"))
+  | find_ch ((v,t)::vts) i xs =
+      if t = Bound i then find_ch vts (i-1) xs
+      else (true, (v, subst_bounds (xs,t)));
+
+fun is_f (Const (@{const_syntax split},_) $ (Abs(x,_,t))) = true
+  | is_f (Abs(x,_,t)) = true
+  | is_f t = false;
+*}
+
+(* assn_tr' & bexp_tr'*)
+ML{*
+fun assn_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | assn_tr' (Const (@{const_syntax inter},_) $ (Const (@{const_syntax Collect},_) $ T1) $
+        (Const (@{const_syntax Collect},_) $ T2)) =
+      Syntax.const @{const_syntax inter} $ dest_abstuple T1 $ dest_abstuple T2
+  | assn_tr' t = t;
+
+fun bexp_tr' (Const (@{const_syntax Collect},_) $ T) = dest_abstuple T
+  | bexp_tr' t = t;
+*}
+
+(*com_tr' *)
+ML{*
+fun mk_assign f =
+  let val (vs, ts) = mk_vts f;
+      val (ch, which) = find_ch (vs~~ts) ((length vs)-1) (rev vs)
+  in
+    if ch then Syntax.const @{syntax_const "_assign"} $ fst which $ snd which
+    else Syntax.const @{const_syntax annskip}
+  end;
+
+fun com_tr' (Const (@{const_syntax Basic},_) $ f) =
+      if is_f f then mk_assign f else Syntax.const @{const_syntax Basic} $ f
+  | com_tr' (Const (@{const_syntax Seq},_) $ c1 $ c2) =
+      Syntax.const @{const_syntax Seq} $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax Cond},_) $ b $ c1 $ c2) =
+      Syntax.const @{const_syntax Cond} $ bexp_tr' b $ com_tr' c1 $ com_tr' c2
+  | com_tr' (Const (@{const_syntax While},_) $ b $ I $ c) =
+      Syntax.const @{const_syntax While} $ bexp_tr' b $ assn_tr' I $ com_tr' c
+  | com_tr' t = t;
+
+fun spec_tr' [p, c, q] =
+  Syntax.const @{syntax_const "_hoare_abort"} $ assn_tr' p $ com_tr' c $ assn_tr' q
+*}
+
+print_translation {* [(@{const_syntax Valid}, spec_tr')] *}
+
+(*** The proof rules ***)
+
+lemma SkipRule: "p \<subseteq> q \<Longrightarrow> Valid p (Basic id) q"
+by (auto simp:Valid_def)
+
+lemma BasicRule: "p \<subseteq> {s. f s \<in> q} \<Longrightarrow> Valid p (Basic f) q"
+by (auto simp:Valid_def)
+
+lemma SeqRule: "Valid P c1 Q \<Longrightarrow> Valid Q c2 R \<Longrightarrow> Valid P (c1;c2) R"
+by (auto simp:Valid_def)
+
+lemma CondRule:
+ "p \<subseteq> {s. (s \<in> b \<longrightarrow> s \<in> w) \<and> (s \<notin> b \<longrightarrow> s \<in> w')}
+  \<Longrightarrow> Valid w c1 q \<Longrightarrow> Valid w' c2 q \<Longrightarrow> Valid p (Cond b c1 c2) q"
+by (fastsimp simp:Valid_def image_def)
+
+lemma iter_aux:
+ "! s s'. Sem c s s' \<longrightarrow> s \<in> Some ` (I \<inter> b) \<longrightarrow> s' \<in> Some ` I \<Longrightarrow>
+  (\<And>s s'. s \<in> Some ` I \<Longrightarrow> iter n b (Sem c) s s' \<Longrightarrow> s' \<in> Some ` (I \<inter> -b))";
+apply(unfold image_def)
+apply(induct n)
+ apply clarsimp
+apply(simp (no_asm_use))
+apply blast
+done
+
+lemma WhileRule:
+ "p \<subseteq> i \<Longrightarrow> Valid (i \<inter> b) c i \<Longrightarrow> i \<inter> (-b) \<subseteq> q \<Longrightarrow> Valid p (While b i c) q"
+apply(simp add:Valid_def)
+apply(simp (no_asm) add:image_def)
+apply clarify
+apply(drule iter_aux)
+  prefer 2 apply assumption
+ apply blast
+apply blast
+done
+
+lemma AbortRule: "p \<subseteq> {s. False} \<Longrightarrow> Valid p Abort q"
+by(auto simp:Valid_def)
+
+
+subsection {* Derivation of the proof rules and, most importantly, the VCG tactic *}
+
+lemma Compl_Collect: "-(Collect b) = {x. ~(b x)}"
+  by blast
+
+use "hoare_tac.ML"
+
+method_setup vcg = {*
+  Scan.succeed (fn ctxt => SIMPLE_METHOD' (hoare_tac ctxt (K all_tac))) *}
+  "verification condition generator"
+
+method_setup vcg_simp = {*
+  Scan.succeed (fn ctxt =>
+    SIMPLE_METHOD' (hoare_tac ctxt (asm_full_simp_tac (simpset_of ctxt)))) *}
+  "verification condition generator plus simplification"
+
+(* Special syntax for guarded statements and guarded array updates: *)
+
+syntax
+  guarded_com :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com"  ("(2_ \<rightarrow>/ _)" 71)
+  array_update :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com"  ("(2_[_] :=/ _)" [70, 65] 61)
+translations
+  "P \<rightarrow> c" == "IF P THEN c ELSE CONST Abort FI"
+  "a[i] := v" => "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
+  (* reverse translation not possible because of duplicate "a" *)
+
+text{* Note: there is no special syntax for guarded array access. Thus
+you must write @{text"j < length a \<rightarrow> a[i] := a!j"}. *}
+
+end
--- a/src/HOL/Hoare/Pointer_Examples.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/Pointer_Examples.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Hoare/Pointers.thy
-    ID:         $Id$
     Author:     Tobias Nipkow
     Copyright   2002 TUM
 
--- a/src/HOL/Hoare/Pointers0.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/Pointers0.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -9,12 +9,12 @@
 - in fact in some case they appear to get (a bit) more complicated.
 *)
 
-theory Pointers0 imports Hoare begin
+theory Pointers0 imports Hoare_Logic begin
 
 subsection "References"
 
-axclass ref < type
-consts Null :: "'a::ref"
+class ref =
+  fixes Null :: 'a
 
 subsection "Field access and update"
 
--- a/src/HOL/Hoare/ROOT.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/ROOT.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -1,8 +1,2 @@
-(*  Title:      HOL/Hoare/ROOT.ML
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   1998-2003 TUM
-*)
 
-use_thys ["Examples", "ExamplesAbort", "Pointers0", "Pointer_Examples",
-  "Pointer_ExamplesAbort", "SchorrWaite", "Separation"];
+use_thy "Hoare";
--- a/src/HOL/Hoare/Separation.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Hoare/Separation.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -12,7 +12,7 @@
 
 *)
 
-theory Separation imports HoareAbort SepLogHeap begin
+theory Separation imports Hoare_Logic_Abort SepLogHeap begin
 
 text{* The semantic definition of a few connectives: *}
 
--- a/src/HOL/IsaMakefile	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/IsaMakefile	Tue Feb 23 14:44:43 2010 -0800
@@ -66,7 +66,6 @@
       TLA-Memory \
   HOL-UNITY \
   HOL-Unix \
-  HOL-W0 \
   HOL-Word-Examples \
   HOL-ZF
     # ^ this is the sort position
@@ -142,7 +141,6 @@
 	@$(ISABELLE_TOOL) usedir -b -f base.ML -d false -g false $(OUT)/Pure HOL-Base
 
 PLAIN_DEPENDENCIES = $(BASE_DEPENDENCIES)\
-  Algebras.thy \
   Complete_Lattice.thy \
   Datatype.thy \
   Extraction.thy \
@@ -385,7 +383,7 @@
 $(LOG)/HOL-Library.gz: $(OUT)/HOL Library/SetsAndFunctions.thy		\
   Library/Abstract_Rat.thy Library/BigO.thy Library/ContNotDenum.thy	\
   Library/Efficient_Nat.thy Library/Sum_Of_Squares.thy			\
-  Library/Sum_Of_Squares/sos_wrapper.ML					\
+  Library/Dlist.thy Library/Sum_Of_Squares/sos_wrapper.ML		\
   Library/Sum_Of_Squares/sum_of_squares.ML Library/Fset.thy		\
   Library/Glbs.thy Library/normarith.ML Library/Executable_Set.thy	\
   Library/Infinite_Set.thy Library/FuncSet.thy				\
@@ -591,9 +589,10 @@
 
 $(LOG)/HOL-Hoare.gz: $(OUT)/HOL Hoare/Arith2.thy Hoare/Examples.thy	\
   Hoare/Hoare.thy Hoare/hoare_tac.ML Hoare/Heap.thy			\
+  Hoare/Hoare_Logic.thy	Hoare/Hoare_Logic_Abort.thy			\
   Hoare/HeapSyntax.thy Hoare/Pointer_Examples.thy Hoare/ROOT.ML		\
   Hoare/ExamplesAbort.thy Hoare/HeapSyntaxAbort.thy			\
-  Hoare/HoareAbort.thy Hoare/SchorrWaite.thy Hoare/Separation.thy	\
+  Hoare/SchorrWaite.thy Hoare/Separation.thy				\
   Hoare/SepLogHeap.thy Hoare/document/root.tex Hoare/document/root.bib
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Hoare
 
@@ -847,14 +846,6 @@
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Prolog
 
 
-## HOL-W0
-
-HOL-W0: HOL $(LOG)/HOL-W0.gz
-
-$(LOG)/HOL-W0.gz: $(OUT)/HOL W0/ROOT.ML W0/W0.thy W0/document/root.tex
-	@$(ISABELLE_TOOL) usedir $(OUT)/HOL W0
-
-
 ## HOL-MicroJava
 
 HOL-MicroJava: HOL $(LOG)/HOL-MicroJava.gz
@@ -964,11 +955,12 @@
   ex/Codegenerator_Pretty.thy ex/Codegenerator_Pretty_Test.thy		\
   ex/Codegenerator_Test.thy ex/Coherent.thy				\
   ex/Efficient_Nat_examples.thy ex/Eval_Examples.thy ex/Fundefs.thy	\
-  ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
-  ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
-  ex/Hilbert_Classical.thy ex/Induction_Schema.thy			\
-  ex/InductiveInvariant.thy ex/InductiveInvariant_examples.thy		\
-  ex/Intuitionistic.thy ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy	\
+  ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy	\
+  ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy		\
+  ex/Higher_Order_Logic.thy ex/Hilbert_Classical.thy			\
+  ex/Induction_Schema.thy ex/InductiveInvariant.thy			\
+  ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy		\
+  ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy				\
   ex/MergeSort.thy ex/Meson_Test.thy ex/MonoidGroup.thy			\
   ex/Multiquote.thy ex/NatSum.thy ex/Numeral.thy ex/PER.thy		\
   ex/Predicate_Compile_ex.thy ex/Predicate_Compile_Quickcheck.thy	\
@@ -1320,7 +1312,7 @@
 		$(LOG)/HOL-SMT-Examples.gz $(LOG)/HOL-SMT.gz		\
 		$(LOG)/HOL-Statespace.gz $(LOG)/HOL-Subst.gz		\
 		$(LOG)/HOL-UNITY.gz $(LOG)/HOL-Unix.gz			\
-		$(LOG)/HOL-W0.gz $(LOG)/HOL-Word-Examples.gz		\
+		$(LOG)/HOL-Word-Examples.gz		\
 		$(LOG)/HOL-Word.gz $(LOG)/HOL-ZF.gz $(LOG)/HOL-ex.gz	\
 		$(LOG)/HOL.gz $(LOG)/HOL4.gz $(LOG)/TLA-Buffer.gz	\
 		$(LOG)/TLA-Inc.gz $(LOG)/TLA-Memory.gz $(LOG)/TLA.gz	\
--- a/src/HOL/Isar_Examples/Group.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Isar_Examples/Group.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -17,15 +17,12 @@
  $\idt{times}$ is provided by the basic HOL theory.
 *}
 
-consts
-  one :: "'a"
-  inverse :: "'a => 'a"
+notation Groups.one ("one")
 
-axclass
-  group < times
-  group_assoc:         "(x * y) * z = x * (y * z)"
-  group_left_one:      "one * x = x"
-  group_left_inverse:  "inverse x * x = one"
+class group = times + one + inverse +
+  assumes group_assoc:         "(x * y) * z = x * (y * z)"
+  assumes group_left_one:      "one * x = x"
+  assumes group_left_inverse:  "inverse x * x = one"
 
 text {*
  The group axioms only state the properties of left one and inverse,
@@ -144,10 +141,10 @@
  \idt{one} :: \alpha)$ are defined like this.
 *}
 
-axclass monoid < times
-  monoid_assoc:       "(x * y) * z = x * (y * z)"
-  monoid_left_one:   "one * x = x"
-  monoid_right_one:  "x * one = x"
+class monoid = times + one +
+  assumes monoid_assoc:       "(x * y) * z = x * (y * z)"
+  assumes monoid_left_one:   "one * x = x"
+  assumes monoid_right_one:  "x * one = x"
 
 text {*
  Groups are \emph{not} yet monoids directly from the definition.  For
--- a/src/HOL/Lattice/Bounds.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Lattice/Bounds.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/Bounds.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
--- a/src/HOL/Lattice/CompleteLattice.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Lattice/CompleteLattice.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/CompleteLattice.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -16,8 +15,8 @@
   bounds (see \S\ref{sec:connect-bounds}).
 *}
 
-axclass complete_lattice \<subseteq> partial_order
-  ex_Inf: "\<exists>inf. is_Inf A inf"
+class complete_lattice =
+  assumes ex_Inf: "\<exists>inf. is_Inf A inf"
 
 theorem ex_Sup: "\<exists>sup::'a::complete_lattice. is_Sup A sup"
 proof -
--- a/src/HOL/Lattice/Lattice.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Lattice/Lattice.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/Lattice.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -15,9 +14,9 @@
   as well).
 *}
 
-axclass lattice \<subseteq> partial_order
-  ex_inf: "\<exists>inf. is_inf x y inf"
-  ex_sup: "\<exists>sup. is_sup x y sup"
+class lattice =
+  assumes ex_inf: "\<exists>inf. is_inf x y inf"
+  assumes ex_sup: "\<exists>sup. is_sup x y sup"
 
 text {*
   The @{text \<sqinter>} (meet) and @{text \<squnion>} (join) operations select such
--- a/src/HOL/Lattice/Orders.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Lattice/Orders.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/Lattice/Orders.thy
-    ID:         $Id$
     Author:     Markus Wenzel, TU Muenchen
 *)
 
@@ -18,21 +17,21 @@
   required to be related (in either direction).
 *}
 
-axclass leq < type
-consts
-  leq :: "'a::leq \<Rightarrow> 'a \<Rightarrow> bool"  (infixl "[=" 50)
+class leq =
+  fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infixl "[=" 50)
+
 notation (xsymbols)
   leq  (infixl "\<sqsubseteq>" 50)
 
-axclass quasi_order < leq
-  leq_refl [intro?]: "x \<sqsubseteq> x"
-  leq_trans [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
+class quasi_order = leq +
+  assumes leq_refl [intro?]: "x \<sqsubseteq> x"
+  assumes leq_trans [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z"
 
-axclass partial_order < quasi_order
-  leq_antisym [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
+class partial_order = quasi_order +
+  assumes leq_antisym [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y"
 
-axclass linear_order < partial_order
-  leq_linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
+class linear_order = partial_order +
+  assumes leq_linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
 
 lemma linear_order_cases:
     "((x::'a::linear_order) \<sqsubseteq> y \<Longrightarrow> C) \<Longrightarrow> (y \<sqsubseteq> x \<Longrightarrow> C) \<Longrightarrow> C"
@@ -54,11 +53,16 @@
 primrec
   undual_dual: "undual (dual x) = x"
 
-instance dual :: (leq) leq ..
+instantiation dual :: (leq) leq
+begin
 
-defs (overloaded)
+definition
   leq_dual_def: "x' \<sqsubseteq> y' \<equiv> undual y' \<sqsubseteq> undual x'"
 
+instance ..
+
+end
+
 lemma undual_leq [iff?]: "(undual x' \<sqsubseteq> undual y') = (y' \<sqsubseteq> x')"
   by (simp add: leq_dual_def)
 
@@ -192,11 +196,16 @@
   \emph{not} be linear in general.
 *}
 
-instance * :: (leq, leq) leq ..
+instantiation * :: (leq, leq) leq
+begin
 
-defs (overloaded)
+definition
   leq_prod_def: "p \<sqsubseteq> q \<equiv> fst p \<sqsubseteq> fst q \<and> snd p \<sqsubseteq> snd q"
 
+instance ..
+
+end
+
 lemma leq_prodI [intro?]:
     "fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q"
   by (unfold leq_prod_def) blast
@@ -249,11 +258,16 @@
   orders need \emph{not} be linear in general.
 *}
 
-instance "fun" :: (type, leq) leq ..
+instantiation "fun" :: (type, leq) leq
+begin
 
-defs (overloaded)
+definition
   leq_fun_def: "f \<sqsubseteq> g \<equiv> \<forall>x. f x \<sqsubseteq> g x"
 
+instance ..
+
+end
+
 lemma leq_funI [intro?]: "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> f \<sqsubseteq> g"
   by (unfold leq_fun_def) blast
 
--- a/src/HOL/Lattices.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Lattices.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -8,7 +8,42 @@
 imports Orderings Groups
 begin
 
-subsection {* Lattices *}
+subsection {* Abstract semilattice *}
+
+text {*
+  This locales provide a basic structure for interpretation into
+  bigger structures;  extensions require careful thinking, otherwise
+  undesired effects may occur due to interpretation.
+*}
+
+locale semilattice = abel_semigroup +
+  assumes idem [simp]: "f a a = a"
+begin
+
+lemma left_idem [simp]:
+  "f a (f a b) = f a b"
+  by (simp add: assoc [symmetric])
+
+end
+
+
+subsection {* Idempotent semigroup *}
+
+class ab_semigroup_idem_mult = ab_semigroup_mult +
+  assumes mult_idem: "x * x = x"
+
+sublocale ab_semigroup_idem_mult < times!: semilattice times proof
+qed (fact mult_idem)
+
+context ab_semigroup_idem_mult
+begin
+
+lemmas mult_left_idem = times.left_idem
+
+end
+
+
+subsection {* Conrete lattices *}
 
 notation
   less_eq  (infix "\<sqsubseteq>" 50) and
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Dlist.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -0,0 +1,256 @@
+(* Author: Florian Haftmann, TU Muenchen *)
+
+header {* Lists with elements distinct as canonical example for datatype invariants *}
+
+theory Dlist
+imports Main Fset
+begin
+
+section {* Prelude *}
+
+text {* Without canonical argument order, higher-order things tend to get confusing quite fast: *}
+
+setup {* Sign.map_naming (Name_Space.add_path "List") *}
+
+primrec member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
+    "member [] y \<longleftrightarrow> False"
+  | "member (x#xs) y \<longleftrightarrow> x = y \<or> member xs y"
+
+lemma member_set:
+  "member = set"
+proof (rule ext)+
+  fix xs :: "'a list" and x :: 'a
+  have "member xs x \<longleftrightarrow> x \<in> set xs" by (induct xs) auto
+  then show "member xs x = set xs x" by (simp add: mem_def)
+qed
+
+lemma not_set_compl:
+  "Not \<circ> set xs = - set xs"
+  by (simp add: fun_Compl_def bool_Compl_def comp_def expand_fun_eq)
+
+primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where
+    "fold f [] s = s"
+  | "fold f (x#xs) s = fold f xs (f x s)"
+
+lemma foldl_fold:
+  "foldl f s xs = List.fold (\<lambda>x s. f s x) xs s"
+  by (induct xs arbitrary: s) simp_all
+
+setup {* Sign.map_naming Name_Space.parent_path *}
+
+
+section {* The type of distinct lists *}
+
+typedef (open) 'a dlist = "{xs::'a list. distinct xs}"
+  morphisms list_of_dlist Abs_dlist
+proof
+  show "[] \<in> ?dlist" by simp
+qed
+
+text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
+
+definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
+  [code del]: "Dlist xs = Abs_dlist (remdups xs)"
+
+lemma distinct_list_of_dlist [simp]:
+  "distinct (list_of_dlist dxs)"
+  using list_of_dlist [of dxs] by simp
+
+lemma list_of_dlist_Dlist [simp]:
+  "list_of_dlist (Dlist xs) = remdups xs"
+  by (simp add: Dlist_def Abs_dlist_inverse)
+
+lemma Dlist_list_of_dlist [simp]:
+  "Dlist (list_of_dlist dxs) = dxs"
+  by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
+
+
+text {* Fundamental operations: *}
+
+definition empty :: "'a dlist" where
+  "empty = Dlist []"
+
+definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
+  "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"
+
+definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
+  "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"
+
+definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
+  "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"
+
+definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
+  "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
+
+
+text {* Derived operations: *}
+
+definition null :: "'a dlist \<Rightarrow> bool" where
+  "null dxs = List.null (list_of_dlist dxs)"
+
+definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where
+  "member dxs = List.member (list_of_dlist dxs)"
+
+definition length :: "'a dlist \<Rightarrow> nat" where
+  "length dxs = List.length (list_of_dlist dxs)"
+
+definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
+  "fold f dxs = List.fold f (list_of_dlist dxs)"
+
+
+section {* Executable version obeying invariant *}
+
+code_abstype Dlist list_of_dlist
+  by simp
+
+lemma list_of_dlist_empty [simp, code abstract]:
+  "list_of_dlist empty = []"
+  by (simp add: empty_def)
+
+lemma list_of_dlist_insert [simp, code abstract]:
+  "list_of_dlist (insert x dxs) = List.insert x (list_of_dlist dxs)"
+  by (simp add: insert_def)
+
+lemma list_of_dlist_remove [simp, code abstract]:
+  "list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)"
+  by (simp add: remove_def)
+
+lemma list_of_dlist_map [simp, code abstract]:
+  "list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))"
+  by (simp add: map_def)
+
+lemma list_of_dlist_filter [simp, code abstract]:
+  "list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"
+  by (simp add: filter_def)
+
+declare null_def [code] member_def [code] length_def [code] fold_def [code] -- {* explicit is better than implicit *}
+
+
+section {* Implementation of sets by distinct lists -- canonical! *}
+
+definition Set :: "'a dlist \<Rightarrow> 'a fset" where
+  "Set dxs = Fset.Set (list_of_dlist dxs)"
+
+definition Coset :: "'a dlist \<Rightarrow> 'a fset" where
+  "Coset dxs = Fset.Coset (list_of_dlist dxs)"
+
+code_datatype Set Coset
+
+declare member_code [code del]
+declare is_empty_Set [code del]
+declare empty_Set [code del]
+declare UNIV_Set [code del]
+declare insert_Set [code del]
+declare remove_Set [code del]
+declare map_Set [code del]
+declare filter_Set [code del]
+declare forall_Set [code del]
+declare exists_Set [code del]
+declare card_Set [code del]
+declare subfset_eq_forall [code del]
+declare subfset_subfset_eq [code del]
+declare eq_fset_subfset_eq [code del]
+declare inter_project [code del]
+declare subtract_remove [code del]
+declare union_insert [code del]
+declare Infimum_inf [code del]
+declare Supremum_sup [code del]
+
+lemma Set_Dlist [simp]:
+  "Set (Dlist xs) = Fset (set xs)"
+  by (simp add: Set_def Fset.Set_def)
+
+lemma Coset_Dlist [simp]:
+  "Coset (Dlist xs) = Fset (- set xs)"
+  by (simp add: Coset_def Fset.Coset_def)
+
+lemma member_Set [simp]:
+  "Fset.member (Set dxs) = List.member (list_of_dlist dxs)"
+  by (simp add: Set_def member_set)
+
+lemma member_Coset [simp]:
+  "Fset.member (Coset dxs) = Not \<circ> List.member (list_of_dlist dxs)"
+  by (simp add: Coset_def member_set not_set_compl)
+
+lemma is_empty_Set [code]:
+  "Fset.is_empty (Set dxs) \<longleftrightarrow> null dxs"
+  by (simp add: null_def null_empty member_set)
+
+lemma bot_code [code]:
+  "bot = Set empty"
+  by (simp add: empty_def)
+
+lemma top_code [code]:
+  "top = Coset empty"
+  by (simp add: empty_def)
+
+lemma insert_code [code]:
+  "Fset.insert x (Set dxs) = Set (insert x dxs)"
+  "Fset.insert x (Coset dxs) = Coset (remove x dxs)"
+  by (simp_all add: insert_def remove_def member_set not_set_compl)
+
+lemma remove_code [code]:
+  "Fset.remove x (Set dxs) = Set (remove x dxs)"
+  "Fset.remove x (Coset dxs) = Coset (insert x dxs)"
+  by (auto simp add: insert_def remove_def member_set not_set_compl)
+
+lemma member_code [code]:
+  "Fset.member (Set dxs) = member dxs"
+  "Fset.member (Coset dxs) = Not \<circ> member dxs"
+  by (simp_all add: member_def)
+
+lemma map_code [code]:
+  "Fset.map f (Set dxs) = Set (map f dxs)"
+  by (simp add: member_set)
+  
+lemma filter_code [code]:
+  "Fset.filter f (Set dxs) = Set (filter f dxs)"
+  by (simp add: member_set)
+
+lemma forall_Set [code]:
+  "Fset.forall P (Set xs) \<longleftrightarrow> list_all P (list_of_dlist xs)"
+  by (simp add: member_set list_all_iff)
+
+lemma exists_Set [code]:
+  "Fset.exists P (Set xs) \<longleftrightarrow> list_ex P (list_of_dlist xs)"
+  by (simp add: member_set list_ex_iff)
+
+lemma card_code [code]:
+  "Fset.card (Set dxs) = length dxs"
+  by (simp add: length_def member_set distinct_card)
+
+lemma foldl_list_of_dlist:
+  "foldl f s (list_of_dlist dxs) = fold (\<lambda>x s. f s x) dxs s"
+  by (simp add: foldl_fold fold_def)
+
+lemma inter_code [code]:
+  "inf A (Set xs) = Set (filter (Fset.member A) xs)"
+  "inf A (Coset xs) = fold Fset.remove xs A"
+  by (simp_all only: Set_def Coset_def foldl_list_of_dlist inter_project list_of_dlist_filter)
+
+lemma subtract_code [code]:
+  "A - Set xs = fold Fset.remove xs A"
+  "A - Coset xs = Set (filter (Fset.member A) xs)"
+  by (simp_all only: Set_def Coset_def foldl_list_of_dlist subtract_remove list_of_dlist_filter)
+
+lemma union_code [code]:
+  "sup (Set xs) A = fold Fset.insert xs A"
+  "sup (Coset xs) A = Coset (filter (Not \<circ> Fset.member A) xs)"
+  by (simp_all only: Set_def Coset_def foldl_list_of_dlist union_insert list_of_dlist_filter)
+
+context complete_lattice
+begin
+
+lemma Infimum_code [code]:
+  "Infimum (Set As) = fold inf As top"
+  by (simp only: Set_def Infimum_inf foldl_list_of_dlist inf.commute)
+
+lemma Supremum_code [code]:
+  "Supremum (Set As) = fold sup As bot"
+  by (simp only: Set_def Supremum_sup foldl_list_of_dlist sup.commute)
+
+end
+
+hide (open) const member fold empty insert remove map filter null member length fold
+
+end
--- a/src/HOL/Library/Library.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Library/Library.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -15,6 +15,7 @@
   ContNotDenum
   Countable
   Diagonalize
+  Dlist
   Efficient_Nat
   Enum
   Eval_Witness
--- a/src/HOL/Library/Multiset.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Library/Multiset.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1212,8 +1212,8 @@
 definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where
   "M' <=# M \<longleftrightarrow> M' <# M \<or> M' = M"
 
-notation (xsymbol) less_multiset (infix "\<subset>#" 50)
-notation (xsymbol) le_multiset (infix "\<subseteq>#" 50)
+notation (xsymbols) less_multiset (infix "\<subset>#" 50)
+notation (xsymbols) le_multiset (infix "\<subseteq>#" 50)
 
 interpretation multiset_order: order le_multiset less_multiset
 proof -
--- a/src/HOL/List.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/List.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -250,7 +250,7 @@
 @{lemma "distinct [2,0,1::nat]" by simp}\\
 @{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
 @{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
-@{lemma "List.insert 3 [0::nat,1,2] = [3, 0,1,2]" by (simp add: List.insert_def)}\\
+@{lemma "List.insert 3 [0::nat,1,2] = [3,0,1,2]" by (simp add: List.insert_def)}\\
 @{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
 @{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
 @{lemma "nth [a,b,c,d] 2 = c" by simp}\\
@@ -2900,10 +2900,14 @@
   "List.insert x [] = [x]"
   by simp
 
-lemma set_insert:
+lemma set_insert [simp]:
   "set (List.insert x xs) = insert x (set xs)"
   by (auto simp add: List.insert_def)
 
+lemma distinct_insert [simp]:
+  "distinct xs \<Longrightarrow> distinct (List.insert x xs)"
+  by (simp add: List.insert_def)
+
 
 subsubsection {* @{text remove1} *}
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mutabelle/ROOT.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -0,0 +1,3 @@
+
+use_thy "MutabelleExtra";
+
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -10,9 +10,11 @@
 val take_random : int -> 'a list -> 'a list
 
 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
-type mtd = string * (theory -> term -> outcome)
+type timing = (string * int) list
 
-type mutant_subentry = term * (string * outcome) list
+type mtd = string * (theory -> term -> outcome * timing)
+
+type mutant_subentry = term * (string * (outcome * timing)) list
 type detailed_entry = string * bool * term * mutant_subentry list
 
 type subentry = string * int * int * int * int * int * int
@@ -52,7 +54,7 @@
 
 (* quickcheck options *)
 (*val quickcheck_generator = "SML"*)
-val iterations = 100
+val iterations = 1
 val size = 5
 
 exception RANDOM;
@@ -75,12 +77,13 @@
   else l + Real.floor (rmod (random ()) (real (h - l + 1)));
 
 datatype outcome = GenuineCex | PotentialCex | NoCex | Donno | Timeout | Error
-type mtd = string * (theory -> term -> outcome)
+type timing = (string * int) list
 
-type mutant_subentry = term * (string * outcome) list
+type mtd = string * (theory -> term -> outcome * timing)
+
+type mutant_subentry = term * (string * (outcome * timing)) list
 type detailed_entry = string * bool * term * mutant_subentry list
 
-
 type subentry = string * int * int * int * int * int * int
 type entry = string * bool * subentry list
 type report = entry list
@@ -94,11 +97,11 @@
 fun invoke_quickcheck quickcheck_generator thy t =
   TimeLimit.timeLimit (Time.fromSeconds (! Auto_Counterexample.time_limit))
       (fn _ =>
-          case Quickcheck.test_term (ProofContext.init thy) false (SOME quickcheck_generator)
+          case Quickcheck.timed_test_term (ProofContext.init thy) false (SOME quickcheck_generator)
                                     size iterations (preprocess thy [] t) of
-            NONE => NoCex
-          | SOME _ => GenuineCex) ()
-  handle TimeLimit.TimeOut => Timeout
+            (NONE, time_report) => (NoCex, time_report)
+          | (SOME _, time_report) => (GenuineCex, time_report)) ()
+  handle TimeLimit.TimeOut => (Timeout, [("timelimit", !Auto_Counterexample.time_limit)])
 
 fun quickcheck_mtd quickcheck_generator =
   ("quickcheck_" ^ quickcheck_generator, invoke_quickcheck quickcheck_generator)
@@ -189,9 +192,9 @@
 
 val forbidden =
  [(* (@{const_name "power"}, "'a"), *)
-  (@{const_name HOL.induct_equal}, "'a"),
-  (@{const_name HOL.induct_implies}, "'a"),
-  (@{const_name HOL.induct_conj}, "'a"),
+  (*(@{const_name induct_equal}, "'a"),
+  (@{const_name induct_implies}, "'a"),
+  (@{const_name induct_conj}, "'a"),*)
   (@{const_name "undefined"}, "'a"),
   (@{const_name "default"}, "'a"),
   (@{const_name "dummy_pattern"}, "'a::{}") (*,
@@ -245,17 +248,26 @@
       Library.nth xs j :: take_random (n - 1) (nth_drop j xs)
     end
 
+fun cpu_time description f =
+  let
+    val start = start_timing ()
+    val result = Exn.capture f ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, (description, time)) end
+
 fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
   let
     val _ = priority ("Invoking " ^ mtd_name)
-    val res = case try (invoke_mtd thy) t of
-                SOME res => res
-              | NONE => (priority ("**** PROBLEMS WITH " ^
-                                 Syntax.string_of_term_global thy t); Error)
+    val ((res, timing), time) = cpu_time "total time"
+      (fn () => case try (invoke_mtd thy) t of
+          SOME (res, timing) => (res, timing)
+        | NONE => (priority ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
+           (Error , [])))
     val _ = priority (" Done")
-  in res end
+  in (res, time :: timing) end
 
 (* theory -> term list -> mtd -> subentry *)
+(*
 fun test_mutants_using_one_method thy mutants (mtd_name, invoke_mtd) =
   let
      val res = map (safe_invoke_mtd thy (mtd_name, invoke_mtd)) mutants
@@ -266,7 +278,7 @@
 
 fun create_entry thy thm exec mutants mtds =
   (Thm.get_name thm, exec, map (test_mutants_using_one_method thy mutants) mtds)
-
+*)
 fun create_detailed_entry thy thm exec mutants mtds =
   let
     fun create_mutant_subentry mutant = (mutant,
@@ -322,15 +334,22 @@
   | string_of_outcome Timeout = "Timeout"
   | string_of_outcome Error = "Error"
 
-fun string_of_mutant_subentry thy (t, results) =
+fun string_of_mutant_subentry thy thm_name (t, results) =
   "mutant: " ^ Syntax.string_of_term_global thy t ^ "\n" ^
-  space_implode "; " (map (fn (mtd_name, outcome) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
+  space_implode "; "
+    (map (fn (mtd_name, (outcome, timing)) => mtd_name ^ ": " ^ string_of_outcome outcome) results) ^
   "\n"
 
+fun string_of_mutant_subentry' thy thm_name (t, results) =
+  "mutant of " ^ thm_name ^ ":" ^
+    cat_lines (map (fn (mtd_name, (outcome, timing)) =>
+      mtd_name ^ ": " ^ string_of_outcome outcome ^ "; " ^
+      space_implode "; " (map (fn (s, t) => (s ^ ": " ^ string_of_int t)) timing)) results)
+
 fun string_of_detailed_entry thy (thm_name, exec, t, mutant_subentries) = 
    thm_name ^ " " ^ (if exec then "[exe]" else "[noexe]") ^ ": " ^
    Syntax.string_of_term_global thy t ^ "\n" ^
-   cat_lines (map (string_of_mutant_subentry thy) mutant_subentries) ^ "\n"
+   cat_lines (map (string_of_mutant_subentry' thy thm_name) mutant_subentries) ^ "\n"
 
 (* subentry -> string *)
 fun string_for_subentry (mtd_name, genuine_cex, potential_cex, no_cex, donno,
--- a/src/HOL/NSA/Hyperreal.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/NSA/Hyperreal.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -7,7 +7,7 @@
 *)
 
 theory Hyperreal
-imports Ln Deriv Taylor Integration HLog
+imports Ln Deriv Taylor HLog
 begin
 
 end
--- a/src/HOL/Nitpick.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Nitpick.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -36,7 +36,6 @@
            and bisim :: "bisim_iterator \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
            and bisim_iterator_max :: bisim_iterator
            and Quot :: "'a \<Rightarrow> 'b"
-           and quot_normal :: "'a \<Rightarrow> 'a"
            and Tha :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
 
 datatype ('a, 'b) pair_box = PairBox 'a 'b
@@ -237,11 +236,10 @@
 setup {* Nitpick_Isar.setup *}
 
 hide (open) const unknown is_unknown undefined_fast_The undefined_fast_Eps bisim 
-    bisim_iterator_max Quot quot_normal Tha PairBox FunBox Word refl' wf'
-    wf_wfrec wf_wfrec' wfrec' 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_eq_frac
-    of_frac
+    bisim_iterator_max Quot Tha PairBox FunBox Word refl' wf' wf_wfrec wf_wfrec'
+    wfrec' 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_eq_frac of_frac
 hide (open) type bisim_iterator pair_box fun_box unsigned_bit signed_bit word
 hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def
     wf'_def wf_wfrec'_def wfrec'_def card'_def setsum'_def fold_graph'_def
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -13,7 +13,7 @@
 
 chapter {* 3. First Steps *}
 
-nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1]
+nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 15 s]
 
 subsection {* 3.1. Propositional Logic *}
 
@@ -70,7 +70,7 @@
 oops
 
 lemma "P Suc"
-nitpick [card = 1-6]
+nitpick
 oops
 
 lemma "P (op +\<Colon>nat\<Rightarrow>nat\<Rightarrow>nat)"
@@ -149,7 +149,7 @@
 "\<lbrakk>even' m; even' n\<rbrakk> \<Longrightarrow> even' (m + n)"
 
 lemma "\<exists>n \<in> {0, 2, 4, 6, 8}. \<not> even' n"
-nitpick [card nat = 10, unary_ints, verbose, show_consts] (* FIXME: should be genuine *)
+nitpick [card nat = 10, unary_ints, verbose, show_consts]
 oops
 
 lemma "even' (n - 2) \<Longrightarrow> even' n"
@@ -210,7 +210,7 @@
 lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>1 \<sigma> t = t"
 nitpick [verbose]
 nitpick [eval = "subst\<^isub>1 \<sigma> t"]
-nitpick [dont_box]
+(* nitpick [dont_box] *)
 oops
 
 primrec subst\<^isub>2 where
@@ -220,7 +220,7 @@
 "subst\<^isub>2 \<sigma> (App t u) = App (subst\<^isub>2 \<sigma> t) (subst\<^isub>2 \<sigma> u)"
 
 lemma "\<not> loose t 0 \<Longrightarrow> subst\<^isub>2 \<sigma> t = t"
-nitpick
+nitpick [card = 1\<midarrow>6]
 sorry
 
 subsection {* 3.11. Scope Monotonicity *}
@@ -243,7 +243,7 @@
 "n \<in> reach \<Longrightarrow> n + 2 \<in> reach"
 
 lemma "n \<in> reach \<Longrightarrow> 2 dvd n"
-nitpick
+nitpick [unary_ints]
 apply (induct set: reach)
   apply auto
  nitpick
@@ -252,7 +252,7 @@
 oops
 
 lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0"
-nitpick
+nitpick [unary_ints]
 apply (induct set: reach)
   apply auto
  nitpick
@@ -289,12 +289,12 @@
           if b \<in> labels t then labels t else (labels t - {a}) \<union> {b}
         else
           if b \<in> labels t then (labels t - {b}) \<union> {a} else labels t)"
-nitpick
+(* nitpick *)
 proof (induct t)
   case Leaf thus ?case by simp
 next
   case (Branch t u) thus ?case
-  nitpick [non_std "'a bin_tree", show_consts]
+  nitpick [non_std, show_all]
   by auto
 qed
 
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -110,12 +110,12 @@
 
 lemma "\<exists>one \<in> {1}. \<exists>two \<in> {2}.
        f5 (\<lambda>a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x"
-nitpick [expect = potential] (* unfortunate *)
+nitpick [expect = genuine]
 oops
 
 lemma "\<exists>two \<in> {2}. \<exists>one \<in> {1}.
        f5 (\<lambda>a. if a = one then 2 else if a = two then 1 else a) (Suc x) = x"
-nitpick [expect = potential] (* unfortunate *)
+nitpick [expect = genuine]
 oops
 
 lemma "\<forall>a. g a = a
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -143,11 +143,11 @@
 by (rule Rep_Sum_inverse)
 
 lemma "0::nat \<equiv> Abs_Nat Zero_Rep"
-(* nitpick [expect = none] FIXME *)
+nitpick [expect = none]
 by (rule Zero_nat_def_raw)
 
 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Rep_Nat n))"
-(* nitpick [expect = none] FIXME *)
+nitpick [expect = none]
 by (rule Suc_def)
 
 lemma "Suc \<equiv> \<lambda>n. Abs_Nat (Suc_Rep (Suc_Rep (Rep_Nat n)))"
--- a/src/HOL/Orderings.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Orderings.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -5,7 +5,7 @@
 header {* Abstract orderings *}
 
 theory Orderings
-imports Algebras
+imports HOL
 uses
   "~~/src/Provers/order.ML"
   "~~/src/Provers/quasi.ML"  (* FIXME unused? *)
--- a/src/HOL/Prolog/Func.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Prolog/Func.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -5,7 +5,7 @@
 header {* Untyped functional language, with call by value semantics *}
 
 theory Func
-imports HOHH Algebras
+imports HOHH
 begin
 
 typedecl tm
--- a/src/HOL/Rings.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Rings.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -13,19 +13,6 @@
 imports Groups
 begin
 
-text {*
-  The theory of partially ordered rings is taken from the books:
-  \begin{itemize}
-  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
-  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
-  \end{itemize}
-  Most of the used notions can also be looked up in 
-  \begin{itemize}
-  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
-  \item \emph{Algebra I} by van der Waerden, Springer.
-  \end{itemize}
-*}
-
 class semiring = ab_semigroup_add + semigroup_mult +
   assumes left_distrib[algebra_simps]: "(a + b) * c = a * c + b * c"
   assumes right_distrib[algebra_simps]: "a * (b + c) = a * b + a * c"
@@ -506,6 +493,19 @@
   assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
   assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
 
+text {*
+  The theory of partially ordered rings is taken from the books:
+  \begin{itemize}
+  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
+  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
+  \end{itemize}
+  Most of the used notions can also be looked up in 
+  \begin{itemize}
+  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
+  \item \emph{Algebra I} by van der Waerden, Springer.
+  \end{itemize}
+*}
+
 class ordered_semiring = mult_mono + semiring_0 + ordered_ab_semigroup_add 
 begin
 
--- a/src/HOL/TLA/Action.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/TLA/Action.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -16,8 +16,8 @@
   'a trfun = "(state * state) => 'a"
   action   = "bool trfun"
 
-instance
-  "*" :: (world, world) world ..
+arities
+  "*" :: (world, world) world
 
 consts
   (** abstract syntax **)
--- a/src/HOL/TLA/Init.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/TLA/Init.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -12,7 +12,7 @@
 begin
 
 typedecl behavior
-instance behavior :: world ..
+arities behavior :: world
 
 types
   temporal = "behavior form"
--- a/src/HOL/TLA/Intensional.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/TLA/Intensional.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -10,8 +10,8 @@
 imports Main
 begin
 
-axclass
-  world < type
+classes world
+classrel world < type
 
 (** abstract syntax **)
 
@@ -171,7 +171,7 @@
   "_liftMem"    :: "[lift, lift] => lift"                ("(_/ \<in> _)" [50, 51] 50)
   "_liftNotMem" :: "[lift, lift] => lift"                ("(_/ \<notin> _)" [50, 51] 50)
 
-axioms
+defs
   Valid_def:   "|- A    ==  ALL w. w |= A"
 
   unl_con:     "LIFT #c w  ==  c"
--- a/src/HOL/TLA/Stfun.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/TLA/Stfun.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -11,7 +11,7 @@
 
 typedecl state
 
-instance state :: world ..
+arities state :: world
 
 types
   'a stfun = "state => 'a"
--- a/src/HOL/Tools/Datatype/datatype_prop.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Datatype/datatype_prop.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -22,6 +22,8 @@
     (string * sort) list -> theory -> term list list
   val make_splits : string list -> descr list ->
     (string * sort) list -> theory -> (term * term) list
+  val make_case_combs : string list -> descr list ->
+    (string * sort) list -> theory -> string -> term list
   val make_weak_case_congs : string list -> descr list ->
     (string * sort) list -> theory -> term list
   val make_case_congs : string list -> descr list ->
--- a/src/HOL/Tools/Function/function.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Function/function.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -97,6 +97,7 @@
           |> addsmps (conceal_partial o Binding.qualify false "partial")
                "psimps" conceal_partial psimp_attribs psimps
           ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps
+          ||> fold_option (Spec_Rules.add Spec_Rules.Equational o (pair fs)) trsimps
           ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"),
                  [Attrib.internal (K (Rule_Cases.case_names cnames)),
                   Attrib.internal (K (Rule_Cases.consumes 1)),
@@ -126,7 +127,7 @@
 val add_function =
   gen_add_function false Specification.check_spec (TypeInfer.anyT HOLogic.typeS)
 val add_function_cmd = gen_add_function true Specification.read_spec "_::type"
-
+                                                
 fun gen_termination_proof prep_term raw_term_opt lthy =
   let
     val term_opt = Option.map (prep_term lthy) raw_term_opt
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -1004,9 +1004,11 @@
   handle Option.Option => raise SYNTAX ("Kodkod.read_next_problems",
                                         "expected number after \"PROBLEM\"")
 
-(* Path.T -> (int * raw_bound list) list * int list *)
+(* Path.T -> bool * ((int * raw_bound list) list * int list) *)
 fun read_output_file path =
-  read_next_problems (Substring.full (File.read path), [], []) |>> rev ||> rev
+  (false, read_next_problems (Substring.full (File.read path), [], [])
+          |>> rev ||> rev)
+  handle IO.Io _ => (true, ([], [])) | OS.SysErr _ => (true, ([], []))
 
 (* The fudge term below is to account for Kodkodi's slow start-up time, which
    is partly due to the JVM and partly due to the ML "bash" function. *)
@@ -1046,7 +1048,7 @@
         (* unit -> unit *)
         fun remove_temporary_files () =
           if overlord then ()
-          else List.app File.rm [in_path, out_path, err_path]
+          else List.app (K () o try File.rm) [in_path, out_path, err_path]
       in
         let
           val ms =
@@ -1076,11 +1078,13 @@
                       " < " ^ File.shell_path in_path ^
                       " > " ^ File.shell_path out_path ^
                       " 2> " ^ File.shell_path err_path)
-              val (ps, nontriv_js) = read_output_file out_path
-                                     |>> map (apfst reindex) ||> map reindex
+              val (io_error, (ps, nontriv_js)) =
+                read_output_file out_path
+                ||> apfst (map (apfst reindex)) ||> apsnd (map reindex)
               val js = triv_js @ nontriv_js
               val first_error =
                 File.fold_lines (fn line => fn "" => line | s => s) err_path ""
+                handle IO.Io _ => "" | OS.SysErr _ => ""
             in
               if null ps then
                 if code = 2 then
@@ -1092,6 +1096,8 @@
                 else if first_error <> "" then
                   Error (first_error |> perhaps (try (unsuffix "."))
                                      |> perhaps (try (unprefix "Error: ")), js)
+                else if io_error then
+                  Error ("I/O error", js)
                 else if code <> 0 then
                   Error ("Unknown error", js)
                 else
@@ -1102,7 +1108,8 @@
         in remove_temporary_files (); outcome end
         handle Exn.Interrupt =>
                let
-                 val nontriv_js = map reindex (snd (read_output_file out_path))
+                 val nontriv_js =
+                   read_output_file out_path |> snd |> snd |> map reindex
                in
                  (remove_temporary_files ();
                   Interrupted (SOME (triv_js @ nontriv_js)))
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -54,6 +54,7 @@
   val numeral_prefix : string
   val ubfp_prefix : string
   val lbfp_prefix : string
+  val quot_normal_prefix : string
   val skolem_prefix : string
   val special_prefix : string
   val uncurry_prefix : string
@@ -173,7 +174,7 @@
   val inverse_axioms_for_rep_fun : theory -> styp -> term list
   val optimized_typedef_axioms : theory -> string * typ list -> term list
   val optimized_quot_type_axioms :
-    theory -> (typ option * bool) list -> string * typ list -> term list
+    Proof.context -> (typ option * bool) list -> string * typ list -> term list
   val def_of_const : theory -> const_table -> styp -> term option
   val fixpoint_kind_of_const :
     theory -> const_table -> string * typ -> fixpoint_kind
@@ -268,6 +269,7 @@
 val step_prefix = nitpick_prefix ^ "step" ^ name_sep
 val ubfp_prefix = nitpick_prefix ^ "ubfp" ^ name_sep
 val lbfp_prefix = nitpick_prefix ^ "lbfp" ^ name_sep
+val quot_normal_prefix = nitpick_prefix ^ "qn" ^ name_sep
 val skolem_prefix = nitpick_prefix ^ "sk"
 val special_prefix = nitpick_prefix ^ "sp"
 val uncurry_prefix = nitpick_prefix ^ "unc"
@@ -277,6 +279,9 @@
 
 (* int -> string *)
 fun sel_prefix_for j = sel_prefix ^ string_of_int j ^ name_sep
+(* Proof.context -> typ -> string *)
+fun quot_normal_name_for_type ctxt T =
+  quot_normal_prefix ^ unyxml (Syntax.string_of_typ ctxt T)
 
 (* string -> string * string *)
 val strip_first_name_sep =
@@ -559,14 +564,15 @@
 (* theory -> typ -> typ -> typ -> typ *)
 fun instantiate_type thy T1 T1' T2 =
   Same.commit (Envir.subst_type_same
-                   (Sign.typ_match thy (Logic.varifyT T1, T1') Vartab.empty))
-              (Logic.varifyT T2)
+                   (Sign.typ_match thy (T1, T1') Vartab.empty)) T2
   handle Type.TYPE_MATCH =>
          raise TYPE ("Nitpick_HOL.instantiate_type", [T1, T1'], [])
+fun varify_and_instantiate_type thy T1 T1' T2 =
+  instantiate_type thy (Logic.varifyT T1) T1' (Logic.varifyT T2)
 
 (* theory -> typ -> typ -> styp *)
 fun repair_constr_type thy body_T' T =
-  instantiate_type thy (body_type T) body_T' T
+  varify_and_instantiate_type thy (body_type T) body_T' T
 
 (* string -> (string * string) list -> theory -> theory *)
 fun register_frac_type frac_s ersaetze thy =
@@ -889,7 +895,8 @@
              [(@{const_name Quot}, rep_type_for_quot_type thy T --> T)]
            else case typedef_info thy s of
              SOME {abs_type, rep_type, Abs_name, ...} =>
-             [(Abs_name, instantiate_type thy abs_type T rep_type --> T)]
+             [(Abs_name,
+               varify_and_instantiate_type thy abs_type T rep_type --> T)]
            | NONE =>
              if T = @{typ ind} then
                [dest_Const @{const Zero_Rep}, dest_Const @{const Suc_Rep}]
@@ -1385,7 +1392,7 @@
     else case typedef_info thy abs_s of
       SOME {abs_type, rep_type, Rep_name, prop_of_Rep, set_name, ...} =>
       let
-        val rep_T = instantiate_type thy abs_type abs_T rep_type
+        val rep_T = varify_and_instantiate_type thy abs_type abs_T rep_type
         val rep_t = Const (Rep_name, abs_T --> rep_T)
         val set_t = Const (set_name, rep_T --> bool_T)
         val set_t' =
@@ -1400,8 +1407,10 @@
       end
     | NONE => []
   end
-fun optimized_quot_type_axioms thy stds abs_z =
+(* Proof.context -> string * typ list -> term list *)
+fun optimized_quot_type_axioms ctxt stds abs_z =
   let
+    val thy = ProofContext.theory_of ctxt
     val abs_T = Type abs_z
     val rep_T = rep_type_for_quot_type thy abs_T
     val equiv_rel = equiv_relation_for_quot_type thy abs_T
@@ -1410,7 +1419,7 @@
     val y_var = Var (("y", 0), rep_T)
     val x = (@{const_name Quot}, rep_T --> abs_T)
     val sel_a_t = select_nth_constr_arg thy stds x a_var 0 rep_T
-    val normal_t = Const (@{const_name quot_normal}, rep_T --> rep_T)
+    val normal_t = Const (quot_normal_name_for_type ctxt abs_T, rep_T --> rep_T)
     val normal_x = normal_t $ x_var
     val normal_y = normal_t $ y_var
     val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
@@ -1639,7 +1648,7 @@
                 in
                   (Abs (Name.uu, rep_T,
                         Const (@{const_name Quot}, rep_T --> abs_T)
-                               $ (Const (@{const_name quot_normal},
+                               $ (Const (quot_normal_name_for_type ctxt abs_T,
                                          rep_T --> rep_T) $ Bound 0)), ts)
                 end
               else if is_quot_rep_fun ctxt x then
@@ -1819,8 +1828,7 @@
                              termination_tacs
            in Synchronized.change cached_wf_props (cons (prop, wf)); wf end
        end)
-    handle List.Empty => false
-         | NO_TRIPLE () => false
+    handle List.Empty => false | NO_TRIPLE () => false
 
 (* The type constraint below is a workaround for a Poly/ML crash. *)
 
@@ -2231,6 +2239,10 @@
        else if String.isPrefix step_prefix s then
          (Const (step_name, T --> T) $ Const (unprefix step_prefix s, T),
           format_type default_format default_format T)
+       else if String.isPrefix quot_normal_prefix s then
+         let val t = Const (nitpick_prefix ^ "normalize quotient type", T) in
+           (t, format_term_type thy def_table formats t)
+         end
        else if String.isPrefix skolem_prefix s then
          let
            val ss = the (AList.lookup (op =) (!skolems) s)
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -1595,12 +1595,7 @@
           KK.Atom (offset_of_type ofs nat_T)
         else
           fold kk_join (map to_integer [u1, u2]) (KK.Rel nat_subtract_rel)
-      | Op2 (Apply, _, R, u1, u2) =>
-        if is_Cst Unrep u2 andalso is_set_type (type_of u1) andalso
-           is_FreeName u1 then
-          false_atom
-        else
-          to_apply R u1 u2
+      | Op2 (Apply, _, R, u1, u2) => to_apply R u1 u2
       | Op2 (Lambda, _, R as Opt (Atom (1, j0)), u1, u2) =>
         to_guard [u1, u2] R (KK.Atom j0)
       | Op2 (Lambda, _, Func (_, Formula Neut), u1, u2) =>
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -95,7 +95,6 @@
   val nickname_of : nut -> string
   val is_skolem_name : nut -> bool
   val is_eval_name : nut -> bool
-  val is_FreeName : nut -> bool
   val is_Cst : cst -> nut -> bool
   val fold_nut : (nut -> 'a -> 'a) -> nut -> 'a -> 'a
   val map_nut : (nut -> nut) -> nut -> nut
@@ -369,8 +368,6 @@
 fun is_eval_name u =
   String.isPrefix eval_prefix (nickname_of u)
   handle NUT ("Nitpick_Nut.nickname_of", _) => false
-fun is_FreeName (FreeName _) = true
-  | is_FreeName _ = false
 (* cst -> nut -> bool *)
 fun is_Cst cst (Cst (cst', _, _)) = (cst = cst')
   | is_Cst _ _ = false
@@ -794,9 +791,9 @@
   end
 
 (* A nut is said to be constructive if whenever it evaluates to unknown in our
-   three-valued logic, it would evaluate to a unrepresentable value ("unrep")
+   three-valued logic, it would evaluate to a unrepresentable value ("Unrep")
    according to the HOL semantics. For example, "Suc n" is constructive if "n"
-   is representable or "Unrep", because unknown implies unrep. *)
+   is representable or "Unrep", because unknown implies "Unrep". *)
 (* nut -> bool *)
 fun is_constructive u =
   is_Cst Unrep u orelse
@@ -819,6 +816,16 @@
 fun unknown_boolean T R =
   Cst (case R of Formula Pos => False | Formula Neg => True | _ => Unknown,
        T, R)
+(* nut -> bool *)
+fun is_fully_representable_set u =
+  not (is_opt_rep (rep_of u)) andalso
+  case u of
+    FreeName _ => true
+  | Op1 (SingletonSet, _, _, _) => true
+  | Op2 (oper, _, _, u1, u2) =>
+    member (op =) [Union, SetDifference, Intersect] oper andalso
+    forall is_fully_representable_set [u1, u2]
+  | _ => false
 
 (* op1 -> typ -> rep -> nut -> nut *)
 fun s_op1 oper T R u1 =
@@ -860,7 +867,7 @@
         if is_constructive u1 then
           Cst (Unrep, T, R)
         else if is_boolean_type T then
-          if is_FreeName u1 then Cst (False, T, Formula Neut)
+          if is_fully_representable_set u1 then Cst (False, T, Formula Neut)
           else unknown_boolean T R
         else case u1 of
           Op2 (Apply, _, _, ConstName (@{const_name List.append}, _, _), _) =>
--- a/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Nitpick/nitpick_preproc.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -293,15 +293,15 @@
         $ do_term new_Ts old_Ts polar t2
       | Const (s as @{const_name The}, T) => do_description_operator s T
       | Const (s as @{const_name Eps}, T) => do_description_operator s T
-      | Const (@{const_name quot_normal}, Type ("fun", [T, _])) =>
-        let val T' = box_type hol_ctxt InFunLHS T in
-          Const (@{const_name quot_normal}, T' --> T')
-        end
       | Const (s as @{const_name Tha}, T) => do_description_operator s T
       | Const (x as (s, T)) =>
         Const (s, if s = @{const_name converse} orelse
                      s = @{const_name trancl} then
                     box_relational_operator_type T
+                  else if String.isPrefix quot_normal_prefix s then
+                    let val T' = box_type hol_ctxt InFunLHS (domain_type T) in
+                      T' --> T'
+                    end
                   else if is_built_in_const thy stds fast_descrs x orelse
                           s = @{const_name Sigma} then
                     T
@@ -1022,8 +1022,9 @@
 
 (* hol_context -> term -> (term list * term list) * (bool * bool) *)
 fun axioms_for_term
-        (hol_ctxt as {thy, max_bisim_depth, stds, user_axioms, fast_descrs,
-                      evals, def_table, nondef_table, user_nondefs, ...}) t =
+        (hol_ctxt as {thy, ctxt, max_bisim_depth, stds, user_axioms,
+                      fast_descrs, evals, def_table, nondef_table, user_nondefs,
+                      ...}) t =
   let
     type accumulator = styp list * (term list * term list)
     (* (term list * term list -> term list)
@@ -1090,7 +1091,8 @@
                else
                  accum |> fold (add_nondef_axiom depth)
                                (nondef_props_for_const thy false nondef_table x)
-                       |> is_funky_typedef thy (range_type T)
+                       |> (is_funky_typedef thy (range_type T) orelse
+                           range_type T = nat_T)
                           ? fold (add_maybe_def_axiom depth)
                                  (nondef_props_for_const thy true
                                                     (extra_table def_table s) x)
@@ -1100,7 +1102,8 @@
                else
                  accum |> fold (add_nondef_axiom depth)
                                (nondef_props_for_const thy false nondef_table x)
-                       |> is_funky_typedef thy (range_type T)
+                       |> (is_funky_typedef thy (range_type T) orelse
+                           range_type T = nat_T)
                           ? fold (add_maybe_def_axiom depth)
                                  (nondef_props_for_const thy true
                                                     (extra_table def_table s) x)
@@ -1134,7 +1137,8 @@
         #> (if is_pure_typedef thy T then
               fold (add_maybe_def_axiom depth) (optimized_typedef_axioms thy z)
             else if is_quot_type thy T then
-              fold (add_def_axiom depth) (optimized_quot_type_axioms thy stds z)
+              fold (add_def_axiom depth)
+                   (optimized_quot_type_axioms ctxt stds z)
             else if max_bisim_depth >= 0 andalso is_codatatype thy T then
               fold (add_maybe_def_axiom depth)
                    (codatatype_bisim_axioms hol_ctxt T)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -8,13 +8,13 @@
 sig
   val setup : theory -> theory
   val preprocess : Predicate_Compile_Aux.options -> string -> theory -> theory
+  val present_graph : bool Unsynchronized.ref
 end;
 
 structure Predicate_Compile (*: PREDICATE_COMPILE*) =
 struct
 
-(* options *)
-val fail_safe_mode = true
+val present_graph = Unsynchronized.ref false
 
 open Predicate_Compile_Aux;
 
@@ -33,10 +33,12 @@
            commas (map (Display.string_of_thm_global thy) intros)) intross)))
   else ()
       
-fun print_specs thy specs =
-  map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
-    ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
-
+fun print_specs options thy specs =
+  if show_intermediate_results options then
+    map (fn (c, thms) => "Constant " ^ c ^ " has specification:\n"
+      ^ (space_implode "\n" (map (Display.string_of_thm_global thy) thms)) ^ "\n") specs
+    |> space_implode "\n" |> tracing
+  else ()
 fun overload_const thy s = the_default s (Option.map fst (AxClass.inst_of_param thy s))
 
 fun map_specs f specs =
@@ -47,15 +49,21 @@
     val _ = print_step options "Compiling predicates to flat introrules..."
     val specs = map (apsnd (map
       (fn th => if is_equationlike th then Predicate_Compile_Data.normalize_equation thy' th else th))) specs
-    val (intross1, thy'') = apfst flat (fold_map Predicate_Compile_Pred.preprocess specs thy')
+    val (intross1, thy'') =
+      apfst flat (fold_map (Predicate_Compile_Pred.preprocess options) specs thy')
     val _ = print_intross options thy'' "Flattened introduction rules: " intross1
     val _ = print_step options "Replacing functions in introrules..."
     val intross2 =
-      if fail_safe_mode then
-        case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
-          SOME intross => intross
-        | NONE => let val _ = warning "Function replacement failed!" in intross1 end
-      else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
+      if function_flattening options then
+        if fail_safe_function_flattening options then
+          case try (map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy''))) intross1 of
+            SOME intross => intross
+          | NONE =>
+            (if show_caught_failures options then tracing "Function replacement failed!" else ();
+            intross1)
+        else map_specs (maps (Predicate_Compile_Fun.rewrite_intro thy'')) intross1
+      else
+        intross1
     val _ = print_intross options thy'' "Introduction rules with replaced functions: " intross2
     val _ = print_step options "Introducing new constants for abstractions at higher-order argument positions..."
     val (intross3, (new_defs, thy''')) = Predicate_Compile_Pred.flat_higher_order_arguments (intross2, thy'')
@@ -70,35 +78,45 @@
   end
 
 fun preprocess_strong_conn_constnames options gr ts thy =
-  let
-    fun get_specs ts = map_filter (fn t =>
-      TermGraph.get_node gr t |>
-      (fn ths => if null ths then NONE else SOME (fst (dest_Const t), ths)))
-      ts
-    val _ = print_step options ("Preprocessing scc of " ^
-      commas (map (Syntax.string_of_term_global thy) ts))
-    val (prednames, funnames) = List.partition (fn t => body_type (fastype_of t) = @{typ bool}) ts
-    (* untangle recursion by defining predicates for all functions *)
-    val _ = print_step options
-      ("Compiling functions (" ^ commas (map (Syntax.string_of_term_global thy) funnames) ^
-        ") to predicates...")
-    val (fun_pred_specs, thy') =
-      if not (null funnames) then Predicate_Compile_Fun.define_predicates
-      (get_specs funnames) thy else ([], thy)
-    val _ = print_specs thy' fun_pred_specs
-    val specs = (get_specs prednames) @ fun_pred_specs
-    val (intross3, thy''') = process_specification options specs thy'
-    val _ = print_intross options thy''' "Introduction rules with new constants: " intross3
-    val intross4 = map_specs (maps remove_pointless_clauses) intross3
-    val _ = print_intross options thy''' "After removing pointless clauses: " intross4
-    val intross5 = map (fn (s, ths) => (overload_const thy''' s, map (AxClass.overload thy''') ths)) intross4
-    val intross6 = map_specs (map (expand_tuples thy''')) intross5
-    val _ = print_intross options thy''' "introduction rules before registering: " intross6
-    val _ = print_step options "Registering introduction rules..."
-    val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
-  in
-    thy''''
-  end;
+  if forall (fn (Const (c, _)) => Predicate_Compile_Core.is_registered thy c) ts then
+    thy
+  else
+    let
+      fun get_specs ts = map_filter (fn t =>
+        TermGraph.get_node gr t |>
+        (fn ths => if null ths then NONE else SOME (fst (dest_Const t), ths)))
+        ts
+      val _ = print_step options ("Preprocessing scc of " ^
+        commas (map (Syntax.string_of_term_global thy) ts))
+      val (prednames, funnames) = List.partition (fn t => body_type (fastype_of t) = @{typ bool}) ts
+      (* untangle recursion by defining predicates for all functions *)
+      val _ = print_step options
+        ("Compiling functions (" ^ commas (map (Syntax.string_of_term_global thy) funnames) ^
+          ") to predicates...")
+      val (fun_pred_specs, thy') =
+        (if function_flattening options andalso (not (null funnames)) then
+          if fail_safe_function_flattening options then
+            case try (Predicate_Compile_Fun.define_predicates (get_specs funnames)) thy of
+              SOME (intross, thy) => (intross, thy)
+            | NONE => ([], thy)
+          else Predicate_Compile_Fun.define_predicates (get_specs funnames) thy
+        else ([], thy))
+        (*||> Theory.checkpoint*)
+      val _ = print_specs options thy' fun_pred_specs
+      val specs = (get_specs prednames) @ fun_pred_specs
+      val (intross3, thy''') = process_specification options specs thy'
+      val _ = print_intross options thy''' "Introduction rules with new constants: " intross3
+      val intross4 = map_specs (maps remove_pointless_clauses) intross3
+      val _ = print_intross options thy''' "After removing pointless clauses: " intross4
+      val intross5 =
+        map (fn (s, ths) => (overload_const thy''' s, map (AxClass.overload thy''') ths)) intross4
+      val intross6 = map_specs (map (expand_tuples thy''')) intross5
+      val _ = print_intross options thy''' "introduction rules before registering: " intross6
+      val _ = print_step options "Registering introduction rules..."
+      val thy'''' = fold Predicate_Compile_Core.register_intros intross6 thy'''
+    in
+      thy''''
+    end;
 
 fun preprocess options t thy =
   let
@@ -106,6 +124,7 @@
     val gr = Output.cond_timeit (!Quickcheck.timing) "preprocess-obtain graph"
           (fn () => Predicate_Compile_Data.obtain_specification_graph options thy t
           |> (fn gr => TermGraph.subgraph (member (op =) (TermGraph.all_succs gr [t])) gr))
+    val _ = if !present_graph then Predicate_Compile_Data.present_graph gr else ()
   in
     Output.cond_timeit (!Quickcheck.timing) "preprocess-process"
       (fn () => (fold_rev (preprocess_strong_conn_constnames options gr)
@@ -128,7 +147,12 @@
       show_modes = chk "show_modes",
       show_mode_inference = chk "show_mode_inference",
       show_compilation = chk "show_compilation",
+      show_caught_failures = false,
       skip_proof = chk "skip_proof",
+      function_flattening = not (chk "no_function_flattening"),
+      fail_safe_function_flattening = false,
+      no_topmost_reordering = false,
+      no_higher_order_predicate = [],
       inductify = chk "inductify",
       compilation = compilation
     }
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -29,6 +29,7 @@
 fun find_indices f xs =
   map_filter (fn (i, true) => SOME i | (i, false) => NONE) (map_index (apsnd f) xs)
 
+fun assert check = if check then () else error "Assertion failed!"
 (* mode *)
 
 datatype mode = Bool | Input | Output | Pair of mode * mode | Fun of mode * mode
@@ -57,21 +58,47 @@
 fun dest_tuple_mode (Pair (mode, mode')) = mode :: dest_tuple_mode mode'
   | dest_tuple_mode _ = []
 
-fun all_modes_of_typ (T as Type ("fun", _)) = 
+
+fun all_modes_of_typ' (T as Type ("fun", _)) = 
+  let
+    val (S, U) = strip_type T
+  in
+    if U = HOLogic.boolT then
+      fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2)
+        (map all_modes_of_typ' S) [Bool]
+    else
+      [Input, Output]
+  end
+  | all_modes_of_typ' (Type ("*", [T1, T2])) = 
+    map_product (curry Pair) (all_modes_of_typ' T1) (all_modes_of_typ' T2)
+  | all_modes_of_typ' _ = [Input, Output]
+
+fun all_modes_of_typ (T as Type ("fun", _)) =
   let
     val (S, U) = strip_type T
   in
     if U = HOLogic.boolT then
       fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2)
-        (map all_modes_of_typ S) [Bool]
+        (map all_modes_of_typ' S) [Bool]
     else
       [Input, Output]
   end
-  | all_modes_of_typ (Type ("*", [T1, T2])) = 
-    map_product (curry Pair) (all_modes_of_typ T1) (all_modes_of_typ T2)
   | all_modes_of_typ (Type ("bool", [])) = [Bool]
-  | all_modes_of_typ _ = [Input, Output]
+  | all_modes_of_typ T = all_modes_of_typ' T
 
+fun all_smodes_of_typ (T as Type ("fun", _)) =
+  let
+    val (S, U) = strip_type T
+    fun all_smodes (Type ("*", [T1, T2])) = 
+      map_product (curry Pair) (all_smodes T1) (all_smodes T2)
+      | all_smodes _ = [Input, Output]
+  in
+    if U = HOLogic.boolT then
+      fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_smodes S) [Bool]
+    else
+      error "all_smodes_of_typ: invalid type for predicate"
+  end
+(*
 fun extract_params arg =
   case fastype_of arg of
     (T as Type ("fun", _)) =>
@@ -86,7 +113,7 @@
         extract_params t1 @ extract_params t2
       end
   | _ => []
-
+*)
 fun ho_arg_modes_of mode =
   let
     fun ho_arg_mode (m as Fun _) =  [m]
@@ -144,9 +171,10 @@
         in
           (comb_option HOLogic.mk_prod (i1, i2), comb_option HOLogic.mk_prod (o1, o2))
         end
-      | split_arg_mode' Input t = (SOME t, NONE)
-      | split_arg_mode' Output t = (NONE,  SOME t)
-      | split_arg_mode' _ _ = error "split_map_mode: mode and term do not match"
+      | split_arg_mode' m t =
+        if eq_mode (m, Input) then (SOME t, NONE)
+        else if eq_mode (m, Output) then (NONE,  SOME t)
+        else error "split_map_mode: mode and term do not match"
   in
     (pairself (map_filter I) o split_list) (map2 split_arg_mode' (strip_fun_mode mode) ts)
   end
@@ -269,7 +297,6 @@
   let
     val T = (Sign.the_const_type thy constname)
   in body_type T = @{typ "bool"} end;
-  
 
 fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
   | is_predT _ = false
@@ -373,7 +400,60 @@
   in
     Logic.list_implies (maps f premises, head)
   end
+
+
+(* split theorems of case expressions *)
+
+(*
+fun has_split_rule_cname @{const_name "nat_case"} = true
+  | has_split_rule_cname @{const_name "list_case"} = true
+  | has_split_rule_cname _ = false
   
+fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true 
+  | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true 
+  | has_split_rule_term thy _ = false
+
+fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
+  | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
+  | has_split_rule_term' thy c = has_split_rule_term thy c
+
+*)
+fun prepare_split_thm ctxt split_thm =
+    (split_thm RS @{thm iffD2})
+    |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
+      @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
+
+fun find_split_thm thy (Const (name, typ)) =
+  let
+    fun split_name str =
+      case first_field "." str
+        of (SOME (field, rest)) => field :: split_name rest
+         | NONE => [str]
+    val splitted_name = split_name name
+  in
+    if length splitted_name > 0 andalso
+       String.isSuffix "_case" (List.last splitted_name)
+    then
+      (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
+      |> space_implode "."
+      |> PureThy.get_thm thy
+      |> SOME
+      handle ERROR msg => NONE
+    else NONE
+  end
+  | find_split_thm _ _ = NONE
+
+
+(* TODO: split rules for let and if are different *)
+fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
+  | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
+  | find_split_thm' thy c = find_split_thm thy c
+
+fun has_split_thm thy t = is_some (find_split_thm thy t)
+
+fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
+
+
 (* lifting term operations to theorems *)
 
 fun map_term thy f th =
@@ -388,7 +468,16 @@
 
 (* Different options for compiler *)
 
-datatype compilation = Pred | Random | Depth_Limited | DSeq | Annotated | Random_DSeq
+datatype compilation = Pred | Random | Depth_Limited | DSeq | Annotated
+  | Pos_Random_DSeq | Neg_Random_DSeq
+
+
+fun negative_compilation_of Pos_Random_DSeq = Neg_Random_DSeq
+  | negative_compilation_of Neg_Random_DSeq = Pos_Random_DSeq
+  | negative_compilation_of c = c
+  
+fun compilation_for_polarity false Pos_Random_DSeq = Neg_Random_DSeq
+  | compilation_for_polarity _ c = c
 
 fun string_of_compilation c = case c of
     Pred => ""
@@ -396,8 +485,9 @@
   | Depth_Limited => "depth limited"
   | DSeq => "dseq"
   | Annotated => "annotated"
-  | Random_DSeq => "random dseq"
-
+  | Pos_Random_DSeq => "pos_random dseq"
+  | Neg_Random_DSeq => "neg_random_dseq"
+  
 (*datatype compilation_options =
   Pred | Random of int | Depth_Limited of int | DSeq of int | Annotated*)
 
@@ -411,8 +501,12 @@
   show_mode_inference : bool,
   show_modes : bool,
   show_compilation : bool,
+  show_caught_failures : bool,
   skip_proof : bool,
-
+  no_topmost_reordering : bool,
+  function_flattening : bool,
+  fail_safe_function_flattening : bool,
+  no_higher_order_predicate : string list,
   inductify : bool,
   compilation : compilation
 };
@@ -428,8 +522,15 @@
 fun show_modes (Options opt) = #show_modes opt
 fun show_mode_inference (Options opt) = #show_mode_inference opt
 fun show_compilation (Options opt) = #show_compilation opt
+fun show_caught_failures (Options opt) = #show_caught_failures opt
+
 fun skip_proof (Options opt) = #skip_proof opt
 
+fun function_flattening (Options opt) = #function_flattening opt
+fun fail_safe_function_flattening (Options opt) = #fail_safe_function_flattening opt
+fun no_topmost_reordering (Options opt) = #no_topmost_reordering opt
+fun no_higher_order_predicate (Options opt) = #no_higher_order_predicate opt
+
 fun is_inductify (Options opt) = #inductify opt
 
 fun compilation (Options opt) = #compilation opt
@@ -444,18 +545,22 @@
   show_modes = false,
   show_mode_inference = false,
   show_compilation = false,
+  show_caught_failures = false,
   skip_proof = true,
-  
+  no_topmost_reordering = false,
+  function_flattening = false,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
   inductify = false,
   compilation = Pred
 }
 
 val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes",
-  "show_mode_inference", "show_compilation", "skip_proof", "inductify"]
+  "show_mode_inference", "show_compilation", "skip_proof", "inductify", "no_function_flattening"]
 
 val compilation_names = [("pred", Pred),
   (*("random", Random), ("depth_limited", Depth_Limited), ("annotated", Annotated),*)
-  ("dseq", DSeq), ("random_dseq", Random_DSeq)]
+  ("dseq", DSeq), ("random_dseq", Pos_Random_DSeq)]
 
 fun print_step options s =
   if show_steps options then tracing s else ()
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -16,7 +16,7 @@
   val register_intros : string * thm list -> theory -> theory
   val is_registered : theory -> string -> bool
   val function_name_of : Predicate_Compile_Aux.compilation -> theory
-    -> string -> Predicate_Compile_Aux.mode -> string
+    -> string -> bool * Predicate_Compile_Aux.mode -> string
   val predfun_intro_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
   val predfun_elim_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
   val all_preds_of : theory -> string list
@@ -153,7 +153,7 @@
   | mode_of (Term m) = m
   | mode_of (Mode_App (d1, d2)) =
     (case mode_of d1 of Fun (m, m') =>
-        (if m = mode_of d2 then m' else error "mode_of")
+        (if eq_mode (m, mode_of d2) then m' else error "mode_of")
       | _ => error "mode_of2")
   | mode_of (Mode_Pair (d1, d2)) =
     Pair (mode_of d1, mode_of d2)
@@ -182,7 +182,7 @@
 
 type moded_clause = term list * (indprem * mode_derivation) list
 
-type 'a pred_mode_table = (string * (mode * 'a) list) list
+type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
 
 (* book-keeping *)
 
@@ -257,8 +257,9 @@
       ^ "functions defined for predicate " ^ quote name)
   | SOME fun_names => fun_names
 
-fun function_name_of compilation thy name mode =
-  case AList.lookup (op =) (function_names_of compilation thy name) mode of
+fun function_name_of compilation thy name (pol, mode) =
+  case AList.lookup eq_mode
+    (function_names_of (compilation_for_polarity pol compilation) thy name) mode of
     NONE => error ("No " ^ string_of_compilation compilation
       ^ "function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name)
   | SOME function_name => function_name
@@ -296,12 +297,12 @@
   if show_modes options then
     tracing ("Inferred modes:\n" ^
       cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
-        string_of_mode ms)) modes))
+        (fn (p, m) => string_of_mode m ^ (if p then "pos" else "neg")) ms)) modes))
   else ()
 
 fun print_pred_mode_table string_of_entry thy pred_mode_table =
   let
-    fun print_mode pred (mode, entry) =  "mode : " ^ string_of_mode mode
+    fun print_mode pred ((pol, mode), entry) =  "mode : " ^ string_of_mode mode
       ^ string_of_entry pred mode entry
     fun print_pred (pred, modes) =
       "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
@@ -364,7 +365,7 @@
     SOME (s, ms) => (case AList.lookup (op =) modes s of
       SOME modes =>
         let
-          val modes' = modes
+          val modes' = map snd modes
         in
           if not (eq_set eq_mode (ms, modes')) then
             error ("expected modes were not inferred:\n"
@@ -381,7 +382,7 @@
       SOME inferred_ms =>
         let
           val preds_without_modes = map fst (filter (null o snd) (modes @ extra_modes))
-          val modes' = inferred_ms
+          val modes' = map snd inferred_ms
         in
           if not (eq_set eq_mode (ms, modes')) then
             error ("expected modes were not inferred:\n"
@@ -880,8 +881,6 @@
       Random_Sequence_CompFuns.mk_random_dseqT T) $ random
   end;
 
-
-
 (* for external use with interactive mode *)
 val pred_compfuns = PredicateCompFuns.compfuns
 val randompred_compfuns = Random_Sequence_CompFuns.compfuns;
@@ -898,6 +897,8 @@
 
 (** mode analysis **)
 
+(* options for mode analysis  are: #use_random, #reorder_premises *)
+
 fun is_constrt thy =
   let
     val cnstrs = flat (maps
@@ -935,7 +936,7 @@
     in merge (map (fn ks => i::ks) is) is end
   else [[]];
 
-fun print_failed_mode options thy modes p m rs is =
+fun print_failed_mode options thy modes p (pol, m) rs is =
   if show_mode_inference options then
     let
       val _ = tracing ("Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
@@ -943,7 +944,7 @@
     in () end
   else ()
 
-fun error_of p m is =
+fun error_of p (pol, m) is =
   ("  Clauses " ^ commas (map (fn i => string_of_int (i + 1)) is) ^ " of " ^
         p ^ " violates mode " ^ string_of_mode m)
 
@@ -992,7 +993,7 @@
 fun is_invertible_function thy (Const (f, _)) = is_constr thy f
   | is_invertible_function thy _ = false
 
-fun non_invertible_subterms thy (Free _) = []
+fun non_invertible_subterms thy (t as Free _) = []
   | non_invertible_subterms thy t = 
   case (strip_comb t) of (f, args) =>
     if is_invertible_function thy f then
@@ -1029,6 +1030,9 @@
   forall
     (fn t => is_eqT (fastype_of t) andalso forall (member (op =) vs) (term_vs t))
       (non_invertible_subterms thy t)
+  andalso
+    (forall (is_eqT o snd)
+      (inter (fn ((f', _), f) => f = f') vs (Term.add_frees t [])))
 
 fun vars_of_destructable_term thy (Free (x, _)) = [x]
   | vars_of_destructable_term thy t =
@@ -1042,18 +1046,52 @@
 
 fun missing_vars vs t = subtract (op =) vs (term_vs t)
 
-fun derivations_of thy modes vs t Input = 
-    [(Term Input, missing_vars vs t)]
-  | derivations_of thy modes vs t Output =
-    if is_possible_output thy vs t then [(Term Output, [])] else []
-  | derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
+fun output_terms (Const ("Pair", _) $ t1 $ t2, Mode_Pair (d1, d2)) =
+    output_terms (t1, d1)  @ output_terms (t2, d2)
+  | output_terms (t1 $ t2, Mode_App (d1, d2)) =
+    output_terms (t1, d1)  @ output_terms (t2, d2)
+  | output_terms (t, Term Output) = [t]
+  | output_terms _ = []
+
+fun lookup_mode modes (Const (s, T)) =
+   (case (AList.lookup (op =) modes s) of
+      SOME ms => SOME (map (fn m => (Context m, [])) ms)
+    | NONE => NONE)
+  | lookup_mode modes (Free (x, _)) =
+    (case (AList.lookup (op =) modes x) of
+      SOME ms => SOME (map (fn m => (Context m , [])) ms)
+    | NONE => NONE)
+
+fun derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) (Pair (m1, m2)) =
     map_product
       (fn (m1, mvars1) => fn (m2, mvars2) => (Mode_Pair (m1, m2), union (op =) mvars1 mvars2))
         (derivations_of thy modes vs t1 m1) (derivations_of thy modes vs t2 m2)
+  | derivations_of thy modes vs t (m as Fun _) =
+    (*let
+      val (p, args) = strip_comb t
+    in
+      (case lookup_mode modes p of
+        SOME ms => map_filter (fn (Context m, []) => let
+          val ms = strip_fun_mode m
+          val (argms, restms) = chop (length args) ms
+          val m' = fold_rev (curry Fun) restms Bool
+        in
+          if forall (fn m => eq_mode (Input, m)) argms andalso eq_mode (m', mode) then
+            SOME (fold (curry Mode_App) (map Term argms) (Context m), missing_vars vs t)
+          else NONE
+        end) ms
+      | NONE => (if is_all_input mode then [(Context mode, [])] else []))
+    end*)
+    (case try (all_derivations_of thy modes vs) t  of
+      SOME derivs =>
+        filter (fn (d, mvars) => eq_mode (mode_of d, m) andalso null (output_terms (t, d))) derivs
+    | NONE => (if is_all_input m then [(Context m, [])] else []))
   | derivations_of thy modes vs t m =
-    (case try (all_derivations_of thy modes vs) t of
-      SOME derivs => filter (fn (d, mvars) => mode_of d = m) derivs
-    | NONE => (if is_all_input m then [(Context m, [])] else []))
+    if eq_mode (m, Input) then
+      [(Term Input, missing_vars vs t)]
+    else if eq_mode (m, Output) then
+      (if is_possible_output thy vs t then [(Term Output, [])] else [])
+    else []
 and all_derivations_of thy modes vs (Const ("Pair", _) $ t1 $ t2) =
   let
     val derivs1 = all_derivations_of thy modes vs t1
@@ -1073,14 +1111,8 @@
           (Mode_App (d1, d2), union (op =) mvars1 mvars2)) (derivations_of thy modes vs t2 m')
         | _ => error "Something went wrong") derivs1
   end
-  | all_derivations_of thy modes vs (Const (s, T)) =
-    (case (AList.lookup (op =) modes s) of
-      SOME ms => map (fn m => (Context m, [])) ms
-    | NONE => error ("No mode for constant " ^ s))
-  | all_derivations_of _ modes vs (Free (x, _)) =
-    (case (AList.lookup (op =) modes x) of
-      SOME ms => map (fn m => (Context m , [])) ms
-    | NONE => error ("No mode for parameter variable " ^ x))
+  | all_derivations_of thy modes vs (Const (s, T)) = the (lookup_mode modes (Const (s, T)))
+  | all_derivations_of thy modes vs (Free (x, T)) = the (lookup_mode modes (Free (x, T)))
   | all_derivations_of _ modes vs _ = error "all_derivations_of"
 
 fun rev_option_ord ord (NONE, NONE) = EQUAL
@@ -1097,7 +1129,7 @@
     SOME (s, _) =>
       (case AList.lookup (op =) modes s of
         SOME ms =>
-          (case AList.lookup (op =) ms (head_mode_of deriv) of
+          (case AList.lookup (op =) (map (fn ((p, m), r) => (m, r)) ms) (head_mode_of deriv) of
             SOME r => r
           | NONE => false)
       | NONE => false)
@@ -1146,51 +1178,56 @@
   tracing ("modes: " ^ (commas (map (fn (s, ms) => s ^ ": " ^
     commas (map (fn (m, r) => string_of_mode m ^ (if r then " random " else " not ")) ms)) modes)))
 
-fun select_mode_prem' thy modes vs ps =
+fun select_mode_prem mode_analysis_options thy pol (modes, (pos_modes, neg_modes)) vs ps =
   let
-    val modes' = map (fn (s, ms) => (s, map fst ms)) modes
+    fun choose_mode_of_prem (Prem t) = partial_hd
+        (sort (deriv_ord2 thy modes t) (all_derivations_of thy pos_modes vs t))
+      | choose_mode_of_prem (Sidecond t) = SOME (Context Bool, missing_vars vs t)
+      | choose_mode_of_prem (Negprem t) = partial_hd
+          (sort (deriv_ord2 thy modes t) (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
+             (all_derivations_of thy neg_modes vs t)))
+      | choose_mode_of_prem p = error ("choose_mode_of_prem: " ^ string_of_prem thy p)
   in
-    partial_hd (sort (premise_ord thy modes) (ps ~~ map
-    (fn Prem t =>
-      partial_hd
-        (sort (deriv_ord2 thy modes t) (all_derivations_of thy modes' vs t))
-     | Sidecond t => SOME (Context Bool, missing_vars vs t)
-     | Negprem t =>
-         partial_hd
-          (sort (deriv_ord2 thy modes t) (filter (fn (d, missing_vars) => is_all_input (head_mode_of d))
-             (all_derivations_of thy modes' vs t)))
-     | p => error (string_of_prem thy p))
-    ps))
+    if #reorder_premises mode_analysis_options then
+      partial_hd (sort (premise_ord thy modes) (ps ~~ map choose_mode_of_prem ps))
+    else
+      SOME (hd ps, choose_mode_of_prem (hd ps))
   end
 
-fun check_mode_clause' use_random thy param_vs modes mode (ts, ps) =
+fun check_mode_clause' mode_analysis_options thy param_vs (modes :
+  (string * ((bool * mode) * bool) list) list) ((pol, mode) : bool * mode) (ts, ps) =
   let
     val vTs = distinct (op =) (fold Term.add_frees (map term_of_prem ps) (fold Term.add_frees ts []))
-    val modes' = modes @ (param_vs ~~ map (fn x => [(x, false)]) (ho_arg_modes_of mode))
-    val (in_ts, out_ts) = split_mode mode ts    
+    val modes' = modes @ (param_vs ~~ map (fn x => [((true, x), false), ((false, x), false)]) (ho_arg_modes_of mode))
+    fun retrieve_modes_of_pol pol = map (fn (s, ms) =>
+      (s, map_filter (fn ((p, m), r) => if p = pol then SOME m else NONE | _ => NONE) ms))
+    val (pos_modes', neg_modes') =
+      if #infer_pos_and_neg_modes mode_analysis_options then
+        (retrieve_modes_of_pol pol modes', retrieve_modes_of_pol (not pol) modes')
+      else
+        let
+          val modes = map (fn (s, ms) => (s, map (fn ((p, m), r) => m) ms)) modes'
+        in (modes, modes) end
+    val (in_ts, out_ts) = split_mode mode ts
     val in_vs = maps (vars_of_destructable_term thy) in_ts
     val out_vs = terms_vs out_ts
+    fun known_vs_after p vs = (case p of
+        Prem t => union (op =) vs (term_vs t)
+      | Sidecond t => union (op =) vs (term_vs t)
+      | Negprem t => union (op =) vs (term_vs t)
+      | _ => error "I do not know")
     fun check_mode_prems acc_ps rnd vs [] = SOME (acc_ps, vs, rnd)
       | check_mode_prems acc_ps rnd vs ps =
-        (case select_mode_prem' thy modes' vs ps of
-          SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd (*TODO: uses random? *)
-            (case p of
-                Prem t => union (op =) vs (term_vs t)
-              | Sidecond t => vs
-              | Negprem t => union (op =) vs (term_vs t)
-              | _ => error "I do not know")
-            (filter_out (equal p) ps)
+        (case
+          (select_mode_prem mode_analysis_options thy pol (modes', (pos_modes', neg_modes')) vs ps) of
+          SOME (p, SOME (deriv, [])) => check_mode_prems ((p, deriv) :: acc_ps) rnd
+            (known_vs_after p vs) (filter_out (equal p) ps)
         | SOME (p, SOME (deriv, missing_vars)) =>
-          if use_random then
+          if #use_random mode_analysis_options andalso pol then
             check_mode_prems ((p, deriv) :: (map
-              (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output)) missing_vars)
-                @ acc_ps) true
-            (case p of
-                Prem t => union (op =) vs (term_vs t)
-              | Sidecond t => union (op =) vs (term_vs t)
-              | Negprem t => union (op =) vs (term_vs t)
-              | _ => error "I do not know")
-            (filter_out (equal p) ps)
+              (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
+                (distinct (op =) missing_vars))
+                @ acc_ps) true (known_vs_after p vs) (filter_out (equal p) ps)
           else NONE
         | SOME (p, NONE) => NONE
         | NONE => NONE)
@@ -1201,11 +1238,11 @@
       if forall (is_constructable thy vs) (in_ts @ out_ts) then
         SOME (ts, rev acc_ps, rnd)
       else
-        if use_random then
+        if #use_random mode_analysis_options andalso pol then
           let
-            val generators = map
+             val generators = map
               (fn v => (Generator (v, the (AList.lookup (op =) vTs v)), Term Output))
-                (subtract (op =) vs (terms_vs out_ts))
+                (subtract (op =) vs (terms_vs (in_ts @ out_ts)))
           in
             SOME (ts, rev (generators @ acc_ps), true)
           end
@@ -1215,66 +1252,120 @@
 
 datatype result = Success of bool | Error of string
 
-fun check_modes_pred' use_random options thy param_vs clauses modes (p, ms) =
+fun check_modes_pred' mode_analysis_options options thy param_vs clauses modes (p, (ms : ((bool * mode) * bool) list)) =
   let
     fun split xs =
       let
         fun split' [] (ys, zs) = (rev ys, rev zs)
           | split' ((m, Error z) :: xs) (ys, zs) = split' xs (ys, z :: zs)
-          | split' ((m, Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
+          | split' (((m : bool * mode), Success rnd) :: xs) (ys, zs) = split' xs ((m, rnd) :: ys, zs)
        in
          split' xs ([], [])
        end
     val rs = these (AList.lookup (op =) clauses p)
     fun check_mode m =
       let
-        val res = map (check_mode_clause' use_random thy param_vs modes m) rs
+        val res = Output.cond_timeit false "work part of check_mode for one mode" (fn _ => 
+          map (check_mode_clause' mode_analysis_options thy param_vs modes m) rs)
       in
+        Output.cond_timeit false "aux part of check_mode for one mode" (fn _ => 
         case find_indices is_none res of
           [] => Success (exists (fn SOME (_, _, true) => true | _ => false) res)
-        | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is))
+        | is => (print_failed_mode options thy modes p m rs is; Error (error_of p m is)))
       end
-    val res = map (fn (m, _) => (m, check_mode m)) ms
+    val _ = if show_mode_inference options then
+        tracing ("checking " ^ string_of_int (length ms) ^ " modes ...")
+      else ()
+    val res = Output.cond_timeit false "check_mode" (fn _ => map (fn (m, _) => (m, check_mode m)) ms)
     val (ms', errors) = split res
   in
-    ((p, ms'), errors)
+    ((p, (ms' : ((bool * mode) * bool) list)), errors)
   end;
 
-fun get_modes_pred' use_random thy param_vs clauses modes (p, ms) =
+fun get_modes_pred' mode_analysis_options thy param_vs clauses modes (p, ms) =
   let
     val rs = these (AList.lookup (op =) clauses p)
   in
     (p, map (fn (m, rnd) =>
-      (m, map ((fn (ts, ps, rnd) => (ts, ps)) o the o check_mode_clause' use_random thy param_vs modes m) rs)) ms)
+      (m, map
+        ((fn (ts, ps, rnd) => (ts, ps)) o the o
+          check_mode_clause' mode_analysis_options thy param_vs modes m) rs)) ms)
   end;
 
-fun fixp f x =
+fun fixp f (x : (string * ((bool * mode) * bool) list) list) =
   let val y = f x
   in if x = y then x else fixp f y end;
 
-fun fixp_with_state f (x, state) =
+fun fixp_with_state f (x : (string * ((bool * mode) * bool) list) list, state) =
   let
     val (y, state') = f (x, state)
   in
     if x = y then (y, state') else fixp_with_state f (y, state')
   end
 
-fun infer_modes use_random options preds extra_modes param_vs clauses thy =
+fun string_of_ext_mode ((pol, mode), rnd) =
+  string_of_mode mode ^ "(" ^ (if pol then "pos" else "neg") ^ ", "
+  ^ (if rnd then "rnd" else "nornd") ^ ")"
+
+fun print_extra_modes options modes =
+  if show_mode_inference options then
+    tracing ("Modes of inferred predicates: " ^
+      cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map string_of_ext_mode ms)) modes))
+  else ()
+
+fun infer_modes mode_analysis_options options compilation preds all_modes param_vs clauses thy =
   let
-    val all_modes = map (fn (s, T) => (s, map (rpair false) (all_modes_of_typ T))) preds
-    fun needs_random s m = (m, member (op =) (#needs_random (the_pred_data thy s)) m)
-    val extra_modes = map (fn (s, ms) => (s, map (needs_random s) ms)) extra_modes
-    val (modes, errors) =
-      fixp_with_state (fn (modes, errors) =>
+    val collect_errors = false
+    fun appair f (x1, x2) (y1, y2) = (f x1 y1, f x2 y2)
+    fun needs_random s (false, m) = ((false, m), false)
+      | needs_random s (true, m) = ((true, m), member (op =) (#needs_random (the_pred_data thy s)) m)
+    fun add_polarity_and_random_bit s b ms = map (fn m => needs_random s (b, m)) ms
+    val prednames = map fst preds
+    (* extramodes contains all modes of all constants, should we only use the necessary ones
+       - what is the impact on performance? *)
+    val extra_modes =
+      if #infer_pos_and_neg_modes mode_analysis_options then
         let
-          val res = map
-            (check_modes_pred' use_random options thy param_vs clauses (modes @ extra_modes)) modes
-        in (map fst res, errors @ maps snd res) end)
-          (all_modes, [])
+          val pos_extra_modes =
+            all_modes_of compilation thy |> filter_out (fn (name, _) => member (op =) prednames name)
+          val neg_extra_modes =
+            all_modes_of (negative_compilation_of compilation) thy
+            |> filter_out (fn (name, _) => member (op =) prednames name)
+        in
+          map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)
+                @ add_polarity_and_random_bit s false (the (AList.lookup (op =) neg_extra_modes s))))
+            pos_extra_modes
+        end
+      else
+        map (fn (s, ms) => (s, (add_polarity_and_random_bit s true ms)))
+          (all_modes_of compilation thy |> filter_out (fn (name, _) => member (op =) prednames name))
+    val _ = print_extra_modes options extra_modes
+    val start_modes =
+      if #infer_pos_and_neg_modes mode_analysis_options then
+        map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms @
+          (map (fn m => ((false, m), false)) ms))) all_modes
+      else
+        map (fn (s, ms) => (s, map (fn m => ((true, m), false)) ms)) all_modes
+    fun iteration modes = map
+      (check_modes_pred' mode_analysis_options options thy param_vs clauses (modes @ extra_modes))
+        modes
+    val ((modes : (string * ((bool * mode) * bool) list) list), errors) =
+      Output.cond_timeit false "Fixpount computation of mode analysis" (fn () =>
+      if collect_errors then
+        fixp_with_state (fn (modes, errors) =>
+          let
+            val (modes', new_errors) = split_list (iteration modes)
+          in (modes', errors @ flat new_errors) end) (start_modes, [])
+        else
+          (fixp (fn modes => map fst (iteration modes)) start_modes, []))
+    val moded_clauses = map (get_modes_pred' mode_analysis_options thy param_vs clauses
+      (modes @ extra_modes)) modes
     val thy' = fold (fn (s, ms) => if member (op =) (map fst preds) s then
-      set_needs_random s (map fst (filter (fn (_, rnd) => rnd = true) ms)) else I) modes thy
+      set_needs_random s (map_filter (fn ((true, m), true) => SOME m | _ => NONE) ms) else I)
+      modes thy
+
   in
-    ((map (get_modes_pred' use_random thy param_vs clauses (modes @ extra_modes)) modes, errors), thy')
+    ((moded_clauses, errors), thy')
   end;
 
 (* term construction *)
@@ -1414,14 +1505,25 @@
        (v', mk_bot compfuns U')]))
   end;
 
-fun compile_expr compilation_modifiers compfuns thy (t, deriv) additional_arguments =
+fun string_of_tderiv thy (t, deriv) = 
+  (case (t, deriv) of
+    (t1 $ t2, Mode_App (deriv1, deriv2)) =>
+      string_of_tderiv thy (t1, deriv1) ^ " $ " ^ string_of_tderiv thy (t2, deriv2)
+  | (Const ("Pair", _) $ t1 $ t2, Mode_Pair (deriv1, deriv2)) =>
+    "(" ^ string_of_tderiv thy (t1, deriv1) ^ ", " ^ string_of_tderiv thy (t2, deriv2) ^ ")"
+  | (t, Term Input) => Syntax.string_of_term_global thy t ^ "[Input]"
+  | (t, Term Output) => Syntax.string_of_term_global thy t ^ "[Output]"
+  | (t, Context m) => Syntax.string_of_term_global thy t ^ "[" ^ string_of_mode m ^ "]")
+
+fun compile_expr compilation_modifiers compfuns thy pol (t, deriv) additional_arguments =
   let
     fun expr_of (t, deriv) =
       (case (t, deriv) of
         (t, Term Input) => SOME t
       | (t, Term Output) => NONE
       | (Const (name, T), Context mode) =>
-        SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy name mode,
+        SOME (Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy name
+          (pol, mode),
           Comp_Mod.funT_of compilation_modifiers mode T))
       | (Free (s, T), Context m) =>
         SOME (Free (s, Comp_Mod.funT_of compilation_modifiers m T))
@@ -1446,7 +1548,7 @@
   end
 
 fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments
-  mode inp (ts, moded_ps) =
+  (pol, mode) inp (ts, moded_ps) =
   let
     val iss = ho_arg_modes_of mode
     val compile_match = compile_match compilation_modifiers compfuns
@@ -1479,7 +1581,7 @@
                  let
                    val u =
                      compile_expr compilation_modifiers compfuns thy
-                       (t, deriv) additional_arguments'
+                       pol (t, deriv) additional_arguments'
                    val (_, out_ts''') = split_mode mode (snd (strip_comb t))
                    val rest = compile_prems out_ts''' vs' names'' ps
                  in
@@ -1489,7 +1591,7 @@
                  let
                    val u = mk_not compfuns
                      (compile_expr compilation_modifiers compfuns thy
-                       (t, deriv) additional_arguments')
+                       (not pol) (t, deriv) additional_arguments')
                    val (_, out_ts''') = split_mode mode (snd (strip_comb t))
                    val rest = compile_prems out_ts''' vs' names'' ps
                  in
@@ -1506,6 +1608,7 @@
              | Generator (v, T) =>
                  let
                    val u = mk_random T
+                   
                    val rest = compile_prems [Free (v, T)]  vs' names'' ps;
                  in
                    (u, rest)
@@ -1519,7 +1622,7 @@
     mk_bind compfuns (mk_single compfuns inp, prem_t)
   end
 
-fun compile_pred compilation_modifiers thy all_vs param_vs s T mode moded_cls =
+fun compile_pred compilation_modifiers thy all_vs param_vs s T (pol, mode) moded_cls =
   let
     (* TODO: val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers
       (all_vs @ param_vs)
@@ -1547,14 +1650,14 @@
       (fn t as Free (x, _) => if member (op =) param_vs x then NONE else SOME t | t => SOME t)) in_ts
     val cl_ts =
       map (compile_clause compilation_modifiers compfuns
-        thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts')) moded_cls;
+        thy all_vs param_vs additional_arguments (pol, mode) (HOLogic.mk_tuple in_ts')) moded_cls;
     val compilation = Comp_Mod.wrap_compilation compilation_modifiers compfuns
       s T mode additional_arguments
       (if null cl_ts then
         mk_bot compfuns (HOLogic.mk_tupleT outTs)
       else foldr1 (mk_sup compfuns) cl_ts)
     val fun_const =
-      Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy s mode, funT)
+      Const (function_name_of (Comp_Mod.compilation compilation_modifiers) thy s (pol, mode), funT)
   in
     HOLogic.mk_Trueprop
       (HOLogic.mk_eq (list_comb (fun_const, in_ts @ additional_arguments), compilation))
@@ -1583,13 +1686,20 @@
   | strip_split_abs (Abs (_, _, t)) = strip_split_abs t
   | strip_split_abs t = t
 
-fun mk_args is_eval (Pair (m1, m2), Type ("*", [T1, T2])) names =
-    let
-      val (t1, names') = mk_args is_eval (m1, T1) names
-      val (t2, names'') = mk_args is_eval (m2, T2) names'
-    in
-      (HOLogic.mk_prod (t1, t2), names'')
-    end
+fun mk_args is_eval (m as Pair (m1, m2), T as Type ("*", [T1, T2])) names =
+    if eq_mode (m, Input) orelse eq_mode (m, Output) then
+      let
+        val x = Name.variant names "x"
+      in
+        (Free (x, T), x :: names)
+      end
+    else
+      let
+        val (t1, names') = mk_args is_eval (m1, T1) names
+        val (t2, names'') = mk_args is_eval (m2, T2) names'
+      in
+        (HOLogic.mk_prod (t1, t2), names'')
+      end
   | mk_args is_eval ((m as Fun _), T) names =
     let
       val funT = funT_of PredicateCompFuns.compfuns m T
@@ -1828,11 +1938,11 @@
 
 (* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
 
-fun prove_sidecond thy modes t =
+fun prove_sidecond thy t =
   let
     fun preds_of t nameTs = case strip_comb t of 
       (f as Const (name, T), args) =>
-        if AList.defined (op =) modes name then (name, T) :: nameTs
+        if is_registered thy name then (name, T) :: nameTs
           else fold preds_of args nameTs
       | _ => nameTs
     val preds = preds_of t []
@@ -1847,7 +1957,7 @@
     (* need better control here! *)
   end
 
-fun prove_clause options thy nargs modes mode (_, clauses) (ts, moded_ps) =
+fun prove_clause options thy nargs mode (_, clauses) (ts, moded_ps) =
   let
     val (in_ts, clause_out_ts) = split_mode mode ts;
     fun prove_prems out_ts [] =
@@ -1911,7 +2021,7 @@
           | Sidecond t =>
            rtac @{thm if_predI} 1
            THEN print_tac' options "before sidecond:"
-           THEN prove_sidecond thy modes t
+           THEN prove_sidecond thy t
            THEN print_tac' options "after sidecond:"
            THEN prove_prems [] ps)
       in (prove_match options thy out_ts)
@@ -1929,7 +2039,7 @@
   | select_sup _ 1 = [rtac @{thm supI1}]
   | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
 
-fun prove_one_direction options thy clauses preds modes pred mode moded_clauses =
+fun prove_one_direction options thy clauses preds pred mode moded_clauses =
   let
     val T = the (AList.lookup (op =) preds pred)
     val nargs = length (binder_types T)
@@ -1942,7 +2052,7 @@
     THEN (EVERY (map
            (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
              (1 upto (length moded_clauses))))
-    THEN (EVERY (map2 (prove_clause options thy nargs modes mode) clauses moded_clauses))
+    THEN (EVERY (map2 (prove_clause options thy nargs mode) clauses moded_clauses))
     THEN print_tac' options "proved one direction"
   end;
 
@@ -2026,10 +2136,10 @@
 (* FIXME: what is this for? *)
 (* replace defined by has_mode thy pred *)
 (* TODO: rewrite function *)
-fun prove_sidecond2 thy modes t = let
+fun prove_sidecond2 thy t = let
   fun preds_of t nameTs = case strip_comb t of 
     (f as Const (name, T), args) =>
-      if AList.defined (op =) modes name then (name, T) :: nameTs
+      if is_registered thy name then (name, T) :: nameTs
         else fold preds_of args nameTs
     | _ => nameTs
   val preds = preds_of t []
@@ -2044,7 +2154,7 @@
    THEN print_tac "after sidecond2 simplification"
    end
   
-fun prove_clause2 thy modes pred mode (ts, ps) i =
+fun prove_clause2 thy pred mode (ts, ps) i =
   let
     val pred_intro_rule = nth (intros_of thy pred) (i - 1)
     val (in_ts, clause_out_ts) = split_mode mode ts;
@@ -2102,7 +2212,7 @@
         | Sidecond t =>
           etac @{thm bindE} 1
           THEN etac @{thm if_predE} 1
-          THEN prove_sidecond2 thy modes t 
+          THEN prove_sidecond2 thy t
           THEN prove_prems2 [] ps)
       in print_tac "before prove_match2:"
          THEN prove_match2 thy out_ts
@@ -2119,11 +2229,11 @@
     THEN prems_tac
   end;
  
-fun prove_other_direction options thy modes pred mode moded_clauses =
+fun prove_other_direction options thy pred mode moded_clauses =
   let
     fun prove_clause clause i =
       (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
-      THEN (prove_clause2 thy modes pred mode clause i)
+      THEN (prove_clause2 thy pred mode clause i)
   in
     (DETERM (TRY (rtac @{thm unit.induct} 1)))
      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
@@ -2136,7 +2246,7 @@
 
 (** proof procedure **)
 
-fun prove_pred options thy clauses preds modes pred mode (moded_clauses, compiled_term) =
+fun prove_pred options thy clauses preds pred (pol, mode) (moded_clauses, compiled_term) =
   let
     val ctxt = ProofContext.init thy
     val clauses = case AList.lookup (op =) clauses pred of SOME rs => rs | NONE => []
@@ -2146,9 +2256,9 @@
         (fn _ =>
         rtac @{thm pred_iffI} 1
         THEN print_tac' options "after pred_iffI"
-        THEN prove_one_direction options thy clauses preds modes pred mode moded_clauses
+        THEN prove_one_direction options thy clauses preds pred mode moded_clauses
         THEN print_tac' options "proved one direction"
-        THEN prove_other_direction options thy modes pred mode moded_clauses
+        THEN prove_other_direction options thy pred mode moded_clauses
         THEN print_tac' options "proved other direction")
       else (fn _ => Skip_Proof.cheat_tac thy))
   end;
@@ -2173,11 +2283,11 @@
   map_preds_modes (fn pred => compile_pred comp_modifiers thy all_vs param_vs pred
       (the (AList.lookup (op =) preds pred))) moded_clauses
 
-fun prove options thy clauses preds modes moded_clauses compiled_terms =
-  map_preds_modes (prove_pred options thy clauses preds modes)
+fun prove options thy clauses preds moded_clauses compiled_terms =
+  map_preds_modes (prove_pred options thy clauses preds)
     (join_preds_modes moded_clauses compiled_terms)
 
-fun prove_by_skip options thy _ _ _ _ compiled_terms =
+fun prove_by_skip options thy _ _ _ compiled_terms =
   map_preds_modes
     (fn pred => fn mode => fn t => Drule.export_without_context (Skip_Proof.make_thm thy t))
     compiled_terms
@@ -2204,9 +2314,13 @@
     val ([preds, intrs], _) = fold_burrow (Variable.import_terms false) [preds, intrs]
       (ProofContext.init thy)
     val preds = map dest_Const preds
-    val extra_modes =
-      all_modes_of compilation thy |> filter_out (fn (name, _) => member (op =) prednames name)
     val all_vs = terms_vs intrs
+    val all_modes = 
+      map (fn (s, T) =>
+        (s,
+            (if member (op =) (no_higher_order_predicate options) s then
+               (all_smodes_of_typ T)
+            else (all_modes_of_typ T)))) preds
     val params =
       case intrs of
         [] =>
@@ -2219,8 +2333,12 @@
           in
             map2 (curry Free) param_names paramTs
           end
-      | (intr :: _) => maps extract_params
-          (snd (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr))))
+      | (intr :: _) =>
+        let
+          val (p, args) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl intr)) 
+        in
+          ho_args_of (hd (the (AList.lookup (op =) all_modes (fst (dest_Const p))))) args
+        end
     val param_vs = map (fst o dest_Free) params
     fun add_clause intr clauses =
       let
@@ -2232,7 +2350,7 @@
       end;
     val clauses = fold add_clause intrs []
   in
-    (preds, all_vs, param_vs, extra_modes, clauses)
+    (preds, all_vs, param_vs, all_modes, clauses)
   end;
 
 (* sanity check of introduction rules *)
@@ -2294,7 +2412,7 @@
             val arg_names = Name.variant_list []
               (map (fn i => "x" ^ string_of_int i) (1 upto length Ts))
             val args = map2 (curry Free) arg_names Ts
-            val predfun = Const (function_name_of Pred thy predname full_mode,
+            val predfun = Const (function_name_of Pred thy predname (true, full_mode),
               Ts ---> PredicateCompFuns.mk_predT @{typ unit})
             val rhs = @{term Predicate.holds} $ (list_comb (predfun, args))
             val eq_term = HOLogic.mk_Trueprop
@@ -2317,20 +2435,20 @@
 
 datatype steps = Steps of
   {
-  define_functions : options -> (string * typ) list -> string * mode list -> theory -> theory,
-  infer_modes : options -> (string * typ) list -> (string * mode list) list
+  define_functions : options -> (string * typ) list -> string * (bool * mode) list -> theory -> theory,
+  (*infer_modes : options -> (string * typ) list -> (string * mode list) list
     -> string list -> (string * (term list * indprem list) list) list
-    -> theory -> ((moded_clause list pred_mode_table * string list) * theory),
-  prove : options -> theory -> (string * (term list * indprem list) list) list
-    -> (string * typ) list -> (string * mode list) list
+    -> theory -> ((moded_clause list pred_mode_table * string list) * theory),*)
+  prove : options -> theory -> (string * (term list * indprem list) list) list -> (string * typ) list
     -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
   add_code_equations : theory -> (string * typ) list
     -> (string * thm list) list -> (string * thm list) list,
   comp_modifiers : Comp_Mod.comp_modifiers,
+  use_random : bool,
   qname : bstring
   }
 
-fun add_equations_of steps options prednames thy =
+fun add_equations_of steps mode_analysis_options options prednames thy =
   let
     fun dest_steps (Steps s) = s
     val _ = print_step options
@@ -2338,14 +2456,20 @@
       (*val _ = check_intros_elim_match thy prednames*)
       (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
     val compilation = Comp_Mod.compilation (#comp_modifiers (dest_steps steps))
-    val (preds, all_vs, param_vs, extra_modes, clauses) =
+    val _ =
+      if show_intermediate_results options then
+        tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
+      else ()
+    val (preds, all_vs, param_vs, all_modes, clauses) =
       prepare_intrs options compilation thy prednames (maps (intros_of thy) prednames)
     val _ = print_step options "Infering modes..."
     val ((moded_clauses, errors), thy') =
-      #infer_modes (dest_steps steps) options preds extra_modes param_vs clauses thy
+      (*Output.cond_timeit true "Infering modes"
+      (fn _ =>*) infer_modes mode_analysis_options
+        options compilation preds all_modes param_vs clauses thy
     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
     val _ = check_expected_modes preds options modes
-    val _ = check_proposed_modes preds options modes extra_modes errors
+    (*val _ = check_proposed_modes preds options modes (fst extra_modes) errors*)
     val _ = print_modes options thy' modes
     val _ = print_step options "Defining executable functions..."
     val thy'' = fold (#define_functions (dest_steps steps) options preds) modes thy'
@@ -2355,8 +2479,8 @@
       compile_preds (#comp_modifiers (dest_steps steps)) thy'' all_vs param_vs preds moded_clauses
     val _ = print_compiled_terms options thy'' compiled_terms
     val _ = print_step options "Proving equations..."
-    val result_thms = #prove (dest_steps steps) options thy'' clauses preds (extra_modes @ modes)
-      moded_clauses compiled_terms
+    val result_thms =
+      #prove (dest_steps steps) options thy'' clauses preds moded_clauses compiled_terms
     val result_thms' = #add_code_equations (dest_steps steps) thy'' preds
       (maps_modes result_thms)
     val qname = #qname (dest_steps steps)
@@ -2398,7 +2522,14 @@
     val thy'' = fold_rev
       (fn preds => fn thy =>
         if not (forall (defined thy) preds) then
-          add_equations_of steps options preds thy
+          let
+            val mode_analysis_options = {use_random = #use_random (dest_steps steps),
+              reorder_premises =
+                not (no_topmost_reordering options andalso not (null (inter (op =) preds names))),
+              infer_pos_and_neg_modes = #use_random (dest_steps steps)}
+          in
+            add_equations_of steps mode_analysis_options options preds thy
+          end
         else thy)
       scc thy' |> Theory.checkpoint
   in thy'' end
@@ -2468,11 +2599,15 @@
   }
 
 val add_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes false,
-  define_functions = create_definitions,
+  (Steps {
+  define_functions =
+    fn options => fn preds => fn (s, modes) =>
+      create_definitions
+      options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
   prove = prove,
   add_code_equations = add_code_equations,
   comp_modifiers = predicate_comp_modifiers,
+  use_random = false,
   qname = "equation"})
 
 val annotated_comp_modifiers = Comp_Mod.Comp_Modifiers
@@ -2499,9 +2634,9 @@
   transform_additional_arguments = K I : (indprem -> term list -> term list)
   }
 
-val random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
+val pos_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
   {
-  compilation = Random_DSeq,
+  compilation = Pos_Random_DSeq,
   function_name_prefix = "random_dseq_",
   compfuns = Random_Sequence_CompFuns.compfuns,
   additional_arguments = K [],
@@ -2510,6 +2645,17 @@
   transform_additional_arguments = K I : (indprem -> term list -> term list)
   }
 
+val neg_random_dseq_comp_modifiers = Comp_Mod.Comp_Modifiers
+  {
+  compilation = Neg_Random_DSeq,
+  function_name_prefix = "random_dseq_neg_",
+  compfuns = Random_Sequence_CompFuns.compfuns,
+  additional_arguments = K [],
+  wrap_compilation = K (K (K (K (K I))))
+   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+  transform_additional_arguments = K I : (indprem -> term list -> term list)
+  }
+
 (*
 val add_depth_limited_equations = gen_add_equations
   (Steps {infer_modes = infer_modes,
@@ -2521,11 +2667,15 @@
   qname = "depth_limited_equation"})
 *)
 val add_annotated_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes false,
-  define_functions = define_functions annotated_comp_modifiers PredicateCompFuns.compfuns,
+  (Steps {
+  define_functions =
+    fn options => fn preds => fn (s, modes) =>
+      define_functions annotated_comp_modifiers PredicateCompFuns.compfuns options preds
+      (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
   prove = prove_by_skip,
   add_code_equations = K (K I),
   comp_modifiers = annotated_comp_modifiers,
+  use_random = false,
   qname = "annotated_equation"})
 (*
 val add_quickcheck_equations = gen_add_equations
@@ -2538,19 +2688,33 @@
   qname = "random_equation"})
 *)
 val add_dseq_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes false,
-  define_functions = define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns,
+  (Steps {
+  define_functions =
+  fn options => fn preds => fn (s, modes) =>
+    define_functions dseq_comp_modifiers DSequence_CompFuns.compfuns
+    options preds (s, map_filter (fn (true, m) => SOME m | _ => NONE) modes),
   prove = prove_by_skip,
   add_code_equations = K (K I),
   comp_modifiers = dseq_comp_modifiers,
+  use_random = false,
   qname = "dseq_equation"})
 
 val add_random_dseq_equations = gen_add_equations
-  (Steps {infer_modes = infer_modes true,
-  define_functions = define_functions random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns,
+  (Steps {
+  define_functions =
+    fn options => fn preds => fn (s, modes) =>
+    let
+      val pos_modes = map_filter (fn (true, m) => SOME m | _ => NONE) modes
+      val neg_modes = map_filter (fn (false, m) => SOME m | _ => NONE) modes
+    in define_functions pos_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
+      options preds (s, pos_modes)
+      #> define_functions neg_random_dseq_comp_modifiers Random_Sequence_CompFuns.compfuns
+      options preds (s, neg_modes)
+    end,
   prove = prove_by_skip,
   add_code_equations = K (K I),
-  comp_modifiers = random_dseq_comp_modifiers,
+  comp_modifiers = pos_random_dseq_comp_modifiers,
+  use_random = true,
   qname = "random_dseq_equation"})
 
 
@@ -2700,8 +2864,8 @@
           | Depth_Limited => depth_limited_comp_modifiers
           | Annotated => annotated_comp_modifiers*)
           | DSeq => dseq_comp_modifiers
-          | Random_DSeq => random_dseq_comp_modifiers
-        val t_pred = compile_expr comp_modifiers compfuns thy (body, deriv) additional_arguments;
+          | Random_DSeq => pos_random_dseq_comp_modifiers
+        val t_pred = compile_expr comp_modifiers compfuns thy true (body, deriv) additional_arguments;
         val T_pred = dest_predT compfuns (fastype_of t_pred)
         val arrange = split_lambda (HOLogic.mk_tuple outargs) output_tuple
       in
@@ -2717,7 +2881,7 @@
       case compilation of
         Random => RandomPredCompFuns.compfuns
       | DSeq => DSequence_CompFuns.compfuns
-      | Random_DSeq => Random_Sequence_CompFuns.compfuns
+      | Pos_Random_DSeq => Random_Sequence_CompFuns.compfuns
       | _ => PredicateCompFuns.compfuns
     val t = analyze_compr thy compfuns param_user_modes options t_compr;
     val T = dest_predT compfuns (fastype_of t);
@@ -2729,7 +2893,7 @@
           (Code_Eval.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
             (fn proc => fn g => fn s => g s |>> Predicate.map proc) thy t' []
             |> Random_Engine.run))
-      | Random_DSeq =>
+      | Pos_Random_DSeq =>
           let
             val [nrandom, size, depth] = arguments
           in
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -6,11 +6,17 @@
 
 signature PREDICATE_COMPILE_DATA =
 sig
-  type specification_table;
-  (*val make_const_spec_table : Predicate_Compile_Aux.options -> theory -> specification_table*)
-  val get_specification : theory -> term -> thm list
+  val ignore_consts : string list -> theory -> theory
+  val keep_functions : string list -> theory -> theory
+  val keep_function : theory -> string -> bool
+  val processed_specs : theory -> string -> (string * thm list) list option
+  val store_processed_specs : (string * (string * thm list) list) -> theory -> theory
+  
+  val get_specification : Predicate_Compile_Aux.options -> theory -> term -> thm list
   val obtain_specification_graph :
     Predicate_Compile_Aux.options -> theory -> term -> thm list TermGraph.T
+    
+  val present_graph : thm list TermGraph.T -> unit
   val normalize_equation : theory -> thm -> thm
 end;
 
@@ -22,20 +28,39 @@
 structure Data = Theory_Data
 (
   type T =
-    {const_spec_table : thm list Symtab.table};
+    {ignore_consts : unit Symtab.table,
+     keep_functions : unit Symtab.table,
+     processed_specs : ((string * thm list) list) Symtab.table};
   val empty =
-    {const_spec_table = Symtab.empty};
+    {ignore_consts = Symtab.empty,
+     keep_functions = Symtab.empty,
+     processed_specs =  Symtab.empty};
   val extend = I;
   fun merge
-    ({const_spec_table = const_spec_table1},
-     {const_spec_table = const_spec_table2}) =
-     {const_spec_table = Symtab.merge (K true) (const_spec_table1, const_spec_table2)}
+    ({ignore_consts = c1, keep_functions = k1, processed_specs = s1},
+     {ignore_consts = c2, keep_functions = k2, processed_specs = s2}) =
+     {ignore_consts = Symtab.merge (K true) (c1, c2),
+      keep_functions = Symtab.merge (K true) (k1, k2),
+      processed_specs = Symtab.merge (K true) (s1, s2)}
 );
 
-fun mk_data c = {const_spec_table = c}
-fun map_data f {const_spec_table = c} = mk_data (f c)
+
+
+fun mk_data (c, k, s) = {ignore_consts = c, keep_functions = k, processed_specs = s}
+fun map_data f {ignore_consts = c, keep_functions = k, processed_specs = s} = mk_data (f (c, k, s))
+
+fun ignore_consts cs = Data.map (map_data (apfst3 (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
+
+fun keep_functions cs = Data.map (map_data (apsnd3 (fold (fn c => Symtab.insert (op =) (c, ())) cs)))
 
-type specification_table = thm list Symtab.table
+fun keep_function thy = Symtab.defined (#keep_functions (Data.get thy))
+
+fun processed_specs thy = Symtab.lookup (#processed_specs (Data.get thy))
+
+fun store_processed_specs (constname, specs) =
+  Data.map (map_data (aptrd3 (Symtab.update_new (constname, specs))))
+(* *)
+
 
 fun defining_term_of_introrule_term t =
   let
@@ -120,17 +145,11 @@
     val t' = Pattern.rewrite_term thy rewr [] t
     val tac = Skip_Proof.cheat_tac thy
     val th'' = Goal.prove ctxt (Term.add_free_names t' []) [] t' (fn _ => tac)
-    val th''' = LocalDefs.unfold ctxt [@{thm split_conv}] th''
+    val th''' = LocalDefs.unfold ctxt [@{thm split_conv}, @{thm fst_conv}, @{thm snd_conv}] th''
   in
     th'''
   end;
 
-fun normalize_equation thy th =
-  mk_meta_equation th
-  |> Predicate_Compile_Set.unfold_set_notation
-  |> full_fun_cong_expand
-  |> split_all_pairs thy
-  |> tap check_equation_format
 
 fun inline_equations thy th =
   let
@@ -143,69 +162,58 @@
   in
     th'
   end
-(*
-fun store_thm_in_table options ignore thy th=
-  let
-    val th = th
-      |> inline_equations options thy
-      |> Predicate_Compile_Set.unfold_set_notation
-      |> AxClass.unoverload thy
-    val (const, th) =
-      if is_equationlike th then
-        let
-          val eq = normalize_equation thy th
-        in
-          (defining_const_of_equation eq, eq)
-        end
-      else if is_introlike th then (defining_const_of_introrule th, th)
-      else error "store_thm: unexpected definition format"
-  in
-    if ignore const then I else Symtab.cons_list (const, th)
-  end
+
+fun normalize_equation thy th =
+  mk_meta_equation th
+  |> full_fun_cong_expand
+  |> split_all_pairs thy
+  |> tap check_equation_format
+  |> inline_equations thy
 
-fun make_const_spec_table options thy =
+fun normalize_intros thy th =
+  split_all_pairs thy th
+  |> inline_equations thy
+
+fun normalize thy th =
+  if is_equationlike th then
+    normalize_equation thy th
+  else
+    normalize_intros thy th
+
+fun get_specification options thy t =
   let
-    fun store ignore f =
-      fold (store_thm_in_table options ignore thy)
-        (map (Thm.transfer thy) (f ))
-    val table = Symtab.empty
-      |> store (K false) Predicate_Compile_Alternative_Defs.get
-    val ignore = Symtab.defined table
-  in
-    table
-    |> store ignore (fn ctxt => maps
-      else [])
-        
-    |> store ignore Nitpick_Simps.get
-    |> store ignore Nitpick_Intros.get
-  end
-
-fun get_specification table constname =
-  case Symtab.lookup table constname of
-    SOME thms => thms                  
-  | NONE => error ("get_specification: lookup of constant " ^ quote constname ^ " failed")
-*)
-
-fun get_specification thy t =
-  Output.cond_timeit true "get_specification" (fn () =>
-  let
+    (*val (c, T) = dest_Const t
+    val t = Const (AxClass.unoverload_const thy (c, T), T)*)
+    val _ = if show_steps options then
+        tracing ("getting specification of " ^ Syntax.string_of_term_global thy t ^
+          " with type " ^ Syntax.string_of_typ_global thy (fastype_of t))
+      else ()
     val ctxt = ProofContext.init thy
     fun filtering th =
       if is_equationlike th andalso
-        defining_const_of_equation (normalize_equation thy th) = (fst (dest_Const t)) then
+        defining_const_of_equation (normalize_equation thy th) = fst (dest_Const t) then
         SOME (normalize_equation thy th)
       else
         if is_introlike th andalso defining_term_of_introrule th = t then
           SOME th
         else
           NONE
-  in
-    case map_filter filtering (map (Thm.transfer thy) (Predicate_Compile_Alternative_Defs.get ctxt))
+    val spec = case map_filter filtering (map (normalize thy o Thm.transfer thy) (Predicate_Compile_Alternative_Defs.get ctxt))
      of [] => (case Spec_Rules.retrieve ctxt t
-       of [] => rev (map_filter filtering (map (Thm.transfer thy) (Nitpick_Intros.get ctxt)))
+       of [] => (case rev ( 
+         (map_filter filtering (map (normalize_intros thy o Thm.transfer thy)
+           (Nitpick_Intros.get ctxt))))
+         of [] => error ("No specification for " ^ (Syntax.string_of_term_global thy t))
+          | ths => ths)
        | ((_, (_, ths)) :: _) => map (normalize_equation thy o Thm.transfer thy) ths)
      | ths => rev ths
-  end)
+    val _ =
+      if show_intermediate_results options then
+        Output.tracing (commas (map (Display.string_of_thm_global thy) spec))
+      else ()
+  in
+    spec
+  end
 
 val logic_operator_names =
   [@{const_name "=="}, 
@@ -216,7 +224,8 @@
    @{const_name "op -->"},
    @{const_name "All"},
    @{const_name "Ex"}, 
-   @{const_name "op &"}]
+   @{const_name "op &"},
+   @{const_name "op |"}]
 
 fun special_cases (c, T) = member (op =) [
   @{const_name Product_Type.Unity},
@@ -233,7 +242,11 @@
   @{const_name Int.Bit1},
   @{const_name Int.Pls},
   @{const_name Int.zero_int_inst.zero_int},
-  @{const_name List.filter}] c
+  @{const_name List.filter},
+  @{const_name HOL.If},
+  @{const_name Groups.minus}
+  ] c
+
 
 fun print_specification options thy constname specs = 
   if show_intermediate_results options then
@@ -254,19 +267,43 @@
       |> filter_out has_code_pred_intros
       |> filter_out case_consts
       |> filter_out special_cases
+      |> filter_out (fn (c, _) => Symtab.defined (#ignore_consts (Data.get thy)) c)
+      |> map (fn (c, _) => (c, Sign.the_const_constraint thy c))
       |> map Const
       (*
       |> filter is_defining_constname*)
     fun extend t =
       let
-        val specs = get_specification thy t
-          |> map (inline_equations thy)
+        val specs = get_specification options thy t
           (*|> Predicate_Compile_Set.unfold_set_notation*)
         (*val _ = print_specification options thy constname specs*)
       in (specs, defiants_of specs) end;
   in
     TermGraph.extend extend t TermGraph.empty
   end;
-  
+
+
+fun present_graph gr =
+  let
+    fun eq_cname (Const (c1, _), Const (c2, _)) = (c1 = c2)
+    fun string_of_const (Const (c, _)) = c
+      | string_of_const _ = error "string_of_const: unexpected term"
+    val constss = TermGraph.strong_conn gr;
+    val mapping = Termtab.empty |> fold (fn consts => fold (fn const =>
+      Termtab.update (const, consts)) consts) constss;
+    fun succs consts = consts
+      |> maps (TermGraph.imm_succs gr)
+      |> subtract eq_cname consts
+      |> map (the o Termtab.lookup mapping)
+      |> distinct (eq_list eq_cname);
+    val conn = [] |> fold (fn consts => cons (consts, succs consts)) constss;
+    
+    fun namify consts = map string_of_const consts
+      |> commas;
+    val prgr = map (fn (consts, constss) =>
+      { name = namify consts, ID = namify consts, dir = "", unfold = true,
+        path = "", parents = map namify constss }) conn;
+  in Present.display_graph prgr end;
+
 
 end;
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -9,6 +9,8 @@
   val define_predicates : (string * thm list) list -> theory -> (string * thm list) list * theory
   val rewrite_intro : theory -> thm -> thm list
   val pred_of_function : theory -> string -> string option
+  
+  val add_function_predicate_translation : (term * term) -> theory -> theory
 end;
 
 structure Predicate_Compile_Fun : PREDICATE_COMPILE_FUN =
@@ -16,19 +18,36 @@
 
 open Predicate_Compile_Aux;
 
-(* Table from constant name (string) to term of inductive predicate *)
-structure Pred_Compile_Preproc = Theory_Data
+(* Table from function to inductive predicate *)
+structure Fun_Pred = Theory_Data
 (
-  type T = string Symtab.table;
-  val empty = Symtab.empty;
+  type T = (term * term) Item_Net.T;
+  val empty = Item_Net.init (op aconv o pairself fst) (single o fst);
   val extend = I;
-  fun merge data : T = Symtab.merge (op =) data;   (* FIXME handle Symtab.DUP ?? *)
+  val merge = Item_Net.merge;
 )
 
-fun pred_of_function thy name = Symtab.lookup (Pred_Compile_Preproc.get thy) name
+fun lookup thy net t =
+  case Item_Net.retrieve net t of
+    [] => NONE
+  | [(f, p)] =>
+    let
+      val subst = Pattern.match thy (f, t) (Vartab.empty, Vartab.empty)
+    in
+      SOME (Envir.subst_term subst p)
+    end
+  | _ => error ("Multiple matches possible for lookup of " ^ Syntax.string_of_term_global thy t)
 
-fun defined thy = Symtab.defined (Pred_Compile_Preproc.get thy) 
+fun pred_of_function thy name =
+  case Item_Net.retrieve (Fun_Pred.get thy) (Const (name, Term.dummyT)) of
+    [] => NONE
+  | [(f, p)] => SOME (fst (dest_Const p))
+  | _ => error ("Multiple matches possible for lookup of constant " ^ name)
 
+fun defined_const thy name = is_some (pred_of_function thy name)
+
+fun add_function_predicate_translation (f, p) =
+  Fun_Pred.map (Item_Net.update (f, p))
 
 fun transform_ho_typ (T as Type ("fun", _)) =
   let
@@ -63,27 +82,6 @@
       (Free (Long_Name.base_name name ^ "P", pred_type T))
   end
 
-fun mk_param thy lookup_pred (t as Free (v, _)) = lookup_pred t
-  | mk_param thy lookup_pred t =
-  if Predicate_Compile_Aux.is_predT (fastype_of t) then
-    t
-  else
-    let
-      val (vs, body) = strip_abs t
-      val names = Term.add_free_names body []
-      val vs_names = Name.variant_list names (map fst vs)
-      val vs' = map2 (curry Free) vs_names (map snd vs)
-      val body' = subst_bounds (rev vs', body)
-      val (f, args) = strip_comb body'
-      val resname = Name.variant (vs_names @ names) "res"
-      val resvar = Free (resname, body_type (fastype_of body'))
-      (*val P = case try lookup_pred f of SOME P => P | NONE => error "mk_param"
-      val pred_body = list_comb (P, args @ [resvar])
-      *)
-      val pred_body = HOLogic.mk_eq (body', resvar)
-      val param = fold_rev lambda (vs' @ [resvar]) pred_body
-    in param end
-    
 (* creates the list of premises for every intro rule *)
 (* theory -> term -> (string list, term list list) *)
 
@@ -92,22 +90,6 @@
   val (func, args) = strip_comb lhs
 in ((func, args), rhs) end;
 
-fun string_of_typ T = Syntax.string_of_typ_global @{theory} T
-
-fun string_of_term t =
-  case t of
-    Const (c, T) => "Const (" ^ c ^ ", " ^ string_of_typ T ^ ")"
-  | Free (c, T) => "Free (" ^ c ^ ", " ^ string_of_typ T ^ ")"
-  | Var ((c, i), T) => "Var ((" ^ c ^ ", " ^ string_of_int i ^ "), " ^ string_of_typ T ^ ")"
-  | Bound i => "Bound " ^ string_of_int i
-  | Abs (x, T, t) => "Abs (" ^ x ^ ", " ^ string_of_typ T ^ ", " ^ string_of_term t ^ ")"
-  | t1 $ t2 => "(" ^ string_of_term t1 ^ ") $ (" ^ string_of_term t2 ^ ")"
-  
-fun ind_package_get_nparams thy name =
-  case try (Inductive.the_inductive (ProofContext.init thy)) name of
-    SOME (_, result) => length (Inductive.params_of (#raw_induct result))
-  | NONE => error ("No such predicate: " ^ quote name) 
-
 (* TODO: does not work with higher order functions yet *)
 fun mk_rewr_eq (func, pred) =
   let
@@ -122,49 +104,6 @@
       (HOLogic.mk_eq (res, list_comb (func, args)), list_comb (pred, args @ [res]))
   end;
 
-fun has_split_rule_cname @{const_name "nat_case"} = true
-  | has_split_rule_cname @{const_name "list_case"} = true
-  | has_split_rule_cname _ = false
-  
-fun has_split_rule_term thy (Const (@{const_name "nat_case"}, _)) = true 
-  | has_split_rule_term thy (Const (@{const_name "list_case"}, _)) = true 
-  | has_split_rule_term thy _ = false
-
-fun has_split_rule_term' thy (Const (@{const_name "If"}, _)) = true
-  | has_split_rule_term' thy (Const (@{const_name "Let"}, _)) = true
-  | has_split_rule_term' thy c = has_split_rule_term thy c
-  
-fun prepare_split_thm ctxt split_thm =
-    (split_thm RS @{thm iffD2})
-    |> LocalDefs.unfold ctxt [@{thm atomize_conjL[symmetric]},
-      @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}]
-
-fun find_split_thm thy (Const (name, typ)) =
-  let
-    fun split_name str =
-      case first_field "." str
-        of (SOME (field, rest)) => field :: split_name rest
-         | NONE => [str]
-    val splitted_name = split_name name
-  in
-    if length splitted_name > 0 andalso
-       String.isSuffix "_case" (List.last splitted_name)
-    then
-      (List.take (splitted_name, length splitted_name - 1)) @ ["split"]
-      |> space_implode "."
-      |> PureThy.get_thm thy
-      |> SOME
-      handle ERROR msg => NONE
-    else NONE
-  end
-  | find_split_thm _ _ = NONE
-
-fun find_split_thm' thy (Const (@{const_name "If"}, _)) = SOME @{thm split_if}
-  | find_split_thm' thy (Const (@{const_name "Let"}, _)) = SOME @{thm refl} (* TODO *)
-  | find_split_thm' thy c = find_split_thm thy c
-
-fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t)
-
 fun folds_map f xs y =
   let
     fun folds_map' acc [] y = [(rev acc, y)]
@@ -174,23 +113,91 @@
       folds_map' [] xs y
     end;
 
-fun mk_prems thy (lookup_pred, get_nparams) t (names, prems) =
+fun keep_functions thy t =
+  case try dest_Const (fst (strip_comb t)) of
+    SOME (c, _) => Predicate_Compile_Data.keep_function thy c
+  | _ => false
+
+fun mk_prems thy lookup_pred t (names, prems) =
   let
     fun mk_prems' (t as Const (name, T)) (names, prems) =
-      if is_constr thy name orelse (is_none (try lookup_pred t)) then
+      (if is_constr thy name orelse (is_none (lookup_pred t)) then
         [(t, (names, prems))]
-      else [(lookup_pred t, (names, prems))]
+      else
+       (*(if is_none (try lookup_pred t) then
+          [(Abs ("uu", fastype_of t, HOLogic.mk_eq (t, Bound 0)), (names, prems))]
+        else*) [(the (lookup_pred t), (names, prems))])
     | mk_prems' (t as Free (f, T)) (names, prems) = 
-      [(lookup_pred t, (names, prems))]
+      (case lookup_pred t of
+        SOME t' => [(t', (names, prems))]
+      | NONE => [(t, (names, prems))])
     | mk_prems' (t as Abs _) (names, prems) =
       if Predicate_Compile_Aux.is_predT (fastype_of t) then
-      [(t, (names, prems))] else error "mk_prems': Abs "
-      (* mk_param *)
+        ([(Envir.eta_contract t, (names, prems))])
+      else
+        let
+          val (vars, body) = strip_abs t
+          val _ = assert (fastype_of body = body_type (fastype_of body))
+          val absnames = Name.variant_list names (map fst vars)
+          val frees = map2 (curry Free) absnames (map snd vars)
+          val body' = subst_bounds (rev frees, body)
+          val resname = Name.variant (absnames @ names) "res"
+          val resvar = Free (resname, fastype_of body)
+          val t = mk_prems' body' ([], [])
+            |> map (fn (res, (inner_names, inner_prems)) =>
+              let
+                fun mk_exists (x, T) t = HOLogic.mk_exists (x, T, t)
+                val vTs = 
+                  fold Term.add_frees inner_prems []
+                  |> filter (fn (x, T) => member (op =) inner_names x)
+                val t = 
+                  fold mk_exists vTs
+                  (foldr1 HOLogic.mk_conj (HOLogic.mk_eq (resvar, res) ::
+                    map HOLogic.dest_Trueprop inner_prems))
+              in
+                t
+              end)
+              |> foldr1 HOLogic.mk_disj
+              |> fold lambda (resvar :: rev frees)
+        in
+          [(t, (names, prems))]
+        end
     | mk_prems' t (names, prems) =
-      if Predicate_Compile_Aux.is_constrt thy t then
+      if Predicate_Compile_Aux.is_constrt thy t orelse keep_functions thy t then
         [(t, (names, prems))]
       else
-        if has_split_rule_term' thy (fst (strip_comb t)) then
+        case (fst (strip_comb t)) of
+          Const (@{const_name "If"}, _) =>
+            (let
+              val (_, [B, x, y]) = strip_comb t
+            in
+              (mk_prems' x (names, prems)
+              |> map (fn (res, (names, prems)) => (res, (names, (HOLogic.mk_Trueprop B) :: prems))))
+              @ (mk_prems' y (names, prems)
+              |> map (fn (res, (names, prems)) =>
+                (res, (names, (HOLogic.mk_Trueprop (HOLogic.mk_not B)) :: prems))))
+            end)
+        | Const (@{const_name "Let"}, _) => 
+            (let
+              val (_, [f, g]) = strip_comb t
+            in
+              mk_prems' f (names, prems)
+              |> maps (fn (res, (names, prems)) =>
+                mk_prems' (betapply (g, res)) (names, prems))
+            end)
+        | Const (@{const_name "split"}, _) => 
+            (let
+              val (_, [g, res]) = strip_comb t
+              val [res1, res2] = Name.variant_list names ["res1", "res2"]
+              val (T1, T2) = HOLogic.dest_prodT (fastype_of res)
+              val (resv1, resv2) = (Free (res1, T1), Free (res2, T2))
+            in
+              mk_prems' (betapplys (g, [resv1, resv2]))
+              (res1 :: res2 :: names,
+              HOLogic.mk_Trueprop (HOLogic.mk_eq (res, HOLogic.mk_prod (resv1, resv2))) :: prems)
+            end)
+        | _ =>
+        if has_split_thm thy (fst (strip_comb t)) then
           let
             val (f, args) = strip_comb t
             val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
@@ -208,8 +215,15 @@
                 val vars = map Free (var_names ~~ (map snd vTs))
                 val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
                 val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+                val (lhss : term list, rhss) =
+                  split_list (map (HOLogic.dest_eq o HOLogic.dest_Trueprop) prems')
               in
-                mk_prems' inner_t (var_names @ names, prems' @ prems)
+                folds_map mk_prems' lhss (var_names @ names, prems)
+                |> map (fn (ress, (names, prems)) =>
+                  let
+                    val prems' = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (ress ~~ rhss)
+                  in (names, prems' @ prems) end)
+                |> maps (mk_prems' inner_t)
               end
           in
             maps mk_prems_of_assm assms
@@ -219,53 +233,77 @@
             val (f, args) = strip_comb t
             (* TODO: special procedure for higher-order functions: split arguments in
               simple types and function types *)
+            val args = map (Pattern.eta_long []) args
             val resname = Name.variant names "res"
             val resvar = Free (resname, body_type (fastype_of t))
+            val _ = assert (fastype_of t = body_type (fastype_of t))
             val names' = resname :: names
             fun mk_prems'' (t as Const (c, _)) =
-              if is_constr thy c orelse (is_none (try lookup_pred t)) then
+              if is_constr thy c orelse (is_none (lookup_pred t)) then
+                let
+                  val _ = ()(*tracing ("not translating function " ^ Syntax.string_of_term_global thy t)*)
+                in
                 folds_map mk_prems' args (names', prems) |>
                 map
                   (fn (argvs, (names'', prems')) =>
                   let
                     val prem = HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
                   in (names'', prem :: prems') end)
+                end
               else
                 let
-                  val pred = lookup_pred t
-                  val nparams = get_nparams pred
-                  val (params, args) = chop nparams args
-                  val params' = map (mk_param thy lookup_pred) params
+                  (* lookup_pred is falsch für polymorphe Argumente und bool. *)
+                  val pred = the (lookup_pred t)
+                  val Ts = binder_types (fastype_of pred)
                 in
                   folds_map mk_prems' args (names', prems)
                   |> map (fn (argvs, (names'', prems')) =>
                     let
-                      val prem = HOLogic.mk_Trueprop (list_comb (pred, params' @ argvs @ [resvar]))
+                      fun lift_arg T t =
+                        if (fastype_of t) = T then t
+                        else
+                          let
+                            val _ = assert (T =
+                              (binder_types (fastype_of t) @ [@{typ bool}] ---> @{typ bool}))
+                            fun mk_if T (b, t, e) =
+                              Const (@{const_name If}, @{typ bool} --> T --> T --> T) $ b $ t $ e
+                            val Ts = binder_types (fastype_of t)
+                            val t = 
+                            list_abs (map (pair "x") Ts @ [("b", @{typ bool})],
+                              mk_if @{typ bool} (list_comb (t, map Bound (length Ts downto 1)),
+                              HOLogic.mk_eq (@{term True}, Bound 0),
+                              HOLogic.mk_eq (@{term False}, Bound 0)))
+                          in
+                            t
+                          end
+                      (*val _ = tracing ("Ts: " ^ commas (map (Syntax.string_of_typ_global thy) Ts))
+                      val _ = map2 check_arity Ts (map fastype_of (argvs @ [resvar]))*)
+                      val argvs' = map2 lift_arg (fst (split_last Ts)) argvs
+                      val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs' @ [resvar]))
                     in (names'', prem :: prems') end)
                 end
             | mk_prems'' (t as Free (_, _)) =
-                let
-                  (* higher order argument call *)
-                  val pred = lookup_pred t
-                in
-                  folds_map mk_prems' args (resname :: names, prems)
-                  |> map (fn (argvs, (names', prems')) =>
-                     let
-                       val prem = HOLogic.mk_Trueprop (list_comb (pred, argvs @ [resvar]))
-                     in (names', prem :: prems') end)
-                end
+              folds_map mk_prems' args (names', prems) |>
+                map
+                  (fn (argvs, (names'', prems')) =>
+                  let
+                    val prem = 
+                      case lookup_pred t of
+                        NONE => HOLogic.mk_Trueprop (HOLogic.mk_eq (resvar, list_comb (f, argvs)))
+                      | SOME p => HOLogic.mk_Trueprop (list_comb (p, argvs @ [resvar]))
+                  in (names'', prem :: prems') end)
             | mk_prems'' t =
               error ("Invalid term: " ^ Syntax.string_of_term_global thy t)
           in
             map (pair resvar) (mk_prems'' f)
           end
   in
-    mk_prems' t (names, prems)
+    mk_prems' (Pattern.eta_long [] t) (names, prems)
   end;
 
 (* assumption: mutual recursive predicates all have the same parameters. *)  
 fun define_predicates specs thy =
-  if forall (fn (const, _) => member (op =) (Symtab.keys (Pred_Compile_Preproc.get thy)) const) specs then
+  if forall (fn (const, _) => defined_const thy const) specs then
     ([], thy)
   else
   let
@@ -275,36 +313,20 @@
       (* create prednames *)
     val ((funs, argss), rhss) = map_split dest_code_eqn eqns |>> split_list
     val argss' = map (map transform_ho_arg) argss
-    val pnames = map dest_Free (distinct (op =) (maps (filter (is_funtype o fastype_of)) argss'))
+    (* TODO: higher order arguments also occur in tuples! *)
+    val ho_argss = distinct (op =) (maps (filter (is_funtype o fastype_of)) argss)
+    val params = distinct (op =) (maps (filter (is_funtype o fastype_of)) argss')
+    val pnames = map dest_Free params
     val preds = map pred_of funs
     val prednames = map (fst o dest_Free) preds
     val funnames = map (fst o dest_Const) funs
     val fun_pred_names = (funnames ~~ prednames)  
       (* mapping from term (Free or Const) to term *)
-    fun lookup_pred (Const (name, T)) =
-      (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
-          SOME c => Const (c, pred_type T)
-        | NONE =>
-          (case AList.lookup op = fun_pred_names name of
-            SOME f => Free (f, pred_type T)
-          | NONE => Const (name, T)))
-      | lookup_pred (Free (name, T)) =
-        if member op = (map fst pnames) name then
-          Free (name, transform_ho_typ T)
-        else
-          Free (name, T)
-      | lookup_pred t =
-         error ("lookup function is not defined for " ^ Syntax.string_of_term_global thy t)
-     
-        (* mapping from term (predicate term, not function term!) to int *)
-    fun get_nparams (Const (name, _)) =
-      the_default 0 (try (ind_package_get_nparams thy) name)
-    | get_nparams (Free (name, _)) =
-        (if member op = prednames name then
-          length pnames
-        else 0)
-    | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
-  
+    fun map_Free f = Free o f o dest_Free
+    val net = fold Item_Net.update
+      ((funs ~~ preds) @ (ho_argss ~~ params))
+        (Fun_Pred.get thy)
+    fun lookup_pred t = lookup thy net t
     (* create intro rules *)
   
     fun mk_intros ((func, pred), (args, rhs)) =
@@ -314,14 +336,15 @@
       else
         let
           val names = Term.add_free_names rhs []
-        in mk_prems thy (lookup_pred, get_nparams) rhs (names, [])
+        in mk_prems thy lookup_pred rhs (names, [])
           |> map (fn (resultt, (names', prems)) =>
             Logic.list_implies (prems, HOLogic.mk_Trueprop (list_comb (pred, args @ [resultt]))))
         end
     fun mk_rewr_thm (func, pred) = @{thm refl}
   in
-    case try (maps mk_intros) ((funs ~~ preds) ~~ (argss' ~~ rhss)) of
-      NONE => ([], thy) 
+    case (*try *)SOME (maps mk_intros ((funs ~~ preds) ~~ (argss' ~~ rhss))) of
+      NONE =>
+        let val _ = tracing "error occured!" in ([], thy) end
     | SOME intr_ts =>
         if is_some (try (map (cterm_of thy)) intr_ts) then
           let
@@ -333,53 +356,59 @@
                   no_elim = false, no_ind = false, skip_mono = false, fork_mono = false}
                 (map (fn (s, T) =>
                   ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds)))
-                pnames
+                []
                 (map (fn x => (Attrib.empty_binding, x)) intr_ts)
                 []
               ||> Sign.restore_naming thy
             val prednames = map (fst o dest_Const) (#preds ind_result)
             (* val rewr_thms = map mk_rewr_eq ((distinct (op =) funs) ~~ (#preds ind_result)) *)
             (* add constants to my table *)
+            
             val specs = map (fn predname => (predname, filter (Predicate_Compile_Aux.is_intro predname)
               (#intrs ind_result))) prednames
+            (*
             val thy'' = Pred_Compile_Preproc.map (fold Symtab.update_new (consts ~~ prednames)) thy'
+            *)
+            
+            val thy'' = Fun_Pred.map
+              (fold Item_Net.update (map (apfst Logic.varify)
+                (distinct (op =) funs ~~ (#preds ind_result)))) thy'
+            (*val _ = print_specs thy'' specs*)
           in
             (specs, thy'')
           end
         else
           let
-            val _ = tracing "Introduction rules of function_predicate are not welltyped"
+            val _ = Output.tracing (
+            "Introduction rules of function_predicate are not welltyped: " ^
+              commas (map (Syntax.string_of_term_global thy) intr_ts))
           in ([], thy) end
   end
 
 fun rewrite_intro thy intro =
   let
-    fun lookup_pred (Const (name, T)) =
+    (*val _ = tracing ("Rewriting intro with registered mapping for: " ^
+      commas (Symtab.keys (Pred_Compile_Preproc.get thy)))*)
+    (*fun lookup_pred (Const (name, T)) =
       (case (Symtab.lookup (Pred_Compile_Preproc.get thy) name) of
-        SOME c => Const (c, pred_type T)
-      | NONE => error ("Function " ^ name ^ " is not inductified"))
-    | lookup_pred (Free (name, T)) = Free (name, T)
-    | lookup_pred _ = error "lookup function is not defined!"
-
-    fun get_nparams (Const (name, _)) =
-      the_default 0 (try (ind_package_get_nparams thy) name)
-    | get_nparams (Free _) = 0
-    | get_nparams t = error ("No parameters for " ^ (Syntax.string_of_term_global thy t))
-    
+        SOME c => SOME (Const (c, pred_type T))
+      | NONE => NONE)
+    | lookup_pred _ = NONE
+    *)
+    fun lookup_pred t = lookup thy (Fun_Pred.get thy) t
     val intro_t = (Logic.unvarify o prop_of) intro
     val (prems, concl) = Logic.strip_horn intro_t
     val frees = map fst (Term.add_frees intro_t [])
     fun rewrite prem names =
       let
+        (*val _ = tracing ("Rewriting premise " ^ Syntax.string_of_term_global thy prem ^ "...")*)
         val t = (HOLogic.dest_Trueprop prem)
         val (lit, mk_lit) = case try HOLogic.dest_not t of
             SOME t => (t, HOLogic.mk_not)
           | NONE => (t, I)
-        val (P, args) = (strip_comb lit) 
+        val (P, args) = (strip_comb lit)
       in
-        folds_map (
-          fn t => if (is_funtype (fastype_of t)) then (fn x => [(t, x)])
-            else mk_prems thy (lookup_pred, get_nparams) t) args (names, [])
+        folds_map (mk_prems thy lookup_pred) args (names, [])
         |> map (fn (resargs, (names', prems')) =>
           let
             val prem' = HOLogic.mk_Trueprop (mk_lit (list_comb (P, resargs)))
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_pred.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -7,23 +7,85 @@
 signature PREDICATE_COMPILE_PRED =
 sig
   (* preprocesses an equation to a set of intro rules; defines new constants *)
-  (*
-  val preprocess_pred_equation : thm -> theory -> thm list * theory
-  *)
-  val preprocess : string -> theory -> (thm list list * theory) 
-  (* output is the term list of clauses of an unknown predicate *)
-  val preprocess_term : term -> theory -> (term list * theory)
-  
-  (*val rewrite : thm -> thm*)
-  
+  val preprocess : Predicate_Compile_Aux.options -> (string * thm list) -> theory
+    -> ((string * thm list) list * theory) 
+  val flat_higher_order_arguments : ((string * thm list) list * theory)
+    -> ((string * thm list) list * ((string * thm list) list * theory))
 end;
 
-(* : PREDICATE_COMPILE_PREPROC_PRED *)  (* FIXME *)
-structure Predicate_Compile_Pred =
+
+structure Predicate_Compile_Pred : PREDICATE_COMPILE_PRED =
 struct
 
 open Predicate_Compile_Aux
 
+
+fun datatype_names_of_case_name thy case_name =
+  map (#1 o #2) (#descr (the (Datatype_Data.info_of_case thy case_name)))
+
+fun make_case_rewrites new_type_names descr sorts thy =
+  let
+    val case_combs = Datatype_Prop.make_case_combs new_type_names descr sorts thy "f";
+    fun make comb =
+      let
+        val Type ("fun", [T, T']) = fastype_of comb;
+        val (Const (case_name, _), fs) = strip_comb comb
+        val used = Term.add_tfree_names comb []
+        val U = TFree (Name.variant used "'t", HOLogic.typeS)
+        val x = Free ("x", T)
+        val f = Free ("f", T' --> U)
+        fun apply_f f' =
+          let
+            val Ts = binder_types (fastype_of f')
+            val bs = map Bound ((length Ts - 1) downto 0)
+          in
+            fold (curry absdummy) (rev Ts) (f $ (list_comb (f', bs)))
+          end
+        val fs' = map apply_f fs
+        val case_c' = Const (case_name, (map fastype_of fs') @ [T] ---> U)
+      in
+        HOLogic.mk_eq (f $ (comb $ x), list_comb (case_c', fs') $ x)
+      end
+  in
+    map make case_combs
+  end
+
+fun case_rewrites thy Tcon =
+  let
+    val info = Datatype.the_info thy Tcon
+    val descr = #descr info
+    val sorts = #sorts info
+    val typ_names = the_default [Tcon] (#alt_names info)
+  in
+    map (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop)
+      (make_case_rewrites typ_names [descr] sorts thy)
+  end
+
+fun instantiated_case_rewrites thy Tcon =
+  let
+    val rew_ths = case_rewrites thy Tcon
+    val ctxt = ProofContext.init thy
+    fun instantiate th =
+    let
+      val f = (fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of th))))))
+      val Type ("fun", [uninst_T, uninst_T']) = fastype_of f
+      val ([tname, tname', uname, yname], ctxt') = Variable.add_fixes ["'t", "'t'", "'u", "y"] ctxt
+      val T = TFree (tname, HOLogic.typeS)
+      val T' = TFree (tname', HOLogic.typeS)
+      val U = TFree (uname, HOLogic.typeS)
+      val y = Free (yname, U)
+      val f' = absdummy (U --> T', Bound 0 $ y)
+      val th' = Thm.certify_instantiate
+        ([(dest_TVar uninst_T, U --> T'), (dest_TVar uninst_T', T')],
+         [((fst (dest_Var f), (U --> T') --> T'), f')]) th
+      val [th'] = Variable.export ctxt' ctxt [th']
+   in
+     th'
+   end
+ in
+   map instantiate rew_ths
+ end
+
 fun is_compound ((Const ("Not", _)) $ t) =
     error "is_compound: Negation should not occur; preprocessing is defect"
   | is_compound ((Const ("Ex", _)) $ _) = true
@@ -35,6 +97,7 @@
 fun flatten constname atom (defs, thy) =
   if is_compound atom then
     let
+      val atom = Envir.beta_norm (Pattern.eta_long [] atom)
       val constname = Name.variant (map (Long_Name.base_name o fst) defs)
         ((Long_Name.base_name constname) ^ "_aux")
       val full_constname = Sign.full_bname thy constname
@@ -50,7 +113,82 @@
       (lhs, ((full_constname, [definition]) :: defs, thy'))
     end
   else
-    (atom, (defs, thy))
+    (case (fst (strip_comb atom)) of
+      (Const (@{const_name If}, _)) => let
+          val if_beta = @{lemma "(if c then x else y) z = (if c then x z else y z)" by simp}
+          val atom' = MetaSimplifier.rewrite_term thy
+            (map (fn th => th RS @{thm eq_reflection}) [@{thm if_bool_eq_disj}, if_beta]) [] atom
+          val _ = assert (not (atom = atom'))
+        in
+          flatten constname atom' (defs, thy)
+        end
+    | _ =>  
+      if (has_split_thm thy (fst (strip_comb atom))) then
+        let
+          val (f, args) = strip_comb atom
+          val split_thm = prepare_split_thm (ProofContext.init thy) (the (find_split_thm' thy f))
+          (* TODO: contextify things - this line is to unvarify the split_thm *)
+          (*val ((_, [isplit_thm]), _) = Variable.import true [split_thm] (ProofContext.init thy)*)
+          val (assms, concl) = Logic.strip_horn (Thm.prop_of split_thm)
+          val (P, [split_t]) = strip_comb (HOLogic.dest_Trueprop concl) 
+          val Tcons = datatype_names_of_case_name thy (fst (dest_Const f))
+          val ths = maps (instantiated_case_rewrites thy) Tcons
+          val atom = MetaSimplifier.rewrite_term thy
+            (map (fn th => th RS @{thm eq_reflection}) ths) [] atom
+          val (f, args) = strip_comb atom
+          val subst = Pattern.match thy (split_t, atom) (Vartab.empty, Vartab.empty)
+          val (_, split_args) = strip_comb split_t
+          val match = split_args ~~ args
+          
+          (*
+          fun mk_prems_of_assm assm =
+            let
+              val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
+              val names = [] (* TODO *)
+              val var_names = Name.variant_list names (map fst vTs)
+              val vars = map Free (var_names ~~ (map snd vTs))
+              val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
+              val (HOLogic.dest_eq (HOLogic.dest_Trueprop prem))
+              val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+            in
+              (*mk_prems' inner_t (var_names @ names, prems' @ prems)*) error "asda"
+            end
+          *)
+          val names = Term.add_free_names atom []
+          val frees = map Free (Term.add_frees atom [])
+          val constname = Name.variant (map (Long_Name.base_name o fst) defs)
+            ((Long_Name.base_name constname) ^ "_aux")
+          val full_constname = Sign.full_bname thy constname
+          val constT = map fastype_of frees ---> HOLogic.boolT
+          val lhs = list_comb (Const (full_constname, constT), frees)
+          fun new_def assm =
+            let
+              val (vTs, assm') = strip_all (Envir.beta_norm (Envir.subst_term subst assm))
+              val var_names = Name.variant_list names (map fst vTs)
+              val vars = map Free (var_names ~~ (map snd vTs))
+              val (prems', pre_res) = Logic.strip_horn (subst_bounds (rev vars, assm'))
+              fun mk_subst prem =
+                let
+                  val (Free (x, T), r) = HOLogic.dest_eq (HOLogic.dest_Trueprop prem)
+                in
+                  ((x, T), r)
+                end
+              val subst = map mk_subst prems'
+              val (_, [inner_t]) = strip_comb (HOLogic.dest_Trueprop pre_res)
+              val def = Logic.mk_equals (lhs, inner_t)
+            in
+              Envir.expand_term_frees subst def
+            end
+         val new_defs = map new_def assms
+         val (definition, thy') = thy
+          |> Sign.add_consts_i [(Binding.name constname, constT, NoSyn)]
+          |> PureThy.add_axioms (map_index
+              (fn (i, t) => ((Binding.name (constname ^ "_def" ^ string_of_int i), t), [])) new_defs)
+        in
+          (lhs, ((full_constname, definition) :: defs, thy'))
+        end
+      else
+        (atom, (defs, thy)))
 
 fun flatten_intros constname intros thy =
   let
@@ -107,30 +245,60 @@
 
 val rewrite =
   Simplifier.simplify (HOL_basic_ss addsimps [@{thm Ball_def}, @{thm Bex_def}])
-  #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}]) 
+  #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
   #> Conv.fconv_rule nnf_conv 
   #> Simplifier.simplify (HOL_basic_ss addsimps [@{thm ex_disj_distrib}])
 
-val rewrite_intros =
-(*  Simplifier.simplify (HOL_basic_ss addsimps @{thms HOL.simp_thms(9)}) *)
-  Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm not_not}])
-  
-fun preprocess (constname, specs) thy =
+fun split_conjs thy t =
+  let 
+    fun split_conjunctions (Const (@{const_name "op &"}, _) $ t1 $ t2) =
+    (split_conjunctions t1) @ (split_conjunctions t2)
+    | split_conjunctions t = [t]
+  in
+    map HOLogic.mk_Trueprop (split_conjunctions (HOLogic.dest_Trueprop t))
+  end
+
+fun rewrite_intros thy =
+  Simplifier.full_simplify (HOL_basic_ss addsimps [@{thm all_not_ex}])
+  #> Simplifier.full_simplify (HOL_basic_ss addsimps @{thms bool_simps} addsimps @{thms nnf_simps})
+  #> map_term thy (maps_premises (split_conjs thy))
+
+fun print_specs options thy msg ths =
+  if show_intermediate_results options then
+    (tracing (msg); tracing (commas (map (Display.string_of_thm_global thy) ths)))
+  else
+    ()
+(*
+fun split_cases thy th =
   let
-    val ctxt = ProofContext.init thy
+    
+  in
+    map_term thy th
+  end
+*)
+fun preprocess options (constname, specs) thy =
+(*  case Predicate_Compile_Data.processed_specs thy constname of
+    SOME specss => (specss, thy)
+  | NONE =>*)
+    let
+      val ctxt = ProofContext.init thy
       val intros =
-      if forall is_pred_equation specs then 
-        introrulify thy (map rewrite specs)
-      else if forall (is_intro constname) specs then
-        map rewrite_intros specs
-      else
-        error ("unexpected specification for constant " ^ quote constname ^ ":\n"
-          ^ commas (map (quote o Display.string_of_thm_global thy) specs))
-    val (intros', (local_defs, thy')) = flatten_intros constname intros thy
-    val (intross, thy'') = fold_map preprocess local_defs thy'
-  in
-    ((constname, intros') :: flat intross,thy'')
-  end;
+        if forall is_pred_equation specs then 
+          map (map_term thy (maps_premises (split_conjs thy))) (introrulify thy (map rewrite specs))
+        else if forall (is_intro constname) specs then
+          map (rewrite_intros thy) specs
+        else
+          error ("unexpected specification for constant " ^ quote constname ^ ":\n"
+            ^ commas (map (quote o Display.string_of_thm_global thy) specs))
+      val _ = print_specs options thy "normalized intros" intros
+      (*val intros = maps (split_cases thy) intros*)
+      val (intros', (local_defs, thy')) = flatten_intros constname intros thy
+      val (intross, thy'') = fold_map (preprocess options) local_defs thy'
+      val full_spec = (constname, intros') :: flat intross
+      (*val thy''' = Predicate_Compile_Data.store_processed_specs (constname, full_spec) thy''*)
+    in
+      (full_spec, thy'')
+    end;
 
 fun preprocess_term t thy = error "preprocess_pred_term: to implement" 
 
@@ -166,7 +334,8 @@
           else
             (arg, (new_defs, thy))
         
-        val (args', (new_defs', thy')) = fold_map replace_abs_arg args (new_defs, thy)
+        val (args', (new_defs', thy')) = fold_map replace_abs_arg
+          (map Envir.beta_eta_contract args) (new_defs, thy)
       in
         (list_comb (pred, args'), (new_defs', thy'))
       end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -6,10 +6,18 @@
 
 signature PREDICATE_COMPILE_QUICKCHECK =
 sig
-  val quickcheck : Proof.context -> term -> int -> term list option
+  (*val quickcheck : Proof.context -> term -> int -> term list option*)
   val test_ref :
     ((unit -> int -> int -> int * int -> term list DSequence.dseq * (int * int)) option) Unsynchronized.ref
   val tracing : bool Unsynchronized.ref;
+  val quickcheck_compile_term : bool -> bool -> Proof.context -> term -> int -> term list option
+(*  val test_term : Proof.context -> bool -> int -> int -> int -> int -> term -> *)
+  val quiet : bool Unsynchronized.ref;
+  val nrandom : int Unsynchronized.ref;
+  val depth : int Unsynchronized.ref;
+  val debug : bool Unsynchronized.ref;
+  val function_flattening : bool Unsynchronized.ref;
+  val no_higher_order_predicate : string list Unsynchronized.ref;
 end;
 
 structure Predicate_Compile_Quickcheck : PREDICATE_COMPILE_QUICKCHECK =
@@ -24,21 +32,106 @@
 
 val target = "Quickcheck"
 
+val quiet = Unsynchronized.ref false;
+
+val nrandom = Unsynchronized.ref 2;
+
+val depth = Unsynchronized.ref 8;
+
+val debug = Unsynchronized.ref false;
+val function_flattening = Unsynchronized.ref true;
+
+
+val no_higher_order_predicate = Unsynchronized.ref [];
+
 val options = Options {
   expected_modes = NONE,
   proposed_modes = NONE,
   proposed_names = [],
+  show_steps = false,
+  show_intermediate_results = false,
+  show_proof_trace = false,
+  show_modes = false,
+  show_mode_inference = false,
+  show_compilation = false,
+  show_caught_failures = false,
+  skip_proof = false,
+  compilation = Random,
+  inductify = true,
+  function_flattening = true,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
+  no_topmost_reordering = true
+}
+
+val debug_options = Options {
+  expected_modes = NONE,
+  proposed_modes = NONE,
+  proposed_names = [],
   show_steps = true,
   show_intermediate_results = true,
   show_proof_trace = false,
-  show_modes = false,
-  show_mode_inference = false,
+  show_modes = true,
+  show_mode_inference = true,
   show_compilation = false,
+  show_caught_failures = true,
   skip_proof = false,
   compilation = Random,
-  inductify = false
+  inductify = true,
+  function_flattening = true,
+  fail_safe_function_flattening = false,
+  no_higher_order_predicate = [],
+  no_topmost_reordering = true
 }
 
+
+fun set_function_flattening b
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f, 
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re}) =
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = b,
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re})
+
+fun set_fail_safe_function_flattening b
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f, 
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re}) =
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f,
+    fail_safe_function_flattening = b, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re})
+
+fun set_no_higher_order_predicate ss
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f, 
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = no_ho,
+    no_topmost_reordering = re}) =
+  (Options { expected_modes = e_m, proposed_modes = p_m, proposed_names = p_n, show_steps = s_s,
+    show_intermediate_results = s_ir, show_proof_trace = s_pt, show_modes = s_m,
+    show_mode_inference = s_mi, show_compilation = s_c, show_caught_failures = s_cf, skip_proof = s_p,
+    compilation = c, inductify = i, function_flattening = f_f,
+    fail_safe_function_flattening = fs_ff, no_higher_order_predicate = ss, no_topmost_reordering = re})
+
+
+fun get_options () = 
+  set_no_higher_order_predicate (!no_higher_order_predicate)
+    (set_function_flattening (!function_flattening)
+      (if !debug then debug_options else options))
+
 fun dest_compfuns (Predicate_Compile_Core.CompilationFuns funs) = funs
 val mk_predT = #mk_predT (dest_compfuns Predicate_Compile_Core.pred_compfuns)
 val mk_randompredT = #mk_predT (dest_compfuns Predicate_Compile_Core.randompred_compfuns)
@@ -63,13 +156,15 @@
 
 fun strip_horn A = (strip_imp_prems A, strip_imp_concl A);
 
-fun quickcheck ctxt t =
+fun cpu_time description f =
   let
-    (*val () =
-      if !tracing then
-        tracing ("Starting quickcheck with " ^ (Syntax.string_of_term ctxt t))
-      else
-        ()*)
+    val start = start_timing ()
+    val result = Exn.capture f ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, (description, time)) end
+
+fun compile_term options ctxt t =
+  let
     val ctxt' = ProofContext.theory (Context.copy_thy) ctxt
     val thy = (ProofContext.theory_of ctxt') 
     val (vs, t') = strip_abs t
@@ -82,44 +177,73 @@
     val thy1 = Sign.add_consts_i [(Binding.name constname, constT, NoSyn)] thy
     val const = Const (full_constname, constT)
     val t = Logic.list_implies
-      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),                               
+      (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]),
        HOLogic.mk_Trueprop (list_comb (Const (full_constname, constT), map Free vs')))
     val tac = fn _ => Skip_Proof.cheat_tac thy1
     val intro = Goal.prove (ProofContext.init thy1) (map fst vs') [] t tac
-    (*val _ = tracing (Display.string_of_thm ctxt' intro)*)
-    val thy2 = (*Output.cond_timeit (!Quickcheck.timing) "predicate intros"
-      (fn () => *)(Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1)
-    val thy3 = (*Output.cond_timeit (!Quickcheck.timing) "predicate preprocessing"
-        (fn () =>*) (Predicate_Compile.preprocess options const thy2)
-    val thy4 = Output.cond_timeit (!Quickcheck.timing) "random_dseq compilation"
+    val thy2 = Context.theory_map (Predicate_Compile_Alternative_Defs.add_thm intro) thy1
+    val (thy3, preproc_time) =  cpu_time "predicate preprocessing"
+        (fn () => Predicate_Compile.preprocess options const thy2)
+    val (thy4, core_comp_time) = cpu_time "random_dseq core compilation"
         (fn () => Predicate_Compile_Core.add_random_dseq_equations options [full_constname] thy3)
-    (*val depth_limited_modes = Predicate_Compile_Core.modes_of Depth_Limited thy'' full_constname*)
-    val modes = Predicate_Compile_Core.modes_of Random_DSeq thy4 full_constname
+    val _ = Predicate_Compile_Core.print_all_modes Pos_Random_DSeq thy4
+    val modes = Predicate_Compile_Core.modes_of Pos_Random_DSeq thy4 full_constname
     val output_mode = fold_rev (curry Fun) (map (K Output) (binder_types constT)) Bool
     val prog =
       if member eq_mode modes output_mode then
         let
-          val name = Predicate_Compile_Core.function_name_of Random_DSeq thy4 full_constname output_mode
+          val name = Predicate_Compile_Core.function_name_of Pos_Random_DSeq thy4
+            full_constname (true, output_mode)
           val T = (mk_randompredT (HOLogic.mk_tupleT (map snd vs')))
         in
           Const (name, T)
         end
-      (*else if member (op =) depth_limited_modes ([], []) then
-        let
-          val name = Predicate_Compile_Core.depth_limited_function_name_of thy'' full_constname ([], [])
-          val T = @{typ code_numeral} --> (mk_predT (HOLogic.mk_tupleT (map snd vs')))
-        in lift_pred (Const (name, T) $ Bound 0) end*)
-      else error "Predicate Compile Quickcheck failed"
+      else error ("Predicate Compile Quickcheck failed: " ^ commas (map string_of_mode modes))
     val qc_term = mk_bind (prog,
       mk_split_lambda (map Free vs') (mk_return (HOLogic.mk_list @{typ term}
       (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
     val compilation =
-      Code_Eval.eval NONE ("Predicate_Compile_Quickcheck.test_ref", test_ref)
+      Code_Eval.eval (SOME target) ("Predicate_Compile_Quickcheck.test_ref", test_ref)
         (fn proc => fn g => fn n => fn size => fn s => g n size s |>> (DSequence.map o map) proc)
         thy4 qc_term []
   in
-    (fn size =>
-      Option.map fst (DSequence.yield (compilation size size |> Random_Engine.run) size true))
+    (fn size => fn nrandom => fn depth =>
+      Option.map fst (DSequence.yield (compilation nrandom size |> Random_Engine.run) depth true))
+  end
+
+fun try_upto quiet f i =
+  let
+    fun try' j =
+      if j <= i then
+        let
+          val _ = priority ("Executing depth " ^ string_of_int j)
+        in
+          case f j handle Match => (if quiet then ()
+             else warning "Exception Match raised during quickcheck"; NONE)
+          of NONE => try' (j + 1) | SOME q => SOME q
+        end
+      else
+        NONE
+  in
+    try' 0
+  end
+
+(* quickcheck interface functions *)
+
+fun compile_term' options ctxt t =
+  let
+    val c = compile_term options ctxt t
+  in
+    (fn size => try_upto (!quiet) (c size (!nrandom)) (!depth))
+  end
+
+fun quickcheck_compile_term function_flattening fail_safe_function_flattening ctxt t =
+  let
+     val options =
+       set_fail_safe_function_flattening fail_safe_function_flattening
+         (set_function_flattening function_flattening (get_options ()))
+  in
+    compile_term' options ctxt t
   end
 
 end;
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -37,7 +37,9 @@
   val equiv_rules_get: Proof.context -> thm list
   val equiv_rules_add: attribute
   val rsp_rules_get: Proof.context -> thm list
+  val rsp_rules_add: attribute
   val prs_rules_get: Proof.context -> thm list
+  val prs_rules_add: attribute
   val id_simps_get: Proof.context -> thm list
   val quotient_rules_get: Proof.context -> thm list
   val quotient_rules_add: attribute
@@ -241,6 +243,7 @@
    val description = "Respectfulness theorems.")
 
 val rsp_rules_get = RspRules.get
+val rsp_rules_add = RspRules.add
 
 (* preservation theorems *)
 structure PrsRules = Named_Thms
@@ -248,6 +251,7 @@
    val description = "Preservation theorems.")
 
 val prs_rules_get = PrsRules.get
+val prs_rules_add = PrsRules.add
 
 (* id simplification theorems *)
 structure IdSimps = Named_Thms
--- a/src/HOL/Typerep.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/Typerep.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -70,7 +70,8 @@
 
 add_typerep @{type_name fun}
 #> Typedef.interpretation ensure_typerep
-#> Code.type_interpretation (ensure_typerep o fst)
+#> Code.datatype_interpretation (ensure_typerep o fst)
+#> Code.abstype_interpretation (ensure_typerep o fst)
 
 end
 *}
--- a/src/HOL/W0/README.html	Tue Feb 23 14:44:24 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<!-- $Id$ -->
-
-<HTML>
-
-<HEAD>
-  <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
-  <TITLE>HOL/W0/README</TITLE>
-</HEAD>
-
-<BODY>
-
-<H1>Type Inference for MiniML (without <tt>let</tt>)</H1>
-
-This theory defines the type inference rules and the type inference algorithm
-<em>W</em> for simply-typed lambda terms due to Milner. It proves the
-soundness and completeness of <em>W</em> w.r.t. to the rules. An optimized
-version <em>I</em> is shown to implement <em>W</em>.
-
-<P>
-
-A report describing the theory is found here:<br>
-<A HREF = "http://www4.informatik.tu-muenchen.de/~nipkow/pubs/tphol96.html">
-Formal Verification of Algorithm W: The Monomorphic Case</A>.
-
-<P>
-
-<B>NOTE:</B> This theory has been superseded by a more recent development
-which formalizes type inference for a language including <tt>let</tt>. For
-details click <A HREF="../MiniML/index.html">here</A>.
-</BODY>
-</HTML>
--- a/src/HOL/W0/ROOT.ML	Tue Feb 23 14:44:24 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-use_thys ["W0"];
--- a/src/HOL/W0/W0.thy	Tue Feb 23 14:44:24 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,925 +0,0 @@
-(*  Title:      HOL/W0/W0.thy
-    ID:         $Id$
-    Author:     Dieter Nazareth, Tobias Nipkow, Thomas Stauner, Markus Wenzel
-*)
-
-theory W0
-imports Main
-begin
-
-section {* Universal error monad *}
-
-datatype 'a maybe = Ok 'a | Fail
-
-definition
-  bind :: "'a maybe \<Rightarrow> ('a \<Rightarrow> 'b maybe) \<Rightarrow> 'b maybe"    (infixl "\<bind>" 60) where
-  "m \<bind> f = (case m of Ok r \<Rightarrow> f r | Fail \<Rightarrow> Fail)"
-
-syntax
-  "_bind" :: "patterns \<Rightarrow> 'a maybe \<Rightarrow> 'b \<Rightarrow> 'c"    ("(_ := _;//_)" 0)
-translations
-  "P := E; F" == "E \<bind> (\<lambda>P. F)"
-
-lemma bind_Ok [simp]: "(Ok s) \<bind> f = (f s)"
-  by (simp add: bind_def)
-
-lemma bind_Fail [simp]: "Fail \<bind> f = Fail"
-  by (simp add: bind_def)
-
-lemma split_bind:
-    "P (res \<bind> f) = ((res = Fail \<longrightarrow> P Fail) \<and> (\<forall>s. res = Ok s \<longrightarrow> P (f s)))"
-  by (induct res) simp_all
-
-lemma split_bind_asm:
-  "P (res \<bind> f) = (\<not> (res = Fail \<and> \<not> P Fail \<or> (\<exists>s. res = Ok s \<and> \<not> P (f s))))"
-  by (simp split: split_bind)
-
-lemmas bind_splits = split_bind split_bind_asm
-
-lemma bind_eq_Fail [simp]:
-  "((m \<bind> f) = Fail) = ((m = Fail) \<or> (\<exists>p. m = Ok p \<and> f p = Fail))"
-  by (simp split: split_bind)
-
-lemma rotate_Ok: "(y = Ok x) = (Ok x = y)"
-  by (rule eq_sym_conv)
-
-
-section {* MiniML-types and type substitutions *}
-
-axclass type_struct \<subseteq> type
-  -- {* new class for structures containing type variables *}
-
-datatype "typ" = TVar nat | TFun "typ" "typ"    (infixr "->" 70)
-  -- {* type expressions *}
-
-types subst = "nat => typ"
-  -- {* type variable substitution *}
-
-instance "typ" :: type_struct ..
-instance list :: (type_struct) type_struct ..
-instance "fun" :: (type, type_struct) type_struct ..
-
-
-subsection {* Substitutions *}
-
-consts
-  app_subst :: "subst \<Rightarrow> 'a::type_struct \<Rightarrow> 'a::type_struct"    ("$")
-  -- {* extension of substitution to type structures *}
-primrec (app_subst_typ)
-  app_subst_TVar: "$s (TVar n) = s n"
-  app_subst_Fun: "$s (t1 -> t2) = $s t1 -> $s t2"
-
-defs (overloaded)
-  app_subst_list: "$s \<equiv> map ($s)"
-
-consts
-  free_tv :: "'a::type_struct \<Rightarrow> nat set"
-  -- {* @{text "free_tv s"}: the type variables occuring freely in the type structure @{text s} *}
-
-primrec (free_tv_typ)
-  "free_tv (TVar m) = {m}"
-  "free_tv (t1 -> t2) = free_tv t1 \<union> free_tv t2"
-
-primrec (free_tv_list)
-  "free_tv [] = {}"
-  "free_tv (x # xs) = free_tv x \<union> free_tv xs"
-
-definition
-  dom :: "subst \<Rightarrow> nat set" where
-  "dom s = {n. s n \<noteq> TVar n}"
-  -- {* domain of a substitution *}
-
-definition
-  cod :: "subst \<Rightarrow> nat set" where
-  "cod s = (\<Union>m \<in> dom s. free_tv (s m))"
-  -- {* codomain of a substitutions: the introduced variables *}
-
-defs (overloaded)
-  free_tv_subst: "free_tv s \<equiv> dom s \<union> cod s"
-
-text {*
-  @{text "new_tv s n"} checks whether @{text n} is a new type variable
-  wrt.\ a type structure @{text s}, i.e.\ whether @{text n} is greater
-  than any type variable occuring in the type structure.
-*}
-
-definition
-  new_tv :: "nat \<Rightarrow> 'a::type_struct \<Rightarrow> bool" where
-  "new_tv n ts = (\<forall>m. m \<in> free_tv ts \<longrightarrow> m < n)"
-
-
-subsubsection {* Identity substitution *}
-
-definition
-  id_subst :: subst where
-  "id_subst = (\<lambda>n. TVar n)"
-
-lemma app_subst_id_te [simp]:
-  "$id_subst = (\<lambda>t::typ. t)"
-  -- {* application of @{text id_subst} does not change type expression *}
-proof
-  fix t :: "typ"
-  show "$id_subst t = t"
-    by (induct t) (simp_all add: id_subst_def)
-qed
-
-lemma app_subst_id_tel [simp]: "$id_subst = (\<lambda>ts::typ list. ts)"
-  -- {* application of @{text id_subst} does not change list of type expressions *}
-proof
-  fix ts :: "typ list"
-  show "$id_subst ts = ts"
-    by (induct ts) (simp_all add: app_subst_list)
-qed
-
-lemma o_id_subst [simp]: "$s o id_subst = s"
-  by (rule ext) (simp add: id_subst_def)
-
-lemma dom_id_subst [simp]: "dom id_subst = {}"
-  by (simp add: dom_def id_subst_def)
-
-lemma cod_id_subst [simp]: "cod id_subst = {}"
-  by (simp add: cod_def)
-
-lemma free_tv_id_subst [simp]: "free_tv id_subst = {}"
-  by (simp add: free_tv_subst)
-
-
-lemma cod_app_subst [simp]:
-  assumes free: "v \<in> free_tv (s n)"
-    and neq: "v \<noteq> n"
-  shows "v \<in> cod s"
-proof -
-  have "s n \<noteq> TVar n"
-  proof
-    assume "s n = TVar n"
-    with free have "v = n" by simp
-    with neq show False ..
-  qed
-  with free show ?thesis
-    by (auto simp add: dom_def cod_def)
-qed
-
-lemma subst_comp_te: "$g ($f t :: typ) = $(\<lambda>x. $g (f x)) t"
-  -- {* composition of substitutions *}
-  by (induct t) simp_all
-
-lemma subst_comp_tel: "$g ($f ts :: typ list) = $(\<lambda>x. $g (f x)) ts"
-  by (induct ts) (simp_all add: app_subst_list subst_comp_te)
-
-
-lemma app_subst_Nil [simp]: "$s [] = []"
-  by (simp add: app_subst_list)
-
-lemma app_subst_Cons [simp]: "$s (t # ts) = ($s t) # ($s ts)"
-  by (simp add: app_subst_list)
-
-lemma new_tv_TVar [simp]: "new_tv n (TVar m) = (m < n)"
-  by (simp add: new_tv_def)
-
-lemma new_tv_Fun [simp]:
-  "new_tv n (t1 -> t2) = (new_tv n t1 \<and> new_tv n t2)"
-  by (auto simp add: new_tv_def)
-
-lemma new_tv_Nil [simp]: "new_tv n []"
-  by (simp add: new_tv_def)
-
-lemma new_tv_Cons [simp]: "new_tv n (t # ts) = (new_tv n t \<and> new_tv n ts)"
-  by (auto simp add: new_tv_def)
-
-lemma new_tv_id_subst [simp]: "new_tv n id_subst"
-  by (simp add: id_subst_def new_tv_def free_tv_subst dom_def cod_def)
-
-lemma new_tv_subst:
-  "new_tv n s =
-    ((\<forall>m. n \<le> m \<longrightarrow> s m = TVar m) \<and>
-     (\<forall>l. l < n \<longrightarrow> new_tv n (s l)))"
-  apply (unfold new_tv_def)
-  apply (tactic "safe_tac HOL_cs")
-  -- {* @{text \<Longrightarrow>} *}
-    apply (tactic {* fast_tac (HOL_cs addDs [@{thm leD}] addss (@{simpset}
-      addsimps [thm "free_tv_subst", thm "dom_def"])) 1 *})
-   apply (subgoal_tac "m \<in> cod s \<or> s l = TVar l")
-    apply (tactic "safe_tac HOL_cs")
-     apply (tactic {* fast_tac (HOL_cs addDs [UnI2] addss (@{simpset}
-       addsimps [thm "free_tv_subst"])) 1 *})
-    apply (drule_tac P = "\<lambda>x. m \<in> free_tv x" in subst, assumption)
-    apply simp
-    apply (unfold free_tv_subst cod_def dom_def)
-    apply clarsimp
-  apply safe
-  apply metis
-  apply (metis linorder_not_less)+
-  done
-
-lemma new_tv_list: "new_tv n x = (\<forall>y \<in> set x. new_tv n y)"
-  by (induct x) simp_all
-
-lemma subst_te_new_tv [simp]:
-  "new_tv n (t::typ) \<Longrightarrow> $(\<lambda>x. if x = n then t' else s x) t = $s t"
-  -- {* substitution affects only variables occurring freely *}
-  by (induct t) simp_all
-
-lemma subst_tel_new_tv [simp]:
-  "new_tv n (ts::typ list) \<Longrightarrow> $(\<lambda>x. if x = n then t else s x) ts = $s ts"
-  by (induct ts) simp_all
-
-lemma new_tv_le: "n \<le> m \<Longrightarrow> new_tv n (t::typ) \<Longrightarrow> new_tv m t"
-  -- {* all greater variables are also new *}
-proof (induct t)
-  case (TVar n)
-  then show ?case by (auto intro: less_le_trans)
-next
-  case TFun
-  then show ?case by simp
-qed
-
-lemma [simp]: "new_tv n t \<Longrightarrow> new_tv (Suc n) (t::typ)"
-  by (rule lessI [THEN less_imp_le [THEN new_tv_le]])
-
-lemma new_tv_list_le:
-  assumes "n \<le> m"
-  shows "new_tv n (ts::typ list) \<Longrightarrow> new_tv m ts"
-proof (induct ts)
-  case Nil
-  then show ?case by simp
-next
-  case Cons
-  with `n \<le> m` show ?case by (auto intro: new_tv_le)
-qed
-
-lemma [simp]: "new_tv n ts \<Longrightarrow> new_tv (Suc n) (ts::typ list)"
-  by (rule lessI [THEN less_imp_le [THEN new_tv_list_le]])
-
-lemma new_tv_subst_le: "n \<le> m \<Longrightarrow> new_tv n (s::subst) \<Longrightarrow> new_tv m s"
-  apply (simp add: new_tv_subst)
-  apply clarify
-  apply (rule_tac P = "l < n" and Q = "n <= l" in disjE)
-    apply clarify
-    apply (simp_all add: new_tv_le)
-  done
-
-lemma [simp]: "new_tv n s \<Longrightarrow> new_tv (Suc n) (s::subst)"
-  by (rule lessI [THEN less_imp_le [THEN new_tv_subst_le]])
-
-lemma new_tv_subst_var:
-    "n < m \<Longrightarrow> new_tv m (s::subst) \<Longrightarrow> new_tv m (s n)"
-  -- {* @{text new_tv} property remains if a substitution is applied *}
-  by (simp add: new_tv_subst)
-
-lemma new_tv_subst_te [simp]:
-    "new_tv n s \<Longrightarrow> new_tv n (t::typ) \<Longrightarrow> new_tv n ($s t)"
-  by (induct t) (auto simp add: new_tv_subst)
-
-lemma new_tv_subst_tel [simp]:
-    "new_tv n s \<Longrightarrow> new_tv n (ts::typ list) \<Longrightarrow> new_tv n ($s ts)"
-  by (induct ts) (fastsimp simp add: new_tv_subst)+
-
-lemma new_tv_Suc_list: "new_tv n ts --> new_tv (Suc n) (TVar n # ts)"
-  -- {* auxilliary lemma *}
-  by (simp add: new_tv_list)
-
-lemma new_tv_subst_comp_1 [simp]:
-    "new_tv n (s::subst) \<Longrightarrow> new_tv n r \<Longrightarrow> new_tv n ($r o s)"
-  -- {* composition of substitutions preserves @{text new_tv} proposition *}
-  by (simp add: new_tv_subst)
-
-lemma new_tv_subst_comp_2 [simp]:
-    "new_tv n (s::subst) \<Longrightarrow> new_tv n r \<Longrightarrow> new_tv n (\<lambda>v. $r (s v))"
-  by (simp add: new_tv_subst)
-
-lemma new_tv_not_free_tv [simp]: "new_tv n ts \<Longrightarrow> n \<notin> free_tv ts"
-  -- {* new type variables do not occur freely in a type structure *}
-  by (auto simp add: new_tv_def)
-
-lemma ftv_mem_sub_ftv_list [simp]:
-    "(t::typ) \<in> set ts \<Longrightarrow> free_tv t \<subseteq> free_tv ts"
-  by (induct ts) auto
-
-text {*
-  If two substitutions yield the same result if applied to a type
-  structure the substitutions coincide on the free type variables
-  occurring in the type structure.
-*}
-
-lemma eq_subst_te_eq_free:
-    "$s1 (t::typ) = $s2 t \<Longrightarrow> n \<in> free_tv t \<Longrightarrow> s1 n = s2 n"
-  by (induct t) auto
-
-lemma eq_free_eq_subst_te:
-    "(\<forall>n. n \<in> free_tv t --> s1 n = s2 n) \<Longrightarrow> $s1 (t::typ) = $s2 t"
-  by (induct t) auto
-
-lemma eq_subst_tel_eq_free:
-    "$s1 (ts::typ list) = $s2 ts \<Longrightarrow> n \<in> free_tv ts \<Longrightarrow> s1 n = s2 n"
-  by (induct ts) (auto intro: eq_subst_te_eq_free)
-
-lemma eq_free_eq_subst_tel:
-    "(\<forall>n. n \<in> free_tv ts --> s1 n = s2 n) \<Longrightarrow> $s1 (ts::typ list) = $s2 ts"
-  by (induct ts) (auto intro: eq_free_eq_subst_te)
-
-text {*
-  \medskip Some useful lemmas.
-*}
-
-lemma codD: "v \<in> cod s \<Longrightarrow> v \<in> free_tv s"
-  by (simp add: free_tv_subst)
-
-lemma not_free_impl_id: "x \<notin> free_tv s \<Longrightarrow> s x = TVar x"
-  by (simp add: free_tv_subst dom_def)
-
-lemma free_tv_le_new_tv: "new_tv n t \<Longrightarrow> m \<in> free_tv t \<Longrightarrow> m < n"
-  by (unfold new_tv_def) fast
-
-lemma free_tv_subst_var: "free_tv (s (v::nat)) \<le> insert v (cod s)"
-  by (cases "v \<in> dom s") (auto simp add: cod_def dom_def)
-
-lemma free_tv_app_subst_te: "free_tv ($s (t::typ)) \<subseteq> cod s \<union> free_tv t"
-  by (induct t) (auto simp add: free_tv_subst_var)
-
-lemma free_tv_app_subst_tel: "free_tv ($s (ts::typ list)) \<subseteq> cod s \<union> free_tv ts"
-  apply (induct ts)
-   apply simp
-  apply (cut_tac free_tv_app_subst_te)
-  apply fastsimp
-  done
-
-lemma free_tv_comp_subst:
-    "free_tv (\<lambda>u::nat. $s1 (s2 u) :: typ) \<subseteq> free_tv s1 \<union> free_tv s2"
-  apply (unfold free_tv_subst dom_def)
-  apply (auto dest!: free_tv_subst_var [THEN subsetD] free_tv_app_subst_te [THEN subsetD]
-    simp add: cod_def dom_def simp del: bex_simps)
-  done
-
-
-subsection {* Most general unifiers *}
-
-consts
-  mgu :: "typ \<Rightarrow> typ \<Rightarrow> subst maybe"
-axioms
-  mgu_eq [simp]: "mgu t1 t2 = Ok u \<Longrightarrow> $u t1 = $u t2"
-  mgu_mg [simp]: "mgu t1 t2 = Ok u \<Longrightarrow> $s t1 = $s t2 \<Longrightarrow> \<exists>r. s = $r o u"
-  mgu_Ok: "$s t1 = $s t2 \<Longrightarrow> \<exists>u. mgu t1 t2 = Ok u"
-  mgu_free [simp]: "mgu t1 t2 = Ok u \<Longrightarrow> free_tv u \<subseteq> free_tv t1 \<union> free_tv t2"
-
-lemma mgu_new: "mgu t1 t2 = Ok u \<Longrightarrow> new_tv n t1 \<Longrightarrow> new_tv n t2 \<Longrightarrow> new_tv n u"
-  -- {* @{text mgu} does not introduce new type variables *}
-  by (unfold new_tv_def) (blast dest: mgu_free)
-
-
-section {* Mini-ML with type inference rules *}
-
-datatype
-  expr = Var nat | Abs expr | App expr expr
-
-
-text {* Type inference rules. *}
-
-inductive
-  has_type :: "typ list \<Rightarrow> expr \<Rightarrow> typ \<Rightarrow> bool"  ("((_) |-/ (_) :: (_))" [60, 0, 60] 60)
-  where
-    Var: "n < length a \<Longrightarrow> a |- Var n :: a ! n"
-  | Abs: "t1#a |- e :: t2 \<Longrightarrow> a |- Abs e :: t1 -> t2"
-  | App: "a |- e1 :: t2 -> t1 \<Longrightarrow> a |- e2 :: t2
-      \<Longrightarrow> a |- App e1 e2 :: t1"
-
-
-text {* Type assigment is closed wrt.\ substitution. *}
-
-lemma has_type_subst_closed: "a |- e :: t ==> $s a |- e :: $s t"
-proof (induct set: has_type)
-  case (Var n a)
-  then have "n < length (map ($ s) a)" by simp
-  then have "map ($ s) a |- Var n :: map ($ s) a ! n"
-    by (rule has_type.Var)
-  also have "map ($ s) a ! n = $ s (a ! n)"
-    by (rule nth_map) (rule Var)
-  also have "map ($ s) a = $ s a"
-    by (simp only: app_subst_list)
-  finally show ?case .
-next
-  case (Abs t1 a e t2)
-  then have "$ s t1 # map ($ s) a |- e :: $ s t2"
-    by (simp add: app_subst_list)
-  then have "map ($ s) a |- Abs e :: $ s t1 -> $ s t2"
-    by (rule has_type.Abs)
-  then show ?case
-    by (simp add: app_subst_list)
-next
-  case App
-  then show ?case by (simp add: has_type.App)
-qed
-
-
-section {* Correctness and completeness of the type inference algorithm W *}
-
-consts
-  "\<W>" :: "expr \<Rightarrow> typ list \<Rightarrow> nat \<Rightarrow> (subst \<times> typ \<times> nat) maybe"
-primrec
-  "\<W> (Var i) a n =
-    (if i < length a then Ok (id_subst, a ! i, n) else Fail)"
-  "\<W> (Abs e) a n =
-    ((s, t, m) := \<W> e (TVar n # a) (Suc n);
-     Ok (s, (s n) -> t, m))"
-  "\<W> (App e1 e2) a n =
-    ((s1, t1, m1) := \<W> e1 a n;
-     (s2, t2, m2) := \<W> e2 ($s1 a) m1;
-     u := mgu ($ s2 t1) (t2 -> TVar m2);
-     Ok ($u o $s2 o s1, $u (TVar m2), Suc m2))"
-
-theorem W_correct: "Ok (s, t, m) = \<W> e a n ==> $s a |- e :: t"
-proof (induct e arbitrary: a s t m n)
-  case (Var i)
-  from `Ok (s, t, m) = \<W> (Var i) a n`
-  show "$s a |- Var i :: t" by (simp add: has_type.Var split: if_splits)
-next
-  case (Abs e)
-  from `Ok (s, t, m) = \<W> (Abs e) a n`
-  obtain t' where "t = s n -> t'"
-      and "Ok (s, t', m) = \<W> e (TVar n # a) (Suc n)"
-    by (auto split: bind_splits)
-  with Abs.hyps show "$s a |- Abs e :: t"
-    by (force intro: has_type.Abs)
-next
-  case (App e1 e2)
-  from `Ok (s, t, m) = \<W> (App e1 e2) a n`
-  obtain s1 t1 n1 s2 t2 n2 u where
-          s: "s = $u o $s2 o s1"
-      and t: "t = u n2"
-      and mgu_ok: "mgu ($s2 t1) (t2 -> TVar n2) = Ok u"
-      and W1_ok: "Ok (s1, t1, n1) = \<W> e1 a n"
-      and W2_ok: "Ok (s2, t2, n2) = \<W> e2 ($s1 a) n1"
-    by (auto split: bind_splits simp: that)
-  show "$s a |- App e1 e2 :: t"
-  proof (rule has_type.App)
-    from s have s': "$u ($s2 ($s1 a)) = $s a"
-      by (simp add: subst_comp_tel o_def)
-    show "$s a |- e1 :: $u t2 -> t"
-    proof -
-      from W1_ok have "$s1 a |- e1 :: t1" by (rule App.hyps(1))
-      then have "$u ($s2 ($s1 a)) |- e1 :: $u ($s2 t1)"
-        by (intro has_type_subst_closed)
-      with s' t mgu_ok show ?thesis by simp
-    qed
-    show "$s a |- e2 :: $u t2"
-    proof -
-      from W2_ok have "$s2 ($s1 a) |- e2 :: t2" by (rule App.hyps(2))
-      then have "$u ($s2 ($s1 a)) |- e2 :: $u t2"
-        by (rule has_type_subst_closed)
-      with s' show ?thesis by simp
-    qed
-  qed
-qed
-
-
-inductive_cases has_type_casesE:
-  "s |- Var n :: t"
-  "s |- Abs e :: t"
-  "s |- App e1 e2 ::t"
-
-
-lemmas [simp] = Suc_le_lessD
-  and [simp del] = less_imp_le ex_simps all_simps
-
-lemma W_var_ge [simp]: "!!a n s t m. \<W> e a n = Ok (s, t, m) \<Longrightarrow> n \<le> m"
-  -- {* the resulting type variable is always greater or equal than the given one *}
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply clarsimp
-   txt {* case @{text "Abs e"} *}
-   apply (simp split add: split_bind)
-   apply (fast dest: Suc_leD)
-  txt {* case @{text "App e1 e2"} *}
-  apply (simp (no_asm) split add: split_bind)
-  apply (intro strip)
-  apply (rename_tac s t na sa ta nb sb)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = n in allE)
-  apply (erule_tac x = "$s a" in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = t in allE)
-  apply (erule_tac x = na in allE)
-  apply (erule_tac x = na in allE)
-  apply (simp add: eq_sym_conv)
-  done
-
-lemma W_var_geD: "Ok (s, t, m) = \<W> e a n \<Longrightarrow> n \<le> m"
-  by (simp add: eq_sym_conv)
-
-lemma new_tv_W: "!!n a s t m.
-  new_tv n a \<Longrightarrow> \<W> e a n = Ok (s, t, m) \<Longrightarrow> new_tv m s & new_tv m t"
-  -- {* resulting type variable is new *}
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply clarsimp
-    apply (force elim: list_ball_nth simp add: id_subst_def new_tv_list new_tv_subst)
-   txt {* case @{text "Abs e"} *}
-   apply (simp (no_asm) add: new_tv_subst new_tv_Suc_list split add: split_bind)
-   apply (intro strip)
-   apply (erule_tac x = "Suc n" in allE)
-   apply (erule_tac x = "TVar n # a" in allE)
-   apply (fastsimp simp add: new_tv_subst new_tv_Suc_list)
-  txt {* case @{text "App e1 e2"} *}
-  apply (simp (no_asm) split add: split_bind)
-  apply (intro strip)
-  apply (rename_tac s t na sa ta nb sb)
-  apply (erule_tac x = n in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = t in allE)
-  apply (erule_tac x = na in allE)
-  apply (erule_tac x = na in allE)
-  apply (simp add: eq_sym_conv)
-  apply (erule_tac x = "$s a" in allE)
-  apply (erule_tac x = sa in allE)
-  apply (erule_tac x = ta in allE)
-  apply (erule_tac x = nb in allE)
-  apply (simp add: o_def rotate_Ok)
-  apply (rule conjI)
-   apply (rule new_tv_subst_comp_2)
-    apply (rule new_tv_subst_comp_2)
-     apply (rule lessI [THEN less_imp_le, THEN new_tv_subst_le])
-     apply (rule_tac n = na in new_tv_subst_le)
-      apply (simp add: rotate_Ok)
-     apply (simp (no_asm_simp))
-    apply (fast dest: W_var_geD intro: new_tv_list_le new_tv_subst_tel
-      lessI [THEN less_imp_le, THEN new_tv_subst_le])
-   apply (erule sym [THEN mgu_new])
-    apply (best dest: W_var_geD intro: new_tv_subst_te new_tv_list_le new_tv_subst_tel
-      lessI [THEN less_imp_le, THEN new_tv_le] lessI [THEN less_imp_le, THEN new_tv_subst_le]
-      new_tv_le)
-   apply (tactic {* fast_tac (HOL_cs addDs [thm "W_var_geD"]
-     addIs [thm "new_tv_list_le", thm "new_tv_subst_tel", thm "new_tv_le"]
-     addss @{simpset}) 1 *})
-  apply (rule lessI [THEN new_tv_subst_var])
-  apply (erule sym [THEN mgu_new])
-    apply (bestsimp intro!: lessI [THEN less_imp_le, THEN new_tv_le] new_tv_subst_te
-      dest!: W_var_geD intro: new_tv_list_le new_tv_subst_tel
-        lessI [THEN less_imp_le, THEN new_tv_subst_le] new_tv_le)
-  apply (tactic {* fast_tac (HOL_cs addDs [thm "W_var_geD"]
-    addIs [thm "new_tv_list_le", thm "new_tv_subst_tel", thm "new_tv_le"]
-    addss @{simpset}) 1 *})
-  done
-
-lemma free_tv_W: "!!n a s t m v. \<W> e a n = Ok (s, t, m) \<Longrightarrow>
-    (v \<in> free_tv s \<or> v \<in> free_tv t) \<Longrightarrow> v < n \<Longrightarrow> v \<in> free_tv a"
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply clarsimp
-    apply (tactic {* fast_tac (HOL_cs addIs [thm "nth_mem", subsetD, thm "ftv_mem_sub_ftv_list"]) 1 *})
-   txt {* case @{text "Abs e"} *}
-   apply (simp add: free_tv_subst split add: split_bind)
-   apply (intro strip)
-   apply (rename_tac s t n1 v)
-   apply (erule_tac x = "Suc n" in allE)
-   apply (erule_tac x = "TVar n # a" in allE)
-   apply (erule_tac x = s in allE)
-   apply (erule_tac x = t in allE)
-   apply (erule_tac x = n1 in allE)
-   apply (erule_tac x = v in allE)
-   apply (force elim!: allE intro: cod_app_subst)
-  txt {* case @{text "App e1 e2"} *}
-  apply (simp (no_asm) split add: split_bind)
-  apply (intro strip)
-  apply (rename_tac s t n1 s1 t1 n2 s3 v)
-  apply (erule_tac x = n in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = t in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = v in allE)
-  txt {* second case *}
-  apply (erule_tac x = "$ s a" in allE)
-  apply (erule_tac x = s1 in allE)
-  apply (erule_tac x = t1 in allE)
-  apply (erule_tac x = n2 in allE)
-  apply (erule_tac x = v in allE)
-  apply (tactic "safe_tac (empty_cs addSIs [conjI, impI] addSEs [conjE])")
-   apply (simp add: rotate_Ok o_def)
-   apply (drule W_var_geD)
-   apply (drule W_var_geD)
-   apply (frule less_le_trans, assumption)
-   apply (fastsimp dest: free_tv_comp_subst [THEN subsetD] sym [THEN mgu_free] codD
-     free_tv_app_subst_te [THEN subsetD] free_tv_app_subst_tel [THEN subsetD] subsetD elim: UnE)
-  apply simp
-  apply (drule sym [THEN W_var_geD])
-  apply (drule sym [THEN W_var_geD])
-  apply (frule less_le_trans, assumption)
-  apply (tactic {* fast_tac (HOL_cs addDs [thm "mgu_free", thm "codD",
-    thm "free_tv_subst_var" RS subsetD,
-    thm "free_tv_app_subst_te" RS subsetD,
-    thm "free_tv_app_subst_tel" RS subsetD, @{thm less_le_trans}, subsetD]
-    addSEs [UnE] addss (@{simpset} setSolver unsafe_solver)) 1 *})
-      -- {* builtin arithmetic in simpset messes things up *}
-  done
-
-text {*
-  \medskip Completeness of @{text \<W>} wrt.\ @{text has_type}.
-*}
-
-lemma W_complete_aux: "!!s' a t' n. $s' a |- e :: t' \<Longrightarrow> new_tv n a \<Longrightarrow>
-    (\<exists>s t. (\<exists>m. \<W> e a n = Ok (s, t, m)) \<and> (\<exists>r. $s' a = $r ($s a) \<and> t' = $r t))"
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply (intro strip)
-    apply (simp (no_asm) cong add: conj_cong)
-    apply (erule has_type_casesE)
-    apply (simp add: eq_sym_conv app_subst_list)
-    apply (rule_tac x = s' in exI)
-    apply simp
-   txt {* case @{text "Abs e"} *}
-   apply (intro strip)
-   apply (erule has_type_casesE)
-   apply (erule_tac x = "\<lambda>x. if x = n then t1 else (s' x)" in allE)
-   apply (erule_tac x = "TVar n # a" in allE)
-   apply (erule_tac x = t2 in allE)
-   apply (erule_tac x = "Suc n" in allE)
-   apply (fastsimp cong add: conj_cong split add: split_bind)
-  txt {* case @{text "App e1 e2"} *}
-  apply (intro strip)
-  apply (erule has_type_casesE)
-  apply (erule_tac x = s' in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = "t2 -> t'" in allE)
-  apply (erule_tac x = n in allE)
-  apply (tactic "safe_tac HOL_cs")
-  apply (erule_tac x = r in allE)
-  apply (erule_tac x = "$s a" in allE)
-  apply (erule_tac x = t2 in allE)
-  apply (erule_tac x = m in allE)
-  apply simp
-  apply (tactic "safe_tac HOL_cs")
-   apply (tactic {* fast_tac (HOL_cs addIs [sym RS thm "W_var_geD",
-     thm "new_tv_W" RS conjunct1, thm "new_tv_list_le", thm "new_tv_subst_tel"]) 1 *})
-  apply (subgoal_tac
-    "$(\<lambda>x. if x = ma then t' else (if x \<in> free_tv t - free_tv sa then r x
-      else ra x)) ($ sa t) =
-    $(\<lambda>x. if x = ma then t' else (if x \<in> free_tv t - free_tv sa then r x
-      else ra x)) (ta -> (TVar ma))")
-   apply (rule_tac [2] t = "$(\<lambda>x. if x = ma then t'
-     else (if x \<in> (free_tv t - free_tv sa) then r x else ra x)) ($sa t)" and
-     s = "($ ra ta) -> t'" in ssubst)
-    prefer 2
-    apply (simp add: subst_comp_te)
-    apply (rule eq_free_eq_subst_te)
-    apply (intro strip)
-    apply (subgoal_tac "na \<noteq> ma")
-     prefer 2
-     apply (fast dest: new_tv_W sym [THEN W_var_geD] new_tv_not_free_tv new_tv_le)
-    apply (case_tac "na \<in> free_tv sa")
-     txt {* @{text "na \<notin> free_tv sa"} *}
-     prefer 2
-     apply (frule not_free_impl_id)
-     apply simp
-    txt {* @{text "na \<in> free_tv sa"} *}
-    apply (drule_tac ts1 = "$s a" and r = "$ r ($ s a)" in subst_comp_tel [THEN [2] trans])
-    apply (drule_tac eq_subst_tel_eq_free)
-     apply (fast intro: free_tv_W free_tv_le_new_tv dest: new_tv_W)
-    apply simp
-    apply (case_tac "na \<in> dom sa")
-     prefer 2
-     txt {* @{text "na \<noteq> dom sa"} *}
-     apply (simp add: dom_def)
-    txt {* @{text "na \<in> dom sa"} *}
-    apply (rule eq_free_eq_subst_te)
-    apply (intro strip)
-    apply (subgoal_tac "nb \<noteq> ma")
-     prefer 2
-     apply (frule new_tv_W, assumption)
-     apply (erule conjE)
-     apply (drule new_tv_subst_tel)
-      apply (fast intro: new_tv_list_le dest: sym [THEN W_var_geD])
-     apply (fastsimp dest: new_tv_W new_tv_not_free_tv simp add: cod_def free_tv_subst)
-    apply (fastsimp simp add: cod_def free_tv_subst)
-   prefer 2
-   apply (simp (no_asm))
-   apply (rule eq_free_eq_subst_te)
-   apply (intro strip)
-   apply (subgoal_tac "na \<noteq> ma")
-    prefer 2
-    apply (frule new_tv_W, assumption)
-    apply (erule conjE)
-    apply (drule sym [THEN W_var_geD])
-    apply (fast dest: new_tv_list_le new_tv_subst_tel new_tv_W new_tv_not_free_tv)
-   apply (case_tac "na \<in> free_tv t - free_tv sa")
-    prefer 2
-    txt {* case @{text "na \<notin> free_tv t - free_tv sa"} *}
-    apply simp
-    defer
-    txt {* case @{text "na \<in> free_tv t - free_tv sa"} *}
-    apply simp
-    apply (drule_tac ts1 = "$s a" and r = "$ r ($ s a)" in subst_comp_tel [THEN [2] trans])
-    apply (drule eq_subst_tel_eq_free)
-     apply (fast intro: free_tv_W free_tv_le_new_tv dest: new_tv_W)
-    apply (simp add: free_tv_subst dom_def)
-   prefer 2 apply fast
-  apply (simp (no_asm_simp) split add: split_bind)
-  apply (tactic "safe_tac HOL_cs")
-   apply (drule mgu_Ok)
-   apply fastsimp
-  apply (drule mgu_mg, assumption)
-  apply (erule exE)
-  apply (rule_tac x = rb in exI)
-  apply (rule conjI)
-   prefer 2
-   apply (drule_tac x = ma in fun_cong)
-   apply (simp add: eq_sym_conv)
-  apply (simp (no_asm) add: o_def subst_comp_tel [symmetric])
-  apply (rule subst_comp_tel [symmetric, THEN [2] trans])
-  apply (simp add: o_def eq_sym_conv)
-  apply (rule eq_free_eq_subst_tel)
-  apply (tactic "safe_tac HOL_cs")
-  apply (subgoal_tac "ma \<noteq> na")
-   prefer 2
-   apply (frule new_tv_W, assumption)
-   apply (erule conjE)
-   apply (drule new_tv_subst_tel)
-    apply (fast intro: new_tv_list_le dest: sym [THEN W_var_geD])
-   apply (frule_tac n = m in new_tv_W, assumption)
-   apply (erule conjE)
-   apply (drule free_tv_app_subst_tel [THEN subsetD])
-   apply (auto dest: W_var_geD [OF sym] new_tv_list_le
-     codD new_tv_not_free_tv)
-  apply (case_tac "na \<in> free_tv t - free_tv sa")
-   prefer 2
-   txt {* case @{text "na \<notin> free_tv t - free_tv sa"} *}
-   apply simp
-   defer
-   txt {* case @{text "na \<in> free_tv t - free_tv sa"} *}
-   apply simp
-   apply (drule free_tv_app_subst_tel [THEN subsetD])
-   apply (fastsimp dest: codD subst_comp_tel [THEN [2] trans]
-     eq_subst_tel_eq_free simp add: free_tv_subst dom_def)
-  done
-
-lemma W_complete: "[] |- e :: t' ==>
-    \<exists>s t. (\<exists>m. \<W> e [] n = Ok (s, t, m)) \<and> (\<exists>r. t' = $r t)"
-  apply (cut_tac a = "[]" and s' = id_subst and e = e and t' = t' in W_complete_aux)
-    apply simp_all
-  done
-
-
-section {* Equivalence of W and I *}
-
-text {*
-  Recursive definition of type inference algorithm @{text \<I>} for
-  Mini-ML.
-*}
-
-consts
-  "\<I>" :: "expr \<Rightarrow> typ list \<Rightarrow> nat \<Rightarrow> subst \<Rightarrow> (subst \<times> typ \<times> nat) maybe"
-primrec
-  "\<I> (Var i) a n s = (if i < length a then Ok (s, a ! i, n) else Fail)"
-  "\<I> (Abs e) a n s = ((s, t, m) := \<I> e (TVar n # a) (Suc n) s;
-    Ok (s, TVar n -> t, m))"
-  "\<I> (App e1 e2) a n s =
-    ((s1, t1, m1) := \<I> e1 a n s;
-    (s2, t2, m2) := \<I> e2 a m1 s1;
-    u := mgu ($s2 t1) ($s2 t2 -> TVar m2);
-    Ok($u o s2, TVar m2, Suc m2))"
-
-text {* \medskip Correctness. *}
-
-lemma I_correct_wrt_W: "!!a m s s' t n.
-    new_tv m a \<and> new_tv m s \<Longrightarrow> \<I> e a m s = Ok (s', t, n) \<Longrightarrow>
-    \<exists>r. \<W> e ($s a) m = Ok (r, $s' t, n) \<and> s' = ($r o s)"
-  apply (atomize (full))
-  apply (induct e)
-    txt {* case @{text "Var n"} *}
-    apply (simp add: app_subst_list split: split_if)
-   txt {* case @{text "Abs e"} *}
-   apply (tactic {* asm_full_simp_tac
-     (@{simpset} setloop (split_inside_tac [thm "split_bind"])) 1 *})
-   apply (intro strip)
-   apply (rule conjI)
-    apply (intro strip)
-    apply (erule allE)+
-    apply (erule impE)
-     prefer 2 apply (fastsimp simp add: new_tv_subst)
-    apply (tactic {* fast_tac (HOL_cs addIs [thm "new_tv_Suc_list" RS mp,
-      thm "new_tv_subst_le", @{thm less_imp_le}, @{thm lessI}]) 1 *})
-   apply (intro strip)
-   apply (erule allE)+
-   apply (erule impE)
-    prefer 2 apply (fastsimp simp add: new_tv_subst)
-   apply (tactic {* fast_tac (HOL_cs addIs [thm "new_tv_Suc_list" RS mp,
-     thm "new_tv_subst_le", @{thm less_imp_le}, @{thm lessI}]) 1 *})
-  txt {* case @{text "App e1 e2"} *}
-  apply (tactic {* simp_tac (@{simpset} setloop (split_inside_tac [thm "split_bind"])) 1 *})
-  apply (intro strip)
-  apply (rename_tac s1' t1 n1 s2' t2 n2 sa)
-  apply (rule conjI)
-   apply fastsimp
-  apply (intro strip)
-  apply (rename_tac s1 t1' n1')
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = m in allE)
-  apply (erule_tac x = s in allE)
-  apply (erule_tac x = s1' in allE)
-  apply (erule_tac x = t1 in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = a in allE)
-  apply (erule_tac x = n1 in allE)
-  apply (erule_tac x = s1' in allE)
-  apply (erule_tac x = s2' in allE)
-  apply (erule_tac x = t2 in allE)
-  apply (erule_tac x = n2 in allE)
-  apply (rule conjI)
-   apply (intro strip)
-   apply (rule notI)
-   apply simp
-   apply (erule impE)
-    apply (frule new_tv_subst_tel, assumption)
-    apply (drule_tac a = "$s a" in new_tv_W, assumption)
-    apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-   apply (fastsimp simp add: subst_comp_tel)
-  apply (intro strip)
-  apply (rename_tac s2 t2' n2')
-  apply (rule conjI)
-   apply (intro strip)
-   apply (rule notI)
-   apply simp
-   apply (erule impE)
-   apply (frule new_tv_subst_tel, assumption)
-   apply (drule_tac a = "$s a" in new_tv_W, assumption)
-    apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-   apply (fastsimp simp add: subst_comp_tel subst_comp_te)
-  apply (intro strip)
-  apply (erule (1) notE impE)
-  apply (erule (1) notE impE)
-  apply (erule exE)
-  apply (erule conjE)
-  apply (erule impE)
-   apply (frule new_tv_subst_tel, assumption)
-   apply (drule_tac a = "$s a" in new_tv_W, assumption)
-   apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-  apply (erule (1) notE impE)
-  apply (erule exE conjE)+
-  apply (simp (asm_lr) add: subst_comp_tel subst_comp_te o_def, (erule conjE)+, hypsubst)+
-  apply (subgoal_tac "new_tv n2 s \<and> new_tv n2 r \<and> new_tv n2 ra")
-   apply (simp add: new_tv_subst)
-  apply (frule new_tv_subst_tel, assumption)
-  apply (drule_tac a = "$s a" in new_tv_W, assumption)
-  apply (tactic "safe_tac HOL_cs")
-    apply (bestsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-   apply (fastsimp dest: sym [THEN W_var_geD] new_tv_subst_le new_tv_list_le)
-  apply (drule_tac e = e1 in sym [THEN W_var_geD])
-  apply (drule new_tv_subst_tel, assumption)
-  apply (drule_tac ts = "$s a" in new_tv_list_le, assumption)
-  apply (drule new_tv_subst_tel, assumption)
-  apply (bestsimp dest: new_tv_W simp add: subst_comp_tel)
-  done
-
-lemma I_complete_wrt_W: "!!a m s.
-    new_tv m a \<and> new_tv m s \<Longrightarrow> \<I> e a m s = Fail \<Longrightarrow> \<W> e ($s a) m = Fail"
-  apply (atomize (full))
-  apply (induct e)
-    apply (simp add: app_subst_list)
-   apply (simp (no_asm))
-   apply (intro strip)
-   apply (subgoal_tac "TVar m # $s a = $s (TVar m # a)")
-    apply (tactic {* asm_simp_tac (HOL_ss addsimps
-      [thm "new_tv_Suc_list", @{thm lessI} RS @{thm less_imp_le} RS thm "new_tv_subst_le"]) 1 *})
-    apply (erule conjE)
-    apply (drule new_tv_not_free_tv [THEN not_free_impl_id])
-    apply (simp (no_asm_simp))
-  apply (simp (no_asm_simp))
-  apply (intro strip)
-  apply (erule exE)+
-  apply (erule conjE)+
-  apply (drule I_correct_wrt_W [COMP swap_prems_rl])
-   apply fast
-  apply (erule exE)
-  apply (erule conjE)
-  apply hypsubst
-  apply (simp (no_asm_simp))
-  apply (erule disjE)
-   apply (rule disjI1)
-   apply (simp (no_asm_use) add: o_def subst_comp_tel)
-   apply (erule allE, erule allE, erule allE, erule impE, erule_tac [2] impE,
-     erule_tac [2] asm_rl, erule_tac [2] asm_rl)
-   apply (rule conjI)
-    apply (fast intro: W_var_ge [THEN new_tv_list_le])
-   apply (rule new_tv_subst_comp_2)
-    apply (fast intro: W_var_ge [THEN new_tv_subst_le])
-   apply (fast intro!: new_tv_subst_tel intro: new_tv_W [THEN conjunct1])
-  apply (rule disjI2)
-  apply (erule exE)+
-  apply (erule conjE)
-  apply (drule I_correct_wrt_W [COMP swap_prems_rl])
-   apply (rule conjI)
-   apply (fast intro: W_var_ge [THEN new_tv_list_le])
-   apply (rule new_tv_subst_comp_1)
-   apply (fast intro: W_var_ge [THEN new_tv_subst_le])
-   apply (fast intro!: new_tv_subst_tel intro: new_tv_W [THEN conjunct1])
-  apply (erule exE)
-  apply (erule conjE)
-  apply hypsubst
-  apply (simp add: o_def subst_comp_te [symmetric] subst_comp_tel [symmetric])
-  done
-
-end
--- a/src/HOL/W0/document/root.tex	Tue Feb 23 14:44:24 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-
-\documentclass[11pt,a4paper]{article}
-\usepackage{isabelle,isabellesym}
-\usepackage{pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-
-\newcommand{\isasymbind}{\textsf{bind}}
-
-\begin{document}
-
-\title{Type inference for let-free MiniML}
-\author{Dieter Nazareth, Tobias Nipkow, Thomas Stauner, Markus Wenzel}
-\maketitle
-
-\tableofcontents
-
-\parindent 0pt\parskip 0.5ex
-\input{session}
-
-%\bibliographystyle{abbrv}
-%\bibliography{root}
-
-\end{document}
--- a/src/HOL/ex/Codegenerator_Candidates.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/Codegenerator_Candidates.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -8,6 +8,8 @@
   Complex_Main
   AssocList
   Binomial
+  "~~/src/HOL/Decision_Procs/Commutative_Ring_Complete"
+  Dlist
   Fset
   Enum
   List_Prefix
@@ -17,12 +19,11 @@
   Permutation
   "~~/src/HOL/Number_Theory/Primes"
   Product_ord
+  "~~/src/HOL/ex/Records"
   SetsAndFunctions
   Tree
   While_Combinator
   Word
-  "~~/src/HOL/Decision_Procs/Commutative_Ring_Complete"
-  "~~/src/HOL/ex/Records"
 begin
 
 inductive sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Gauge_Integration.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -0,0 +1,573 @@
+(*  Author:     Jacques D. Fleuriot, University of Edinburgh
+    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
+
+    Replaced by ~~/src/HOL/Multivariate_Analysis/Real_Integral.thy .
+*)
+
+header{*Theory of Integration on real intervals*}
+
+theory Gauge_Integration
+imports Complex_Main
+begin
+
+text {*
+
+\textbf{Attention}: This theory defines the Integration on real
+intervals.  This is just a example theory for historical / expository interests.
+A better replacement is found in the Multivariate Analysis library. This defines
+the gauge integral on real vector spaces and in the Real Integral theory
+is a specialization to the integral on arbitrary real intervals.  The
+Multivariate Analysis package also provides a better support for analysis on
+integrals.
+
+*}
+
+text{*We follow John Harrison in formalizing the Gauge integral.*}
+
+subsection {* Gauges *}
+
+definition
+  gauge :: "[real set, real => real] => bool" where
+  [code del]:"gauge E g = (\<forall>x\<in>E. 0 < g(x))"
+
+
+subsection {* Gauge-fine divisions *}
+
+inductive
+  fine :: "[real \<Rightarrow> real, real \<times> real, (real \<times> real \<times> real) list] \<Rightarrow> bool"
+for
+  \<delta> :: "real \<Rightarrow> real"
+where
+  fine_Nil:
+    "fine \<delta> (a, a) []"
+| fine_Cons:
+    "\<lbrakk>fine \<delta> (b, c) D; a < b; a \<le> x; x \<le> b; b - a < \<delta> x\<rbrakk>
+      \<Longrightarrow> fine \<delta> (a, c) ((a, x, b) # D)"
+
+lemmas fine_induct [induct set: fine] =
+  fine.induct [of "\<delta>" "(a,b)" "D" "split P", unfolded split_conv, standard]
+
+lemma fine_single:
+  "\<lbrakk>a < b; a \<le> x; x \<le> b; b - a < \<delta> x\<rbrakk> \<Longrightarrow> fine \<delta> (a, b) [(a, x, b)]"
+by (rule fine_Cons [OF fine_Nil])
+
+lemma fine_append:
+  "\<lbrakk>fine \<delta> (a, b) D; fine \<delta> (b, c) D'\<rbrakk> \<Longrightarrow> fine \<delta> (a, c) (D @ D')"
+by (induct set: fine, simp, simp add: fine_Cons)
+
+lemma fine_imp_le: "fine \<delta> (a, b) D \<Longrightarrow> a \<le> b"
+by (induct set: fine, simp_all)
+
+lemma nonempty_fine_imp_less: "\<lbrakk>fine \<delta> (a, b) D; D \<noteq> []\<rbrakk> \<Longrightarrow> a < b"
+apply (induct set: fine, simp)
+apply (drule fine_imp_le, simp)
+done
+
+lemma empty_fine_imp_eq: "\<lbrakk>fine \<delta> (a, b) D; D = []\<rbrakk> \<Longrightarrow> a = b"
+by (induct set: fine, simp_all)
+
+lemma fine_eq: "fine \<delta> (a, b) D \<Longrightarrow> a = b \<longleftrightarrow> D = []"
+apply (cases "D = []")
+apply (drule (1) empty_fine_imp_eq, simp)
+apply (drule (1) nonempty_fine_imp_less, simp)
+done
+
+lemma mem_fine:
+  "\<lbrakk>fine \<delta> (a, b) D; (u, x, v) \<in> set D\<rbrakk> \<Longrightarrow> u < v \<and> u \<le> x \<and> x \<le> v"
+by (induct set: fine, simp, force)
+
+lemma mem_fine2: "\<lbrakk>fine \<delta> (a, b) D; (u, z, v) \<in> set D\<rbrakk> \<Longrightarrow> a \<le> u \<and> v \<le> b"
+apply (induct arbitrary: z u v set: fine, auto)
+apply (simp add: fine_imp_le)
+apply (erule order_trans [OF less_imp_le], simp)
+done
+
+lemma mem_fine3: "\<lbrakk>fine \<delta> (a, b) D; (u, z, v) \<in> set D\<rbrakk> \<Longrightarrow> v - u < \<delta> z"
+by (induct arbitrary: z u v set: fine) auto
+
+lemma BOLZANO:
+  fixes P :: "real \<Rightarrow> real \<Rightarrow> bool"
+  assumes 1: "a \<le> b"
+  assumes 2: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c"
+  assumes 3: "\<And>x. \<exists>d>0. \<forall>a b. a \<le> x & x \<le> b & (b-a) < d \<longrightarrow> P a b"
+  shows "P a b"
+apply (subgoal_tac "split P (a,b)", simp)
+apply (rule lemma_BOLZANO [OF _ _ 1])
+apply (clarify, erule (3) 2)
+apply (clarify, rule 3)
+done
+
+text{*We can always find a division that is fine wrt any gauge*}
+
+lemma fine_exists:
+  assumes "a \<le> b" and "gauge {a..b} \<delta>" shows "\<exists>D. fine \<delta> (a, b) D"
+proof -
+  {
+    fix u v :: real assume "u \<le> v"
+    have "a \<le> u \<Longrightarrow> v \<le> b \<Longrightarrow> \<exists>D. fine \<delta> (u, v) D"
+      apply (induct u v rule: BOLZANO, rule `u \<le> v`)
+       apply (simp, fast intro: fine_append)
+      apply (case_tac "a \<le> x \<and> x \<le> b")
+       apply (rule_tac x="\<delta> x" in exI)
+       apply (rule conjI)
+        apply (simp add: `gauge {a..b} \<delta>` [unfolded gauge_def])
+       apply (clarify, rename_tac u v)
+       apply (case_tac "u = v")
+        apply (fast intro: fine_Nil)
+       apply (subgoal_tac "u < v", fast intro: fine_single, simp)
+      apply (rule_tac x="1" in exI, clarsimp)
+      done
+  }
+  with `a \<le> b` show ?thesis by auto
+qed
+
+lemma fine_covers_all:
+  assumes "fine \<delta> (a, c) D" and "a < x" and "x \<le> c"
+  shows "\<exists> N < length D. \<forall> d t e. D ! N = (d,t,e) \<longrightarrow> d < x \<and> x \<le> e"
+  using assms
+proof (induct set: fine)
+  case (2 b c D a t)
+  thus ?case
+  proof (cases "b < x")
+    case True
+    with 2 obtain N where *: "N < length D"
+      and **: "\<And> d t e. D ! N = (d,t,e) \<Longrightarrow> d < x \<and> x \<le> e" by auto
+    hence "Suc N < length ((a,t,b)#D) \<and>
+           (\<forall> d t' e. ((a,t,b)#D) ! Suc N = (d,t',e) \<longrightarrow> d < x \<and> x \<le> e)" by auto
+    thus ?thesis by auto
+  next
+    case False with 2
+    have "0 < length ((a,t,b)#D) \<and>
+           (\<forall> d t' e. ((a,t,b)#D) ! 0 = (d,t',e) \<longrightarrow> d < x \<and> x \<le> e)" by auto
+    thus ?thesis by auto
+  qed
+qed auto
+
+lemma fine_append_split:
+  assumes "fine \<delta> (a,b) D" and "D2 \<noteq> []" and "D = D1 @ D2"
+  shows "fine \<delta> (a,fst (hd D2)) D1" (is "?fine1")
+  and "fine \<delta> (fst (hd D2), b) D2" (is "?fine2")
+proof -
+  from assms
+  have "?fine1 \<and> ?fine2"
+  proof (induct arbitrary: D1 D2)
+    case (2 b c D a' x D1 D2)
+    note induct = this
+
+    thus ?case
+    proof (cases D1)
+      case Nil
+      hence "fst (hd D2) = a'" using 2 by auto
+      with fine_Cons[OF `fine \<delta> (b,c) D` induct(3,4,5)] Nil induct
+      show ?thesis by (auto intro: fine_Nil)
+    next
+      case (Cons d1 D1')
+      with induct(2)[OF `D2 \<noteq> []`, of D1'] induct(8)
+      have "fine \<delta> (b, fst (hd D2)) D1'" and "fine \<delta> (fst (hd D2), c) D2" and
+        "d1 = (a', x, b)" by auto
+      with fine_Cons[OF this(1) induct(3,4,5), OF induct(6)] Cons
+      show ?thesis by auto
+    qed
+  qed auto
+  thus ?fine1 and ?fine2 by auto
+qed
+
+lemma fine_\<delta>_expand:
+  assumes "fine \<delta> (a,b) D"
+  and "\<And> x. \<lbrakk> a \<le> x ; x \<le> b \<rbrakk> \<Longrightarrow> \<delta> x \<le> \<delta>' x"
+  shows "fine \<delta>' (a,b) D"
+using assms proof induct
+  case 1 show ?case by (rule fine_Nil)
+next
+  case (2 b c D a x)
+  show ?case
+  proof (rule fine_Cons)
+    show "fine \<delta>' (b,c) D" using 2 by auto
+    from fine_imp_le[OF 2(1)] 2(6) `x \<le> b`
+    show "b - a < \<delta>' x"
+      using 2(7)[OF `a \<le> x`] by auto
+  qed (auto simp add: 2)
+qed
+
+lemma fine_single_boundaries:
+  assumes "fine \<delta> (a,b) D" and "D = [(d, t, e)]"
+  shows "a = d \<and> b = e"
+using assms proof induct
+  case (2 b c  D a x)
+  hence "D = []" and "a = d" and "b = e" by auto
+  moreover
+  from `fine \<delta> (b,c) D` `D = []` have "b = c"
+    by (rule empty_fine_imp_eq)
+  ultimately show ?case by simp
+qed auto
+
+lemma fine_listsum_eq_diff:
+  fixes f :: "real \<Rightarrow> real"
+  shows "fine \<delta> (a, b) D \<Longrightarrow> (\<Sum>(u, x, v)\<leftarrow>D. f v - f u) = f b - f a"
+by (induct set: fine) simp_all
+
+text{*Lemmas about combining gauges*}
+
+lemma gauge_min:
+     "[| gauge(E) g1; gauge(E) g2 |]
+      ==> gauge(E) (%x. min (g1(x)) (g2(x)))"
+by (simp add: gauge_def)
+
+lemma fine_min:
+      "fine (%x. min (g1(x)) (g2(x))) (a,b) D
+       ==> fine(g1) (a,b) D & fine(g2) (a,b) D"
+apply (erule fine.induct)
+apply (simp add: fine_Nil)
+apply (simp add: fine_Cons)
+done
+
+subsection {* Riemann sum *}
+
+definition
+  rsum :: "[(real \<times> real \<times> real) list, real \<Rightarrow> real] \<Rightarrow> real" where
+  "rsum D f = (\<Sum>(u, x, v)\<leftarrow>D. f x * (v - u))"
+
+lemma rsum_Nil [simp]: "rsum [] f = 0"
+unfolding rsum_def by simp
+
+lemma rsum_Cons [simp]: "rsum ((u, x, v) # D) f = f x * (v - u) + rsum D f"
+unfolding rsum_def by simp
+
+lemma rsum_zero [simp]: "rsum D (\<lambda>x. 0) = 0"
+by (induct D, auto)
+
+lemma rsum_left_distrib: "rsum D f * c = rsum D (\<lambda>x. f x * c)"
+by (induct D, auto simp add: algebra_simps)
+
+lemma rsum_right_distrib: "c * rsum D f = rsum D (\<lambda>x. c * f x)"
+by (induct D, auto simp add: algebra_simps)
+
+lemma rsum_add: "rsum D (\<lambda>x. f x + g x) =  rsum D f + rsum D g"
+by (induct D, auto simp add: algebra_simps)
+
+lemma rsum_append: "rsum (D1 @ D2) f = rsum D1 f + rsum D2 f"
+unfolding rsum_def map_append listsum_append ..
+
+
+subsection {* Gauge integrability (definite) *}
+
+definition
+  Integral :: "[(real*real),real=>real,real] => bool" where
+  [code del]: "Integral = (%(a,b) f k. \<forall>e > 0.
+                               (\<exists>\<delta>. gauge {a .. b} \<delta> &
+                               (\<forall>D. fine \<delta> (a,b) D -->
+                                         \<bar>rsum D f - k\<bar> < e)))"
+
+lemma Integral_def2:
+  "Integral = (%(a,b) f k. \<forall>e>0. (\<exists>\<delta>. gauge {a..b} \<delta> &
+                               (\<forall>D. fine \<delta> (a,b) D -->
+                                         \<bar>rsum D f - k\<bar> \<le> e)))"
+unfolding Integral_def
+apply (safe intro!: ext)
+apply (fast intro: less_imp_le)
+apply (drule_tac x="e/2" in spec)
+apply force
+done
+
+text{*The integral is unique if it exists*}
+
+lemma Integral_unique:
+    "[| a \<le> b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2"
+apply (simp add: Integral_def)
+apply (drule_tac x = "\<bar>k1 - k2\<bar> /2" in spec)+
+apply auto
+apply (drule gauge_min, assumption)
+apply (drule_tac \<delta> = "%x. min (\<delta> x) (\<delta>' x)"
+       in fine_exists, assumption, auto)
+apply (drule fine_min)
+apply (drule spec)+
+apply auto
+apply (subgoal_tac "\<bar>(rsum D f - k2) - (rsum D f - k1)\<bar> < \<bar>k1 - k2\<bar>")
+apply arith
+apply (drule add_strict_mono, assumption)
+apply (auto simp only: left_distrib [symmetric] mult_2_right [symmetric] 
+                mult_less_cancel_right)
+done
+
+lemma Integral_zero [simp]: "Integral(a,a) f 0"
+apply (auto simp add: Integral_def)
+apply (rule_tac x = "%x. 1" in exI)
+apply (auto dest: fine_eq simp add: gauge_def rsum_def)
+done
+
+lemma fine_rsum_const: "fine \<delta> (a,b) D \<Longrightarrow> rsum D (\<lambda>x. c) = (c * (b - a))"
+unfolding rsum_def
+by (induct set: fine, auto simp add: algebra_simps)
+
+lemma Integral_eq_diff_bounds: "a \<le> b ==> Integral(a,b) (%x. 1) (b - a)"
+apply (cases "a = b", simp)
+apply (simp add: Integral_def, clarify)
+apply (rule_tac x = "%x. b - a" in exI)
+apply (rule conjI, simp add: gauge_def)
+apply (clarify)
+apply (subst fine_rsum_const, assumption, simp)
+done
+
+lemma Integral_mult_const: "a \<le> b ==> Integral(a,b) (%x. c)  (c*(b - a))"
+apply (cases "a = b", simp)
+apply (simp add: Integral_def, clarify)
+apply (rule_tac x = "%x. b - a" in exI)
+apply (rule conjI, simp add: gauge_def)
+apply (clarify)
+apply (subst fine_rsum_const, assumption, simp)
+done
+
+lemma Integral_mult:
+     "[| a \<le> b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)"
+apply (auto simp add: order_le_less
+            dest: Integral_unique [OF order_refl Integral_zero])
+apply (auto simp add: Integral_def setsum_right_distrib[symmetric] mult_assoc)
+apply (case_tac "c = 0", force)
+apply (drule_tac x = "e/abs c" in spec)
+apply (simp add: divide_pos_pos)
+apply clarify
+apply (rule_tac x="\<delta>" in exI, clarify)
+apply (drule_tac x="D" in spec, clarify)
+apply (simp add: pos_less_divide_eq abs_mult [symmetric]
+                 algebra_simps rsum_right_distrib)
+done
+
+lemma Integral_add:
+  assumes "Integral (a, b) f x1"
+  assumes "Integral (b, c) f x2"
+  assumes "a \<le> b" and "b \<le> c"
+  shows "Integral (a, c) f (x1 + x2)"
+proof (cases "a < b \<and> b < c", simp only: Integral_def split_conv, rule allI, rule impI)
+  fix \<epsilon> :: real assume "0 < \<epsilon>"
+  hence "0 < \<epsilon> / 2" by auto
+
+  assume "a < b \<and> b < c"
+  hence "a < b" and "b < c" by auto
+
+  from `Integral (a, b) f x1`[simplified Integral_def split_conv,
+                              rule_format, OF `0 < \<epsilon>/2`]
+  obtain \<delta>1 where \<delta>1_gauge: "gauge {a..b} \<delta>1"
+    and I1: "\<And> D. fine \<delta>1 (a,b) D \<Longrightarrow> \<bar> rsum D f - x1 \<bar> < (\<epsilon> / 2)" by auto
+
+  from `Integral (b, c) f x2`[simplified Integral_def split_conv,
+                              rule_format, OF `0 < \<epsilon>/2`]
+  obtain \<delta>2 where \<delta>2_gauge: "gauge {b..c} \<delta>2"
+    and I2: "\<And> D. fine \<delta>2 (b,c) D \<Longrightarrow> \<bar> rsum D f - x2 \<bar> < (\<epsilon> / 2)" by auto
+
+  def \<delta> \<equiv> "\<lambda> x. if x < b then min (\<delta>1 x) (b - x)
+           else if x = b then min (\<delta>1 b) (\<delta>2 b)
+                         else min (\<delta>2 x) (x - b)"
+
+  have "gauge {a..c} \<delta>"
+    using \<delta>1_gauge \<delta>2_gauge unfolding \<delta>_def gauge_def by auto
+  moreover {
+    fix D :: "(real \<times> real \<times> real) list"
+    assume fine: "fine \<delta> (a,c) D"
+    from fine_covers_all[OF this `a < b` `b \<le> c`]
+    obtain N where "N < length D"
+      and *: "\<forall> d t e. D ! N = (d, t, e) \<longrightarrow> d < b \<and> b \<le> e"
+      by auto
+    obtain d t e where D_eq: "D ! N = (d, t, e)" by (cases "D!N", auto)
+    with * have "d < b" and "b \<le> e" by auto
+    have in_D: "(d, t, e) \<in> set D"
+      using D_eq[symmetric] using `N < length D` by auto
+
+    from mem_fine[OF fine in_D]
+    have "d < e" and "d \<le> t" and "t \<le> e" by auto
+
+    have "t = b"
+    proof (rule ccontr)
+      assume "t \<noteq> b"
+      with mem_fine3[OF fine in_D] `b \<le> e` `d \<le> t` `t \<le> e` `d < b` \<delta>_def
+      show False by (cases "t < b") auto
+    qed
+
+    let ?D1 = "take N D"
+    let ?D2 = "drop N D"
+    def D1 \<equiv> "take N D @ [(d, t, b)]"
+    def D2 \<equiv> "(if b = e then [] else [(b, t, e)]) @ drop (Suc N) D"
+
+    have "D \<noteq> []" using `N < length D` by auto
+    from hd_drop_conv_nth[OF this `N < length D`]
+    have "fst (hd ?D2) = d" using `D ! N = (d, t, e)` by auto
+    with fine_append_split[OF _ _ append_take_drop_id[symmetric]]
+    have fine1: "fine \<delta> (a,d) ?D1" and fine2: "fine \<delta> (d,c) ?D2"
+      using `N < length D` fine by auto
+
+    have "fine \<delta>1 (a,b) D1" unfolding D1_def
+    proof (rule fine_append)
+      show "fine \<delta>1 (a, d) ?D1"
+      proof (rule fine1[THEN fine_\<delta>_expand])
+        fix x assume "a \<le> x" "x \<le> d"
+        hence "x \<le> b" using `d < b` `x \<le> d` by auto
+        thus "\<delta> x \<le> \<delta>1 x" unfolding \<delta>_def by auto
+      qed
+
+      have "b - d < \<delta>1 t"
+        using mem_fine3[OF fine in_D] \<delta>_def `b \<le> e` `t = b` by auto
+      from `d < b` `d \<le> t` `t = b` this
+      show "fine \<delta>1 (d, b) [(d, t, b)]" using fine_single by auto
+    qed
+    note rsum1 = I1[OF this]
+
+    have drop_split: "drop N D = [D ! N] @ drop (Suc N) D"
+      using nth_drop'[OF `N < length D`] by simp
+
+    have fine2: "fine \<delta>2 (e,c) (drop (Suc N) D)"
+    proof (cases "drop (Suc N) D = []")
+      case True
+      note * = fine2[simplified drop_split True D_eq append_Nil2]
+      have "e = c" using fine_single_boundaries[OF * refl] by auto
+      thus ?thesis unfolding True using fine_Nil by auto
+    next
+      case False
+      note * = fine_append_split[OF fine2 False drop_split]
+      from fine_single_boundaries[OF *(1)]
+      have "fst (hd (drop (Suc N) D)) = e" using D_eq by auto
+      with *(2) have "fine \<delta> (e,c) (drop (Suc N) D)" by auto
+      thus ?thesis
+      proof (rule fine_\<delta>_expand)
+        fix x assume "e \<le> x" and "x \<le> c"
+        thus "\<delta> x \<le> \<delta>2 x" using `b \<le> e` unfolding \<delta>_def by auto
+      qed
+    qed
+
+    have "fine \<delta>2 (b, c) D2"
+    proof (cases "e = b")
+      case True thus ?thesis using fine2 by (simp add: D1_def D2_def)
+    next
+      case False
+      have "e - b < \<delta>2 b"
+        using mem_fine3[OF fine in_D] \<delta>_def `d < b` `t = b` by auto
+      with False `t = b` `b \<le> e`
+      show ?thesis using D2_def
+        by (auto intro!: fine_append[OF _ fine2] fine_single
+               simp del: append_Cons)
+    qed
+    note rsum2 = I2[OF this]
+
+    have "rsum D f = rsum (take N D) f + rsum [D ! N] f + rsum (drop (Suc N) D) f"
+      using rsum_append[symmetric] nth_drop'[OF `N < length D`] by auto
+    also have "\<dots> = rsum D1 f + rsum D2 f"
+      by (cases "b = e", auto simp add: D1_def D2_def D_eq rsum_append algebra_simps)
+    finally have "\<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>"
+      using add_strict_mono[OF rsum1 rsum2] by simp
+  }
+  ultimately show "\<exists> \<delta>. gauge {a .. c} \<delta> \<and>
+    (\<forall>D. fine \<delta> (a,c) D \<longrightarrow> \<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>)"
+    by blast
+next
+  case False
+  hence "a = b \<or> b = c" using `a \<le> b` and `b \<le> c` by auto
+  thus ?thesis
+  proof (rule disjE)
+    assume "a = b" hence "x1 = 0"
+      using `Integral (a, b) f x1` Integral_zero Integral_unique[of a b] by auto
+    thus ?thesis using `a = b` `Integral (b, c) f x2` by auto
+  next
+    assume "b = c" hence "x2 = 0"
+      using `Integral (b, c) f x2` Integral_zero Integral_unique[of b c] by auto
+    thus ?thesis using `b = c` `Integral (a, b) f x1` by auto
+  qed
+qed
+
+text{*Fundamental theorem of calculus (Part I)*}
+
+text{*"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988 *}
+
+lemma strad1:
+       "\<lbrakk>\<forall>z::real. z \<noteq> x \<and> \<bar>z - x\<bar> < s \<longrightarrow>
+             \<bar>(f z - f x) / (z - x) - f' x\<bar> < e/2;
+        0 < s; 0 < e; a \<le> x; x \<le> b\<rbrakk>
+       \<Longrightarrow> \<forall>z. \<bar>z - x\<bar> < s -->\<bar>f z - f x - f' x * (z - x)\<bar> \<le> e/2 * \<bar>z - x\<bar>"
+apply clarify
+apply (case_tac "z = x", simp)
+apply (drule_tac x = z in spec)
+apply (rule_tac z1 = "\<bar>inverse (z - x)\<bar>" 
+       in real_mult_le_cancel_iff2 [THEN iffD1])
+ apply simp
+apply (simp del: abs_inverse abs_mult add: abs_mult [symmetric]
+          mult_assoc [symmetric])
+apply (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) 
+                    = (f z - f x) / (z - x) - f' x")
+ apply (simp add: abs_mult [symmetric] mult_ac diff_minus)
+apply (subst mult_commute)
+apply (simp add: left_distrib diff_minus)
+apply (simp add: mult_assoc divide_inverse)
+apply (simp add: left_distrib)
+done
+
+lemma lemma_straddle:
+  assumes f': "\<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x)" and "0 < e"
+  shows "\<exists>g. gauge {a..b} g &
+                (\<forall>x u v. a \<le> u & u \<le> x & x \<le> v & v \<le> b & (v - u) < g(x)
+                  --> \<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))"
+proof -
+  have "\<forall>x\<in>{a..b}.
+        (\<exists>d > 0. \<forall>u v. u \<le> x & x \<le> v & (v - u) < d --> 
+                       \<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))"
+  proof (clarsimp)
+    fix x :: real assume "a \<le> x" and "x \<le> b"
+    with f' have "DERIV f x :> f'(x)" by simp
+    then have "\<forall>r>0. \<exists>s>0. \<forall>z. z \<noteq> x \<and> \<bar>z - x\<bar> < s \<longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < r"
+      by (simp add: DERIV_iff2 LIM_eq)
+    with `0 < e` obtain s
+    where "\<forall>z. z \<noteq> x \<and> \<bar>z - x\<bar> < s \<longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < e/2" and "0 < s"
+      by (drule_tac x="e/2" in spec, auto)
+    then have strad [rule_format]:
+        "\<forall>z. \<bar>z - x\<bar> < s --> \<bar>f z - f x - f' x * (z - x)\<bar> \<le> e/2 * \<bar>z - x\<bar>"
+      using `0 < e` `a \<le> x` `x \<le> b` by (rule strad1)
+    show "\<exists>d>0. \<forall>u v. u \<le> x \<and> x \<le> v \<and> v - u < d \<longrightarrow> \<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)"
+    proof (safe intro!: exI)
+      show "0 < s" by fact
+    next
+      fix u v :: real assume "u \<le> x" and "x \<le> v" and "v - u < s"
+      have "\<bar>f v - f u - f' x * (v - u)\<bar> =
+            \<bar>(f v - f x - f' x * (v - x)) + (f x - f u - f' x * (x - u))\<bar>"
+        by (simp add: right_diff_distrib)
+      also have "\<dots> \<le> \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f x - f u - f' x * (x - u)\<bar>"
+        by (rule abs_triangle_ineq)
+      also have "\<dots> = \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f u - f x - f' x * (u - x)\<bar>"
+        by (simp add: right_diff_distrib)
+      also have "\<dots> \<le> (e/2) * \<bar>v - x\<bar> + (e/2) * \<bar>u - x\<bar>"
+        using `u \<le> x` `x \<le> v` `v - u < s` by (intro add_mono strad, simp_all)
+      also have "\<dots> \<le> e * (v - u) / 2 + e * (v - u) / 2"
+        using `u \<le> x` `x \<le> v` `0 < e` by (intro add_mono, simp_all)
+      also have "\<dots> = e * (v - u)"
+        by simp
+      finally show "\<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)" .
+    qed
+  qed
+  thus ?thesis
+    by (simp add: gauge_def) (drule bchoice, auto)
+qed
+
+lemma fundamental_theorem_of_calculus:
+  "\<lbrakk> a \<le> b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) \<rbrakk>
+             \<Longrightarrow> Integral(a,b) f' (f(b) - f(a))"
+ apply (drule order_le_imp_less_or_eq, auto)
+ apply (auto simp add: Integral_def2)
+ apply (drule_tac e = "e / (b - a)" in lemma_straddle)
+  apply (simp add: divide_pos_pos)
+ apply clarify
+ apply (rule_tac x="g" in exI, clarify)
+ apply (clarsimp simp add: rsum_def)
+ apply (frule fine_listsum_eq_diff [where f=f])
+ apply (erule subst)
+ apply (subst listsum_subtractf [symmetric])
+ apply (rule listsum_abs [THEN order_trans])
+ apply (subst map_map [unfolded o_def])
+ apply (subgoal_tac "e = (\<Sum>(u, x, v)\<leftarrow>D. (e / (b - a)) * (v - u))")
+  apply (erule ssubst)
+  apply (simp add: abs_minus_commute)
+  apply (rule listsum_mono)
+  apply (clarify, rename_tac u x v)
+  apply ((drule spec)+, erule mp)
+  apply (simp add: mem_fine mem_fine2 mem_fine3)
+ apply (frule fine_listsum_eq_diff [where f="\<lambda>x. x"])
+ apply (simp only: split_def)
+ apply (subst listsum_const_mult)
+ apply simp
+done
+
+end
--- a/src/HOL/ex/PER.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/PER.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,5 +1,4 @@
 (*  Title:      HOL/ex/PER.thy
-    ID:         $Id$
     Author:     Oscar Slotosch and Markus Wenzel, TU Muenchen
 *)
 
@@ -30,12 +29,10 @@
   but not necessarily reflexive.
 *}
 
-consts
-  eqv :: "'a => 'a => bool"    (infixl "\<sim>" 50)
-
-axclass partial_equiv < type
-  partial_equiv_sym [elim?]: "x \<sim> y ==> y \<sim> x"
-  partial_equiv_trans [trans]: "x \<sim> y ==> y \<sim> z ==> x \<sim> z"
+class partial_equiv =
+  fixes eqv :: "'a => 'a => bool"    (infixl "\<sim>" 50)
+  assumes partial_equiv_sym [elim?]: "x \<sim> y ==> y \<sim> x"
+  assumes partial_equiv_trans [trans]: "x \<sim> y ==> y \<sim> z ==> x \<sim> z"
 
 text {*
   \medskip The domain of a partial equivalence relation is the set of
@@ -70,7 +67,10 @@
   structural one corresponding to the congruence property.
 *}
 
-defs (overloaded)
+instantiation "fun" :: (partial_equiv, partial_equiv) partial_equiv
+begin
+
+definition
   eqv_fun_def: "f \<sim> g == \<forall>x \<in> domain. \<forall>y \<in> domain. x \<sim> y --> f x \<sim> g y"
 
 lemma partial_equiv_funI [intro?]:
@@ -86,8 +86,7 @@
   spaces (in \emph{both} argument positions).
 *}
 
-instance "fun" :: (partial_equiv, partial_equiv) partial_equiv
-proof
+instance proof
   fix f g h :: "'a::partial_equiv => 'b::partial_equiv"
   assume fg: "f \<sim> g"
   show "g \<sim> f"
@@ -110,6 +109,8 @@
   qed
 qed
 
+end
+
 
 subsection {* Total equivalence *}
 
@@ -120,8 +121,8 @@
   symmetric.
 *}
 
-axclass equiv < partial_equiv
-  eqv_refl [intro]: "x \<sim> x"
+class equiv =
+  assumes eqv_refl [intro]: "x \<sim> x"
 
 text {*
   On total equivalences all elements are reflexive, and congruence
@@ -150,7 +151,7 @@
   \emph{equivalence classes} over elements of the base type @{typ 'a}.
 *}
 
-typedef 'a quot = "{{x. a \<sim> x}| a::'a. True}"
+typedef 'a quot = "{{x. a \<sim> x}| a::'a::partial_equiv. True}"
   by blast
 
 lemma quotI [intro]: "{x. a \<sim> x} \<in> quot"
--- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -2,6 +2,21 @@
 imports "../Predicate_Compile"
 begin
 
+section {* Common constants *}
+
+declare HOL.if_bool_eq_disj[code_pred_inline]
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name Let}] *}
+
+section {* Pairs *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name fst}, @{const_name snd}, @{const_name split}] *}
+
+section {* Bounded quantifiers *}
+
+declare Ball_def[code_pred_inline]
+declare Bex_def[code_pred_inline]
+
 section {* Set operations *}
 
 declare Collect_def[code_pred_inline]
@@ -9,13 +24,37 @@
 
 declare eq_reflection[OF empty_def, code_pred_inline]
 declare insert_code[code_pred_def]
+
+declare subset_iff[code_pred_inline]
+
+declare Int_def[code_pred_inline]
 declare eq_reflection[OF Un_def, code_pred_inline]
 declare eq_reflection[OF UNION_def, code_pred_inline]
 
+lemma Diff[code_pred_inline]:
+  "(A - B) = (%x. A x \<and> \<not> B x)"
+by (auto simp add: mem_def)
 
+lemma set_equality[code_pred_inline]:
+  "(A = B) = ((\<forall>x. A x \<longrightarrow> B x) \<and> (\<forall>x. B x \<longrightarrow> A x))"
+by (fastsimp simp add: mem_def)
+
+section {* Setup for Numerals *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name number_of}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name number_of}] *}
+
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
 
 section {* Alternative list definitions *}
+
+text {* size simps are not yet added to the Spec_Rules interface. So they are just added manually here! *}
  
+lemma [code_pred_def]:
+  "length [] = 0"
+  "length (x # xs) = Suc (length xs)"
+by auto
+
 subsection {* Alternative rules for set *}
 
 lemma set_ConsI1 [code_pred_intro]:
--- a/src/HOL/ex/Predicate_Compile_Quickcheck.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/Predicate_Compile_Quickcheck.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -3,11 +3,14 @@
 header {* A Prototype of Quickcheck based on the Predicate Compiler *}
 
 theory Predicate_Compile_Quickcheck
-imports "../Predicate_Compile"
+imports "../Predicate_Compile" Quickcheck Predicate_Compile_Alternative_Defs
 uses "../Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
 begin
 
-setup {* Quickcheck.add_generator ("predicate_compile", Predicate_Compile_Quickcheck.quickcheck) *}
+setup {* Quickcheck.add_generator ("predicate_compile_wo_ff", Predicate_Compile_Quickcheck.quickcheck_compile_term false true) *}
+setup {* Quickcheck.add_generator ("predicate_compile_ff_fs", Predicate_Compile_Quickcheck.quickcheck_compile_term true true) *}
+setup {* Quickcheck.add_generator ("predicate_compile_ff_nofs", Predicate_Compile_Quickcheck.quickcheck_compile_term true false) *}
+
 (*
 datatype alphabet = a | b
 
--- a/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1,45 +1,39 @@
 theory Predicate_Compile_Quickcheck_ex
 imports Predicate_Compile_Quickcheck
-  Predicate_Compile_Alternative_Defs
 begin
 
-ML {* Predicate_Compile_Alternative_Defs.get *}
-
 section {* Sets *}
-(*
+
 lemma "x \<in> {(1::nat)} ==> False"
-quickcheck[generator=predicate_compile, iterations=10]
+quickcheck[generator=predicate_compile_wo_ff, iterations=10]
 oops
-*)
-(* TODO: some error with doubled negation *)
-(*
+
 lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x \<noteq> Suc 0"
-quickcheck[generator=predicate_compile]
+quickcheck[generator=predicate_compile_wo_ff]
 oops
-*)
-(*
+
 lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x = Suc 0"
-quickcheck[generator=predicate_compile]
+quickcheck[generator=predicate_compile_wo_ff]
 oops
-*) 
+ 
 lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x <= Suc 0"
-(*quickcheck[generator=predicate_compile]*)
+quickcheck[generator=predicate_compile_wo_ff]
 oops
 
 section {* Numerals *}
-(*
+
 lemma
   "x \<in> {1, 2, (3::nat)} ==> x = 1 \<or> x = 2"
-quickcheck[generator=predicate_compile]
+quickcheck[generator=predicate_compile_wo_ff]
 oops
-*)
+
 lemma "x \<in> {1, 2, (3::nat)} ==> x < 3"
-(*quickcheck[generator=predicate_compile]*)
+quickcheck[generator=predicate_compile_wo_ff]
 oops
 
 lemma
   "x \<in> {1, 2} \<union> {3, 4} ==> x = (1::nat) \<or> x = (2::nat)"
-(*quickcheck[generator=predicate_compile]*)
+quickcheck[generator=predicate_compile_wo_ff]
 oops
 
 section {* Context Free Grammar *}
@@ -53,33 +47,15 @@
 | "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
 | "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
 | "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
-(*
-code_pred [random_dseq inductify] "S\<^isub>1p" .
-*)
-(*thm B\<^isub>1p.random_dseq_equation*)
-(*
-values [random_dseq 2, 2, 4] 10 "{x. S\<^isub>1p x}"
-values [random_dseq 1, 1, 5] 20 "{x. S\<^isub>1p x}"
 
-ML {* set ML_Context.trace *}
-*)
-ML {* set Toplevel.debug *}
-(*
-quickcheck[generator = predicate_compile, size = 10, iterations = 1]
-oops
-*)
-ML {* Spec_Rules.get *}
-ML {* Item_Net.retrieve *}
-local_setup {* Local_Theory.checkpoint *}
-ML {* Predicate_Compile_Data.get_specification @{theory} @{term "append"} *}
 lemma
-  "w \<in> S\<^isub>1p \<Longrightarrow> w = []"
-quickcheck[generator = predicate_compile, iterations=1]
+  "w \<in> S\<^isub>1 \<Longrightarrow> w = []"
+quickcheck[generator = predicate_compile_ff_nofs, iterations=1]
 oops
 
 theorem S\<^isub>1_sound:
-"w \<in> S\<^isub>1p \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=15]
+"w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
+quickcheck[generator=predicate_compile_ff_nofs, size=15]
 oops
 
 
@@ -90,7 +66,7 @@
 | "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
 | "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
 | "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
-
+(*
 code_pred [random_dseq inductify] S\<^isub>2 .
 thm S\<^isub>2.random_dseq_equation
 thm A\<^isub>2.random_dseq_equation
@@ -118,10 +94,10 @@
 "w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"
 quickcheck[generator=predicate_compile, size = 10, iterations = 1]
 oops
-
+*)
 theorem S\<^isub>2_sound:
 "w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=15, iterations=1]
+quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10]
 oops
 
 inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
@@ -141,17 +117,17 @@
 
 lemma S\<^isub>3_sound:
 "w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator=predicate_compile, size=10, iterations=10]
+quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10]
 oops
 
 lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"
-quickcheck[size=10, generator = predicate_compile]
+quickcheck[size=10, generator = predicate_compile_ff_fs]
 oops
 
 theorem S\<^isub>3_complete:
 "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3"
 (*quickcheck[generator=SML]*)
-quickcheck[generator=predicate_compile, size=10, iterations=100]
+quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100]
 oops
 
 
@@ -166,20 +142,23 @@
 
 theorem S\<^isub>4_sound:
 "w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
-quickcheck[generator = predicate_compile, size=5, iterations=1]
+quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
 oops
 
 theorem S\<^isub>4_complete:
 "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4"
-quickcheck[generator = predicate_compile, size=5, iterations=1]
+quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1]
 oops
 
-hide const b
+hide const a b
 
 subsection {* Lexicographic order *}
+(* TODO *)
+(*
 lemma
   "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
-
+oops
+*)
 subsection {* IMP *}
 
 types
@@ -208,7 +187,7 @@
 
 lemma
   "exec c s s' ==> exec (Seq c c) s s'"
-quickcheck[generator = predicate_compile, size=3, iterations=1]
+(*quickcheck[generator = predicate_compile_wo_ff, size=2, iterations=10]*)
 oops
 
 subsection {* Lambda *}
@@ -263,28 +242,9 @@
 
 lemma
   "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
-quickcheck[generator = predicate_compile, size = 7, iterations = 10]
+quickcheck[generator = predicate_compile_ff_fs, size = 7, iterations = 10]
 oops
 
-(*
-code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
-thm typing.equation
-
-code_pred (modes: i => i => bool,  i => o => bool as reduce') beta .
-thm beta.equation
-
-values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
-
-definition "reduce t = Predicate.the (reduce' t)"
-
-value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
-
-code_pred [random] typing .
-code_pred [random_dseq] typing .
-
-(*values [random] 1 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
-*)*)
-
 subsection {* JAD *}
 
 definition matrix :: "('a :: semiring_0) list list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
@@ -300,9 +260,17 @@
 lemma [code_pred_intro]:
   "matrix [] 0 m"
   "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
-sorry
+proof -
+  show "matrix [] 0 m" unfolding matrix_def by auto
+next
+  show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m"
+    unfolding matrix_def by auto
+qed
 
-code_pred [random_dseq inductify] matrix sorry
+code_pred [random_dseq inductify] matrix
+  apply (cases x)
+  unfolding matrix_def apply fastsimp
+  apply fastsimp done
 
 
 values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}"
@@ -344,10 +312,10 @@
 
 definition
   "length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)"
-
+(*
 definition
   "transpose M = [map (\<lambda> xs. xs ! i) (takeWhile (\<lambda> xs. i < length xs) M). i \<leftarrow> [0 ..< length (M ! 0)]]"
-
+*)
 definition
   "inflate upds = foldr (\<lambda> (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)"
 
@@ -356,15 +324,14 @@
 
 definition
   "jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (\<lambda> (i, x). v ! i * x)))"
-ML {* ML_Context.trace := false *}
 
 lemma "matrix (M::int list list) rs cs \<Longrightarrow> False"
-quickcheck[generator = predicate_compile, size = 6]
+quickcheck[generator = predicate_compile_ff_nofs, size = 6]
 oops
 
 lemma
   "\<lbrakk> matrix M rs cs ; length v = cs \<rbrakk> \<Longrightarrow> jad_mv v (jad M) = mv M v"
-(*quickcheck[generator = predicate_compile]*)
+quickcheck[generator = predicate_compile_wo_ff]
 oops
 
 end
\ No newline at end of file
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -252,10 +252,12 @@
   "one_or_two = {Suc 0, (Suc (Suc 0))}"
 
 code_pred [inductify] one_or_two .
+
 code_pred [dseq] one_or_two .
-(*code_pred [random_dseq] one_or_two .*)
+code_pred [random_dseq] one_or_two .
+thm one_or_two.dseq_equation
 values [expected "{Suc 0::nat, 2::nat}"] "{x. one_or_two x}"
-(*values [random_dseq 1,1,2] "{x. one_or_two x}"*)
+values [random_dseq 0,0,10] 3 "{x. one_or_two x}"
 
 inductive one_or_two' :: "nat => bool"
 where
@@ -269,12 +271,12 @@
 
 definition one_or_two'':
   "one_or_two'' == {1, (2::nat)}"
-ML {* prop_of @{thm one_or_two''} *}
-(*code_pred [inductify] one_or_two'' .
+
+code_pred [inductify] one_or_two'' .
 thm one_or_two''.equation
 
 values "{x. one_or_two'' x}"
-*)
+
 subsection {* even predicate *}
 
 inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
@@ -779,6 +781,25 @@
 thm divmod_rel.equation
 value [code] "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
 
+subsection {* Transforming predicate logic into logic programs *}
+
+subsection {* Transforming functions into logic programs *}
+definition
+  "case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"
+
+code_pred [inductify] case_f .
+thm case_fP.equation
+thm case_fP.intros
+
+fun fold_map_idx where
+  "fold_map_idx f i y [] = (y, [])"
+| "fold_map_idx f i y (x # xs) =
+ (let (y', x') = f i y x; (y'', xs') = fold_map_idx f (Suc i) y' xs
+ in (y'', x' # xs'))"
+
+text {* mode analysis explores thousand modes - this is infeasible at the moment... *}
+(*code_pred [inductify, show_steps] fold_map_idx .*)
+
 subsection {* Minimum *}
 
 definition Min
@@ -883,9 +904,16 @@
 
 
 values [random_dseq 1, 2, 5] 10 "{(n, xs, ys::int list). lexn (%(x, y). x <= y) n (xs, ys)}"
-
-
-code_pred [inductify] lenlex .
+thm lenlex_conv
+thm lex_conv
+declare list.size(3,4)[code_pred_def]
+(*code_pred [inductify, show_steps, show_intermediate_results] length .*)
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name Orderings.top_class.top}] *}
+code_pred [inductify] lex .
+thm lex.equation
+thm lex_def
+declare lenlex_conv[code_pred_def]
+code_pred [inductify, show_steps, show_intermediate_results] lenlex .
 thm lenlex.equation
 
 code_pred [random_dseq inductify] lenlex .
@@ -893,10 +921,10 @@
 
 values [random_dseq 4, 2, 4] 100 "{(xs, ys::int list). lenlex (%(x, y). x <= y) (xs, ys)}"
 thm lists.intros
-(*
+
 code_pred [inductify] lists .
-*)
-(*thm lists.equation*)
+
+thm lists.equation
 
 subsection {* AVL Tree *}
 
@@ -974,13 +1002,17 @@
   (o * o => bool) => i => bool,
   (i * o => bool) => i => bool) [inductify] Domain .
 thm Domain.equation
+thm Range_def
+
 code_pred (modes:
   (o * o => bool) => o => bool,
   (o * o => bool) => i => bool,
   (o * i => bool) => i => bool) [inductify] Range .
 thm Range.equation
+
 code_pred [inductify] Field .
 thm Field.equation
+
 (*thm refl_on_def
 code_pred [inductify] refl_on .
 thm refl_on.equation*)
@@ -992,9 +1024,10 @@
 thm trans.equation
 code_pred [inductify] single_valued .
 thm single_valued.equation
-code_pred [inductify] inv_image .
+thm inv_image_def
+(*code_pred [inductify] inv_image .
 thm inv_image.equation
-
+*)
 subsection {* Inverting list functions *}
 
 (*code_pred [inductify] length .
--- a/src/HOL/ex/ROOT.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/ROOT.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -67,7 +67,8 @@
   "Quickcheck_Examples",
   "Landau",
   "Execute_Choice",
-  "Summation"
+  "Summation",
+  "Gauge_Integration"
 ];
 
 HTML.with_charset "utf-8" (no_document use_thys)
--- a/src/HOL/ex/Refute_Examples.thy	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/HOL/ex/Refute_Examples.thy	Tue Feb 23 14:44:43 2010 -0800
@@ -1417,29 +1417,20 @@
 
 (*****************************************************************************)
 
-subsubsection {* Axiomatic type classes and overloading *}
+subsubsection {* Type classes and overloading *}
 
 text {* A type class without axioms: *}
 
-axclass classA
+class classA
 
 lemma "P (x::'a::classA)"
   refute
 oops
 
-text {* The axiom of this type class does not contain any type variables: *}
-
-axclass classB
-  classB_ax: "P | ~ P"
-
-lemma "P (x::'a::classB)"
-  refute
-oops
-
 text {* An axiom with a type variable (denoting types which have at least two elements): *}
 
-axclass classC < type
-  classC_ax: "\<exists>x y. x \<noteq> y"
+class classC =
+  assumes classC_ax: "\<exists>x y. x \<noteq> y"
 
 lemma "P (x::'a::classC)"
   refute
@@ -1451,11 +1442,9 @@
 
 text {* A type class for which a constant is defined: *}
 
-consts
-  classD_const :: "'a \<Rightarrow> 'a"
-
-axclass classD < type
-  classD_ax: "classD_const (classD_const x) = classD_const x"
+class classD =
+  fixes classD_const :: "'a \<Rightarrow> 'a"
+  assumes classD_ax: "classD_const (classD_const x) = classD_const x"
 
 lemma "P (x::'a::classD)"
   refute
@@ -1463,16 +1452,12 @@
 
 text {* A type class with multiple superclasses: *}
 
-axclass classE < classC, classD
+class classE = classC + classD
 
 lemma "P (x::'a::classE)"
   refute
 oops
 
-lemma "P (x::'a::{classB, classE})"
-  refute
-oops
-
 text {* OFCLASS: *}
 
 lemma "OFCLASS('a::type, type_class)"
@@ -1485,12 +1470,6 @@
   apply intro_classes
 done
 
-lemma "OFCLASS('a, classB_class)"
-  refute  -- {* no countermodel exists *}
-  apply intro_classes
-  apply simp
-done
-
 lemma "OFCLASS('a::type, classC_class)"
   refute
 oops
--- a/src/Pure/Isar/class_target.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/Pure/Isar/class_target.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -56,11 +56,6 @@
   (*tactics*)
   val intro_classes_tac: thm list -> tactic
   val default_intro_tac: Proof.context -> thm list -> tactic
-
-  (*old axclass layer*)
-  val axclass_cmd: binding * xstring list
-    -> (Attrib.binding * string list) list
-    -> theory -> class * theory
 end;
 
 structure Class_Target : CLASS_TARGET =
@@ -629,24 +624,5 @@
   Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac))
     "apply some intro/elim rule"));
 
-
-(** old axclass command **)
-
-fun axclass_cmd (class, raw_superclasses) raw_specs thy =
-  let
-    val _ = Output.legacy_feature "command \"axclass\" deprecated; use \"class\" instead.";
-    val ctxt = ProofContext.init thy;
-    val superclasses = map (Sign.read_class thy) raw_superclasses;
-    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
-      raw_specs;
-    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
-          raw_specs)
-      |> snd
-      |> (map o map) fst;
-  in
-    AxClass.define_class (class, superclasses) []
-      (name_atts ~~ axiomss) thy
-  end;
-
 end;
 
--- a/src/Pure/Isar/code.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/Pure/Isar/code.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -49,10 +49,13 @@
   val add_signature_cmd: string * string -> theory -> theory
   val add_datatype: (string * typ) list -> theory -> theory
   val add_datatype_cmd: string list -> theory -> theory
+  val datatype_interpretation:
+    (string * ((string * sort) list * (string * typ list) list)
+      -> theory -> theory) -> theory -> theory
   val add_abstype: string * typ -> string * typ -> theory -> Proof.state
   val add_abstype_cmd: string -> string -> theory -> Proof.state
-  val type_interpretation:
-    (string * ((string * sort) list * (string * typ list) list)
+  val abstype_interpretation:
+    (string * ((string * sort) list * ((string * typ) * (string * thm)))
       -> theory -> theory) -> theory -> theory
   val add_eqn: thm -> theory -> theory
   val add_nbe_eqn: thm -> theory -> theory
@@ -63,8 +66,8 @@
   val del_eqns: string -> theory -> theory
   val add_case: thm -> theory -> theory
   val add_undefined: string -> theory -> theory
-  val get_datatype: theory -> string -> ((string * sort) list * (string * typ list) list)
-  val get_datatype_of_constr_or_abstr: theory -> string -> (string * bool) option
+  val get_type: theory -> string -> ((string * sort) list * (string * typ list) list)
+  val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option
   val is_constr: theory -> string -> bool
   val is_abstr: theory -> string -> bool
   val get_cert: theory -> ((thm * bool) list -> (thm * bool) list) -> string -> cert
@@ -163,21 +166,21 @@
   signatures: int Symtab.table * typ Symtab.table,
   functions: ((bool * fun_spec) * (serial * fun_spec) list) Symtab.table
     (*with explicit history*),
-  datatypes: ((serial * ((string * sort) list * typ_spec)) list) Symtab.table
+  types: ((serial * ((string * sort) list * typ_spec)) list) Symtab.table
     (*with explicit history*),
   cases: (int * (int * string list)) Symtab.table * unit Symtab.table
 };
 
-fun make_spec (history_concluded, ((signatures, functions), (datatypes, cases))) =
+fun make_spec (history_concluded, ((signatures, functions), (types, cases))) =
   Spec { history_concluded = history_concluded,
-    signatures = signatures, functions = functions, datatypes = datatypes, cases = cases };
+    signatures = signatures, functions = functions, types = types, cases = cases };
 fun map_spec f (Spec { history_concluded = history_concluded, signatures = signatures,
-  functions = functions, datatypes = datatypes, cases = cases }) =
-  make_spec (f (history_concluded, ((signatures, functions), (datatypes, cases))));
+  functions = functions, types = types, cases = cases }) =
+  make_spec (f (history_concluded, ((signatures, functions), (types, cases))));
 fun merge_spec (Spec { history_concluded = _, signatures = (tycos1, sigs1), functions = functions1,
-    datatypes = datatypes1, cases = (cases1, undefs1) },
+    types = types1, cases = (cases1, undefs1) },
   Spec { history_concluded = _, signatures = (tycos2, sigs2), functions = functions2,
-    datatypes = datatypes2, cases = (cases2, undefs2) }) =
+    types = types2, cases = (cases2, undefs2) }) =
   let
     val signatures = (Symtab.merge (op =) (tycos1, tycos2),
       Symtab.merge typ_equiv (sigs1, sigs2));
@@ -190,15 +193,15 @@
           then raw_history else filtered_history;
       in ((false, (snd o hd) history), history) end;
     val functions = Symtab.join (K merge_functions) (functions1, functions2);
-    val datatypes = Symtab.join (K (AList.merge (op =) (K true))) (datatypes1, datatypes2);
+    val types = Symtab.join (K (AList.merge (op =) (K true))) (types1, types2);
     val cases = (Symtab.merge (K true) (cases1, cases2),
       Symtab.merge (K true) (undefs1, undefs2));
-  in make_spec (false, ((signatures, functions), (datatypes, cases))) end;
+  in make_spec (false, ((signatures, functions), (types, cases))) end;
 
 fun history_concluded (Spec { history_concluded, ... }) = history_concluded;
 fun the_signatures (Spec { signatures, ... }) = signatures;
 fun the_functions (Spec { functions, ... }) = functions;
-fun the_datatypes (Spec { datatypes, ... }) = datatypes;
+fun the_types (Spec { types, ... }) = types;
 fun the_cases (Spec { cases, ... }) = cases;
 val map_history_concluded = map_spec o apfst;
 val map_signatures = map_spec o apsnd o apfst o apfst;
@@ -423,11 +426,11 @@
       $ Free ("x", ty_abs)), Free ("x", ty_abs));
   in (tyco, (vs ~~ sorts, ((fst abs_ty, ty), (rep, cert)))) end;    
 
-fun get_datatype_entry thy tyco = case these (Symtab.lookup ((the_datatypes o the_exec) thy) tyco)
+fun get_type_entry thy tyco = case these (Symtab.lookup ((the_types o the_exec) thy) tyco)
  of (_, entry) :: _ => SOME entry
   | _ => NONE;
 
-fun get_datatype_spec thy tyco = case get_datatype_entry thy tyco
+fun get_type_spec thy tyco = case get_type_entry thy tyco
  of SOME (vs, spec) => apfst (pair vs) (constructors_of spec)
   | NONE => arity_number thy tyco
       |> Name.invents Name.context Name.aT
@@ -435,23 +438,23 @@
       |> rpair []
       |> rpair false;
 
-fun get_abstype_spec thy tyco = case get_datatype_entry thy tyco
+fun get_abstype_spec thy tyco = case get_type_entry thy tyco
  of SOME (vs, Abstractor spec) => (vs, spec)
   | NONE => error ("Not an abstract type: " ^ tyco);
  
-fun get_datatype thy = fst o get_datatype_spec thy;
+fun get_type thy = fst o get_type_spec thy;
 
-fun get_datatype_of_constr_or_abstr thy c =
+fun get_type_of_constr_or_abstr thy c =
   case (snd o strip_type o const_typ thy) c
-   of Type (tyco, _) => let val ((vs, cos), abstract) = get_datatype_spec thy tyco
+   of Type (tyco, _) => let val ((vs, cos), abstract) = get_type_spec thy tyco
         in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end
     | _ => NONE;
 
-fun is_constr thy c = case get_datatype_of_constr_or_abstr thy c
+fun is_constr thy c = case get_type_of_constr_or_abstr thy c
  of SOME (_, false) => true
    | _ => false;
 
-fun is_abstr thy c = case get_datatype_of_constr_or_abstr thy c
+fun is_abstr thy c = case get_type_of_constr_or_abstr thy c
  of SOME (_, true) => true
    | _ => false;
 
@@ -952,7 +955,7 @@
       |> Symtab.dest
       |> (map o apsnd) (snd o fst)
       |> sort (string_ord o pairself fst);
-    val datatypes = the_datatypes exec
+    val datatypes = the_types exec
       |> Symtab.dest
       |> map (fn (tyco, (_, (vs, spec)) :: _) =>
           ((tyco, vs), constructors_of spec))
@@ -1088,24 +1091,21 @@
   (map_exec_purge o map_cases o apsnd) (Symtab.update (c, ())) thy;
 
 
-(* datatypes *)
+(* types *)
 
-structure Type_Interpretation =
-  Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
-
-fun register_datatype (tyco, vs_spec) thy =
+fun register_type (tyco, vs_spec) thy =
   let
     val (old_constrs, some_old_proj) =
-      case these (Symtab.lookup ((the_datatypes o the_exec) thy) tyco)
+      case these (Symtab.lookup ((the_types o the_exec) thy) tyco)
        of (_, (_, Constructors cos)) :: _ => (map fst cos, NONE)
         | (_, (_, Abstractor (_, (co, _)))) :: _ => ([], SOME co)
         | [] => ([], NONE)
     val outdated_funs = case some_old_proj
-     of NONE => []
+     of NONE => old_constrs
       | SOME old_proj => Symtab.fold
           (fn (c, ((_, spec), _)) => if member (op =) (the_list (associated_abstype spec)) tyco
             then insert (op =) c else I)
-            ((the_functions o the_exec) thy) [old_proj];
+            ((the_functions o the_exec) thy) (old_proj :: old_constrs);
     fun drop_outdated_cases cases = fold Symtab.delete_safe
       (Symtab.fold (fn (c, (_, (_, cos))) =>
         if exists (member (op =) old_constrs) cos
@@ -1116,13 +1116,15 @@
     |> map_exec_purge
         ((map_typs o Symtab.map_default (tyco, [])) (cons (serial (), vs_spec))
         #> (map_cases o apfst) drop_outdated_cases)
-    |> Type_Interpretation.data (tyco, serial ())
   end;
 
-fun type_interpretation f = Type_Interpretation.interpretation
-  (fn (tyco, _) => fn thy => f (tyco, get_datatype thy tyco) thy);
+fun unoverload_const_typ thy (c, ty) = (AxClass.unoverload_const thy (c, ty), ty);
 
-fun unoverload_const_typ thy (c, ty) = (AxClass.unoverload_const thy (c, ty), ty);
+structure Datatype_Interpretation =
+  Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
+
+fun datatype_interpretation f = Datatype_Interpretation.interpretation
+  (fn (tyco, _) => fn thy => f (tyco, get_type thy tyco) thy);
 
 fun add_datatype proto_constrs thy =
   let
@@ -1131,20 +1133,29 @@
   in
     thy
     |> fold (del_eqns o fst) constrs
-    |> register_datatype (tyco, (vs, Constructors cos))
+    |> register_type (tyco, (vs, Constructors cos))
+    |> Datatype_Interpretation.data (tyco, serial ())
   end;
 
 fun add_datatype_cmd raw_constrs thy =
   add_datatype (map (read_bare_const thy) raw_constrs) thy;
 
+structure Abstype_Interpretation =
+  Interpretation(type T = string * serial val eq = eq_snd (op =) : T * T -> bool);
+
+fun abstype_interpretation f = Abstype_Interpretation.interpretation
+  (fn (tyco, _) => fn thy => f (tyco, get_abstype_spec thy tyco) thy);
+
 fun add_abstype proto_abs proto_rep thy =
   let
     val (abs, rep) = pairself (unoverload_const_typ thy) (proto_abs, proto_rep);
     val (tyco, (vs, (abs_ty as (abs, ty), (rep, cert_prop)))) = abstype_cert thy abs (fst rep);
     fun after_qed [[cert]] = ProofContext.theory
-      (register_datatype (tyco, (vs, Abstractor (abs_ty, (rep, cert))))
+      (del_eqns abs
+      #> register_type (tyco, (vs, Abstractor (abs_ty, (rep, cert))))
       #> change_fun_spec false rep ((K o Proj)
-        (map_types Logic.varifyT (mk_proj tyco vs ty abs rep), tyco)));
+        (map_types Logic.varifyT (mk_proj tyco vs ty abs rep), tyco))
+      #> Abstype_Interpretation.data (tyco, serial ()));
   in
     thy
     |> ProofContext.init
@@ -1188,7 +1199,7 @@
            (mk_attribute o code_target_attr))
       || Scan.succeed (mk_attribute add_warning_eqn);
   in
-    Type_Interpretation.init
+    Datatype_Interpretation.init
     #> Attrib.setup (Binding.name "code") (Scan.lift code_attribute_parser)
         "declare theorems for code generation"
   end));
--- a/src/Pure/Isar/isar_syn.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/Pure/Isar/isar_syn.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -99,13 +99,6 @@
   OuterSyntax.command "defaultsort" "declare default sort" K.thy_decl
     (P.sort >> (Toplevel.theory o Sign.add_defsort));
 
-val _ =
-  OuterSyntax.command "axclass" "define axiomatic type class" K.thy_decl
-    (P.binding -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
-        P.!!! (P.list1 P.xname)) []
-        -- Scan.repeat (SpecParse.thm_name ":" -- (P.prop >> single))
-      >> (fn (x, y) => Toplevel.theory (snd o Class.axclass_cmd x y)));
-
 
 (* types *)
 
--- a/src/Tools/Code/code_eval.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/Tools/Code/code_eval.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -130,7 +130,7 @@
     val thy = ProofContext.theory_of background;
     val tyco = Sign.intern_type thy raw_tyco;
     val constrs = map (Code.check_const thy) raw_constrs;
-    val constrs' = (map fst o snd o Code.get_datatype thy) tyco;
+    val constrs' = (map fst o snd o Code.get_type thy) tyco;
     val _ = if eq_set (op =) (constrs, constrs') then ()
       else error ("Type " ^ quote tyco ^ ": given constructors diverge from real constructors")
     val is_first = is_first_occ background;
--- a/src/Tools/Code/code_thingol.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/Tools/Code/code_thingol.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -256,7 +256,7 @@
     | thyname :: _ => thyname;
   fun thyname_of_const thy c = case AxClass.class_of_param thy c
    of SOME class => thyname_of_class thy class
-    | NONE => (case Code.get_datatype_of_constr_or_abstr thy c
+    | NONE => (case Code.get_type_of_constr_or_abstr thy c
        of SOME (tyco, _) => Codegen.thyname_of_type thy tyco
         | NONE => Codegen.thyname_of_const thy c);
   fun purify_base "==>" = "follows"
@@ -543,7 +543,7 @@
   let
     val stmt_datatype =
       let
-        val (vs, cos) = Code.get_datatype thy tyco;
+        val (vs, cos) = Code.get_type thy tyco;
       in
         fold_map (translate_tyvar_sort thy algbr eqngr) vs
         ##>> fold_map (fn (c, tys) =>
@@ -569,7 +569,7 @@
         ##>> fold_map (translate_eqn thy algbr eqngr) eqns
         #>> (fn info => Fun (c, info))
       end;
-    val stmt_const = case Code.get_datatype_of_constr_or_abstr thy c
+    val stmt_const = case Code.get_type_of_constr_or_abstr thy c
      of SOME (tyco, _) => stmt_datatypecons tyco
       | NONE => (case AxClass.class_of_param thy c
          of SOME class => stmt_classparam class
--- a/src/Tools/quickcheck.ML	Tue Feb 23 14:44:24 2010 -0800
+++ b/src/Tools/quickcheck.ML	Tue Feb 23 14:44:43 2010 -0800
@@ -10,6 +10,8 @@
   val timing : bool Unsynchronized.ref
   val test_term: Proof.context -> bool -> string option -> int -> int -> term ->
     (string * term) list option
+  val timed_test_term: Proof.context -> bool -> string option -> int -> int -> term ->
+    ((string * term) list option * (string * int) list)
   val add_generator: string * (Proof.context -> term -> int -> term list option) -> theory -> theory
   val setup: theory -> theory
   val quickcheck: (string * string) list -> int -> Proof.state -> (string * term) list option
@@ -97,13 +99,20 @@
     val frees = Term.add_frees t [];
   in (map fst frees, list_abs_free (frees, t)) end
 
-fun test_term ctxt quiet generator_name size i t =
+fun cpu_time description f =
+  let
+    val start = start_timing ()
+    val result = Exn.capture f ()
+    val time = Time.toMilliseconds (#cpu (end_timing start))
+  in (Exn.release result, (description, time)) end
+
+fun timed_test_term ctxt quiet generator_name size i t =
   let
     val (names, t') = prep_test_term t;
-    val testers = (*cond_timeit (!timing) "quickcheck compilation"
-      (fn () => *)(case generator_name
+    val (testers, comp_time) = cpu_time "quickcheck compilation"
+      (fn () => (case generator_name
        of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t'
-        | SOME name => [mk_tester_select name ctxt t']);
+        | SOME name => [mk_tester_select name ctxt t']));
     fun iterate f 0 = NONE
       | iterate f j = case f () handle Match => (if quiet then ()
              else warning "Exception Match raised during quickcheck"; NONE)
@@ -117,13 +126,17 @@
       else (if quiet then () else priority ("Test data size: " ^ string_of_int k);
         case with_testers k testers
          of NONE => with_size (k + 1) | SOME q => SOME q);
+    val (result, exec_time) = cpu_time "quickcheck execution"
+      (fn () => case with_size 1
+        of NONE => NONE
+        | SOME ts => SOME (names ~~ ts))
   in
-    cond_timeit (!timing) "quickcheck execution"
-    (fn () => case with_size 1
-      of NONE => NONE
-      | SOME ts => SOME (names ~~ ts))
+    (result, [exec_time, comp_time])
   end;
 
+fun test_term ctxt quiet generator_name size i t =
+  fst (timed_test_term ctxt quiet generator_name size i t)
+
 fun monomorphic_term thy insts default_T = 
   let
     fun subst (T as TFree (v, S)) =