src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy
author wenzelm
Thu, 23 Aug 2012 15:06:15 +0200
changeset 48907 5c4275c3b5b8
parent 44821 a92f65e174cf
permissions -rw-r--r--
turned 'boogie_open' into thy_load command, without any declarations of 'uses';
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Boogie/Examples/Boogie_Max_Stepwise.thy
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     3
*)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     4
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     5
header {* Boogie example: get the greatest element of a list *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     6
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     7
theory Boogie_Max_Stepwise
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     8
imports Boogie
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
     9
begin
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    10
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    11
text {*
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    12
We prove correct the verification condition generated from the following
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    13
Boogie code:
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    14
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    15
\begin{verbatim}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    16
procedure max(array : [int]int, length : int)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    17
  returns (max : int);
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    18
  requires (0 < length);
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    19
  ensures (forall i: int :: 0 <= i && i < length ==> array[i] <= max);
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    20
  ensures (exists i: int :: 0 <= i && i < length ==> array[i] == max);
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    21
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    22
implementation max(array : [int]int, length : int) returns (max : int)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    23
{
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    24
  var p : int, k : int;
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    25
  max := array[0];
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    26
  k := 0;
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    27
  p := 1;
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    28
  while (p < length)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    29
    invariant (forall i: int :: 0 <= i && i < p ==> array[i] <= max);
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    30
    invariant (array[k] == max);
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    31
  {
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    32
    if (array[p] > max) { max := array[p]; k := p; }
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    33
    p := p + 1;
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    34
  }
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    35
}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    36
\end{verbatim}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    37
*}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    38
40580
0592d3a39c08 require the b2i file ending in the boogie_open command (for consistency with the theory header)
boehmes
parents: 40540
diff changeset
    39
boogie_open "Boogie_Max.b2i"
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    40
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    41
boogie_status -- {* gives an overview of all verification conditions *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    42
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    43
boogie_status max -- {* shows the names of all unproved assertions
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    44
  of the verification condition @{text max} *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    45
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    46
boogie_status (full) max -- {* shows the state of all assertions
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    47
  of the verification condition @{text max} *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    48
34182
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    49
boogie_status (paths) max -- {* shows the names of all unproved assertions,
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    50
  grouped by program paths, of the verification condition @{text max} *}
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    51
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    52
boogie_status (full paths) max -- {* shows the state of all assertions
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    53
  of the verification condition @{text max}, grouped by program paths,
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    54
  with already proved or irrelevant assertions written in parentheses *}
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    55
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    56
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    57
text {* Let's find out which assertions of @{text max} are hard to prove: *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    58
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 34182
diff changeset
    59
boogie_status (narrow step_timeout: 4 final_timeout: 15) max
34182
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    60
  (simp add: labels, (fast | metis)?)
35323
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 34182
diff changeset
    61
  -- {* The argument @{text step_timeout} is optional, restricting the runtime
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 34182
diff changeset
    62
        of each intermediate narrowing step by the given number of seconds. *}
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 34182
diff changeset
    63
  -- {* The argument @{text final_timeout} is optional, restricting the
259931828ecc separated narrowing timeouts for intermediate and final steps
boehmes
parents: 34182
diff changeset
    64
        runtime of the final narrowing steps by the given number of seconds. *}
34182
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    65
  -- {* The last argument should be a method to be applied at each narrowing
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    66
        step. Note that different methods may be composed to one method by
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    67
        a language similar to regular expressions. *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    68
34182
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    69
boogie_status (scan timeout: 4) max
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    70
  (simp add: labels, (fast | metis)?)
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    71
  -- {* finds the first hard assertion wrt. to the given method *}
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    72
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    73
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    74
text {* Now, let's prove the three hard assertions of @{text max}: *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    75
34182
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
    76
boogie_vc max (examine: L_14_5c L_14_5b L_14_5a) 
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    77
proof boogie_cases
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    78
  case L_14_5c
44821
a92f65e174cf avoid using legacy theorem names
huffman
parents: 40580
diff changeset
    79
  thus ?case by (metis less_le_not_le zle_add1_eq_le less_linear)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    80
next
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    81
  case L_14_5b
44821
a92f65e174cf avoid using legacy theorem names
huffman
parents: 40580
diff changeset
    82
  thus ?case by (metis less_le_not_le xt1(10) linorder_linear zless_add1_eq)
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    83
next
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    84
  case L_14_5a
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    85
  note max0 = `max0 = array 0`
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    86
  {
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    87
    fix i :: int
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    88
    assume "0 \<le> i \<and> i < 1"
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    89
    hence "i = 0" by simp
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    90
    with max0 have "array i \<le> max0" by simp
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    91
  }
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    92
  thus ?case by simp
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    93
qed
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    94
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    95
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    96
boogie_status max -- {* the above proved assertions are not shown anymore *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    97
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    98
boogie_status (full) max -- {* states the above proved assertions as proved
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
    99
  and all other assertions as not proved *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   100
34182
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
   101
boogie_status (paths) max
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
   102
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
   103
boogie_status (full paths) max
69eddb55588e updated example
boehmes
parents: 34089
diff changeset
   104
34068
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   105
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   106
text {* Now we prove the remaining assertions all at once: *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   107
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   108
boogie_vc max by (auto simp add: labels)
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   109
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   110
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   111
boogie_status -- {* @{text max} has been completely proved *}
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   112
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   113
boogie_end
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   114
a78307d72e58 make assertion labels unique already when loading a verification condition,
boehmes
parents:
diff changeset
   115
end