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