src/HOL/Tools/ATP_Manager/atp_manager.ML
author wenzelm
Tue, 04 Aug 2009 19:20:24 +0200
changeset 32327 0971cc0b6a57
parent 31793 src/HOL/Tools/atp_manager.ML@7c10b13d49fe
child 32451 8f0dc876fb1b
permissions -rw-r--r--
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
     1 (*  Title:      HOL/Tools/ATP_Manager/atp_manager.ML
     2     Author:     Fabian Immler, TU Muenchen
     3 
     4 ATP threads are registered here.
     5 Threads with the same birth-time are seen as one group.
     6 All threads of a group are killed when one thread of it has been successful,
     7 or after a certain time,
     8 or when the maximum number of threads exceeds; then the oldest thread is killed.
     9 *)
    10 
    11 signature ATP_MANAGER =
    12 sig
    13   val get_atps: unit -> string
    14   val set_atps: string -> unit
    15   val get_max_atps: unit -> int
    16   val set_max_atps: int -> unit
    17   val get_timeout: unit -> int
    18   val set_timeout: int -> unit
    19   val get_full_types: unit -> bool
    20   val set_full_types: bool -> unit
    21   val kill: unit -> unit
    22   val info: unit -> unit
    23   val messages: int option -> unit
    24   type prover = int -> (thm * (string * int)) list option ->
    25     (thm * (string * int)) list option -> string -> int ->
    26     Proof.context * (thm list * thm) ->
    27     bool * string * string * string vector * (thm * (string * int)) list
    28   val add_prover: string -> prover -> theory -> theory
    29   val print_provers: theory -> unit
    30   val get_prover: string -> theory -> prover option
    31   val sledgehammer: string list -> Proof.state -> unit
    32 end;
    33 
    34 structure AtpManager: ATP_MANAGER =
    35 struct
    36 
    37 (** preferences **)
    38 
    39 val message_store_limit = 20;
    40 val message_display_limit = 5;
    41 
    42 local
    43 
    44 val atps = ref "e remote_vampire";
    45 val max_atps = ref 5;   (* ~1 means infinite number of atps *)
    46 val timeout = ref 60;
    47 val full_types = ref false;
    48 
    49 in
    50 
    51 fun get_atps () = CRITICAL (fn () => ! atps);
    52 fun set_atps str = CRITICAL (fn () => atps := str);
    53 
    54 fun get_max_atps () = CRITICAL (fn () => ! max_atps);
    55 fun set_max_atps number = CRITICAL (fn () => max_atps := number);
    56 
    57 fun get_timeout () = CRITICAL (fn () => ! timeout);
    58 fun set_timeout time = CRITICAL (fn () => timeout := time);
    59 
    60 fun get_full_types () = CRITICAL (fn () => ! full_types);
    61 fun set_full_types bool = CRITICAL (fn () => full_types := bool);
    62 
    63 val _ =
    64   ProofGeneralPgip.add_preference Preferences.category_proof
    65     (Preferences.string_pref atps
    66       "ATP: provers" "Default automatic provers (separated by whitespace)");
    67 
    68 val _ =
    69   ProofGeneralPgip.add_preference Preferences.category_proof
    70     (Preferences.int_pref max_atps
    71       "ATP: maximum number" "How many provers may run in parallel");
    72 
    73 val _ =
    74   ProofGeneralPgip.add_preference Preferences.category_proof
    75     (Preferences.int_pref timeout
    76       "ATP: timeout" "ATPs will be interrupted after this time (in seconds)");
    77 
    78 val _ =
    79   ProofGeneralPgip.add_preference Preferences.category_proof
    80     (Preferences.bool_pref full_types
    81       "ATP: full types" "ATPs will use full type information");
    82 
    83 end;
    84 
    85 
    86 
    87 (** thread management **)
    88 
    89 (* data structures over threads *)
    90 
    91 structure ThreadHeap = HeapFun
    92 (
    93   type elem = Time.time * Thread.thread;
    94   fun ord ((a, _), (b, _)) = Time.compare (a, b);
    95 );
    96 
    97 fun lookup_thread xs = AList.lookup Thread.equal xs;
    98 fun delete_thread xs = AList.delete Thread.equal xs;
    99 fun update_thread xs = AList.update Thread.equal xs;
   100 
   101 
   102 (* state of thread manager *)
   103 
   104 datatype T = State of
   105  {managing_thread: Thread.thread option,
   106   timeout_heap: ThreadHeap.T,
   107   oldest_heap: ThreadHeap.T,
   108   active: (Thread.thread * (Time.time * Time.time * string)) list,
   109   cancelling: (Thread.thread * (Time.time * Time.time * string)) list,
   110   messages: string list,
   111   store: string list};
   112 
   113 fun make_state managing_thread timeout_heap oldest_heap active cancelling messages store =
   114   State {managing_thread = managing_thread, timeout_heap = timeout_heap, oldest_heap = oldest_heap,
   115     active = active, cancelling = cancelling, messages = messages, store = store};
   116 
   117 val state = Synchronized.var "atp_manager"
   118   (make_state NONE ThreadHeap.empty ThreadHeap.empty [] [] [] []);
   119 
   120 
   121 (* unregister thread *)
   122 
   123 fun unregister (success, message) thread = Synchronized.change state
   124   (fn state as State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   125     (case lookup_thread active thread of
   126       SOME (birthtime, _, description) =>
   127         let
   128           val (group, active') =
   129             if success then List.partition (fn (_, (tb, _, _)) => tb = birthtime) active
   130             else List.partition (fn (th, _) => Thread.equal (th, thread)) active
   131 
   132           val now = Time.now ()
   133           val cancelling' =
   134             fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) group cancelling
   135 
   136           val message' = description ^ "\n" ^ message ^
   137             (if length group <= 1 then ""
   138              else "\nInterrupted " ^ string_of_int (length group - 1) ^ " other group members")
   139           val store' = message' ::
   140             (if length store <= message_store_limit then store
   141              else #1 (chop message_store_limit store))
   142         in make_state
   143           managing_thread timeout_heap oldest_heap active' cancelling' (message' :: messages) store'
   144         end
   145     | NONE => state));
   146 
   147 
   148 (* kill excessive atp threads *)
   149 
   150 fun excessive_atps active =
   151   let val max = get_max_atps ()
   152   in length active > max andalso max > ~1 end;
   153 
   154 local
   155 
   156 fun kill_oldest () =
   157   let exception Unchanged in
   158     Synchronized.change_result state
   159       (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   160         if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active)
   161         then raise Unchanged
   162         else
   163           let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap
   164           in (oldest_thread,
   165           make_state managing_thread timeout_heap oldest_heap' active cancelling messages store) end)
   166       |> unregister (false, "Interrupted (maximum number of ATPs exceeded)")
   167     handle Unchanged => ()
   168   end;
   169 
   170 in
   171 
   172 fun kill_excessive () =
   173   let val State {active, ...} = Synchronized.value state
   174   in if excessive_atps active then (kill_oldest (); kill_excessive ()) else () end;
   175 
   176 end;
   177 
   178 fun print_new_messages () =
   179   let val to_print = Synchronized.change_result state
   180     (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   181       (messages, make_state managing_thread timeout_heap oldest_heap active cancelling [] store))
   182   in
   183     if null to_print then ()
   184     else priority ("Sledgehammer: " ^ space_implode "\n\n" to_print)
   185   end;
   186 
   187 
   188 (* start a watching thread -- only one may exist *)
   189 
   190 fun check_thread_manager () = Synchronized.change state
   191   (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   192     if (case managing_thread of SOME thread => Thread.isActive thread | NONE => false)
   193     then make_state managing_thread timeout_heap oldest_heap active cancelling messages store
   194     else let val managing_thread = SOME (SimpleThread.fork false (fn () =>
   195       let
   196         val min_wait_time = Time.fromMilliseconds 300
   197         val max_wait_time = Time.fromSeconds 10
   198 
   199         (* wait for next thread to cancel, or maximum*)
   200         fun time_limit (State {timeout_heap, ...}) =
   201           (case try ThreadHeap.min timeout_heap of
   202             NONE => SOME (Time.+ (Time.now (), max_wait_time))
   203           | SOME (time, _) => SOME time)
   204 
   205         (* action: find threads whose timeout is reached, and interrupt cancelling threads *)
   206         fun action (State {managing_thread, timeout_heap, oldest_heap, active, cancelling,
   207                            messages, store}) =
   208           let val (timeout_threads, timeout_heap') =
   209             ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap
   210           in
   211             if null timeout_threads andalso null cancelling andalso not (excessive_atps active)
   212             then NONE
   213             else
   214               let
   215                 val _ = List.app (SimpleThread.interrupt o #1) cancelling
   216                 val cancelling' = filter (Thread.isActive o #1) cancelling
   217                 val state' = make_state
   218                   managing_thread timeout_heap' oldest_heap active cancelling' messages store
   219               in SOME (map #2 timeout_threads, state') end
   220           end
   221       in
   222         while Synchronized.change_result state
   223           (fn st as
   224             State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   225             if (null active) andalso (null cancelling) andalso (null messages)
   226             then (false, make_state NONE timeout_heap oldest_heap active cancelling messages store)
   227             else (true, st))
   228         do
   229           (Synchronized.timed_access state time_limit action
   230             |> these
   231             |> List.app (unregister (false, "Interrupted (reached timeout)"));
   232             kill_excessive ();
   233             print_new_messages ();
   234             (*give threads time to respond to interrupt*)
   235             OS.Process.sleep min_wait_time)
   236       end))
   237     in make_state managing_thread timeout_heap oldest_heap active cancelling messages store end);
   238 
   239 
   240 (* thread is registered here by sledgehammer *)
   241 
   242 fun register birthtime deadtime (thread, desc) =
   243  (Synchronized.change state
   244     (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   245       let
   246         val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap
   247         val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap
   248         val active' = update_thread (thread, (birthtime, deadtime, desc)) active
   249       in make_state managing_thread timeout_heap' oldest_heap' active' cancelling messages store end);
   250   check_thread_manager ());
   251 
   252 
   253 
   254 (** user commands **)
   255 
   256 (* kill: move all threads to cancelling *)
   257 
   258 fun kill () = Synchronized.change state
   259   (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   260     let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active
   261     in make_state
   262       managing_thread timeout_heap oldest_heap [] (formerly_active @ cancelling) messages store
   263     end);
   264 
   265 
   266 (* ATP info *)
   267 
   268 fun info () =
   269   let
   270     val State {active, cancelling, ...} = Synchronized.value state
   271 
   272     fun running_info (_, (birth_time, dead_time, desc)) = "Running: "
   273         ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), birth_time))
   274         ^ " s  --  "
   275         ^ (string_of_int o Time.toSeconds) (Time.- (dead_time, Time.now ()))
   276         ^ " s to live:\n" ^ desc
   277     fun cancelling_info (_, (_, dead_time, desc)) = "Trying to interrupt thread since "
   278         ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), dead_time))
   279         ^ " s:\n" ^ desc
   280 
   281     val running =
   282       if null active then "No ATPs running."
   283       else space_implode "\n\n" ("Running ATPs:" :: map running_info active)
   284     val interrupting =
   285       if null cancelling then ""
   286       else space_implode "\n\n"
   287         ("Trying to interrupt the following ATPs:" :: map cancelling_info cancelling)
   288 
   289   in writeln (running ^ "\n" ^ interrupting) end;
   290 
   291 fun messages opt_limit =
   292   let
   293     val limit = the_default message_display_limit opt_limit;
   294     val State {store = msgs, ...} = Synchronized.value state
   295     val header = "Recent ATP messages" ^
   296       (if length msgs <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
   297   in writeln (space_implode "\n\n" (header :: #1 (chop limit msgs))) end;
   298 
   299 
   300 
   301 (** The Sledgehammer **)
   302 
   303 (* named provers *)
   304 
   305 type prover = int -> (thm * (string * int)) list option ->
   306   (thm * (string * int)) list option -> string -> int ->
   307   Proof.context * (thm list * thm) ->
   308   bool * string * string * string vector * (thm * (string * int)) list
   309 
   310 fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
   311 
   312 structure Provers = TheoryDataFun
   313 (
   314   type T = (prover * stamp) Symtab.table
   315   val empty = Symtab.empty
   316   val copy = I
   317   val extend = I
   318   fun merge _ tabs : T = Symtab.merge (eq_snd op =) tabs
   319     handle Symtab.DUP dup => err_dup_prover dup
   320 );
   321 
   322 fun add_prover name prover thy =
   323   Provers.map (Symtab.update_new (name, (prover, stamp ()))) thy
   324     handle Symtab.DUP dup => err_dup_prover dup;
   325 
   326 fun print_provers thy = Pretty.writeln
   327   (Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy))));
   328 
   329 fun get_prover name thy = case Symtab.lookup (Provers.get thy) name of
   330   NONE => NONE
   331 | SOME (prover, _) => SOME prover;
   332 
   333 (* start prover thread *)
   334 
   335 fun start_prover name birthtime deadtime i proof_state =
   336   (case get_prover name (Proof.theory_of proof_state) of
   337     NONE => warning ("Unknown external prover: " ^ quote name)
   338   | SOME prover =>
   339       let
   340         val (ctxt, (_, goal)) = Proof.get_goal proof_state
   341         val desc =
   342           "external prover " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^
   343             Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))
   344         val _ = SimpleThread.fork true (fn () =>
   345           let
   346             val _ = register birthtime deadtime (Thread.self (), desc)
   347             val result =
   348               let val (success, message, _, _, _) =
   349                 prover (get_timeout ()) NONE NONE name i (Proof.get_goal proof_state)
   350               in (success, message) end
   351               handle ResHolClause.TOO_TRIVIAL
   352                 => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
   353               | ERROR msg
   354                 => (false, "Error: " ^ msg)
   355             val _ = unregister result (Thread.self ())
   356           in () end handle Interrupt => ())
   357       in () end);
   358 
   359 
   360 (* sledghammer for first subgoal *)
   361 
   362 fun sledgehammer names proof_state =
   363   let
   364     val provers =
   365       if null names then String.tokens (Symbol.is_ascii_blank o String.str) (get_atps ())
   366       else names
   367     val birthtime = Time.now ()
   368     val deadtime = Time.+ (birthtime, Time.fromSeconds (get_timeout ()))
   369   in List.app (fn name => start_prover name birthtime deadtime 1 proof_state) provers end;
   370 
   371 
   372 
   373 (** Isar command syntax **)
   374 
   375 local structure K = OuterKeyword and P = OuterParse in
   376 
   377 val _ =
   378   OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag
   379     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
   380 
   381 val _ =
   382   OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag
   383     (Scan.succeed (Toplevel.no_timing o Toplevel.imperative info));
   384 
   385 val _ =
   386   OuterSyntax.improper_command "atp_messages" "print recent messages issued by managed provers" K.diag
   387     (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
   388       (fn limit => Toplevel.no_timing o Toplevel.imperative (fn () => messages limit)));
   389 
   390 val _ =
   391   OuterSyntax.improper_command "print_atps" "print external provers" K.diag
   392     (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o
   393       Toplevel.keep (print_provers o Toplevel.theory_of)));
   394 
   395 val _ =
   396   OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag
   397     (Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o
   398       Toplevel.keep (sledgehammer names o Toplevel.proof_of)));
   399 
   400 end;
   401 
   402 end;
   403