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