| author | desharna | 
| Tue, 28 Sep 2021 10:47:18 +0200 | |
| changeset 74370 | d8dc8fdc46fc | 
| parent 69597 | ff784d5a5bfb | 
| child 74382 | 8d0294d877bd | 
| 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 | 
| 66551 | 11 | val hosmtlibC: SMT_Util.class | 
| 56090 | 12 | val add_logic: int * (term list -> string option) -> Context.generic -> Context.generic | 
| 66551 | 13 | val translate_config: SMT_Util.order -> Proof.context -> SMT_Translate.config | 
| 58482 
7836013951e6
simplified and repaired veriT index handling code
 blanchet parents: 
58429diff
changeset | 14 | val assert_name_of_index: int -> string | 
| 57203 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 15 | val assert_index_of_name: string -> int | 
| 57704 | 16 | val assert_prefix : string | 
| 57229 | 17 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 18 | |
| 58061 | 19 | structure SMTLIB_Interface: SMTLIB_INTERFACE = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | struct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 21 | |
| 66551 | 22 | val hoC = ["ho"] | 
| 23 | ||
| 24 | val smtlibC = ["smtlib"] (* SMT-LIB 2 *) | |
| 25 | val hosmtlibC = smtlibC @ hoC (* possibly SMT-LIB 3 *) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 26 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 27 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 28 | (* builtins *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 29 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 30 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 31 | fun int_num _ i = SOME (string_of_int i) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | |
| 58061 | 33 | fun is_linear [t] = SMT_Util.is_number t | 
| 34 | | 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 | 35 | | is_linear _ = false | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 36 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | fun times _ _ ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 38 |     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 | 39 |     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 | 40 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 41 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 42 | val setup_builtins = | 
| 66551 | 43 | SMT_Builtin.add_builtin_typ hosmtlibC | 
| 69593 | 44 |     (\<^typ>\<open>'a => 'b\<close>, fn Type (\<^type_name>\<open>fun\<close>, Ts) => SOME ("->", Ts), K (K NONE)) #>
 | 
| 58061 | 45 | fold (SMT_Builtin.add_builtin_typ smtlibC) [ | 
| 69593 | 46 |     (\<^typ>\<open>bool\<close>, K (SOME ("Bool", [])), K (K NONE)),
 | 
| 47 |     (\<^typ>\<open>int\<close>, K (SOME ("Int", [])), int_num)] #>
 | |
| 58061 | 48 | fold (SMT_Builtin.add_builtin_fun' smtlibC) [ | 
| 69597 | 49 | (\<^const>\<open>True\<close>, "true"), | 
| 50 | (\<^const>\<open>False\<close>, "false"), | |
| 51 | (\<^const>\<open>Not\<close>, "not"), | |
| 52 | (\<^const>\<open>HOL.conj\<close>, "and"), | |
| 53 | (\<^const>\<open>HOL.disj\<close>, "or"), | |
| 54 | (\<^const>\<open>HOL.implies\<close>, "=>"), | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 55 |     (@{const HOL.eq ('a)}, "="),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 56 |     (@{const If ('a)}, "ite"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 57 |     (@{const less (int)}, "<"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 58 |     (@{const less_eq (int)}, "<="),
 | 
| 57711 
caadd484dec6
Changing ~ into - for unuary minus (not supported by veriT)
 fleury parents: 
57704diff
changeset | 59 |     (@{const uminus (int)}, "-"),
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 |     (@{const plus (int)}, "+"),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 |     (@{const minus (int)}, "-")] #>
 | 
| 58061 | 62 | SMT_Builtin.add_builtin_fun smtlibC | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 |     (Term.dest_Const @{const times (int)}, times)
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 64 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 65 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 66 | |
| 
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 | (* serialization *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 69 | |
| 57238 | 70 | (** logic **) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 71 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 72 | 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 | 73 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 74 | structure Logics = Generic_Data | 
| 
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 | type T = (int * (term list -> string option)) list | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 77 | val empty = [] | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 | val extend = I | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 | 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 | 80 | ) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 81 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 82 | 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 | 83 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | fun choose_logic ctxt ts = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 85 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 86 | fun choose [] = "AUFLIA" | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 | | 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 | 88 | 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 | 89 | (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 | 90 | "" => "" (* 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 | 91 | | 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 | 92 | end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 93 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 94 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 95 | (** serialization **) | 
| 
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 | fun var i = "?v" ^ string_of_int i | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 98 | |
| 66551 | 99 | fun tree_of_sterm l (SMT_Translate.SVar (i, [])) = SMTLIB.Sym (var (l - i - 1)) | 
| 100 | | tree_of_sterm l (SMT_Translate.SVar (i, ts)) = | |
| 101 | SMTLIB.S (SMTLIB.Sym (var (l - i - 1)) :: map (tree_of_sterm l) ts) | |
| 102 | | tree_of_sterm _ (SMT_Translate.SConst (n, [])) = SMTLIB.Sym n | |
| 103 | | tree_of_sterm l (SMT_Translate.SConst (n, ts)) = | |
| 58061 | 104 | SMTLIB.S (SMTLIB.Sym n :: map (tree_of_sterm l) ts) | 
| 105 | | 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 | 106 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 107 | val l' = l + length ss | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 108 | |
| 58061 | 109 | fun quant_name SMT_Translate.SForall = "forall" | 
| 110 | | quant_name SMT_Translate.SExists = "exists" | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 111 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 112 | 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 | 113 | [SMTLIB.Key keyword, SMTLIB.S (map (tree_of_sterm l') ps)] | 
| 58061 | 114 | fun trees_of_pat (SMT_Translate.SPat ps) = gen_trees_of_pat "pattern" ps | 
| 115 | | 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 | 116 | fun tree_of_pats [] t = t | 
| 58061 | 117 | | 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 | 118 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 119 | val vs = map_index (fn (i, ty) => | 
| 58061 | 120 | 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 | 121 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 122 | val body = t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 123 | |> tree_of_sterm l' | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56091diff
changeset | 124 | |> tree_of_pats pats | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 125 | in | 
| 58061 | 126 | 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 | 127 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 128 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 129 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 130 | fun sctrarg (sel, typ) = "(" ^ sel ^ " " ^ typ ^ ")"
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 131 | 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 | 132 | 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 | 133 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 134 | 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 | 135 | |
| 58061 | 136 | 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 | 137 | |
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 138 | val assert_prefix = "a" | 
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 139 | |
| 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 140 | 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 | 141 | 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 | 142 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 143 | fun sdtyp (fp, dtyps) = | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 144 |   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 | 145 | (space_implode "\n " (map sdatatype dtyps))) | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 146 | |
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 147 | fun serialize smt_options comments {logic, sorts, dtyps, funcs} ts =
 | 
| 59015 | 148 | let | 
| 67405 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 wenzelm parents: 
67399diff
changeset | 149 |     val unsat_core = member (op =) smt_options (":produce-unsat-cores", "true")
 | 
| 59015 | 150 | in | 
| 151 | Buffer.empty | |
| 152 | |> fold (Buffer.add o enclose "; " "\n") comments | |
| 153 |     |> fold (fn (k, v) => Buffer.add ("(set-option " ^ k ^ " " ^ v ^ ")\n")) smt_options
 | |
| 154 | |> Buffer.add logic | |
| 155 | |> fold (Buffer.add o enclose "(declare-sort " " 0)\n") (sort fast_string_ord sorts) | |
| 156 | |> fold sdtyp (AList.coalesce (op =) dtyps) | |
| 157 | |> 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 | 158 | (sort (fast_string_ord o apply2 fst) funcs) | 
| 59015 | 159 | |> fold (fn (i, t) => Buffer.add (enclose "(assert " ")\n" | 
| 160 | (SMTLIB.str_of (named_sterm (assert_name_of_index i) (tree_of_sterm 0 t))))) (map_index I ts) | |
| 161 | |> Buffer.add "(check-sat)\n" | |
| 162 | |> Buffer.add (if unsat_core then "(get-unsat-core)\n" else "(get-proof)\n") | |
| 163 | |> Buffer.content | |
| 164 | end | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 165 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 166 | (* interface *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 167 | |
| 66551 | 168 | fun translate_config order ctxt = {
 | 
| 169 | order = order, | |
| 57238 | 170 | logic = choose_logic ctxt, | 
| 58360 | 171 | fp_kinds = [], | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 172 | serialize = serialize} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 173 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 174 | val _ = Theory.setup (Context.theory_map | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 175 | (setup_builtins #> | 
| 66551 | 176 | SMT_Translate.add_config (smtlibC, translate_config SMT_Util.First_Order) #> | 
| 177 | SMT_Translate.add_config (hosmtlibC, translate_config SMT_Util.Higher_Order))) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 178 | |
| 57229 | 179 | end; |