src/HOL/Tools/Nitpick/nitpick.ML
author blanchet
Wed, 18 Apr 2012 22:40:25 +0200
changeset 48431 e30323bfc93c
parent 48430 366838a5e235
child 48433 a72239723ae8
permissions -rw-r--r--
added SZS status wrappers in TPTP mode
     1 (*  Title:      HOL/Tools/Nitpick/nitpick.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2008, 2009, 2010
     4 
     5 Finite model generation for HOL formulas using Kodkod.
     6 *)
     7 
     8 signature NITPICK =
     9 sig
    10   type styp = Nitpick_Util.styp
    11   type term_postprocessor = Nitpick_Model.term_postprocessor
    12 
    13   datatype mode = Auto_Try | Try | Normal | TPTP
    14 
    15   type params =
    16     {cards_assigns: (typ option * int list) list,
    17      maxes_assigns: (styp option * int list) list,
    18      iters_assigns: (styp option * int list) list,
    19      bitss: int list,
    20      bisim_depths: int list,
    21      boxes: (typ option * bool option) list,
    22      finitizes: (typ option * bool option) list,
    23      monos: (typ option * bool option) list,
    24      stds: (typ option * bool) list,
    25      wfs: (styp option * bool option) list,
    26      sat_solver: string,
    27      blocking: bool,
    28      falsify: bool,
    29      debug: bool,
    30      verbose: bool,
    31      overlord: bool,
    32      user_axioms: bool option,
    33      assms: bool,
    34      whacks: term list,
    35      merge_type_vars: bool,
    36      binary_ints: bool option,
    37      destroy_constrs: bool,
    38      specialize: bool,
    39      star_linear_preds: bool,
    40      total_consts: bool option,
    41      needs: term list option,
    42      peephole_optim: bool,
    43      datatype_sym_break: int,
    44      kodkod_sym_break: int,
    45      timeout: Time.time option,
    46      tac_timeout: Time.time option,
    47      max_threads: int,
    48      show_datatypes: bool,
    49      show_skolems: bool,
    50      show_consts: bool,
    51      evals: term list,
    52      formats: (term option * int list) list,
    53      atomss: (typ option * string list) list,
    54      max_potential: int,
    55      max_genuine: int,
    56      check_potential: bool,
    57      check_genuine: bool,
    58      batch_size: int,
    59      expect: string}
    60 
    61   val genuineN : string
    62   val quasi_genuineN : string
    63   val potentialN : string
    64   val noneN : string
    65   val unknownN : string
    66   val register_frac_type : string -> (string * string) list -> theory -> theory
    67   val unregister_frac_type : string -> theory -> theory
    68   val register_codatatype : typ -> string -> styp list -> theory -> theory
    69   val unregister_codatatype : typ -> theory -> theory
    70   val register_term_postprocessor :
    71     typ -> term_postprocessor -> theory -> theory
    72   val unregister_term_postprocessor : typ -> theory -> theory
    73   val pick_nits_in_term :
    74     Proof.state -> params -> mode -> int -> int -> int -> (term * term) list
    75     -> term list -> term -> string * Proof.state
    76   val pick_nits_in_subgoal :
    77     Proof.state -> params -> mode -> int -> int -> string * Proof.state
    78 end;
    79 
    80 structure Nitpick : NITPICK =
    81 struct
    82 
    83 open Nitpick_Util
    84 open Nitpick_HOL
    85 open Nitpick_Preproc
    86 open Nitpick_Mono
    87 open Nitpick_Scope
    88 open Nitpick_Peephole
    89 open Nitpick_Rep
    90 open Nitpick_Nut
    91 open Nitpick_Kodkod
    92 open Nitpick_Model
    93 
    94 structure KK = Kodkod
    95 
    96 datatype mode = Auto_Try | Try | Normal | TPTP
    97 
    98 fun is_mode_nt mode = (mode = Normal orelse mode = TPTP)
    99 
   100 type params =
   101   {cards_assigns: (typ option * int list) list,
   102    maxes_assigns: (styp option * int list) list,
   103    iters_assigns: (styp option * int list) list,
   104    bitss: int list,
   105    bisim_depths: int list,
   106    boxes: (typ option * bool option) list,
   107    finitizes: (typ option * bool option) list,
   108    monos: (typ option * bool option) list,
   109    stds: (typ option * bool) list,
   110    wfs: (styp option * bool option) list,
   111    sat_solver: string,
   112    blocking: bool,
   113    falsify: bool,
   114    debug: bool,
   115    verbose: bool,
   116    overlord: bool,
   117    user_axioms: bool option,
   118    assms: bool,
   119    whacks: term list,
   120    merge_type_vars: bool,
   121    binary_ints: bool option,
   122    destroy_constrs: bool,
   123    specialize: bool,
   124    star_linear_preds: bool,
   125    total_consts: bool option,
   126    needs: term list option,
   127    peephole_optim: bool,
   128    datatype_sym_break: int,
   129    kodkod_sym_break: int,
   130    timeout: Time.time option,
   131    tac_timeout: Time.time option,
   132    max_threads: int,
   133    show_datatypes: bool,
   134    show_skolems: bool,
   135    show_consts: bool,
   136    evals: term list,
   137    formats: (term option * int list) list,
   138    atomss: (typ option * string list) list,
   139    max_potential: int,
   140    max_genuine: int,
   141    check_potential: bool,
   142    check_genuine: bool,
   143    batch_size: int,
   144    expect: string}
   145 
   146 val genuineN = "genuine"
   147 val quasi_genuineN = "quasi_genuine"
   148 val potentialN = "potential"
   149 val noneN = "none"
   150 val unknownN = "unknown"
   151 
   152 (* TODO: eliminate these historical aliases *)
   153 val register_frac_type = Nitpick_HOL.register_frac_type_global
   154 val unregister_frac_type = Nitpick_HOL.unregister_frac_type_global
   155 val register_codatatype = Nitpick_HOL.register_codatatype_global
   156 val unregister_codatatype = Nitpick_HOL.unregister_codatatype_global
   157 val register_term_postprocessor =
   158   Nitpick_Model.register_term_postprocessor_global
   159 val unregister_term_postprocessor =
   160   Nitpick_Model.unregister_term_postprocessor_global
   161 
   162 type problem_extension =
   163   {free_names: nut list,
   164    sel_names: nut list,
   165    nonsel_names: nut list,
   166    rel_table: nut NameTable.table,
   167    unsound: bool,
   168    scope: scope}
   169 
   170 type rich_problem = KK.problem * problem_extension
   171 
   172 fun pretties_for_formulas _ _ [] = []
   173   | pretties_for_formulas ctxt s ts =
   174     [Pretty.str (s ^ plural_s_for_list ts ^ ":"),
   175      Pretty.indent indent_size (Pretty.chunks
   176          (map2 (fn j => fn t =>
   177                    Pretty.block [t |> shorten_names_in_term
   178                                    |> Syntax.pretty_term ctxt,
   179                                  Pretty.str (if j = 1 then "." else ";")])
   180                (length ts downto 1) ts))]
   181 
   182 fun install_java_message () =
   183   "Nitpick requires Java Development Kit 1.6/1.7 via ISABELLE_JDK_HOME setting."
   184 fun install_kodkodi_message () =
   185   "Nitpick requires the external Java program Kodkodi. To install it, download \
   186   \the package from \"http://www21.in.tum.de/~blanchet/#software\" and add the \
   187   \\"kodkodi-x.y.z\" directory's full path to " ^
   188   Path.print (Path.expand (Path.explode "$ISABELLE_HOME_USER/etc/components")) ^
   189   " on a line of its own."
   190 
   191 val max_unsound_delay_ms = 200
   192 val max_unsound_delay_percent = 2
   193 
   194 fun unsound_delay_for_timeout NONE = max_unsound_delay_ms
   195   | unsound_delay_for_timeout (SOME timeout) =
   196     Int.max (0, Int.min (max_unsound_delay_ms,
   197                          Time.toMilliseconds timeout
   198                          * max_unsound_delay_percent div 100))
   199 
   200 fun passed_deadline NONE = false
   201   | passed_deadline (SOME time) = Time.compare (Time.now (), time) <> LESS
   202 
   203 fun none_true assigns = forall (not_equal (SOME true) o snd) assigns
   204 
   205 fun has_lonely_bool_var (@{const Pure.conjunction}
   206                          $ (@{const Trueprop} $ Free _) $ _) = true
   207   | has_lonely_bool_var _ = false
   208 
   209 val syntactic_sorts =
   210   @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,equal}"} @
   211   @{sort numeral}
   212 fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) =
   213     subset (op =) (S, syntactic_sorts)
   214   | has_tfree_syntactic_sort _ = false
   215 val has_syntactic_sorts = exists_type (exists_subtype has_tfree_syntactic_sort)
   216 
   217 fun plazy f = Pretty.blk (0, pstrs (f ()))
   218 
   219 fun pick_them_nits_in_term deadline state (params : params) mode i n step
   220                            subst assm_ts orig_t =
   221   let
   222     val timer = Timer.startRealTimer ()
   223     val thy = Proof.theory_of state
   224     val ctxt = Proof.context_of state
   225 (* FIXME: reintroduce code before new release:
   226 
   227     val nitpick_thy = Thy_Info.get_theory "Nitpick"
   228     val _ = Theory.subthy (nitpick_thy, thy) orelse
   229             error "You must import the theory \"Nitpick\" to use Nitpick"
   230 *)
   231     val {cards_assigns, maxes_assigns, iters_assigns, bitss, bisim_depths,
   232          boxes, finitizes, monos, stds, wfs, sat_solver, falsify, debug,
   233          verbose, overlord, user_axioms, assms, whacks, merge_type_vars,
   234          binary_ints, destroy_constrs, specialize, star_linear_preds,
   235          total_consts, needs, peephole_optim, datatype_sym_break,
   236          kodkod_sym_break, tac_timeout, max_threads, show_datatypes,
   237          show_skolems, show_consts, evals, formats, atomss, max_potential,
   238          max_genuine, check_potential, check_genuine, batch_size, ...} = params
   239     val state_ref = Unsynchronized.ref state
   240     fun pprint print =
   241       if mode = Auto_Try then
   242         Unsynchronized.change state_ref o Proof.goal_message o K
   243         o Pretty.chunks o cons (Pretty.str "") o single
   244         o Pretty.mark Isabelle_Markup.hilite
   245       else
   246         print o Pretty.string_of
   247     val pprint_a = pprint Output.urgent_message
   248     fun pprint_nt f = () |> is_mode_nt mode ? pprint Output.urgent_message o f
   249     fun pprint_v f = () |> verbose ? pprint Output.urgent_message o f
   250     fun pprint_d f = () |> debug ? pprint tracing o f
   251     val print = pprint Output.urgent_message o curry Pretty.blk 0 o pstrs
   252     fun print_t f = () |> mode = TPTP ? print o f
   253 (*
   254     val print_g = pprint tracing o Pretty.str
   255 *)
   256     val print_nt = pprint_nt o K o plazy
   257     val print_v = pprint_v o K o plazy
   258 
   259     fun check_deadline () =
   260       if passed_deadline deadline then raise TimeLimit.TimeOut else ()
   261 
   262     val assm_ts = if assms orelse mode <> Normal then assm_ts else []
   263     val _ =
   264       if step = 0 then
   265         print_nt (fn () => "Nitpicking formula...")
   266       else
   267         pprint_nt (fn () => Pretty.chunks (
   268             pretties_for_formulas ctxt ("Nitpicking " ^
   269                 (if i <> 1 orelse n <> 1 then
   270                    "subgoal " ^ string_of_int i ^ " of " ^ string_of_int n
   271                  else
   272                    "goal")) [Logic.list_implies (assm_ts, orig_t)]))
   273     val _ = print_v (enclose "Timestamp: " "." o Date.fmt "%H:%M:%S"
   274                      o Date.fromTimeLocal o Time.now)
   275     val neg_t = if falsify then Logic.mk_implies (orig_t, @{prop False})
   276                 else orig_t
   277     val conj_ts = neg_t :: assm_ts @ evals @ these needs
   278     val tfree_table =
   279       if merge_type_vars then merged_type_var_table_for_terms thy conj_ts
   280       else []
   281     val merge_tfrees = merge_type_vars_in_term thy merge_type_vars tfree_table
   282     val neg_t = neg_t |> merge_tfrees
   283     val assm_ts = assm_ts |> map merge_tfrees
   284     val evals = evals |> map merge_tfrees
   285     val needs = needs |> Option.map (map merge_tfrees)
   286     val conj_ts = neg_t :: assm_ts @ evals @ these needs
   287     val original_max_potential = max_potential
   288     val original_max_genuine = max_genuine
   289     val max_bisim_depth = fold Integer.max bisim_depths ~1
   290     val case_names = case_const_names ctxt stds
   291     val defs = all_defs_of thy subst
   292     val nondefs = all_nondefs_of ctxt subst
   293     val def_tables = const_def_tables ctxt subst defs
   294     val nondef_table = const_nondef_table nondefs
   295     val simp_table = Unsynchronized.ref (const_simp_table ctxt subst)
   296     val psimp_table = const_psimp_table ctxt subst
   297     val choice_spec_table = const_choice_spec_table ctxt subst
   298     val nondefs =
   299       nondefs |> filter_out (is_choice_spec_axiom thy choice_spec_table)
   300     val intro_table = inductive_intro_table ctxt subst def_tables
   301     val ground_thm_table = ground_theorem_table thy
   302     val ersatz_table = ersatz_table ctxt
   303     val hol_ctxt as {wf_cache, ...} =
   304       {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
   305        stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug,
   306        whacks = whacks, binary_ints = binary_ints,
   307        destroy_constrs = destroy_constrs, specialize = specialize,
   308        star_linear_preds = star_linear_preds, total_consts = total_consts,
   309        needs = needs, tac_timeout = tac_timeout, evals = evals,
   310        case_names = case_names, def_tables = def_tables,
   311        nondef_table = nondef_table, nondefs = nondefs, simp_table = simp_table,
   312        psimp_table = psimp_table, choice_spec_table = choice_spec_table,
   313        intro_table = intro_table, ground_thm_table = ground_thm_table,
   314        ersatz_table = ersatz_table, skolems = Unsynchronized.ref [],
   315        special_funs = Unsynchronized.ref [],
   316        unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref [],
   317        constr_cache = Unsynchronized.ref []}
   318     val pseudo_frees = []
   319     val real_frees = fold Term.add_frees conj_ts []
   320     val _ = null (fold Term.add_tvars conj_ts []) orelse
   321             error "Nitpick cannot handle goals with schematic type variables."
   322     val (nondef_ts, def_ts, need_ts, got_all_mono_user_axioms,
   323          no_poly_user_axioms, binarize) =
   324       preprocess_formulas hol_ctxt assm_ts neg_t
   325     val got_all_user_axioms =
   326       got_all_mono_user_axioms andalso no_poly_user_axioms
   327 
   328     fun print_wf (x, (gfp, wf)) =
   329       pprint_nt (fn () => Pretty.blk (0,
   330           pstrs ("The " ^ (if gfp then "co" else "") ^ "inductive predicate \"")
   331           @ Syntax.pretty_term ctxt (Const x) ::
   332           pstrs (if wf then
   333                    "\" was proved well-founded. Nitpick can compute it \
   334                    \efficiently."
   335                  else
   336                    "\" could not be proved well-founded. Nitpick might need to \
   337                    \unroll it.")))
   338     val _ = if verbose then List.app print_wf (!wf_cache) else ()
   339     val das_wort_formula = if falsify then "Negated conjecture" else "Formula"
   340     val _ =
   341       pprint_d (fn () => Pretty.chunks
   342           (pretties_for_formulas ctxt das_wort_formula [hd nondef_ts] @
   343            pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @
   344            pretties_for_formulas ctxt "Relevant nondefinitional axiom"
   345                                  (tl nondef_ts)))
   346     val _ = List.app (ignore o Term.type_of) (nondef_ts @ def_ts)
   347             handle TYPE (_, Ts, ts) =>
   348                    raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts)
   349 
   350     val nondef_us = nondef_ts |> map (nut_from_term hol_ctxt Eq)
   351     val def_us = def_ts |> map (nut_from_term hol_ctxt DefEq)
   352     val need_us = need_ts |> map (nut_from_term hol_ctxt Eq)
   353     val (free_names, const_names) =
   354       fold add_free_and_const_names (nondef_us @ def_us @ need_us) ([], [])
   355     val (sel_names, nonsel_names) =
   356       List.partition (is_sel o nickname_of) const_names
   357     val sound_finitizes = none_true finitizes
   358     val standard = forall snd stds
   359 (*
   360     val _ =
   361       List.app (print_g o string_for_nut ctxt) (nondef_us @ def_us @ need_us)
   362 *)
   363 
   364     val unique_scope = forall (curry (op =) 1 o length o snd) cards_assigns
   365     fun monotonicity_message Ts extra =
   366       let val pretties = map (pretty_maybe_quote o pretty_for_type ctxt) Ts in
   367         Pretty.blk (0,
   368           pstrs ("The type" ^ plural_s_for_list pretties ^ " ") @
   369           pretty_serial_commas "and" pretties @
   370           pstrs (" " ^
   371                  (if none_true monos then
   372                     "passed the monotonicity test"
   373                   else
   374                     (if length pretties = 1 then "is" else "are") ^
   375                     " considered monotonic") ^
   376                  ". " ^ extra))
   377       end
   378     fun is_type_fundamentally_monotonic T =
   379       (is_datatype ctxt stds T andalso not (is_quot_type ctxt T) andalso
   380        (not (is_pure_typedef ctxt T) orelse is_univ_typedef ctxt T)) orelse
   381       is_number_type ctxt T orelse is_bit_type T
   382     fun is_type_actually_monotonic T =
   383       formulas_monotonic hol_ctxt binarize T (nondef_ts, def_ts)
   384     fun is_type_kind_of_monotonic T =
   385       case triple_lookup (type_match thy) monos T of
   386         SOME (SOME false) => false
   387       | _ => is_type_actually_monotonic T
   388     fun is_type_monotonic T =
   389       unique_scope orelse
   390       case triple_lookup (type_match thy) monos T of
   391         SOME (SOME b) => b
   392       | _ => is_type_fundamentally_monotonic T orelse
   393              is_type_actually_monotonic T
   394     fun is_deep_datatype_finitizable T =
   395       triple_lookup (type_match thy) finitizes T = SOME (SOME true)
   396     fun is_shallow_datatype_finitizable (T as Type (@{type_name fun_box}, _)) =
   397         is_type_kind_of_monotonic T
   398       | is_shallow_datatype_finitizable T =
   399         case triple_lookup (type_match thy) finitizes T of
   400           SOME (SOME b) => b
   401         | _ => is_type_kind_of_monotonic T
   402     fun is_datatype_deep T =
   403       not standard orelse T = @{typ unit} orelse T = nat_T orelse
   404       is_word_type T orelse
   405       exists (curry (op =) T o domain_type o type_of) sel_names
   406     val all_Ts = ground_types_in_terms hol_ctxt binarize (nondef_ts @ def_ts)
   407                  |> sort Term_Ord.typ_ord
   408     val _ =
   409       if verbose andalso binary_ints = SOME true andalso
   410          exists (member (op =) [nat_T, int_T]) all_Ts then
   411         print_v (K "The option \"binary_ints\" will be ignored because of the \
   412                    \presence of rationals, reals, \"Suc\", \"gcd\", or \"lcm\" \
   413                    \in the problem or because of the \"non_std\" option.")
   414       else
   415         ()
   416     val _ =
   417       if mode = Normal andalso
   418          exists (fn Type (@{type_name Datatype.node}, _) => true | _ => false)
   419                 all_Ts then
   420         print_nt (K ("Warning: The problem involves directly or indirectly the \
   421                      \internal type " ^ quote @{type_name Datatype.node} ^
   422                      ". This type is very Nitpick-unfriendly, and its presence \
   423                      \usually indicates either a failure of abstraction or a \
   424                      \quirk in Nitpick."))
   425       else
   426         ()
   427     val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts
   428     val _ =
   429       if verbose andalso not unique_scope then
   430         case filter_out is_type_fundamentally_monotonic mono_Ts of
   431           [] => ()
   432         | interesting_mono_Ts =>
   433           pprint_v (K (monotonicity_message interesting_mono_Ts
   434                           "Nitpick might be able to skip some scopes."))
   435       else
   436         ()
   437     val (deep_dataTs, shallow_dataTs) =
   438       all_Ts |> filter (is_datatype ctxt stds)
   439              |> List.partition is_datatype_deep
   440     val finitizable_dataTs =
   441       (deep_dataTs |> filter_out (is_finite_type hol_ctxt)
   442                    |> filter is_deep_datatype_finitizable) @
   443       (shallow_dataTs |> filter_out (is_finite_type hol_ctxt)
   444                       |> filter is_shallow_datatype_finitizable)
   445     val _ =
   446       if verbose andalso not (null finitizable_dataTs) then
   447         pprint_v (K (monotonicity_message finitizable_dataTs
   448                          "Nitpick can use a more precise finite encoding."))
   449       else
   450         ()
   451     (* This detection code is an ugly hack. Fortunately, it is used only to
   452        provide a hint to the user. *)
   453     fun is_struct_induct_step (name, (Rule_Cases.Case {fixes, assumes, ...}, _)) =
   454       not (null fixes) andalso
   455       exists (String.isSuffix ".hyps" o fst) assumes andalso
   456       exists (exists (curry (op =) name o shortest_name o fst)
   457               o datatype_constrs hol_ctxt) deep_dataTs
   458     val likely_in_struct_induct_step =
   459       exists is_struct_induct_step (Proof_Context.cases_of ctxt)
   460     val _ = if standard andalso likely_in_struct_induct_step then
   461               pprint_nt (fn () => Pretty.blk (0,
   462                   pstrs "Hint: To check that the induction hypothesis is \
   463                         \general enough, try this command: " @
   464                   [Pretty.mark Isabelle_Markup.sendback (Pretty.blk (0,
   465                        pstrs ("nitpick [non_std, show_all]")))] @ pstrs "."))
   466             else
   467               ()
   468 (*
   469     val _ = print_g "Monotonic types:"
   470     val _ = List.app (print_g o string_for_type ctxt) mono_Ts
   471     val _ = print_g "Nonmonotonic types:"
   472     val _ = List.app (print_g o string_for_type ctxt) nonmono_Ts
   473 *)
   474 
   475     val incremental = Int.max (max_potential, max_genuine) >= 2
   476     val actual_sat_solver =
   477       if sat_solver <> "smart" then
   478         if incremental andalso
   479            not (member (op =) (Kodkod_SAT.configured_sat_solvers true)
   480                        sat_solver) then
   481           (print_nt (K ("An incremental SAT solver is required: \"SAT4J\" will \
   482                         \be used instead of " ^ quote sat_solver ^ "."));
   483            "SAT4J")
   484         else
   485           sat_solver
   486       else
   487         Kodkod_SAT.smart_sat_solver_name incremental
   488     val _ =
   489       if sat_solver = "smart" then
   490         print_v (fn () =>
   491             "Using SAT solver " ^ quote actual_sat_solver ^ ". The following" ^
   492             (if incremental then " incremental " else " ") ^
   493             "solvers are configured: " ^
   494             commas_quote (Kodkod_SAT.configured_sat_solvers incremental) ^ ".")
   495       else
   496         ()
   497 
   498     val too_big_scopes = Unsynchronized.ref []
   499 
   500     fun problem_for_scope unsound
   501             (scope as {card_assigns, bits, bisim_depth, datatypes, ofs, ...}) =
   502       let
   503         val _ = not (exists (fn other => scope_less_eq other scope)
   504                             (!too_big_scopes)) orelse
   505                 raise TOO_LARGE ("Nitpick.pick_them_nits_in_term.\
   506                                  \problem_for_scope", "too large scope")
   507 (*
   508         val _ = print_g "Offsets:"
   509         val _ = List.app (fn (T, j0) =>
   510                              print_g (string_for_type ctxt T ^ " = " ^
   511                                     string_of_int j0))
   512                          (Typtab.dest ofs)
   513 *)
   514         val effective_total_consts =
   515           case total_consts of
   516             SOME b => b
   517           | NONE => forall (is_exact_type datatypes true) all_Ts
   518         val main_j0 = offset_of_type ofs bool_T
   519         val (nat_card, nat_j0) = spec_of_type scope nat_T
   520         val (int_card, int_j0) = spec_of_type scope int_T
   521         val _ = (nat_j0 = main_j0 andalso int_j0 = main_j0) orelse
   522                 raise BAD ("Nitpick.pick_them_nits_in_term.problem_for_scope",
   523                            "bad offsets")
   524         val kk = kodkod_constrs peephole_optim nat_card int_card main_j0
   525         val (free_names, rep_table) =
   526           choose_reps_for_free_vars scope pseudo_frees free_names
   527                                     NameTable.empty
   528         val (sel_names, rep_table) = choose_reps_for_all_sels scope rep_table
   529         val (nonsel_names, rep_table) =
   530           choose_reps_for_consts scope effective_total_consts nonsel_names
   531                                  rep_table
   532         val (guiltiest_party, min_highest_arity) =
   533           NameTable.fold (fn (name, R) => fn (s, n) =>
   534                              let val n' = arity_of_rep R in
   535                                if n' > n then (nickname_of name, n') else (s, n)
   536                              end) rep_table ("", 1)
   537         val min_univ_card =
   538           NameTable.fold (Integer.max o min_univ_card_of_rep o snd) rep_table
   539                          (univ_card nat_card int_card main_j0 [] KK.True)
   540         val _ = if debug then ()
   541                 else check_arity guiltiest_party min_univ_card min_highest_arity
   542 
   543         val def_us =
   544           def_us |> map (choose_reps_in_nut scope unsound rep_table true)
   545         val nondef_us =
   546           nondef_us |> map (choose_reps_in_nut scope unsound rep_table false)
   547         val need_us =
   548           need_us |> map (choose_reps_in_nut scope unsound rep_table false)
   549 (*
   550         val _ = List.app (print_g o string_for_nut ctxt)
   551                          (free_names @ sel_names @ nonsel_names @
   552                           nondef_us @ def_us @ need_us)
   553 *)
   554         val (free_rels, pool, rel_table) =
   555           rename_free_vars free_names initial_pool NameTable.empty
   556         val (sel_rels, pool, rel_table) =
   557           rename_free_vars sel_names pool rel_table
   558         val (other_rels, pool, rel_table) =
   559           rename_free_vars nonsel_names pool rel_table
   560         val nondef_us = nondef_us |> map (rename_vars_in_nut pool rel_table)
   561         val def_us = def_us |> map (rename_vars_in_nut pool rel_table)
   562         val need_us = need_us |> map (rename_vars_in_nut pool rel_table)
   563         val nondef_fs = map (kodkod_formula_from_nut ofs kk) nondef_us
   564         val def_fs = map (kodkod_formula_from_nut ofs kk) def_us
   565         val formula = fold (fold s_and) [def_fs, nondef_fs] KK.True
   566         val comment = (if unsound then "unsound" else "sound") ^ "\n" ^
   567                       Print_Mode.setmp [] multiline_string_for_scope scope
   568         val kodkod_sat_solver =
   569           Kodkod_SAT.sat_solver_spec actual_sat_solver |> snd
   570         val bit_width = if bits = 0 then 16 else bits + 1
   571         val delay =
   572           if unsound then
   573             Option.map (fn time => Time.- (time, Time.now ())) deadline
   574             |> unsound_delay_for_timeout
   575           else
   576             0
   577         val settings = [("solver", commas_quote kodkod_sat_solver),
   578                         ("bit_width", string_of_int bit_width),
   579                         ("symmetry_breaking", string_of_int kodkod_sym_break),
   580                         ("sharing", "3"),
   581                         ("flatten", "false"),
   582                         ("delay", signed_string_of_int delay)]
   583         val plain_rels = free_rels @ other_rels
   584         val plain_bounds = map (bound_for_plain_rel ctxt debug) plain_rels
   585         val plain_axioms = map (declarative_axiom_for_plain_rel kk) plain_rels
   586         val need_vals =
   587           map (fn dtype as {typ, ...} =>
   588                   (typ, needed_values_for_datatype need_us ofs dtype)) datatypes
   589         val sel_bounds =
   590           map (bound_for_sel_rel ctxt debug need_vals datatypes) sel_rels
   591         val dtype_axioms =
   592           declarative_axioms_for_datatypes hol_ctxt binarize need_us need_vals
   593               datatype_sym_break bits ofs kk rel_table datatypes
   594         val declarative_axioms = plain_axioms @ dtype_axioms
   595         val univ_card = Int.max (univ_card nat_card int_card main_j0
   596                                      (plain_bounds @ sel_bounds) formula,
   597                                  main_j0 |> bits > 0 ? Integer.add (bits + 1))
   598         val (built_in_bounds, built_in_axioms) =
   599           bounds_and_axioms_for_built_in_rels_in_formulas debug univ_card
   600               nat_card int_card main_j0 (formula :: declarative_axioms)
   601         val bounds = built_in_bounds @ plain_bounds @ sel_bounds
   602                      |> not debug ? merge_bounds
   603         val axioms = built_in_axioms @ declarative_axioms
   604         val highest_arity =
   605           fold Integer.max (map (fst o fst) (maps fst bounds)) 0
   606         val formula = fold_rev s_and axioms formula
   607         val _ = if bits = 0 then () else check_bits bits formula
   608         val _ = if debug orelse formula = KK.False then ()
   609                 else check_arity "" univ_card highest_arity
   610       in
   611         SOME ({comment = comment, settings = settings, univ_card = univ_card,
   612                tuple_assigns = [], bounds = bounds,
   613                int_bounds = if bits = 0 then sequential_int_bounds univ_card
   614                             else pow_of_two_int_bounds bits main_j0,
   615                expr_assigns = [], formula = formula},
   616               {free_names = free_names, sel_names = sel_names,
   617                nonsel_names = nonsel_names, rel_table = rel_table,
   618                unsound = unsound, scope = scope})
   619       end
   620       handle TOO_LARGE (loc, msg) =>
   621              if loc = "Nitpick_Kodkod.check_arity" andalso
   622                 not (Typtab.is_empty ofs) then
   623                problem_for_scope unsound
   624                    {hol_ctxt = hol_ctxt, binarize = binarize,
   625                     card_assigns = card_assigns, bits = bits,
   626                     bisim_depth = bisim_depth, datatypes = datatypes,
   627                     ofs = Typtab.empty}
   628              else if loc = "Nitpick.pick_them_nits_in_term.\
   629                            \problem_for_scope" then
   630                NONE
   631              else
   632                (Unsynchronized.change too_big_scopes (cons scope);
   633                 print_v (fn () =>
   634                     "Limit reached: " ^ msg ^ ". Skipping " ^
   635                     (if unsound then "potential component of " else "") ^
   636                     "scope.");
   637                 NONE)
   638            | TOO_SMALL (_, msg) =>
   639              (print_v (fn () =>
   640                   "Limit reached: " ^ msg ^ ". Skipping " ^
   641                   (if unsound then "potential component of " else "") ^
   642                   "scope.");
   643               NONE)
   644 
   645     val das_wort_model =
   646       (if falsify then "counterexample" else "model")
   647       |> not standard ? prefix "nonstandard "
   648 
   649     val scopes = Unsynchronized.ref []
   650     val generated_scopes = Unsynchronized.ref []
   651     val generated_problems = Unsynchronized.ref ([] : rich_problem list)
   652     val checked_problems = Unsynchronized.ref (SOME [])
   653     val met_potential = Unsynchronized.ref 0
   654 
   655     fun update_checked_problems problems =
   656       List.app (Unsynchronized.change checked_problems o Option.map o cons
   657                 o nth problems)
   658     fun show_kodkod_warning "" = ()
   659       | show_kodkod_warning s = print_nt (fn () => "Kodkod warning: " ^ s ^ ".")
   660 
   661     fun print_and_check_model genuine bounds
   662             ({free_names, sel_names, nonsel_names, rel_table, scope, ...}
   663              : problem_extension) =
   664       let
   665         val (reconstructed_model, codatatypes_ok) =
   666           reconstruct_hol_model
   667               {show_datatypes = show_datatypes, show_skolems = show_skolems,
   668                show_consts = show_consts}
   669               scope formats atomss real_frees pseudo_frees free_names sel_names
   670               nonsel_names rel_table bounds
   671         val genuine_means_genuine =
   672           got_all_user_axioms andalso none_true wfs andalso
   673           sound_finitizes andalso total_consts <> SOME true andalso
   674           codatatypes_ok
   675         fun assms_prop () = Logic.mk_conjunction_list (neg_t :: assm_ts)
   676       in
   677         (print_t (fn () =>
   678              "% SZS status " ^
   679              (if genuine then
   680                 if falsify then "CounterSatisfiable" else "Satisfiable"
   681               else
   682                 "Unknown") ^ "\n" ^
   683              "% SZS output start FiniteModel\n");
   684          pprint_a (Pretty.chunks
   685              [Pretty.blk (0,
   686                   (pstrs ((if mode = Auto_Try then "Auto " else "") ^
   687                           "Nitpick found a" ^
   688                           (if not genuine then " potentially spurious "
   689                            else if genuine_means_genuine then " "
   690                            else " quasi genuine ") ^ das_wort_model) @
   691                    (case pretties_for_scope scope verbose of
   692                       [] => []
   693                     | pretties => pstrs " for " @ pretties) @
   694                    [Pretty.str ":\n"])),
   695               Pretty.indent indent_size reconstructed_model]);
   696          print_t (fn () => "% SZS output stop\n");
   697          if genuine then
   698            (if check_genuine andalso standard then
   699               case prove_hol_model scope tac_timeout free_names sel_names
   700                                    rel_table bounds (assms_prop ()) of
   701                 SOME true =>
   702                 print ("Confirmation by \"auto\": The above " ^
   703                        das_wort_model ^ " is really genuine.")
   704               | SOME false =>
   705                 if genuine_means_genuine then
   706                   error ("A supposedly genuine " ^ das_wort_model ^ " was \
   707                          \shown to be spurious by \"auto\".\nThis should never \
   708                          \happen.\nPlease send a bug report to blanchet\
   709                          \te@in.tum.de.")
   710                  else
   711                    print ("Refutation by \"auto\": The above " ^
   712                           das_wort_model ^ " is spurious.")
   713                | NONE => print "No confirmation by \"auto\"."
   714             else
   715               ();
   716             if not standard andalso likely_in_struct_induct_step then
   717               print "The existence of a nonstandard model suggests that the \
   718                     \induction hypothesis is not general enough or may even be \
   719                     \wrong. See the Nitpick manual's \"Inductive Properties\" \
   720                     \section for details (\"isabelle doc nitpick\")."
   721             else
   722               ();
   723             if has_lonely_bool_var orig_t then
   724               print "Hint: Maybe you forgot a colon after the lemma's name?"
   725             else if has_syntactic_sorts orig_t then
   726               print "Hint: Maybe you forgot a type constraint?"
   727             else
   728               ();
   729             if not genuine_means_genuine then
   730               if no_poly_user_axioms then
   731                 let
   732                   val options =
   733                     [] |> not got_all_mono_user_axioms
   734                           ? cons ("user_axioms", "\"true\"")
   735                        |> not (none_true wfs)
   736                           ? cons ("wf", "\"smart\" or \"false\"")
   737                        |> not sound_finitizes
   738                           ? cons ("finitize", "\"smart\" or \"false\"")
   739                        |> total_consts = SOME true
   740                           ? cons ("total_consts", "\"smart\" or \"false\"")
   741                        |> not codatatypes_ok
   742                           ? cons ("bisim_depth", "a nonnegative value")
   743                   val ss =
   744                     map (fn (name, value) => quote name ^ " set to " ^ value)
   745                         options
   746                 in
   747                   print ("Try again with " ^
   748                          space_implode " " (serial_commas "and" ss) ^
   749                          " to confirm that the " ^ das_wort_model ^
   750                          " is genuine.")
   751                 end
   752               else
   753                 print ("Nitpick is unable to guarantee the authenticity of \
   754                        \the " ^ das_wort_model ^ " in the presence of \
   755                        \polymorphic axioms.")
   756             else
   757               ();
   758             NONE)
   759          else
   760            if not genuine then
   761              (Unsynchronized.inc met_potential;
   762               if check_potential then
   763                 let
   764                   val status = prove_hol_model scope tac_timeout free_names
   765                                                sel_names rel_table bounds
   766                                                (assms_prop ())
   767                 in
   768                   (case status of
   769                      SOME true => print ("Confirmation by \"auto\": The \
   770                                          \above " ^ das_wort_model ^
   771                                          " is genuine.")
   772                    | SOME false => print ("Refutation by \"auto\": The above " ^
   773                                           das_wort_model ^ " is spurious.")
   774                    | NONE => print "No confirmation by \"auto\".");
   775                   status
   776                 end
   777               else
   778                 NONE)
   779            else
   780              NONE)
   781         |> pair genuine_means_genuine
   782       end
   783     fun solve_any_problem (found_really_genuine, max_potential, max_genuine,
   784                            donno) first_time problems =
   785       let
   786         val max_potential = Int.max (0, max_potential)
   787         val max_genuine = Int.max (0, max_genuine)
   788         fun print_and_check genuine (j, bounds) =
   789           print_and_check_model genuine bounds (snd (nth problems j))
   790         val max_solutions = max_potential + max_genuine
   791                             |> not incremental ? Integer.min 1
   792       in
   793         if max_solutions <= 0 then
   794           (found_really_genuine, 0, 0, donno)
   795         else
   796           case KK.solve_any_problem debug overlord deadline max_threads
   797                                     max_solutions (map fst problems) of
   798             KK.JavaNotInstalled =>
   799             (print_nt install_java_message;
   800              (found_really_genuine, max_potential, max_genuine, donno + 1))
   801           | KK.JavaTooOld =>
   802             (print_nt install_java_message;
   803              (found_really_genuine, max_potential, max_genuine, donno + 1))
   804           | KK.KodkodiNotInstalled =>
   805             (print_nt install_kodkodi_message;
   806              (found_really_genuine, max_potential, max_genuine, donno + 1))
   807           | KK.Normal ([], unsat_js, s) =>
   808             (update_checked_problems problems unsat_js; show_kodkod_warning s;
   809              (found_really_genuine, max_potential, max_genuine, donno))
   810           | KK.Normal (sat_ps, unsat_js, s) =>
   811             let
   812               val (lib_ps, con_ps) =
   813                 List.partition (#unsound o snd o nth problems o fst) sat_ps
   814             in
   815               update_checked_problems problems (unsat_js @ map fst lib_ps);
   816               show_kodkod_warning s;
   817               if null con_ps then
   818                 let
   819                   val genuine_codes =
   820                     lib_ps |> take max_potential
   821                            |> map (print_and_check false)
   822                            |> filter (curry (op =) (SOME true) o snd)
   823                   val found_really_genuine =
   824                     found_really_genuine orelse exists fst genuine_codes
   825                   val num_genuine = length genuine_codes
   826                   val max_genuine = max_genuine - num_genuine
   827                   val max_potential = max_potential
   828                                       - (length lib_ps - num_genuine)
   829                 in
   830                   if max_genuine <= 0 then
   831                     (found_really_genuine, 0, 0, donno)
   832                   else
   833                     let
   834                       (* "co_js" is the list of sound problems whose unsound
   835                          pendants couldn't be satisfied and hence that most
   836                          probably can't be satisfied themselves. *)
   837                       val co_js =
   838                         map (fn j => j - 1) unsat_js
   839                         |> filter (fn j =>
   840                                       j >= 0 andalso
   841                                       scopes_equivalent
   842                                           (#scope (snd (nth problems j)),
   843                                            #scope (snd (nth problems (j + 1)))))
   844                       val bye_js = sort_distinct int_ord (map fst sat_ps @
   845                                                           unsat_js @ co_js)
   846                       val problems =
   847                         problems |> filter_out_indices bye_js
   848                                  |> max_potential <= 0
   849                                     ? filter_out (#unsound o snd)
   850                     in
   851                       solve_any_problem (found_really_genuine, max_potential,
   852                                          max_genuine, donno) false problems
   853                     end
   854                 end
   855               else
   856                 let
   857                   val genuine_codes =
   858                     con_ps |> take max_genuine
   859                            |> map (print_and_check true)
   860                   val max_genuine = max_genuine - length genuine_codes
   861                   val found_really_genuine =
   862                     found_really_genuine orelse exists fst genuine_codes
   863                 in
   864                   if max_genuine <= 0 orelse not first_time then
   865                     (found_really_genuine, 0, max_genuine, donno)
   866                   else
   867                     let
   868                       val bye_js = sort_distinct int_ord
   869                                                  (map fst sat_ps @ unsat_js)
   870                       val problems =
   871                         problems |> filter_out_indices bye_js
   872                                  |> filter_out (#unsound o snd)
   873                     in
   874                       solve_any_problem (found_really_genuine, 0, max_genuine,
   875                                          donno) false problems
   876                     end
   877                 end
   878             end
   879           | KK.TimedOut unsat_js =>
   880             (update_checked_problems problems unsat_js; raise TimeLimit.TimeOut)
   881           | KK.Error (s, unsat_js) =>
   882             (update_checked_problems problems unsat_js;
   883              print_v (K ("Kodkod error: " ^ s ^ "."));
   884              (found_really_genuine, max_potential, max_genuine, donno + 1))
   885       end
   886 
   887     fun run_batch j n scopes (found_really_genuine, max_potential, max_genuine,
   888                               donno) =
   889       let
   890         val _ =
   891           if null scopes then
   892             print_nt (K "The scope specification is inconsistent.")
   893           else if verbose then
   894             pprint_nt (fn () => Pretty.chunks
   895                 [Pretty.blk (0,
   896                      pstrs ((if n > 1 then
   897                                "Batch " ^ string_of_int (j + 1) ^ " of " ^
   898                                signed_string_of_int n ^ ": "
   899                              else
   900                                "") ^
   901                             "Trying " ^ string_of_int (length scopes) ^
   902                             " scope" ^ plural_s_for_list scopes ^ ":")),
   903                  Pretty.indent indent_size
   904                      (Pretty.chunks (map2
   905                           (fn j => fn scope =>
   906                               Pretty.block (
   907                                   (case pretties_for_scope scope true of
   908                                      [] => [Pretty.str "Empty"]
   909                                    | pretties => pretties) @
   910                                   [Pretty.str (if j = 1 then "." else ";")]))
   911                           (length scopes downto 1) scopes))])
   912           else
   913             ()
   914         fun add_problem_for_scope (scope, unsound) (problems, donno) =
   915           (check_deadline ();
   916            case problem_for_scope unsound scope of
   917              SOME problem =>
   918              (problems
   919               |> (null problems orelse
   920                   not (KK.problems_equivalent (fst problem, fst (hd problems))))
   921                   ? cons problem, donno)
   922            | NONE => (problems, donno + 1))
   923         val (problems, donno) =
   924           fold add_problem_for_scope
   925                (map_product pair scopes
   926                     ((if max_genuine > 0 then [false] else []) @
   927                      (if max_potential > 0 then [true] else [])))
   928                ([], donno)
   929         val _ = Unsynchronized.change generated_problems (append problems)
   930         val _ = Unsynchronized.change generated_scopes (append scopes)
   931         val _ =
   932           if j + 1 = n then
   933             let
   934               val (unsound_problems, sound_problems) =
   935                 List.partition (#unsound o snd) (!generated_problems)
   936             in
   937               if not (null sound_problems) andalso
   938                  forall (KK.is_problem_trivially_false o fst)
   939                         sound_problems then
   940                 print_nt (fn () =>
   941                     "Warning: The conjecture " ^
   942                     (if falsify then "either trivially holds"
   943                      else "is either trivially unsatisfiable") ^
   944                     " for the given scopes or lies outside Nitpick's supported \
   945                     \fragment." ^
   946                     (if exists (not o KK.is_problem_trivially_false o fst)
   947                                unsound_problems then
   948                        " Only potentially spurious " ^ das_wort_model ^
   949                        "s may be found."
   950                      else
   951                        ""))
   952               else
   953                 ()
   954             end
   955           else
   956             ()
   957       in
   958         solve_any_problem (found_really_genuine, max_potential, max_genuine,
   959                            donno) true (rev problems)
   960       end
   961 
   962     fun scope_count (problems : rich_problem list) scope =
   963       length (filter (curry scopes_equivalent scope o #scope o snd) problems)
   964     fun excipit did_so_and_so =
   965       let
   966         val do_filter =
   967           if !met_potential = max_potential then filter_out (#unsound o snd)
   968           else I
   969         val total = length (!scopes)
   970         val unsat =
   971           fold (fn scope =>
   972                    case scope_count (do_filter (!generated_problems)) scope of
   973                      0 => I
   974                    | n =>
   975                      scope_count (do_filter (these (!checked_problems)))
   976                                  scope = n
   977                      ? Integer.add 1) (!generated_scopes) 0
   978       in
   979         (if mode = TPTP then "% SZS status Unknown\n" else "") ^
   980         "Nitpick " ^ did_so_and_so ^
   981         (if is_some (!checked_problems) andalso total > 0 then
   982            " " ^ string_of_int unsat ^ " of " ^ string_of_int total ^ " scope"
   983            ^ plural_s total
   984          else
   985            "") ^ "."
   986       end
   987 
   988     val (skipped, the_scopes) =
   989       all_scopes hol_ctxt binarize cards_assigns maxes_assigns iters_assigns
   990                  bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs
   991                  finitizable_dataTs
   992     val _ = if skipped > 0 then
   993               print_nt (fn () => "Too many scopes. Skipping " ^
   994                                  string_of_int skipped ^ " scope" ^
   995                                  plural_s skipped ^
   996                                  ". (Consider using \"mono\" or \
   997                                  \\"merge_type_vars\" to prevent this.)")
   998             else
   999               ()
  1000     val _ = scopes := the_scopes
  1001 
  1002     fun run_batches _ _ []
  1003                     (found_really_genuine, max_potential, max_genuine, donno) =
  1004         if donno > 0 andalso max_genuine > 0 then
  1005           (print_nt (fn () => excipit "checked"); unknownN)
  1006         else if max_genuine = original_max_genuine then
  1007           if max_potential = original_max_potential then
  1008             (print_t (K "% SZS status Unknown\n");
  1009              print_nt (fn () =>
  1010                  "Nitpick found no " ^ das_wort_model ^ "." ^
  1011                  (if not standard andalso likely_in_struct_induct_step then
  1012                     " This suggests that the induction hypothesis might be \
  1013                     \general enough to prove this subgoal."
  1014                   else
  1015                     "")); if skipped > 0 then unknownN else noneN)
  1016           else
  1017             (print_nt (fn () =>
  1018                  excipit ("could not find a" ^
  1019                           (if max_genuine = 1 then
  1020                              " better " ^ das_wort_model ^ "."
  1021                            else
  1022                              "ny better " ^ das_wort_model ^ "s.") ^
  1023                           " It checked"));
  1024              potentialN)
  1025         else if found_really_genuine then
  1026           genuineN
  1027         else
  1028           quasi_genuineN
  1029       | run_batches j n (batch :: batches) z =
  1030         let val (z as (_, _, max_genuine, _)) = run_batch j n batch z in
  1031           run_batches (j + 1) n (if max_genuine > 0 then batches else []) z
  1032         end
  1033 
  1034     val batches = batch_list batch_size (!scopes)
  1035     val outcome_code =
  1036       run_batches 0 (length batches) batches
  1037                   (false, max_potential, max_genuine, 0)
  1038       handle TimeLimit.TimeOut =>
  1039              (print_nt (fn () => excipit "ran out of time after checking");
  1040               if !met_potential > 0 then potentialN else unknownN)
  1041     val _ = print_v (fn () =>
  1042                 "Total time: " ^ string_from_time (Timer.checkRealTimer timer) ^
  1043                 ".")
  1044   in (outcome_code, !state_ref) end
  1045 
  1046 (* Give the inner timeout a chance. *)
  1047 val timeout_bonus = seconds 1.0
  1048 
  1049 fun pick_nits_in_term state (params as {debug, timeout, expect, ...}) mode i n
  1050                       step subst assm_ts orig_t =
  1051   let val print_nt = if is_mode_nt mode then Output.urgent_message else K () in
  1052     if getenv "KODKODI" = "" then
  1053       (* The "expect" argument is deliberately ignored if Kodkodi is missing so
  1054          that the "Nitpick_Examples" can be processed on any machine. *)
  1055       (print_nt (Pretty.string_of (plazy install_kodkodi_message));
  1056        ("no_kodkodi", state))
  1057     else
  1058       let
  1059         val unknown_outcome = (unknownN, state)
  1060         val deadline = Option.map (curry Time.+ (Time.now ())) timeout
  1061         val outcome as (outcome_code, _) =
  1062           time_limit (if debug orelse is_none timeout then NONE
  1063                       else SOME (Time.+ (the timeout, timeout_bonus)))
  1064               (pick_them_nits_in_term deadline state params mode i n step subst
  1065                                       assm_ts) orig_t
  1066           handle TOO_LARGE (_, details) =>
  1067                  (print_nt ("Limit reached: " ^ details ^ "."); unknown_outcome)
  1068                | TOO_SMALL (_, details) =>
  1069                  (print_nt ("Limit reached: " ^ details ^ "."); unknown_outcome)
  1070                | Kodkod.SYNTAX (_, details) =>
  1071                  (print_nt ("Malformed Kodkodi output: " ^ details ^ ".");
  1072                   unknown_outcome)
  1073                | TimeLimit.TimeOut =>
  1074                  (print_nt "Nitpick ran out of time."; unknown_outcome)
  1075       in
  1076         if expect = "" orelse outcome_code = expect then outcome
  1077         else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
  1078       end
  1079   end
  1080 
  1081 fun is_fixed_equation ctxt
  1082                       (Const (@{const_name "=="}, _) $ Free (s, _) $ Const _) =
  1083     Variable.is_fixed ctxt s
  1084   | is_fixed_equation _ _ = false
  1085 fun extract_fixed_frees ctxt (assms, t) =
  1086   let
  1087     val (subst, other_assms) =
  1088       List.partition (is_fixed_equation ctxt) assms
  1089       |>> map Logic.dest_equals
  1090   in (subst, other_assms, subst_atomic subst t) end
  1091 
  1092 fun pick_nits_in_subgoal state params mode i step =
  1093   let
  1094     val ctxt = Proof.context_of state
  1095     val t = state |> Proof.raw_goal |> #goal |> prop_of
  1096   in
  1097     case Logic.count_prems t of
  1098       0 => (Output.urgent_message "No subgoal!"; (noneN, state))
  1099     | n =>
  1100       let
  1101         val t = Logic.goal_params t i |> fst
  1102         val assms = map term_of (Assumption.all_assms_of ctxt)
  1103         val (subst, assms, t) = extract_fixed_frees ctxt (assms, t)
  1104       in pick_nits_in_term state params mode i n step subst assms t end
  1105   end
  1106 
  1107 end;