| author | wenzelm | 
| Fri, 22 Jul 2022 15:15:26 +0200 | |
| changeset 75686 | 42f19e398ee4 | 
| parent 75365 | 83940294cc67 | 
| child 78177 | ea7a3cc64df5 | 
| 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 | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 12 | val bvsmlibC: SMT_Util.class | 
| 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 13 | val add_logic: int * (string -> term list -> string option) -> Context.generic -> Context.generic | 
| 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 14 | val del_logic: int * (string -> term list -> string option) -> Context.generic -> Context.generic | 
| 66551 | 15 | val translate_config: SMT_Util.order -> Proof.context -> SMT_Translate.config | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 16 | val assert_name_of_role_and_index: SMT_Util.role -> int -> string | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 17 | val role_and_index_of_assert_name: string -> SMT_Util.role * int | 
| 57229 | 18 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 19 | |
| 58061 | 20 | structure SMTLIB_Interface: SMTLIB_INTERFACE = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 21 | struct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 22 | |
| 66551 | 23 | val hoC = ["ho"] | 
| 24 | ||
| 25 | val smtlibC = ["smtlib"] (* SMT-LIB 2 *) | |
| 26 | val hosmtlibC = smtlibC @ hoC (* possibly SMT-LIB 3 *) | |
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 27 | val bvsmlibC = smtlibC @ ["BV"] (* if BV are supported *) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 28 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 29 | (* builtins *) | 
| 
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 | local | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | fun int_num _ i = SOME (string_of_int i) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 33 | |
| 58061 | 34 | fun is_linear [t] = SMT_Util.is_number t | 
| 35 | | 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 | 36 | | is_linear _ = false | 
| 
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 | fun times _ _ ts = | 
| 74382 | 39 | let val mk = Term.list_comb o pair \<^Const>\<open>times \<^Type>\<open>int\<close>\<close> | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 40 |     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 | 41 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 42 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 43 | val setup_builtins = | 
| 66551 | 44 | SMT_Builtin.add_builtin_typ hosmtlibC | 
| 69593 | 45 |     (\<^typ>\<open>'a => 'b\<close>, fn Type (\<^type_name>\<open>fun\<close>, Ts) => SOME ("->", Ts), K (K NONE)) #>
 | 
| 58061 | 46 | fold (SMT_Builtin.add_builtin_typ smtlibC) [ | 
| 69593 | 47 |     (\<^typ>\<open>bool\<close>, K (SOME ("Bool", [])), K (K NONE)),
 | 
| 48 |     (\<^typ>\<open>int\<close>, K (SOME ("Int", [])), int_num)] #>
 | |
