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