| author | haftmann | 
| Fri, 07 May 2010 09:59:24 +0200 | |
| changeset 36725 | 34c36a5cb808 | 
| parent 35316 | 870dfea4f9c0 | 
| child 72990 | db8f94656024 | 
| permissions | -rw-r--r-- | 
| 13875 | 1  | 
(* Title: HOL/Hoare/HeapSyntax.thy  | 
2  | 
Author: Tobias Nipkow  | 
|
3  | 
Copyright 2002 TUM  | 
|
4  | 
*)  | 
|
5  | 
||
| 
35316
 
870dfea4f9c0
dropped axclass; dropped Id; session theory Hoare.thy
 
haftmann 
parents: 
35113 
diff
changeset
 | 
6  | 
theory HeapSyntax imports Hoare_Logic Heap begin  | 
| 13875 | 7  | 
|
8  | 
subsection "Field access and update"  | 
|
9  | 
||
10  | 
syntax  | 
|
| 35101 | 11  | 
  "_refupdate" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a ref \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
 | 
| 13875 | 12  | 
   ("_/'((_ \<rightarrow> _)')" [1000,0] 900)
 | 
| 35101 | 13  | 
"_fassign" :: "'a ref => id => 'v => 's com"  | 
| 13875 | 14  | 
   ("(2_^._ :=/ _)" [70,1000,65] 61)
 | 
| 35101 | 15  | 
  "_faccess"  :: "'a ref => ('a ref \<Rightarrow> 'v) => 'v"
 | 
| 13875 | 16  | 
   ("_^._" [65,1000] 65)
 | 
17  | 
translations  | 
|
| 35113 | 18  | 
"f(r \<rightarrow> v)" == "f(CONST addr r := v)"  | 
19  | 
"p^.f := e" => "f := f(p \<rightarrow> e)"  | 
|
20  | 
"p^.f" => "f(CONST addr p)"  | 
|
| 13875 | 21  | 
|
22  | 
||
23  | 
declare fun_upd_apply[simp del] fun_upd_same[simp] fun_upd_other[simp]  | 
|
24  | 
||
25  | 
||
26  | 
text "An example due to Suzuki:"  | 
|
27  | 
||
28  | 
lemma "VARS v n  | 
|
29  | 
  {w = Ref w0 & x = Ref x0 & y = Ref y0 & z = Ref z0 &
 | 
|
30  | 
distinct[w0,x0,y0,z0]}  | 
|
31  | 
w^.v := (1::int); w^.n := x;  | 
|
32  | 
x^.v := 2; x^.n := y;  | 
|
33  | 
y^.v := 3; y^.n := z;  | 
|
34  | 
z^.v := 4; x^.n := z  | 
|
35  | 
  {w^.n^.n^.v = 4}"
 | 
|
36  | 
by vcg_simp  | 
|
37  | 
||
38  | 
end  |