| 58061 | 49 | fold (SMT_Builtin.add_builtin_fun' smtlibC) [ | 
| 74382 | 50 | (\<^Const>\<open>True\<close>, "true"), | 
| 51 | (\<^Const>\<open>False\<close>, "false"), | |
| 52 | (\<^Const>\<open>Not\<close>, "not"), | |
| 53 | (\<^Const>\<open>conj\<close>, "and"), | |
| 54 | (\<^Const>\<open>disj\<close>, "or"), | |
| 55 | (\<^Const>\<open>implies\<close>, "=>"), | |
| 56 | (\<^Const>\<open>HOL.eq \<^typ>\<open>'a\<close>\<close>, "="), | |
| 57 | (\<^Const>\<open>If \<^typ>\<open>'a\<close>\<close>, "ite"), | |
| 58 | (\<^Const>\<open>less \<^Type>\<open>int\<close>\<close>, "<"), | |
| 59 | (\<^Const>\<open>less_eq \<^Type>\<open>int\<close>\<close>, "<="), | |
| 60 | (\<^Const>\<open>uminus \<^Type>\<open>int\<close>\<close>, "-"), | |
| 61 | (\<^Const>\<open>plus \<^Type>\<open>int\<close>\<close>, "+"), | |
| 62 | (\<^Const>\<open>minus \<^Type>\<open>int\<close>\<close>, "-")] #> | |
| 58061 | 63 | SMT_Builtin.add_builtin_fun smtlibC | 
| 74382 | 64 | (Term.dest_Const \<^Const>\<open>times \<^Type>\<open>int\<close>\<close>, times) | 
| 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 | end | 
| 
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 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 69 | (* serialization *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 70 | |
| 57238 | 71 | (** logic **) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 72 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 73 | 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 | 74 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 75 | structure Logics = Generic_Data | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 76 | ( | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 77 | type T = (int * (string -> term list -> string option)) list | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 | val empty = [] | 
| 
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) | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 83 | fun del_logic pf = Logics.map (Ord_List.remove fst_int_ord pf) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 | |
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 85 | fun choose_logic ctxt prover ts = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 86 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 | fun choose [] = "AUFLIA" | 
| 74817 
1fd8705503b4
generate problems with correct logic for veriT
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
74561diff
changeset | 88 | | choose ((_, f) :: fs) = (case f prover 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 | 89 | 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 | 90 | (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 | 91 | "" => "" (* 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 | 92 | | 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 | 93 | end | 
| 56078 
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 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 96 | (** serialization **) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 97 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 98 | fun var i = "?v" ^ string_of_int i | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 99 | |
| 66551 | 100 | fun tree_of_sterm l (SMT_Translate.SVar (i, [])) = SMTLIB.Sym (var (l - i - 1)) | 
| 101 | | tree_of_sterm l (SMT_Translate.SVar (i, ts)) = | |
| 102 | SMTLIB.S (SMTLIB.Sym (var (l - i - 1)) :: map (tree_of_sterm l) ts) | |
| 103 | | tree_of_sterm _ (SMT_Translate.SConst (n, [])) = SMTLIB.Sym n | |
| 104 | | tree_of_sterm l (SMT_Translate.SConst (n, ts)) = | |
| 58061 | 105 | SMTLIB.S (SMTLIB.Sym n :: map (tree_of_sterm l) ts) | 
| 106 | | 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 | 107 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 108 | val l' = l + length ss | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 109 | |
| 58061 | 110 | fun quant_name SMT_Translate.SForall = "forall" | 
| 111 | | quant_name SMT_Translate.SExists = "exists" | |
| 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 | 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 | 114 | [SMTLIB.Key keyword, SMTLIB.S (map (tree_of_sterm l') ps)] | 
| 58061 | 115 | fun trees_of_pat (SMT_Translate.SPat ps) = gen_trees_of_pat "pattern" ps | 
| 116 | | 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 | 117 | fun tree_of_pats [] t = t | 
| 58061 | 118 | | 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 | 119 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 120 | val vs = map_index (fn (i, ty) => | 
| 58061 | 121 | 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 | 122 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 123 | val body = t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 124 | |> tree_of_sterm l' | 
| 57165 
7b1bf424ec5f
removed SMT weights -- their impact was very inconclusive anyway
 blanchet parents: 
56091diff
changeset | 125 | |> tree_of_pats pats | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 126 | in | 
| 58061 | 127 | 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 | 128 | end | 
| 
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 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 131 | fun sctrarg (sel, typ) = "(" ^ sel ^ " " ^ typ ^ ")"
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 132 | 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 | 133 | 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 | 134 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 135 | 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 | 136 | |
| 58061 | 137 | 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 | 138 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 139 | val conjecture_prefix = "conjecture" (* FUDGE *) | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 140 | val hypothesis_prefix = "hypothesis" (* FUDGE *) | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 141 | val axiom_prefix = "axiom" (* FUDGE *) | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 142 | |
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 143 | fun assert_prefix_of_role SMT_Util.Conjecture = conjecture_prefix | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 144 | | assert_prefix_of_role SMT_Util.Hypothesis = hypothesis_prefix | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 145 | | assert_prefix_of_role SMT_Util.Axiom = axiom_prefix | 
| 57203 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 146 | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 147 | fun assert_name_of_role_and_index role i = assert_prefix_of_role role ^ string_of_int i | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 148 | |
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 149 | fun unprefix_axiom s = | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 150 | try (pair SMT_Util.Conjecture o unprefix conjecture_prefix) s | 
| 75365 | 151 | |> curry merge_options (try (pair SMT_Util.Hypothesis o unprefix hypothesis_prefix) s) | 
| 152 | |> curry merge_options (try (pair SMT_Util.Axiom o unprefix axiom_prefix) s) | |
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 153 | |> the | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 154 | |
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 155 | fun role_and_index_of_assert_name s = | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 156 | apsnd (the_default ~1 o Int.fromString) (unprefix_axiom s) | 
| 57203 
171fa7d56c4f
generate ':named' attributes in SMT-LIB 2 problems, to facilitate reconstruction
 blanchet parents: 
57165diff
changeset | 157 | |
| 58429 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 158 | fun sdtyp (fp, dtyps) = | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 159 |   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 | 160 | (space_implode "\n " (map sdatatype dtyps))) | 
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 161 | |
| 
0b94858325a5
interleave (co)datatypes in the right order w.r.t. dependencies
 blanchet parents: 
58361diff
changeset | 162 | fun serialize smt_options comments {logic, sorts, dtyps, funcs} ts =
 | 
| 59015 | 163 | let | 
| 67405 
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
 wenzelm parents: 
67399diff
changeset | 164 |     val unsat_core = member (op =) smt_options (":produce-unsat-cores", "true")
 | 
| 59015 | 165 | in | 
| 166 | Buffer.empty | |
| 167 | |> fold (Buffer.add o enclose "; " "\n") comments | |
| 168 |     |> fold (fn (k, v) => Buffer.add ("(set-option " ^ k ^ " " ^ v ^ ")\n")) smt_options
 | |
| 169 | |> Buffer.add logic | |
| 170 | |> fold (Buffer.add o enclose "(declare-sort " " 0)\n") (sort fast_string_ord sorts) | |
| 171 | |> fold sdtyp (AList.coalesce (op =) dtyps) | |
| 172 | |> 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 | 173 | (sort (fast_string_ord o apply2 fst) funcs) | 
| 75274 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 174 | |> fold (fn (i, (role, t)) => Buffer.add (enclose "(assert " ")\n" | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 175 | (SMTLIB.str_of (named_sterm (assert_name_of_role_and_index role i) (tree_of_sterm 0 t))))) | 
| 
e89709b80b6e
used more descriptive assert names in SMT-Lib output
 desharna parents: 
74817diff
changeset | 176 | (map_index I ts) | 
| 59015 | 177 | |> Buffer.add "(check-sat)\n" | 
| 178 | |> Buffer.add (if unsat_core then "(get-unsat-core)\n" else "(get-proof)\n") | |
| 179 | |> Buffer.content | |
| 180 | end | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 181 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 182 | (* interface *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 183 | |
| 66551 | 184 | fun translate_config order ctxt = {
 | 
| 185 | order = order, | |
| 57238 | 186 | logic = choose_logic ctxt, | 
| 58360 | 187 | fp_kinds = [], | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 188 | serialize = serialize} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 189 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 190 | val _ = Theory.setup (Context.theory_map | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 191 | (setup_builtins #> | 
| 66551 | 192 | SMT_Translate.add_config (smtlibC, translate_config SMT_Util.First_Order) #> | 
| 193 | 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 | 194 | |
| 57229 | 195 | end; |