src/HOL/Tools/Mirabelle/mirabelle_arith.ML
author desharna
Fri, 17 Dec 2021 09:51:37 +0100
changeset 74948 15ce207f69c8
parent 74515 64c0d78d2f19
child 75003 f21e7e6172a0
permissions -rw-r--r--
added support for initialization messages to Mirabelle

(*  Title:      HOL/Mirabelle/Tools/mirabelle_arith.ML
    Author:     Jasmin Blanchette, TU Munich
    Author:     Sascha Boehme, TU Munich
    Author:     Makarius
    Author:     Martin Desharnais, UniBw Munich

Mirabelle action: "arith".
*)

structure Mirabelle_Arith: MIRABELLE_ACTION =
struct

fun make_action ({timeout, ...} : Mirabelle.action_context) =
  let
    fun run_action ({pre, ...} : Mirabelle.command) =
      (case Timing.timing (Mirabelle.can_apply timeout Arith_Data.arith_tac) pre of
        ({cpu, ...}, true) => "succeeded (" ^ Time.toString cpu ^ " s)"
      | (_, false) => "failed")
  in ("", {run_action = run_action, finalize = K ""}) end

val () = Mirabelle.register_action "arith" make_action

end