| author | wenzelm | 
| Wed, 12 Oct 2016 15:48:05 +0200 | |
| changeset 64168 | e573b985390c | 
| parent 63102 | 71059cf60658 | 
| child 64304 | 96bc94c87a81 | 
| permissions | -rw-r--r-- | 
| 58061 | 1  | 
(* Title: HOL/Tools/SMT/smt_solver.ML  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
2  | 
Author: Sascha Boehme, TU Muenchen  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
4  | 
SMT solvers registry and SMT tactic.  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
6  | 
|
| 58061 | 7  | 
signature SMT_SOLVER =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
9  | 
(*configuration*)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
10  | 
datatype outcome = Unsat | Sat | Unknown  | 
| 57163 | 11  | 
|
12  | 
type parsed_proof =  | 
|
| 58061 | 13  | 
    {outcome: SMT_Failure.failure option,
 | 
| 60201 | 14  | 
fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option,  | 
| 57163 | 15  | 
atp_proof: unit -> (term, string) ATP_Proof.atp_step list}  | 
16  | 
||
17  | 
type solver_config =  | 
|
18  | 
    {name: string,
 | 
|
| 58061 | 19  | 
class: Proof.context -> SMT_Util.class,  | 
| 57163 | 20  | 
avail: unit -> bool,  | 
21  | 
command: unit -> string list,  | 
|
22  | 
options: Proof.context -> string list,  | 
|
| 57239 | 23  | 
smt_options: (string * string) list,  | 
| 57163 | 24  | 
default_max_relevant: int,  | 
25  | 
outcome: string -> string list -> outcome * string list,  | 
|
| 58061 | 26  | 
parse_proof: (Proof.context -> SMT_Translate.replay_data ->  | 
| 57164 | 27  | 
((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->  | 
28  | 
parsed_proof) option,  | 
|
| 58061 | 29  | 
replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option}  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
31  | 
(*registry*)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
32  | 
val add_solver: solver_config -> theory -> theory  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
33  | 
val default_max_relevant: Proof.context -> string -> int  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
35  | 
(*filter*)  | 
| 58061 | 36  | 
val smt_filter: Proof.context -> thm -> ((string * ATP_Problem_Generate.stature) * thm) list ->  | 
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
37  | 
int -> Time.time -> parsed_proof  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
39  | 
(*tactic*)  | 
| 58061 | 40  | 
val smt_tac: Proof.context -> thm list -> int -> tactic  | 
41  | 
val smt_tac': Proof.context -> thm list -> int -> tactic  | 
|
| 57229 | 42  | 
end;  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
43  | 
|
| 58061 | 44  | 
structure SMT_Solver: SMT_SOLVER =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
45  | 
struct  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
47  | 
(* interface to external solvers *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
49  | 
local  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
50  | 
|
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
51  | 
fun make_command command options problem_path proof_path =  | 
| 
62549
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
52  | 
"(exec 2>&1;" :: map File.bash_string (command () @ options) @  | 
| 
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
53  | 
[File.bash_path problem_path, ")", ">", File.bash_path proof_path]  | 
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
54  | 
|> space_implode " "  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
55  | 
|
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
56  | 
fun with_trace ctxt msg f x =  | 
| 58061 | 57  | 
let val _ = SMT_Config.trace_msg ctxt (fn () => msg) ()  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
58  | 
in f x end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
60  | 
fun run ctxt name mk_cmd input =  | 
| 58061 | 61  | 
(case SMT_Config.certificates_of ctxt of  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
62  | 
NONE =>  | 
| 58061 | 63  | 
if not (SMT_Config.is_available ctxt name) then  | 
| 56131 | 64  | 
        error ("The SMT solver " ^ quote name ^ " is not installed")
 | 
| 58061 | 65  | 
else if Config.get ctxt SMT_Config.debug_files = "" then  | 
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
66  | 
        with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
67  | 
else  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
68  | 
let  | 
| 58061 | 69  | 
val base_path = Path.explode (Config.get ctxt SMT_Config.debug_files)  | 
70  | 
val in_path = Path.ext "smt_in" base_path  | 
|
71  | 
val out_path = Path.ext "smt_out" base_path  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
72  | 
in Cache_IO.raw_run mk_cmd input in_path out_path end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
73  | 
| SOME certs =>  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
74  | 
(case Cache_IO.lookup certs input of  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
75  | 
(NONE, key) =>  | 
| 58061 | 76  | 
if Config.get ctxt SMT_Config.read_only_certificates then  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
77  | 
            error ("Bad certificate cache: missing certificate")
 | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
78  | 
else  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
79  | 
Cache_IO.run_and_cache certs key mk_cmd input  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
80  | 
| (SOME output, _) =>  | 
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
81  | 
          with_trace ctxt ("Using cached certificate from " ^
 | 
| 
62549
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
61268 
diff
changeset
 | 
82  | 
Path.print (Cache_IO.cache_path_of certs) ^ " ...") I output))  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
83  | 
|
| 
63102
 
71059cf60658
better handling of veriT's 'unknown' status
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
62549 
diff
changeset
 | 
84  | 
(* Z3 returns 1 if "get-proof" or "get-model" fails. veriT returns 255. *)  | 
| 
 
71059cf60658
better handling of veriT's 'unknown' status
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
62549 
diff
changeset
 | 
85  | 
val normal_return_codes = [0, 1, 255]  | 
| 56087 | 86  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
87  | 
fun run_solver ctxt name mk_cmd input =  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
88  | 
let  | 
| 57164 | 89  | 
fun pretty tag lines = Pretty.string_of (Pretty.big_list tag (map Pretty.str lines))  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
90  | 
|
| 58061 | 91  | 
val _ = SMT_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
92  | 
|
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
93  | 
    val {redirected_output = res, output = err, return_code} =
 | 
| 58061 | 94  | 
SMT_Config.with_timeout ctxt (run ctxt name mk_cmd) input  | 
95  | 
val _ = SMT_Config.trace_msg ctxt (pretty "Solver:") err  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
96  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
97  | 
val output = fst (take_suffix (equal "") res)  | 
| 58061 | 98  | 
val _ = SMT_Config.trace_msg ctxt (pretty "Result:") output  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
99  | 
|
| 56087 | 100  | 
val _ = member (op =) normal_return_codes return_code orelse  | 
| 58061 | 101  | 
raise SMT_Failure.SMT (SMT_Failure.Abnormal_Termination return_code)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
102  | 
in output end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
104  | 
fun trace_assms ctxt =  | 
| 58061 | 105  | 
SMT_Config.trace_msg ctxt (Pretty.string_of o  | 
| 61268 | 106  | 
Pretty.big_list "Assertions:" o map (Thm.pretty_thm ctxt o snd))  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
107  | 
|
| 58061 | 108  | 
fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT_Translate.replay_data) =
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
109  | 
let  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
110  | 
fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
111  | 
fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
112  | 
fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
113  | 
in  | 
| 58061 | 114  | 
SMT_Config.trace_msg ctxt (fn () =>  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
115  | 
Pretty.string_of (Pretty.big_list "Names:" [  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
116  | 
Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
117  | 
Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
118  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
120  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
121  | 
|
| 57239 | 122  | 
fun invoke name command smt_options ithms ctxt =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
123  | 
let  | 
| 58061 | 124  | 
val options = SMT_Config.solver_options_of ctxt  | 
| 56112 | 125  | 
val comments = [space_implode " " options]  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
126  | 
|
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
127  | 
    val (str, replay_data as {context = ctxt', ...}) =
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
128  | 
ithms  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
129  | 
|> tap (trace_assms ctxt)  | 
| 58061 | 130  | 
|> SMT_Translate.translate ctxt smt_options comments  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
131  | 
||> tap trace_replay_data  | 
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
132  | 
in (run_solver ctxt' name (make_command command options) str, replay_data) end  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
133  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
134  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
137  | 
(* configuration *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
139  | 
datatype outcome = Unsat | Sat | Unknown  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
140  | 
|
| 57163 | 141  | 
type parsed_proof =  | 
| 58061 | 142  | 
  {outcome: SMT_Failure.failure option,
 | 
| 60201 | 143  | 
fact_ids: (int * ((string * ATP_Problem_Generate.stature) * thm)) list option,  | 
| 57163 | 144  | 
atp_proof: unit -> (term, string) ATP_Proof.atp_step list}  | 
145  | 
||
146  | 
type solver_config =  | 
|
147  | 
  {name: string,
 | 
|
| 58061 | 148  | 
class: Proof.context -> SMT_Util.class,  | 
| 57163 | 149  | 
avail: unit -> bool,  | 
150  | 
command: unit -> string list,  | 
|
151  | 
options: Proof.context -> string list,  | 
|
| 57239 | 152  | 
smt_options: (string * string) list,  | 
| 57163 | 153  | 
default_max_relevant: int,  | 
154  | 
outcome: string -> string list -> outcome * string list,  | 
|
| 58061 | 155  | 
parse_proof: (Proof.context -> SMT_Translate.replay_data ->  | 
| 57164 | 156  | 
((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->  | 
157  | 
parsed_proof) option,  | 
|
| 58061 | 158  | 
replay: (Proof.context -> SMT_Translate.replay_data -> string list -> thm) option}  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
159  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
160  | 
|
| 56106 | 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 =  | 
|
| 58061 | 170  | 
if has_topsort (Thm.prop_of thm) then (SMT_Normalize.drop_fact_warning ctxt thm; TrueI) else thm  | 
| 56106 | 171  | 
|
172  | 
||
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
173  | 
(* registry *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
174  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
175  | 
type solver_info = {
 | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
176  | 
command: unit -> string list,  | 
| 57239 | 177  | 
smt_options: (string * string) list,  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
178  | 
default_max_relevant: int,  | 
| 58061 | 179  | 
parse_proof: Proof.context -> SMT_Translate.replay_data ->  | 
| 57164 | 180  | 
((string * ATP_Problem_Generate.stature) * thm) list -> term list -> term -> string list ->  | 
181  | 
parsed_proof,  | 
|
| 58061 | 182  | 
replay: Proof.context -> SMT_Translate.replay_data -> string list -> thm}  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
184  | 
structure Solvers = Generic_Data  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
185  | 
(  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
186  | 
type T = solver_info Symtab.table  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
187  | 
val empty = Symtab.empty  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
188  | 
val extend = I  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
189  | 
fun merge data = Symtab.merge (K true) data  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
190  | 
)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
192  | 
local  | 
| 57164 | 193  | 
fun parse_proof outcome parse_proof0 outer_ctxt replay_data xfacts prems concl output =  | 
| 57157 | 194  | 
(case outcome output of  | 
| 57164 | 195  | 
(Unsat, lines) =>  | 
196  | 
(case parse_proof0 of  | 
|
197  | 
SOME pp => pp outer_ctxt replay_data xfacts prems concl lines  | 
|
| 60201 | 198  | 
        | NONE => {outcome = NONE, fact_ids = NONE, atp_proof = K []})
 | 
| 58061 | 199  | 
| (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat)))  | 
| 57157 | 200  | 
|
201  | 
fun replay outcome replay0 oracle outer_ctxt  | 
|
| 58061 | 202  | 
      (replay_data as {context = ctxt, ...} : SMT_Translate.replay_data) output =
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
203  | 
(case outcome output of  | 
| 57164 | 204  | 
(Unsat, lines) =>  | 
| 58061 | 205  | 
if not (Config.get ctxt SMT_Config.oracle) andalso is_some replay0  | 
| 57164 | 206  | 
then the replay0 outer_ctxt replay_data lines  | 
| 57157 | 207  | 
else oracle ()  | 
| 58061 | 208  | 
| (result, _) => raise SMT_Failure.SMT (SMT_Failure.Counterexample (result = Sat)))  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
209  | 
|
| 59632 | 210  | 
  val cfalse = Thm.cterm_of @{context} @{prop False}
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
211  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
212  | 
|
| 57239 | 213  | 
fun add_solver ({name, class, avail, command, options, smt_options, default_max_relevant, outcome,
 | 
| 57164 | 214  | 
parse_proof = parse_proof0, replay = replay0} : solver_config) =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
215  | 
let  | 
| 57157 | 216  | 
    fun solver oracle = {
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
217  | 
command = command,  | 
| 57239 | 218  | 
smt_options = smt_options,  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
219  | 
default_max_relevant = default_max_relevant,  | 
| 57157 | 220  | 
parse_proof = parse_proof (outcome name) parse_proof0,  | 
221  | 
replay = replay (outcome name) replay0 oracle}  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
222  | 
|
| 
56132
 
64eeda68e693
delayed construction of command (and of noncommercial check) + tuning
 
blanchet 
parents: 
56131 
diff
changeset
 | 
223  | 
    val info = {name = name, class = class, avail = avail, options = options}
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
224  | 
in  | 
| 57157 | 225  | 
Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, oracle) =>  | 
226  | 
Context.theory_map (Solvers.map (Symtab.update_new (name, solver oracle)))) #>  | 
|
| 58061 | 227  | 
Context.theory_map (SMT_Config.add_solver info)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
228  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
229  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
230  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
231  | 
|
| 56125 | 232  | 
fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
233  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
234  | 
fun name_and_info_of ctxt =  | 
| 58061 | 235  | 
let val name = SMT_Config.solver_of ctxt  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
236  | 
in (name, get_info ctxt name) end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
237  | 
|
| 57163 | 238  | 
val default_max_relevant = #default_max_relevant oo get_info  | 
| 57157 | 239  | 
|
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
240  | 
fun apply_solver_and_replay ctxt thms0 =  | 
| 57157 | 241  | 
let  | 
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
242  | 
val thms = map (check_topsort ctxt) thms0  | 
| 57239 | 243  | 
    val (name, {command, smt_options, replay, ...}) = name_and_info_of ctxt
 | 
244  | 
val (output, replay_data) =  | 
|
| 58061 | 245  | 
invoke name command smt_options (SMT_Normalize.normalize ctxt thms) ctxt  | 
| 57157 | 246  | 
in replay ctxt replay_data output end  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
247  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
249  | 
(* filter *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
250  | 
|
| 58061 | 251  | 
fun smt_filter ctxt0 goal xfacts i time_limit =  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
252  | 
let  | 
| 58061 | 253  | 
val ctxt = ctxt0 |> Config.put SMT_Config.timeout (Time.toReal time_limit)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
254  | 
|
| 
60695
 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 
wenzelm 
parents: 
60201 
diff
changeset
 | 
255  | 
    val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i NONE goal
 | 
| 56106 | 256  | 
    fun negate ct = Thm.dest_comb ct ||> Thm.apply @{cterm Not} |-> Thm.apply
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
257  | 
val cprop =  | 
| 58061 | 258  | 
(case try negate (Thm.rhs_of (SMT_Normalize.atomize_conv ctxt concl)) of  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
259  | 
SOME ct => ct  | 
| 
61033
 
fd7fe96ca7b9
generate proper error instead of exception if goal cannot be atomized
 
blanchet 
parents: 
60752 
diff
changeset
 | 
260  | 
| NONE => raise SMT_Failure.SMT (SMT_Failure.Other_Failure "cannot atomize goal"))  | 
| 56082 | 261  | 
|
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
262  | 
val conjecture = Thm.assume cprop  | 
| 
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
263  | 
val facts = map snd xfacts  | 
| 
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
264  | 
val thms = conjecture :: prems @ facts  | 
| 
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
265  | 
val thms' = map (check_topsort ctxt) thms  | 
| 57159 | 266  | 
|
| 57239 | 267  | 
    val (name, {command, smt_options, parse_proof, ...}) = name_and_info_of ctxt
 | 
268  | 
val (output, replay_data) =  | 
|
| 58061 | 269  | 
invoke name command smt_options (SMT_Normalize.normalize ctxt thms') ctxt  | 
| 57164 | 270  | 
in  | 
271  | 
parse_proof ctxt replay_data xfacts (map Thm.prop_of prems) (Thm.term_of concl) output  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
272  | 
end  | 
| 60201 | 273  | 
  handle SMT_Failure.SMT fail => {outcome = SOME fail, fact_ids = NONE, atp_proof = K []}
 | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
274  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
275  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
276  | 
(* SMT tactic *)  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
277  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
278  | 
local  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
279  | 
fun str_of ctxt fail =  | 
| 58061 | 280  | 
"Solver " ^ SMT_Config.solver_of ctxt ^ ": " ^ SMT_Failure.string_of_failure fail  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
281  | 
|
| 
57165
 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 
blanchet 
parents: 
57164 
diff
changeset
 | 
282  | 
fun safe_solve ctxt facts = SOME (apply_solver_and_replay ctxt facts)  | 
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
283  | 
handle  | 
| 58061 | 284  | 
SMT_Failure.SMT (fail as SMT_Failure.Counterexample _) =>  | 
285  | 
(SMT_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)  | 
|
286  | 
| SMT_Failure.SMT (fail as SMT_Failure.Time_Out) =>  | 
|
287  | 
        error ("SMT: Solver " ^ quote (SMT_Config.solver_of ctxt) ^ ": " ^
 | 
|
288  | 
SMT_Failure.string_of_failure fail ^ " (setting the " ^  | 
|
289  | 
"configuration option " ^ quote (Config.name_of SMT_Config.timeout) ^ " might help)")  | 
|
290  | 
| SMT_Failure.SMT fail => error (str_of ctxt fail)  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
291  | 
|
| 60752 | 292  | 
fun resolve ctxt (SOME thm) = resolve_tac ctxt [thm] 1  | 
293  | 
| resolve _ NONE = no_tac  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
294  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
295  | 
fun tac prove ctxt rules =  | 
| 58061 | 296  | 
CONVERSION (SMT_Normalize.atomize_conv ctxt)  | 
| 60752 | 297  | 
    THEN' resolve_tac ctxt @{thms ccontr}
 | 
298  | 
    THEN' SUBPROOF (fn {context = ctxt', prems, ...} =>
 | 
|
299  | 
resolve ctxt' (prove ctxt' (rules @ prems))) ctxt  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
300  | 
in  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
301  | 
|
| 58061 | 302  | 
val smt_tac = tac safe_solve  | 
303  | 
val smt_tac' = tac (SOME oo apply_solver_and_replay)  | 
|
| 
56078
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
304  | 
|
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
305  | 
end  | 
| 
 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 
blanchet 
parents:  
diff
changeset
 | 
306  | 
|
| 57229 | 307  | 
end;  |