src/HOL/Tools/SMT2/smt2_solver.ML
changeset 57159 24cbdebba35a
parent 57158 f028d93798e6
child 57163 7fc7de3b387e
--- a/src/HOL/Tools/SMT2/smt2_solver.ML	Mon Jun 02 17:34:26 2014 +0200
+++ b/src/HOL/Tools/SMT2/smt2_solver.ML	Mon Jun 02 17:34:27 2014 +0200
@@ -26,10 +26,10 @@
   val default_max_relevant: Proof.context -> string -> int
 
   (*filter*)
-  val smt2_filter: Proof.context -> thm -> ('a * (int option * thm)) list -> int -> Time.time ->
-    {outcome: SMT2_Failure.failure option, rewrite_rules: thm list, conjecture_id: int,
-     prem_ids: int list, helper_ids: (int * thm) list, fact_ids: (int * ('a * thm)) list,
-     z3_proof: Z3_New_Proof.z3_step list}
+  val smt2_filter: Proof.context -> thm -> ((string * 'a) * (int option * thm)) list -> int ->
+    Time.time ->
+    {outcome: SMT2_Failure.failure option, fact_ids: (int * ((string * 'a) * thm)) list,
+     atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
 
   (*tactic*)
   val smt2_tac: Proof.context -> thm list -> int -> tactic
@@ -273,20 +273,28 @@
     |> (fn ((iidths0, z3_proof), {rewrite_rules, ...}) =>
       let
         val iidths = if can_filter ctxt then iidths0 else map (apsnd (apfst (K no_id))) iwthms
+
         fun id_of_index i = the_default no_id (Option.map fst (AList.lookup (op =) iidths i))
+
+        val conjecture_id = id_of_index conjecture_i
+        val prem_ids = map id_of_index (prems_i upto facts_i - 1)
+        val helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths
+        val fact_ids = map_filter (fn (i, (id, _)) =>
+          try (apsnd (apsnd snd o nth xwfacts)) (id, i - facts_i)) iidths
+        val fact_helper_ts =
+          map (fn (_, th) => (ATP_Util.short_thm_name ctxt th, prop_of th)) helper_ids @
+          map (fn (_, ((s, _), th)) => (s, prop_of th)) fact_ids
+        val fact_helper_ids =
+          map (fn (id, th) => (id, ATP_Util.short_thm_name ctxt th)) helper_ids @
+          map (fn (id, ((name, _), _)) => (id, name)) fact_ids
       in
-        {outcome = NONE,
-         rewrite_rules = rewrite_rules,
-         conjecture_id = id_of_index conjecture_i,
-         prem_ids = map id_of_index (prems_i upto facts_i - 1),
-         helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths,
-         fact_ids = map_filter (fn (i, (id, _)) =>
-           try (apsnd (apsnd snd o nth xwfacts)) (id, i - facts_i)) iidths,
-         z3_proof = z3_proof}
+        {outcome = NONE, fact_ids = fact_ids,
+         atp_proof = fn () => Z3_New_Isar.atp_proof_of_z3_proof ctxt rewrite_rules
+           (map Thm.prop_of prems) (Thm.term_of concl) fact_helper_ts prem_ids conjecture_id
+           fact_helper_ids z3_proof}
       end)
   end
-  handle SMT2_Failure.SMT fail => {outcome = SOME fail, rewrite_rules = [], conjecture_id = no_id,
-    prem_ids = [], helper_ids = [], fact_ids = [], z3_proof = []}
+  handle SMT2_Failure.SMT fail => {outcome = SOME fail, fact_ids = [], atp_proof = K []}
 
 
 (* SMT tactic *)