merged
authorblanchet
Thu, 29 Oct 2009 12:09:32 +0100
changeset 33566 1c62ac4ef6d1
parent 33300 939ca97f5a11 (current diff)
parent 33565 5fad8e36dfb1 (diff)
child 33567 3b8fc89a52b7
merged
src/HOL/IsaMakefile
src/HOL/Tools/Nitpick/nitpick.ML
src/HOL/Tools/Nitpick/nitpick_isar.ML
src/Tools/quickcheck.ML
--- a/doc-src/Nitpick/nitpick.tex	Thu Oct 29 11:26:47 2009 +0100
+++ b/doc-src/Nitpick/nitpick.tex	Thu Oct 29 12:09:32 2009 +0100
@@ -2,7 +2,7 @@
 \usepackage[T1]{fontenc}
 \usepackage{amsmath}
 \usepackage{amssymb}
-\usepackage[french,english]{babel}
+\usepackage[english,french]{babel}
 \usepackage{color}
 \usepackage{graphicx}
 %\usepackage{mathpazo}
@@ -40,6 +40,8 @@
 
 \begin{document}
 
+\selectlanguage{english}
+
 \title{\includegraphics[scale=0.5]{isabelle_nitpick} \\[4ex]
 Picking Nits \\[\smallskipamount]
 \Large A User's Guide to Nitpick for Isabelle/HOL 2010}
@@ -112,6 +114,13 @@
 machine called \texttt{java}. The examples presented in this manual can be found
 in Isabelle's \texttt{src/HOL/Nitpick\_Examples/Manual\_Nits.thy} theory.
 
+Throughout this manual, we will explicitly invoke the \textbf{nitpick} command.
+Nitpick also provides an automatic mode that can be enabled using the
+``Auto Nitpick'' option from the ``Isabelle'' menu in Proof General. In this
+mode, Nitpick is run on every newly entered theorem, much like Auto Quickcheck.
+The collective time limit for Auto Nitpick and Auto Quickcheck can be set using
+the ``Auto Counterexample Time Limit'' option.
+
 \newbox\boxA
 \setbox\boxA=\hbox{\texttt{nospam}}
 
@@ -154,16 +163,6 @@
 configured SAT solvers in Isabelle (e.g., for Refute), these will also be
 available to Nitpick.
 
-Throughout this manual, we will explicitly invoke the \textbf{nitpick} command.
-Nitpick also provides an automatic mode that can be enabled by specifying
-
-\prew
-\textbf{nitpick\_params} [\textit{auto}]
-\postw
-
-at the beginning of the theory file. In this mode, Nitpick is run for up to 5
-seconds (by default) on every newly entered theorem, much like Auto Quickcheck.
-
 \subsection{Propositional Logic}
 \label{propositional-logic}
 
@@ -1595,6 +1594,17 @@
 (\S\ref{authentication}), optimizations
 (\S\ref{optimizations}), and timeouts (\S\ref{timeouts}).
 
+You can instruct Nitpick to run automatically on newly entered theorems by
+enabling the ``Auto Nitpick'' option from the ``Isabelle'' menu in Proof
+General. For automatic runs, \textit{user\_axioms} (\S\ref{mode-of-operation})
+and \textit{assms} (\S\ref{mode-of-operation}) are implicitly enabled,
+\textit{blocking} (\S\ref{mode-of-operation}), \textit{verbose}
+(\S\ref{output-format}), and \textit{debug} (\S\ref{output-format}) are
+disabled, \textit{max\_potential} (\S\ref{output-format}) is taken to be 0, and
+\textit{timeout} (\S\ref{timeouts}) is superseded by the ``Auto Counterexample
+Time Limit'' in Proof General's ``Isabelle'' menu. Nitpick's output is also more
+concise.
+
 The number of options can be overwhelming at first glance. Do not let that worry
 you: Nitpick's defaults have been chosen so that it almost always does the right
 thing, and the most important options have been covered in context in
@@ -1622,35 +1632,19 @@
 \end{enum}
 
 Default values are indicated in square brackets. Boolean options have a negated
-counterpart (e.g., \textit{auto} vs.\ \textit{no\_auto}). When setting Boolean
-options, ``= \textit{true}'' may be omitted.
+counterpart (e.g., \textit{blocking} vs.\ \textit{no\_blocking}). When setting
+Boolean options, ``= \textit{true}'' may be omitted.
 
 \subsection{Mode of Operation}
 \label{mode-of-operation}
 
 \begin{enum}
-\opfalse{auto}{no\_auto}
-Specifies whether Nitpick should be run automatically on newly entered theorems.
-For automatic runs, \textit{user\_axioms} (\S\ref{mode-of-operation}) and
-\textit{assms} (\S\ref{mode-of-operation}) are implicitly enabled,
-\textit{blocking} (\S\ref{mode-of-operation}), \textit{verbose}
-(\S\ref{output-format}), and \textit{debug} (\S\ref{output-format}) are
-disabled, \textit{max\_potential} (\S\ref{output-format}) is taken to be 0, and
-\textit{auto\_timeout} (\S\ref{timeouts}) is used as the time limit instead of
-\textit{timeout} (\S\ref{timeouts}). The output is also more concise.
-
-\nopagebreak
-{\small See also \textit{auto\_timeout} (\S\ref{timeouts}).}
-
 \optrue{blocking}{non\_blocking}
 Specifies whether the \textbf{nitpick} command should operate synchronously.
 The asynchronous (non-blocking) mode lets the user start proving the putative
 theorem while Nitpick looks for a counterexample, but it can also be more
 confusing. For technical reasons, automatic runs currently always block.
 
-\nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation}).}
-
 \optrue{falsify}{satisfy}
 Specifies whether Nitpick should look for falsifying examples (countermodels) or
 satisfying examples (models). This manual assumes throughout that
@@ -1670,16 +1664,15 @@
 considered.
 
 \nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation}), \textit{assms}
-(\S\ref{mode-of-operation}), and \textit{debug} (\S\ref{output-format}).}
+{\small See also \textit{assms} (\S\ref{mode-of-operation}) and \textit{debug}
+(\S\ref{output-format}).}
 
 \optrue{assms}{no\_assms}
 Specifies whether the relevant assumptions in structured proof should be
 considered. The option is implicitly enabled for automatic runs.
 
 \nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation})
-and \textit{user\_axioms} (\S\ref{mode-of-operation}).}
+{\small See also \textit{user\_axioms} (\S\ref{mode-of-operation}).}
 
 \opfalse{overlord}{no\_overlord}
 Specifies whether Nitpick should put its temporary files in
@@ -1836,14 +1829,14 @@
 
 \nopagebreak
 {\small See also \textit{card} (\S\ref{scope-of-search}),
-\textit{coalesce\_type\_vars} (\S\ref{scope-of-search}), and \textit{verbose}
+\textit{merge\_type\_vars} (\S\ref{scope-of-search}), and \textit{verbose}
 (\S\ref{output-format}).}
 
 \opsmart{mono}{non\_box}
 Specifies the default monotonicity setting to use. This can be overridden on a
 per-type basis using the \textit{mono}~\qty{type} option described above.
 
