author | wenzelm |
Mon, 13 Apr 2020 22:08:14 +0200 | |
changeset 71751 | abf3e80bd815 |
parent 67613 | ce654b0e6d69 |
child 72990 | db8f94656024 |
permissions | -rw-r--r-- |
1476 | 1 |
(* Title: HOL/Hoare/Examples.thy |
2 |
Author: Norbert Galm |
|
5646 | 3 |
Copyright 1998 TUM |
67410 | 4 |
*) |
1335 | 5 |
|
67410 | 6 |
chapter \<open>Various examples\<close> |
1335 | 7 |
|
35316
870dfea4f9c0
dropped axclass; dropped Id; session theory Hoare.thy
haftmann
parents:
16796
diff
changeset
|
8 |
theory Examples imports Hoare_Logic Arith2 begin |
13682 | 9 |
|
67410 | 10 |
section \<open>ARITHMETIC\<close> |
13682 | 11 |
|
67410 | 12 |
subsection \<open>multiplication by successive addition\<close> |
13682 | 13 |
|
13737 | 14 |
lemma multiply_by_add: "VARS m s a b |
67613 | 15 |
{a=A \<and> b=B} |
13682 | 16 |
m := 0; s := 0; |
67613 | 17 |
WHILE m\<noteq>a |
18 |
INV {s=m*b \<and> a=A \<and> b=B} |
|
13682 | 19 |
DO s := s+b; m := m+(1::nat) OD |
20 |
{s = A*B}" |
|
21 |
by vcg_simp |
|
22 |
||
63106 | 23 |
lemma multiply_by_add_time: "VARS m s a b t |
67613 | 24 |
{a=A \<and> b=B \<and> t=0} |
63106 | 25 |
m := 0; t := t+1; s := 0; t := t+1; |
67613 | 26 |
WHILE m\<noteq>a |
27 |
INV {s=m*b \<and> a=A \<and> b=B \<and> t = 2*m + 2} |
|
63106 | 28 |
DO s := s+b; t := t+1; m := m+(1::nat); t := t+1 OD |
29 |
{s = A*B \<and> t = 2*A + 2}" |
|
30 |
by vcg_simp |
|
31 |
||
32 |
lemma multiply_by_add2: "VARS M N P :: int |
|
67613 | 33 |
{m=M \<and> n=N} |
13789 | 34 |
IF M < 0 THEN M := -M; N := -N ELSE SKIP FI; |
35 |
P := 0; |
|
36 |
WHILE 0 < M |
|
67613 | 37 |
INV {0 \<le> M \<and> (\<exists>p. p = (if m<0 then -m else m) & p*N = m*n & P = (p-M)*N)} |
13789 | 38 |
DO P := P+N; M := M - 1 OD |
39 |
{P = m*n}" |
|
40 |
apply vcg_simp |
|
63106 | 41 |
apply (auto simp add:int_distrib) |
42 |
done |
|
43 |
||
44 |
lemma multiply_by_add2_time: "VARS M N P t :: int |
|
67613 | 45 |
{m=M \<and> n=N \<and> t=0} |
63106 | 46 |
IF M < 0 THEN M := -M; t := t+1; N := -N; t := t+1 ELSE SKIP FI; |
47 |
P := 0; t := t+1; |
|
48 |
WHILE 0 < M |
|
67613 | 49 |
INV {0 \<le> M & (\<exists>p. p = (if m<0 then -m else m) & p*N = m*n & P = (p-M)*N & t \<ge> 0 & t \<le> 2*(p-M)+3)} |
63106 | 50 |
DO P := P+N; t := t+1; M := M - 1; t := t+1 OD |
51 |
{P = m*n & t \<le> 2*abs m + 3}" |
|
52 |
apply vcg_simp |
|
53 |
apply (auto simp add:int_distrib) |
|
13789 | 54 |
done |
13682 | 55 |
|
67410 | 56 |
|
57 |
subsection \<open>Euclid's algorithm for GCD\<close> |
|
13682 | 58 |
|
13737 | 59 |
lemma Euclid_GCD: "VARS a b |
13682 | 60 |
{0<A & 0<B} |
61 |
a := A; b := B; |
|
13857 | 62 |
WHILE a \<noteq> b |
13682 | 63 |
INV {0<a & 0<b & gcd A B = gcd a b} |
64 |
DO IF a<b THEN b := b-a ELSE a := a-b FI OD |
|
65 |
{a = gcd A B}" |
|
66 |
apply vcg |
|
67410 | 67 |
\<comment> \<open>Now prove the verification conditions\<close> |
13682 | 68 |
apply auto |
69 |
apply(simp add: gcd_diff_r less_imp_le) |
|
16796 | 70 |
apply(simp add: linorder_not_less gcd_diff_l) |
13682 | 71 |
apply(erule gcd_nnn) |
72 |
done |
|
73 |
||
63106 | 74 |
lemma Euclid_GCD_time: "VARS a b t |
75 |
{0<A & 0<B & t=0} |
|
76 |
a := A; t := t+1; b := B; t := t+1; |
|
77 |
WHILE a \<noteq> b |
|
78 |
INV {0<a & 0<b & gcd A B = gcd a b & a\<le>A & b\<le>B & t \<le> max A B - max a b + 2} |
|
79 |
DO IF a<b THEN b := b-a; t := t+1 ELSE a := a-b; t := t+1 FI OD |
|
80 |
{a = gcd A B & t \<le> max A B + 2}" |
|
81 |
apply vcg |
|
67410 | 82 |
\<comment> \<open>Now prove the verification conditions\<close> |
63106 | 83 |
apply auto |
84 |
apply(simp add: gcd_diff_r less_imp_le) |
|
85 |
apply(simp add: linorder_not_less gcd_diff_l) |
|
86 |
apply(erule gcd_nnn) |
|
87 |
done |
|
88 |
||
67410 | 89 |
|
90 |
subsection \<open>Dijkstra's extension of Euclid's algorithm for simultaneous GCD and SCM\<close> |
|
91 |
||
92 |
text \<open> |
|
93 |
From E.W. Disjkstra. Selected Writings on Computing, p 98 (EWD474), |
|
94 |
where it is given without the invariant. Instead of defining \<open>scm\<close> |
|
95 |
explicitly we have used the theorem \<open>scm x y = x * y / gcd x y\<close> and avoided |
|
96 |
division by mupltiplying with \<open>gcd x y\<close>. |
|
97 |
\<close> |
|
13682 | 98 |
|
99 |
lemmas distribs = |
|
100 |
diff_mult_distrib diff_mult_distrib2 add_mult_distrib add_mult_distrib2 |
|
101 |
||
13737 | 102 |
lemma gcd_scm: "VARS a b x y |
13682 | 103 |
{0<A & 0<B & a=A & b=B & x=B & y=A} |
104 |
WHILE a ~= b |
|
105 |
INV {0<a & 0<b & gcd A B = gcd a b & 2*A*B = a*x + b*y} |
|
106 |
DO IF a<b THEN (b := b-a; x := x+y) ELSE (a := a-b; y := y+x) FI OD |
|
107 |
{a = gcd A B & 2*A*B = a*(x+y)}" |
|
108 |
apply vcg |
|
109 |
apply simp |
|
16796 | 110 |
apply(simp add: distribs gcd_diff_r linorder_not_less gcd_diff_l) |
13682 | 111 |
apply(simp add: distribs gcd_nnn) |
112 |
done |
|
113 |
||
67410 | 114 |
|
115 |
subsection \<open>Power by iterated squaring and multiplication\<close> |
|
13682 | 116 |
|
13737 | 117 |
lemma power_by_mult: "VARS a b c |
13682 | 118 |
{a=A & b=B} |
119 |
c := (1::nat); |
|
120 |
WHILE b ~= 0 |
|
121 |
INV {A^B = c * a^b} |
|
122 |
DO WHILE b mod 2 = 0 |
|
123 |
INV {A^B = c * a^b} |
|
124 |
DO a := a*a; b := b div 2 OD; |
|
125 |
c := c*a; b := b - 1 |
|
126 |
OD |
|
127 |
{c = A^B}" |
|
128 |
apply vcg_simp |
|
129 |
apply(case_tac "b") |
|
42154 | 130 |
apply simp |
13682 | 131 |
apply simp |
132 |
done |
|
133 |
||
67410 | 134 |
|
135 |
subsection \<open>Factorial\<close> |
|
13682 | 136 |
|
13737 | 137 |
lemma factorial: "VARS a b |
13682 | 138 |
{a=A} |
139 |
b := 1; |
|
63106 | 140 |
WHILE a > 0 |
13682 | 141 |
INV {fac A = b * fac a} |
142 |
DO b := b*a; a := a - 1 OD |
|
143 |
{b = fac A}" |
|
144 |
apply vcg_simp |
|
145 |
apply(clarsimp split: nat_diff_split) |
|
146 |
done |
|
147 |
||
63106 | 148 |
lemma factorial_time: "VARS a b t |
149 |
{a=A & t=0} |
|
150 |
b := 1; t := t+1; |
|
151 |
WHILE a > 0 |
|
152 |
INV {fac A = b * fac a & a \<le> A & t = 2*(A-a)+1} |
|
153 |
DO b := b*a; t := t+1; a := a - 1; t := t+1 OD |
|
154 |
{b = fac A & t = 2*A + 1}" |
|
155 |
apply vcg_simp |
|
156 |
apply(clarsimp split: nat_diff_split) |
|
157 |
done |
|
158 |
||
13684 | 159 |
lemma [simp]: "1 \<le> i \<Longrightarrow> fac (i - Suc 0) * i = fac i" |
160 |
by(induct i, simp_all) |
|
161 |
||
63106 | 162 |
lemma factorial2: "VARS i f |
13684 | 163 |
{True} |
164 |
i := (1::nat); f := 1; |
|
165 |
WHILE i <= n INV {f = fac(i - 1) & 1 <= i & i <= n+1} |
|
166 |
DO f := f*i; i := i+1 OD |
|
167 |
{f = fac n}" |
|
168 |
apply vcg_simp |
|
169 |
apply(subgoal_tac "i = Suc n") |
|
170 |
apply simp |
|
171 |
apply arith |
|
172 |
done |
|
13682 | 173 |
|
63106 | 174 |
lemma factorial2_time: "VARS i f t |
175 |
{t=0} |
|
176 |
i := (1::nat); t := t+1; f := 1; t := t+1; |
|
177 |
WHILE i \<le> n INV {f = fac(i - 1) & 1 \<le> i & i \<le> n+1 & t = 2*(i-1)+2} |
|
178 |
DO f := f*i; t := t+1; i := i+1; t := t+1 OD |
|
179 |
{f = fac n & t = 2*n+2}" |
|
180 |
apply vcg_simp |
|
181 |
apply auto |
|
182 |
apply(subgoal_tac "i = Suc n") |
|
183 |
apply simp |
|
184 |
apply arith |
|
185 |
done |
|
186 |
||
13682 | 187 |
|
67410 | 188 |
subsection \<open>Square root\<close> |
189 |
||
190 |
\<comment> \<open>the easy way:\<close> |
|
13682 | 191 |
|
13737 | 192 |
lemma sqrt: "VARS r x |
13682 | 193 |
{True} |
63106 | 194 |
r := (0::nat); |
195 |
WHILE (r+1)*(r+1) <= X |
|
196 |
INV {r*r \<le> X} |
|
13682 | 197 |
DO r := r+1 OD |
198 |
{r*r <= X & X < (r+1)*(r+1)}" |
|
199 |
apply vcg_simp |
|
200 |
done |
|
201 |
||
63106 | 202 |
lemma sqrt_time: "VARS r t |
203 |
{t=0} |
|
204 |
r := (0::nat); t := t+1; |
|
205 |
WHILE (r+1)*(r+1) <= X |
|
206 |
INV {r*r \<le> X & t = r+1} |
|
207 |
DO r := r+1; t := t+1 OD |
|
208 |
{r*r <= X & X < (r+1)*(r+1) & (t-1)*(t-1) \<le> X}" |
|
209 |
apply vcg_simp |
|
210 |
done |
|
211 |
||
67410 | 212 |
\<comment> \<open>without multiplication\<close> |
63106 | 213 |
lemma sqrt_without_multiplication: "VARS u w r |
214 |
{x=X} |
|
215 |
u := 1; w := 1; r := (0::nat); |
|
216 |
WHILE w <= X |
|
217 |
INV {u = r+r+1 & w = (r+1)*(r+1) & r*r <= X} |
|
13682 | 218 |
DO r := r + 1; w := w + u + 2; u := u + 2 OD |
219 |
{r*r <= X & X < (r+1)*(r+1)}" |
|
220 |
apply vcg_simp |
|
221 |
done |
|
222 |
||
223 |
||
67410 | 224 |
section \<open>LISTS\<close> |
13682 | 225 |
|
13737 | 226 |
lemma imperative_reverse: "VARS y x |
13682 | 227 |
{x=X} |
228 |
y:=[]; |
|
229 |
WHILE x ~= [] |
|
230 |
INV {rev(x)@y = rev(X)} |
|
231 |
DO y := (hd x # y); x := tl x OD |
|
232 |
{y=rev(X)}" |
|
233 |
apply vcg_simp |
|
234 |
apply(simp add: neq_Nil_conv) |
|
235 |
apply auto |
|
236 |
done |
|
237 |
||
63106 | 238 |
lemma imperative_reverse_time: "VARS y x t |
239 |
{x=X & t=0} |
|
240 |
y:=[]; t := t+1; |
|
241 |
WHILE x ~= [] |
|
242 |
INV {rev(x)@y = rev(X) & t = 2*(length y) + 1} |
|
243 |
DO y := (hd x # y); t := t+1; x := tl x; t := t+1 OD |
|
244 |
{y=rev(X) & t = 2*length X + 1}" |
|
245 |
apply vcg_simp |
|
246 |
apply(simp add: neq_Nil_conv) |
|
247 |
apply auto |
|
248 |
done |
|
249 |
||
13737 | 250 |
lemma imperative_append: "VARS x y |
13682 | 251 |
{x=X & y=Y} |
252 |
x := rev(x); |
|
253 |
WHILE x~=[] |
|
254 |
INV {rev(x)@y = X@Y} |
|
255 |
DO y := (hd x # y); |
|
256 |
x := tl x |
|
257 |
OD |
|
258 |
{y = X@Y}" |
|
259 |
apply vcg_simp |
|
260 |
apply(simp add: neq_Nil_conv) |
|
261 |
apply auto |
|
262 |
done |
|
263 |
||
63106 | 264 |
lemma imperative_append_time_no_rev: "VARS x y t |
265 |
{x=X & y=Y} |
|
266 |
x := rev(x); t := 0; |
|
267 |
WHILE x~=[] |
|
268 |
INV {rev(x)@y = X@Y & length x \<le> length X & t = 2 * (length X - length x)} |
|
269 |
DO y := (hd x # y); t := t+1; |
|
270 |
x := tl x; t := t+1 |
|
271 |
OD |
|
272 |
{y = X@Y & t = 2 * length X}" |
|
273 |
apply vcg_simp |
|
274 |
apply(simp add: neq_Nil_conv) |
|
275 |
apply auto |
|
276 |
done |
|
277 |
||
13682 | 278 |
|
67410 | 279 |
section \<open>ARRAYS\<close> |
13682 | 280 |
|
67410 | 281 |
subsection \<open>Search for a key\<close> |
282 |
||
13737 | 283 |
lemma zero_search: "VARS A i |
13682 | 284 |
{True} |
285 |
i := 0; |
|
67613 | 286 |
WHILE i < length A & A!i \<noteq> key |
287 |
INV {\<forall>j. j<i --> A!j \<noteq> key} |
|
13682 | 288 |
DO i := i+1 OD |
289 |
{(i < length A --> A!i = key) & |
|
67613 | 290 |
(i = length A --> (\<forall>j. j < length A \<longrightarrow> A!j \<noteq> key))}" |
13682 | 291 |
apply vcg_simp |
292 |
apply(blast elim!: less_SucE) |
|
293 |
done |
|
294 |
||
63106 | 295 |
lemma zero_search_time: "VARS A i t |
296 |
{t=0} |
|
297 |
i := 0; t := t+1; |
|
67613 | 298 |
WHILE i < length A \<and> A!i \<noteq> key |
299 |
INV {(\<forall>j. j<i \<longrightarrow> A!j \<noteq> key) \<and> i \<le> length A \<and> t = i+1} |
|
63106 | 300 |
DO i := i+1; t := t+1 OD |
67613 | 301 |
{(i < length A \<longrightarrow> A!i = key) \<and> |
302 |
(i = length A \<longrightarrow> (\<forall>j. j < length A --> A!j \<noteq> key)) \<and> t \<le> length A + 1}" |
|
63106 | 303 |
apply vcg_simp |
304 |
apply(blast elim!: less_SucE) |
|
305 |
done |
|
306 |
||
67410 | 307 |
text \<open> |
308 |
The \<open>partition\<close> procedure for quicksort. |
|
309 |
\<^item> \<open>A\<close> is the array to be sorted (modelled as a list). |
|
310 |
\<^item> Elements of \<open>A\<close> must be of class order to infer at the end |
|
311 |
that the elements between u and l are equal to pivot. |
|
13682 | 312 |
|
67410 | 313 |
Ambiguity warnings of parser are due to \<open>:=\<close> being used |
314 |
both for assignment and list update. |
|
315 |
\<close> |
|
13682 | 316 |
lemma lem: "m - Suc 0 < n ==> m < Suc n" |
317 |
by arith |
|
318 |
||
319 |
||
320 |
lemma Partition: |
|
67613 | 321 |
"[| leq == \<lambda>A i. \<forall>k. k<i \<longrightarrow> A!k \<le> pivot; |
322 |
geq == \<lambda>A i. \<forall>k. i<k \<and> k<length A \<longrightarrow> pivot \<le> A!k |] ==> |
|
13737 | 323 |
VARS A u l |
13682 | 324 |
{0 < length(A::('a::order)list)} |
325 |
l := 0; u := length A - Suc 0; |
|
67613 | 326 |
WHILE l \<le> u |
327 |
INV {leq A l \<and> geq A u \<and> u<length A \<and> l\<le>length A} |
|
328 |
DO WHILE l < length A \<and> A!l \<le> pivot |
|
329 |
INV {leq A l & geq A u \<and> u<length A \<and> l\<le>length A} |
|
13682 | 330 |
DO l := l+1 OD; |
67613 | 331 |
WHILE 0 < u & pivot \<le> A!u |
332 |
INV {leq A l & geq A u \<and> u<length A \<and> l\<le>length A} |
|
13682 | 333 |
DO u := u - 1 OD; |
67613 | 334 |
IF l \<le> u THEN A := A[l := A!u, u := A!l] ELSE SKIP FI |
13682 | 335 |
OD |
67613 | 336 |
{leq A u & (\<forall>k. u<k \<and> k<l --> A!k = pivot) \<and> geq A l}" |
67410 | 337 |
\<comment> \<open>expand and delete abbreviations first\<close> |
58860 | 338 |
apply (simp) |
13682 | 339 |
apply (erule thin_rl)+ |
340 |
apply vcg_simp |
|
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
341 |
apply (force simp: neq_Nil_conv) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
342 |
apply (blast elim!: less_SucE intro: Suc_leI) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
343 |
apply (blast elim!: less_SucE intro: less_imp_diff_less dest: lem) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
344 |
apply (force simp: nth_list_update) |
13682 | 345 |
done |
346 |
||
62390 | 347 |
end |