src/HOL/Tools/SMT2/smt2_solver.ML
changeset 56090 34bd10a9a2ad
parent 56087 2cd8fcb4804d
child 56094 2adbc6e4cd8f
--- a/src/HOL/Tools/SMT2/smt2_solver.ML	Thu Mar 13 13:18:13 2014 +0100
+++ b/src/HOL/Tools/SMT2/smt2_solver.ML	Thu Mar 13 13:18:14 2014 +0100
@@ -10,7 +10,7 @@
   datatype outcome = Unsat | Sat | Unknown
   type solver_config = {
     name: string,
-    class: Proof.context -> SMT2_Utils.class,
+    class: Proof.context -> SMT2_Util.class,
     avail: unit -> bool,
     command: unit -> string list,
     options: Proof.context -> string list,
@@ -142,7 +142,7 @@
 
 type solver_config = {
   name: string,
-  class: Proof.context -> SMT2_Utils.class,
+  class: Proof.context -> SMT2_Util.class,
   avail: unit -> bool,
   command: unit -> string list,
   options: Proof.context -> string list,