src/HOL/Tools/try.ML
author boehmes
Mon, 08 Nov 2010 12:13:44 +0100
changeset 40424 7550b2cba1cb
parent 40301 bf39a257b3d3
child 40931 061b8257ab9f
permissions -rw-r--r--
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/try.ML
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
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     7
signature TRY =
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
     8
sig
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
     9
  val auto : bool Unsynchronized.ref
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    10
  val invoke_try : Time.time option -> Proof.state -> bool
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    11
  val setup : theory -> theory
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    12
end;
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    13
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    14
structure Try : TRY =
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    15
struct
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    16
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    17
val auto = Unsynchronized.ref false
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    18
39333
c277c79fb9db add Proof General option
blanchet
parents: 39332
diff changeset
    19
val _ =
c277c79fb9db add Proof General option
blanchet
parents: 39332
diff changeset
    20
  ProofGeneralPgip.add_preference Preferences.category_tracing
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
    21
      (Preferences.bool_pref auto "auto-try" "Try standard proof methods.")
39333
c277c79fb9db add Proof General option
blanchet
parents: 39332
diff changeset
    22
40301
bf39a257b3d3 simplified some time constants;
wenzelm
parents: 40222
diff changeset
    23
val default_timeout = seconds 5.0
38944
827c98e8ba8b fiddling with "try"
blanchet
parents: 38942
diff changeset
    24
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    25
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
    26
  let val {goal, ...} = Proof.goal st in
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    27
    case (case timeout_opt of
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    28
            SOME timeout => TimeLimit.timeLimit timeout
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    29
          | NONE => fn f => fn x => f x) (Seq.pull o tac) (pre st) of
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    30
      SOME (x, _) => nprems_of (post x) < nprems_of goal
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    31
    | NONE => false
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    32
  end
40113
1f61f0826e8a handle timeouts (to prevent failure from other threads);
blanchet
parents: 39719
diff changeset
    33
  handle TimeLimit.TimeOut => false
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    34
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    35
fun do_generic timeout_opt command pre post apply st =
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    36
  let val timer = Timer.startRealTimer () in
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    37
    if can_apply timeout_opt pre post apply st then
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    38
      SOME (command, Time.toMilliseconds (Timer.checkRealTimer timer))
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    39
    else
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    40
      NONE
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    41
  end
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    42
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    43
fun named_method thy name =
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    44
  Method.method thy (Args.src ((name, []), Position.none))
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    45
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    46
fun apply_named_method_on_first_goal name ctxt =
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    47
  let val thy = ProofContext.theory_of ctxt in
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    48
    Proof.refine (Method.SelectGoals (1, Method.Basic (named_method thy name)))
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    49
  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
    50
  handle ERROR _ => K Seq.empty (* e.g., the method isn't available yet *)
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    51
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
    52
fun do_named_method (name, (all_goals, run_if_auto)) auto timeout_opt st =
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
    53
  if not auto orelse run_if_auto then
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
    54
    do_generic timeout_opt
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
    55
               (name ^ (if all_goals andalso
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
    56
                           nprems_of (#goal (Proof.goal st)) > 1 then
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
    57
                          "[1]"
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
    58
                        else
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
    59
                          "")) I (#goal o Proof.goal)
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
    60
               (apply_named_method_on_first_goal name (Proof.context_of st)) st
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
    61
  else
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
    62
    NONE
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    63
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
    64
(* name * (all_goals, run_if_auto) *)
39547
5df45da44bfb reorder proof methods and take out "best";
blanchet
parents: 39336
diff changeset
    65
val named_methods =
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
    66
  [("simp", (false, true)),
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
    67
   ("auto", (true, true)),
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
    68
   ("fast", (false, false)),
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
    69
   ("fastsimp", (false, false)),
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
    70
   ("force", (false, false)),
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
    71
   ("blast", (false, true)),
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
    72
   ("metis", (false, true)),
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
   ("linarith", (false, true)),
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
    74
   ("presburger", (false, true))]
39547
5df45da44bfb reorder proof methods and take out "best";
blanchet
parents: 39336
diff changeset
    75
val do_methods = map do_named_method named_methods
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    76
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    77
fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    78
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
    79
fun do_try auto timeout_opt st =
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
    80
  case do_methods |> Par_List.map (fn f => f auto timeout_opt st)
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    81
                  |> map_filter I |> sort (int_ord o pairself snd) of
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    82
    [] => (if auto then () else writeln "No proof found."; (false, st))
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
    83
  | xs as (s, _) :: _ =>
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    84
    let
39332
538b94dc62de make Try's output more concise
blanchet
parents: 39331
diff changeset
    85
      val xs = xs |> map swap |> AList.coalesce (op =)
538b94dc62de make Try's output more concise
blanchet
parents: 39331
diff changeset
    86
                  |> map (swap o apsnd commas)
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    87
      val message =
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    88
        (if auto then "Auto Try found a proof" else "Try this command") ^ ": " ^
38946
da5e4f182f69 distinguish between "by" and "apply"
blanchet
parents: 38944
diff changeset
    89
        Markup.markup Markup.sendback
da5e4f182f69 distinguish between "by" and "apply"
blanchet
parents: 38944
diff changeset
    90
            ((if nprems_of (#goal (Proof.goal st)) = 1 then "by" else "apply") ^
da5e4f182f69 distinguish between "by" and "apply"
blanchet
parents: 38944
diff changeset
    91
             " " ^ s) ^
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
    92
        "\n(" ^ space_implode "; " (map time_string xs) ^ ").\n"
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    93
    in
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    94
      (true, st |> (if auto then
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    95
                      Proof.goal_message
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    96
                          (fn () => Pretty.chunks [Pretty.str "",
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    97
                                    Pretty.markup Markup.hilite
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    98
                                                  [Pretty.str message]])
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
    99
                    else
40132
7ee65dbffa31 renamed Output.priority to Output.urgent_message to emphasize its special role more clearly;
wenzelm
parents: 40113
diff changeset
   100
                      tap (fn _ => Output.urgent_message message)))
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   101
    end
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   102
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
   103
val invoke_try = fst oo do_try false
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   104
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   105
val tryN = "try"
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   106
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   107
val _ =
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   108
  Outer_Syntax.improper_command tryN
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   109
      "try a combination of proof methods" Keyword.diag
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
   110
      (Scan.succeed (Toplevel.keep (K () o do_try false (SOME default_timeout)
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
   111
                                    o Toplevel.proof_of)))
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   112
39336
1899349a5026 change signature of "Try.invoke_try" to make it more flexible
blanchet
parents: 39334
diff changeset
   113
fun auto_try st = if not (!auto) then (false, st) else do_try true NONE st
39331
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   114
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   115
val setup = Auto_Tools.register_tool (tryN, auto_try)
8b1969d603c0 added Auto Try to the mix of automatic tools
blanchet
parents: 39222
diff changeset
   116
38942
e10c11971fa7 "try" -- a new diagnosis tool that tries to apply several methods in parallel
blanchet
parents:
diff changeset
   117
end;