src/HOL/Tools/try0.ML
author wenzelm
Sat, 30 Dec 2023 21:54:08 +0100
changeset 79396 d08460213400
parent 78239 4fe65149f3fd
child 79742 2e4518e8a36b
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
46641
8801a24f9e9a renamed 'try_methods' to 'try0'
blanchet
parents: 45666
diff changeset
     1
(*  Title:      HOL/Tools/try0.ML
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     3
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     4
Try a combination of proof methods.
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     5
*)
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     6
46641
8801a24f9e9a renamed 'try_methods' to 'try0'
blanchet
parents: 45666
diff changeset
     7
signature TRY0 =
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     8
sig
43020
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
     9
  val noneN : string
55177
b7ca9f98faca made 'try0' (more) silent
blanchet
parents: 54431
diff changeset
    10
  val silence_methods : bool -> Proof.context -> Proof.context
55179
blanchet
parents: 55178
diff changeset
    11
  val try0 : Time.time option -> string list * string list * string list * string list ->
blanchet
parents: 55178
diff changeset
    12
    Proof.state -> bool
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    13
end;
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    14
46641
8801a24f9e9a renamed 'try_methods' to 'try0'
blanchet
parents: 45666
diff changeset
    15
structure Try0 : TRY0 =
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    16
struct
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    17
55179
blanchet
parents: 55178
diff changeset
    18
val noneN = "none";
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
    19
55179
blanchet
parents: 55178
diff changeset
    20
datatype mode = Auto_Try | Try | Normal;
43020
abb5d1f907e4 added "try" command, to launch Solve Direct, Quickcheck, Nitpick, Sledgehammer, and Try Methods
blanchet
parents: 43018
diff changeset
    21
55179
blanchet
parents: 55178
diff changeset
    22
val default_timeout = seconds 5.0;
38944
827c98e8ba8b fiddling with "try"
blanchet
parents: 38942
diff changeset
    23
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    24
fun can_apply timeout_opt pre post tac st =
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    25
  let val {goal, ...} = Proof.goal st in
55182
dd1e95e67b30 more robust w.r.t. exceptions raised by proof methods
blanchet
parents: 55181
diff changeset
    26
    (case (case timeout_opt of
74870
d54b3c96ee50 more robust physical timeout (despite 1bea05713dde), especially relevant for quickcheck where large unary numerals may cause excessive heap allocations and resulting GC is better included in the timing;
wenzelm
parents: 74508
diff changeset
    27
            SOME timeout => Timeout.apply_physical timeout
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    28
          | NONE => fn f => fn x => f x) (Seq.pull o tac) (pre st) of
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59184
diff changeset
    29
      SOME (x, _) => Thm.nprems_of (post x) < Thm.nprems_of goal
55182
dd1e95e67b30 more robust w.r.t. exceptions raised by proof methods
blanchet
parents: 55181
diff changeset
    30
    | NONE => false)
dd1e95e67b30 more robust w.r.t. exceptions raised by proof methods
blanchet
parents: 55181
diff changeset
    31
  end;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    32
55181
blanchet
parents: 55179
diff changeset
    33
fun apply_generic timeout_opt name command pre post apply st =
73386
3fb201ca8fb5 tuned --- more elementary Time operations;
wenzelm
parents: 69593
diff changeset
    34
  let val time_start = Time.now () in
55182
dd1e95e67b30 more robust w.r.t. exceptions raised by proof methods
blanchet
parents: 55181
diff changeset
    35
    if try (can_apply timeout_opt pre post apply) st = SOME true then
73386
3fb201ca8fb5 tuned --- more elementary Time operations;
wenzelm
parents: 69593
diff changeset
    36
      SOME (name, command, Time.toMilliseconds (Time.now () - time_start))
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    37
    else
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    38
      NONE
55179
blanchet
parents: 55178
diff changeset
    39
  end;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    40
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58903
diff changeset
    41
fun parse_method keywords s =
57918
f5d73caba4e5 tuned signature according to Scala version -- prefer explicit argument;
wenzelm
parents: 56982
diff changeset
    42
  enclose "(" ")" s
59083
88b0b1f28adc tuned signature;
wenzelm
parents: 59058
diff changeset
    43
  |> Token.explode keywords Position.start
57918
f5d73caba4e5 tuned signature according to Scala version -- prefer explicit argument;
wenzelm
parents: 56982
diff changeset
    44
  |> filter Token.is_proper
