| author | wenzelm | 
| Sat, 03 Sep 2011 19:39:16 +0200 | |
| changeset 44674 | bad4f9158c80 | 
| parent 41601 | fda8511006f9 | 
| child 47152 | 446cfc760ccf | 
| permissions | -rw-r--r-- | 
| 33419 | 1 | (* Title: HOL/Boogie/Examples/VCC_Max.thy | 
| 2 | Author: Sascha Boehme, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Boogie example from VCC: get the greatest element of a list *}
 | |
| 6 | ||
| 7 | theory VCC_Max | |
| 8 | imports Boogie | |
| 40540 | 9 | uses ("VCC_Max.b2i")
 | 
| 33419 | 10 | begin | 
| 11 | ||
| 12 | text {*
 | |
| 13 | We prove correct the verification condition generated from the following | |
| 14 | C code: | |
| 15 | ||
| 16 | \begin{verbatim}
 | |
| 17 | #include "vcc.h" | |
| 18 | ||
| 19 | typedef unsigned char byte; | |
| 20 | typedef unsigned long ulong; | |
| 21 | ||
| 22 | static byte maximum(byte arr[], ulong len) | |
| 23 | requires(wrapped(as_array(arr, len))) | |
| 24 | requires (0 < len && len < (1UI64 << 40)) | |
| 25 | ensures (forall(ulong i; i<len ==> arr[i] <= result)) | |
| 26 | ensures (exists(ulong i; i<len && arr[i] == result)) | |
| 27 | {
 | |
| 28 | byte max = arr[0]; | |
| 29 | ulong p; | |
| 30 | spec(ulong witness = 0;) | |
| 31 | ||
| 32 | for (p = 1; p < len; p++) | |
| 33 | invariant(p <= len) | |
| 34 | invariant(forall(ulong i; i < p ==> arr[i] <= max)) | |
| 35 | invariant(witness < len && arr[witness] == max) | |
| 36 |   {
 | |
| 37 | if (arr[p] > max) | |
| 38 |     {
 | |
| 39 | max = arr[p]; | |
| 40 | speconly(witness = p;) | |
| 41 | } | |
| 42 | } | |
| 43 | return max; | |
| 44 | } | |
| 45 | \end{verbatim}
 | |
| 46 | *} | |
| 47 | ||
| 40580 
0592d3a39c08
require the b2i file ending in the boogie_open command (for consistency with the theory header)
 boehmes parents: 
40540diff
changeset | 48 | boogie_open (quiet) "VCC_Max.b2i" | 
| 33419 | 49 | |
| 40513 | 50 | declare [[smt_certificates="VCC_Max.certs"]] | 
| 36081 
70deefb6c093
renamed "smt_record" to "smt_fixed" (somewhat more expressive) and inverted its semantics
 boehmes parents: 
35153diff
changeset | 51 | declare [[smt_fixed=true]] | 
| 41601 | 52 | declare [[smt_oracle=false]] | 
| 34993 
bf3b8462732b
updated examples due to changes in the way SMT certificates are stored
 boehmes parents: 
34068diff
changeset | 53 | |
| 33419 | 54 | boogie_status | 
| 55 | ||
| 33445 
f0c78a28e18e
shorter names for variables and verification conditions,
 boehmes parents: 
33419diff
changeset | 56 | boogie_vc maximum | 
| 34068 
a78307d72e58
make assertion labels unique already when loading a verification condition,
 boehmes parents: 
33445diff
changeset | 57 | by boogie | 
| 33419 | 58 | |
| 59 | boogie_end | |
| 60 | ||
| 35153 
5e8935678ee4
include solver arguments as comments in SMT problem files (to distinguish different results from the same problem when caching results)
 boehmes parents: 
34993diff
changeset | 61 | end |