author | sultana |
Tue, 17 Apr 2012 16:14:06 +0100 | |
changeset 47509 | 6f215c2ebd72 |
parent 46950 | d0181abdbdac |
child 48891 | c0eafbd55de3 |
permissions | -rw-r--r-- |
33419 | 1 |
(* Title: HOL/Boogie/Boogie.thy |
2 |
Author: Sascha Boehme, TU Muenchen |
|
3 |
*) |
|
4 |
||
5 |
header {* Integration of the Boogie program verifier *} |
|
6 |
||
7 |
theory Boogie |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
35105
diff
changeset
|
8 |
imports Word |
46950
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
45294
diff
changeset
|
9 |
keywords |
d0181abdbdac
declare command keywords via theory header, including strict checking outside Pure;
wenzelm
parents:
45294
diff
changeset
|
10 |
"boogie_open" "boogie_end" :: thy_decl and "boogie_vc" :: thy_goal and "boogie_status" :: diag |
33419 | 11 |
uses |
12 |
("Tools/boogie_vcs.ML") |
|
13 |
("Tools/boogie_loader.ML") |
|
34181
003333ffa543
merged verification condition structure and term representation in one datatype,
boehmes
parents:
34068
diff
changeset
|
14 |
("Tools/boogie_tactics.ML") |
33419 | 15 |
("Tools/boogie_commands.ML") |
16 |
begin |
|
17 |
||
33444 | 18 |
text {* |
19 |
HOL-Boogie and its applications are described in: |
|
20 |
\begin{itemize} |
|
21 |
||
22 |
\item Sascha B"ohme, K. Rustan M. Leino, and Burkhart Wolff. |
|
23 |
HOL-Boogie --- An Interactive Prover for the Boogie Program-Verifier. |
|
24 |
Theorem Proving in Higher Order Logics, 2008. |
|
25 |
||
26 |
\item Sascha B"ohme, Micha{\l} Moskal, Wolfram Schulte, and Burkhart Wolff. |
|
27 |
HOL-Boogie --- An Interactive Prover-Backend for the Verifying C Compiler. |
|
28 |
Journal of Automated Reasoning, 2009. |
|
29 |
||
30 |
\end{itemize} |
|
31 |
*} |
|
32 |
||
33 |
||
34068
a78307d72e58
make assertion labels unique already when loading a verification condition,
boehmes
parents:
33465
diff
changeset
|
34 |
|
33419 | 35 |
section {* Built-in types and functions of Boogie *} |
36 |
||
37 |
subsection {* Labels *} |
|
38 |
||
39 |
text {* |
|
40 |
See "Generating error traces from verification-condition counterexamples" |
|
41 |
by Leino e.a. (2004) for a description of Boogie's labelling mechanism and |
|
42 |
semantics. |
|
43 |
*} |
|
44 |
||
45 |
definition assert_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "assert_at l P = P" |
|
46 |
definition block_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "block_at l P = P" |
|
47 |
||
48 |
lemmas labels = assert_at_def block_at_def |
|
49 |
||
50 |
||
51 |
subsection {* Integer arithmetics *} |
|
52 |
||
53 |
text {* |
|
54 |
The operations @{text div} and @{text mod} are built-in in Boogie, but |
|
55 |
without a particular semantics due to different interpretations in |
|
56 |
programming languages. We assume that each application comes with a |
|
57 |
proper axiomatization. |
|
58 |
*} |
|
59 |
||
60 |
axiomatization |
|
61 |
boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and |
|
62 |
boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70) |
|
63 |
||
64 |
||
65 |
||
66 |
section {* Boogie environment *} |
|
67 |
||
68 |
text {* |
|
69 |
Proving Boogie-generated verification conditions happens inside |
|
70 |
a Boogie environment: |
|
71 |
||
72 |
boogie_open "b2i file without extension" |
|
73 |
boogie_vc "verification condition 1" ... |
|
74 |
boogie_vc "verification condition 2" ... |
|
75 |
boogie_vc "verification condition 3" ... |
|
76 |
boogie_end |
|
77 |
||
78 |
See the Boogie examples for more details. |
|
79 |
||
80 |
At most one Boogie environment should occur per theory, |
|
81 |
otherwise unexpected effects may arise. |
|
82 |
*} |
|
83 |
||
84 |
||
34068
a78307d72e58
make assertion labels unique already when loading a verification condition,
boehmes
parents:
33465
diff
changeset
|
85 |
|
33419 | 86 |
section {* Setup *} |
87 |
||
88 |
ML {* |
|
89 |
structure Boogie_Axioms = Named_Thms |
|
90 |
( |
|
45294 | 91 |
val name = @{binding boogie} |
33465 | 92 |
val description = |
93 |
"Boogie background axioms loaded along with Boogie verification conditions" |
|
33419 | 94 |
) |
95 |
*} |
|
96 |
setup Boogie_Axioms.setup |
|
97 |
||
98 |
use "Tools/boogie_vcs.ML" |
|
99 |
use "Tools/boogie_loader.ML" |
|
34181
003333ffa543
merged verification condition structure and term representation in one datatype,
boehmes
parents:
34068
diff
changeset
|
100 |
use "Tools/boogie_tactics.ML" |
003333ffa543
merged verification condition structure and term representation in one datatype,
boehmes
parents:
34068
diff
changeset
|
101 |
setup Boogie_Tactics.setup |
003333ffa543
merged verification condition structure and term representation in one datatype,
boehmes
parents:
34068
diff
changeset
|
102 |
|
33419 | 103 |
use "Tools/boogie_commands.ML" |
104 |
setup Boogie_Commands.setup |
|
105 |
||
106 |
end |