| 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
 | 
|  |      8 | imports SMT
 | 
|  |      9 | uses
 | 
|  |     10 |   ("Tools/boogie_vcs.ML")
 | 
|  |     11 |   ("Tools/boogie_loader.ML")
 | 
|  |     12 |   ("Tools/boogie_commands.ML")
 | 
|  |     13 |   ("Tools/boogie_split.ML")
 | 
|  |     14 | begin
 | 
|  |     15 | 
 | 
| 33444 |     16 | text {*
 | 
|  |     17 | HOL-Boogie and its applications are described in:
 | 
|  |     18 | \begin{itemize}
 | 
|  |     19 | 
 | 
|  |     20 | \item Sascha B"ohme, K. Rustan M. Leino, and Burkhart Wolff.
 | 
|  |     21 | HOL-Boogie --- An Interactive Prover for the Boogie Program-Verifier.
 | 
|  |     22 | Theorem Proving in Higher Order Logics, 2008.
 | 
|  |     23 | 
 | 
|  |     24 | \item Sascha B"ohme, Micha{\l} Moskal, Wolfram Schulte, and Burkhart Wolff.
 | 
|  |     25 | HOL-Boogie --- An Interactive Prover-Backend for the Verifying C Compiler.
 | 
|  |     26 | Journal of Automated Reasoning, 2009.
 | 
|  |     27 | 
 | 
|  |     28 | \end{itemize}
 | 
|  |     29 | *}
 | 
|  |     30 | 
 | 
|  |     31 | 
 | 
| 33419 |     32 | section {* Built-in types and functions of Boogie *}
 | 
|  |     33 | 
 | 
|  |     34 | subsection {* Labels *}
 | 
|  |     35 | 
 | 
|  |     36 | text {*
 | 
|  |     37 | See "Generating error traces from verification-condition counterexamples"
 | 
|  |     38 | by Leino e.a. (2004) for a description of Boogie's labelling mechanism and
 | 
|  |     39 | semantics.
 | 
|  |     40 | *}
 | 
|  |     41 | 
 | 
|  |     42 | definition assert_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "assert_at l P = P"
 | 
|  |     43 | definition block_at :: "bool \<Rightarrow> bool \<Rightarrow> bool" where "block_at l P = P"
 | 
|  |     44 | 
 | 
|  |     45 | lemmas labels = assert_at_def block_at_def
 | 
|  |     46 | 
 | 
|  |     47 | 
 | 
|  |     48 | subsection {* Arrays *}
 | 
|  |     49 | 
 | 
|  |     50 | abbreviation (input) boogie_select :: "('i \<Rightarrow> 'v) \<Rightarrow> 'i \<Rightarrow> 'v"
 | 
|  |     51 | where "boogie_select \<equiv> (\<lambda>f x. f x)"
 | 
|  |     52 | 
 | 
|  |     53 | abbreviation (input) boogie_store :: 
 | 
|  |     54 |   "('i \<Rightarrow> 'v) \<Rightarrow> 'i \<Rightarrow> 'v \<Rightarrow> 'i \<Rightarrow> 'v"
 | 
|  |     55 | where "boogie_store \<equiv> fun_upd"
 | 
|  |     56 | 
 | 
|  |     57 | 
 | 
|  |     58 | subsection {* Integer arithmetics *}
 | 
|  |     59 | 
 | 
|  |     60 | text {*
 | 
|  |     61 | The operations @{text div} and @{text mod} are built-in in Boogie, but
 | 
|  |     62 | without a particular semantics due to different interpretations in
 | 
|  |     63 | programming languages. We assume that each application comes with a
 | 
|  |     64 | proper axiomatization.
 | 
|  |     65 | *}
 | 
|  |     66 | 
 | 
|  |     67 | axiomatization
 | 
|  |     68 |   boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and
 | 
|  |     69 |   boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70)
 | 
|  |     70 | 
 | 
|  |     71 | 
 | 
|  |     72 | subsection {* Bitvectors *}
 | 
|  |     73 | 
 | 
|  |     74 | text {*
 | 
|  |     75 | Boogie's and Z3's built-in bitvector functions are modelled with
 | 
|  |     76 | functions of the HOL-Word library and the SMT theory of bitvectors.
 | 
|  |     77 | Every of the following bitvector functions is supported by the SMT
 | 
|  |     78 | binding.
 | 
|  |     79 | *}
 | 
|  |     80 | 
 | 
|  |     81 | abbreviation (input) boogie_bv_concat :: 
 | 
|  |     82 |   "'a::len0 word \<Rightarrow> 'b::len0 word \<Rightarrow> 'c::len0 word"
 | 
|  |     83 | where "boogie_bv_concat \<equiv> (\<lambda>w1 w2. word_cat w1 w2)"
 | 
|  |     84 | 
 | 