-\opfalse{coalesce\_type\_vars}{dont\_coalesce\_type\_vars}
+\opfalse{merge\_type\_vars}{dont\_merge\_type\_vars}
 Specifies whether type variables with the same sort constraints should be
 merged. Setting this option to \textit{true} can reduce the number of scopes
 tried and the size of the generated Kodkod formulas, but it also diminishes the
@@ -1861,9 +1854,6 @@
 option is useful to determine which scopes are tried or which SAT solver is
 used. This option is implicitly disabled for automatic runs.
 
-\nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation}).}
-
 \opfalse{debug}{no\_debug}
 Specifies whether Nitpick should display additional debugging information beyond
 what \textit{verbose} already displays. Enabling \textit{debug} also enables
@@ -1871,8 +1861,8 @@
 option is implicitly disabled for automatic runs.
 
 \nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation}), \textit{overlord}
-(\S\ref{mode-of-operation}), and \textit{batch\_size} (\S\ref{optimizations}).}
+{\small See also \textit{overlord} (\S\ref{mode-of-operation}) and
+\textit{batch\_size} (\S\ref{optimizations}).}
 
 \optrue{show\_skolems}{hide\_skolem}
 Specifies whether the values of Skolem constants should be displayed as part of
@@ -1910,8 +1900,7 @@
 enabled.
 
 \nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation}),
-\textit{check\_potential} (\S\ref{authentication}), and
+{\small See also \textit{check\_potential} (\S\ref{authentication}) and
 \textit{sat\_solver} (\S\ref{optimizations}).}
 
 \opt{max\_genuine}{int}{$\mathbf{1}$}
@@ -1968,8 +1957,7 @@
 shown to be genuine, Nitpick displays a message to this effect and terminates.
 
 \nopagebreak
-{\small See also \textit{max\_potential} (\S\ref{output-format}) and
-\textit{auto\_timeout} (\S\ref{timeouts}).}
+{\small See also \textit{max\_potential} (\S\ref{output-format}).}
 
 \opfalse{check\_genuine}{trust\_genuine}
 Specifies whether genuine and likely genuine counterexamples should be given to
@@ -1979,8 +1967,7 @@
 \texttt{blan{\color{white}nospam}\kern-\wd\boxA{}chette@in.tum.de}.
 
 \nopagebreak
-{\small See also \textit{max\_genuine} (\S\ref{output-format}) and
-\textit{auto\_timeout} (\S\ref{timeouts}).}
+{\small See also \textit{max\_genuine} (\S\ref{output-format}).}
 
 \ops{expect}{string}
 Specifies the expected outcome, which must be one of the following:
@@ -2206,21 +2193,14 @@
 Specifies the maximum amount of time that the \textbf{nitpick} command should
 spend looking for a counterexample. Nitpick tries to honor this constraint as
 well as it can but offers no guarantees. For automatic runs,
-\textit{auto\_timeout} is used instead.
+\textit{timeout} is ignored; instead, Auto Quickcheck and Auto Nitpick share
+a time slot whose length is specified by the ``Auto Counterexample Time
+Limit'' option in Proof General.
 
 \nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation})
-and \textit{max\_threads} (\S\ref{optimizations}).}
-
-\opt{auto\_timeout}{time}{$\mathbf{5}$ s}
-Specifies the maximum amount of time that Nitpick should use to find a
-counterexample when running automatically. Nitpick tries to honor this
-constraint as well as it can but offers no guarantees.
-
-\nopagebreak
-{\small See also \textit{auto} (\S\ref{mode-of-operation}).}
-
-\opt{tac\_timeout}{time}{$\mathbf{500}$ ms}
+{\small See also \textit{max\_threads} (\S\ref{optimizations}).}
+
+\opt{tac\_timeout}{time}{$\mathbf{500}$\,ms}
 Specifies the maximum amount of time that the \textit{auto} tactic should use
 when checking a counterexample, and similarly that \textit{lexicographic\_order}
 and \textit{sizechange} should use when checking whether a (co)in\-duc\-tive
@@ -2398,11 +2378,11 @@
 & \textit{subgoal}\end{aligned}$
 \postw
 
+\let\antiq=\textrm
+
 \subsection{Registration of Coinductive Datatypes}
 \label{registration-of-coinductive-datatypes}
 
-\let\antiq=\textrm
-
 If you have defined a custom coinductive datatype, you can tell Nitpick about
 it, so that it can use an efficient Kodkod axiomatization similar to the one it
 uses for lazy lists. The interface for registering and unregistering coinductive
@@ -2467,6 +2447,12 @@
 representation of functions synthesized by Isabelle, which is an implementation
 detail.
 
+\item[$\bullet$] Nitpick maintains a global cache of wellfoundedness conditions,
+which can become invalid if you change the definition of an inductive predicate
+that is registered in the cache. To clear the cache,
+run Nitpick with the \textit{tac\_timeout} option set to a new value (e.g.,
+501$\,\textit{ms}$).
+
 \item[$\bullet$] Nitpick produces spurious counterexamples when invoked after a
 \textbf{guess} command in a structured proof.
 
Binary file doc-src/gfx/isabelle_nitpick.pdf has changed
--- a/src/HOL/IsaMakefile	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/IsaMakefile	Thu Oct 29 12:09:32 2009 +0100
@@ -102,6 +102,7 @@
   $(SRC)/Provers/hypsubst.ML \
   $(SRC)/Provers/quantifier1.ML \
   $(SRC)/Provers/splitter.ML \
+  $(SRC)/Tools/Auto_Counterexample.thy \
   $(SRC)/Tools/Code/code_haskell.ML \
   $(SRC)/Tools/Code/code_ml.ML \
   $(SRC)/Tools/Code/code_preproc.ML \
@@ -114,6 +115,7 @@
   $(SRC)/Tools/IsaPlanner/rw_tools.ML \
   $(SRC)/Tools/IsaPlanner/zipper.ML \
   $(SRC)/Tools/atomize_elim.ML \
+  $(SRC)/Tools/auto_counterexample.ML \
   $(SRC)/Tools/auto_solve.ML \
   $(SRC)/Tools/coherent.ML \
   $(SRC)/Tools/cong_tac.ML \
--- a/src/HOL/Main.thy	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Main.thy	Thu Oct 29 12:09:32 2009 +0100
@@ -1,7 +1,7 @@
 header {* Main HOL *}
 
 theory Main
