| author | wenzelm | 
| Sat, 20 Jan 2024 16:23:51 +0100 | |
| changeset 79504 | 958d7b118c7b | 
| parent 72990 | db8f94656024 | 
| permissions | -rw-r--r-- | 
| 13875 | 1  | 
(* Title: HOL/Hoare/Heap.thy  | 
2  | 
Author: Tobias Nipkow  | 
|
3  | 
Copyright 2002 TUM  | 
|
4  | 
*)  | 
|
5  | 
||
| 72990 | 6  | 
section \<open>Pointers, heaps and heap abstractions\<close>  | 
7  | 
||
8  | 
text \<open>See the paper by Mehta and Nipkow.\<close>  | 
|
9  | 
||
10  | 
theory Heap  | 
|
11  | 
imports Main  | 
|
12  | 
begin  | 
|
| 13875 | 13  | 
|
14  | 
subsection "References"  | 
|
15  | 
||
| 58310 | 16  | 
datatype 'a ref = Null | Ref 'a  | 
| 13875 | 17  | 
|
| 67613 | 18  | 
lemma not_Null_eq [iff]: "(x \<noteq> Null) = (\<exists>y. x = Ref y)"  | 
| 13875 | 19  | 
by (induct x) auto  | 
20  | 
||
| 67613 | 21  | 
lemma not_Ref_eq [iff]: "(\<forall>y. x \<noteq> Ref y) = (x = Null)"  | 
| 13875 | 22  | 
by (induct x) auto  | 
23  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
20503 
diff
changeset
 | 
24  | 
primrec addr :: "'a ref \<Rightarrow> 'a" where  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
20503 
diff
changeset
 | 
25  | 
"addr (Ref a) = a"  | 
| 13875 | 26  | 
|
27  | 
||
| 72990 | 28  | 
subsection "The heap"  | 
| 13875 | 29  | 
|
| 72990 | 30  | 
subsubsection "Paths in the heap"  | 
| 13875 | 31  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
20503 
diff
changeset
 | 
32  | 
primrec Path :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool" where
 | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
20503 
diff
changeset
 | 
33  | 
"Path h x [] y \<longleftrightarrow> x = y"  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
20503 
diff
changeset
 | 
