author | wenzelm |
Sun, 31 Dec 2023 19:24:37 +0100 | |
changeset 79409 | e1895596e1b9 |
parent 72990 | db8f94656024 |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: HOL/Hoare/SepLogHeap.thy |
14074 | 2 |
Author: Tobias Nipkow |
3 |
Copyright 2002 TUM |
|
4 |
*) |
|
5 |
||
72990 | 6 |
section \<open>Heap abstractions for Separation Logic\<close> |
7 |
||
8 |
text \<open>(at the moment only Path and List)\<close> |
|
9 |
||
18576 | 10 |
theory SepLogHeap |
72990 | 11 |
imports Main |
18576 | 12 |
begin |
18447 | 13 |
|
42174 | 14 |
type_synonym heap = "(nat \<Rightarrow> nat option)" |
14074 | 15 |
|
62042 | 16 |
text\<open>\<open>Some\<close> means allocated, \<open>None\<close> means |
17 |
free. Address \<open>0\<close> serves as the null reference.\<close> |
|
14074 | 18 |
|
72990 | 19 |
|
14074 | 20 |
subsection "Paths in the heap" |
21 |
||
38353 | 22 |
primrec Path :: "heap \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool" |
23 |
where |
|
24 |
"Path h x [] y = (x = y)" |
|
25 |
| "Path h x (a#as) y = (x\<noteq>0 \<and> a=x \<and> (\<exists>b. h x = Some b \<and> Path h b as y))" |
|
14074 | 26 |
|
27 |
lemma [iff]: "Path h 0 xs y = (xs = [] \<and> y = 0)" |
|
16972 | 28 |
by (cases xs) simp_all |
14074 | 29 |
|
30 |
lemma [simp]: "x\<noteq>0 \<Longrightarrow> Path h x as z = |
|
31 |
(as = [] \<and> z = x \<or> (\<exists>y bs. as = x#bs \<and> h x = Some y & Path h y bs z))" |
|
16972 | 32 |
by (cases as) auto |
14074 | 33 |
|
34 |
lemma [simp]: "\<And>x. Path f x (as@bs) z = (\<exists>y. Path f x as y \<and> Path f y bs z)" |
|
16972 | 35 |
by (induct as) auto |
14074 | 36 |
|
37 |
lemma Path_upd[simp]: |
|
38 |
"\<And>x. u \<notin> set as \<Longrightarrow> Path (f(u := v)) x as y = Path f x as y" |
|
16972 | 39 |
by (induct as) simp_all |
14074 | 40 |
|
41 |
||
42 |
subsection "Lists on the heap" |
|
43 |
||
38353 | 44 |
definition List :: "heap \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool" |
45 |
where "List h x as = Path h x as 0" |
|
14074 | 46 |
|
47 |
lemma [simp]: "List h x [] = (x = 0)" |
|
16972 | 48 |
by (simp add: List_def) |
14074 | 49 |
|
50 |
lemma [simp]: |
|
51 |
"List h x (a#as) = (x\<noteq>0 \<and> a=x \<and> (\<exists>y. h x = Some y \<and> List h y as))" |
|
16972 | 52 |
by (simp add: List_def) |
14074 | 53 |
|
54 |
lemma [simp]: "List h 0 as = (as = [])" |
|
16972 | 55 |
by (cases as) simp_all |
14074 | 56 |
|
57 |
lemma List_non_null: "a\<noteq>0 \<Longrightarrow> |
|
58 |
List h a as = (\<exists>b bs. as = a#bs \<and> h a = Some b \<and> List h b bs)" |
|
16972 | 59 |
by (cases as) simp_all |
14074 | 60 |
|
61 |
theorem notin_List_update[simp]: |
|
62 |
"\<And>x. a \<notin> set as \<Longrightarrow> List (h(a := y)) x as = List h x as" |
|
16972 | 63 |
by (induct as) simp_all |
14074 | 64 |
|
65 |
lemma List_unique: "\<And>x bs. List h x as \<Longrightarrow> List h x bs \<Longrightarrow> as = bs" |
|
16972 | 66 |
by (induct as) (auto simp add:List_non_null) |
14074 | 67 |
|
68 |
lemma List_unique1: "List h p as \<Longrightarrow> \<exists>!as. List h p as" |
|
16972 | 69 |
by (blast intro: List_unique) |
14074 | 70 |
|
71 |
lemma List_app: "\<And>x. List h x (as@bs) = (\<exists>y. Path h x as y \<and> List h y bs)" |
|
16972 | 72 |
by (induct as) auto |
14074 | 73 |
|
74 |
lemma List_hd_not_in_tl[simp]: "List h b as \<Longrightarrow> h a = Some b \<Longrightarrow> a \<notin> set as" |
|
75 |
apply (clarsimp simp add:in_set_conv_decomp) |
|
76 |
apply(frule List_app[THEN iffD1]) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
42174
diff
changeset
|
77 |
apply(fastforce dest: List_unique) |
14074 | 78 |
done |
79 |
||
80 |
lemma List_distinct[simp]: "\<And>x. List h x as \<Longrightarrow> distinct as" |
|
16972 | 81 |
by (induct as) (auto dest:List_hd_not_in_tl) |
14074 | 82 |
|
83 |
lemma list_in_heap: "\<And>p. List h p ps \<Longrightarrow> set ps \<subseteq> dom h" |
|
16972 | 84 |
by (induct ps) auto |
14074 | 85 |
|
86 |
lemma list_ortho_sum1[simp]: |
|
87 |
"\<And>p. \<lbrakk> List h1 p ps; dom h1 \<inter> dom h2 = {}\<rbrakk> \<Longrightarrow> List (h1++h2) p ps" |
|
16972 | 88 |
by (induct ps) (auto simp add:map_add_def split:option.split) |
14074 | 89 |
|
18447 | 90 |
|
14074 | 91 |
lemma list_ortho_sum2[simp]: |
92 |
"\<And>p. \<lbrakk> List h2 p ps; dom h1 \<inter> dom h2 = {}\<rbrakk> \<Longrightarrow> List (h1++h2) p ps" |
|
16972 | 93 |
by (induct ps) (auto simp add:map_add_def split:option.split) |
14074 | 94 |
|
95 |
end |