-imports Plain Nitpick Predicate_Compile Recdef
+imports Plain Nitpick
 begin
 
 text {*
--- a/src/HOL/Nitpick.thy	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Nitpick.thy	Thu Oct 29 12:09:32 2009 +0100
@@ -8,7 +8,7 @@
 header {* Nitpick: Yet Another Counterexample Generator for Isabelle/HOL *}
 
 theory Nitpick
-imports Map SAT
+imports Map Quickcheck SAT
 uses ("Tools/Nitpick/kodkod.ML")
      ("Tools/Nitpick/kodkod_sat.ML")
      ("Tools/Nitpick/nitpick_util.ML")
@@ -117,12 +117,16 @@
  apply (simp only: unit.cases)
 by simp
 
+declare unit.cases [nitpick_simp del]
+
 lemma nat_case_def [nitpick_def]:
 "nat_case x f n \<equiv> if n = 0 then x else f (n - 1)"
 apply (rule eq_reflection)
 by (case_tac n) auto
 
-lemmas dvd_def = dvd_eq_mod_eq_0 [THEN eq_reflection, nitpick_def]
+declare nat.cases [nitpick_simp del]
+
+lemmas [nitpick_def] = dvd_eq_mod_eq_0 [THEN eq_reflection]
 
 lemma list_size_simp [nitpick_simp]:
 "list_size f xs = (if xs = [] then 0
@@ -206,6 +210,21 @@
 definition of_frac :: "'a \<Rightarrow> 'b\<Colon>{inverse,ring_1}" where
 "of_frac q \<equiv> of_int (num q) / of_int (denom q)"
 
+(* While Nitpick normally avoids to unfold definitions for locales, it
+   unfortunately needs to unfold them when dealing with the following built-in
+   constants. A cleaner approach would be to change "Nitpick_HOL" and
+   "Nitpick_Nits" so that they handle the unexpanded overloaded constants
+   directly, but this is slightly more tricky to implement. *)
+lemmas [nitpick_def] = div_int_inst.div_int div_int_inst.mod_int
+    div_nat_inst.div_nat div_nat_inst.mod_nat lower_semilattice_fun_inst.inf_fun
+    minus_fun_inst.minus_fun minus_int_inst.minus_int minus_nat_inst.minus_nat
+    one_int_inst.one_int one_nat_inst.one_nat ord_fun_inst.less_eq_fun
+    ord_int_inst.less_eq_int ord_int_inst.less_int ord_nat_inst.less_eq_nat
+    ord_nat_inst.less_nat plus_int_inst.plus_int plus_nat_inst.plus_nat
+    times_int_inst.times_int times_nat_inst.times_nat uminus_int_inst.uminus_int
+    upper_semilattice_fun_inst.sup_fun zero_int_inst.zero_int
+    zero_nat_inst.zero_nat
+
 use "Tools/Nitpick/kodkod.ML"
 use "Tools/Nitpick/kodkod_sat.ML"
 use "Tools/Nitpick/nitpick_util.ML"
@@ -222,6 +241,8 @@
 use "Tools/Nitpick/nitpick_tests.ML"
 use "Tools/Nitpick/minipick.ML"
 
+setup {* Nitpick_Isar.setup *}
+
 hide (open) const unknown undefined_fast_The undefined_fast_Eps bisim 
     bisim_iterator_max Tha 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
@@ -230,10 +251,10 @@
 hide (open) type bisim_iterator pair_box fun_box
 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
-    The_psimp Eps_psimp unit_case_def nat_case_def dvd_def list_size_simp
-    nat_gcd_def nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def
-    one_frac_def num_def denom_def norm_frac_def frac_def plus_frac_def
-    times_frac_def uminus_frac_def number_of_frac_def inverse_frac_def
-    less_eq_frac_def of_frac_def
+    The_psimp Eps_psimp unit_case_def nat_case_def list_size_simp nat_gcd_def
+    nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def one_frac_def
+    num_def denom_def norm_frac_def frac_def plus_frac_def times_frac_def
+    uminus_frac_def number_of_frac_def inverse_frac_def less_eq_frac_def
+    of_frac_def
 
 end
--- a/src/HOL/Quickcheck.thy	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Quickcheck.thy	Thu Oct 29 12:09:32 2009 +0100
@@ -126,6 +126,8 @@
   shows "random_aux k = rhs k"
   using assms by (rule code_numeral.induct)
 
+setup {* Quickcheck.setup *}
+
 subsection {* the Random-Predicate Monad *} 
 
 types 'a randompred = "Random.seed \<Rightarrow> ('a Predicate.pred \<times> Random.seed)"
--- a/src/HOL/Tools/Nitpick/HISTORY	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/HISTORY	Thu Oct 29 12:09:32 2009 +0100
@@ -6,7 +6,11 @@
     "nitpick_ind_intro" to "nitpick_intro"
   * Replaced "special_depth" and "skolemize_depth" options by "specialize"
     and "skolemize"
+  * Renamed "coalesce_type_vars" to "merge_type_vars"
+  * Optimized Kodkod encoding of datatypes whose constructors don't appear in
+    the formula to falsify
   * Fixed monotonicity check
+  * Fixed error in display of uncurried constants
 
 Version 1.2.2 (16 Oct 2009)
 
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -23,7 +23,7 @@
     overlord: bool,
     user_axioms: bool option,
     assms: bool,
-    coalesce_type_vars: bool,
+    merge_type_vars: bool,
     destroy_constrs: bool,
     specialize: bool,
     skolemize: bool,
@@ -88,7 +88,7 @@
   overlord: bool,
   user_axioms: bool option,
   assms: bool,
-  coalesce_type_vars: bool,
+  merge_type_vars: bool,
   destroy_constrs: bool,
   specialize: bool,
   skolemize: bool,
@@ -175,7 +175,7 @@
     val ctxt = Proof.context_of state
     val {cards_assigns, maxes_assigns, iters_assigns, bisim_depths, boxes,
          monos, wfs, sat_solver, blocking, falsify, debug, verbose, overlord,
-         user_axioms, assms, coalesce_type_vars, destroy_constrs, specialize,
+         user_axioms, assms, merge_type_vars, destroy_constrs, specialize,
          skolemize, star_linear_preds, uncurry, fast_descrs, peephole_optim,
          tac_timeout, sym_break, sharing_depth, flatten_props, max_threads,
          show_skolems, show_datatypes, show_consts, evals, formats,
@@ -187,7 +187,7 @@
     val pprint =
       if auto then
         Unsynchronized.change state_ref o Proof.goal_message o K
-        o curry Pretty.blk 0 o cons (Pretty.str "") o single
+        o Pretty.chunks o cons (Pretty.str "") o single
         o Pretty.mark Markup.hilite
       else
         priority o Pretty.string_of
@@ -220,7 +220,7 @@
                     neg_t
     val (assms_t, evals) =
       assms_t :: evals
-      |> coalesce_type_vars ? coalesce_type_vars_in_terms
+      |> merge_type_vars ? merge_type_vars_in_terms
       |> hd pairf tl
     val original_max_potential = max_potential
     val original_max_genuine = max_genuine
@@ -283,6 +283,21 @@
             handle TYPE (_, Ts, ts) =>
                    raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts)
 
+    val core_u = nut_from_term thy fast_descrs (!special_funs) Eq core_t
+    val def_us = map (nut_from_term thy fast_descrs (!special_funs) DefEq)
+                     def_ts
+    val nondef_us = map (nut_from_term thy fast_descrs (!special_funs) Eq)
+                        nondef_ts
+    val (free_names, const_names) =
+      fold add_free_and_const_names (core_u :: def_us @ nondef_us) ([], [])
+    val (sel_names, nonsel_names) =
+      List.partition (is_sel o nickname_of) const_names
+    val would_be_genuine = got_all_user_axioms andalso none_true wfs
+(*
+    val _ = List.app (priority o string_for_nut ctxt)
+                     (core_u :: def_us @ nondef_us)
+*)
+
     val unique_scope = forall (equal 1 o length o snd) cards_assigns
     (* typ -> bool *)
     fun is_free_type_monotonic T =
@@ -298,6 +313,8 @@
         not (is_pure_typedef thy T) orelse is_univ_typedef thy T
         orelse is_number_type thy T
         orelse formulas_monotonic ext_ctxt T def_ts nondef_ts core_t
