src/HOL/Tools/Nitpick/nitpick_isar.ML
author blanchet
Thu, 03 Mar 2011 11:20:48 +0100
changeset 42747 03f699556955
parent 42746 e3cd0dce9b1a
child 42864 bd6296de1432
permissions -rw-r--r--
simplify "need" option's syntax
     1 (*  Title:      HOL/Tools/Nitpick/nitpick_isar.ML
     2     Author:     Jasmin Blanchette, TU Muenchen
     3     Copyright   2008, 2009, 2010
     4 
     5 Adds the "nitpick" and "nitpick_params" commands to Isabelle/Isar's outer
     6 syntax.
     7 *)
     8 
     9 signature NITPICK_ISAR =
    10 sig
    11   type params = Nitpick.params
    12 
    13   val auto : bool Unsynchronized.ref
    14   val default_params : theory -> (string * string) list -> params
    15   val setup : theory -> theory
    16 end;
    17 
    18 structure Nitpick_Isar : NITPICK_ISAR =
    19 struct
    20 
    21 open Nitpick_Util
    22 open Nitpick_HOL
    23 open Nitpick_Rep
    24 open Nitpick_Nut
    25 open Nitpick
    26 
    27 val auto = Unsynchronized.ref false
    28 
    29 (* Maximum number of scopes for Auto Nitpick. Be frugal since it has to share
    30    its time slot with several other automatic tools. *)
    31 val max_auto_scopes = 6
    32 
    33 val _ =
    34   ProofGeneralPgip.add_preference Preferences.category_tracing
    35       (Preferences.bool_pref auto "auto-nitpick"
    36            "Run Nitpick automatically.")
    37 
    38 type raw_param = string * string list
    39 
    40 val default_default_params =
    41   [("card", "1\<midarrow>10"),
    42    ("iter", "0,1,2,4,8,12,16,20,24,28"),
    43    ("bits", "1,2,3,4,6,8,10,12,14,16"),
    44    ("bisim_depth", "9"),
    45    ("box", "smart"),
    46    ("finitize", "smart"),
    47    ("mono", "smart"),
    48    ("std", "true"),
    49    ("wf", "smart"),
    50    ("sat_solver", "smart"),
    51    ("batch_size", "smart"),
    52    ("blocking", "true"),
    53    ("falsify", "true"),
    54    ("user_axioms", "smart"),
    55    ("assms", "true"),
    56    ("merge_type_vars", "false"),
    57    ("binary_ints", "smart"),
    58    ("destroy_constrs", "true"),
    59    ("specialize", "true"),
    60    ("star_linear_preds", "true"),
    61    ("total_consts", "smart"),
    62    ("peephole_optim", "true"),
    63    ("datatype_sym_break", "5"),
    64    ("kodkod_sym_break", "15"),
    65    ("timeout", "30"),
    66    ("tac_timeout", "0.5"),
    67    ("max_threads", "0"),
    68    ("debug", "false"),
    69    ("verbose", "false"),
    70    ("overlord", "false"),
    71    ("show_datatypes", "false"),
    72    ("show_consts", "false"),
    73    ("format", "1"),
    74    ("max_potential", "1"),
    75    ("max_genuine", "1"),
    76    ("check_potential", "false"),
    77    ("check_genuine", "false")]
    78 
    79 val negated_params =
    80   [("dont_box", "box"),
    81    ("dont_finitize", "finitize"),
    82    ("non_mono", "mono"),
    83    ("non_std", "std"),
    84    ("non_wf", "wf"),
    85    ("non_blocking", "blocking"),
    86    ("satisfy", "falsify"),
    87    ("no_user_axioms", "user_axioms"),
    88    ("no_assms", "assms"),
    89    ("dont_merge_type_vars", "merge_type_vars"),
    90    ("unary_ints", "binary_ints"),
    91    ("dont_destroy_constrs", "destroy_constrs"),
    92    ("dont_specialize", "specialize"),
    93    ("dont_star_linear_preds", "star_linear_preds"),
    94    ("partial_consts", "total_consts"),
    95    ("no_peephole_optim", "peephole_optim"),
    96    ("no_debug", "debug"),
    97    ("quiet", "verbose"),
    98    ("no_overlord", "overlord"),
    99    ("hide_datatypes", "show_datatypes"),
   100    ("hide_consts", "show_consts"),
   101    ("trust_potential", "check_potential"),
   102    ("trust_genuine", "check_genuine")]
   103 
   104 fun is_known_raw_param s =
   105   AList.defined (op =) default_default_params s orelse
   106   AList.defined (op =) negated_params s orelse
   107   member (op =) ["max", "show_all", "whack", "eval", "need", "atoms",
   108                  "expect"] s orelse
   109   exists (fn p => String.isPrefix (p ^ " ") s)
   110          ["card", "max", "iter", "box", "dont_box", "finitize", "dont_finitize",
   111           "mono", "non_mono", "std", "non_std", "wf", "non_wf", "format",
   112           "atoms"]
   113 
   114 fun check_raw_param (s, _) =
   115   if is_known_raw_param s then ()
   116   else error ("Unknown parameter: " ^ quote s ^ ".")
   117 
   118 fun unnegate_param_name name =
   119   case AList.lookup (op =) negated_params name of
   120     NONE => if String.isPrefix "dont_" name then SOME (unprefix "dont_" name)
   121             else if String.isPrefix "non_" name then SOME (unprefix "non_" name)
   122             else NONE
   123   | some_name => some_name
   124 fun normalize_raw_param (name, value) =
   125   case unnegate_param_name name of
   126     SOME name' => [(name', case value of
   127                              ["false"] => ["true"]
   128                            | ["true"] => ["false"]
   129                            | [] => ["false"]
   130                            | _ => value)]
   131   | NONE => if name = "show_all" then
   132               [("show_datatypes", value), ("show_consts", value)]
   133             else
   134               [(name, value)]
   135 
   136 structure Data = Theory_Data
   137 (
   138   type T = raw_param list
   139   val empty = map (apsnd single) default_default_params
   140   val extend = I
   141   fun merge data = AList.merge (op =) (K true) data
   142 )
   143 
   144 val set_default_raw_param =
   145   Data.map o fold (AList.update (op =)) o normalize_raw_param
   146 val default_raw_params = Data.get
   147 
   148 fun is_punctuation s = (s = "," orelse s = "-" orelse s = "\<midarrow>")
   149 
   150 fun stringify_raw_param_value [] = ""
   151   | stringify_raw_param_value [s] = s
   152   | stringify_raw_param_value (s1 :: s2 :: ss) =
   153     s1 ^ (if is_punctuation s1 orelse is_punctuation s2 then "" else " ") ^
   154     stringify_raw_param_value (s2 :: ss)
   155 
   156 fun maxed_int_from_string min_int s = Int.max (min_int, the (Int.fromString s))
   157 
   158 fun extract_params ctxt auto default_params override_params =
   159   let
   160     val override_params = maps normalize_raw_param override_params
   161     val raw_params = rev override_params @ rev default_params
   162     val raw_lookup = AList.lookup (op =) raw_params
   163     val lookup = Option.map stringify_raw_param_value o raw_lookup
   164     val lookup_string = the_default "" o lookup
   165     fun general_lookup_bool option default_value name =
   166       case lookup name of
   167         SOME s => parse_bool_option option name s
   168       | NONE => default_value
   169     val lookup_bool = the o general_lookup_bool false (SOME false)
   170     val lookup_bool_option = general_lookup_bool true NONE
   171     fun do_int name value =
   172       case value of
   173         SOME s => (case Int.fromString s of
   174                      SOME i => i
   175                    | NONE => error ("Parameter " ^ quote name ^
   176                                     " must be assigned an integer value."))
   177       | NONE => 0
   178     fun lookup_int name = do_int name (lookup name)
   179     fun lookup_int_option name =
   180       case lookup name of
   181         SOME "smart" => NONE
   182       | value => SOME (do_int name value)
   183     fun int_range_from_string name min_int s =
   184       let
   185         val (k1, k2) =
   186           (case space_explode "-" s of
   187              [s] => the_default (s, s) (first_field "\<midarrow>" s)
   188            | ["", s2] => ("-" ^ s2, "-" ^ s2)
   189            | [s1, s2] => (s1, s2)
   190            | _ => raise Option)
   191           |> pairself (maxed_int_from_string min_int)
   192       in if k1 <= k2 then k1 upto k2 else k1 downto k2 end
   193       handle Option.Option =>
   194              error ("Parameter " ^ quote name ^
   195                     " must be assigned a sequence of integers.")
   196     fun int_seq_from_string name min_int s =
   197       maps (int_range_from_string name min_int) (space_explode "," s)
   198     fun lookup_int_seq name min_int =
   199       case lookup name of
   200         SOME s => (case int_seq_from_string name min_int s of
   201                      [] => [min_int]
   202                    | value => value)
   203       | NONE => [min_int]
   204     fun lookup_assigns read prefix default convert =
   205       (NONE, convert (the_default default (lookup prefix)))
   206       :: map (fn (name, value) =>
   207                  (SOME (read (String.extract (name, size prefix + 1, NONE))),
   208                   convert (stringify_raw_param_value value)))
   209              (filter (String.isPrefix (prefix ^ " ") o fst) raw_params)
   210     fun lookup_ints_assigns read prefix min_int =
   211       lookup_assigns read prefix (signed_string_of_int min_int)
   212                      (int_seq_from_string prefix min_int)
   213     fun lookup_bool_assigns read prefix =
   214       lookup_assigns read prefix "" (the o parse_bool_option false prefix)
   215     fun lookup_bool_option_assigns read prefix =
   216       lookup_assigns read prefix "" (parse_bool_option true prefix)
   217     fun lookup_strings_assigns read prefix =
   218       lookup_assigns read prefix "" (space_explode " ")
   219     fun lookup_time name =
   220       case lookup name of
   221         SOME s => parse_time_option name s
   222       | NONE => NONE
   223     val read_type_polymorphic =
   224       Syntax.read_typ ctxt #> Logic.mk_type
   225       #> singleton (Variable.polymorphic ctxt) #> Logic.dest_type
   226     val read_term_polymorphic =
   227       Syntax.read_term ctxt #> singleton (Variable.polymorphic ctxt)
   228     val lookup_term_list_option_polymorphic =
   229       AList.lookup (op =) raw_params #> Option.map (map read_term_polymorphic)
   230     val read_const_polymorphic = read_term_polymorphic #> dest_Const
   231     val cards_assigns = lookup_ints_assigns read_type_polymorphic "card" 1
   232                         |> auto ? map (apsnd (take max_auto_scopes))
   233     val maxes_assigns = lookup_ints_assigns read_const_polymorphic "max" ~1
   234     val iters_assigns = lookup_ints_assigns read_const_polymorphic "iter" 0
   235     val bitss = lookup_int_seq "bits" 1
   236     val bisim_depths = lookup_int_seq "bisim_depth" ~1
   237     val boxes = lookup_bool_option_assigns read_type_polymorphic "box"
   238     val finitizes = lookup_bool_option_assigns read_type_polymorphic "finitize"
   239     val monos = if auto then [(NONE, SOME true)]
   240                 else lookup_bool_option_assigns read_type_polymorphic "mono"
   241     val stds = lookup_bool_assigns read_type_polymorphic "std"
   242     val wfs = lookup_bool_option_assigns read_const_polymorphic "wf"
   243     val sat_solver = lookup_string "sat_solver"
   244     val blocking = auto orelse lookup_bool "blocking"
   245     val falsify = lookup_bool "falsify"
   246     val debug = not auto andalso lookup_bool "debug"
   247     val verbose = debug orelse (not auto andalso lookup_bool "verbose")
   248     val overlord = lookup_bool "overlord"
   249     val user_axioms = lookup_bool_option "user_axioms"
   250     val assms = lookup_bool "assms"
   251     val whacks = lookup_term_list_option_polymorphic "whack" |> these
   252     val merge_type_vars = lookup_bool "merge_type_vars"
   253     val binary_ints = lookup_bool_option "binary_ints"
   254     val destroy_constrs = lookup_bool "destroy_constrs"
   255     val specialize = lookup_bool "specialize"
   256     val star_linear_preds = lookup_bool "star_linear_preds"
   257     val total_consts = lookup_bool_option "total_consts"
   258     val needs = lookup_term_list_option_polymorphic "need"
   259     val peephole_optim = lookup_bool "peephole_optim"
   260     val datatype_sym_break = lookup_int "datatype_sym_break"
   261     val kodkod_sym_break = lookup_int "kodkod_sym_break"
   262     val timeout = if auto then NONE else lookup_time "timeout"
   263     val tac_timeout = lookup_time "tac_timeout"
   264     val max_threads = if auto then 1 else Int.max (0, lookup_int "max_threads")
   265     val show_datatypes = debug orelse lookup_bool "show_datatypes"
   266     val show_consts = debug orelse lookup_bool "show_consts"
   267     val evals = lookup_term_list_option_polymorphic "eval" |> these
   268     val formats = lookup_ints_assigns read_term_polymorphic "format" 0
   269     val atomss = lookup_strings_assigns read_type_polymorphic "atoms"
   270     val max_potential =
   271       if auto then 0 else Int.max (0, lookup_int "max_potential")
   272     val max_genuine = Int.max (0, lookup_int "max_genuine")
   273     val check_potential = lookup_bool "check_potential"
   274     val check_genuine = lookup_bool "check_genuine"
   275     val batch_size =
   276       case lookup_int_option "batch_size" of
   277         SOME n => Int.max (1, n)
   278       | NONE => if debug then 1 else 50
   279     val expect = lookup_string "expect"
   280   in
   281     {cards_assigns = cards_assigns, maxes_assigns = maxes_assigns,
   282      iters_assigns = iters_assigns, bitss = bitss, bisim_depths = bisim_depths,
   283      boxes = boxes, finitizes = finitizes, monos = monos, stds = stds,
   284      wfs = wfs, sat_solver = sat_solver, blocking = blocking, falsify = falsify,
   285      debug = debug, verbose = verbose, overlord = overlord,
   286      user_axioms = user_axioms, assms = assms, whacks = whacks,
   287      merge_type_vars = merge_type_vars, binary_ints = binary_ints,
   288      destroy_constrs = destroy_constrs, specialize = specialize,
   289      star_linear_preds = star_linear_preds, total_consts = total_consts,
   290      needs = needs, peephole_optim = peephole_optim,
   291      datatype_sym_break = datatype_sym_break,
   292      kodkod_sym_break = kodkod_sym_break, timeout = timeout,
   293      tac_timeout = tac_timeout, max_threads = max_threads,
   294      show_datatypes = show_datatypes, show_consts = show_consts,
   295      evals = evals, formats = formats, atomss = atomss,
   296      max_potential = max_potential, max_genuine = max_genuine,
   297      check_potential = check_potential, check_genuine = check_genuine,
   298      batch_size = batch_size, expect = expect}
   299   end
   300 
   301 fun default_params thy =
   302   extract_params (ProofContext.init_global thy) false (default_raw_params thy)
   303   o map (apsnd single)
   304 
   305 val parse_key = Scan.repeat1 Parse.typ_group >> space_implode " "
   306 val parse_value =
   307   Scan.repeat1 (Parse.minus >> single
   308                 || Scan.repeat1 (Scan.unless Parse.minus
   309                                              (Parse.name || Parse.float_number))
   310                 || Parse.$$$ "," |-- Parse.number >> prefix "," >> single)
   311   >> flat
   312 val parse_param = parse_key -- Scan.optional (Parse.$$$ "=" |-- parse_value) []
   313 val parse_params =
   314   Scan.optional (Parse.$$$ "[" |-- Parse.list parse_param --| Parse.$$$ "]") []
   315 
   316 fun handle_exceptions ctxt f x =
   317   f x
   318   handle ARG (loc, details) =>
   319          error ("Bad argument(s) to " ^ quote loc ^ ": " ^ details ^ ".")
   320        | BAD (loc, details) =>
   321          error ("Internal error (" ^ quote loc ^ "): " ^ details ^ ".")
   322        | NOT_SUPPORTED details =>
   323          (warning ("Unsupported case: " ^ details ^ "."); x)
   324        | NUT (loc, us) =>
   325          error ("Invalid intermediate term" ^ plural_s_for_list us ^
   326                 " (" ^ quote loc ^ "): " ^
   327                 commas (map (string_for_nut ctxt) us) ^ ".")
   328        | REP (loc, Rs) =>
   329          error ("Invalid representation" ^ plural_s_for_list Rs ^
   330                 " (" ^ quote loc ^ "): " ^ commas (map string_for_rep Rs) ^ ".")
   331        | TERM (loc, ts) =>
   332          error ("Invalid term" ^ plural_s_for_list ts ^
   333                 " (" ^ quote loc ^ "): " ^
   334                 commas (map (Syntax.string_of_term ctxt) ts) ^ ".")
   335        | TYPE (loc, Ts, ts) =>
   336          error ("Invalid type" ^ plural_s_for_list Ts ^
   337                 (if null ts then
   338                    ""
   339                  else
   340                    " for term" ^ plural_s_for_list ts ^ " " ^
   341                    commas (map (quote o Syntax.string_of_term ctxt) ts)) ^
   342                 " (" ^ quote loc ^ "): " ^
   343                 commas (map (Syntax.string_of_typ ctxt) Ts) ^ ".")
   344        | Refute.REFUTE (loc, details) =>
   345          error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^ ".")
   346 
   347 fun pick_nits override_params auto i step state =
   348   let
   349     val thy = Proof.theory_of state
   350     val ctxt = Proof.context_of state
   351     val _ = List.app check_raw_param override_params
   352     val params as {blocking, debug, ...} =
   353       extract_params ctxt auto (default_raw_params thy) override_params
   354     fun go () =
   355       (false, state)
   356       |> (if auto then perhaps o try
   357           else if debug then fn f => fn x => f x
   358           else handle_exceptions ctxt)
   359          (fn (_, state) => pick_nits_in_subgoal state params auto i step
   360                            |>> curry (op =) "genuine")
   361   in if blocking then go () else Future.fork (tap go) |> K (false, state) end
   362 
   363 fun nitpick_trans (params, i) =
   364   Toplevel.keep (fn st =>
   365       (pick_nits params false i (Toplevel.proof_position_of st)
   366                  (Toplevel.proof_of st); ()))
   367 
   368 fun string_for_raw_param (name, value) =
   369   name ^ " = " ^ stringify_raw_param_value value
   370 
   371 fun nitpick_params_trans params =
   372   Toplevel.theory
   373       (fold set_default_raw_param params
   374        #> tap (fn thy =>
   375                   writeln ("Default parameters for Nitpick:\n" ^
   376                            (case rev (default_raw_params thy) of
   377                               [] => "none"
   378                             | params =>
   379                               (map check_raw_param params;
   380                                params |> map string_for_raw_param
   381                                       |> sort_strings |> cat_lines)))))
   382 
   383 val parse_nitpick_command =
   384   (parse_params -- Scan.optional Parse.nat 1) #>> nitpick_trans
   385 val parse_nitpick_params_command = parse_params #>> nitpick_params_trans
   386 
   387 val _ = Outer_Syntax.improper_command "nitpick"
   388             "try to find a counterexample for a given subgoal using Nitpick"
   389             Keyword.diag parse_nitpick_command
   390 val _ = Outer_Syntax.command "nitpick_params"
   391             "set and display the default parameters for Nitpick"
   392             Keyword.thy_decl parse_nitpick_params_command
   393 
   394 val auto_nitpick = pick_nits [] true 1 0
   395 
   396 val setup = Auto_Tools.register_tool (auto, auto_nitpick)
   397 
   398 end;