34  | 
| "Path h x (a#as) y \<longleftrightarrow> x = Ref a \<and> Path h (h a) as y"  | 
| 13875 | 35  | 
|
36  | 
lemma [iff]: "Path h Null xs y = (xs = [] \<and> y = Null)"  | 
|
37  | 
apply(case_tac xs)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
38  | 
apply fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
39  | 
apply fastforce  | 
| 13875 | 40  | 
done  | 
41  | 
||
42  | 
lemma [simp]: "Path h (Ref a) as z =  | 
|
43  | 
(as = [] \<and> z = Ref a \<or> (\<exists>bs. as = a#bs \<and> Path h (h a) bs z))"  | 
|
44  | 
apply(case_tac as)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
45  | 
apply fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
46  | 
apply fastforce  | 
| 13875 | 47  | 
done  | 
48  | 
||
49  | 
lemma [simp]: "\<And>x. Path f x (as@bs) z = (\<exists>y. Path f x as y \<and> Path f y bs z)"  | 
|
50  | 
by(induct as, simp+)  | 
|
51  | 
||
52  | 
lemma Path_upd[simp]:  | 
|
53  | 
"\<And>x. u \<notin> set as \<Longrightarrow> Path (f(u := v)) x as y = Path f x as y"  | 
|
54  | 
by(induct as, simp, simp add:eq_sym_conv)  | 
|
55  | 
||
56  | 
lemma Path_snoc:  | 
|
57  | 
"Path (f(a := q)) p as (Ref a) \<Longrightarrow> Path (f(a := q)) p (as @ [a]) q"  | 
|
58  | 
by simp  | 
|
59  | 
||
60  | 
||
| 72990 | 61  | 
subsubsection "Non-repeating paths"  | 
| 19399 | 62  | 
|
| 38353 | 63  | 
definition distPath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool"
 | 
64  | 
where "distPath h x as y \<longleftrightarrow> Path h x as y \<and> distinct as"  | 
|
| 19399 | 65  | 
|
| 69597 | 66  | 
text\<open>The term \<^term>\<open>distPath h x as y\<close> expresses the fact that a  | 
67  | 
non-repeating path \<^term>\<open>as\<close> connects location \<^term>\<open>x\<close> to location  | 
|
68  | 
\<^term>\<open>y\<close> by means of the \<^term>\<open>h\<close> field. In the case where \<open>x  | 
|
69  | 
= y\<close>, and there is a cycle from \<^term>\<open>x\<close> to itself, \<^term>\<open>as\<close> can  | 
|
70  | 
be both \<^term>\<open>[]\<close> and the non-repeating list of nodes in the  | 
|
| 62042 | 71  | 
cycle.\<close>  | 
| 19399 | 72  | 
|
73  | 
lemma neq_dP: "p \<noteq> q \<Longrightarrow> Path h p Ps q \<Longrightarrow> distinct Ps \<Longrightarrow>  | 
|
| 67613 | 74  | 
\<exists>a Qs. p = Ref a \<and> Ps = a#Qs \<and> a \<notin> set Qs"  | 
| 19399 | 75  | 
by (case_tac Ps, auto)  | 
76  | 
||
77  | 
||
78  | 
lemma neq_dP_disp: "\<lbrakk> p \<noteq> q; distPath h p Ps q \<rbrakk> \<Longrightarrow>  | 
|
| 67613 | 79  | 
\<exists>a Qs. p = Ref a \<and> Ps = a#Qs \<and> a \<notin> set Qs"  | 
| 19399 | 80  | 
apply (simp only:distPath_def)  | 
81  | 
by (case_tac Ps, auto)  | 
|
82  | 
||
83  | 
||
| 72990 | 84  | 
subsubsection "Lists on the heap"  | 
| 13875 | 85  | 
|
| 72990 | 86  | 
paragraph "Relational abstraction"  | 
| 13875 | 87  | 
|
| 38353 | 88  | 
definition List :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
89  | 
where "List h x as = Path h x as Null"  | 
|
| 13875 | 90  | 
|
91  | 
lemma [simp]: "List h x [] = (x = Null)"  | 
|
92  | 
by(simp add:List_def)  | 
|
93  | 
||
94  | 
lemma [simp]: "List h x (a#as) = (x = Ref a \<and> List h (h a) as)"  | 
|
95  | 
by(simp add:List_def)  | 
|
96  | 
||
97  | 
lemma [simp]: "List h Null as = (as = [])"  | 
|
98  | 
by(case_tac as, simp_all)  | 
|
99  | 
||
100  | 
lemma List_Ref[simp]: "List h (Ref a) as = (\<exists>bs. as = a#bs \<and> List h (h a) bs)"  | 
|
101  | 
by(case_tac as, simp_all, fast)  | 
|
102  | 
||
103  | 
theorem notin_List_update[simp]:  | 
|
104  | 
"\<And>x. a \<notin> set as \<Longrightarrow> List (h(a := y)) x as = List h x as"  | 
|
105  | 
apply(induct as)  | 
|
106  | 
apply simp  | 
|
107  | 
apply(clarsimp simp add:fun_upd_apply)  | 
|
108  | 
done  | 
|
109  | 
||
110  | 
lemma List_unique: "\<And>x bs. List h x as \<Longrightarrow> List h x bs \<Longrightarrow> as = bs"  | 
|
111  | 
by(induct as, simp, clarsimp)  | 
|
112  | 
||
113  | 
lemma List_unique1: "List h p as \<Longrightarrow> \<exists>!as. List h p as"  | 
|
114  | 
by(blast intro:List_unique)  | 
|
115  | 
||
116  | 
lemma List_app: "\<And>x. List h x (as@bs) = (\<exists>y. Path h x as y \<and> List h y bs)"  | 
|
117  | 
by(induct as, simp, clarsimp)  | 
|
118  | 
||
119  | 
lemma List_hd_not_in_tl[simp]: "List h (h a) as \<Longrightarrow> a \<notin> set as"  | 
|
120  | 
apply (clarsimp simp add:in_set_conv_decomp)  | 
|
121  | 
apply(frule List_app[THEN iffD1])  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
122  | 
apply(fastforce dest: List_unique)  | 
| 13875 | 123  | 
done  | 
124  | 
||
125  | 
lemma List_distinct[simp]: "\<And>x. List h x as \<Longrightarrow> distinct as"  | 
|
126  | 
apply(induct as, simp)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
127  | 
apply(fastforce dest:List_hd_not_in_tl)  | 
| 13875 | 128  | 
done  | 
129  | 
||
| 19399 | 130  | 
lemma Path_is_List:  | 
131  | 
"\<lbrakk>Path h b Ps (Ref a); a \<notin> set Ps\<rbrakk> \<Longrightarrow> List (h(a := Null)) b (Ps @ [a])"  | 
|
| 20503 | 132  | 
apply (induct Ps arbitrary: b)  | 
| 19399 | 133  | 
apply (auto simp add:fun_upd_apply)  | 
134  | 
done  | 
|
135  | 
||
136  | 
||
| 72990 | 137  | 
subsubsection "Functional abstraction"  | 
| 13875 | 138  | 
|
| 38353 | 139  | 
definition islist :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> bool"
 | 
140  | 
where "islist h p \<longleftrightarrow> (\<exists>as. List h p as)"  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
20503 
diff
changeset
 | 
141  | 
|
| 38353 | 142  | 
definition list :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list"
 | 
143  | 
where "list h p = (SOME as. List h p as)"  | 
|
| 13875 | 144  | 
|
145  | 
lemma List_conv_islist_list: "List h p as = (islist h p \<and> as = list h p)"  | 
|
146  | 
apply(simp add:islist_def list_def)  | 
|
147  | 
apply(rule iffI)  | 
|
148  | 
apply(rule conjI)  | 
|
149  | 
apply blast  | 
|
150  | 
apply(subst some1_equality)  | 
|
151  | 
apply(erule List_unique1)  | 
|
152  | 
apply assumption  | 
|
153  | 
apply(rule refl)  | 
|
154  | 
apply simp  | 
|
155  | 
apply(rule someI_ex)  | 
|
156  | 
apply fast  | 
|
157  | 
done  | 
|
158  | 
||
159  | 
lemma [simp]: "islist h Null"  | 
|
160  | 
by(simp add:islist_def)  | 
|
161  | 
||
162  | 
lemma [simp]: "islist h (Ref a) = islist h (h a)"  | 
|
163  | 
by(simp add:islist_def)  | 
|
164  | 
||
165  | 
lemma [simp]: "list h Null = []"  | 
|
166  | 
by(simp add:list_def)  | 
|
167  | 
||
168  | 
lemma list_Ref_conv[simp]:  | 
|
169  | 
"islist h (h a) \<Longrightarrow> list h (Ref a) = a # list h (h a)"  | 
|
170  | 
apply(insert List_Ref[of h])  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
38353 
diff
changeset
 | 
171  | 
apply(fastforce simp:List_conv_islist_list)  | 
| 13875 | 172  | 
done  | 
173  | 
||
174  | 
lemma [simp]: "islist h (h a) \<Longrightarrow> a \<notin> set(list h (h a))"  | 
|
175  | 
apply(insert List_hd_not_in_tl[of h])  | 
|
176  | 
apply(simp add:List_conv_islist_list)  | 
|
177  | 
done  | 
|
178  | 
||
179  | 
lemma list_upd_conv[simp]:  | 
|
180  | 
"islist h p \<Longrightarrow> y \<notin> set(list h p) \<Longrightarrow> list (h(y := q)) p = list h p"  | 
|
181  | 
apply(drule notin_List_update[of _ _ h q p])  | 
|
182  | 
apply(simp add:List_conv_islist_list)  | 
|
183  | 
done  | 
|
184  | 
||
185  | 
lemma islist_upd[simp]:  | 
|
186  | 
"islist h p \<Longrightarrow> y \<notin> set(list h p) \<Longrightarrow> islist (h(y := q)) p"  | 
|
187  | 
apply(frule notin_List_update[of _ _ h q p])  | 
|
188  | 
apply(simp add:List_conv_islist_list)  | 
|
189  | 
done  | 
|
190  | 
||
191  | 
end  |