author | wenzelm |
Fri, 08 Dec 2023 16:01:42 +0100 | |
changeset 79209 | fe24edf8acda |
parent 72990 | db8f94656024 |
permissions | -rw-r--r-- |
13875 | 1 |
(* Title: HOL/Hoare/Pointer_ExamplesAbort.thy |
2 |
Author: Tobias Nipkow |
|
3 |
Copyright 2002 TUM |
|
4 |
*) |
|
5 |
||
72990 | 6 |
section \<open>Examples of verifications of pointer programs\<close> |
13875 | 7 |
|
72990 | 8 |
theory Pointer_ExamplesAbort |
9 |
imports HeapSyntaxAbort |
|
10 |
begin |
|
13875 | 11 |
|
72990 | 12 |
subsection "Verifications" |
13 |
||
14 |
subsubsection "List reversal" |
|
13875 | 15 |
|
16 |
text "Interestingly, this proof is the same as for the unguarded program:" |
|
17 |
||
18 |
lemma "VARS tl p q r |
|
19 |
{List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}} |
|
20 |
WHILE p \<noteq> Null |
|
21 |
INV {\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and> |
|
22 |
rev ps @ qs = rev Ps @ Qs} |
|
23 |
DO r := p; (p \<noteq> Null \<rightarrow> p := p^.tl); r^.tl := q; q := r OD |
|
24 |
{List tl q (rev Ps @ Qs)}" |
|
25 |
apply vcg_simp |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
38353
diff
changeset
|
26 |
apply fastforce |
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
38353
diff
changeset
|
27 |
apply(fastforce intro:notin_List_update[THEN iffD2]) |
13875 | 28 |
done |
29 |
||
30 |
end |