src/HOL/Mirabelle/Tools/mirabelle_arith.ML
author boehmes
Sat, 05 Sep 2009 11:45:57 +0200
changeset 32521 f20cc66b2c74
parent 32518 e3c4e337196c
child 32564 378528d2f7eb
permissions -rw-r--r--
added initialization and cleanup of actions, added option to suppress Isabelle output, sledgehammer action produces its own report (no need for additional perl script)

(* Title:  mirabelle_arith.ML
   Author: Jasmin Blanchette and Sascha Boehme
*)

structure Mirabelle_Arith : MIRABELLE_ACTION =
struct

fun arith_tag id = "#" ^ string_of_int id ^ " arith: "

fun init _ = I
fun done _ _ = ()

fun run id {pre, timeout, log, ...} =
  if Mirabelle.can_apply timeout Arith_Data.arith_tac pre
  then log (arith_tag id ^ "succeeded")
  else ()
  handle TimeLimit.TimeOut => log (arith_tag id ^ "timeout")

fun invoke _ = Mirabelle.register (init, Mirabelle.catch arith_tag run, done)

end