src/HOL/Boogie/Examples/Boogie_Max.thy
author boehmes
Fri, 13 Nov 2009 15:11:41 +0100
changeset 33663 a84fd6385832
parent 33445 f0c78a28e18e
child 34068 a78307d72e58
permissions -rw-r--r--
adapted proofs due to changes in HOL-Boogie
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Boogie/Examples/Boogie_Dijkstra.thy
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     3
*)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     4
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     5
header {* Boogie example: get the greatest element of a list *}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     6
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     7
theory Boogie_Max
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     8
imports Boogie
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
     9
begin
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    10
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    11
text {*
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    12
We prove correct the verification condition generated from the following
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    13
Boogie code:
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    14
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    15
\begin{verbatim}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    16
procedure max(array : [int]int, length : int)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    17
  returns (max : int);
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    18
  requires (0 < length);
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    19
  ensures (forall i: int :: 0 <= i && i < length ==> array[i] <= max);
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    20
  ensures (exists i: int :: 0 <= i && i < length ==> array[i] == max);
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    21
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    22
implementation max(array : [int]int, length : int) returns (max : int)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    23
{
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    24
  var p : int, k : int;
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    25
  max := array[0];
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    26
  k := 0;
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    27
  p := 1;
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    28
  while (p < length)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    29
    invariant (forall i: int :: 0 <= i && i < p ==> array[i] <= max);
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    30
    invariant (array[k] == max);
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    31
  {
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    32
    if (array[p] > max) { max := array[p]; k := p; }
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    33
    p := p + 1;
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    34
  }
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    35
}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    36
\end{verbatim}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    37
*}
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    38
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    39
boogie_open "~~/src/HOL/Boogie/Examples/Boogie_Max"
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    40
33424
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    41
text {*
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    42
Approach 1: Discharge the verification condition fully automatically by SMT:
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    43
*}
33445
f0c78a28e18e shorter names for variables and verification conditions,
boehmes
parents: 33424
diff changeset
    44
boogie_vc max
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    45
  unfolding labels
33445
f0c78a28e18e shorter names for variables and verification conditions,
boehmes
parents: 33424
diff changeset
    46
  using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/Boogie_max"]]
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    47
  using [[z3_proofs=true]]
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    48
  by (smt boogie)
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    49
33424
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    50
text {*
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    51
Approach 2: Split the verification condition, try to solve the splinters by
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    52
a selection of automated proof tools, and show the remaining subgoals by an
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    53
explicit proof. This approach is most useful in an interactive debug-and-fix
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    54
mode. 
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    55
*}
33445
f0c78a28e18e shorter names for variables and verification conditions,
boehmes
parents: 33424
diff changeset
    56
boogie_vc max
33424
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    57
proof (split_vc (verbose) try: fast simp)
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    58
  print_cases
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    59
  case L_14_5c
33663
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    60
  thus ?case by (metis less_le_not_le zle_add1_eq_le zless_linear)
33424
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    61
next
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    62
  case L_14_5b
33663
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    63
  thus ?case by (metis less_le_not_le xt1(10) zle_linear zless_add1_eq)
33424
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    64
next
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    65
  case L_14_5a
33663
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    66
  note max0 = `max0 = array 0`
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    67
  {
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    68
    fix i :: int
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    69
    assume "0 \<le> i \<and> i < 1"
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    70
    hence "i = 0" by simp
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    71
    with max0 have "array i \<le> max0" by simp
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    72
  }
a84fd6385832 adapted proofs due to changes in HOL-Boogie
boehmes
parents: 33445
diff changeset
    73
  thus ?case by simp
33424
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    74
qed
a3b002e2cd55 proper and unique case names for the split_vc method,
boehmes
parents: 33419
diff changeset
    75
33419
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    76
boogie_end
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    77
8ae45e87b992 added HOL-Boogie
boehmes
parents:
diff changeset
    78
end