author | wenzelm |
Sun, 07 Nov 2021 19:53:37 +0100 | |
changeset 74726 | 33ed2eb06d68 |
parent 74401 | 1aa05eee4e8b |
permissions | -rw-r--r-- |
52722 | 1 |
(* Title: HOL/SMT_Examples/boogie.ML |
2 |
Author: Sascha Boehme, TU Muenchen |
|
3 |
||
52734 | 4 |
Proving Boogie-generated verification conditions. |
52722 | 5 |
*) |
6 |
||
7 |
signature BOOGIE = |
|
8 |
sig |
|
55788
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
9 |
val boogie_prove: theory -> string list -> unit |
57232 | 10 |
end; |
52722 | 11 |
|
12 |
structure Boogie: BOOGIE = |
|
13 |
struct |
|
14 |
||
15 |
(* utility functions *) |
|
16 |
||
17 |
val as_int = fst o read_int o raw_explode |
|
18 |
||
19 |
val isabelle_name = |
|
52732 | 20 |
let |
52722 | 21 |
fun purge s = if Symbol.is_letter s orelse Symbol.is_digit s then s else |
22 |
(case s of |
|
23 |
"." => "_o_" |
|
24 |
| "_" => "_n_" |
|
25 |
| "$" => "_S_" |
|
26 |
| "@" => "_G_" |
|
27 |
| "#" => "_H_" |
|
28 |
| "^" => "_T_" |
|
29 |
| _ => ("_" ^ string_of_int (ord s) ^ "_")) |
|
30 |
in prefix "b_" o translate_string purge end |
|
31 |
||
32 |
||
33 |
(* context *) |
|
34 |
||
35 |
type context = |
|
36 |
typ Symtab.table * (term * bool) Symtab.table * term list * term list |
|
37 |
||
38 |
val empty_context: context = (Symtab.empty, Symtab.empty, [], []) |
|
39 |
||
40 |
fun add_type name (tds, fds, axs, vcs) = |
|
41 |
let |
|
69597 | 42 |
val T = TFree (isabelle_name name, \<^sort>\<open>type\<close>) |
52722 | 43 |
val tds' = Symtab.update (name, T) tds |
44 |
in (tds', fds, axs, vcs) end |
|
45 |
||
46 |
fun add_func name Ts T unique (tds, fds, axs, vcs) = |
|
47 |
let |
|
48 |
val t = Free (isabelle_name name, Ts ---> T) |
|
49 |
val fds' = Symtab.update (name, (t, unique)) fds |
|
50 |
in (tds, fds', axs, vcs) end |
|
51 |
||
52 |
fun add_axiom t (tds, fds, axs, vcs) = (tds, fds, t :: axs, vcs) |
|
53 |
||
54 |
fun add_vc t (tds, fds, axs, vcs) = (tds, fds, axs, t :: vcs) |
|
55 |
||
56 |
fun lookup_type (tds, _, _, _) name = |
|
57 |
(case Symtab.lookup tds name of |
|
58 |
SOME T => T |
|
59 |
| NONE => error "Undeclared type") |
|
60 |
||
61 |
fun lookup_func (_, fds, _, _) name = |
|
62 |
(case Symtab.lookup fds name of |
|
63 |
SOME t_unique => t_unique |
|
64 |
| NONE => error "Undeclared function") |
|
65 |
||
66 |
||
67 |
(* constructors *) |
|
68 |
||
69 |
fun mk_var name T = Free ("V_" ^ isabelle_name name, T) |
|
70 |
||
74401 | 71 |
fun mk_arrayT (Ts, T) = \<^Type>\<open>fun \<open>HOLogic.mk_tupleT Ts\<close> T\<close> |
52722 | 72 |
|
73 |
fun mk_binary t (t1, t2) = t $ t1 $ t2 |
|
74 |
||
75 |
fun mk_nary _ t [] = t |
|
76 |
| mk_nary f _ ts = uncurry (fold_rev f) (split_last ts) |
|
77 |
||
74401 | 78 |
fun mk_distinct [] = \<^Const>\<open>True\<close> |
79 |
| mk_distinct [_] = \<^Const>\<open>True\<close> |
|
52722 | 80 |
| mk_distinct (t :: ts) = |
81 |
let |
|
82 |
fun mk_noteq u u' = |
|
83 |
HOLogic.mk_conj (HOLogic.mk_not (HOLogic.mk_eq (t, u)), u') |
|
84 |
in fold_rev mk_noteq ts (mk_distinct ts) end |
|
85 |
||
86 |
fun mk_store m k v = |
|
87 |
let |
|
74401 | 88 |
val kT = Term.fastype_of k |
52722 | 89 |
val vT = Term.fastype_of v |
74401 | 90 |
in \<^Const>\<open>fun_upd kT vT for m k v\<close> end |
52722 | 91 |
|
92 |
fun mk_quant q (Free (x, T)) t = q T $ absfree (x, T) t |
|
93 |
| mk_quant _ t _ = raise TERM ("bad variable", [t]) |
|
94 |
||
74401 | 95 |
val patternT = \<^Type>\<open>SMT.pattern\<close> |
52722 | 96 |
|
74401 | 97 |
fun mk_pat t = \<^Const>\<open>SMT.pat \<open>Term.fastype_of t\<close> for t\<close> |
52722 | 98 |
|
99 |
fun mk_pattern [] = raise TERM ("mk_pattern", []) |
|
58061 | 100 |
| mk_pattern ts = SMT_Util.mk_symb_list patternT (map mk_pat ts) |
52722 | 101 |
|
102 |
fun mk_trigger [] t = t |
|
103 |
| mk_trigger pss t = |
|
74401 | 104 |
\<^Const>\<open>SMT.trigger\<close> $ |
69597 | 105 |
SMT_Util.mk_symb_list \<^typ>\<open>SMT.pattern SMT.symb_list\<close> (map mk_pattern pss) $ t |
52722 | 106 |
|
107 |
||
108 |
(* parser *) |
|
109 |
||
110 |
fun repeat f n ls = |
|
111 |
let fun apply (xs, ls) = f ls |>> (fn x => x :: xs) |
|
112 |
in funpow (as_int n) apply ([], ls) |>> rev end |
|
113 |
||
74401 | 114 |
fun parse_type _ (["bool"] :: ls) = (\<^Type>\<open>bool\<close>, ls) |
115 |
| parse_type _ (["int"] :: ls) = (\<^Type>\<open>int\<close>, ls) |
|
52722 | 116 |
| parse_type cx (["array", arity] :: ls) = |
117 |
repeat (parse_type cx) arity ls |>> mk_arrayT o split_last |
|
118 |
| parse_type cx (("type-con" :: name :: _) :: ls) = (lookup_type cx name, ls) |
|
119 |
| parse_type _ _ = error "Bad type" |
|
120 |
||
74401 | 121 |
fun parse_expr _ (["true"] :: ls) = (\<^Const>\<open>True\<close>, ls) |
122 |
| parse_expr _ (["false"] :: ls) = (\<^Const>\<open>False\<close>, ls) |
|
52722 | 123 |
| parse_expr cx (["not"] :: ls) = parse_expr cx ls |>> HOLogic.mk_not |
74401 | 124 |
| parse_expr cx (["and", n] :: ls) = parse_nary_expr cx n HOLogic.mk_conj \<^Const>\<open>True\<close> ls |
125 |
| parse_expr cx (["or", n] :: ls) = parse_nary_expr cx n HOLogic.mk_disj \<^Const>\<open>False\<close> ls |
|
126 |
| parse_expr cx (["implies"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>implies\<close>) ls |
|
52722 | 127 |
| parse_expr cx (["="] :: ls) = parse_bin_expr cx HOLogic.mk_eq ls |
128 |
| parse_expr cx (["var", name] :: ls) = parse_type cx ls |>> mk_var name |
|
129 |
| parse_expr cx (["fun", name, n] :: ls) = |
|
130 |
let val (t, _) = lookup_func cx name |
|
131 |
in repeat (parse_expr cx) n ls |>> curry Term.list_comb t end |
|
132 |
| parse_expr cx (("label" :: _) :: ls) = parse_expr cx ls |
|
74401 | 133 |
| parse_expr _ (["int-num", n] :: ls) = (HOLogic.mk_number \<^Type>\<open>int\<close> (as_int n), ls) |
134 |
| parse_expr cx (["<"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>less \<^Type>\<open>int\<close>\<close>) ls |
|
135 |
| parse_expr cx (["<="] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>less_eq \<^Type>\<open>int\<close>\<close>) ls |
|
136 |
| parse_expr cx ([">"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>less \<^Type>\<open>int\<close>\<close> o swap) ls |
|
137 |
| parse_expr cx ([">="] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>less_eq \<^Type>\<open>int\<close>\<close> o swap) ls |
|
138 |
| parse_expr cx (["+"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>plus \<^Type>\<open>int\<close>\<close>) ls |
|
139 |
| parse_expr cx (["-"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>minus \<^Type>\<open>int\<close>\<close>) ls |
|
140 |
| parse_expr cx (["*"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>times \<^Type>\<open>int\<close>\<close>) ls |
|
141 |
| parse_expr cx (["/"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>boogie_div\<close>) ls |
|
142 |
| parse_expr cx (["%"] :: ls) = parse_bin_expr cx (mk_binary \<^Const>\<open>boogie_mod\<close>) ls |
|
52722 | 143 |
| parse_expr cx (["select", n] :: ls) = |
144 |
repeat (parse_expr cx) n ls |
|
145 |
|>> (fn ts => hd ts $ HOLogic.mk_tuple (tl ts)) |
|
146 |
| parse_expr cx (["store", n] :: ls) = |
|
147 |
repeat (parse_expr cx) n ls |
|
148 |
|>> split_last |
|
149 |
|>> (fn (ts, t) => mk_store (hd ts) (HOLogic.mk_tuple (tl ts)) t) |
|
150 |
| parse_expr cx (["forall", vars, pats, atts] :: ls) = |
|
151 |
parse_quant cx HOLogic.all_const vars pats atts ls |
|
152 |
| parse_expr cx (["exists", vars, pats, atts] :: ls) = |
|
153 |
parse_quant cx HOLogic.exists_const vars pats atts ls |
|
154 |
| parse_expr _ _ = error "Bad expression" |
|
155 |
||
156 |
and parse_bin_expr cx f ls = ls |> parse_expr cx ||>> parse_expr cx |>> f |
|
157 |
||
158 |
and parse_nary_expr cx n f c ls = |
|
159 |
repeat (parse_expr cx) n ls |>> mk_nary (curry f) c |
|
160 |
||
161 |
and parse_quant cx q vars pats atts ls = |
|
162 |
let |
|
163 |
val ((((vs, pss), _), t), ls') = |
|
164 |
ls |
|
165 |
|> repeat (parse_var cx) vars |
|
166 |
||>> repeat (parse_pat cx) pats |
|
167 |
||>> repeat (parse_attr cx) atts |
|
168 |
||>> parse_expr cx |
|
169 |
in (fold_rev (mk_quant q) vs (mk_trigger pss t), ls') end |
|
170 |
||
171 |
and parse_var cx (["var", name] :: ls) = parse_type cx ls |>> mk_var name |
|
172 |
| parse_var _ _ = error "Bad variable" |
|
173 |
||
174 |
and parse_pat cx (["pat", n] :: ls) = repeat (parse_expr cx) n ls |
|
175 |
| parse_pat _ _ = error "Bad pattern" |
|
176 |
||
177 |
and parse_attr cx (["attribute", name, n] :: ls) = |
|
178 |
let |
|
179 |
fun attr (["expr-attr"] :: ls) = parse_expr cx ls |>> K () |
|
180 |
| attr (("string-attr" :: _) :: ls) = ((), ls) |
|
181 |
| attr _ = error "Bad attribute value" |
|
182 |
in repeat attr n ls |>> K name end |
|
183 |
| parse_attr _ _ = error "Bad attribute" |
|
184 |
||
185 |
fun parse_func cx arity n ls = |
|
186 |
let |
|
187 |
val ((Ts, atts), ls') = |
|
188 |
ls |> repeat (parse_type cx) arity ||>> repeat (parse_attr cx) n |
|
67405
e9ab4ad7bd15
uniform use of Standard ML op-infix -- eliminated warnings;
wenzelm
parents:
67399
diff
changeset
|
189 |
val unique = member (op =) atts "unique" |
52722 | 190 |
in ((split_last Ts, unique), ls') end |
191 |
||
192 |
fun parse_decl (("type-decl" :: name :: _) :: ls) cx = (ls, add_type name cx) |
|
193 |
| parse_decl (["fun-decl", name, arity, n] :: ls) cx = |
|
194 |
let val (((Ts, T), unique), ls') = parse_func cx arity n ls |
|
195 |
in (ls', add_func name Ts T unique cx) end |
|
196 |
| parse_decl (("axiom" :: _) :: ls) cx = |
|
197 |
let val (t, ls') = parse_expr cx ls |
|
198 |
in (ls', add_axiom t cx) end |
|
199 |
| parse_decl (("var-decl" :: _) :: ls) cx = |
|
200 |
parse_type cx ls |> snd |> rpair cx |
|
201 |
| parse_decl (("vc" :: _) :: ls) cx = |
|
202 |
let val (t, ls') = parse_expr cx ls |
|
203 |
in (ls', add_vc t cx) end |
|
204 |
| parse_decl _ _ = error "Bad declaration" |
|
205 |
||
206 |
fun parse_lines [] cx = cx |
|
207 |
| parse_lines ls cx = parse_decl ls cx |-> parse_lines |
|
208 |
||
209 |
||
55788
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
210 |
(* splitting of text lines into a lists of tokens *) |
52722 | 211 |
|
212 |
fun is_blank c = (c = " " orelse c = "\t" orelse c = "\r" orelse c = "\n") |
|
213 |
||
55788
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
214 |
val token_lines = |
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
215 |
map (String.tokens (is_blank o str)) |
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
216 |
#> filter (fn [] => false | _ => true) |
52722 | 217 |
|
218 |
||
219 |
(* proving verification conditions *) |
|
220 |
||
221 |
fun add_unique_axioms (tds, fds, axs, vcs) = |
|
222 |
Symtab.fold (fn (_, (t, true)) => cons t | _ => I) fds [] |
|
223 |
|> map (swap o Term.dest_Free) |
|
224 |
|> AList.group (op =) |
|
225 |
|> map (fn (T, ns) => mk_distinct (map (Free o rpair T) ns)) |
|
226 |
|> (fn axs' => (tds, fds, axs' @ axs, vcs)) |
|
227 |
||
228 |
fun build_proof_context thy (tds, fds, axs, vcs) = |
|
229 |
let |
|
230 |
val vc = |
|
231 |
(case vcs of |
|
232 |
[vc] => vc |
|
233 |
| _ => error "Bad number of verification conditions") |
|
234 |
in |
|
235 |
Proof_Context.init_global thy |
|
236 |
|> Symtab.fold (fn (_, T) => Variable.declare_typ T) tds |
|
237 |
|> Symtab.fold (fn (_, (t, _)) => Variable.declare_term t) fds |
|
238 |
|> fold Variable.declare_term axs |
|
239 |
|> fold Variable.declare_term vcs |
|
240 |
|> pair (map HOLogic.mk_Trueprop axs, HOLogic.mk_Trueprop vc) |
|
241 |
end |
|
242 |
||
243 |
val boogie_rules = |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
59936
diff
changeset
|
244 |
[@{thm fst_conv}, @{thm snd_conv}, @{thm prod.collapse}] @ |
52722 | 245 |
[@{thm fun_upd_same}, @{thm fun_upd_apply}] |
246 |
||
247 |
fun boogie_tac ctxt axioms = |
|
58061 | 248 |
ALLGOALS (SMT_Solver.smt_tac ctxt (boogie_rules @ axioms)) |
52722 | 249 |
|
55788
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
250 |
fun boogie_prove thy lines = |
52722 | 251 |
let |
252 |
val ((axioms, vc), ctxt) = |
|
253 |
empty_context |
|
55788
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
254 |
|> parse_lines (token_lines lines) |
52722 | 255 |
|> add_unique_axioms |
54447
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
256 |
|> build_proof_context thy |
52722 | 257 |
|
57232 | 258 |
val _ = Goal.prove ctxt [] axioms vc (fn {prems, context} => boogie_tac context prems) |
52732 | 259 |
val _ = writeln "Verification condition proved successfully" |
52722 | 260 |
|
54447
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
261 |
in () end |
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
262 |
|
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
263 |
|
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
264 |
(* Isar command *) |
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
265 |
|
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
266 |
val _ = |
69597 | 267 |
Outer_Syntax.command \<^command_keyword>\<open>boogie_file\<close> |
54447
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
268 |
"prove verification condition from .b2i file" |
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
69597
diff
changeset
|
269 |
(Resources.provide_parse_file >> (fn get_file => |
54447
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
270 |
Toplevel.theory (fn thy => |
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
271 |
let |
72747
5f9d66155081
clarified theory keywords: loaded_files are determined statically in Scala, but ML needs to do it semantically;
wenzelm
parents:
69597
diff
changeset
|
272 |
val ({lines, ...}, thy') = get_file thy; |
55788
67699e08e969
store blobs / inlined files as separate text lines: smaller values are more healthy for the Poly/ML RTS and allow implicit sharing;
wenzelm
parents:
54447
diff
changeset
|
273 |
val _ = boogie_prove thy' lines; |
54447
019394de2b41
proper thy_load command 'boogie_file' -- avoid direct access to file-system;
wenzelm
parents:
52734
diff
changeset
|
274 |
in thy' end))) |
52722 | 275 |
|
57232 | 276 |
end; |