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