author | paulson <lp15@cam.ac.uk> |
Thu, 26 Sep 2024 14:44:37 +0100 | |
changeset 80948 | 572970d15ab0 |
parent 80914 | d97fdabd9e2b |
child 81189 | 47a0dfee26ea |
permissions | -rw-r--r-- |
13875 | 1 |
(* Title: HOL/Hoare/HeapSyntax.thy |
2 |
Author: Tobias Nipkow |
|
3 |
Copyright 2002 TUM |
|
4 |
*) |
|
5 |
||
72990 | 6 |
section \<open>Heap syntax\<close> |
7 |
||
8 |
theory HeapSyntax |
|
9 |
imports Hoare_Logic Heap |
|
10 |
begin |
|
13875 | 11 |
|
12 |
subsection "Field access and update" |
|
13 |
||
14 |
syntax |
|
35101 | 15 |
"_refupdate" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a ref \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
72990
diff
changeset
|
16 |
(\<open>_/'((_ \<rightarrow> _)')\<close> [1000,0] 900) |
35101 | 17 |
"_fassign" :: "'a ref => id => 'v => 's com" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
72990
diff
changeset
|
18 |
(\<open>(2_^._ :=/ _)\<close> [70,1000,65] 61) |
35101 | 19 |
"_faccess" :: "'a ref => ('a ref \<Rightarrow> 'v) => 'v" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
72990
diff
changeset
|
20 |
(\<open>_^._\<close> [65,1000] 65) |
13875 | 21 |
translations |
35113 | 22 |
"f(r \<rightarrow> v)" == "f(CONST addr r := v)" |
23 |
"p^.f := e" => "f := f(p \<rightarrow> e)" |
|
24 |
"p^.f" => "f(CONST addr p)" |
|
13875 | 25 |
|
26 |
||
27 |
declare fun_upd_apply[simp del] fun_upd_same[simp] fun_upd_other[simp] |
|
28 |
||
29 |
||
30 |
text "An example due to Suzuki:" |
|
31 |
||
32 |
lemma "VARS v n |
|
33 |
{w = Ref w0 & x = Ref x0 & y = Ref y0 & z = Ref z0 & |
|
34 |
distinct[w0,x0,y0,z0]} |
|
35 |
w^.v := (1::int); w^.n := x; |
|
36 |
x^.v := 2; x^.n := y; |
|
37 |
y^.v := 3; y^.n := z; |
|
38 |
z^.v := 4; x^.n := z |
|
39 |
{w^.n^.n^.v = 4}" |
|
40 |
by vcg_simp |
|
41 |
||
42 |
end |