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