src/HOL/Tools/atp_manager.ML
changeset 28571 47d88239658d
parent 28543 637f2808ab64
child 28582 c269a3045fdf
--- a/src/HOL/Tools/atp_manager.ML	Mon Oct 13 13:56:54 2008 +0200
+++ b/src/HOL/Tools/atp_manager.ML	Mon Oct 13 14:04:28 2008 +0200
@@ -2,27 +2,25 @@
     ID:         $Id$
     Author:     Fabian Immler, TU Muenchen
 
-ATP threads have to be registered here.  Threads with the same
-birth-time are seen as one group.  All threads of a group are killed
-when one thread of it has been successful, or after a certain time, or
-when the maximum number of threads exceeds; then the oldest thread is
-killed.
+ATP threads are registered here.
+Threads with the same birth-time are seen as one group.
+All threads of a group are killed when one thread of it has been successful,
+or after a certain time,
+or when the maximum number of threads exceeds; then the oldest thread is killed.
 *)
 
 signature ATP_MANAGER =
 sig
   val kill_all: unit -> unit
-  val info: unit -> unit
+  val info: unit -> string
   val set_atps: string -> unit
   val set_max_atp: int -> unit
   val set_timeout: int -> unit
-  val set_groupkilling: bool -> unit
-  val start: unit -> unit
-  val register: Time.time -> Time.time -> (Thread.thread * string) -> unit
-  val unregister: bool -> unit
-  val add_prover: string -> (Proof.state -> Thread.thread * string) -> theory -> theory
+    
+  val atp_thread: (unit -> 'a option) -> ('a -> string) -> Thread.thread
+  val add_prover: string -> (int -> Proof.state -> Thread.thread) -> theory -> theory
   val print_provers: theory -> unit
-  val sledgehammer: Proof.state -> unit
+  val sledgehammer: string list -> Proof.state -> unit
 end;
 
 structure AtpManager : ATP_MANAGER =
@@ -32,171 +30,189 @@
   (
     type elem = Time.time * Thread.thread;
     fun ord ((a, _), (b, _)) = Time.compare (a, b);
-  );
+  )
 
-  (* create global state of threadmanager *)
-  val timeout_heap = ref ThreadHeap.empty
-  val oldest_heap = ref ThreadHeap.empty
-  (* managed threads *)
-  val active = ref ([] : (Thread.thread * Time.time * Time.time * string) list)
-  val cancelling = ref ([] : (Thread.thread * Time.time * Time.time * string) list)
-  (* settings *)
-  val atps = ref "e,spass"
-  val maximum_atps = ref 5   (* ~1 means infinite number of atps*)
-  val timeout = ref 60
-  val groupkilling = ref true
+  (* state of threadmanager *)
+  datatype T = State of {
+      timeout_heap: ThreadHeap.T,
+      oldest_heap: ThreadHeap.T,
+      (* managed threads *)
+      active: (Thread.thread * (Time.time * Time.time * string)) list,
+      cancelling: (Thread.thread * (Time.time * Time.time * string)) list
+  }
+  val state = ref (State {
+        timeout_heap = ThreadHeap.empty,
+        oldest_heap = ThreadHeap.empty,
+        active = [],
+        cancelling = []})
+      
   (* synchronizing *)
-  val lock = Mutex.mutex () (* to be acquired for changing state *)
-  val state_change = ConditionVar.conditionVar () (* signal when state changes *)
+  val lock = Mutex.mutex ()
+  val state_change = ConditionVar.conditionVar ()
   (* watches over running threads and interrupts them if required *)
   val managing_thread = ref (NONE: Thread.thread option);
 
