Mirabelle: actions are responsible for their log messages, output is better readable
authorboehmes
Tue, 01 Sep 2009 15:20:21 +0200
changeset 32472 7b92a8b8daaf
parent 32470 8ca2f3500dbc
child 32473 6341f907aba4
Mirabelle: actions are responsible for their log messages, output is better readable
src/HOL/Tools/Mirabelle/Tools/mirabelle.ML
src/HOL/Tools/Mirabelle/Tools/mirabelle_arith.ML
src/HOL/Tools/Mirabelle/Tools/mirabelle_metis.ML
src/HOL/Tools/Mirabelle/Tools/mirabelle_sledgehammer.ML
src/HOL/Tools/Mirabelle/lib/scripts/mirabelle.pl
--- a/src/HOL/Tools/Mirabelle/Tools/mirabelle.ML	Tue Sep 01 14:10:38 2009 +0200
+++ b/src/HOL/Tools/Mirabelle/Tools/mirabelle.ML	Tue Sep 01 15:20:21 2009 +0200
@@ -38,8 +38,8 @@
 
 (* Mirabelle core *)
 
-type action = Time.time -> {pre: Proof.state, post: Toplevel.state option} ->
-  string option
+type action = {pre: Proof.state, post: Toplevel.state option,
+  timeout: Time.time, log: string -> unit} -> unit
 
 structure Actions = TheoryDataFun
 (
@@ -67,33 +67,30 @@
   (* FIXME: with multithreading and parallel proofs enabled, we might need to
      encapsulate this inside a critical section *)
 
-fun capture_exns f x =
+fun log_block thy msg = log thy (msg ^ "\n------------------")
+fun log_action thy name msg = log_block thy (name ^ ": " ^ msg)
+
+fun capture_exns logf f x =
   let
     fun f' x = f x
-      handle TimeLimit.TimeOut => SOME "time out"
-           | ERROR msg => SOME ("error: " ^ msg)
-  in (case try f' x of NONE => SOME "exception" | SOME msg => msg) end
+      handle TimeLimit.TimeOut => logf "time out"
+           | ERROR msg => logf ("error: " ^ msg)
+  in (case try f' x of NONE => logf "exception" | _ => ()) end
 
-fun apply_action timeout st (name, action) =
-  Option.map (pair name) (capture_exns (action timeout) st)
+fun apply_actions thy info (pre, post, time) actions =
+  let
+    val _ = log_block thy info
+    fun apply (name, action) =
+      let val st = {pre=pre, post=post, timeout=time, log=log_action thy name}
+      in capture_exns (log_action thy name) action st end
+  in List.app apply actions end
 
 fun in_range _ _ NONE = true
   | in_range l r (SOME i) = (l <= i andalso (r < 0 orelse i <= r))
 
 fun only_within_range thy pos f x =
   let val l = Config.get_thy thy start_line and r = Config.get_thy thy end_line
-  in if in_range l r (Position.line_of pos) then f x else [] end
-
-fun pretty_print pos name msgs =
-  let
-    val str0 = string_of_int o the_default 0
-    val loc = str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos)
-    val head = "at " ^ loc ^ " (" ^ name ^ "):"
-
-    fun pretty_msg (name, msg) = Pretty.block (map Pretty.str [name, ": ", msg])
-  in
-    Pretty.string_of (Pretty.big_list head (map pretty_msg msgs))
-  end
+  in if in_range l r (Position.line_of pos) then f x else () end
 
 in
 
@@ -102,13 +99,15 @@
     val thy = Proof.theory_of pre
     val pos = Toplevel.pos_of tr
     val name = Toplevel.name_of tr
-    val timeout = Time.fromSeconds (Config.get_thy thy timeout)
-    val st = {pre=pre, post=post}
+    val st = (pre, post, Time.fromSeconds (Config.get_thy thy timeout))
+
+    val str0 = string_of_int o the_default 0
+    val loc = str0 (Position.line_of pos) ^ ":" ^ str0 (Position.column_of pos)
+    val info = "\n\nat " ^ loc ^ " (" ^ name ^ "):"
   in
     Actions.get thy
     |> Symtab.dest
-    |> only_within_range thy pos (map_filter (apply_action timeout st))
-    |> (fn [] => () | msgs => log thy (pretty_print pos name msgs))
+    |> only_within_range thy pos (apply_actions thy info st)
   end
 
 end
@@ -128,12 +127,12 @@
 
 val goal_thm_of = snd o snd o Proof.get_goal
 
-fun can_apply timeout tac st =
+fun can_apply time tac st =
   let
     val (ctxt, (facts, goal)) = Proof.get_goal st
     val full_tac = HEADGOAL (Method.insert_tac facts THEN' tac ctxt)
   in
-    (case TimeLimit.timeLimit timeout (Seq.pull o full_tac) goal of
+    (case TimeLimit.timeLimit time (Seq.pull o full_tac) goal of
       SOME (thm, _) => true
     | NONE => false)
   end
--- a/src/HOL/Tools/Mirabelle/Tools/mirabelle_arith.ML	Tue Sep 01 14:10:38 2009 +0200
+++ b/src/HOL/Tools/Mirabelle/Tools/mirabelle_arith.ML	Tue Sep 01 15:20:21 2009 +0200
@@ -5,10 +5,10 @@
 structure Mirabelle_Arith : MIRABELLE_ACTION =
 struct
 
-fun arith_action timeout {pre=st, ...} =
-  if Mirabelle.can_apply timeout Arith_Data.arith_tac st
-  then SOME "succeeded"
-  else NONE
+fun arith_action {pre, timeout, log, ...} =
+  if Mirabelle.can_apply timeout Arith_Data.arith_tac pre
+  then log "succeeded"
+  else ()
 
 fun invoke _ = Mirabelle.register ("arith", arith_action)
 
--- a/src/HOL/Tools/Mirabelle/Tools/mirabelle_metis.ML	Tue Sep 01 14:10:38 2009 +0200
+++ b/src/HOL/Tools/Mirabelle/Tools/mirabelle_metis.ML	Tue Sep 01 15:20:21 2009 +0200
@@ -5,18 +5,19 @@
 structure Mirabelle_Metis : MIRABELLE_ACTION =
 struct
 
-fun metis_action timeout {pre, post} =
+fun metis_action {pre, post, timeout, log} =
   let
     val thms = Mirabelle.theorems_of_sucessful_proof post
     val names = map Thm.get_name thms
+    val add_info = if null names then I else suffix (":\n" ^ commas names)
 
     val facts = Facts.props (ProofContext.facts_of (Proof.context_of pre))
 
     fun metis ctxt = MetisTools.metis_tac ctxt (thms @ facts)
   in
     (if Mirabelle.can_apply timeout metis pre then "succeeded" else "failed")
-    |> suffix (" (" ^ commas names ^ ")")
-    |> SOME
+    |> add_info
+    |> log
   end
 
 fun invoke _ = Mirabelle.register ("metis", metis_action)
--- a/src/HOL/Tools/Mirabelle/Tools/mirabelle_sledgehammer.ML	Tue Sep 01 14:10:38 2009 +0200
+++ b/src/HOL/Tools/Mirabelle/Tools/mirabelle_sledgehammer.ML	Tue Sep 01 15:20:21 2009 +0200
@@ -29,7 +29,7 @@
 val keepK = "keep"
 val metisK = "metis"
 
-fun sledgehammer_action args timeout {pre=st, ...} =
+fun sledgehammer_action args {pre=st, timeout, log, ...} =
   let
     val prover_name =
       AList.lookup (op =) args proverK
@@ -56,25 +56,28 @@
           TimeLimit.timeLimit timeout atp (Proof.get_goal st)
       in (success, message, SOME thm_names) end
       handle ResHolClause.TOO_TRIVIAL => (true, "trivial", SOME [])
-           | ERROR msg => (false, "error: " ^ msg, NONE)
     val (success, message, thm_names) = safe init done sh
       (AList.lookup (op =) args keepK)
+    val _ =
+      if success then log (quote prover_name ^ " succeeded:\n" ^ message)
+      else log (prover_name ^ " failed")
 
     (* try metis *)
     fun get_thms ctxt = maps (thms_of_name ctxt)
     fun metis thms ctxt = MetisTools.metis_tac ctxt thms
-    fun apply_metis thms = "\nApplying metis with these theorems: " ^
-     (if Mirabelle.can_apply timeout (metis thms) st
-      then "success"
-      else "failure")
-    val msg = if not (AList.defined (op =) args metisK) then ""
-      else (apply_metis (get_thms (Proof.context_of st) (these thm_names))
-        handle ERROR m => "\nException: " ^ m)
-  in
-    if success
-    then SOME ("success (" ^ prover_name ^ ": " ^ message ^ ")" ^ msg)
-    else NONE
-  end
+    fun log_metis s =
+      log ("applying metis: " ^ s)
+    fun apply_metis thms =
+      if Mirabelle.can_apply timeout (metis thms) st
+      then "succeeded" else "failed"
+    val _ =
+      if not (AList.defined (op =) args metisK) then ()
+      else
+        these thm_names
+        |> get_thms (Proof.context_of st)
+        |> apply_metis
+        |> log_metis
+  in () end
 
 fun invoke args = Mirabelle.register ("sledgehammer", sledgehammer_action args)
 
--- a/src/HOL/Tools/Mirabelle/lib/scripts/mirabelle.pl	Tue Sep 01 14:10:38 2009 +0200
+++ b/src/HOL/Tools/Mirabelle/lib/scripts/mirabelle.pl	Tue Sep 01 15:20:21 2009 +0200
@@ -120,7 +120,6 @@
 foreach $name (@action_names) {
   print LOG_FILE "  $name\n";
 }
-print LOG_FILE "\n\n";
 close(LOG_FILE);
 
 my $r = system "\"$ENV{'ISABELLE_PROCESS'}\" " .