src/HOL/Hoare/ExamplesTC.thy
author wenzelm
Fri, 08 Dec 2023 16:01:42 +0100
changeset 79209 fe24edf8acda
parent 74505 ce8152fb021b
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
72807
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     1
(* Title:      Examples using Hoare Logic for Total Correctness
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     2
   Author:     Walter Guttmann
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     3
*)
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     4
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     5
section \<open>Examples using Hoare Logic for Total Correctness\<close>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     6
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     7
theory ExamplesTC
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 72807
diff changeset
     8
  imports Hoare_Logic
72807
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
     9
begin
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    10
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    11
text \<open>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    12
This theory demonstrates a few simple partial- and total-correctness proofs.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    13
The first example is taken from HOL/Hoare/Examples.thy written by N. Galm.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    14
We have added the invariant \<open>m \<le> a\<close>.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    15
\<close>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    16
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    17
lemma multiply_by_add: "VARS m s a b
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    18
  {a=A \<and> b=B}
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    19
  m := 0; s := 0;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    20
  WHILE m\<noteq>a
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    21
  INV {s=m*b \<and> a=A \<and> b=B \<and> m\<le>a}
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    22
  DO s := s+b; m := m+(1::nat) OD
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    23
  {s = A*B}"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    24
  by vcg_simp
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    25
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    26
text \<open>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    27
Here is the total-correctness proof for the same program.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    28
It needs the additional invariant \<open>m \<le> a\<close>.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    29
\<close>
74505
ce8152fb021b removed junk;
wenzelm
parents: 74503
diff changeset
    30
72807
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    31
lemma multiply_by_add_tc: "VARS m s a b
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    32
  [a=A \<and> b=B]
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    33
  m := 0; s := 0;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    34
  WHILE m\<noteq>a
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    35
  INV {s=m*b \<and> a=A \<and> b=B \<and> m\<le>a}
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    36
  VAR {a-m}
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    37
  DO s := s+b; m := m+(1::nat) OD
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    38
  [s = A*B]"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    39
  apply vcg_tc_simp
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    40
  by auto
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    41
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    42
text \<open>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    43
Next, we prove partial correctness of a program that computes powers.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    44
\<close>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    45
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    46
lemma power: "VARS (p::int) i
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    47
  { True }
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    48
  p := 1;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    49
  i := 0;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    50
  WHILE i < n
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    51
    INV { p = x^i \<and> i \<le> n }
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    52
     DO p := p * x;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    53
        i := i + 1
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    54
     OD
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    55
  { p = x^n }"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    56
  apply vcg_simp
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    57
  by auto
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    58
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    59
text \<open>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    60
Here is its total-correctness proof.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    61
\<close>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    62
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    63
lemma power_tc: "VARS (p::int) i
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    64
  [ True ]
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    65
  p := 1;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    66
  i := 0;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    67
  WHILE i < n
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    68
    INV { p = x^i \<and> i \<le> n }
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    69
    VAR { n - i }
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    70
     DO p := p * x;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    71
        i := i + 1
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    72
     OD
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    73
  [ p = x^n ]"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    74
  apply vcg_tc
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    75
  by auto
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    76
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    77
text \<open>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    78
The last example is again taken from HOL/Hoare/Examples.thy.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    79
We have modified it to integers so it requires precondition \<open>0 \<le> x\<close>.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    80
\<close>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    81
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    82
lemma sqrt_tc: "VARS r
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    83
  [0 \<le> (x::int)]
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    84
  r := 0;
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    85
  WHILE (r+1)*(r+1) <= x
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    86
  INV {r*r \<le> x}
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
    87
  VAR { nat (x-r)}
72807
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    88
  DO r := r+1 OD
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    89
  [r*r \<le> x \<and> x < (r+1)*(r+1)]"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    90
  apply vcg_tc_simp
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    91
  by (smt (verit) div_pos_pos_trivial mult_less_0_iff nonzero_mult_div_cancel_left)
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    92
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    93
text \<open>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    94
A total-correctness proof allows us to extract a function for further use.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    95
For every input satisfying the precondition the function returns an output satisfying the postcondition.
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    96
\<close>
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    97
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    98
lemma sqrt_exists:
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
    99
  "0 \<le> (x::int) \<Longrightarrow> \<exists>r' . r'*r' \<le> x \<and> x < (r'+1)*(r'+1)"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   100
  using tc_extract_function sqrt_tc by blast
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   101
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   102
definition "sqrt (x::int) \<equiv> (SOME r' . r'*r' \<le> x \<and> x < (r'+1)*(r'+1))"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   103
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   104
lemma sqrt_function:
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   105
  assumes "0 \<le> (x::int)"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   106
    and "r' = sqrt x"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   107
  shows "r'*r' \<le> x \<and> x < (r'+1)*(r'+1)"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   108
proof -
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   109
  let ?P = "\<lambda>r' . r'*r' \<le> x \<and> x < (r'+1)*(r'+1)"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   110
  have "?P (SOME z . ?P z)"
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   111
    by (metis (mono_tags, lifting) assms(1) sqrt_exists some_eq_imp)
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   112
  thus ?thesis
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   113
    using assms(2) sqrt_def by auto
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   114
qed
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   115
74503
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   116
text \<open>Nested loops!\<close>
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   117
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   118
lemma "VARS (i::nat) j
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   119
  [ True ]
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   120
  WHILE 0 < i
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   121
    INV { True }
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   122
    VAR { z = i }
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   123
     DO i := i - 1; j := i;
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   124
        WHILE 0 < j
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   125
        INV { z = i+1 }
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   126
        VAR { j }
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   127
        DO j := j - 1 OD
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   128
     OD
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   129
  [ i \<le> 0 ]"
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   130
  apply vcg_tc
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   131
  by auto
403ce50e6a2a separated commands from annotations to be able to abstract about the latter only
nipkow
parents: 72990
diff changeset
   132
72807
ea189da0ff60 Total correctness examples by Walter Guttmann
nipkow
parents:
diff changeset
   133
end