-  (* move a thread from active to cancelling
-    managing_thread trys to interrupt all threads in cancelling
-
-   call from an environment where a lock has already been aquired *)
-  fun unregister_locked thread =
-    let val entrys = (filter (fn (t,_,_,_) => Thread.equal (t, thread))) (! active)
-    val entrys_update = map (fn (th, tb, _, desc) => (th, tb, Time.now(), desc)) entrys
-    val _ = change cancelling (append entrys_update)
-    val _ = change active (filter_out (fn (t,_,_,_) => Thread.equal (t, thread)))
-    in () end;
+  (* active and cancelling are association lists *)
+  val lookup_th = AList.lookup Thread.equal;
+  val delete_th = AList.delete Thread.equal;
+  val update_th = AList.update Thread.equal;
+    
+    
+  (* waiting for a condition, checking it after time or when signalled
+     then changing state with returning a result *)
+  fun protected_wait_change_result P time f = uninterruptible (fn restore => fn () =>
+      let
+      val _ = Mutex.lock lock;
+      val _ = while not (P (! state)) do ConditionVar.waitUntil (state_change, lock, time (! state));
+      val res = Exn.Result (restore (fn () => change_result state f) ()) handle exn => Exn.Exn exn
+      val _ = ConditionVar.broadcast state_change;
+      val _ = Mutex.unlock lock;
+  in Exn.release res end) ()
+  
+  fun protected_change f = protected_wait_change_result (fn _ => true) (fn _ => Time.zeroTime) f;
+  
+  fun protected f = protected_change (`f);
 
-  (* start a watching thread which runs forever *)
-  (* must *not* be called more than once!! => problem with locks *)
-  fun start () =
+  (* unregister Thread from Threadmanager = move to cancelling *)
+  fun unregister success message thread = protected_change (fn State {timeout_heap, oldest_heap, active, cancelling} => 
     let
-    val new_thread = SimpleThread.fork false (fn () =>
+    val info = lookup_th active thread
+    (* get birthtime of unregistering thread if successful - for group-killing*)
+    val birthtime = case info of NONE => Time.zeroTime
+      | SOME (tb, _, _) => if success then tb else Time.zeroTime
+    (* move unregistering thread to cancelling *)
+    val active' = delete_th thread active
+    val cancelling' = case info of NONE => cancelling
+      | SOME (tb, _, desc) => update_th (thread, (tb, Time.now(), desc)) cancelling
+    (* move all threads of the same group to cancelling *)
+    val group_threads = map_filter (fn (th, (tb, _, desc)) =>
+        if tb = birthtime then SOME (th, (tb, Time.now(), desc)) else NONE)
+      active 
+    val active'' = filter_out (fn (_, (tb, _, _)) => tb = birthtime) active'
+    val cancelling'' = append group_threads cancelling'
+    (* message for user *)
+    val message = case info of NONE => ""
+      | SOME (_, _, desc) => "Sledgehammer: " ^ desc ^ "\n    " ^ message ^
+        (if length group_threads > 1
+        then "\n\nInterrupted " ^ Int.toString (length group_threads - 1) ^ " others of group."
+        else "")
+    in (message, State {timeout_heap = timeout_heap,
+                  oldest_heap = oldest_heap,
+                  active = active'',
+                  cancelling = cancelling''})
+    end)
+    
+  (* start a watching thread which runs forever - only one may exist => checked by register *)
+  fun start () = managing_thread := SOME (SimpleThread.fork false (fn () =>
+    let
+    val min_wait_time = Time.fromMilliseconds 300
+    val max_wait_time = Time.fromSeconds 10
+    (* action is required when there are cancelling threads, or a thread is to be cancelled *)
+    fun P (State {timeout_heap, oldest_heap, active, cancelling}) =
+      length cancelling > 0  orelse
+        (not (ThreadHeap.is_empty timeout_heap) andalso
+        Time.< (#1 (ThreadHeap.min timeout_heap), Time.now ()))
+    (* wait for next thread to cancel, or a maximum of 10 Seconds*)
+    fun time (State {timeout_heap, oldest_heap, active, cancelling}) =
+      if ThreadHeap.is_empty timeout_heap then Time.+ (Time.now(), max_wait_time)
+      else #1 (ThreadHeap.min timeout_heap)
+    (* action: cancel find threads whose timeout is reached, and interrupt cancelling threads *)
+    fun action (State {timeout_heap, oldest_heap, active, cancelling}) =
       let
-      (* never give up lock except for waiting *)
-      val _ = Mutex.lock lock
-      fun wait_for_next_event time =
-        let
-        (* wait for signal or next timeout, give up lock meanwhile *)
-        val _ = ConditionVar.waitUntil (state_change, lock, time)
-        (* move threads with priority less than Time.now() to cancelling *)
-        fun cancelolder heap =
-          if ThreadHeap.is_empty heap then heap else
+      (* find out threads wich reached their timeout *)
+      fun find_threads (heap, threads) =
+        if ThreadHeap.is_empty heap then (heap, threads) else
           let val (mintime, minthread) = ThreadHeap.min heap
-          in
-            if Time.> (mintime, Time.now()) then heap
-            else (unregister_locked minthread;
-            cancelolder (ThreadHeap.delete_min heap))
+          in if Time.> (mintime, Time.now())
+            then (heap, threads)
+            else find_threads (ThreadHeap.delete_min heap, minthread :: threads)
           end
-        val _ = change timeout_heap cancelolder
-        val _ = change cancelling (filter (fn (t,_,_,_) => Thread.isActive t))
-        val _ = map (fn (t, _, _, _) => SimpleThread.interrupt t) (! cancelling)
-        (* if there are threads to cancel, send periodic interrupts *)
-        (* TODO: find out what realtime-values are appropriate *)
-        val next_time =
-          if length (! cancelling) > 0 then
-            Time.+ (Time.now(), Time.fromMilliseconds 300)
-          else if ThreadHeap.is_empty (! timeout_heap) then
-            Time.+ (Time.now(), Time.fromSeconds 10)
-          else
-            #1 (ThreadHeap.min (! timeout_heap))
-          in
-            wait_for_next_event next_time
-          end
-        in wait_for_next_event Time.zeroTime end)
-      in managing_thread := SOME new_thread end
+      val (timeout_heap', timeout_threads) = find_threads (timeout_heap, [])
+      val _ = List.app (SimpleThread.interrupt o fst) cancelling
+      val cancelling' = filter (fn (t, _) => Thread.isActive t) cancelling
+      (* return threads with timeout and changed state *)
+      in (timeout_threads, State {timeout_heap = timeout_heap',
+                    oldest_heap = oldest_heap,
+                    active = active,
+                    cancelling = cancelling'})
+      end
+    in while true do 
+        (* cancel threads found by 'action' *)
+        (map (priority o (unregister false "Interrupted (reached timeout)")) (protected_wait_change_result P time action);
+        (* give threads time to respond to interrupt *)
+        OS.Process.sleep min_wait_time)
+    end))
 
-  (* calling thread registers itself to be managed here with a relative timeout *)
-  fun register birthtime deadtime (thread, name) =
-    let
-    val _ = Mutex.lock lock
-    (* create the atp-managing-thread if this is the first call to register *)
-    val _ =
+  (* calling thread is registered here by sledgehammer *)
+  fun register birthtime deadtime (thread, desc) = protected_change (
+    fn State {timeout_heap, oldest_heap, active, cancelling} => 
+    let val _ = (* create the atp-managing-thread if this is the first call to register *)
       if (case ! managing_thread of SOME thread => Thread.isActive thread | NONE => false)
       then () else start ()
     (* insertion *)
-    val _ = change timeout_heap (ThreadHeap.insert (deadtime, thread))
-    val _ = change oldest_heap (ThreadHeap.insert (birthtime, thread))
-    val _ = change active (cons (thread, birthtime, deadtime, name))
-    (*maximum number of atps must not exceed*)
-    val _ = let
-      fun kill_oldest () =
-        let val (_, oldest_thread) = ThreadHeap.min (!oldest_heap)
-        val _ = change oldest_heap ThreadHeap.delete_min
-        in unregister_locked oldest_thread end
-      in
-        while ! maximum_atps > ~1 andalso length (! active) > ! maximum_atps
-        do kill_oldest ()
-      end
-    (* state of threadmanager changed => signal *)
-    val _ = ConditionVar.signal state_change
-    val _ = Mutex.unlock lock
-    in () end
+    in ((),State {timeout_heap = ThreadHeap.insert (deadtime, thread) timeout_heap,
+                  oldest_heap = ThreadHeap.insert (birthtime, thread) oldest_heap,
+                  active = update_th (thread, (birthtime, deadtime, desc)) active,
+                  cancelling = cancelling})
+    end)
 
-  (* calling Thread unregisters itself from Threadmanager; thread is responsible
-    to terminate after calling this method *)
-  fun unregister success =
-    let val _ = Mutex.lock lock
-    val thread = Thread.self ()
-    (* get birthtime of unregistering thread - for group-killing*)
-    fun get_birthtime [] = Time.zeroTime
-      | get_birthtime ((t,tb,td,desc)::actives) = if Thread.equal (thread, t)
-      then tb
-      else get_birthtime actives
-    val birthtime = get_birthtime (! active)
-    (* remove unregistering thread *)
-    val _ = change active (filter_out (fn (t,_,_,_) => Thread.equal (t, thread)))
-    val _ = if (! groupkilling) andalso success
-      then (* remove all threads of the same group *)
-      let
-      val group_threads = filter (fn (_, tb, _, _) => tb = birthtime) (! active)
-      val _ = change cancelling (append group_threads)
-      val _ = change active (filter_out (fn (_, tb, _, _) => tb = birthtime))
-      in () end
-      else ()
-    val _ = ConditionVar.signal state_change
-    val _ = Mutex.unlock lock
-    in () end;
 
   (* Move all threads to cancelling *)
-  fun kill_all () =
+  fun kill_all () = protected_change (fn State {timeout_heap, oldest_heap, active, cancelling} =>
     let
-    val _ = Mutex.lock lock
-    val _ = change active (map (fn (th, tb, _, desc) => (th, tb, Time.now(), desc)))
-    val _ = change cancelling (append (! active))
-    val _ = active := []
-    val _ = ConditionVar.signal state_change
-    val _ = Mutex.unlock lock
-    in () end;
-
-  fun info () =
+    val active' = (map (fn (th, (tb, _, desc)) => (th, (tb, Time.now(), desc))) active)
+    in ((),State {timeout_heap = timeout_heap,
+                  oldest_heap = oldest_heap,
+                  active = [],
+                  cancelling = append active' cancelling})
+    end)
+    
+  (* give information on running threads *)
+  fun info () = protected (fn State {timeout_heap, oldest_heap, active, cancelling} =>
     let
-    val _ = Mutex.lock lock
-    fun running_info (_, birth_time, dead_time, desc) =
-      priority ("Running: "
+    fun running_info (_, (birth_time, dead_time, desc)) = "Running: "
         ^ ((Int.toString o Time.toSeconds) (Time.- (Time.now(), birth_time)))
         ^ " s  --  "
         ^ ((Int.toString o Time.toSeconds) (Time.- (dead_time, Time.now())))
-        ^ " s to live:\n" ^ desc)
-    fun cancelling_info (_, _, dead_time, desc) =
-      priority ("Trying to interrupt thread since "
+        ^ " s to live:\n" ^ desc
+    fun cancelling_info (_, (_, dead_time, desc)) = "Trying to interrupt thread since "
         ^ (Int.toString o Time.toSeconds) (Time.- (Time.now(), dead_time))
-        ^ " s:\n" ^ desc )
-    val _ = if length (! active) = 0 then [priority "No ATPs running."]
-      else (priority "--- RUNNING ATPs ---";
-      map (fn entry => running_info entry) (! active))
-    val _ = if length (! cancelling) = 0 then []
-      else (priority "--- TRYING TO INTERRUPT FOLLOWING ATPs ---";
-      map (fn entry => cancelling_info entry) (! cancelling))
-    val _ = Mutex.unlock lock
-    in () end;
+        ^ " s:\n" ^ desc
+    val running = if null active then "No ATPs running."
+      else String.concatWith "\n\n" ("--- RUNNING ATPs ---" ::
+      (map (fn entry => running_info entry) active))
+    val interrupting = if null cancelling then ""
+      else String.concatWith "\n\n" ("--- TRYING TO INTERRUPT FOLLOWING ATPs ---" ::
+      (map (fn entry => cancelling_info entry) cancelling))
+    in running ^ "\n" ^ interrupting end)
 
+    
+    (* thread wrapping an atp-call *)   
+    
+    fun atp_thread call_prover produce_answer =
+      SimpleThread.fork true (fn () =>
+        let
+        val result = call_prover ()
+        val message = case result of NONE => "Failed."
+            | SOME result => "Try this command: " ^ produce_answer result
+        in priority (unregister (isSome result) message (Thread.self()))
+        end handle Interrupt => ()
+      )
 
+      
     (* preferences *)
 
-    fun set_max_atp number = CRITICAL (fn () => maximum_atps := number);
+    val atps = ref "e spass"
+    val maximum_atps = ref 5   (* ~1 means infinite number of atps*)
+    val timeout = ref 60
+    val groupkilling = ref true
     fun set_atps str = CRITICAL (fn () => atps := str);
+    fun set_max_atp number = CRITICAL (fn () => maximum_atps := number);
     fun set_timeout time = CRITICAL (fn () => timeout := time);
-    fun set_groupkilling boolean = CRITICAL (fn () => groupkilling := boolean);
 
     val _ = ProofGeneralPgip.add_preference "Proof"
         {name = "ATP - Provers (see print_atps)",
-         descr = "Which external automatic provers (seperated by commas)",
+         descr = "Default automatic provers (seperated by whitespace)",
          default = !atps,
          pgiptype = PgipTypes.Pgipstring,
          get = fn () => !atps,
@@ -228,7 +244,7 @@
 
   structure Provers = TheoryDataFun
   (
-    type T = ((Proof.state -> Thread.thread * string) * stamp) Symtab.table
+    type T = ((int -> Proof.state -> Thread.thread) * stamp) Symtab.table
     val empty = Symtab.empty
     val copy = I
     val extend = I
@@ -243,21 +259,47 @@
   fun print_provers thy = Pretty.writeln
     (Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy))));
 
-  fun run_prover state name =
+  fun prover_desc state subgoal name =
+     let val (ctxt, (chain_ths, goal)) = Proof.get_goal state
+     in "External prover " ^ quote name ^ " for Subgoal " ^ Int.toString subgoal^ ":\n" ^ Syntax.string_of_term ctxt (List.nth(prems_of goal, subgoal-1)) end   
+    
+  fun run_prover state subgoal name =
     (case Symtab.lookup (Provers.get (Proof.theory_of state)) name of
       NONE => (warning ("Unknown external prover: " ^ quote name); NONE)
-    | SOME (prover_fn, _) => SOME (prover_fn state));
-
-
-  (* sledghammer *)
+    | SOME (prover_fn, _) => SOME (prover_fn subgoal state, prover_desc state subgoal name));
 
-  fun sledgehammer state =
+  fun kill_one () = 
     let
-      val proverids = String.tokens (fn c => c = #",") (! atps)
-      val threads_names = map_filter (run_prover state) proverids
-      val birthtime = Time.now()
-      val deadtime = Time.+ (Time.now(), Time.fromSeconds (! timeout))
-      val _ = List.app (register birthtime deadtime) threads_names
+    val oldest = protected_change (fn (s as State {timeout_heap , oldest_heap, active, cancelling}) =>
+      if ThreadHeap.is_empty oldest_heap orelse length active <= !maximum_atps orelse ! maximum_atps = ~1 then (NONE, s)
+      else
+      let val (_, oldest_thread) = ThreadHeap.min (oldest_heap)
+      in (SOME oldest_thread, State {timeout_heap = timeout_heap,
+        oldest_heap = ThreadHeap.delete_min oldest_heap,
+        active = active,
+        cancelling = cancelling}) end)
+    in case oldest of NONE => ()
+      | SOME thread => priority (unregister false "Interrupted (Maximum number of ATPs exceeded)." thread) end
+      
+  fun kill_oldest (State {timeout_heap, oldest_heap, active, cancelling}) =
+    if length active > !maximum_atps andalso !maximum_atps > ~1
+    then let val _ = kill_one () in kill_oldest (! state) end
+    else ()
+      
+    
+  (* sledghammer for first subgoal *)
+
+  fun sledgehammer names proof_state =
+    let
+    val proverids = case names of
+      [] => String.tokens (Symbol.is_ascii_blank o String.str) (! atps)
+      | names => names
+    val threads_names = map_filter (run_prover proof_state 1) proverids
+    val birthtime = Time.now()
+    val deadtime = Time.+ (Time.now(), Time.fromSeconds (! timeout))
+    val _ = List.app (register birthtime deadtime) threads_names
+      (*maximum number of atps must not exceed*)
+    val _ = kill_oldest (! state)
     in () end
 
 
@@ -271,7 +313,7 @@
 
   val _ =
     OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag
-      (Scan.succeed (Toplevel.no_timing o Toplevel.imperative info));
+      (Scan.succeed (Toplevel.no_timing o Toplevel.imperative (priority o info)));
 
   val _ =
     OuterSyntax.improper_command "print_atps" "print external provers" K.diag
@@ -280,8 +322,8 @@
 
   val _ =
     OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag
-      (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_proof o
-        Toplevel.keep (sledgehammer o Toplevel.proof_of)));
+      (Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o
+      Toplevel.keep ((sledgehammer names) o Toplevel.proof_of)));
 
   end;