f5d73caba4e5 tuned signature according to Scala version -- prefer explicit argument;
wenzelm
parents: 56982
diff changeset
    45
  |> Scan.read Token.stopper Method.parse
f5d73caba4e5 tuned signature according to Scala version -- prefer explicit argument;
wenzelm
parents: 56982
diff changeset
    46
  |> (fn SOME (Method.Source src, _) => src | _ => raise Fail "expected Source");
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    47
55742
a989bdaf8121 modernized Method.check_name/check_source (with reports) vs. strict Method.the_method (without interning nor reports), e.g. relevant for semantic completion;
wenzelm
parents: 55182
diff changeset
    48
fun apply_named_method_on_first_goal ctxt =
58928
23d0ffd48006 plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
wenzelm
parents: 58903
diff changeset
    49
  parse_method (Thy_Header.get_keywords' ctxt)
55997
9dc5ce83202c modernized Attrib.check_name/check_src similar to methods (see also a989bdaf8121);
wenzelm
parents: 55765
diff changeset
    50
  #> Method.method_cmd ctxt
55182
dd1e95e67b30 more robust w.r.t. exceptions raised by proof methods
blanchet
parents: 55181
diff changeset
    51
  #> Method.Basic
58005
c28e6bc6635d more compact datatypes;
wenzelm
parents: 57918
diff changeset
    52
  #> (fn m => Method.Combinator (Method.no_combinator_info, Method.Select_Goals 1, [m]))
55182
dd1e95e67b30 more robust w.r.t. exceptions raised by proof methods
blanchet
parents: 55181
diff changeset
    53
  #> Proof.refine;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    54
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    55
fun add_attr_text (NONE, _) s = s
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    56
  | add_attr_text (_, []) s = s
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    57
  | add_attr_text (SOME x, fs) s =
55179
blanchet
parents: 55178
diff changeset
    58
    s ^ " " ^ (if x = "" then "" else x ^ ": ") ^ space_implode " " fs;
blanchet
parents: 55178
diff changeset
    59
42179
24662b614fd4 added support for "dest:" to "try"
blanchet
parents: 41999
diff changeset
    60
fun attrs_text (sx, ix, ex, dx) (ss, is, es, ds) =
55179
blanchet
parents: 55178
diff changeset
    61
  "" |> fold add_attr_text [(sx, ss), (ix, is), (ex, es), (dx, ds)];
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    62
55181
blanchet
parents: 55179
diff changeset
    63
fun apply_named_method (name, ((all_goals, run_if_auto_try), attrs)) mode timeout_opt quad st =
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
    64
  if mode <> Auto_Try orelse run_if_auto_try then
42179
24662b614fd4 added support for "dest:" to "try"
blanchet
parents: 41999
diff changeset
    65
    let val attrs = attrs_text attrs quad in
55181
blanchet
parents: 55179
diff changeset
    66
      apply_generic timeout_opt name
54291
709a2bbd7638 by (auto ...)[1] not by (auto [1])
blanchet
parents: 54248
diff changeset
    67
        ((name ^ attrs |> attrs <> "" ? enclose "(" ")") ^
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59184
diff changeset
    68
         (if all_goals andalso Thm.nprems_of (#goal (Proof.goal st)) > 1 then "[1]" else ""))
54291
709a2bbd7638 by (auto ...)[1] not by (auto [1])
blanchet
parents: 54248
diff changeset
    69
        I (#goal o Proof.goal)
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60350
diff changeset
    70
        (apply_named_method_on_first_goal (Proof.context_of st) (name ^ attrs)
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 60350
diff changeset
    71
          #> Seq.filter_results) st
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    72
    end
40222
cd6d2b0a4096 reintroduced Auto Try, but this time really off by default -- and leave some classical+simp reasoners out for Auto Try (but keep them for Try)
blanchet
parents: 40132
diff changeset
    73
  else
55179
blanchet
parents: 55178
diff changeset
    74
    NONE;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    75
55179
blanchet
parents: 55178
diff changeset
    76
val full_attrs = (SOME "simp", SOME "intro", SOME "elim", SOME "dest");
blanchet
parents: 55178
diff changeset
    77
val clas_attrs = (NONE, SOME "intro", SOME "elim", SOME "dest");
blanchet
parents: 55178
diff changeset
    78
val simp_attrs = (SOME "add", NONE, NONE, NONE);
blanchet
parents: 55178
diff changeset
    79
val metis_attrs = (SOME "", SOME "", SOME "", SOME "");
blanchet
parents: 55178
diff changeset
    80
val no_attrs = (NONE, NONE, NONE, NONE);
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    81
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
    82
(* name * ((all_goals, run_if_auto_try), (simp, intro, elim, dest) *)
39547
5df45da44bfb reorder proof methods and take out "best";
blanchet
parents: 39336
diff changeset
    83
val named_methods =
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    84
  [("simp", ((false, true), simp_attrs)),
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    85
   ("auto", ((true, true), full_attrs)),
55178
318cd8ac1817 added 'algebra' and 'meson' to 'try0'
blanchet
parents: 55177
diff changeset
    86
   ("blast", ((false, true), clas_attrs)),
318cd8ac1817 added 'algebra' and 'meson' to 'try0'
blanchet
parents: 55177
diff changeset
    87
   ("metis", ((false, true), metis_attrs)),
63961
2fd9656c4c82 invoke argo as part of the tried automatic proof methods
boehmes
parents: 63690
diff changeset
    88
   ("argo", ((false, true), no_attrs)),
55178
318cd8ac1817 added 'algebra' and 'meson' to 'try0'
blanchet
parents: 55177
diff changeset
    89
   ("linarith", ((false, true), no_attrs)),
318cd8ac1817 added 'algebra' and 'meson' to 'try0'
blanchet
parents: 55177
diff changeset
    90
   ("presburger", ((false, true), no_attrs)),
318cd8ac1817 added 'algebra' and 'meson' to 'try0'
blanchet
parents: 55177
diff changeset
    91
   ("algebra", ((false, true), no_attrs)),
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    92
   ("fast", ((false, false), clas_attrs)),
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44651
diff changeset
    93
   ("fastforce", ((false, false), full_attrs)),
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
    94
   ("force", ((false, false), full_attrs)),
56850
13a7bca533a3 added 'satx' proof method to Try0
blanchet
parents: 56467
diff changeset
    95
   ("meson", ((false, false), metis_attrs)),
78239
4fe65149f3fd add proof method "order" to command "try0"
desharna
parents: 74870
diff changeset
    96
   ("satx", ((false, false), no_attrs)),
4fe65149f3fd add proof method "order" to command "try0"
desharna
parents: 74870
diff changeset
    97
   ("order", ((false, true), no_attrs))];
55179
blanchet
parents: 55178
diff changeset
    98
55181
blanchet
parents: 55179
diff changeset
    99
val apply_methods = map apply_named_method named_methods;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   100
55179
blanchet
parents: 55178
diff changeset
   101
fun time_string ms = string_of_int ms ^ " ms";
blanchet
parents: 55178
diff changeset
   102
fun tool_time_string (s, ms) = s ^ ": " ^ time_string ms;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   103
55177
b7ca9f98faca made 'try0' (more) silent
blanchet
parents: 54431
diff changeset
   104
(* Makes reconstructor tools as silent as possible. The "set_visible" calls suppresses "Unification
b7ca9f98faca made 'try0' (more) silent
blanchet
parents: 54431
diff changeset
   105
   bound exceeded" warnings and the like. *)
b7ca9f98faca made 'try0' (more) silent
blanchet
parents: 54431
diff changeset
   106
fun silence_methods debug =
b7ca9f98faca made 'try0' (more) silent
blanchet
parents: 54431
diff changeset
   107
  Config.put Metis_Tactic.verbose debug
60275
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   108
  #> not debug ? (fn ctxt =>
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   109
      ctxt
62984
61b32a6d87e9 more silence;
wenzelm
parents: 62969
diff changeset
   110
      |> Simplifier_Trace.disable
60275
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   111
      |> Context_Position.set_visible false
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   112
      |> Config.put Unify.trace_bound (Config.get ctxt Unify.search_bound)
63961
2fd9656c4c82 invoke argo as part of the tried automatic proof methods
boehmes
parents: 63690
diff changeset
   113
      |> Config.put Argo_Tactic.trace "none"
60275
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   114
      |> Proof_Context.background_theory (fn thy =>
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   115
          thy
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   116
          |> Context_Position.set_visible_global false
d8a4fe35da00 silence local Unify.trace_bound as well: existing tools either refer to Proof.context or theory;
wenzelm
parents: 60190
diff changeset
   117
          |> Config.put_global Unify.trace_bound (Config.get_global thy Unify.search_bound)));
55177
b7ca9f98faca made 'try0' (more) silent
blanchet
parents: 54431
diff changeset
   118
55181
blanchet
parents: 55179
diff changeset
   119
fun generic_try0 mode timeout_opt quad st =
41038
9592334001d5 quiet Metis in "try"
blanchet
parents: 40931
diff changeset
   120
  let
56982
51d4189d95cf silence methods even better
blanchet
parents: 56850
diff changeset
   121
    val st = Proof.map_contexts (silence_methods false) st;
55179
blanchet
parents: 55178
diff changeset
   122
    fun trd (_, _, t) = t;
67225
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   123
    fun try_method method = method mode timeout_opt quad st;
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   124
    fun get_message (_, command, ms) = "Found proof: " ^ Active.sendback_markup_command
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   125
      ((if Thm.nprems_of (#goal (Proof.goal st)) = 1 then "by" else "apply") ^ " " ^ command) ^
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   126
      " (" ^ time_string ms ^ ")";
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   127
    val print_step = Option.map (tap (writeln o get_message));
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   128
    val get_results =
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   129
      if mode = Normal
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   130
      then Par_List.map (try_method #> print_step) #> map_filter I #> sort (int_ord o apply2 trd)
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   131
      else Par_List.get_some try_method #> the_list;
41038
9592334001d5 quiet Metis in "try"
blanchet
parents: 40931
diff changeset
   132
  in
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
   133
    if mode = Normal then
55179
blanchet
parents: 55178
diff changeset
   134
      "Trying " ^ space_implode " " (Try.serial_commas "and" (map (quote o fst) named_methods)) ^
blanchet
parents: 55178
diff changeset
   135
      "..."
58843
521cea5fa777 discontinued obsolete Output.urgent_message;
wenzelm
parents: 58842
diff changeset
   136
      |> writeln
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
   137
    else
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
   138
      ();
67225
cb34f5f49a08 have 'try0' display results faster
blanchet
parents: 67149
diff changeset
   139
    (case get_results apply_methods of
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
   140
      [] =>
63690
48a2c88091d7 tuning punctuation in messages output by Isabelle
blanchet
parents: 63518
diff changeset
   141
      (if mode = Normal then writeln "No proof found" else (); (false, (noneN, [])))
54291
709a2bbd7638 by (auto ...)[1] not by (auto [1])
blanchet
parents: 54248
diff changeset
   142
    | xs as (name, command, _) :: _ =>
41038
9592334001d5 quiet Metis in "try"
blanchet
parents: 40931
diff changeset
   143
      let
54291
709a2bbd7638 by (auto ...)[1] not by (auto [1])
blanchet
parents: 54248
diff changeset
   144
        val xs = xs |> map (fn (name, _, n) => (n, name))
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   145
                    |> AList.coalesce (op =)
63690
48a2c88091d7 tuning punctuation in messages output by Isabelle
blanchet
parents: 63518
diff changeset
   146
                    |> map (swap o apsnd commas);
41038
9592334001d5 quiet Metis in "try"
blanchet
parents: 40931
diff changeset
   147
        val message =
43026
0f15575a6465 handle non-auto try cases gracefully in Try Methods
blanchet
parents: 43024
diff changeset
   148
          (case mode of
52970
3e0fe71f3ce1 tuned messages
blanchet
parents: 52717
diff changeset
   149
             Auto_Try => "Auto Try0 found a proof"
3e0fe71f3ce1 tuned messages
blanchet
parents: 52717
diff changeset
   150
           | Try => "Try0 found a proof"
43031
e437d47f419f more concise output
blanchet
parents: 43026
diff changeset
   151
           | Normal => "Try this") ^ ": " ^
63518
ae8fd6fe63a1 tuned signature;
wenzelm
parents: 62984
diff changeset
   152
          Active.sendback_markup_command
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59184
diff changeset
   153
              ((if Thm.nprems_of (#goal (Proof.goal st)) = 1 then "by"
54291
709a2bbd7638 by (auto ...)[1] not by (auto [1])
blanchet
parents: 54248
diff changeset
   154
                else "apply") ^ " " ^ command) ^
54248
c7af3d651658 make 'try0' return faster when invoked as part of 'try'
blanchet
parents: 53052
diff changeset
   155
          (case xs of
63690
48a2c88091d7 tuning punctuation in messages output by Isabelle
blanchet
parents: 63518
diff changeset
   156
            [(_, ms)] => " (" ^ time_string ms ^ ")"
48a2c88091d7 tuning punctuation in messages output by Isabelle
blanchet
parents: 63518
diff changeset
   157
          | xs => "\n(" ^ space_implode "; " (map tool_time_string xs) ^ ")");
41038
9592334001d5 quiet Metis in "try"
blanchet
parents: 40931
diff changeset
   158
      in
59184
830bb7ddb3ab explicit message channels for "state", "information";
wenzelm
parents: 59083
diff changeset
   159
        (true, (name, if mode = Auto_Try then [message] else (writeln message; [])))
54291
709a2bbd7638 by (auto ...)[1] not by (auto [1])
blanchet
parents: 54248
diff changeset
   160
      end)
55179
blanchet
parents: 55178
diff changeset
   161
  end;
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   162
55181
blanchet
parents: 55179
diff changeset
   163
fun try0 timeout_opt = fst oo generic_try0 Normal timeout_opt;
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   164
46641
8801a24f9e9a renamed 'try_methods' to 'try0'
blanchet
parents: 45666
diff changeset
   165
fun try0_trans quad =
60190
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60094
diff changeset
   166
  Toplevel.keep_proof
906de96ba68a allow diagnostic proof commands with skip_proofs;
wenzelm
parents: 60094
diff changeset
   167
    (ignore o generic_try0 Normal (SOME default_timeout) quad o Toplevel.proof_of);
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   168
55179
blanchet
parents: 55178
diff changeset
   169
fun merge_attrs (s1, i1, e1, d1) (s2, i2, e2, d2) = (s1 @ s2, i1 @ i2, e1 @ e2, d1 @ d2);
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   170
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   171
fun string_of_xthm (xref, args) =
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   172
  Facts.string_of_ref xref ^
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67225
diff changeset
   173
  implode (map (enclose "[" "]" o Pretty.unformatted_string_of o Token.pretty_src \<^context>) args);
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   174
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   175
val parse_fact_refs =
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62519
diff changeset
   176
  Scan.repeat1 (Scan.unless (Parse.name -- Args.colon) (Parse.thm >> string_of_xthm));
55179
blanchet
parents: 55178
diff changeset
   177
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   178
val parse_attr =
55179
blanchet
parents: 55178
diff changeset
   179
  Args.$$$ "simp" |-- Args.colon |-- parse_fact_refs >> (fn ss => (ss, [], [], []))
blanchet
parents: 55178
diff changeset
   180
  || Args.$$$ "intro" |-- Args.colon |-- parse_fact_refs >> (fn is => ([], is, [], []))
blanchet
parents: 55178
diff changeset
   181
  || Args.$$$ "elim" |-- Args.colon |-- parse_fact_refs >> (fn es => ([], [], es, []))
blanchet
parents: 55178
diff changeset
   182
  || Args.$$$ "dest" |-- Args.colon |-- parse_fact_refs >> (fn ds => ([], [], [], ds));
blanchet
parents: 55178
diff changeset
   183
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   184
fun parse_attrs x =
55179
blanchet
parents: 55178
diff changeset
   185
  (Args.parens parse_attrs
blanchet
parents: 55178
diff changeset
   186
   || Scan.repeat parse_attr >> (fn quad => fold merge_attrs quad ([], [], [], []))) x;
41999
3c029ef9e0f2 added "simp:", "intro:", and "elim:" to "try" command
blanchet
parents: 41038
diff changeset
   187
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   188
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 63961
diff changeset
   189
  Outer_Syntax.command \<^command_keyword>\<open>try0\<close> "try a combination of proof methods"
55179
blanchet
parents: 55178
diff changeset
   190
    (Scan.optional parse_attrs ([], [], [], []) #>> try0_trans);
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   191
74508
3315c551fe6e clarified signature;
wenzelm
parents: 73386
diff changeset
   192
val _ =
3315c551fe6e clarified signature;
wenzelm
parents: 73386
diff changeset
   193
  Try.tool_setup
3315c551fe6e clarified signature;
wenzelm
parents: 73386
diff changeset
   194
   {name = "try0", weight = 30, auto_option = \<^system_option>\<open>auto_methods\<close>,
3315c551fe6e clarified signature;
wenzelm
parents: 73386
diff changeset
   195
    body = fn auto => generic_try0 (if auto then Auto_Try else Try) NONE ([], [], [], [])};
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   196
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   197
end;