src/HOL/Tools/Nitpick/nitpick.ML
author blanchet
Thu, 18 Feb 2010 18:48:07 +0100
changeset 35220 2bcdae5f4fdb
parent 35190 ce653cc27a94
child 35280 54ab4921f826
permissions -rw-r--r--
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
     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 params = {
    12     cards_assigns: (typ option * int list) list,
    13     maxes_assigns: (styp option * int list) list,
    14     iters_assigns: (styp option * int list) list,
    15     bitss: int list,
    16     bisim_depths: int list,
    17     boxes: (typ option * bool option) list,
    18     monos: (typ option * bool option) list,
    19     stds: (typ option * bool) list,
    20     wfs: (styp option * bool option) list,
    21     sat_solver: string,
    22     blocking: bool,
    23     falsify: bool,
    24     debug: bool,
    25     verbose: bool,
    26     overlord: bool,
    27     user_axioms: bool option,
    28     assms: bool,
    29     merge_type_vars: bool,
    30     binary_ints: bool option,
    31     destroy_constrs: bool,
    32     specialize: bool,
    33     skolemize: bool,
    34     star_linear_preds: bool,
    35     uncurry: bool,
    36     fast_descrs: bool,
    37     peephole_optim: bool,
    38     timeout: Time.time option,
    39     tac_timeout: Time.time option,
    40     sym_break: int,
    41     sharing_depth: int,
    42     flatten_props: bool,
    43     max_threads: int,
    44     show_skolems: bool,
    45     show_datatypes: bool,
    46     show_consts: bool,
    47     evals: term list,
    48     formats: (term option * int list) list,
    49     max_potential: int,
    50     max_genuine: int,
    51     check_potential: bool,
    52     check_genuine: bool,
    53     batch_size: int,
    54     expect: string}
    55 
    56   val register_frac_type : string -> (string * string) list -> theory -> theory
    57   val unregister_frac_type : string -> theory -> theory
    58   val register_codatatype : typ -> string -> styp list -> theory -> theory
    59   val unregister_codatatype : typ -> theory -> theory
    60   val pick_nits_in_term :
    61     Proof.state -> params -> bool -> int -> int -> int -> term list -> term
    62     -> string * Proof.state
    63   val pick_nits_in_subgoal :
    64     Proof.state -> params -> bool -> int -> int -> string * Proof.state
    65 end;
    66 
    67 structure Nitpick : NITPICK =
    68 struct
    69 
    70 open Nitpick_Util
    71 open Nitpick_HOL
    72 open Nitpick_Preproc
    73 open Nitpick_Mono
    74 open Nitpick_Scope
    75 open Nitpick_Peephole
    76 open Nitpick_Rep
    77 open Nitpick_Nut
    78 open Nitpick_Kodkod
    79 open Nitpick_Model
    80 
    81 structure KK = Kodkod
    82 
    83 type params = {
    84   cards_assigns: (typ option * int list) list,
    85   maxes_assigns: (styp option * int list) list,
    86   iters_assigns: (styp option * int list) list,
    87   bitss: int list,
    88   bisim_depths: int list,
    89   boxes: (typ option * bool option) list,
    90   monos: (typ option * bool option) list,
    91   stds: (typ option * bool) list,
    92   wfs: (styp option * bool option) list,
    93   sat_solver: string,
    94   blocking: bool,
    95   falsify: bool,
    96   debug: bool,
    97   verbose: bool,
    98   overlord: bool,
    99   user_axioms: bool option,
   100   assms: bool,
   101   merge_type_vars: bool,
   102   binary_ints: bool option,
   103   destroy_constrs: bool,
   104   specialize: bool,
   105   skolemize: bool,
   106   star_linear_preds: bool,
   107   uncurry: bool,
   108   fast_descrs: bool,
   109   peephole_optim: bool,
   110   timeout: Time.time option,
   111   tac_timeout: Time.time option,
   112   sym_break: int,
   113   sharing_depth: int,
   114   flatten_props: bool,
   115   max_threads: int,
   116   show_skolems: bool,
   117   show_datatypes: bool,
   118   show_consts: bool,
   119   evals: term list,
   120   formats: (term option * int list) list,
   121   max_potential: int,
   122   max_genuine: int,
   123   check_potential: bool,
   124   check_genuine: bool,
   125   batch_size: int,
   126   expect: string}
   127 
   128 type problem_extension = {
   129   free_names: nut list,
   130   sel_names: nut list,
   131   nonsel_names: nut list,
   132   rel_table: nut NameTable.table,
   133   unsound: bool,
   134   scope: scope,
   135   core: KK.formula,
   136   defs: KK.formula list}
   137 
   138 type rich_problem = KK.problem * problem_extension
   139 
   140 (* Proof.context -> string -> term list -> Pretty.T list *)
   141 fun pretties_for_formulas _ _ [] = []
   142   | pretties_for_formulas ctxt s ts =
   143     [Pretty.str (s ^ plural_s_for_list ts ^ ":"),
   144      Pretty.indent indent_size (Pretty.chunks
   145          (map2 (fn j => fn t =>
   146                    Pretty.block [t |> shorten_names_in_term
   147                                    |> Syntax.pretty_term ctxt,
   148                                  Pretty.str (if j = 1 then "." else ";")])
   149                (length ts downto 1) ts))]
   150 
   151 (* unit -> string *)
   152 fun install_kodkodi_message () =
   153   "Nitpick requires the external Java program Kodkodi. To install it, download \
   154   \the package from Isabelle's web page and add the \"kodkodi-x.y.z\" \
   155   \directory's full path to \"" ^
   156   Path.implode (Path.expand (Path.appends
   157       (Path.variable "ISABELLE_HOME_USER" ::
   158        map Path.basic ["etc", "components"]))) ^ "\"."
   159 
   160 val max_unsound_delay_ms = 200
   161 val max_unsound_delay_percent = 2
   162 
   163 (* Time.time option -> int *)
   164 fun unsound_delay_for_timeout NONE = max_unsound_delay_ms
   165   | unsound_delay_for_timeout (SOME timeout) =
   166     Int.max (0, Int.min (max_unsound_delay_ms,
   167                          Time.toMilliseconds timeout
   168                          * max_unsound_delay_percent div 100))
   169 
   170 (* Time.time option -> bool *)
   171 fun passed_deadline NONE = false
   172   | passed_deadline (SOME time) = Time.compare (Time.now (), time) <> LESS
   173 
   174 (* ('a * bool option) list -> bool *)
   175 fun none_true assigns = forall (not_equal (SOME true) o snd) assigns
   176 
   177 val syntactic_sorts =
   178   @{sort "{default,zero,one,plus,minus,uminus,times,inverse,abs,sgn,ord,eq}"} @
   179   @{sort number}
   180 (* typ -> bool *)
   181 fun has_tfree_syntactic_sort (TFree (_, S as _ :: _)) =
   182     subset (op =) (S, syntactic_sorts)
   183   | has_tfree_syntactic_sort _ = false
   184 (* term -> bool *)
   185 val has_syntactic_sorts = exists_type (exists_subtype has_tfree_syntactic_sort)
   186 
   187 (* (unit -> string) -> Pretty.T *)
   188 fun plazy f = Pretty.blk (0, pstrs (f ()))
   189 
   190 (* Time.time -> Proof.state -> params -> bool -> int -> int -> int -> term
   191    -> string * Proof.state *)
   192 fun pick_them_nits_in_term deadline state (params : params) auto i n step
   193                            orig_assm_ts orig_t =
   194   let
   195     val timer = Timer.startRealTimer ()
   196     val thy = Proof.theory_of state
   197     val ctxt = Proof.context_of state
   198 (* FIXME: reintroduce code before new release
   199     val nitpick_thy = ThyInfo.get_theory "Nitpick"
   200     val _ = Theory.subthy (nitpick_thy, thy) orelse
   201             error "You must import the theory \"Nitpick\" to use Nitpick"
   202 *)
   203     val {cards_assigns, maxes_assigns, iters_assigns, bitss, bisim_depths,
   204          boxes, monos, stds, wfs, sat_solver, blocking, falsify, debug, verbose,
   205          overlord, user_axioms, assms, merge_type_vars, binary_ints,
   206          destroy_constrs, specialize, skolemize, star_linear_preds, uncurry,
   207          fast_descrs, peephole_optim, tac_timeout, sym_break, sharing_depth,
   208          flatten_props, max_threads, show_skolems, show_datatypes, show_consts,
   209          evals, formats, max_potential, max_genuine, check_potential,
   210          check_genuine, batch_size, ...} =
   211       params
   212     val state_ref = Unsynchronized.ref state
   213     (* Pretty.T -> unit *)
   214     val pprint =
   215       if auto then
   216         Unsynchronized.change state_ref o Proof.goal_message o K
   217         o Pretty.chunks o cons (Pretty.str "") o single
   218         o Pretty.mark Markup.hilite
   219       else
   220         priority o Pretty.string_of
   221     (* (unit -> Pretty.T) -> unit *)
   222     fun pprint_m f = () |> not auto ? pprint o f
   223     fun pprint_v f = () |> verbose ? pprint o f
   224     fun pprint_d f = () |> debug ? pprint o f
   225     (* string -> unit *)
   226     val print = pprint o curry Pretty.blk 0 o pstrs
   227     (* (unit -> string) -> unit *)
   228     val print_m = pprint_m o K o plazy
   229     val print_v = pprint_v o K o plazy
   230     val print_d = pprint_d o K o plazy
   231 
   232     (* unit -> unit *)
   233     fun check_deadline () =
   234       if debug andalso passed_deadline deadline then raise TimeLimit.TimeOut
   235       else ()
   236     (* unit -> 'a *)
   237     fun do_interrupted () =
   238       if passed_deadline deadline then raise TimeLimit.TimeOut
   239       else raise Interrupt
   240 
   241     val _ =
   242       if step = 0 then
   243         print_m (fn () => "Nitpicking formula...")
   244       else
   245         pprint_m (fn () => Pretty.chunks (
   246             pretties_for_formulas ctxt ("Nitpicking " ^
   247                 (if i <> 1 orelse n <> 1 then
   248                    "subgoal " ^ string_of_int i ^ " of " ^ string_of_int n
   249                  else
   250                    "goal")) [Logic.list_implies (orig_assm_ts, orig_t)]))
   251     val neg_t = if falsify then Logic.mk_implies (orig_t, @{prop False})
   252                 else orig_t
   253     val assms_t = if assms orelse auto then
   254                     Logic.mk_conjunction_list (neg_t :: orig_assm_ts)
   255                   else
   256                     neg_t
   257     val (assms_t, evals) =
   258       assms_t :: evals |> merge_type_vars ? merge_type_vars_in_terms
   259                        |> pairf hd tl
   260     val original_max_potential = max_potential
   261     val original_max_genuine = max_genuine
   262 (*
   263     val _ = priority ("*** " ^ Syntax.string_of_term ctxt orig_t)
   264     val _ = List.app (fn t => priority ("*** " ^ Syntax.string_of_term ctxt t))
   265                      orig_assm_ts
   266 *)
   267     val max_bisim_depth = fold Integer.max bisim_depths ~1
   268     val case_names = case_const_names thy stds
   269     val (defs, built_in_nondefs, user_nondefs) = all_axioms_of thy
   270     val def_table = const_def_table ctxt defs
   271     val nondef_table = const_nondef_table (built_in_nondefs @ user_nondefs)
   272     val simp_table = Unsynchronized.ref (const_simp_table ctxt)
   273     val psimp_table = const_psimp_table ctxt
   274     val intro_table = inductive_intro_table ctxt def_table
   275     val ground_thm_table = ground_theorem_table thy
   276     val ersatz_table = ersatz_table thy
   277     val (hol_ctxt as {wf_cache, ...}) =
   278       {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
   279        stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug,
   280        binary_ints = binary_ints, destroy_constrs = destroy_constrs,
   281        specialize = specialize, skolemize = skolemize,
   282        star_linear_preds = star_linear_preds, uncurry = uncurry,
   283        fast_descrs = fast_descrs, tac_timeout = tac_timeout, evals = evals,
   284        case_names = case_names, def_table = def_table,
   285        nondef_table = nondef_table, user_nondefs = user_nondefs,
   286        simp_table = simp_table, psimp_table = psimp_table,
   287        intro_table = intro_table, ground_thm_table = ground_thm_table,
   288        ersatz_table = ersatz_table, skolems = Unsynchronized.ref [],
   289        special_funs = Unsynchronized.ref [],
   290        unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref [],
   291        constr_cache = Unsynchronized.ref []}
   292     val frees = Term.add_frees assms_t []
   293     val _ = null (Term.add_tvars assms_t []) orelse
   294             raise NOT_SUPPORTED "schematic type variables"
   295     val (((def_ts, nondef_ts), (got_all_mono_user_axioms, no_poly_user_axioms)),
   296          core_t, binarize) = preprocess_term hol_ctxt assms_t
   297     val got_all_user_axioms =
   298       got_all_mono_user_axioms andalso no_poly_user_axioms
   299 
   300     (* styp * (bool * bool) -> unit *)
   301     fun print_wf (x, (gfp, wf)) =
   302       pprint (Pretty.blk (0,
   303           pstrs ("The " ^ (if gfp then "co" else "") ^ "inductive predicate \"")
   304           @ Syntax.pretty_term ctxt (Const x) ::
   305           pstrs (if wf then
   306                    "\" was proved well-founded. Nitpick can compute it \
   307                    \efficiently."
   308                  else
   309                    "\" could not be proved well-founded. Nitpick might need to \
   310                    \unroll it.")))
   311     val _ = if verbose then List.app print_wf (!wf_cache) else ()
   312     val _ =
   313       pprint_d (fn () =>
   314           Pretty.chunks
   315               (pretties_for_formulas ctxt "Preprocessed formula" [core_t] @
   316                pretties_for_formulas ctxt "Relevant definitional axiom" def_ts @
   317                pretties_for_formulas ctxt "Relevant nondefinitional axiom"
   318                                      nondef_ts))
   319     val _ = List.app (ignore o Term.type_of) (core_t :: def_ts @ nondef_ts)
   320             handle TYPE (_, Ts, ts) =>
   321                    raise TYPE ("Nitpick.pick_them_nits_in_term", Ts, ts)
   322 
   323     val core_u = nut_from_term hol_ctxt Eq core_t
   324     val def_us = map (nut_from_term hol_ctxt DefEq) def_ts
   325     val nondef_us = map (nut_from_term hol_ctxt Eq) nondef_ts
   326     val (free_names, const_names) =
   327       fold add_free_and_const_names (core_u :: def_us @ nondef_us) ([], [])
   328     val (sel_names, nonsel_names) =
   329       List.partition (is_sel o nickname_of) const_names
   330     val genuine_means_genuine = got_all_user_axioms andalso none_true wfs
   331     val standard = forall snd stds
   332 (*
   333     val _ = List.app (priority o string_for_nut ctxt)
   334                      (core_u :: def_us @ nondef_us)
   335 *)
   336 
   337     val unique_scope = forall (curry (op =) 1 o length o snd) cards_assigns
   338     (* typ -> bool *)
   339     fun is_type_always_monotonic T =
   340       (is_datatype thy stds T andalso not (is_quot_type thy T) andalso
   341        (not (is_pure_typedef thy T) orelse is_univ_typedef thy T)) orelse
   342       is_number_type thy T orelse is_bit_type T
   343     fun is_type_monotonic T =
   344       unique_scope orelse
   345       case triple_lookup (type_match thy) monos T of
   346         SOME (SOME b) => b
   347       | _ => is_type_always_monotonic T orelse
   348              formulas_monotonic hol_ctxt binarize T Plus def_ts nondef_ts core_t
   349     fun is_deep_datatype T =
   350       is_datatype thy stds T andalso
   351       (not standard orelse T = nat_T orelse is_word_type T orelse
   352        exists (curry (op =) T o domain_type o type_of) sel_names)
   353     val all_Ts = ground_types_in_terms hol_ctxt binarize
   354                                        (core_t :: def_ts @ nondef_ts)
   355                  |> sort TermOrd.typ_ord
   356     val _ = if verbose andalso binary_ints = SOME true andalso
   357                exists (member (op =) [nat_T, int_T]) all_Ts then
   358               print_v (K "The option \"binary_ints\" will be ignored because \
   359                          \of the presence of rationals, reals, \"Suc\", \
   360                          \\"gcd\", or \"lcm\" in the problem.")
   361             else
   362               ()
   363     val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts
   364     val _ =
   365       if verbose andalso not unique_scope then
   366         case filter_out is_type_always_monotonic mono_Ts of
   367           [] => ()
   368         | interesting_mono_Ts =>
   369           print_v (fn () =>
   370                       let
   371                         val ss = map (quote o string_for_type ctxt)
   372                                      interesting_mono_Ts
   373                       in
   374                         "The type" ^ plural_s_for_list ss ^ " " ^
   375                         space_implode " " (serial_commas "and" ss) ^ " " ^
   376                         (if none_true monos then
   377                            "passed the monotonicity test"
   378                          else
   379                            (if length ss = 1 then "is" else "are") ^
   380                            " considered monotonic") ^
   381                         ". Nitpick might be able to skip some scopes."
   382                       end)
   383       else
   384         ()
   385     val deep_dataTs = filter is_deep_datatype all_Ts
   386     (* This detection code is an ugly hack. Fortunately, it is used only to
   387        provide a hint to the user. *)
   388     (* string * (Rule_Cases.T * bool) -> bool *)
   389     fun is_struct_induct_step (name, (Rule_Cases.Case {fixes, assumes, ...}, _)) =
   390       not (null fixes) andalso
   391       exists (String.isSuffix ".hyps" o fst) assumes andalso
   392       exists (exists (curry (op =) name o shortest_name o fst)
   393               o datatype_constrs hol_ctxt) deep_dataTs
   394     val likely_in_struct_induct_step =
   395       exists is_struct_induct_step (ProofContext.cases_of ctxt)
   396     val _ = if standard andalso likely_in_struct_induct_step then
   397               pprint_m (fn () => Pretty.blk (0,
   398                   pstrs "Hint: To check that the induction hypothesis is \
   399                         \general enough, try this command: " @
   400                   [Pretty.mark Markup.sendback (Pretty.blk (0,
   401                        pstrs ("nitpick [non_std, show_all]")))] @ pstrs "."))
   402             else
   403               ()
   404 (*
   405     val _ = priority "Monotonic types:"
   406     val _ = List.app (priority o string_for_type ctxt) mono_Ts
   407     val _ = priority "Nonmonotonic types:"
   408     val _ = List.app (priority o string_for_type ctxt) nonmono_Ts
   409 *)
   410 
   411     val need_incremental = Int.max (max_potential, max_genuine) >= 2
   412     val effective_sat_solver =
   413       if sat_solver <> "smart" then
   414         if need_incremental andalso
   415            not (member (op =) (Kodkod_SAT.configured_sat_solvers true)
   416                       sat_solver) then
   417           (print_m (K ("An incremental SAT solver is required: \"SAT4J\" will \
   418                        \be used instead of " ^ quote sat_solver ^ "."));
   419            "SAT4J")
   420         else
   421           sat_solver
   422       else
   423         Kodkod_SAT.smart_sat_solver_name need_incremental
   424     val _ =
   425       if sat_solver = "smart" then
   426         print_v (fn () => "Using SAT solver " ^ quote effective_sat_solver ^
   427                           ". The following" ^
   428                           (if need_incremental then " incremental " else " ") ^
   429                           "solvers are configured: " ^
   430                           commas (map quote (Kodkod_SAT.configured_sat_solvers
   431                                                        need_incremental)) ^ ".")
   432       else
   433         ()
   434 
   435     val too_big_scopes = Unsynchronized.ref []
   436 
   437     (* bool -> scope -> rich_problem option *)
   438     fun problem_for_scope unsound
   439             (scope as {card_assigns, bits, bisim_depth, datatypes, ofs, ...}) =
   440       let
   441         val _ = not (exists (fn other => scope_less_eq other scope)
   442                             (!too_big_scopes)) orelse
   443                 raise TOO_LARGE ("Nitpick.pick_them_nits_in_term.\
   444                                  \problem_for_scope", "too large scope")
   445 (*
   446         val _ = priority "Offsets:"
   447         val _ = List.app (fn (T, j0) =>
   448                              priority (string_for_type ctxt T ^ " = " ^
   449                                        string_of_int j0))
   450                          (Typtab.dest ofs)
   451 *)
   452         val all_exact = forall (is_exact_type datatypes) all_Ts
   453         (* nut list -> rep NameTable.table -> nut list * rep NameTable.table *)
   454         val repify_consts = choose_reps_for_consts scope all_exact
   455         val main_j0 = offset_of_type ofs bool_T
   456         val (nat_card, nat_j0) = spec_of_type scope nat_T
   457         val (int_card, int_j0) = spec_of_type scope int_T
   458         val _ = (nat_j0 = main_j0 andalso int_j0 = main_j0) orelse
   459                 raise BAD ("Nitpick.pick_them_nits_in_term.problem_for_scope",
   460                            "bad offsets")
   461         val kk = kodkod_constrs peephole_optim nat_card int_card main_j0
   462         val (free_names, rep_table) =
   463           choose_reps_for_free_vars scope free_names NameTable.empty
   464         val (sel_names, rep_table) = choose_reps_for_all_sels scope rep_table
   465         val (nonsel_names, rep_table) = repify_consts nonsel_names rep_table
   466         val min_highest_arity =
   467           NameTable.fold (curry Int.max o arity_of_rep o snd) rep_table 1
   468         val min_univ_card =
   469           NameTable.fold (curry Int.max o min_univ_card_of_rep o snd) rep_table
   470                          (univ_card nat_card int_card main_j0 [] KK.True)
   471         val _ = check_arity min_univ_card min_highest_arity
   472 
   473         val core_u = choose_reps_in_nut scope unsound rep_table false core_u
   474         val def_us = map (choose_reps_in_nut scope unsound rep_table true)
   475                          def_us
   476         val nondef_us = map (choose_reps_in_nut scope unsound rep_table false)
   477                             nondef_us
   478 (*
   479         val _ = List.app (priority o string_for_nut ctxt)
   480                          (free_names @ sel_names @ nonsel_names @
   481                           core_u :: def_us @ nondef_us)
   482 *)
   483         val (free_rels, pool, rel_table) =
   484           rename_free_vars free_names initial_pool NameTable.empty
   485         val (sel_rels, pool, rel_table) =
   486           rename_free_vars sel_names pool rel_table
   487         val (other_rels, pool, rel_table) =
   488           rename_free_vars nonsel_names pool rel_table
   489         val core_u = rename_vars_in_nut pool rel_table core_u
   490         val def_us = map (rename_vars_in_nut pool rel_table) def_us
   491         val nondef_us = map (rename_vars_in_nut pool rel_table) nondef_us
   492         val core_f = kodkod_formula_from_nut bits ofs kk core_u
   493         val def_fs = map (kodkod_formula_from_nut bits ofs kk) def_us
   494         val nondef_fs = map (kodkod_formula_from_nut bits ofs kk) nondef_us
   495         val formula = fold (fold s_and) [def_fs, nondef_fs] core_f
   496         val comment = (if unsound then "unsound" else "sound") ^ "\n" ^
   497                       PrintMode.setmp [] multiline_string_for_scope scope
   498         val kodkod_sat_solver =
   499           Kodkod_SAT.sat_solver_spec effective_sat_solver |> snd
   500         val bit_width = if bits = 0 then 16 else bits + 1
   501         val delay = if unsound then
   502                       Option.map (fn time => Time.- (time, Time.now ()))
   503                                  deadline
   504                       |> unsound_delay_for_timeout
   505                     else
   506                       0
   507         val settings = [("solver", commas (map quote kodkod_sat_solver)),
   508                         ("skolem_depth", "-1"),
   509                         ("bit_width", string_of_int bit_width),
   510                         ("symmetry_breaking", signed_string_of_int sym_break),
   511                         ("sharing", signed_string_of_int sharing_depth),
   512                         ("flatten", Bool.toString flatten_props),
   513                         ("delay", signed_string_of_int delay)]
   514         val plain_rels = free_rels @ other_rels
   515         val plain_bounds = map (bound_for_plain_rel ctxt debug) plain_rels
   516         val plain_axioms = map (declarative_axiom_for_plain_rel kk) plain_rels
   517         val sel_bounds = map (bound_for_sel_rel ctxt debug datatypes) sel_rels
   518         val dtype_axioms =
   519           declarative_axioms_for_datatypes hol_ctxt binarize bits ofs kk
   520                                            rel_table datatypes
   521         val declarative_axioms = plain_axioms @ dtype_axioms
   522         val univ_card = Int.max (univ_card nat_card int_card main_j0
   523                                      (plain_bounds @ sel_bounds) formula,
   524                                  main_j0 |> bits > 0 ? Integer.add (bits + 1))
   525         val built_in_bounds = bounds_for_built_in_rels_in_formula debug
   526                                   univ_card nat_card int_card main_j0 formula
   527         val bounds = built_in_bounds @ plain_bounds @ sel_bounds
   528                      |> not debug ? merge_bounds
   529         val highest_arity =
   530           fold Integer.max (map (fst o fst) (maps fst bounds)) 0
   531         val formula = fold_rev s_and declarative_axioms formula
   532         val _ = if bits = 0 then () else check_bits bits formula
   533         val _ = if formula = KK.False then ()
   534                 else check_arity univ_card highest_arity
   535       in
   536         SOME ({comment = comment, settings = settings, univ_card = univ_card,
   537                tuple_assigns = [], bounds = bounds,
   538                int_bounds =
   539                    if bits = 0 then sequential_int_bounds univ_card
   540                    else pow_of_two_int_bounds bits main_j0 univ_card,
   541                expr_assigns = [], formula = formula},
   542               {free_names = free_names, sel_names = sel_names,
   543                nonsel_names = nonsel_names, rel_table = rel_table,
   544                unsound = unsound, scope = scope, core = core_f,
   545                defs = nondef_fs @ def_fs @ declarative_axioms})
   546       end
   547       handle TOO_LARGE (loc, msg) =>
   548              if loc = "Nitpick_Kodkod.check_arity" andalso
   549                 not (Typtab.is_empty ofs) then
   550                problem_for_scope unsound
   551                    {hol_ctxt = hol_ctxt, binarize = binarize,
   552                     card_assigns = card_assigns, bits = bits,
   553                     bisim_depth = bisim_depth, datatypes = datatypes,
   554                     ofs = Typtab.empty}
   555              else if loc = "Nitpick.pick_them_nits_in_term.\
   556                            \problem_for_scope" then
   557                NONE
   558              else
   559                (Unsynchronized.change too_big_scopes (cons scope);
   560                 print_v (fn () => ("Limit reached: " ^ msg ^
   561                                    ". Skipping " ^ (if unsound then "potential"
   562                                                     else "genuine") ^
   563                                    " component of scope."));
   564                 NONE)
   565            | TOO_SMALL (loc, msg) =>
   566              (print_v (fn () => ("Limit reached: " ^ msg ^
   567                                  ". Skipping " ^ (if unsound then "potential"
   568                                                   else "genuine") ^
   569                                  " component of scope."));
   570               NONE)
   571 
   572     (* int -> (''a * int list list) list -> ''a -> KK.tuple_set *)
   573     fun tuple_set_for_rel univ_card =
   574       KK.TupleSet o map (kk_tuple debug univ_card) o the oo AList.lookup (op =)
   575 
   576     val das_wort_model =
   577       (if falsify then "counterexample" else "model")
   578       |> not standard ? prefix "nonstandard "
   579 
   580     val scopes = Unsynchronized.ref []
   581     val generated_scopes = Unsynchronized.ref []
   582     val generated_problems = Unsynchronized.ref ([] : rich_problem list)
   583     val checked_problems = Unsynchronized.ref (SOME [])
   584     val met_potential = Unsynchronized.ref 0
   585 
   586     (* rich_problem list -> int list -> unit *)
   587     fun update_checked_problems problems =
   588       List.app (Unsynchronized.change checked_problems o Option.map o cons
   589                 o nth problems)
   590 
   591     (* bool -> KK.raw_bound list -> problem_extension -> bool option *)
   592     fun print_and_check_model genuine bounds
   593             ({free_names, sel_names, nonsel_names, rel_table, scope, ...}
   594              : problem_extension) =
   595       let
   596         val (reconstructed_model, codatatypes_ok) =
   597           reconstruct_hol_model {show_skolems = show_skolems,
   598                                  show_datatypes = show_datatypes,
   599                                  show_consts = show_consts}
   600               scope formats frees free_names sel_names nonsel_names rel_table
   601               bounds
   602         val genuine_means_genuine = genuine_means_genuine andalso codatatypes_ok
   603       in
   604         pprint (Pretty.chunks
   605             [Pretty.blk (0,
   606                  (pstrs ("Nitpick found a" ^
   607                          (if not genuine then " potential "
   608                           else if genuine_means_genuine then " "
   609                           else " likely genuine ") ^ das_wort_model) @
   610                   (case pretties_for_scope scope verbose of
   611                      [] => []
   612                    | pretties => pstrs " for " @ pretties) @
   613                   [Pretty.str ":\n"])),
   614              Pretty.indent indent_size reconstructed_model]);
   615         if genuine then
   616           (if check_genuine andalso standard then
   617              (case prove_hol_model scope tac_timeout free_names sel_names
   618                                    rel_table bounds assms_t of
   619                 SOME true => print ("Confirmation by \"auto\": The above " ^
   620                                     das_wort_model ^ " is really genuine.")
   621               | SOME false =>
   622                 if genuine_means_genuine then
   623                   error ("A supposedly genuine " ^ das_wort_model ^ " was \
   624                          \shown to be spurious by \"auto\".\nThis should never \
   625                          \happen.\nPlease send a bug report to blanchet\
   626                          \te@in.tum.de.")
   627                 else
   628                   print ("Refutation by \"auto\": The above " ^ das_wort_model ^
   629                          " is spurious.")
   630               | NONE => print "No confirmation by \"auto\".")
   631            else
   632              ();
   633            if not standard andalso likely_in_struct_induct_step then
   634              print "The existence of a nonstandard model suggests that the \
   635                    \induction hypothesis is not general enough or perhaps even \
   636                    \wrong. See the \"Inductive Properties\" section of the \
   637                    \Nitpick manual for details (\"isabelle doc nitpick\")."
   638            else
   639              ();
   640            if has_syntactic_sorts orig_t then
   641              print "Hint: Maybe you forgot a type constraint?"
   642            else
   643              ();
   644            if not genuine_means_genuine then
   645              if no_poly_user_axioms then
   646                let
   647                  val options =
   648                    [] |> not got_all_mono_user_axioms
   649                          ? cons ("user_axioms", "\"true\"")
   650                       |> not (none_true wfs)
   651                          ? cons ("wf", "\"smart\" or \"false\"")
   652                       |> not codatatypes_ok
   653                          ? cons ("bisim_depth", "a nonnegative value")
   654                  val ss =
   655                    map (fn (name, value) => quote name ^ " set to " ^ value)
   656                        options
   657                in
   658                  print ("Try again with " ^
   659                         space_implode " " (serial_commas "and" ss) ^
   660                         " to confirm that the " ^ das_wort_model ^
   661                         " is genuine.")
   662                end
   663              else
   664                print ("Nitpick is unable to guarantee the authenticity of \
   665                       \the " ^ das_wort_model ^ " in the presence of \
   666                       \polymorphic axioms.")
   667            else
   668              ();
   669            NONE)
   670         else
   671           if not genuine then
   672             (Unsynchronized.inc met_potential;
   673              if check_potential then
   674                let
   675                  val status = prove_hol_model scope tac_timeout free_names
   676                                               sel_names rel_table bounds assms_t
   677                in
   678                  (case status of
   679                     SOME true => print ("Confirmation by \"auto\": The above " ^
   680                                         das_wort_model ^ " is genuine.")
   681                   | SOME false => print ("Refutation by \"auto\": The above " ^
   682                                          das_wort_model ^ " is spurious.")
   683                   | NONE => print "No confirmation by \"auto\".");
   684                  status
   685                end
   686              else
   687                NONE)
   688           else
   689             NONE
   690       end
   691     (* int -> int -> int -> bool -> rich_problem list -> int * int * int *)
   692     fun solve_any_problem max_potential max_genuine donno first_time problems =
   693       let
   694         val max_potential = Int.max (0, max_potential)
   695         val max_genuine = Int.max (0, max_genuine)
   696         (* bool -> int * KK.raw_bound list -> bool option *)
   697         fun print_and_check genuine (j, bounds) =
   698           print_and_check_model genuine bounds (snd (nth problems j))
   699         val max_solutions = max_potential + max_genuine
   700                             |> not need_incremental ? curry Int.min 1
   701       in
   702         if max_solutions <= 0 then
   703           (0, 0, donno)
   704         else
   705           case KK.solve_any_problem overlord deadline max_threads max_solutions
   706                                     (map fst problems) of
   707             KK.NotInstalled =>
   708             (print_m install_kodkodi_message;
   709              (max_potential, max_genuine, donno + 1))
   710           | KK.Normal ([], unsat_js) =>
   711             (update_checked_problems problems unsat_js;
   712              (max_potential, max_genuine, donno))
   713           | KK.Normal (sat_ps, unsat_js) =>
   714             let
   715               val (lib_ps, con_ps) =
   716                 List.partition (#unsound o snd o nth problems o fst) sat_ps
   717             in
   718               update_checked_problems problems (unsat_js @ map fst lib_ps);
   719               if null con_ps then
   720                 let
   721                   val num_genuine = take max_potential lib_ps
   722                                     |> map (print_and_check false)
   723                                     |> filter (curry (op =) (SOME true))
   724                                     |> length
   725                   val max_genuine = max_genuine - num_genuine
   726                   val max_potential = max_potential
   727                                       - (length lib_ps - num_genuine)
   728                 in
   729                   if max_genuine <= 0 then
   730                     (0, 0, donno)
   731                   else
   732                     let
   733                       (* "co_js" is the list of sound problems whose unsound
   734                          pendants couldn't be satisfied and hence that most
   735                          probably can't be satisfied themselves. *)
   736                       val co_js =
   737                         map (fn j => j - 1) unsat_js
   738                         |> filter (fn j =>
   739                                       j >= 0 andalso
   740                                       scopes_equivalent
   741                                           (#scope (snd (nth problems j)))
   742                                           (#scope (snd (nth problems (j + 1)))))
   743                       val bye_js = sort_distinct int_ord (map fst sat_ps @
   744                                                           unsat_js @ co_js)
   745                       val problems =
   746                         problems |> filter_out_indices bye_js
   747                                  |> max_potential <= 0
   748                                     ? filter_out (#unsound o snd)
   749                     in
   750                       solve_any_problem max_potential max_genuine donno false
   751                                         problems
   752                     end
   753                 end
   754               else
   755                 let
   756                   val _ = take max_genuine con_ps
   757                           |> List.app (ignore o print_and_check true)
   758                   val max_genuine = max_genuine - length con_ps
   759                 in
   760                   if max_genuine <= 0 orelse not first_time then
   761                     (0, max_genuine, donno)
   762                   else
   763                     let
   764                       val bye_js = sort_distinct int_ord
   765                                                  (map fst sat_ps @ unsat_js)
   766                       val problems =
   767                         problems |> filter_out_indices bye_js
   768                                  |> filter_out (#unsound o snd)
   769                     in solve_any_problem 0 max_genuine donno false problems end
   770                 end
   771             end
   772           | KK.TimedOut unsat_js =>
   773             (update_checked_problems problems unsat_js; raise TimeLimit.TimeOut)
   774           | KK.Interrupted NONE => (checked_problems := NONE; do_interrupted ())
   775           | KK.Interrupted (SOME unsat_js) =>
   776             (update_checked_problems problems unsat_js; do_interrupted ())
   777           | KK.Error (s, unsat_js) =>
   778             (update_checked_problems problems unsat_js;
   779              print_v (K ("Kodkod error: " ^ s ^ "."));
   780              (max_potential, max_genuine, donno + 1))
   781       end
   782 
   783     (* int -> int -> scope list -> int * int * int -> int * int * int *)
   784     fun run_batch j n scopes (max_potential, max_genuine, donno) =
   785       let
   786         val _ =
   787           if null scopes then
   788             print_m (K "The scope specification is inconsistent.")
   789           else if verbose then
   790             pprint (Pretty.chunks
   791                 [Pretty.blk (0,
   792                      pstrs ((if n > 1 then
   793                                "Batch " ^ string_of_int (j + 1) ^ " of " ^
   794                                signed_string_of_int n ^ ": "
   795                              else
   796                                "") ^
   797                             "Trying " ^ string_of_int (length scopes) ^
   798                             " scope" ^ plural_s_for_list scopes ^ ":")),
   799                  Pretty.indent indent_size
   800                      (Pretty.chunks (map2
   801                           (fn j => fn scope =>
   802                               Pretty.block (
   803                                   (case pretties_for_scope scope true of
   804                                      [] => [Pretty.str "Empty"]
   805                                    | pretties => pretties) @
   806                                   [Pretty.str (if j = 1 then "." else ";")]))
   807                           (length scopes downto 1) scopes))])
   808           else
   809             ()
   810         (* scope * bool -> rich_problem list * bool
   811            -> rich_problem list * bool *)
   812         fun add_problem_for_scope (scope as {datatypes, ...}, unsound)
   813                                   (problems, donno) =
   814           (check_deadline ();
   815            case problem_for_scope unsound scope of
   816              SOME problem =>
   817              (problems
   818               |> (null problems orelse
   819                   not (KK.problems_equivalent (fst problem)
   820                                               (fst (hd problems))))
   821                   ? cons problem, donno)
   822            | NONE => (problems, donno + 1))
   823         val (problems, donno) =
   824           fold add_problem_for_scope
   825                (map_product pair scopes
   826                     ((if max_genuine > 0 then [false] else []) @
   827                      (if max_potential > 0 then [true] else [])))
   828                ([], donno)
   829         val _ = Unsynchronized.change generated_problems (append problems)
   830         val _ = Unsynchronized.change generated_scopes (append scopes)
   831         val _ =
   832           if j + 1 = n then
   833             let
   834               val (unsound_problems, sound_problems) =
   835                 List.partition (#unsound o snd) (!generated_problems)
   836             in
   837               if not (null sound_problems) andalso
   838                  forall (KK.is_problem_trivially_false o fst)
   839                         sound_problems then
   840                 print_m (fn () =>
   841                     "Warning: The conjecture either trivially holds for the \
   842                     \given scopes or (more likely) lies outside Nitpick's \
   843                     \supported fragment." ^
   844                     (if exists (not o KK.is_problem_trivially_false o fst)
   845                                unsound_problems then
   846                        " Only potential counterexamples may be found."
   847                      else
   848                        ""))
   849               else
   850                 ()
   851             end
   852           else
   853             ()
   854       in
   855         solve_any_problem max_potential max_genuine donno true (rev problems)
   856       end
   857 
   858     (* rich_problem list -> scope -> int *)
   859     fun scope_count (problems : rich_problem list) scope =
   860       length (filter (scopes_equivalent scope o #scope o snd) problems)
   861     (* string -> string *)
   862     fun excipit did_so_and_so =
   863       let
   864         (* rich_problem list -> rich_problem list *)
   865         val do_filter =
   866           if !met_potential = max_potential then filter_out (#unsound o snd)
   867           else I
   868         val total = length (!scopes)
   869         val unsat =
   870           fold (fn scope =>
   871                    case scope_count (do_filter (!generated_problems)) scope of
   872                      0 => I
   873                    | n =>
   874                      scope_count (do_filter (these (!checked_problems)))
   875                                  scope = n
   876                      ? Integer.add 1) (!generated_scopes) 0
   877       in
   878         "Nitpick " ^ did_so_and_so ^
   879         (if is_some (!checked_problems) andalso total > 0 then
   880            " after checking " ^
   881            string_of_int (Int.min (total - 1, unsat)) ^ " of " ^
   882            string_of_int total ^ " scope" ^ plural_s total
   883          else
   884            "") ^ "."
   885       end
   886 
   887     (* int -> int -> scope list -> int * int * int -> KK.outcome *)
   888     fun run_batches _ _ [] (max_potential, max_genuine, donno) =
   889         if donno > 0 andalso max_genuine > 0 then
   890           (print_m (fn () => excipit "ran out of some resource"); "unknown")
   891         else if max_genuine = original_max_genuine then
   892           if max_potential = original_max_potential then
   893             (print_m (fn () =>
   894                  "Nitpick found no " ^ das_wort_model ^ "." ^
   895                  (if not standard andalso likely_in_struct_induct_step then
   896                     " This suggests that the induction hypothesis might be \
   897                     \general enough to prove this subgoal."
   898                   else
   899                     "")); "none")
   900           else
   901             (print_m (fn () =>
   902                  "Nitpick could not find a" ^
   903                  (if max_genuine = 1 then " better " ^ das_wort_model ^ "."
   904                   else "ny better " ^ das_wort_model ^ "s.")); "potential")
   905         else
   906           if genuine_means_genuine then "genuine" else "likely_genuine"
   907       | run_batches j n (batch :: batches) z =
   908         let val (z as (_, max_genuine, _)) = run_batch j n batch z in
   909           run_batches (j + 1) n (if max_genuine > 0 then batches else []) z
   910         end
   911 
   912     val (skipped, the_scopes) =
   913       all_scopes hol_ctxt binarize sym_break cards_assigns maxes_assigns
   914                  iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs
   915     val _ = if skipped > 0 then
   916               print_m (fn () => "Too many scopes. Skipping " ^
   917                                 string_of_int skipped ^ " scope" ^
   918                                 plural_s skipped ^
   919                                 ". (Consider using \"mono\" or \
   920                                 \\"merge_type_vars\" to prevent this.)")
   921             else
   922               ()
   923     val _ = scopes := the_scopes
   924 
   925     val batches = batch_list batch_size (!scopes)
   926     val outcome_code =
   927       (run_batches 0 (length batches) batches (max_potential, max_genuine, 0)
   928        handle Exn.Interrupt => do_interrupted ())
   929       handle TimeLimit.TimeOut =>
   930              (print_m (fn () => excipit "ran out of time");
   931               if !met_potential > 0 then "potential" else "unknown")
   932            | Exn.Interrupt => if auto orelse debug then raise Interrupt
   933                               else error (excipit "was interrupted")
   934     val _ = print_v (fn () => "Total time: " ^
   935                               signed_string_of_int (Time.toMilliseconds
   936                                     (Timer.checkRealTimer timer)) ^ " ms.")
   937   in (outcome_code, !state_ref) end
   938   handle Exn.Interrupt =>
   939          if auto orelse #debug params then
   940            raise Interrupt
   941          else
   942            if passed_deadline deadline then
   943              (priority "Nitpick ran out of time."; ("unknown", state))
   944            else
   945              error "Nitpick was interrupted."
   946 
   947 (* Proof.state -> params -> bool -> int -> int -> int -> term
   948    -> string * Proof.state *)
   949 fun pick_nits_in_term state (params as {debug, timeout, expect, ...}) auto i n
   950                       step orig_assm_ts orig_t =
   951   if getenv "KODKODI" = "" then
   952     (if auto then ()
   953      else warning (Pretty.string_of (plazy install_kodkodi_message));
   954      ("unknown", state))
   955   else
   956     let
   957       val deadline = Option.map (curry Time.+ (Time.now ())) timeout
   958       val outcome as (outcome_code, _) =
   959         time_limit (if debug then NONE else timeout)
   960             (pick_them_nits_in_term deadline state params auto i n step
   961                                     orig_assm_ts) orig_t
   962     in
   963       if expect = "" orelse outcome_code = expect then outcome
   964       else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
   965     end
   966 
   967 (* Proof.state -> params -> bool -> int -> int -> string * Proof.state *)
   968 fun pick_nits_in_subgoal state params auto i step =
   969   let
   970     val ctxt = Proof.context_of state
   971     val t = state |> Proof.raw_goal |> #goal |> prop_of
   972   in
   973     case Logic.count_prems t of
   974       0 => (priority "No subgoal!"; ("none", state))
   975     | n =>
   976       let
   977         val assms = map term_of (Assumption.all_assms_of ctxt)
   978         val (t, frees) = Logic.goal_params t i
   979       in
   980         pick_nits_in_term state params auto i n step assms
   981                           (subst_bounds (frees, t))
   982       end
   983   end
   984 
   985 end;