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