src/HOL/Tools/SMT2/smt2_solver.ML
changeset 58061 3d060f43accb
parent 58060 835b5443b978
child 58062 f4d8987656b9
equal deleted inserted replaced
58060:835b5443b978 58061:3d060f43accb
     1 (*  Title:      HOL/Tools/SMT2/smt2_solver.ML
       
     2     Author:     Sascha Boehme, TU Muenchen
       
     3 
       
     4 SMT solvers registry and SMT tactic.
       
     5 *)
       
     6 
       
     7 signature SMT2_SOLVER =
       
     8 sig
       
     9   (*configuration*)
       
    10   datatype outcome = Unsat | Sat | Unknown
       
    11 
       
    12   type parsed_proof =
       
    13     {outcome: SMT2_Failure.failure option,
       
    14      fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list,
       
    15      atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
       
    16 
       
    17   type solver_config =
       
    18     {name: string,
       
    19      class: Proof.context -> SMT2_Util.class,
       
    20      avail: unit -> bool,
       
    21      command: unit -> string list,
       
    22      options: Proof.context -> string list,
       
    23      smt_options: (string * string) list,
       
    24      default_max_relevant: int,
       
    25      outcome: string -> string list -> outcome * string list,
       
    26      parse_proof: (Proof.context -> SMT2_Translate.replay_data ->
       
    27        ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
       
    28        parsed_proof) option,
       
    29      replay: (Proof.context -> SMT2_Translate.replay_data -> string list -> thm) option}
       
    30 
       
    31   (*registry*)
       
    32   val add_solver: solver_config -> theory -> theory
       
    33   val default_max_relevant: Proof.context -> string -> int
       
    34 
       
    35   (*filter*)
       
    36   val smt2_filter: Proof.context -> thm -> ((string * ATP_Problem_Generate.stature) * thm) list ->
       
    37     int -> Time.time -> parsed_proof
       
    38 
       
    39   (*tactic*)
       
    40   val smt2_tac: Proof.context -> thm list -> int -> tactic
       
    41   val smt2_tac': Proof.context -> thm list -> int -> tactic
       
    42 end;
       
    43 
       
    44 structure SMT2_Solver: SMT2_SOLVER =
       
    45 struct
       
    46 
       
    47 (* interface to external solvers *)
       
    48 
       
    49 local
       
    50 
       
    51 fun make_command command options problem_path proof_path =
       
    52   "(exec 2>&1;" :: map File.shell_quote (command () @ options) @
       
    53   [File.shell_path problem_path, ")", ">", File.shell_path proof_path]
       
    54   |> space_implode " "
       
    55 
       
    56 fun with_trace ctxt msg f x =
       
    57   let val _ = SMT2_Config.trace_msg ctxt (fn () => msg) ()
       
    58   in f x end
       
    59 
       
    60 fun run ctxt name mk_cmd input =
       
    61   (case SMT2_Config.certificates_of ctxt of
       
    62     NONE =>
       
    63       if not (SMT2_Config.is_available ctxt name) then
       
    64         error ("The SMT solver " ^ quote name ^ " is not installed")
       
    65       else if Config.get ctxt SMT2_Config.debug_files = "" then
       
    66         with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input
       
    67       else
       
    68         let
       
    69           val base_path = Path.explode (Config.get ctxt SMT2_Config.debug_files)
       
    70           val in_path = Path.ext "smt2_in" base_path
       
    71           val out_path = Path.ext "smt2_out" base_path
       
    72         in Cache_IO.raw_run mk_cmd input in_path out_path end
       
    73   | SOME certs =>
       
    74       (case Cache_IO.lookup certs input of
       
    75         (NONE, key) =>
       
    76           if Config.get ctxt SMT2_Config.read_only_certificates then
       
    77             error ("Bad certificate cache: missing certificate")
       
    78           else
       
    79             Cache_IO.run_and_cache certs key mk_cmd input
       
    80       | (SOME output, _) =>
       
    81           with_trace ctxt ("Using cached certificate from " ^
       
    82             File.shell_path (Cache_IO.cache_path_of certs) ^ " ...") I output))
       
    83 
       
    84 (* Z3 returns 1 if "get-model" or "get-model" fails *)
       
    85 val normal_return_codes = [0, 1]
       
    86 
       
    87 fun run_solver ctxt name mk_cmd input =
       
    88   let
       
    89     fun pretty tag lines = Pretty.string_of (Pretty.big_list tag (map Pretty.str lines))
       
    90 
       
    91     val _ = SMT2_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
       
    92 
       
    93     val {redirected_output = res, output = err, return_code} =
       
    94       SMT2_Config.with_timeout ctxt (run ctxt name mk_cmd) input
       
    95     val _ = SMT2_Config.trace_msg ctxt (pretty "Solver:") err
       
    96 
       
    97     val output = fst (take_suffix (equal "") res)
       
    98     val _ = SMT2_Config.trace_msg ctxt (pretty "Result:") output
       
    99 
       
   100     val _ = member (op =) normal_return_codes return_code orelse
       
   101       raise SMT2_Failure.SMT (SMT2_Failure.Abnormal_Termination return_code)
       
   102   in output end
       
   103 
       
   104 fun trace_assms ctxt =
       
   105   SMT2_Config.trace_msg ctxt (Pretty.string_of o
       
   106     Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd))
       
   107 
       
   108 fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT2_Translate.replay_data) =
       
   109   let
       
   110     fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
       
   111     fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
       
   112     fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
       
   113   in
       
   114     SMT2_Config.trace_msg ctxt (fn () =>
       
   115       Pretty.string_of (Pretty.big_list "Names:" [
       
   116         Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),
       
   117         Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
       
   118   end
       
   119 
       
   120 in
       
   121 
       
   122 fun invoke name command smt_options ithms ctxt =
       
   123   let
       
   124     val options = SMT2_Config.solver_options_of ctxt
       
   125     val comments = [space_implode " " options]
       
   126 
       
   127     val (str, replay_data as {context = ctxt', ...}) =
       
   128       ithms
       
   129       |> tap (trace_assms ctxt)
       
   130       |> SMT2_Translate.translate ctxt smt_options comments
       
   131       ||> tap trace_replay_data
       
   132   in (run_solver ctxt' name (make_command command options) str, replay_data) end
       
   133 
       
   134 end
       
   135 
       
   136 
       
   137 (* configuration *)
       
   138 
       
   139 datatype outcome = Unsat | Sat | Unknown
       
   140 
       
   141 type parsed_proof =
       
   142   {outcome: SMT2_Failure.failure option,
       
   143    fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list,
       
   144    atp_proof: unit -> (term, string) ATP_Proof.atp_step list}
       
   145 
       
   146 type solver_config =
       
   147   {name: string,
       
   148    class: Proof.context -> SMT2_Util.class,
       
   149    avail: unit -> bool,
       
   150    command: unit -> string list,
       
   151    options: Proof.context -> string list,
       
   152    smt_options: (string * string) list,
       
   153    default_max_relevant: int,
       
   154    outcome: string -> string list -> outcome * string list,
       
   155    parse_proof: (Proof.context -> SMT2_Translate.replay_data ->
       
   156      ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
       
   157      parsed_proof) option,
       
   158    replay: (Proof.context -> SMT2_Translate.replay_data -> string list -> thm) option}
       
   159 
       
   160 
       
   161 (* check well-sortedness *)
       
   162 
       
   163 val has_topsort = Term.exists_type (Term.exists_subtype (fn
       
   164     TFree (_, []) => true
       
   165   | TVar (_, []) => true
       
   166   | _ => false))
       
   167 
       
   168 (* top sorts cause problems with atomization *)
       
   169 fun check_topsort ctxt thm =
       
   170   if has_topsort (Thm.prop_of thm) then (SMT2_Normalize.drop_fact_warning ctxt thm; TrueI) else thm
       
   171 
       
   172 
       
   173 (* registry *)
       
   174 
       
   175 type solver_info = {
       
   176   command: unit -> string list,
       
   177   smt_options: (string * string) list,
       
   178   default_max_relevant: int,
       
   179   parse_proof: Proof.context -> SMT2_Translate.replay_data ->
       
   180     ((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->
       
   181     parsed_proof,
       
   182   replay: Proof.context -> SMT2_Translate.replay_data -> string list -> thm}
       
   183 
       
   184 structure Solvers = Generic_Data
       
   185 (
       
   186   type T = solver_info Symtab.table
       
   187   val empty = Symtab.empty
       
   188   val extend = I
       
   189   fun merge data = Symtab.merge (K true) data
       
   190 )
       
   191 
       
   192 local
       
   193   fun parse_proof outcome parse_proof0 outer_ctxt replay_data xfacts prems concl output =
       
   194     (case outcome output of
       
   195       (Unsat, lines) =>
       
   196         (case parse_proof0 of
       
   197           SOME pp => pp outer_ctxt replay_data xfacts prems concl lines
       
   198         | NONE => {outcome = NONE, fact_ids = [], atp_proof = K []})
       
   199     | (result, _) => raise SMT2_Failure.SMT (SMT2_Failure.Counterexample (result = Sat)))
       
   200 
       
   201   fun replay outcome replay0 oracle outer_ctxt
       
   202       (replay_data as {context = ctxt, ...} : SMT2_Translate.replay_data) output =
       
   203     (case outcome output of
       
   204       (Unsat, lines) =>
       
   205         if not (Config.get ctxt SMT2_Config.oracle) andalso is_some replay0
       
   206         then the replay0 outer_ctxt replay_data lines
       
   207         else oracle ()
       
   208     | (result, _) => raise SMT2_Failure.SMT (SMT2_Failure.Counterexample (result = Sat)))
       
   209 
       
   210   val cfalse = Thm.cterm_of @{theory} @{prop False}
       
   211 in
       
   212 
       
   213 fun add_solver ({name, class, avail, command, options, smt_options, default_max_relevant, outcome,
       
   214     parse_proof = parse_proof0, replay = replay0} : solver_config) =
       
   215   let
       
   216     fun solver oracle = {
       
   217       command = command,
       
   218       smt_options = smt_options,
       
   219       default_max_relevant = default_max_relevant,
       
   220       parse_proof = parse_proof (outcome name) parse_proof0,
       
   221       replay = replay (outcome name) replay0 oracle}
       
   222 
       
   223     val info = {name = name, class = class, avail = avail, options = options}
       
   224   in
       
   225     Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, oracle) =>
       
   226     Context.theory_map (Solvers.map (Symtab.update_new (name, solver oracle)))) #>
       
   227     Context.theory_map (SMT2_Config.add_solver info)
       
   228   end
       
   229 
       
   230 end
       
   231 
       
   232 fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
       
   233 
       
   234 fun name_and_info_of ctxt =
       
   235   let val name = SMT2_Config.solver_of ctxt
       
   236   in (name, get_info ctxt name) end
       
   237 
       
   238 val default_max_relevant = #default_max_relevant oo get_info
       
   239 
       
   240 fun apply_solver_and_replay ctxt thms0 =
       
   241   let
       
   242     val thms = map (check_topsort ctxt) thms0
       
   243     val (name, {command, smt_options, replay, ...}) = name_and_info_of ctxt
       
   244     val (output, replay_data) =
       
   245       invoke name command smt_options (SMT2_Normalize.normalize ctxt thms) ctxt
       
   246   in replay ctxt replay_data output end
       
   247 
       
   248 
       
   249 (* filter *)
       
   250 
       
   251 fun smt2_filter ctxt0 goal xfacts i time_limit =
       
   252   let
       
   253     val ctxt = ctxt0 |> Config.put SMT2_Config.timeout (Time.toReal time_limit)
       
   254 
       
   255     val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i goal
       
   256     fun negate ct = Thm.dest_comb ct ||> Thm.apply @{cterm Not} |-> Thm.apply
       
   257     val cprop =
       
   258       (case try negate (Thm.rhs_of (SMT2_Normalize.atomize_conv ctxt concl)) of
       
   259         SOME ct => ct
       
   260       | NONE => raise SMT2_Failure.SMT (SMT2_Failure.Other_Failure "goal is not a HOL term"))
       
   261 
       
   262     val conjecture = Thm.assume cprop
       
   263     val facts = map snd xfacts
       
   264     val thms = conjecture :: prems @ facts
       
   265     val thms' = map (check_topsort ctxt) thms
       
   266 
       
   267     val (name, {command, smt_options, parse_proof, ...}) = name_and_info_of ctxt
       
   268     val (output, replay_data) =
       
   269       invoke name command smt_options (SMT2_Normalize.normalize ctxt thms') ctxt
       
   270   in
       
   271     parse_proof ctxt replay_data xfacts (map Thm.prop_of prems) (Thm.term_of concl) output
       
   272   end
       
   273   handle SMT2_Failure.SMT fail => {outcome = SOME fail, fact_ids = [], atp_proof = K []}
       
   274 
       
   275 
       
   276 (* SMT tactic *)
       
   277 
       
   278 local
       
   279   fun str_of ctxt fail =
       
   280     "Solver " ^ SMT2_Config.solver_of ctxt ^ ": " ^ SMT2_Failure.string_of_failure fail
       
   281 
       
   282   fun safe_solve ctxt facts = SOME (apply_solver_and_replay ctxt facts)
       
   283     handle
       
   284       SMT2_Failure.SMT (fail as SMT2_Failure.Counterexample _) =>
       
   285         (SMT2_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
       
   286     | SMT2_Failure.SMT (fail as SMT2_Failure.Time_Out) =>
       
   287         error ("SMT: Solver " ^ quote (SMT2_Config.solver_of ctxt) ^ ": " ^
       
   288           SMT2_Failure.string_of_failure fail ^ " (setting the " ^
       
   289           "configuration option " ^ quote (Config.name_of SMT2_Config.timeout) ^ " might help)")
       
   290     | SMT2_Failure.SMT fail => error (str_of ctxt fail)
       
   291 
       
   292   fun resolve (SOME thm) = rtac thm 1
       
   293     | resolve NONE = no_tac
       
   294 
       
   295   fun tac prove ctxt rules =
       
   296     CONVERSION (SMT2_Normalize.atomize_conv ctxt)
       
   297     THEN' rtac @{thm ccontr}
       
   298     THEN' SUBPROOF (fn {context = ctxt, prems, ...} => resolve (prove ctxt (rules @ prems))) ctxt
       
   299 in
       
   300 
       
   301 val smt2_tac = tac safe_solve
       
   302 val smt2_tac' = tac (SOME oo apply_solver_and_replay)
       
   303 
       
   304 end
       
   305 
       
   306 end;