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