| author | wenzelm | 
| Sat, 23 Feb 2013 22:00:12 +0100 | |
| changeset 51258 | 28b60ee75ef8 | 
| parent 50662 | b1f4291eb916 | 
| 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 | |
| 9 | begin | |
| 10 | ||
| 11 | text {*
 | |
| 12 | We prove correct the verification condition generated from the following | |
| 13 | C code: | |
| 14 | ||
| 15 | \begin{verbatim}
 | |
| 16 | #include "vcc.h" | |
| 17 | ||
| 18 | typedef unsigned char byte; | |
| 19 | typedef unsigned long ulong; | |
| 20 | ||
| 21 | static byte maximum(byte arr[], ulong len) | |
| 22 | requires(wrapped(as_array(arr, len))) | |
| 23 | requires (0 < len && len < (1UI64 << 40)) | |
| 24 | ensures (forall(ulong i; i<len ==> arr[i] <= result)) | |
| 25 | ensures (exists(ulong i; i<len && arr[i] == result)) | |
| 26 | {
 | |
| 27 | byte max = arr[0]; | |
| 28 | ulong p; | |
| 29 | spec(ulong witness = 0;) | |
| 30 | ||
| 31 | for (p = 1; p < len; p++) | |
| 32 | invariant(p <= len) | |
| 33 | invariant(forall(ulong i; i < p ==> arr[i] <= max)) | |
| 34 | invariant(witness < len && arr[witness] == max) | |
| 35 |   {
 | |
| 50662 
b1f4291eb916
regenerate certificates -- and use oracle in a few cases where the Z3 proof parser can't deal with Z3 3.2 proofs
 blanchet parents: 
48907diff
changeset | 36 | if (arr[p] > max) | 
| 33419 | 37 |     {
 | 
| 38 | max = arr[p]; | |
| 39 | speconly(witness = p;) | |
| 40 | } | |
| 41 | } | |
| 42 | return max; | |
| 43 | } | |
| 44 | \end{verbatim}
 | |
| 45 | *} | |
| 46 | ||
| 40580 
0592d3a39c08
require the b2i file ending in the boogie_open command (for consistency with the theory header)
 boehmes parents: 
40540diff
changeset | 47 | boogie_open (quiet) "VCC_Max.b2i" | 
| 33419 | 48 | |
| 50662 
b1f4291eb916
regenerate certificates -- and use oracle in a few cases where the Z3 proof parser can't deal with Z3 3.2 proofs
 blanchet parents: 
48907diff
changeset | 49 | declare [[smt_oracle = false]] | 
| 
b1f4291eb916
regenerate certificates -- and use oracle in a few cases where the Z3 proof parser can't deal with Z3 3.2 proofs
 blanchet parents: 
48907diff
changeset | 50 | declare [[z3_with_extensions = true]] | 
| 47152 
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
 blanchet parents: 
41601diff
changeset | 51 | declare [[smt_certificates = "VCC_Max.certs"]] | 
| 
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
 blanchet parents: 
41601diff
changeset | 52 | declare [[smt_read_only_certificates = true]] | 
| 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 |