src/HOL/Mirabelle/Actions/mirabelle_arith.ML
changeset 48348 3fabf352243e
parent 32567 de411627a985
equal deleted inserted replaced
48347:92d1c566ebbf 48348:3fabf352243e
       
     1 (*  Title:      HOL/Mirabelle/Actions/mirabelle_arith.ML
       
     2     Author:     Jasmin Blanchette and Sascha Boehme, TU Munich
       
     3 *)
       
     4 
       
     5 structure Mirabelle_Arith : MIRABELLE_ACTION =
       
     6 struct
       
     7 
       
     8 fun arith_tag id = "#" ^ string_of_int id ^ " arith: "
       
     9 
       
    10 fun init _ = I
       
    11 fun done _ _ = ()
       
    12 
       
    13 fun run id ({pre, timeout, log, ...}: Mirabelle.run_args) =
       
    14   if Mirabelle.can_apply timeout Arith_Data.arith_tac pre
       
    15   then log (arith_tag id ^ "succeeded")
       
    16   else ()
       
    17   handle TimeLimit.TimeOut => log (arith_tag id ^ "timeout")
       
    18 
       
    19 fun invoke _ = Mirabelle.register (init, Mirabelle.catch arith_tag run, done)
       
    20 
       
    21 end