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