+    fun is_datatype_shallow T =
+      not (exists (equal T o domain_type o type_of) sel_names)
     val Ts = ground_types_in_terms ext_ctxt (core_t :: def_ts @ nondef_ts)
              |> sort TermOrd.typ_ord
     val (all_dataTs, all_free_Ts) =
@@ -326,7 +343,7 @@
         ()
     val mono_Ts = mono_dataTs @ mono_free_Ts
     val nonmono_Ts = nonmono_dataTs @ nonmono_free_Ts
-
+    val shallow_dataTs = filter is_datatype_shallow mono_dataTs
 (*
     val _ = priority "Monotonic datatypes:"
     val _ = List.app (priority o string_for_type ctxt) mono_dataTs
@@ -338,19 +355,6 @@
     val _ = List.app (priority o string_for_type ctxt) nonmono_free_Ts
 *)
 
-    val core_u = nut_from_term thy fast_descrs (!special_funs) Eq core_t
-    val def_us = map (nut_from_term thy fast_descrs (!special_funs) DefEq)
-                     def_ts
-    val nondef_us = map (nut_from_term thy fast_descrs (!special_funs) Eq)
-                        nondef_ts
-    val (free_names, const_names) =
-      fold add_free_and_const_names (core_u :: def_us @ nondef_us) ([], [])
-    val nonsel_names = filter_out (is_sel o nickname_of) const_names
-    val would_be_genuine = got_all_user_axioms andalso none_true wfs
-(*
-    val _ = List.app (priority o string_for_nut ctxt)
-                     (core_u :: def_us @ nondef_us)
-*)
     val need_incremental = Int.max (max_potential, max_genuine) >= 2
     val effective_sat_solver =
       if sat_solver <> "smart" then
@@ -778,11 +782,9 @@
                    case scope_count (do_filter (!generated_problems)) scope of
                      0 => I
                    | n =>
