| author | wenzelm | 
| Tue, 09 May 2017 21:43:46 +0200 | |
| changeset 65796 | 7d1c5150af70 | 
| parent 59058 | a78612c67ec0 | 
| child 66134 | a1fb6beb2731 | 
| permissions | -rw-r--r-- | 
| 58061 | 1 | (* Title: HOL/Tools/SMT/smtlib_interface.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 | Author: Jasmin Blanchette, TU Muenchen | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 4 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 5 | Interface to SMT solvers based on the SMT-LIB 2 format. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 6 | *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 7 | |
| 58061 | 8 | signature SMTLIB_INTERFACE = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 9 | sig | 
| 58061 | 10 | val smtlibC: SMT_Util.class | 
| 56090 | 11 | val add_logic: int * (term list -> string option) -> Context.generic -> Context.generic | 
| 58061 | 12 | val translate_config: Proof.context -> SMT_Translate.config | 
| 58482 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58429diff
changeset | 13 | val assert_name_of_index: int -> string | 
| 57203 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 14 | val assert_index_of_name: string -> int | 
| 57704 | 15 | val assert_prefix : string | 
| 57229 | 16 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 17 | |
| 58061 | 18 | structure SMTLIB_Interface: SMTLIB_INTERFACE = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 19 | struct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | |
| 58061 | 21 | val smtlibC = ["smtlib"] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 22 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 23 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 24 | (* builtins *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 25 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 26 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 27 | fun int_num _ i = SOME (string_of_int i) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 28 | |
| 58061 | 29 | fun is_linear [t] = SMT_Util.is_number t | 
| 30 | | is_linear [t, u] = SMT_Util.is_number t orelse SMT_Util.is_number u | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 31 | | is_linear _ = false | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 33 | fun times _ _ ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 34 |     let val mk = Term.list_comb o pair @{const times (int)}
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 35 |     in if is_linear ts then SOME ("*", 2, ts, mk) else NONE end
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 36 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 38 | val setup_builtins = | 
| 58061 | 39 | fold (SMT_Builtin.add_builtin_typ smtlibC) [ | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 40 |     (@{typ bool}, K (SOME "Bool"), K (K NONE)),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 41 |     (@{typ int}, K (SOME "Int"), int_num)] #>
 | 
| 58061 | 42 | fold (SMT_Builtin.add_builtin_fun' smtlibC) [ | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 43 |     (@{const True}, "true"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 44 |     (@{const False}, "false"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 45 |     (@{const Not}, "not"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 46 |     (@{const HOL.conj}, "and"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 47 |     (@{const HOL.disj}, "or"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 48 |     (@{const HOL.implies}, "=>"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 49 |     (@{const HOL.eq ('a)}, "="),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 50 |     (@{const If ('a)}, "ite"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 51 |     (@{const less (int)}, "<"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 52 |     (@{const less_eq (int)}, "<="),
 | 
| 57711 
caadd484dec6
Changing ~ into - for unuary minus (not supported by veriT)
 fleury parents: 
57704diff
changeset | 53 |     (@{const uminus (int)}, "-"),
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 54 |     (@{const plus (int)}, "+"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 55 |     (@{const minus (int)}, "-")] #>
 | 
| 58061 | 56 | SMT_Builtin.add_builtin_fun smtlibC | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 57 |     (Term.dest_Const @{const times (int)}, times)
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 58 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 59 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 | (* serialization *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 | |
| 57238 | 64 | (** logic **) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 65 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 66 | fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 67 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 68 | structure Logics = Generic_Data | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 69 | ( | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 70 | type T = (int * (term list -> string option)) list | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 71 | val empty = [] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 72 | val extend = I | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 73 | fun merge data = Ord_List.merge fst_int_ord data | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 74 | ) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 75 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 76 | fun add_logic pf = Logics.map (Ord_List.insert fst_int_ord pf) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 77 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 | fun choose_logic ctxt ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 80 | fun choose [] = "AUFLIA" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 81 | | choose ((_, f) :: fs) = (case f ts of SOME s => s | NONE => choose fs) | 
| 57553 
2a6c31ac1c9a
don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
 blanchet parents: 
57239diff
changeset | 82 | in | 
| 
2a6c31ac1c9a
don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
 blanchet parents: 
57239diff
changeset | 83 | (case choose (Logics.get (Context.Proof ctxt)) of | 
| 
2a6c31ac1c9a
don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
 blanchet parents: 
57239diff
changeset | 84 | "" => "" (* for default Z3 logic, a subset of everything *) | 
| 
2a6c31ac1c9a
don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
 blanchet parents: 
57239diff
changeset | 85 | | logic => "(set-logic " ^ logic ^ ")\n") | 
| 
2a6c31ac1c9a
don't generate a 'set-logic' command when generating problems in a non-standard (but Z3-supported) union-of-everything logic
 blanchet parents: 
57239diff
changeset | 86 | end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 88 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 89 | (** serialization **) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 90 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 91 | fun var i = "?v" ^ string_of_int i | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 92 | |
| 58061 | 93 | fun tree_of_sterm l (SMT_Translate.SVar i) = SMTLIB.Sym (var (l - i - 1)) | 
| 94 | | tree_of_sterm _ (SMT_Translate.SApp (n, [])) = SMTLIB.Sym n | |
| 95 | | tree_of_sterm l (SMT_Translate.SApp (n, ts)) = | |
| 96 | SMTLIB.S (SMTLIB.Sym n :: map (tree_of_sterm l) ts) | |
| 97 | | tree_of_sterm _ (SMT_Translate.SLet _) = | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 98 | raise Fail "SMT-LIB: unsupported let expression" | 
| 58061 | 99 | | tree_of_sterm l (SMT_Translate.SQua (q, ss, pats, t)) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 100 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 101 | val l' = l + length ss | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 102 | |
| 58061 | 103 | fun quant_name SMT_Translate.SForall = "forall" | 
| 104 | | quant_name SMT_Translate.SExists = "exists" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 105 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 106 | fun gen_trees_of_pat keyword ps = | 
| 59021 
b29281d6d1db
always generate patterns as lists of lists, to avoid confusing CVC4 (and stick to what SMT-LIB 2 actually says)
 blanchet parents: 
59015diff
changeset | 107 | [SMTLIB.Key keyword, SMTLIB.S (map (tree_of_sterm l') ps)] | 
| 58061 | 108 | fun trees_of_pat (SMT_Translate.SPat ps) = gen_trees_of_pat "pattern" ps | 
| 109 | | trees_of_pat (SMT_Translate.SNoPat ps) = gen_trees_of_pat "no-pattern" ps | |
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56091diff
changeset | 110 | fun tree_of_pats [] t = t | 
| 58061 | 111 | | tree_of_pats pats t = SMTLIB.S (SMTLIB.Sym "!" :: t :: maps trees_of_pat pats) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 112 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 113 | val vs = map_index (fn (i, ty) => | 
| 58061 | 114 | SMTLIB.S [SMTLIB.Sym (var (l + i)), SMTLIB.Sym ty]) ss | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 115 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 116 | val body = t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 117 | |> tree_of_sterm l' | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56091diff
changeset | 118 | |> tree_of_pats pats | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 119 | in | 
| 58061 | 120 | SMTLIB.S [SMTLIB.Sym (quant_name q), SMTLIB.S vs, body] | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 121 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 122 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 123 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 124 | fun sctrarg (sel, typ) = "(" ^ sel ^ " " ^ typ ^ ")"
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 125 | fun sctr (name, args) = enclose "(" ")" (space_implode " " (name :: map sctrarg args))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 126 | fun sdatatype (name, ctrs) = enclose "(" ")" (space_implode " " (name :: map sctr ctrs))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 127 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 128 | fun string_of_fun (f, (ss, s)) = f ^ " (" ^ space_implode " " ss ^ ") " ^ s
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 129 | |
| 58061 | 130 | fun named_sterm s t = SMTLIB.S [SMTLIB.Sym "!", t, SMTLIB.Key "named", SMTLIB.Sym s] | 
| 57203 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 131 | |
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 132 | val assert_prefix = "a" | 
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 133 | |
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 134 | fun assert_name_of_index i = assert_prefix ^ string_of_int i | 
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 135 | fun assert_index_of_name s = the_default ~1 (Int.fromString (unprefix assert_prefix s)) | 
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 136 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 137 | fun sdtyp (fp, dtyps) = | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 138 |   Buffer.add (enclose ("(declare-" ^ BNF_FP_Util.co_prefix fp ^ "datatypes () (") "))\n"
 | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 139 | (space_implode "\n " (map sdatatype dtyps))) | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 140 | |
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 141 | fun serialize smt_options comments {logic, sorts, dtyps, funcs} ts =
 | 
| 59015 | 142 | let | 
| 143 |     val unsat_core = member (op =) smt_options (":produce-unsat-cores", "true")
 | |
| 144 | in | |
| 145 | Buffer.empty | |
| 146 | |> fold (Buffer.add o enclose "; " "\n") comments | |
| 147 |     |> fold (fn (k, v) => Buffer.add ("(set-option " ^ k ^ " " ^ v ^ ")\n")) smt_options
 | |
| 148 | |> Buffer.add logic | |
| 149 | |> fold (Buffer.add o enclose "(declare-sort " " 0)\n") (sort fast_string_ord sorts) | |
| 150 | |> fold sdtyp (AList.coalesce (op =) dtyps) | |
| 151 | |> fold (Buffer.add o enclose "(declare-fun " ")\n" o string_of_fun) | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
59021diff
changeset | 152 | (sort (fast_string_ord o apply2 fst) funcs) | 
| 59015 | 153 | |> fold (fn (i, t) => Buffer.add (enclose "(assert " ")\n" | 
| 154 | (SMTLIB.str_of (named_sterm (assert_name_of_index i) (tree_of_sterm 0 t))))) (map_index I ts) | |
| 155 | |> Buffer.add "(check-sat)\n" | |
| 156 | |> Buffer.add (if unsat_core then "(get-unsat-core)\n" else "(get-proof)\n") | |
| 157 | |> Buffer.content | |
| 158 | end | |
| 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 | (* interface *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 161 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 162 | fun translate_config ctxt = {
 | 
| 57238 | 163 | logic = choose_logic ctxt, | 
| 58360 | 164 | fp_kinds = [], | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 165 | serialize = serialize} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 166 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 167 | val _ = Theory.setup (Context.theory_map | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 168 | (setup_builtins #> | 
| 58061 | 169 | SMT_Translate.add_config (smtlibC, translate_config))) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 170 | |
| 57229 | 171 | end; |