diff -r dca8d06ecbba -r 8cecd655eef4 src/HOL/SMT_Examples/boogie.ML --- a/src/HOL/SMT_Examples/boogie.ML Thu Jun 12 01:00:49 2014 +0200 +++ b/src/HOL/SMT_Examples/boogie.ML Thu Jun 12 01:00:49 2014 +0200 @@ -7,7 +7,7 @@ signature BOOGIE = sig val boogie_prove: theory -> string list -> unit -end +end; structure Boogie: BOOGIE = struct @@ -16,7 +16,6 @@ val as_int = fst o read_int o raw_explode - val isabelle_name = let fun purge s = if Symbol.is_letter s orelse Symbol.is_digit s then s else @@ -31,7 +30,6 @@ in prefix "b_" o translate_string purge end - (* context *) type context = @@ -39,54 +37,44 @@ val empty_context: context = (Symtab.empty, Symtab.empty, [], []) - fun add_type name (tds, fds, axs, vcs) = let val T = TFree (isabelle_name name, @{sort type}) val tds' = Symtab.update (name, T) tds in (tds', fds, axs, vcs) end - fun add_func name Ts T unique (tds, fds, axs, vcs) = let val t = Free (isabelle_name name, Ts ---> T) val fds' = Symtab.update (name, (t, unique)) fds in (tds, fds', axs, vcs) end - fun add_axiom t (tds, fds, axs, vcs) = (tds, fds, t :: axs, vcs) fun add_vc t (tds, fds, axs, vcs) = (tds, fds, axs, t :: vcs) - fun lookup_type (tds, _, _, _) name = (case Symtab.lookup tds name of SOME T => T | NONE => error "Undeclared type") - fun lookup_func (_, fds, _, _) name = (case Symtab.lookup fds name of SOME t_unique => t_unique | NONE => error "Undeclared function") - (* constructors *) fun mk_var name T = Free ("V_" ^ isabelle_name name, T) - fun mk_arrayT (Ts, T) = Type (@{type_name "fun"}, [HOLogic.mk_tupleT Ts, T]) - fun mk_binary t (t1, t2) = t $ t1 $ t2 - fun mk_nary _ t [] = t | mk_nary f _ ts = uncurry (fold_rev f) (split_last ts) - fun mk_distinct [] = @{const HOL.True} | mk_distinct [_] = @{const HOL.True} | mk_distinct (t :: ts) = @@ -95,34 +83,27 @@ HOLogic.mk_conj (HOLogic.mk_not (HOLogic.mk_eq (t, u)), u') in fold_rev mk_noteq ts (mk_distinct ts) end - fun mk_store m k v = let val mT = Term.fastype_of m and kT = Term.fastype_of k val vT = Term.fastype_of v in Const (@{const_name fun_upd}, mT --> kT --> vT --> mT) $ m $ k $ v end - fun mk_quant q (Free (x, T)) t = q T $ absfree (x, T) t | mk_quant _ t _ = raise TERM ("bad variable", [t]) - -fun mk_list T = HOLogic.mk_list T - - val patternT = @{typ "SMT2.pattern"} fun mk_pat t = Const (@{const_name "SMT2.pat"}, Term.fastype_of t --> patternT) $ t fun mk_pattern [] = raise TERM ("mk_pattern", []) - | mk_pattern ts = mk_list patternT (map mk_pat ts) - + | mk_pattern ts = SMT2_Util.mk_symb_list patternT (map mk_pat ts) fun mk_trigger [] t = t | mk_trigger pss t = @{term "SMT2.trigger"} $ - mk_list @{typ "SMT2.pattern list"} (map mk_pattern pss) $ t + SMT2_Util.mk_symb_list @{typ "SMT2.pattern SMT2.symb_list"} (map mk_pattern pss) $ t (* parser *) @@ -131,7 +112,6 @@ let fun apply (xs, ls) = f ls |>> (fn x => x :: xs) in funpow (as_int n) apply ([], ls) |>> rev end - fun parse_type _ (["bool"] :: ls) = (@{typ bool}, ls) | parse_type _ (["int"] :: ls) = (@{typ int}, ls) | parse_type cx (["array", arity] :: ls) = @@ -139,42 +119,29 @@ | parse_type cx (("type-con" :: name :: _) :: ls) = (lookup_type cx name, ls) | parse_type _ _ = error "Bad type" - fun parse_expr _ (["true"] :: ls) = (@{term True}, ls) | parse_expr _ (["false"] :: ls) = (@{term False}, ls) | parse_expr cx (["not"] :: ls) = parse_expr cx ls |>> HOLogic.mk_not - | parse_expr cx (["and", n] :: ls) = - parse_nary_expr cx n HOLogic.mk_conj @{term True} ls - | parse_expr cx (["or", n] :: ls) = - parse_nary_expr cx n HOLogic.mk_disj @{term False} ls - | parse_expr cx (["implies"] :: ls) = - parse_bin_expr cx (mk_binary @{term HOL.implies}) ls + | parse_expr cx (["and", n] :: ls) = parse_nary_expr cx n HOLogic.mk_conj @{term True} ls + | parse_expr cx (["or", n] :: ls) = parse_nary_expr cx n HOLogic.mk_disj @{term False} ls + | parse_expr cx (["implies"] :: ls) = parse_bin_expr cx (mk_binary @{term HOL.implies}) ls | parse_expr cx (["="] :: ls) = parse_bin_expr cx HOLogic.mk_eq ls | parse_expr cx (["var", name] :: ls) = parse_type cx ls |>> mk_var name | parse_expr cx (["fun", name, n] :: ls) = let val (t, _) = lookup_func cx name in repeat (parse_expr cx) n ls |>> curry Term.list_comb t end | parse_expr cx (("label" :: _) :: ls) = parse_expr cx ls - | parse_expr _ (["int-num", n] :: ls) = - (HOLogic.mk_number @{typ int} (as_int n), ls) - | parse_expr cx (["<"] :: ls) = - parse_bin_expr cx (mk_binary @{term "op < :: int => _"}) ls - | parse_expr cx (["<="] :: ls) = - parse_bin_expr cx (mk_binary @{term "op <= :: int => _"}) ls - | parse_expr cx ([">"] :: ls) = - parse_bin_expr cx (mk_binary @{term "op < :: int => _"}o swap) ls + | parse_expr _ (["int-num", n] :: ls) = (HOLogic.mk_number @{typ int} (as_int n), ls) + | parse_expr cx (["<"] :: ls) = parse_bin_expr cx (mk_binary @{term "op < :: int => _"}) ls + | parse_expr cx (["<="] :: ls) = parse_bin_expr cx (mk_binary @{term "op <= :: int => _"}) ls + | parse_expr cx ([">"] :: ls) = parse_bin_expr cx (mk_binary @{term "op < :: int => _"}o swap) ls | parse_expr cx ([">="] :: ls) = parse_bin_expr cx (mk_binary @{term "op <= :: int => _"} o swap) ls - | parse_expr cx (["+"] :: ls) = - parse_bin_expr cx (mk_binary @{term "op + :: int => _"}) ls - | parse_expr cx (["-"] :: ls) = - parse_bin_expr cx (mk_binary @{term "op - :: int => _"}) ls - | parse_expr cx (["*"] :: ls) = - parse_bin_expr cx (mk_binary @{term "op * :: int => _"}) ls - | parse_expr cx (["/"] :: ls) = - parse_bin_expr cx (mk_binary @{term boogie_div}) ls - | parse_expr cx (["%"] :: ls) = - parse_bin_expr cx (mk_binary @{term boogie_mod}) ls + | parse_expr cx (["+"] :: ls) = parse_bin_expr cx (mk_binary @{term "op + :: int => _"}) ls + | parse_expr cx (["-"] :: ls) = parse_bin_expr cx (mk_binary @{term "op - :: int => _"}) ls + | parse_expr cx (["*"] :: ls) = parse_bin_expr cx (mk_binary @{term "op * :: int => _"}) ls + | parse_expr cx (["/"] :: ls) = parse_bin_expr cx (mk_binary @{term boogie_div}) ls + | parse_expr cx (["%"] :: ls) = parse_bin_expr cx (mk_binary @{term boogie_mod}) ls | parse_expr cx (["select", n] :: ls) = repeat (parse_expr cx) n ls |>> (fn ts => hd ts $ HOLogic.mk_tuple (tl ts)) @@ -188,14 +155,11 @@ parse_quant cx HOLogic.exists_const vars pats atts ls | parse_expr _ _ = error "Bad expression" - and parse_bin_expr cx f ls = ls |> parse_expr cx ||>> parse_expr cx |>> f - and parse_nary_expr cx n f c ls = repeat (parse_expr cx) n ls |>> mk_nary (curry f) c - and parse_quant cx q vars pats atts ls = let val ((((vs, pss), _), t), ls') = @@ -206,15 +170,12 @@ ||>> parse_expr cx in (fold_rev (mk_quant q) vs (mk_trigger pss t), ls') end - and parse_var cx (["var", name] :: ls) = parse_type cx ls |>> mk_var name | parse_var _ _ = error "Bad variable" - and parse_pat cx (["pat", n] :: ls) = repeat (parse_expr cx) n ls | parse_pat _ _ = error "Bad pattern" - and parse_attr cx (["attribute", name, n] :: ls) = let fun attr (["expr-attr"] :: ls) = parse_expr cx ls |>> K () @@ -223,7 +184,6 @@ in repeat attr n ls |>> K name end | parse_attr _ _ = error "Bad attribute" - fun parse_func cx arity n ls = let val ((Ts, atts), ls') = @@ -231,7 +191,6 @@ val unique = member (op =) atts "unique" in ((split_last Ts, unique), ls') end - fun parse_decl (("type-decl" :: name :: _) :: ls) cx = (ls, add_type name cx) | parse_decl (["fun-decl", name, arity, n] :: ls) cx = let val (((Ts, T), unique), ls') = parse_func cx arity n ls @@ -246,12 +205,10 @@ in (ls', add_vc t cx) end | parse_decl _ _ = error "Bad declaration" - fun parse_lines [] cx = cx | parse_lines ls cx = parse_decl ls cx |-> parse_lines - (* splitting of text lines into a lists of tokens *) fun is_blank c = (c = " " orelse c = "\t" orelse c = "\r" orelse c = "\n") @@ -261,7 +218,6 @@ #> filter (fn [] => false | _ => true) - (* proving verification conditions *) fun add_unique_axioms (tds, fds, axs, vcs) = @@ -271,7 +227,6 @@ |> map (fn (T, ns) => mk_distinct (map (Free o rpair T) ns)) |> (fn axs' => (tds, fds, axs' @ axs, vcs)) - fun build_proof_context thy (tds, fds, axs, vcs) = let val vc = @@ -287,16 +242,13 @@ |> pair (map HOLogic.mk_Trueprop axs, HOLogic.mk_Trueprop vc) end - val boogie_rules = [@{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] @ [@{thm fun_upd_same}, @{thm fun_upd_apply}] - fun boogie_tac ctxt axioms = ALLGOALS (SMT2_Solver.smt2_tac ctxt (boogie_rules @ axioms)) - fun boogie_prove thy lines = let val ((axioms, vc), ctxt) = @@ -305,8 +257,7 @@ |> add_unique_axioms |> build_proof_context thy - val _ = Goal.prove ctxt [] axioms vc (fn {prems, context} => - boogie_tac context prems) + val _ = Goal.prove ctxt [] axioms vc (fn {prems, context} => boogie_tac context prems) val _ = writeln "Verification condition proved successfully" in () end @@ -324,4 +275,4 @@ val _ = boogie_prove thy' lines; in thy' end))) -end +end;