| author | blanchet | 
| Tue, 01 Jan 2013 11:36:30 +0100 | |
| changeset 50662 | b1f4291eb916 | 
| parent 48907 | 5c4275c3b5b8 | 
| permissions | -rw-r--r-- | 
| 34068 
a78307d72e58
make assertion labels unique already when loading a verification condition,
 boehmes parents: 
33663diff
changeset | 1 | (* Title: HOL/Boogie/Examples/Boogie_Max.thy | 
| 33419 | 2 | Author: Sascha Boehme, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Boogie example: get the greatest element of a list *}
 | |
| 6 | ||
| 7 | theory Boogie_Max | |
| 8 | imports Boogie | |
| 9 | begin | |
| 10 | ||
| 11 | text {*
 | |
| 12 | We prove correct the verification condition generated from the following | |
| 13 | Boogie code: | |
| 14 | ||
| 15 | \begin{verbatim}
 | |
| 16 | procedure max(array : [int]int, length : int) | |
| 17 | returns (max : int); | |
| 18 | requires (0 < length); | |
| 19 | ensures (forall i: int :: 0 <= i && i < length ==> array[i] <= max); | |
| 20 | ensures (exists i: int :: 0 <= i && i < length ==> array[i] == max); | |
| 21 | ||
| 22 | implementation max(array : [int]int, length : int) returns (max : int) | |
| 23 | {
 | |
| 24 | var p : int, k : int; | |
| 25 | max := array[0]; | |
| 26 | k := 0; | |
| 27 | p := 1; | |
| 28 | while (p < length) | |
| 29 | invariant (forall i: int :: 0 <= i && i < p ==> array[i] <= max); | |
| 30 | invariant (array[k] == max); | |
| 31 |   {
 | |
| 32 |     if (array[p] > max) { max := array[p]; k := p; }
 | |
| 33 | p := p + 1; | |
| 34 | } | |
| 35 | } | |
| 36 | \end{verbatim}
 | |
| 37 | *} | |
| 38 | ||
| 40580 
0592d3a39c08
require the b2i file ending in the boogie_open command (for consistency with the theory header)
 boehmes parents: 
40540diff
changeset | 39 | boogie_open "Boogie_Max.b2i" | 
| 33419 | 40 | |
| 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 | 41 | declare [[smt_oracle = false]] | 
| 47152 
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
 blanchet parents: 
41601diff
changeset | 42 | declare [[smt_certificates = "Boogie_Max.certs"]] | 
| 
446cfc760ccf
renamed "smt_fixed" to "smt_read_only_certificates"
 blanchet parents: 
41601diff
changeset | 43 | declare [[smt_read_only_certificates = true]] | 
| 34993 
bf3b8462732b
updated examples due to changes in the way SMT certificates are stored
 boehmes parents: 
34068diff
changeset | 44 | |
| 33445 
f0c78a28e18e
shorter names for variables and verification conditions,
 boehmes parents: 
33424diff
changeset | 45 | boogie_vc max | 
| 34068 
a78307d72e58
make assertion labels unique already when loading a verification condition,
 boehmes parents: 
33663diff
changeset | 46 | by boogie | 
| 33424 
a3b002e2cd55
proper and unique case names for the split_vc method,
 boehmes parents: 
33419diff
changeset | 47 | |
| 33419 | 48 | boogie_end | 
| 49 | ||
| 50 | end |