|  |     85 | abbreviation (input) boogie_bv_extract :: 
 | 
|  |     86 |   "nat \<Rightarrow> nat \<Rightarrow> 'a::len0 word \<Rightarrow> 'b::len0 word"
 | 
|  |     87 | where "boogie_bv_extract \<equiv> (\<lambda>mb lb w. slice lb w)"
 | 
|  |     88 | 
 | 
|  |     89 | abbreviation (input) z3_bvnot :: "'a::len0 word \<Rightarrow> 'a word"
 | 
|  |     90 | where "z3_bvnot \<equiv> (\<lambda>w. NOT w)"
 | 
|  |     91 | 
 | 
|  |     92 | abbreviation (input) z3_bvand :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |     93 | where "z3_bvand \<equiv> (\<lambda>w1 w2. w1 AND w2)"
 | 
|  |     94 | 
 | 
|  |     95 | abbreviation (input) z3_bvor :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |     96 | where "z3_bvor \<equiv> (\<lambda>w1 w2. w1 OR w2)"
 | 
|  |     97 | 
 | 
|  |     98 | abbreviation (input) z3_bvxor :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |     99 | where "z3_bvxor \<equiv> (\<lambda>w1 w2. w1 XOR w2)"
 | 
|  |    100 | 
 | 
|  |    101 | abbreviation (input) z3_bvneg :: "'a::len0 word \<Rightarrow> 'a word"
 | 
|  |    102 | where "z3_bvneg \<equiv> (\<lambda>w. - w)"
 | 
|  |    103 | 
 | 
|  |    104 | abbreviation (input) z3_bvadd :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    105 | where "z3_bvadd \<equiv> (\<lambda>w1 w2. w1 + w2)"
 | 
|  |    106 | 
 | 
|  |    107 | abbreviation (input) z3_bvsub :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    108 | where "z3_bvsub \<equiv> (\<lambda>w1 w2. w1 - w2)"
 | 
|  |    109 | 
 | 
|  |    110 | abbreviation (input) z3_bvmul :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    111 | where "z3_bvmul \<equiv> (\<lambda>w1 w2. w1 * w2)"
 | 
|  |    112 | 
 | 
|  |    113 | abbreviation (input) z3_bvudiv :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    114 | where "z3_bvudiv \<equiv> (\<lambda>w1 w2. w1 div w2)"
 | 
|  |    115 | 
 | 
|  |    116 | abbreviation (input) z3_bvurem :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    117 | where "z3_bvurem \<equiv> (\<lambda>w1 w2. w1 mod w2)"
 | 
|  |    118 | 
 | 
|  |    119 | abbreviation (input) z3_bvsdiv :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    120 | where "z3_bvsdiv \<equiv> (\<lambda>w1 w2. w1 sdiv w2)"
 | 
|  |    121 | 
 | 
|  |    122 | abbreviation (input) z3_bvsrem :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    123 | where "z3_bvsrem \<equiv> (\<lambda>w1 w2. w1 srem w2)"
 | 
|  |    124 | 
 | 
|  |    125 | abbreviation (input) z3_bvshl :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    126 | where "z3_bvshl \<equiv> (\<lambda>w1 w2. bv_shl w1 w2)"
 | 
|  |    127 | 
 | 
|  |    128 | abbreviation (input) z3_bvlshr :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    129 | where "z3_bvlshr \<equiv> (\<lambda>w1 w2. bv_lshr w1 w2)"
 | 
|  |    130 | 
 | 
|  |    131 | abbreviation (input) z3_bvashr :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
 | 
|  |    132 | where "z3_bvashr \<equiv> (\<lambda>w1 w2. bv_ashr w1 w2)"
 | 
|  |    133 | 
 | 
|  |    134 | abbreviation (input) z3_sign_extend :: "'a::len word \<Rightarrow> 'b::len word"
 | 
|  |    135 | where "z3_sign_extend \<equiv> (\<lambda>w. scast w)"
 | 
|  |    136 | 
 | 
|  |    137 | abbreviation (input) z3_zero_extend :: "'a::len0 word \<Rightarrow> 'b::len0 word"
 | 
|  |    138 | where "z3_zero_extend \<equiv> (\<lambda>w. ucast w)"
 | 
|  |    139 | 
 | 
|  |    140 | abbreviation (input) z3_rotate_left :: "nat \<Rightarrow> 'a::len0 word \<Rightarrow> 'a word"
 | 
|  |    141 | where "z3_rotate_left \<equiv> (\<lambda>n w. word_rotl n w)"
 | 
|  |    142 | 
 | 
|  |    143 | abbreviation (input) z3_rotate_right :: "nat \<Rightarrow> 'a::len0 word \<Rightarrow> 'a word"
 | 
|  |    144 | where "z3_rotate_right \<equiv> (\<lambda>n w. word_rotr n w)"
 | 
