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