-                     if scope_count (do_filter (these (!checked_problems)))
-                                    scope = n then
-                       Integer.add 1
-                     else
-                       I) (!generated_scopes) 0
+                     scope_count (do_filter (these (!checked_problems)))
+                                 scope = n
+                     ? Integer.add 1) (!generated_scopes) 0
       in
         "Nitpick " ^ did_so_and_so ^
         (if is_some (!checked_problems) andalso total > 0 then
@@ -814,6 +816,7 @@
 
     val _ = scopes := all_scopes ext_ctxt sym_break cards_assigns maxes_assigns
                                  iters_assigns bisim_depths mono_Ts nonmono_Ts
+                                 shallow_dataTs
     val batches = batch_list batch_size (!scopes)
     val outcome_code =
       (run_batches 0 (length batches) batches (max_potential, max_genuine, 0)
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -125,7 +125,7 @@
   val is_inductive_pred : extended_context -> styp -> bool
   val is_constr_pattern_lhs : theory -> term -> bool
   val is_constr_pattern_formula : theory -> term -> bool
-  val coalesce_type_vars_in_terms : term list -> term list
+  val merge_type_vars_in_terms : term list -> term list
   val ground_types_in_type : extended_context -> typ -> typ list
   val ground_types_in_terms : extended_context -> term list -> typ list
   val format_type : int list -> int list -> typ -> typ
@@ -1016,24 +1016,6 @@
       | do_eq _ = false
   in do_eq end
 
-(* This table is not pretty. A better approach would be to avoid expanding the
-   operators to their low-level definitions, but this would require dealing with
-   overloading. *)
-val built_in_built_in_defs =
-  [@{thm div_int_inst.div_int}, @{thm div_int_inst.mod_int},
-   @{thm div_nat_inst.div_nat}, @{thm div_nat_inst.mod_nat},
-   @{thm lower_semilattice_fun_inst.inf_fun}, @{thm minus_fun_inst.minus_fun},
-   @{thm minus_int_inst.minus_int}, @{thm minus_nat_inst.minus_nat},
-   @{thm one_int_inst.one_int}, @{thm one_nat_inst.one_nat},
-   @{thm ord_fun_inst.less_eq_fun}, @{thm ord_int_inst.less_eq_int},
-   @{thm ord_int_inst.less_int}, @{thm ord_nat_inst.less_eq_nat},
-   @{thm ord_nat_inst.less_nat}, @{thm plus_int_inst.plus_int},
-   @{thm plus_nat_inst.plus_nat}, @{thm times_int_inst.times_int},
-   @{thm times_nat_inst.times_nat}, @{thm uminus_int_inst.uminus_int},
-   @{thm upper_semilattice_fun_inst.sup_fun}, @{thm zero_int_inst.zero_int},
-   @{thm zero_nat_inst.zero_nat}]
-  |> map prop_of
-
 (* theory -> term list * term list * term list *)
 fun all_axioms_of thy =
   let
@@ -1054,8 +1036,7 @@
     val (built_in_nondefs, user_nondefs) =
       List.partition (is_typedef_axiom thy false) user_nondefs
       |>> append built_in_nondefs
-    val defs = built_in_built_in_defs @
-               (thy |> PureThy.all_thms_of
+    val defs = (thy |> PureThy.all_thms_of
                     |> filter (equal Thm.definitionK o Thm.get_kind o snd)
                     |> map (prop_of o snd) |> filter is_plain_definition) @
                user_defs @ built_in_defs
@@ -1309,10 +1290,6 @@
     |> fold_rev (curry absdummy) (func_Ts @ [dataT])
   end
 
-val redefined_in_Nitpick_thy =
-  [@{const_name option_case}, @{const_name nat_case}, @{const_name list_case},
-   @{const_name list_size}]
-
 (* theory -> string -> typ -> typ -> term -> term *)
 fun optimized_record_get thy s rec_T res_T t =
   let val constr_x = the_single (datatype_constrs thy rec_T) in
@@ -1382,7 +1359,6 @@
   (is_real_equational_fun ext_ctxt orf is_real_inductive_pred ext_ctxt
    orf (String.isPrefix ubfp_prefix orf String.isPrefix lbfp_prefix) o fst)
   andf (not o has_trivial_definition thy def_table)
-  andf (not o member (op =) redefined_in_Nitpick_thy o fst)
 
 (* term * term -> term *)
 fun s_betapply (Const (@{const_name If}, _) $ @{const True} $ t, _) = t
@@ -1630,6 +1606,7 @@
                                  (tac ctxt (auto_tac (clasimpset_of ctxt))))
        #> the #> Goal.finish ctxt) goal
 
+val max_cached_wfs = 100
 val cached_timeout = Unsynchronized.ref (SOME Time.zeroTime)
 val cached_wf_props : (term * bool) list Unsynchronized.ref =
   Unsynchronized.ref []
@@ -1663,8 +1640,11 @@
                  else
                    ()
        in
-         if tac_timeout = (!cached_timeout) then ()
-         else (cached_wf_props := []; cached_timeout := tac_timeout);
+         if tac_timeout = (!cached_timeout)
+            andalso length (!cached_wf_props) < max_cached_wfs then
+           ()
+         else
+           (cached_wf_props := []; cached_timeout := tac_timeout);
          case AList.lookup (op =) (!cached_wf_props) prop of
            SOME wf => wf
          | NONE =>
@@ -1879,9 +1859,7 @@
 (* extended_context -> styp -> term list *)
 fun raw_equational_fun_axioms (ext_ctxt as {thy, fast_descrs, simp_table,
                                             psimp_table, ...}) (x as (s, _)) =
-  if s mem redefined_in_Nitpick_thy then
-    []
-  else case def_props_for_const thy fast_descrs (!simp_table) x of
+  case def_props_for_const thy fast_descrs (!simp_table) x of
     [] => (case def_props_for_const thy fast_descrs psimp_table x of
              [] => [inductive_pred_axiom ext_ctxt x]
            | psimps => psimps)
@@ -1890,7 +1868,7 @@
 val equational_fun_axioms = map extensionalize oo raw_equational_fun_axioms
 
 (* term list -> term list *)
-fun coalesce_type_vars_in_terms ts =
+fun merge_type_vars_in_terms ts =
   let
     (* typ -> (sort * string) list -> (sort * string) list *)
     fun add_type (TFree (s, S)) table =
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -10,7 +10,9 @@
 sig
   type params = Nitpick.params
 
+  val auto: bool Unsynchronized.ref
   val default_params : theory -> (string * string) list -> params
+  val setup : theory -> theory
 end
 
 structure Nitpick_Isar : NITPICK_ISAR =
@@ -22,6 +24,14 @@
 open Nitpick_Nut
 open Nitpick
 
+val auto = Unsynchronized.ref false;
+
+val _ = ProofGeneralPgip.add_preference Preferences.category_tracing
+            (setmp_CRITICAL auto false
+                 (fn () => Preferences.bool_pref auto
+                               "auto-nitpick"
+                               "Whether to run Nitpick automatically.") ())
+
 type raw_param = string * string list
 
 val default_default_params =
@@ -33,12 +43,11 @@
    ("wf", ["smart"]),
    ("sat_solver", ["smart"]),
    ("batch_size", ["smart"]),
-   ("auto", ["false"]),
    ("blocking", ["true"]),
    ("falsify", ["true"]),
    ("user_axioms", ["smart"]),
    ("assms", ["true"]),
-   ("coalesce_type_vars", ["false"]),
+   ("merge_type_vars", ["false"]),
    ("destroy_constrs", ["true"]),
    ("specialize", ["true"]),
    ("skolemize", ["true"]),
@@ -47,7 +56,6 @@
    ("fast_descrs", ["true"]),
    ("peephole_optim", ["true"]),
    ("timeout", ["30 s"]),
-   ("auto_timeout", ["5 s"]),
    ("tac_timeout", ["500 ms"]),
    ("sym_break", ["20"]),
    ("sharing_depth", ["3"]),
@@ -70,12 +78,11 @@
   [("dont_box", "box"),
    ("non_mono", "mono"),
    ("non_wf", "wf"),
-   ("no_auto", "auto"),
    ("non_blocking", "blocking"),
    ("satisfy", "falsify"),
    ("no_user_axioms", "user_axioms"),
    ("no_assms", "assms"),
-   ("dont_coalesce_type_vars", "coalesce_type_vars"),
+   ("dont_merge_type_vars", "merge_type_vars"),
    ("dont_destroy_constrs", "destroy_constrs"),
    ("dont_specialize", "specialize"),
    ("dont_skolemize", "skolemize"),
@@ -126,31 +133,22 @@
   | NONE => (name, value)
 
 structure TheoryData = TheoryDataFun(
-  type T = {params: raw_param list, registered_auto: bool}
-  val empty = {params = rev default_default_params, registered_auto = false}
+  type T = {params: raw_param list}
+  val empty = {params = rev default_default_params}
   val copy = I
   val extend = I
-  fun merge _ ({params = ps1, registered_auto = a1},
-               {params = ps2, registered_auto = a2}) =
-    {params = AList.merge (op =) (op =) (ps1, ps2),
-     registered_auto = a1 orelse a2})
+  fun merge _ ({params = ps1}, {params = ps2}) =
+    {params = AList.merge (op =) (op =) (ps1, ps2)})
 
 (* raw_param -> theory -> theory *)
 fun set_default_raw_param param thy =
-  let val {params, registered_auto} = TheoryData.get thy in
+  let val {params} = TheoryData.get thy in
     TheoryData.put
-      {params = AList.update (op =) (unnegate_raw_param param) params,
-       registered_auto = registered_auto} thy
+        {params = AList.update (op =) (unnegate_raw_param param) params} thy
   end
 (* theory -> raw_param list *)
 val default_raw_params = #params o TheoryData.get
 
-(* theory -> theory *)
-fun set_registered_auto thy =
-  TheoryData.put {params = default_raw_params thy, registered_auto = true} thy
-(* theory -> bool *)
-val is_registered_auto = #registered_auto o TheoryData.get
-
 (* string -> bool *)
 fun is_punctuation s = (s = "," orelse s = "-" orelse s = "\<midarrow>")
 
@@ -305,7 +303,7 @@
     val overlord = lookup_bool "overlord"
     val user_axioms = lookup_bool_option "user_axioms"
     val assms = lookup_bool "assms"
-    val coalesce_type_vars = lookup_bool "coalesce_type_vars"
+    val merge_type_vars = lookup_bool "merge_type_vars"
     val destroy_constrs = lookup_bool "destroy_constrs"
     val specialize = lookup_bool "specialize"
     val skolemize = lookup_bool "skolemize"
@@ -313,8 +311,7 @@
     val uncurry = lookup_bool "uncurry"
     val fast_descrs = lookup_bool "fast_descrs"
     val peephole_optim = lookup_bool "peephole_optim"
-    val timeout = if auto then lookup_time "auto_timeout"
-                  else lookup_time "timeout"
+    val timeout = if auto then NONE else lookup_time "timeout"
     val tac_timeout = lookup_time "tac_timeout"
     val sym_break = Int.max (0, lookup_int "sym_break")
     val sharing_depth = Int.max (1, lookup_int "sharing_depth")
@@ -326,8 +323,8 @@
     val show_consts = show_all orelse lookup_bool "show_consts"
     val formats = lookup_ints_assigns read_term_polymorphic "format" 0
     val evals = lookup_term_list "eval"
-    val max_potential = if auto then 0
-                        else Int.max (0, lookup_int "max_potential")
+    val max_potential =
+      if auto then 0 else Int.max (0, lookup_int "max_potential")
     val max_genuine = Int.max (0, lookup_int "max_genuine")
     val check_potential = lookup_bool "check_potential"
     val check_genuine = lookup_bool "check_genuine"
@@ -341,7 +338,7 @@
      monos = monos, wfs = wfs, sat_solver = sat_solver, blocking = blocking,
      falsify = falsify, debug = debug, verbose = verbose, overlord = overlord,
      user_axioms = user_axioms, assms = assms,
-     coalesce_type_vars = coalesce_type_vars, destroy_constrs = destroy_constrs,
+     merge_type_vars = merge_type_vars, destroy_constrs = destroy_constrs,
      specialize = specialize, skolemize = skolemize,
      star_linear_preds = star_linear_preds, uncurry = uncurry,
      fast_descrs = fast_descrs, peephole_optim = peephole_optim,
@@ -412,7 +409,7 @@
        | Refute.REFUTE (loc, details) =>
          error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^ ".")
 
-(* raw_param list -> bool -> int -> Proof.state -> Proof.state *)
+(* raw_param list -> bool -> int -> Proof.state -> bool * Proof.state *)
 fun pick_nits override_params auto subgoal state =
   let
     val thy = Proof.theory_of state
@@ -421,70 +418,49 @@
     val _ = List.app check_raw_param override_params
     val params as {blocking, debug, ...} =
       extract_params ctxt auto (default_raw_params thy) override_params
-    (* unit -> Proof.state *)
+    (* unit -> bool * Proof.state *)
     fun go () =
-      (if auto then perhaps o try
-       else if debug then fn f => fn x => f x
-       else handle_exceptions ctxt)
-      (fn state => pick_nits_in_subgoal state params auto subgoal |> snd)
-      state
+      (false, state)
+      |> (if auto then perhaps o try
+          else if debug then fn f => fn x => f x
+          else handle_exceptions ctxt)
+         (fn (_, state) => pick_nits_in_subgoal state params auto subgoal
+                           |>> equal "genuine")
   in
-    if auto orelse blocking then
-      go ()
-    else
-      (SimpleThread.fork true (fn () => (go (); ()));
-       state)
+    if auto orelse blocking then go ()
+    else (SimpleThread.fork true (fn () => (go (); ())); (false, state))
   end
 
 (* (TableFun().key * string list) list option * int option
    -> Toplevel.transition -> Toplevel.transition *)
 fun nitpick_trans (opt_params, opt_subgoal) =
   Toplevel.keep (K ()
-      o pick_nits (these opt_params) false (the_default 1 opt_subgoal)
+      o snd o pick_nits (these opt_params) false (the_default 1 opt_subgoal)
       o Toplevel.proof_of)
 
 (* raw_param -> string *)
 fun string_for_raw_param (name, value) =
   name ^ " = " ^ stringify_raw_param_value value
 
-(* bool -> Proof.state -> Proof.state *)
-fun pick_nits_auto interactive state =
-  let val thy = Proof.theory_of state in
-    ((interactive andalso not (!Toplevel.quiet)
-      andalso the (general_lookup_bool false (default_raw_params thy)
-                  (SOME false) "auto"))
-     ? pick_nits [] true 0) state
-  end
-
-(* theory -> theory *)
-fun register_auto thy =
-  (not (is_registered_auto thy)
-   ? (set_registered_auto
-      #> Context.theory_map (Specification.add_theorem_hook pick_nits_auto)))
-  thy
-
 (* (TableFun().key * string) list option -> Toplevel.transition
    -> Toplevel.transition *)
 fun nitpick_params_trans opt_params =
   Toplevel.theory
-      (fn thy =>
-          let val thy = fold set_default_raw_param (these opt_params) thy in
-            writeln ("Default parameters for Nitpick:\n" ^
-                     (case rev (default_raw_params thy) of
-                        [] => "none"
-                      | params =>
-                        (map check_raw_param params;
-                         params |> map string_for_raw_param |> sort_strings
-                                |> cat_lines)));
-            register_auto thy
-          end)
+      (fold set_default_raw_param (these opt_params)
+       #> tap (fn thy => 
+                  writeln ("Default parameters for Nitpick:\n" ^
+                           (case rev (default_raw_params thy) of
+                              [] => "none"
+                            | params =>
+                              (map check_raw_param params;
+                               params |> map string_for_raw_param
+                                      |> sort_strings |> cat_lines)))))
 
 (* OuterParse.token list
    -> (Toplevel.transition -> Toplevel.transition) * OuterParse.token list *)
-fun scan_nitpick_command tokens =
-  (scan_params -- Scan.option OuterParse.nat) tokens |>> nitpick_trans
-fun scan_nitpick_params_command tokens =
-  scan_params tokens |>> nitpick_params_trans
+val scan_nitpick_command =
+  (scan_params -- Scan.option OuterParse.nat) #>> nitpick_trans
+val scan_nitpick_params_command = scan_params #>> nitpick_params_trans
 
 val _ = OuterSyntax.improper_command "nitpick"
             "try to find a counterexample for a given subgoal using Kodkod"
@@ -493,4 +469,10 @@
             "set and display the default parameters for Nitpick"
             OuterKeyword.thy_decl scan_nitpick_params_command
 
+(* Proof.state -> bool * Proof.state *)
+fun auto_nitpick state =
+  if not (!auto) then (false, state) else pick_nits [] true 1 state
+
+val setup = Auto_Counterexample.register_tool ("nitpick", auto_nitpick)
+
 end;
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -716,8 +716,8 @@
 (* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec list
    -> dtype_spec -> nfa_entry option *)
 fun nfa_entry_for_datatype _ _ _ _ ({co = true, ...} : dtype_spec) = NONE
-  | nfa_entry_for_datatype ext_ctxt kk rel_table dtypes
-                           ({typ, constrs, ...} : dtype_spec) =
+  | nfa_entry_for_datatype _ _ _ _ {shallow = true, ...} = NONE
+  | nfa_entry_for_datatype ext_ctxt kk rel_table dtypes {typ, constrs, ...} =
     SOME (typ, maps (nfa_transitions_for_constr ext_ctxt kk rel_table dtypes
                      o #const) constrs)
 
@@ -884,12 +884,13 @@
 
 (* extended_context -> int Typtab.table -> kodkod_constrs -> nut NameTable.table
    -> dtype_spec -> Kodkod.formula list *)
-fun other_axioms_for_datatype ext_ctxt ofs kk rel_table (dtype as {typ, ...}) =
-  let val j0 = offset_of_type ofs typ in
-    sel_axioms_for_datatype ext_ctxt j0 kk rel_table dtype @
-    uniqueness_axioms_for_datatype ext_ctxt kk rel_table dtype @
-    partition_axioms_for_datatype j0 kk rel_table dtype
-  end
+fun other_axioms_for_datatype _ _ _ _ {shallow = true, ...} = []
+  | other_axioms_for_datatype ext_ctxt ofs kk rel_table (dtype as {typ, ...}) =
+    let val j0 = offset_of_type ofs typ in
+      sel_axioms_for_datatype ext_ctxt j0 kk rel_table dtype @
+      uniqueness_axioms_for_datatype ext_ctxt kk rel_table dtype @
+      partition_axioms_for_datatype j0 kk rel_table dtype
+    end
 
 (* extended_context -> int Typtab.table -> kodkod_constrs -> nut NameTable.table
    -> dtype_spec list -> Kodkod.formula list *)
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -133,12 +133,14 @@
       | aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
   in apsnd (pairself rev) o aux end
 
-(* typ -> term -> term list * term *)
-fun break_in_two (Type ("*", [T1, T2]))
-                 (Const (@{const_name Pair}, _) $ t1 $ t2) =
-    break_in_two T2 t2 |>> cons t1
-  | break_in_two _ (Const (@{const_name Pair}, _) $ t1 $ t2) = ([t1], t2)
-  | break_in_two _ t = raise TERM ("Nitpick_Model.break_in_two", [t])
+(* typ -> typ -> typ -> term -> term * term *)
+fun break_in_two T T1 T2 t =
+  let
+    val ps = HOLogic.flat_tupleT_paths T
+    val cut = length (HOLogic.strip_tupleT T1)
+    val (ps1, ps2) = pairself HOLogic.flat_tupleT_paths (T1, T2)
+    val (ts1, ts2) = t |> HOLogic.strip_ptuple ps |> chop cut
+  in (HOLogic.mk_ptuple ps1 T1 ts1, HOLogic.mk_ptuple ps2 T2 ts2) end
 (* typ -> term -> term -> term *)
 fun pair_up (Type ("*", [T1', T2']))
             (t1 as Const (@{const_name Pair},
@@ -153,12 +155,12 @@
 (* typ -> typ -> typ -> term -> term *)
 fun typecast_fun (Type ("fun", [T1', T2'])) T1 T2 t =
     let
-      (* typ -> typ -> typ -> term -> term *)
-      fun do_curry T1a T1b T2 t =
+      (* typ -> typ -> typ -> typ -> term -> term *)
+      fun do_curry T1 T1a T1b T2 t =
         let
           val (maybe_opt, ps) = dest_plain_fun t
           val ps =
-            ps |>> map (break_in_two T1a #>> mk_flat_tuple T1a)
+            ps |>> map (break_in_two T1 T1a T1b)
                |> uncurry (map2 (fn (t1a, t1b) => fn t2 => (t1a, (t1b, t2))))
                |> AList.coalesce (op =)
                |> map (apsnd (make_plain_fun maybe_opt T1b T2))
@@ -185,7 +187,7 @@
         case factor_out_types T1' T1 of
           ((_, NONE), (_, NONE)) => t |> do_arrow T1' T2' T1 T2
         | ((_, NONE), (T1a, SOME T1b)) =>
-          t |> do_curry T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2)
+          t |> do_curry T1 T1a T1b T2 |> do_arrow T1' T2' T1a (T1b --> T2)
         | ((T1a', SOME T1b'), (_, NONE)) =>
           t |> do_arrow T1a' (T1b' --> T2') T1 T2 |> do_uncurry T1' T2'
         | _ => raise TYPE ("Nitpick_Model.typecast_fun.do_fun", [T1, T1'], [])
@@ -329,7 +331,8 @@
           HOLogic.mk_number nat_T j
         else case datatype_spec datatypes T of
           NONE => atom for_auto T j
-        | SOME {constrs, co, ...} =>
+        | SOME {shallow = true, ...} => atom for_auto T j
+        | SOME {co, constrs, ...} =>
           let
             (* styp -> int list *)
             fun tuples_for_const (s, T) =
@@ -600,11 +603,12 @@
     (* typ -> dtype_spec list *)
     fun integer_datatype T =
       [{typ = T, card = card_of_type card_assigns T, co = false,
-        precise = false, constrs = []}]
+        precise = false, shallow = false, constrs = []}]
       handle TYPE ("Nitpick_HOL.card_of_type", _, _) => []
     val (codatatypes, datatypes) =
-      List.partition #co datatypes
-      ||> append (integer_datatype nat_T @ integer_datatype int_T)
+      datatypes |> filter_out #shallow
+                |> List.partition #co
+                ||> append (integer_datatype nat_T @ integer_datatype int_T)
     val block_of_datatypes =
       if show_datatypes andalso not (null datatypes) then
         [Pretty.big_list ("Datatype" ^ plural_s_for_list datatypes ^ ":")
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -749,8 +749,9 @@
            (~1 upto num_sels_for_constr_type T - 1)
 (* scope -> dtype_spec -> nut list * rep NameTable.table
    -> nut list * rep NameTable.table *)
-fun choose_rep_for_sels_of_datatype scope ({constrs, ...} : dtype_spec) =
-  fold_rev (choose_rep_for_sels_for_constr scope o #const) constrs
+fun choose_rep_for_sels_of_datatype _ ({shallow = true, ...} : dtype_spec) = I
+  | choose_rep_for_sels_of_datatype scope {constrs, ...} =
+    fold_rev (choose_rep_for_sels_for_constr scope o #const) constrs
 (* scope -> rep NameTable.table -> nut list * rep NameTable.table *)
 fun choose_reps_for_all_sels (scope as {datatypes, ...}) =
   fold (choose_rep_for_sels_of_datatype scope) datatypes o pair []
--- a/src/HOL/Tools/Nitpick/nitpick_scope.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -22,6 +22,7 @@
     card: int,
     co: bool,
     precise: bool,
+    shallow: bool,
     constrs: constr_spec list}
 
   type scope = {
@@ -44,7 +45,7 @@
   val all_scopes :
     extended_context -> int -> (typ option * int list) list
     -> (styp option * int list) list -> (styp option * int list) list
-    -> int list -> typ list -> typ list -> scope list
+    -> int list -> typ list -> typ list -> typ list -> scope list
 end;
 
 structure Nitpick_Scope : NITPICK_SCOPE =
@@ -66,6 +67,7 @@
   card: int,
   co: bool,
   precise: bool,
+  shallow: bool,
   constrs: constr_spec list}
 
 type scope = {
@@ -126,7 +128,7 @@
       card_assigns |> filter_out (equal @{typ bisim_iterator} o fst)
                    |> List.partition (is_fp_iterator_type o fst)
     val (unimportant_card_asgns, important_card_asgns) =
-      card_asgns |> List.partition ((is_datatype thy orf is_integer_type) o fst)
+      card_asgns |> List.partition ((is_integer_type orf is_datatype thy) o fst)
     val cards =
       map (fn (T, k) => quote (string_for_type ctxt T) ^ " = " ^
                         string_of_int k)
@@ -431,10 +433,11 @@
      explicit_max = max, total = total} :: constrs
   end
 
-(* extended_context -> scope_desc -> typ * int -> dtype_spec *)
-fun datatype_spec_from_scope_descriptor (ext_ctxt as {thy, ...})
+(* extended_context -> typ list -> scope_desc -> typ * int -> dtype_spec *)
+fun datatype_spec_from_scope_descriptor (ext_ctxt as {thy, ...}) shallow_dataTs
                                         (desc as (card_asgns, _)) (T, card) =
   let
+    val shallow = T mem shallow_dataTs
     val co = is_codatatype thy T
     val xs = boxed_datatype_constrs ext_ctxt T
     val self_recs = map (is_self_recursive_constr_type o snd) xs
@@ -448,14 +451,18 @@
     val constrs =
       fold_rev (add_constr_spec desc co card sum_dom_cards num_self_recs
                                 num_non_self_recs) (self_recs ~~ xs) []
-  in {typ = T, card = card, co = co, precise = precise, constrs = constrs} end
+  in
+    {typ = T, card = card, co = co, precise = precise, shallow = shallow,
+     constrs = constrs}
+  end
 
-(* extended_context -> int -> scope_desc -> scope *)
-fun scope_from_descriptor (ext_ctxt as {thy, ...}) sym_break
+(* extended_context -> int -> typ list -> scope_desc -> scope *)
+fun scope_from_descriptor (ext_ctxt as {thy, ...}) sym_break shallow_dataTs
                           (desc as (card_asgns, _)) =
   let
-    val datatypes = map (datatype_spec_from_scope_descriptor ext_ctxt desc)
-                        (filter (is_datatype thy o fst) card_asgns)
+    val datatypes =
+      map (datatype_spec_from_scope_descriptor ext_ctxt shallow_dataTs desc)
+          (filter (is_datatype thy o fst) card_asgns)
     val bisim_depth = card_of_type card_asgns @{typ bisim_iterator} - 1
   in
     {ext_ctxt = ext_ctxt, card_assigns = card_asgns, datatypes = datatypes,
@@ -480,9 +487,9 @@
 
 (* extended_context -> int -> (typ option * int list) list
    -> (styp option * int list) list -> (styp option * int list) list -> int list
-   -> typ list -> typ list -> scope list *)
+   -> typ list -> typ list -> typ list -> scope list *)
 fun all_scopes (ext_ctxt as {thy, ctxt, ...}) sym_break cards_asgns maxes_asgns
-               iters_asgns bisim_depths mono_Ts nonmono_Ts =
+               iters_asgns bisim_depths mono_Ts nonmono_Ts shallow_dataTs =
   let
     val cards_asgns = fix_cards_assigns_wrt_boxing thy mono_Ts cards_asgns
     val blocks = blocks_for_types ctxt cards_asgns maxes_asgns iters_asgns
@@ -492,7 +499,7 @@
                 |> map_filter (scope_descriptor_from_combination thy blocks)
   in
     descs |> length descs <= distinct_threshold ? distinct (op =)
-          |> map (scope_from_descriptor ext_ctxt sym_break)
+          |> map (scope_from_descriptor ext_ctxt sym_break shallow_dataTs)
   end
 
 end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Auto_Counterexample.thy	Thu Oct 29 12:09:32 2009 +0100
@@ -0,0 +1,15 @@
+(*  Title:   Tools/Auto_Counterexample.thy
+    Author:  Jasmin Blanchette, TU Muenchen
+
+Counterexample Search Unit (do not abbreviate!).
+*)
+
+header {* Counterexample Search Unit *}
+
+theory Auto_Counterexample
+imports Pure
+uses
+  "~~/src/Tools/auto_counterexample.ML"
+begin
+
+end
--- a/src/Tools/Code_Generator.thy	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/Tools/Code_Generator.thy	Thu Oct 29 12:09:32 2009 +0100
@@ -5,7 +5,7 @@
 header {* Loading the code generator modules *}
 
 theory Code_Generator
-imports Pure
+imports Auto_Counterexample
 uses
   "~~/src/Tools/value.ML"
   "~~/src/Tools/quickcheck.ML"
@@ -25,4 +25,4 @@
   #> Nbe.setup
 *}
 
-end
\ No newline at end of file
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/auto_counterexample.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -0,0 +1,59 @@
+(*  Title:      Tools/auto_counterexample.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Counterexample Search Unit (do not abbreviate!).
+*)
+
+signature AUTO_COUNTEREXAMPLE =
+sig
+  val time_limit: int Unsynchronized.ref
+
+  val register_tool :
+    string * (Proof.state -> bool * Proof.state) -> theory -> theory
+end;
+
+structure Auto_Counterexample : AUTO_COUNTEREXAMPLE =
+struct
+
+(* preferences *)
+
+val time_limit = Unsynchronized.ref 2500
+
+val _ =
+  ProofGeneralPgip.add_preference Preferences.category_tracing
+    (Preferences.nat_pref time_limit
+      "auto-counterexample-time-limit"
+      "Time limit for automatic counterexample search (in milliseconds).")
+
+
+(* configuration *)
+
+structure Data = TheoryDataFun(
+  type T = (string * (Proof.state -> bool * Proof.state)) list
+  val empty = []
+  val copy = I
+  val extend = I
+  fun merge _ (tools1, tools2) = AList.merge (op =) (K true) (tools1, tools2)
+)
+
+val register_tool = Data.map o AList.update (op =)
+
+
+(* automatic testing *)
+
+val _ = Context.>> (Specification.add_theorem_hook (fn interact => fn state =>
+  case !time_limit of
+    0 => state
+  | ms =>
+    TimeLimit.timeLimit (Time.fromMilliseconds ms)
+        (fn () =>
+            if interact andalso not (!Toplevel.quiet) then
+              fold_rev (fn (_, tool) => fn (true, state) => (true, state)
+                                         | (false, state) => tool state)
+                   (Data.get (Proof.theory_of state)) (false, state) |> snd
+            else
+              state) ()
+    handle TimeLimit.TimeOut =>
+           (warning "Automatic counterexample search timed out."; state)))
+
+end;
--- a/src/Tools/quickcheck.ML	Thu Oct 29 11:26:47 2009 +0100
+++ b/src/Tools/quickcheck.ML	Thu Oct 29 12:09:32 2009 +0100
@@ -7,10 +7,10 @@
 signature QUICKCHECK =
 sig
   val auto: bool Unsynchronized.ref
-  val auto_time_limit: int Unsynchronized.ref
   val test_term: Proof.context -> bool -> string option -> int -> int -> term ->
     (string * term) list option
   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
 end;
 
@@ -20,20 +20,13 @@
 (* preferences *)
 
 val auto = Unsynchronized.ref false;
-val auto_time_limit = Unsynchronized.ref 2500;
 
 val _ =
   ProofGeneralPgip.add_preference Preferences.category_tracing
   (setmp_CRITICAL auto true (fn () =>
     Preferences.bool_pref auto
       "auto-quickcheck"
-      "Whether to enable quickcheck automatically.") ());
-
-val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
-    (Preferences.nat_pref auto_time_limit
-      "auto-quickcheck-time-limit"
-      "Time limit for automatic quickcheck (in milliseconds).");
+      "Whether to run Quickcheck automatically.") ());
 
 
 (* quickcheck configuration -- default parameters, test generators *)
@@ -159,28 +152,26 @@
 
 (* automatic testing *)
 
-val _ = Context.>> (Specification.add_theorem_hook (fn int => fn state =>
-  let
-    val ctxt = Proof.context_of state;
-    val assms = map term_of (Assumption.all_assms_of ctxt);
-    val Test_Params { size, iterations, default_type } =
-      (snd o Data.get o Proof.theory_of) state;
-    fun test () =
-      let
-        val res = TimeLimit.timeLimit (Time.fromMilliseconds (!auto_time_limit))
-          (try (test_goal true NONE size iterations default_type [] 1 assms)) state;
-      in
-        case res of
-          NONE => state
-        | SOME NONE => state
-        | SOME cex => Proof.goal_message (fn () => Pretty.chunks [Pretty.str "",
-            Pretty.mark Markup.hilite (pretty_counterex ctxt cex)]) state
-      end handle TimeLimit.TimeOut => (warning "Auto quickcheck: timeout."; state);
-  in
-    if int andalso !auto andalso not (!Toplevel.quiet)
-    then test ()
-    else state
-  end));
+fun auto_quickcheck state =
+  if not (!auto) then
+    (false, state)
+  else
+    let
+      val ctxt = Proof.context_of state;
+      val assms = map term_of (Assumption.all_assms_of ctxt);
+      val Test_Params { size, iterations, default_type } =
+        (snd o Data.get o Proof.theory_of) state;
+      val res =
+        try (test_goal true NONE size iterations default_type [] 1 assms) state;
+    in
+      case res of
+        NONE => (false, state)
+      | SOME NONE => (false, state)
+      | SOME cex => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "",
+          Pretty.mark Markup.hilite (pretty_counterex ctxt cex)])) state)
+    end
+
+val setup = Auto_Counterexample.register_tool ("quickcheck", auto_quickcheck)
 
 
 (* Isar commands *)
@@ -251,4 +242,3 @@
 
 
 val auto_quickcheck = Quickcheck.auto;
-val auto_quickcheck_time_limit = Quickcheck.auto_time_limit;