|  |    145 | 
 | 
|  |    146 | abbreviation (input) z3_bvult :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool"
 | 
|  |    147 | where "z3_bvult \<equiv> (\<lambda>w1 w2. w1 < w2)"
 | 
|  |    148 | 
 | 
|  |    149 | abbreviation (input) z3_bvule :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool"
 | 
|  |    150 | where "z3_bvule \<equiv> (\<lambda>w1 w2. w1 \<le> w2)"
 | 
|  |    151 | 
 | 
|  |    152 | abbreviation (input) z3_bvugt :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool" 
 | 
|  |    153 | where "z3_bvugt \<equiv> (\<lambda>w1 w2. w1 > w2)"
 | 
|  |    154 | 
 | 
|  |    155 | abbreviation (input) z3_bvuge :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool" 
 | 
|  |    156 | where "z3_bvuge \<equiv> (\<lambda>w1 w2. w1 \<ge> w2)"
 | 
|  |    157 | 
 | 
|  |    158 | abbreviation (input) z3_bvslt :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
 | 
|  |    159 | where "z3_bvslt \<equiv> (\<lambda>w1 w2. w1 <s w2)"
 | 
|  |    160 | 
 | 
|  |    161 | abbreviation (input) z3_bvsle :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
 | 
|  |    162 | where "z3_bvsle \<equiv> (\<lambda>w1 w2. w1 <=s w2)"
 | 
|  |    163 | 
 | 
|  |    164 | abbreviation (input) z3_bvsgt :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
 | 
|  |    165 | where "z3_bvsgt \<equiv> (\<lambda>w1 w2. w2 <s w1)"
 | 
|  |    166 | 
 | 
|  |    167 | abbreviation (input) z3_bvsge :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
 | 
|  |    168 | where "z3_bvsge \<equiv> (\<lambda>w1 w2. w2 <=s w1)"
 | 
|  |    169 | 
 | 
|  |    170 | 
 | 
|  |    171 | section {* Boogie environment *}
 | 
|  |    172 | 
 | 
|  |    173 | text {*
 | 
|  |    174 | Proving Boogie-generated verification conditions happens inside
 | 
|  |    175 | a Boogie environment:
 | 
|  |    176 | 
 | 
|  |    177 |   boogie_open "b2i file without extension"
 | 
|  |    178 |   boogie_vc "verification condition 1" ...
 | 
|  |    179 |   boogie_vc "verification condition 2" ...
 | 
|  |    180 |   boogie_vc "verification condition 3" ...
 | 
|  |    181 |   boogie_end
 | 
|  |    182 | 
 | 
|  |    183 | See the Boogie examples for more details.
 | 
|  |    184 |  
 | 
|  |    185 | At most one Boogie environment should occur per theory,
 | 
|  |    186 | otherwise unexpected effects may arise.
 | 
|  |    187 | *}
 | 
|  |    188 | 
 | 
|  |    189 | 
 | 
|  |    190 | section {* Setup *}
 | 
|  |    191 | 
 | 
|  |    192 | ML {*
 | 
|  |    193 | structure Boogie_Axioms = Named_Thms
 | 
|  |    194 | (
 | 
|  |    195 |   val name = "boogie"
 | 
| 33465 |    196 |   val description =
 | 
|  |    197 |     "Boogie background axioms loaded along with Boogie verification conditions"
 | 
| 33419 |    198 | )
 | 
|  |    199 | *}
 | 
|  |    200 | setup Boogie_Axioms.setup
 | 
|  |    201 | 
 | 
|  |    202 | text {*
 | 
|  |    203 | Opening a Boogie environment causes the following list of theorems to be
 | 
|  |    204 | enhanced by all theorems found in Boogie_Axioms.
 | 
|  |    205 | *}
 | 
|  |    206 | ML {*
 | 
|  |    207 | structure Split_VC_SMT_Rules = Named_Thms
 | 
|  |    208 | (
 | 
|  |    209 |   val name = "split_vc_smt"
 | 
| 33465 |    210 |   val description =
 | 
|  |    211 |     "theorems given to the SMT sub-tactic of the split_vc method"
 | 
| 33419 |    212 | )
 | 
|  |    213 | *}
 | 
|  |    214 | setup Split_VC_SMT_Rules.setup
 | 
|  |    215 | 
 | 
|  |    216 | use "Tools/boogie_vcs.ML"
 | 
|  |    217 | use "Tools/boogie_loader.ML"
 | 
|  |    218 | use "Tools/boogie_commands.ML"
 | 
|  |    219 | setup Boogie_Commands.setup
 | 
|  |    220 | 
 | 
|  |    221 | use "Tools/boogie_split.ML"
 | 
|  |    222 | setup Boogie_Split.setup
 | 
|  |    223 | 
 | 
|  |    224 | end
 |