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