src/HOL/Tools/atp_manager.ML
changeset 32327 0971cc0b6a57
parent 32326 9d70ecf11b7a
child 32328 f2fd9da84bac
child 32329 1527ff8c2dfb
     1.1 --- a/src/HOL/Tools/atp_manager.ML	Tue Aug 04 16:13:16 2009 +0200
     1.2 +++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.3 @@ -1,403 +0,0 @@
     1.4 -(*  Title:      HOL/Tools/atp_manager.ML
     1.5 -    Author:     Fabian Immler, TU Muenchen
     1.6 -
     1.7 -ATP threads are registered here.
     1.8 -Threads with the same birth-time are seen as one group.
     1.9 -All threads of a group are killed when one thread of it has been successful,
    1.10 -or after a certain time,
    1.11 -or when the maximum number of threads exceeds; then the oldest thread is killed.
    1.12 -*)
    1.13 -
    1.14 -signature ATP_MANAGER =
    1.15 -sig
    1.16 -  val get_atps: unit -> string
    1.17 -  val set_atps: string -> unit
    1.18 -  val get_max_atps: unit -> int
    1.19 -  val set_max_atps: int -> unit
    1.20 -  val get_timeout: unit -> int
    1.21 -  val set_timeout: int -> unit
    1.22 -  val get_full_types: unit -> bool
    1.23 -  val set_full_types: bool -> unit
    1.24 -  val kill: unit -> unit
    1.25 -  val info: unit -> unit
    1.26 -  val messages: int option -> unit
    1.27 -  type prover = int -> (thm * (string * int)) list option ->
    1.28 -    (thm * (string * int)) list option -> string -> int ->
    1.29 -    Proof.context * (thm list * thm) ->
    1.30 -    bool * string * string * string vector * (thm * (string * int)) list
    1.31 -  val add_prover: string -> prover -> theory -> theory
    1.32 -  val print_provers: theory -> unit
    1.33 -  val get_prover: string -> theory -> prover option
    1.34 -  val sledgehammer: string list -> Proof.state -> unit
    1.35 -end;
    1.36 -
    1.37 -structure AtpManager: ATP_MANAGER =
    1.38 -struct
    1.39 -
    1.40 -(** preferences **)
    1.41 -
    1.42 -val message_store_limit = 20;
    1.43 -val message_display_limit = 5;
    1.44 -
    1.45 -local
    1.46 -
    1.47 -val atps = ref "e remote_vampire";
    1.48 -val max_atps = ref 5;   (* ~1 means infinite number of atps *)
    1.49 -val timeout = ref 60;
    1.50 -val full_types = ref false;
    1.51 -
    1.52 -in
    1.53 -
    1.54 -fun get_atps () = CRITICAL (fn () => ! atps);
    1.55 -fun set_atps str = CRITICAL (fn () => atps := str);
    1.56 -
    1.57 -fun get_max_atps () = CRITICAL (fn () => ! max_atps);
    1.58 -fun set_max_atps number = CRITICAL (fn () => max_atps := number);
    1.59 -
    1.60 -fun get_timeout () = CRITICAL (fn () => ! timeout);
    1.61 -fun set_timeout time = CRITICAL (fn () => timeout := time);
    1.62 -
    1.63 -fun get_full_types () = CRITICAL (fn () => ! full_types);
    1.64 -fun set_full_types bool = CRITICAL (fn () => full_types := bool);
    1.65 -
    1.66 -val _ =
    1.67 -  ProofGeneralPgip.add_preference Preferences.category_proof
    1.68 -    (Preferences.string_pref atps
    1.69 -      "ATP: provers" "Default automatic provers (separated by whitespace)");
    1.70 -
    1.71 -val _ =
    1.72 -  ProofGeneralPgip.add_preference Preferences.category_proof
    1.73 -    (Preferences.int_pref max_atps
    1.74 -      "ATP: maximum number" "How many provers may run in parallel");
    1.75 -
    1.76 -val _ =
    1.77 -  ProofGeneralPgip.add_preference Preferences.category_proof
    1.78 -    (Preferences.int_pref timeout
    1.79 -      "ATP: timeout" "ATPs will be interrupted after this time (in seconds)");
    1.80 -
    1.81 -val _ =
    1.82 -  ProofGeneralPgip.add_preference Preferences.category_proof
    1.83 -    (Preferences.bool_pref full_types
    1.84 -      "ATP: full types" "ATPs will use full type information");
    1.85 -
    1.86 -end;
    1.87 -
    1.88 -
    1.89 -
    1.90 -(** thread management **)
    1.91 -
    1.92 -(* data structures over threads *)
    1.93 -
    1.94 -structure ThreadHeap = HeapFun
    1.95 -(
    1.96 -  type elem = Time.time * Thread.thread;
    1.97 -  fun ord ((a, _), (b, _)) = Time.compare (a, b);
    1.98 -);
    1.99 -
   1.100 -fun lookup_thread xs = AList.lookup Thread.equal xs;
   1.101 -fun delete_thread xs = AList.delete Thread.equal xs;
   1.102 -fun update_thread xs = AList.update Thread.equal xs;
   1.103 -
   1.104 -
   1.105 -(* state of thread manager *)
   1.106 -
   1.107 -datatype T = State of
   1.108 - {managing_thread: Thread.thread option,
   1.109 -  timeout_heap: ThreadHeap.T,
   1.110 -  oldest_heap: ThreadHeap.T,
   1.111 -  active: (Thread.thread * (Time.time * Time.time * string)) list,
   1.112 -  cancelling: (Thread.thread * (Time.time * Time.time * string)) list,
   1.113 -  messages: string list,
   1.114 -  store: string list};
   1.115 -
   1.116 -fun make_state managing_thread timeout_heap oldest_heap active cancelling messages store =
   1.117 -  State {managing_thread = managing_thread, timeout_heap = timeout_heap, oldest_heap = oldest_heap,
   1.118 -    active = active, cancelling = cancelling, messages = messages, store = store};
   1.119 -
   1.120 -val state = Synchronized.var "atp_manager"
   1.121 -  (make_state NONE ThreadHeap.empty ThreadHeap.empty [] [] [] []);
   1.122 -
   1.123 -
   1.124 -(* unregister thread *)
   1.125 -
   1.126 -fun unregister (success, message) thread = Synchronized.change state
   1.127 -  (fn state as State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.128 -    (case lookup_thread active thread of
   1.129 -      SOME (birthtime, _, description) =>
   1.130 -        let
   1.131 -          val (group, active') =
   1.132 -            if success then List.partition (fn (_, (tb, _, _)) => tb = birthtime) active
   1.133 -            else List.partition (fn (th, _) => Thread.equal (th, thread)) active
   1.134 -
   1.135 -          val now = Time.now ()
   1.136 -          val cancelling' =
   1.137 -            fold (fn (th, (tb, _, desc)) => update_thread (th, (tb, now, desc))) group cancelling
   1.138 -
   1.139 -          val message' = description ^ "\n" ^ message ^
   1.140 -            (if length group <= 1 then ""
   1.141 -             else "\nInterrupted " ^ string_of_int (length group - 1) ^ " other group members")
   1.142 -          val store' = message' ::
   1.143 -            (if length store <= message_store_limit then store
   1.144 -             else #1 (chop message_store_limit store))
   1.145 -        in make_state
   1.146 -          managing_thread timeout_heap oldest_heap active' cancelling' (message' :: messages) store'
   1.147 -        end
   1.148 -    | NONE => state));
   1.149 -
   1.150 -
   1.151 -(* kill excessive atp threads *)
   1.152 -
   1.153 -fun excessive_atps active =
   1.154 -  let val max = get_max_atps ()
   1.155 -  in length active > max andalso max > ~1 end;
   1.156 -
   1.157 -local
   1.158 -
   1.159 -fun kill_oldest () =
   1.160 -  let exception Unchanged in
   1.161 -    Synchronized.change_result state
   1.162 -      (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.163 -        if ThreadHeap.is_empty oldest_heap orelse not (excessive_atps active)
   1.164 -        then raise Unchanged
   1.165 -        else
   1.166 -          let val ((_, oldest_thread), oldest_heap') = ThreadHeap.min_elem oldest_heap
   1.167 -          in (oldest_thread,
   1.168 -          make_state managing_thread timeout_heap oldest_heap' active cancelling messages store) end)
   1.169 -      |> unregister (false, "Interrupted (maximum number of ATPs exceeded)")
   1.170 -    handle Unchanged => ()
   1.171 -  end;
   1.172 -
   1.173 -in
   1.174 -
   1.175 -fun kill_excessive () =
   1.176 -  let val State {active, ...} = Synchronized.value state
   1.177 -  in if excessive_atps active then (kill_oldest (); kill_excessive ()) else () end;
   1.178 -
   1.179 -end;
   1.180 -
   1.181 -fun print_new_messages () =
   1.182 -  let val to_print = Synchronized.change_result state
   1.183 -    (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.184 -      (messages, make_state managing_thread timeout_heap oldest_heap active cancelling [] store))
   1.185 -  in
   1.186 -    if null to_print then ()
   1.187 -    else priority ("Sledgehammer: " ^ space_implode "\n\n" to_print)
   1.188 -  end;
   1.189 -
   1.190 -
   1.191 -(* start a watching thread -- only one may exist *)
   1.192 -
   1.193 -fun check_thread_manager () = Synchronized.change state
   1.194 -  (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.195 -    if (case managing_thread of SOME thread => Thread.isActive thread | NONE => false)
   1.196 -    then make_state managing_thread timeout_heap oldest_heap active cancelling messages store
   1.197 -    else let val managing_thread = SOME (SimpleThread.fork false (fn () =>
   1.198 -      let
   1.199 -        val min_wait_time = Time.fromMilliseconds 300
   1.200 -        val max_wait_time = Time.fromSeconds 10
   1.201 -
   1.202 -        (* wait for next thread to cancel, or maximum*)
   1.203 -        fun time_limit (State {timeout_heap, ...}) =
   1.204 -          (case try ThreadHeap.min timeout_heap of
   1.205 -            NONE => SOME (Time.+ (Time.now (), max_wait_time))
   1.206 -          | SOME (time, _) => SOME time)
   1.207 -
   1.208 -        (* action: find threads whose timeout is reached, and interrupt cancelling threads *)
   1.209 -        fun action (State {managing_thread, timeout_heap, oldest_heap, active, cancelling,
   1.210 -                           messages, store}) =
   1.211 -          let val (timeout_threads, timeout_heap') =
   1.212 -            ThreadHeap.upto (Time.now (), Thread.self ()) timeout_heap
   1.213 -          in
   1.214 -            if null timeout_threads andalso null cancelling andalso not (excessive_atps active)
   1.215 -            then NONE
   1.216 -            else
   1.217 -              let
   1.218 -                val _ = List.app (SimpleThread.interrupt o #1) cancelling
   1.219 -                val cancelling' = filter (Thread.isActive o #1) cancelling
   1.220 -                val state' = make_state
   1.221 -                  managing_thread timeout_heap' oldest_heap active cancelling' messages store
   1.222 -              in SOME (map #2 timeout_threads, state') end
   1.223 -          end
   1.224 -      in
   1.225 -        while Synchronized.change_result state
   1.226 -          (fn st as
   1.227 -            State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.228 -            if (null active) andalso (null cancelling) andalso (null messages)
   1.229 -            then (false, make_state NONE timeout_heap oldest_heap active cancelling messages store)
   1.230 -            else (true, st))
   1.231 -        do
   1.232 -          (Synchronized.timed_access state time_limit action
   1.233 -            |> these
   1.234 -            |> List.app (unregister (false, "Interrupted (reached timeout)"));
   1.235 -            kill_excessive ();
   1.236 -            print_new_messages ();
   1.237 -            (*give threads time to respond to interrupt*)
   1.238 -            OS.Process.sleep min_wait_time)
   1.239 -      end))
   1.240 -    in make_state managing_thread timeout_heap oldest_heap active cancelling messages store end);
   1.241 -
   1.242 -
   1.243 -(* thread is registered here by sledgehammer *)
   1.244 -
   1.245 -fun register birthtime deadtime (thread, desc) =
   1.246 - (Synchronized.change state
   1.247 -    (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.248 -      let
   1.249 -        val timeout_heap' = ThreadHeap.insert (deadtime, thread) timeout_heap
   1.250 -        val oldest_heap' = ThreadHeap.insert (birthtime, thread) oldest_heap
   1.251 -        val active' = update_thread (thread, (birthtime, deadtime, desc)) active
   1.252 -      in make_state managing_thread timeout_heap' oldest_heap' active' cancelling messages store end);
   1.253 -  check_thread_manager ());
   1.254 -
   1.255 -
   1.256 -
   1.257 -(** user commands **)
   1.258 -
   1.259 -(* kill: move all threads to cancelling *)
   1.260 -
   1.261 -fun kill () = Synchronized.change state
   1.262 -  (fn State {managing_thread, timeout_heap, oldest_heap, active, cancelling, messages, store} =>
   1.263 -    let val formerly_active = map (fn (th, (tb, _, desc)) => (th, (tb, Time.now (), desc))) active
   1.264 -    in make_state
   1.265 -      managing_thread timeout_heap oldest_heap [] (formerly_active @ cancelling) messages store
   1.266 -    end);
   1.267 -
   1.268 -
   1.269 -(* ATP info *)
   1.270 -
   1.271 -fun info () =
   1.272 -  let
   1.273 -    val State {active, cancelling, ...} = Synchronized.value state
   1.274 -
   1.275 -    fun running_info (_, (birth_time, dead_time, desc)) = "Running: "
   1.276 -        ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), birth_time))
   1.277 -        ^ " s  --  "
   1.278 -        ^ (string_of_int o Time.toSeconds) (Time.- (dead_time, Time.now ()))
   1.279 -        ^ " s to live:\n" ^ desc
   1.280 -    fun cancelling_info (_, (_, dead_time, desc)) = "Trying to interrupt thread since "
   1.281 -        ^ (string_of_int o Time.toSeconds) (Time.- (Time.now (), dead_time))
   1.282 -        ^ " s:\n" ^ desc
   1.283 -
   1.284 -    val running =
   1.285 -      if null active then "No ATPs running."
   1.286 -      else space_implode "\n\n" ("Running ATPs:" :: map running_info active)
   1.287 -    val interrupting =
   1.288 -      if null cancelling then ""
   1.289 -      else space_implode "\n\n"
   1.290 -        ("Trying to interrupt the following ATPs:" :: map cancelling_info cancelling)
   1.291 -
   1.292 -  in writeln (running ^ "\n" ^ interrupting) end;
   1.293 -
   1.294 -fun messages opt_limit =
   1.295 -  let
   1.296 -    val limit = the_default message_display_limit opt_limit;
   1.297 -    val State {store = msgs, ...} = Synchronized.value state
   1.298 -    val header = "Recent ATP messages" ^
   1.299 -      (if length msgs <= limit then ":" else " (" ^ string_of_int limit ^ " displayed):");
   1.300 -  in writeln (space_implode "\n\n" (header :: #1 (chop limit msgs))) end;
   1.301 -
   1.302 -
   1.303 -
   1.304 -(** The Sledgehammer **)
   1.305 -
   1.306 -(* named provers *)
   1.307 -
   1.308 -type prover = int -> (thm * (string * int)) list option ->
   1.309 -  (thm * (string * int)) list option -> string -> int ->
   1.310 -  Proof.context * (thm list * thm) ->
   1.311 -  bool * string * string * string vector * (thm * (string * int)) list
   1.312 -
   1.313 -fun err_dup_prover name = error ("Duplicate prover: " ^ quote name);
   1.314 -
   1.315 -structure Provers = TheoryDataFun
   1.316 -(
   1.317 -  type T = (prover * stamp) Symtab.table
   1.318 -  val empty = Symtab.empty
   1.319 -  val copy = I
   1.320 -  val extend = I
   1.321 -  fun merge _ tabs : T = Symtab.merge (eq_snd op =) tabs
   1.322 -    handle Symtab.DUP dup => err_dup_prover dup
   1.323 -);
   1.324 -
   1.325 -fun add_prover name prover thy =
   1.326 -  Provers.map (Symtab.update_new (name, (prover, stamp ()))) thy
   1.327 -    handle Symtab.DUP dup => err_dup_prover dup;
   1.328 -
   1.329 -fun print_provers thy = Pretty.writeln
   1.330 -  (Pretty.strs ("external provers:" :: sort_strings (Symtab.keys (Provers.get thy))));
   1.331 -
   1.332 -fun get_prover name thy = case Symtab.lookup (Provers.get thy) name of
   1.333 -  NONE => NONE
   1.334 -| SOME (prover, _) => SOME prover;
   1.335 -
   1.336 -(* start prover thread *)
   1.337 -
   1.338 -fun start_prover name birthtime deadtime i proof_state =
   1.339 -  (case get_prover name (Proof.theory_of proof_state) of
   1.340 -    NONE => warning ("Unknown external prover: " ^ quote name)
   1.341 -  | SOME prover =>
   1.342 -      let
   1.343 -        val (ctxt, (_, goal)) = Proof.get_goal proof_state
   1.344 -        val desc =
   1.345 -          "external prover " ^ quote name ^ " for subgoal " ^ string_of_int i ^ ":\n" ^
   1.346 -            Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i))
   1.347 -        val _ = SimpleThread.fork true (fn () =>
   1.348 -          let
   1.349 -            val _ = register birthtime deadtime (Thread.self (), desc)
   1.350 -            val result =
   1.351 -              let val (success, message, _, _, _) =
   1.352 -                prover (get_timeout ()) NONE NONE name i (Proof.get_goal proof_state)
   1.353 -              in (success, message) end
   1.354 -              handle ResHolClause.TOO_TRIVIAL
   1.355 -                => (true, "Empty clause: Try this command: " ^ Markup.markup Markup.sendback "apply metis")
   1.356 -              | ERROR msg
   1.357 -                => (false, "Error: " ^ msg)
   1.358 -            val _ = unregister result (Thread.self ())
   1.359 -          in () end handle Interrupt => ())
   1.360 -      in () end);
   1.361 -
   1.362 -
   1.363 -(* sledghammer for first subgoal *)
   1.364 -
   1.365 -fun sledgehammer names proof_state =
   1.366 -  let
   1.367 -    val provers =
   1.368 -      if null names then String.tokens (Symbol.is_ascii_blank o String.str) (get_atps ())
   1.369 -      else names
   1.370 -    val birthtime = Time.now ()
   1.371 -    val deadtime = Time.+ (birthtime, Time.fromSeconds (get_timeout ()))
   1.372 -  in List.app (fn name => start_prover name birthtime deadtime 1 proof_state) provers end;
   1.373 -
   1.374 -
   1.375 -
   1.376 -(** Isar command syntax **)
   1.377 -
   1.378 -local structure K = OuterKeyword and P = OuterParse in
   1.379 -
   1.380 -val _ =
   1.381 -  OuterSyntax.improper_command "atp_kill" "kill all managed provers" K.diag
   1.382 -    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative kill));
   1.383 -
   1.384 -val _ =
   1.385 -  OuterSyntax.improper_command "atp_info" "print information about managed provers" K.diag
   1.386 -    (Scan.succeed (Toplevel.no_timing o Toplevel.imperative info));
   1.387 -
   1.388 -val _ =
   1.389 -  OuterSyntax.improper_command "atp_messages" "print recent messages issued by managed provers" K.diag
   1.390 -    (Scan.option (P.$$$ "(" |-- P.nat --| P.$$$ ")") >>
   1.391 -      (fn limit => Toplevel.no_timing o Toplevel.imperative (fn () => messages limit)));
   1.392 -
   1.393 -val _ =
   1.394 -  OuterSyntax.improper_command "print_atps" "print external provers" K.diag
   1.395 -    (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory o
   1.396 -      Toplevel.keep (print_provers o Toplevel.theory_of)));
   1.397 -
   1.398 -val _ =
   1.399 -  OuterSyntax.command "sledgehammer" "call all automatic theorem provers" K.diag
   1.400 -    (Scan.repeat P.xname >> (fn names => Toplevel.no_timing o Toplevel.unknown_proof o
   1.401 -      Toplevel.keep (sledgehammer names o Toplevel.proof_of)));
   1.402 -
   1.403 -end;
   1.404 -
   1.405 -end;
   1.406 -