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