boehmes@36898: (* Title: HOL/Tools/SMT/smtlib_interface.ML boehmes@36898: Author: Sascha Boehme, TU Muenchen boehmes@36898: boehmes@36898: Interface to SMT solvers based on the SMT-LIB format. boehmes@36898: *) boehmes@36898: boehmes@36898: signature SMTLIB_INTERFACE = boehmes@36898: sig boehmes@41059: val smtlibC: SMT_Config.class boehmes@41059: val add_logic: int * (term list -> string option) -> Context.generic -> boehmes@36899: Context.generic boehmes@41059: val setup: theory -> theory boehmes@36898: val interface: SMT_Solver.interface boehmes@36898: end boehmes@36898: boehmes@36898: structure SMTLIB_Interface: SMTLIB_INTERFACE = boehmes@36898: struct boehmes@36898: boehmes@41059: structure B = SMT_Builtin boehmes@36898: structure N = SMT_Normalize boehmes@36898: structure T = SMT_Translate boehmes@36898: boehmes@41059: val smtlibC = ["smtlib"] boehmes@36898: boehmes@36898: boehmes@41059: boehmes@41059: (* facts about uninterpreted constants *) boehmes@36898: boehmes@36898: infix 2 ?? boehmes@40161: fun (ex ?? f) irules = irules |> exists (ex o Thm.prop_of o snd) irules ? f boehmes@36898: boehmes@36898: boehmes@41059: (** pairs **) boehmes@36898: boehmes@36898: val pair_rules = [@{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] boehmes@36898: haftmann@37678: val pair_type = (fn Type (@{type_name Product_Type.prod}, _) => true | _ => false) boehmes@36898: val exists_pair_type = Term.exists_type (Term.exists_subtype pair_type) boehmes@36898: boehmes@40161: val add_pair_rules = exists_pair_type ?? append (map (pair ~1) pair_rules) boehmes@36898: boehmes@36898: boehmes@41059: (** function update **) boehmes@36898: boehmes@36898: val fun_upd_rules = [@{thm fun_upd_same}, @{thm fun_upd_apply}] boehmes@36898: boehmes@36898: val is_fun_upd = (fn Const (@{const_name fun_upd}, _) => true | _ => false) boehmes@36898: val exists_fun_upd = Term.exists_subterm is_fun_upd boehmes@36898: boehmes@40161: val add_fun_upd_rules = exists_fun_upd ?? append (map (pair ~1) fun_upd_rules) boehmes@36898: boehmes@36898: boehmes@41059: (** abs/min/max **) boehmes@36898: boehmes@36898: val exists_abs_min_max = Term.exists_subterm (fn boehmes@36898: Const (@{const_name abs}, _) => true boehmes@36898: | Const (@{const_name min}, _) => true boehmes@36898: | Const (@{const_name max}, _) => true boehmes@36898: | _ => false) boehmes@36898: boehmes@37786: val unfold_abs_conv = Conv.rewr_conv (mk_meta_eq @{thm abs_if}) boehmes@37786: val unfold_min_conv = Conv.rewr_conv (mk_meta_eq @{thm min_def}) boehmes@37786: val unfold_max_conv = Conv.rewr_conv (mk_meta_eq @{thm max_def}) boehmes@36898: boehmes@36898: fun expand_conv cv = N.eta_expand_conv (K cv) boehmes@36898: fun expand2_conv cv = N.eta_expand_conv (N.eta_expand_conv (K cv)) boehmes@36898: boehmes@36898: fun unfold_def_conv ctxt ct = boehmes@36898: (case Thm.term_of ct of boehmes@36898: Const (@{const_name abs}, _) $ _ => unfold_abs_conv boehmes@36898: | Const (@{const_name abs}, _) => expand_conv unfold_abs_conv ctxt boehmes@36898: | Const (@{const_name min}, _) $ _ $ _ => unfold_min_conv boehmes@36898: | Const (@{const_name min}, _) $ _ => expand_conv unfold_min_conv ctxt boehmes@36898: | Const (@{const_name min}, _) => expand2_conv unfold_min_conv ctxt boehmes@36898: | Const (@{const_name max}, _) $ _ $ _ => unfold_max_conv boehmes@36898: | Const (@{const_name max}, _) $ _ => expand_conv unfold_max_conv ctxt boehmes@36898: | Const (@{const_name max}, _) => expand2_conv unfold_max_conv ctxt boehmes@36898: | _ => Conv.all_conv) ct boehmes@36898: boehmes@36898: fun unfold_abs_min_max_defs ctxt thm = boehmes@36898: if exists_abs_min_max (Thm.prop_of thm) wenzelm@36936: then Conv.fconv_rule (Conv.top_conv unfold_def_conv ctxt) thm boehmes@36898: else thm boehmes@36898: boehmes@36898: boehmes@41059: (** include additional facts **) boehmes@36898: boehmes@40161: fun extra_norm has_datatypes irules ctxt = boehmes@40161: irules boehmes@39298: |> not has_datatypes ? add_pair_rules boehmes@36898: |> add_fun_upd_rules boehmes@40161: |> map (apsnd (unfold_abs_min_max_defs ctxt)) boehmes@36898: |> rpair ctxt boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41059: (* builtins *) boehmes@36898: boehmes@41059: local boehmes@41059: fun int_num _ i = SOME (string_of_int i) boehmes@36899: boehmes@41059: fun distinct _ _ [t] = boehmes@41059: (case try HOLogic.dest_list t of boehmes@41059: SOME (ts as _ :: _) => SOME ("distinct", ts) boehmes@41059: | _ => NONE) boehmes@41059: | distinct _ _ _ = NONE boehmes@41059: in boehmes@36898: boehmes@41059: val setup = boehmes@41059: B.add_builtin_typ smtlibC (@{typ int}, K (SOME "Int"), int_num) #> boehmes@41059: fold (B.add_builtin_fun' smtlibC) [ boehmes@41059: (@{const True}, "true"), boehmes@41059: (@{const False}, "false"), boehmes@41059: (* FIXME: we do not test if these constants are fully applied! *) boehmes@41059: (@{const Not}, "not"), boehmes@41059: (@{const HOL.conj}, "and"), boehmes@41059: (@{const HOL.disj}, "or"), boehmes@41059: (@{const HOL.implies}, "implies"), boehmes@41059: (@{const HOL.eq (bool)}, "iff"), boehmes@41059: (@{const HOL.eq ('a)}, "="), boehmes@41059: (@{const If (bool)}, "if_then_else"), boehmes@41059: (@{const If ('a)}, "ite"), boehmes@41059: (@{const less (int)}, "<"), boehmes@41059: (@{const less_eq (int)}, "<="), boehmes@41059: (@{const uminus (int)}, "~"), boehmes@41059: (@{const plus (int)}, "+"), boehmes@41059: (@{const minus (int)}, "-"), boehmes@41059: (@{const times (int)}, "*") ] #> boehmes@41059: B.add_builtin_fun smtlibC (Term.dest_Const @{const distinct ('a)}, distinct) boehmes@36898: boehmes@41059: end boehmes@36898: boehmes@36898: boehmes@36898: boehmes@41059: (* serialization *) boehmes@36898: boehmes@41059: (** header **) boehmes@41059: boehmes@41059: fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2) boehmes@36899: boehmes@36899: structure Logics = Generic_Data boehmes@36899: ( boehmes@36899: type T = (int * (term list -> string option)) list boehmes@36899: val empty = [] boehmes@36899: val extend = I wenzelm@39687: fun merge (bs1, bs2) = Ord_List.union fst_int_ord bs2 bs1 boehmes@36899: ) boehmes@36899: boehmes@41059: fun add_logic pf = Logics.map (Ord_List.insert fst_int_ord pf) boehmes@36899: boehmes@36899: fun choose_logic ctxt ts = boehmes@36899: let boehmes@36899: fun choose [] = "AUFLIA" boehmes@41059: | choose ((_, f) :: fs) = (case f ts of SOME s => s | NONE => choose fs) boehmes@41059: in [":logic " ^ choose (Logics.get (Context.Proof ctxt))] end boehmes@36899: boehmes@36899: boehmes@36899: (* serialization *) boehmes@36899: boehmes@36898: val add = Buffer.add boehmes@36898: fun sep f = add " " #> f boehmes@36898: fun enclose l r f = sep (add l #> f #> add r) boehmes@36898: val par = enclose "(" ")" boehmes@36898: fun app n f = (fn [] => sep (add n) | xs => par (add n #> fold f xs)) boehmes@36898: fun line f = f #> add "\n" boehmes@36898: boehmes@36898: fun var i = add "?v" #> add (string_of_int i) boehmes@36898: boehmes@36898: fun sterm l (T.SVar i) = sep (var (l - i - 1)) boehmes@36898: | sterm l (T.SApp (n, ts)) = app n (sterm l) ts boehmes@36898: | sterm _ (T.SLet _) = raise Fail "SMT-LIB: unsupported let expression" boehmes@40664: | sterm l (T.SQua (q, ss, ps, w, t)) = boehmes@36898: let boehmes@36898: val quant = add o (fn T.SForall => "forall" | T.SExists => "exists") boehmes@36898: val vs = map_index (apfst (Integer.add l)) ss boehmes@36898: fun var_decl (i, s) = par (var i #> sep (add s)) boehmes@36898: val sub = sterm (l + length ss) boehmes@36898: fun pat kind ts = sep (add kind #> enclose "{" " }" (fold sub ts)) boehmes@36898: fun pats (T.SPat ts) = pat ":pat" ts boehmes@36898: | pats (T.SNoPat ts) = pat ":nopat" ts boehmes@40664: fun weight NONE = I boehmes@40664: | weight (SOME i) = boehmes@40664: sep (add ":weight { " #> add (string_of_int i) #> add " }") boehmes@40664: in boehmes@40664: par (quant q #> fold var_decl vs #> sub t #> fold pats ps #> weight w) boehmes@40664: end boehmes@36898: boehmes@37155: fun ssort sorts = sort fast_string_ord sorts boehmes@37155: fun fsort funcs = sort (prod_ord fast_string_ord (K EQUAL)) funcs boehmes@37155: boehmes@39298: fun sdatatypes decls = boehmes@39298: let boehmes@39298: fun con (n, []) = add n boehmes@39298: | con (n, sels) = par (add n #> boehmes@39298: fold (fn (n, s) => sep (par (add n #> sep (add s)))) sels) boehmes@39298: fun dtyp (n, decl) = add n #> fold (sep o con) decl boehmes@39298: in line (add ":datatypes " #> par (fold (par o dtyp) decls)) end boehmes@39298: boehmes@39298: fun serialize comments {header, sorts, dtyps, funcs} ts = boehmes@36898: Buffer.empty boehmes@36898: |> line (add "(benchmark Isabelle") boehmes@36898: |> line (add ":status unknown") boehmes@36899: |> fold (line o add) header boehmes@36898: |> length sorts > 0 ? boehmes@37155: line (add ":extrasorts" #> par (fold (sep o add) (ssort sorts))) boehmes@39298: |> fold sdatatypes dtyps boehmes@36898: |> length funcs > 0 ? ( boehmes@36898: line (add ":extrafuns" #> add " (") #> boehmes@36898: fold (fn (f, (ss, s)) => boehmes@37155: line (sep (app f (sep o add) (ss @ [s])))) (fsort funcs) #> boehmes@36898: line (add ")")) boehmes@36898: |> fold (fn t => line (add ":assumption" #> sterm 0 t)) ts boehmes@36898: |> line (add ":formula true)") boehmes@36898: |> fold (fn str => line (add "; " #> add str)) comments boehmes@36898: |> Buffer.content boehmes@36898: boehmes@36898: boehmes@36898: boehmes@36899: (** interfaces **) boehmes@36898: boehmes@36898: val interface = { boehmes@41059: class = smtlibC, boehmes@40162: extra_norm = extra_norm, boehmes@36898: translate = { boehmes@36898: prefixes = { boehmes@36898: sort_prefix = "S", boehmes@36898: func_prefix = "f"}, boehmes@36899: header = choose_logic, boehmes@41059: is_fol = true, boehmes@41059: has_datatypes = false, boehmes@36898: serialize = serialize}} boehmes@36898: boehmes@36898: end