25
|
1 |
header {*
|
|
2 |
Generic Separation Logic
|
|
3 |
*}
|
|
4 |
|
|
5 |
theory Hoare_gen
|
|
6 |
imports Main
|
|
7 |
"../progtut/Tactical" "../Separation_Algebra/Sep_Tactics"
|
|
8 |
begin
|
|
9 |
|
|
10 |
ML_file "../Separation_Algebra/sep_tactics.ML"
|
|
11 |
|
|
12 |
definition pasrt :: "bool \<Rightarrow> (('a::sep_algebra) \<Rightarrow> bool)" ("<_>" [72] 71)
|
|
13 |
where "pasrt b = (\<lambda> s . s = 0 \<and> b)"
|
|
14 |
|
|
15 |
lemma sep_conj_cond1: "(p \<and>* <cond> \<and>* q) = (<cond> \<and>* p \<and>* q)"
|
|
16 |
by(simp add: sep_conj_ac)
|
|
17 |
|
|
18 |
lemma sep_conj_cond2: "(p \<and>* <cond>) = (<cond> \<and>* p)"
|
|
19 |
by(simp add: sep_conj_ac)
|
|
20 |
|
|
21 |
lemma sep_conj_cond3: "((<cond> \<and>* p) \<and>* r) = (<cond> \<and>* p \<and>* r)"
|
|
22 |
by (metis sep.mult_assoc)
|
|
23 |
|
|
24 |
lemmas sep_conj_cond = sep_conj_cond1 sep_conj_cond2 sep_conj_cond3
|
|
25 |
|
|
26 |
lemma cond_true_eq[simp]: "<True> = \<box>"
|
|
27 |
by(unfold sep_empty_def pasrt_def, auto)
|
|
28 |
|
|
29 |
lemma cond_true_eq1: "(<True> \<and>* p) = p"
|
|
30 |
by(simp)
|
|
31 |
|
|
32 |
lemma false_simp [simp]: "<False> = sep_false" (* move *)
|
|
33 |
by (simp add:pasrt_def)
|
|
34 |
|
|
35 |
lemma cond_true_eq2: "(p \<and>* <True>) = p"
|
|
36 |
by simp
|
|
37 |
|
|
38 |
lemma condD: "(<b> ** r) s \<Longrightarrow> b \<and> r s"
|
|
39 |
by (unfold sep_conj_def pasrt_def, auto)
|
|
40 |
|
|
41 |
locale sep_exec =
|
|
42 |
fixes step :: "'conf \<Rightarrow> 'conf"
|
|
43 |
and recse:: "'conf \<Rightarrow> 'a::sep_algebra"
|
|
44 |
begin
|
|
45 |
|
|
46 |
definition "run n = step ^^ n"
|
|
47 |
|
|
48 |
lemma run_add: "run (n1 + n2) s = run n1 (run n2 s)"
|
|
49 |
apply (unfold run_def)
|
|
50 |
by (metis funpow_add o_apply)
|
|
51 |
|
|
52 |
definition
|
|
53 |
Hoare_gen :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
|
|
54 |
("(\<lbrace>(1_)\<rbrace> / (_)/ \<lbrace>(1_)\<rbrace>)" 50)
|
|
55 |
where
|
|
56 |
"\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace> \<equiv>
|
|
57 |
(\<forall> s r. (p**c**r) (recse s) \<longrightarrow> (\<exists> k. ((q ** c ** r) (recse (run (Suc k) s)))))"
|
|
58 |
|
|
59 |
lemma HoareI [case_names Pre]:
|
|
60 |
assumes h: "\<And> r s. (p**c**r) (recse s) \<Longrightarrow> (\<exists> k. ((q ** c ** r) (recse (run (Suc k) s))))"
|
|
61 |
shows "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
|
|
62 |
using h
|
|
63 |
by (unfold Hoare_gen_def, auto)
|
|
64 |
|
|
65 |
lemma frame_rule:
|
|
66 |
assumes h: "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
|
|
67 |
shows "\<lbrace> p ** r \<rbrace> c \<lbrace> q ** r \<rbrace>"
|
|
68 |
proof(induct rule: HoareI)
|
|
69 |
case (Pre r' s')
|
|
70 |
hence "(p \<and>* c \<and>* r \<and>* r') (recse s')" by (auto simp:sep_conj_ac)
|
|
71 |
from h[unfolded Hoare_gen_def, rule_format, OF this]
|
|
72 |
show ?case
|
|
73 |
by (metis sep_conj_assoc sep_conj_left_commute)
|
|
74 |
qed
|
|
75 |
|
|
76 |
lemma sequencing:
|
|
77 |
assumes h1: "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
78 |
and h2: "\<lbrace>q\<rbrace> c \<lbrace>r\<rbrace>"
|
|
79 |
shows "\<lbrace>p\<rbrace> c \<lbrace>r\<rbrace>"
|
|
80 |
proof(induct rule:HoareI)
|
|
81 |
case (Pre r' s')
|
|
82 |
from h1[unfolded Hoare_gen_def, rule_format, OF Pre]
|
|
83 |
obtain k1 where "(q \<and>* c \<and>* r') (recse (run (Suc k1) s'))" by auto
|
|
84 |
from h2[unfolded Hoare_gen_def, rule_format, OF this]
|
|
85 |
obtain k2 where "(r \<and>* c \<and>* r') (recse (run (Suc k2) (run (Suc k1) s')))" by auto
|
|
86 |
thus ?case
|
|
87 |
apply (rule_tac x = "Suc (k1 + k2)" in exI)
|
|
88 |
by (metis add_Suc_right nat_add_commute sep_exec.run_add)
|
|
89 |
qed
|
|
90 |
|
|
91 |
lemma pre_stren:
|
|
92 |
assumes h1: "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
93 |
and h2: "\<And>s. r s \<Longrightarrow> p s"
|
|
94 |
shows "\<lbrace>r\<rbrace> c \<lbrace>q\<rbrace>"
|
|
95 |
proof(induct rule:HoareI)
|
|
96 |
case (Pre r' s')
|
|
97 |
with h2
|
|
98 |
have "(p \<and>* c \<and>* r') (recse s')"
|
|
99 |
by (metis sep_conj_impl1)
|
|
100 |
from h1[unfolded Hoare_gen_def, rule_format, OF this]
|
|
101 |
show ?case .
|
|
102 |
qed
|
|
103 |
|
|
104 |
lemma post_weaken:
|
|
105 |
assumes h1: "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
106 |
and h2: "\<And> s. q s \<Longrightarrow> r s"
|
|
107 |
shows "\<lbrace>p\<rbrace> c \<lbrace>r\<rbrace>"
|
|
108 |
proof(induct rule:HoareI)
|
|
109 |
case (Pre r' s')
|
|
110 |
from h1[unfolded Hoare_gen_def, rule_format, OF this]
|
|
111 |
obtain k where "(q \<and>* c \<and>* r') (recse (run (Suc k) s'))" by blast
|
|
112 |
with h2
|
|
113 |
show ?case
|
|
114 |
by (metis sep_conj_impl1)
|
|
115 |
qed
|
|
116 |
|
|
117 |
lemma hoare_adjust:
|
|
118 |
assumes h1: "\<lbrace>p1\<rbrace> c \<lbrace>q1\<rbrace>"
|
|
119 |
and h2: "\<And>s. p s \<Longrightarrow> p1 s"
|
|
120 |
and h3: "\<And>s. q1 s \<Longrightarrow> q s"
|
|
121 |
shows "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
122 |
using h1 h2 h3 post_weaken pre_stren
|
|
123 |
by (metis)
|
|
124 |
|
|
125 |
lemma code_exI:
|
|
126 |
assumes h: "\<And> k. \<lbrace>p\<rbrace> c(k) \<lbrace>q\<rbrace>"
|
|
127 |
shows "\<lbrace>p\<rbrace> EXS k. c(k) \<lbrace>q\<rbrace>"
|
|
128 |
proof(unfold pred_ex_def, induct rule:HoareI)
|
|
129 |
case (Pre r' s')
|
|
130 |
then obtain k where "(p \<and>* (\<lambda> s. c k s) \<and>* r') (recse s')"
|
|
131 |
by (auto elim!:sep_conjE intro!:sep_conjI)
|
|
132 |
from h[unfolded Hoare_gen_def, rule_format, OF this]
|
|
133 |
show ?case
|
|
134 |
by (auto elim!:sep_conjE intro!:sep_conjI)
|
|
135 |
qed
|
|
136 |
|
|
137 |
lemma code_extension:
|
|
138 |
assumes h: "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
|
|
139 |
shows "\<lbrace> p \<rbrace> c ** e \<lbrace> q \<rbrace>"
|
|
140 |
proof(induct rule:HoareI)
|
|
141 |
case (Pre r' s')
|
|
142 |
hence "(p \<and>* c \<and>* e \<and>* r') (recse s')" by (auto simp:sep_conj_ac)
|
|
143 |
from h[unfolded Hoare_gen_def, rule_format, OF this]
|
|
144 |
show ?case
|
|
145 |
by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
|
|
146 |
qed
|
|
147 |
|
|
148 |
lemma code_extension1:
|
|
149 |
assumes h: "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
|
|
150 |
shows "\<lbrace> p \<rbrace> e ** c \<lbrace> q \<rbrace>"
|
|
151 |
by (metis code_extension h sep.mult_commute)
|
|
152 |
|
|
153 |
lemma composition:
|
|
154 |
assumes h1: "\<lbrace>p\<rbrace> c1 \<lbrace>q\<rbrace>"
|
|
155 |
and h2: "\<lbrace>q\<rbrace> c2 \<lbrace>r\<rbrace>"
|
|
156 |
shows "\<lbrace>p\<rbrace> c1 ** c2 \<lbrace>r\<rbrace>"
|
|
157 |
proof(induct rule:HoareI)
|
|
158 |
case (Pre r' s')
|
|
159 |
hence "(p \<and>* c1 \<and>* c2 \<and>* r') (recse s')" by (auto simp:sep_conj_ac)
|
|
160 |
from h1[unfolded Hoare_gen_def, rule_format, OF this]
|
|
161 |
obtain k1 where "(q \<and>* c2 \<and>* c1 \<and>* r') (recse (run (Suc k1) s'))"
|
|
162 |
by (auto simp:sep_conj_ac)
|
|
163 |
from h2[unfolded Hoare_gen_def, rule_format, OF this, folded run_add]
|
|
164 |
show ?case
|
|
165 |
by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
|
|
166 |
qed
|
|
167 |
|
|
168 |
|
|
169 |
definition
|
|
170 |
IHoare :: "('b::sep_algebra \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow>
|
|
171 |
('b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
|
|
172 |
("(1_).(\<lbrace>(1_)\<rbrace> / (_)/ \<lbrace>(1_)\<rbrace>)" 50)
|
|
173 |
where
|
|
174 |
"I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace> = (\<forall>s'. \<lbrace> EXS s. <P s> \<and>* <(s ## s')> \<and>* I(s + s')\<rbrace> c
|
|
175 |
\<lbrace> EXS s. <Q s> \<and>* <(s ## s')> \<and>* I(s + s')\<rbrace>)"
|
|
176 |
|
|
177 |
lemma IHoareI [case_names IPre]:
|
|
178 |
assumes h: "\<And>s' s r cnf . (<P s> \<and>* <(s ## s')> \<and>* I(s + s') \<and>* c \<and>* r) (recse cnf) \<Longrightarrow>
|
|
179 |
(\<exists>k t. (<Q t> \<and>* <(t ## s')> \<and>* I(t + s') \<and>* c \<and>* r) (recse (run (Suc k) cnf)))"
|
|
180 |
shows "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
181 |
unfolding IHoare_def
|
|
182 |
proof
|
|
183 |
fix s'
|
|
184 |
show " \<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace> c
|
|
185 |
\<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
|
|
186 |
proof(unfold pred_ex_def, induct rule:HoareI)
|
|
187 |
case (Pre r s)
|
|
188 |
then obtain sa where "(<P sa> \<and>* <(sa ## s')> \<and>* I (sa + s') \<and>* c \<and>* r) (recse s)"
|
|
189 |
by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
|
|
190 |
hence " (\<exists>k t. (<Q t> \<and>* <(t##s')> \<and>* I(t + s') \<and>* c \<and>* r) (recse (run (Suc k) s)))"
|
|
191 |
by (rule h)
|
|
192 |
then obtain k t where h2: "(<Q t> \<and>* <(t ## s')> \<and>* I(t + s') \<and>* c \<and>* r) (recse (run (Suc k) s))"
|
|
193 |
by auto
|
|
194 |
thus "\<exists>k. ((\<lambda>s. \<exists>sa. (<Q sa> \<and>* <(sa ## s')> \<and>* I (sa + s')) s) \<and>* c \<and>* r) (recse (run (Suc k) s))"
|
|
195 |
apply (rule_tac x = k in exI)
|
|
196 |
by (auto intro!:sep_conjI elim!:sep_conjE simp:sep_conj_ac)
|
|
197 |
qed
|
|
198 |
qed
|
|
199 |
|
|
200 |
lemma I_frame_rule:
|
|
201 |
assumes h: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
202 |
shows "I. \<lbrace>P \<and>* R\<rbrace> c \<lbrace>Q \<and>* R\<rbrace>"
|
|
203 |
proof(induct rule:IHoareI)
|
|
204 |
case (IPre s' s r cnf)
|
|
205 |
hence "((<(P \<and>* R) s> \<and>* <(s##s')> \<and>* I (s + s')) \<and>* c \<and>* r) (recse cnf)"
|
|
206 |
by (auto simp:sep_conj_ac)
|
|
207 |
then obtain s1 s2
|
|
208 |
where h1: "((<P s1> \<and>* <((s1 + s2) ## s')> \<and>*I (s1 + s2 + s')) \<and>* c \<and>* r) (recse cnf)"
|
|
209 |
"s1 ## s2" "s1 + s2 ## s'" "P s1" "R s2"
|
|
210 |
by (unfold pasrt_def, auto elim!:sep_conjE intro!:sep_conjI)
|
|
211 |
hence "((EXS s. <P s> \<and>* <(s ## s2 +s')> \<and>*I (s + (s2 + s'))) \<and>* c \<and>* r) (recse cnf)"
|
|
212 |
apply (sep_cancel, unfold pred_ex_def, auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE)
|
|
213 |
apply (rule_tac x = s1 in exI, unfold pasrt_def,
|
|
214 |
auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE simp:sep_conj_ac)
|
|
215 |
by (metis sep_add_assoc sep_add_disjD)
|
|
216 |
from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this]
|
|
217 |
obtain k s where
|
|
218 |
"((<Q s> \<and>* <(s ## s2 + s')> \<and>* I (s + (s2 + s'))) \<and>* c \<and>* r) (recse (run (Suc k) cnf))"
|
|
219 |
by (unfold pasrt_def pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI)
|
|
220 |
thus ?case
|
|
221 |
proof(rule_tac x = k in exI, rule_tac x = "s + s2" in exI, sep_cancel+)
|
|
222 |
fix h ha
|
|
223 |
assume hh: "(<Q s> \<and>* <(s ## s2 + s')> \<and>* I (s + (s2 + s'))) ha"
|
|
224 |
show " (<(Q \<and>* R) (s + s2)> \<and>* <(s + s2 ## s')> \<and>* I (s + s2 + s')) ha"
|
|
225 |
proof -
|
|
226 |
from hh have h0: "s ## s2 + s'"
|
|
227 |
by (metis pasrt_def sep_conjD)
|
|
228 |
with h1(2, 3)
|
|
229 |
have h2: "s + s2 ## s'"
|
|
230 |
by (metis sep_add_disjD sep_disj_addI1)
|
|
231 |
moreover from h1(2, 3) h2 have h3: "(s + (s2 + s')) = (s + s2 + s')"
|
|
232 |
by (metis `s ## s2 + s'` sep_add_assoc sep_add_disjD sep_disj_addD1)
|
|
233 |
moreover from hh have "Q s" by (metis pasrt_def sep_conjD)
|
|
234 |
moreover from h0 h1(2) h1(3) have "s ## s2"
|
|
235 |
by (metis sep_add_disjD sep_disj_addD)
|
|
236 |
moreover note h1(5)
|
|
237 |
ultimately show ?thesis
|
|
238 |
by (smt h0 hh sep_conjI)
|
|
239 |
qed
|
|
240 |
qed
|
|
241 |
qed
|
|
242 |
|
|
243 |
lemma I_sequencing:
|
|
244 |
assumes h1: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
245 |
and h2: "I. \<lbrace>Q\<rbrace> c \<lbrace>R\<rbrace>"
|
|
246 |
shows "I. \<lbrace>P\<rbrace> c \<lbrace>R\<rbrace>"
|
|
247 |
using h1 h2 sequencing
|
|
248 |
by (smt IHoare_def)
|
|
249 |
|
|
250 |
lemma I_pre_stren:
|
|
251 |
assumes h1: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
252 |
and h2: "\<And>s. R s \<Longrightarrow> P s"
|
|
253 |
shows "I. \<lbrace>R\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
254 |
proof(unfold IHoare_def, default)
|
|
255 |
fix s'
|
|
256 |
show "\<lbrace>EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace> c
|
|
257 |
\<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
|
|
258 |
proof(rule pre_stren)
|
|
259 |
from h1[unfolded IHoare_def, rule_format, of s']
|
|
260 |
show "\<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace> c
|
|
261 |
\<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>" .
|
|
262 |
next
|
|
263 |
fix s
|
|
264 |
show "(EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')) s \<Longrightarrow>
|
|
265 |
(EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')) s"
|
|
266 |
apply (unfold pred_ex_def, clarify)
|
|
267 |
apply (rule_tac x = sa in exI, sep_cancel+)
|
|
268 |
by (insert h2, auto simp:pasrt_def)
|
|
269 |
qed
|
|
270 |
qed
|
|
271 |
|
|
272 |
lemma I_post_weaken:
|
|
273 |
assumes h1: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
274 |
and h2: "\<And> s. Q s \<Longrightarrow> R s"
|
|
275 |
shows "I. \<lbrace>P\<rbrace> c \<lbrace>R\<rbrace>"
|
|
276 |
proof(unfold IHoare_def, default)
|
|
277 |
fix s'
|
|
278 |
show "\<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace> c
|
|
279 |
\<lbrace>EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
|
|
280 |
proof(rule post_weaken)
|
|
281 |
from h1[unfolded IHoare_def, rule_format, of s']
|
|
282 |
show "\<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace> c
|
|
283 |
\<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>" .
|
|
284 |
next
|
|
285 |
fix s
|
|
286 |
show "(EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')) s \<Longrightarrow>
|
|
287 |
(EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')) s"
|
|
288 |
apply (unfold pred_ex_def, clarify)
|
|
289 |
apply (rule_tac x = sa in exI, sep_cancel+)
|
|
290 |
by (insert h2, auto simp:pasrt_def)
|
|
291 |
qed
|
|
292 |
qed
|
|
293 |
|
|
294 |
lemma I_hoare_adjust:
|
|
295 |
assumes h1: "I. \<lbrace>P1\<rbrace> c \<lbrace>Q1\<rbrace>"
|
|
296 |
and h2: "\<And>s. P s \<Longrightarrow> P1 s"
|
|
297 |
and h3: "\<And>s. Q1 s \<Longrightarrow> Q s"
|
|
298 |
shows "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
299 |
using h1 h2 h3 I_post_weaken I_pre_stren
|
|
300 |
by (metis)
|
|
301 |
|
|
302 |
lemma I_code_exI:
|
|
303 |
assumes h: "\<And> k. I. \<lbrace>P\<rbrace> c(k) \<lbrace>Q\<rbrace>"
|
|
304 |
shows "I. \<lbrace>P\<rbrace> EXS k. c(k) \<lbrace>Q\<rbrace>"
|
|
305 |
using h
|
|
306 |
by (smt IHoare_def code_exI)
|
|
307 |
|
|
308 |
lemma I_code_extension:
|
|
309 |
assumes h: "I. \<lbrace> P \<rbrace> c \<lbrace> Q \<rbrace>"
|
|
310 |
shows "I. \<lbrace> P \<rbrace> c ** e \<lbrace> Q \<rbrace>"
|
|
311 |
using h
|
|
312 |
by (smt IHoare_def sep_exec.code_extension)
|
|
313 |
|
|
314 |
lemma I_composition:
|
|
315 |
assumes h1: "I. \<lbrace>P\<rbrace> c1 \<lbrace>Q\<rbrace>"
|
|
316 |
and h2: "I. \<lbrace>Q\<rbrace> c2 \<lbrace>R\<rbrace>"
|
|
317 |
shows "I. \<lbrace>P\<rbrace> c1 ** c2 \<lbrace>R\<rbrace>"
|
|
318 |
using h1 h2
|
|
319 |
by (smt IHoare_def sep_exec.composition)
|
|
320 |
|
|
321 |
lemma pre_condI:
|
|
322 |
assumes h: "cond \<Longrightarrow> \<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
323 |
shows "\<lbrace><cond> \<and>* p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
324 |
proof(induct rule:HoareI)
|
|
325 |
case (Pre r s)
|
|
326 |
hence "cond" "(p \<and>* c \<and>* r) (recse s)"
|
|
327 |
apply (metis pasrt_def sep_conjD)
|
|
328 |
by (smt Pre.hyps pasrt_def sep_add_zero sep_conj_commute sep_conj_def)
|
|
329 |
from h[OF this(1), unfolded Hoare_gen_def, rule_format, OF this(2)]
|
|
330 |
show ?case .
|
|
331 |
qed
|
|
332 |
|
|
333 |
lemma I_pre_condI:
|
|
334 |
assumes h: "cond \<Longrightarrow> I.\<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
335 |
shows "I.\<lbrace><cond> \<and>* P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
336 |
using h
|
|
337 |
by (smt IHoareI condD cond_true_eq2 sep.mult_commute)
|
|
338 |
|
|
339 |
lemma code_condI:
|
|
340 |
assumes h: "b \<Longrightarrow> \<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
|
|
341 |
shows "\<lbrace>p\<rbrace> <b>**c \<lbrace>q\<rbrace>"
|
|
342 |
proof(induct rule: HoareI)
|
|
343 |
case (Pre r s)
|
|
344 |
hence h1: "b" "(p \<and>* c \<and>* r) (recse s)"
|
|
345 |
apply (metis condD sep_conjD sep_conj_assoc)
|
|
346 |
by (smt Pre.hyps condD sep_conj_impl)
|
|
347 |
from h[OF h1(1), unfolded Hoare_gen_def, rule_format, OF h1(2)]
|
|
348 |
and h1(1)
|
|
349 |
show ?case
|
|
350 |
by (metis (full_types) cond_true_eq1)
|
|
351 |
qed
|
|
352 |
|
|
353 |
lemma I_code_condI:
|
|
354 |
assumes h: "b \<Longrightarrow> I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
355 |
shows "I.\<lbrace>P\<rbrace> <b>**c \<lbrace>Q\<rbrace>"
|
|
356 |
using h
|
|
357 |
by (smt IHoareI condD cond_true_eq2 sep.mult_commute sep_conj_cond1)
|
|
358 |
|
|
359 |
lemma precond_exI:
|
|
360 |
assumes h:"\<And>x. \<lbrace>p x\<rbrace> c \<lbrace>q\<rbrace>"
|
|
361 |
shows "\<lbrace>EXS x. p x\<rbrace> c \<lbrace>q\<rbrace>"
|
|
362 |
proof(induct rule:HoareI)
|
|
363 |
case (Pre r s)
|
|
364 |
then obtain x where "(p x \<and>* c \<and>* r) (recse s)"
|
|
365 |
by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI)
|
|
366 |
from h[of x, unfolded Hoare_gen_def, rule_format, OF this]
|
|
367 |
show ?case .
|
|
368 |
qed
|
|
369 |
|
|
370 |
lemma I_precond_exI:
|
|
371 |
assumes h:"\<And>x. I. \<lbrace>P x\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
372 |
shows "I.\<lbrace>EXS x. P x\<rbrace> c \<lbrace>Q\<rbrace>"
|
|
373 |
proof(induct rule:IHoareI)
|
|
374 |
case (IPre s' s r cnf)
|
|
375 |
then obtain x
|
|
376 |
where "((EXS s. <P x s> \<and>* <(s ## s')> \<and>* I (s + s')) \<and>* c \<and>* r) (recse cnf)"
|
|
377 |
by ( auto elim!:sep_conjE intro!:sep_conjI simp:pred_ex_def pasrt_def sep_conj_ac)
|
|
378 |
from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this]
|
|
379 |
obtain k t
|
|
380 |
where "((<Q t> \<and>* <(t ## s')> \<and>* I (t + s')) \<and>* c \<and>* r) (recse (run (Suc k) cnf))"
|
|
381 |
by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI)
|
|
382 |
thus ?case
|
|
383 |
by (auto simp:sep_conj_ac)
|
|
384 |
qed
|
|
385 |
|
|
386 |
lemma hoare_sep_false:
|
|
387 |
"\<lbrace>sep_false\<rbrace> c
|
|
388 |
\<lbrace>q\<rbrace>"
|
|
389 |
by(unfold Hoare_gen_def, clarify, simp)
|
|
390 |
|
|
391 |
lemma I_hoare_sep_false:
|
|
392 |
"I. \<lbrace>sep_false\<rbrace> c
|
|
393 |
\<lbrace>Q\<rbrace>"
|
|
394 |
by (smt IHoareI condD)
|
|
395 |
|
|
396 |
end
|
|
397 |
|
|
398 |
instantiation set :: (type)sep_algebra
|
|
399 |
begin
|
|
400 |
|
|
401 |
definition set_zero_def: "0 = {}"
|
|
402 |
|
|
403 |
definition plus_set_def: "s1 + s2 = s1 \<union> s2"
|
|
404 |
|
|
405 |
definition sep_disj_set_def: "sep_disj s1 s2 = (s1 \<inter> s2 = {})"
|
|
406 |
|
|
407 |
lemmas set_ins_def = sep_disj_set_def plus_set_def set_zero_def
|
|
408 |
|
|
409 |
instance
|
|
410 |
apply(default)
|
|
411 |
apply(simp add:set_ins_def)
|
|
412 |
apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
|
|
413 |
apply (metis inf_commute)
|
|
414 |
apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
|
|
415 |
apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
|
|
416 |
apply (metis sup_commute)
|
|
417 |
apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
|
|
418 |
apply (metis (lifting) Un_assoc)
|
|
419 |
apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
|
|
420 |
apply (metis (lifting) Int_Un_distrib Un_empty inf_sup_distrib1 sup_eq_bot_iff)
|
|
421 |
apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
|
|
422 |
by (metis (lifting) Int_Un_distrib Int_Un_distrib2 sup_eq_bot_iff)
|
|
423 |
end
|
|
424 |
|
|
425 |
section {* A big operator of infinite separation conjunction *}
|
|
426 |
|
|
427 |
definition "fam_conj I cpt s = (\<exists> p. (s = (\<Union> i \<in> I. p(i))) \<and>
|
|
428 |
(\<forall> i \<in> I. cpt i (p i)) \<and>
|
|
429 |
(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> p(i) ## p(j)))"
|
|
430 |
|
|
431 |
lemma fam_conj_zero_simp: "fam_conj {} cpt = <True>"
|
|
432 |
proof
|
|
433 |
fix s
|
|
434 |
show "fam_conj {} cpt s = (<True>) s"
|
|
435 |
proof
|
|
436 |
assume "fam_conj {} cpt s"
|
|
437 |
then obtain p where "s = (\<Union> i \<in> {}. p(i))"
|
|
438 |
by (unfold fam_conj_def, auto)
|
|
439 |
hence "s = {}" by auto
|
|
440 |
thus "(<True>) s" by (metis pasrt_def set_zero_def)
|
|
441 |
next
|
|
442 |
assume "(<True>) s"
|
|
443 |
hence eq_s: "s = {}" by (metis pasrt_def set_zero_def)
|
|
444 |
let ?p = "\<lambda> i. {}"
|
|
445 |
have "(s = (\<Union> i \<in> {}. ?p(i)))" by (unfold eq_s, auto)
|
|
446 |
moreover have "(\<forall> i \<in> {}. cpt i (?p i))" by auto
|
|
447 |
moreover have "(\<forall> i \<in> {}. \<forall> j \<in> {}. i \<noteq> j \<longrightarrow> ?p(i) ## ?p(j))" by auto
|
|
448 |
ultimately show "fam_conj {} cpt s"
|
|
449 |
by (unfold eq_s fam_conj_def, auto)
|
|
450 |
qed
|
|
451 |
qed
|
|
452 |
|
|
453 |
lemma fam_conj_disj_simp_pre:
|
|
454 |
assumes h1: "I = I1 + I2"
|
|
455 |
and h2: "I1 ## I2"
|
|
456 |
shows "fam_conj I cpt = (fam_conj I1 cpt \<and>* fam_conj I2 cpt)"
|
|
457 |
proof
|
|
458 |
fix s
|
|
459 |
let ?fm = "fam_conj I cpt" and ?fm1 = "fam_conj I1 cpt" and ?fm2 = "fam_conj I2 cpt"
|
|
460 |
show "?fm s = (?fm1 \<and>* ?fm2) s"
|
|
461 |
proof
|
|
462 |
assume "?fm s"
|
|
463 |
then obtain p where pre:
|
|
464 |
"s = (\<Union> i \<in> I. p(i))"
|
|
465 |
"(\<forall> i \<in> I. cpt i (p i))"
|
|
466 |
"(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> p(i) ## p(j))"
|
|
467 |
unfolding fam_conj_def by metis
|
|
468 |
from pre(1) h1 h2 have "s = (\<Union> i \<in> I1. p(i)) + (\<Union> i \<in> I2. p(i))"
|
|
469 |
by (auto simp:set_ins_def)
|
|
470 |
moreover from pre h1 have "?fm1 (\<Union> i \<in> I1. p(i))"
|
|
471 |
by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def)
|
|
472 |
moreover from pre h1 have "?fm2 (\<Union> i \<in> I2. p(i))"
|
|
473 |
by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def)
|
|
474 |
moreover have "(\<Union> i \<in> I1. p(i)) ## (\<Union> i \<in> I2. p(i))"
|
|
475 |
proof -
|
|
476 |
{ fix x xa xb
|
|
477 |
assume h: "I1 \<inter> I2 = {}"
|
|
478 |
"\<forall>i\<in>I1 \<union> I2. \<forall>j\<in>I1 \<union> I2. i \<noteq> j \<longrightarrow> p i \<inter> p j = {}"
|
|
479 |
"xa \<in> I1" "x \<in> p xa" "xb \<in> I2" "x \<in> p xb"
|
|
480 |
have "p xa \<inter> p xb = {}"
|
|
481 |
proof(rule h(2)[rule_format])
|
|
482 |
from h(1, 3, 5) show "xa \<in> I1 \<union> I2" by auto
|
|
483 |
next
|
|
484 |
from h(1, 3, 5) show "xb \<in> I1 \<union> I2" by auto
|
|
485 |
next
|
|
486 |
from h(1, 3, 5) show "xa \<noteq> xb" by auto
|
|
487 |
qed with h(4, 6) have False by auto
|
|
488 |
} with h1 h2 pre(3) show ?thesis by (auto simp:set_ins_def)
|
|
489 |
qed
|
|
490 |
ultimately show "(?fm1 \<and>* ?fm2) s" by (auto intro!:sep_conjI)
|
|
491 |
next
|
|
492 |
assume "(?fm1 \<and>* ?fm2) s"
|
|
493 |
then obtain s1 s2 where pre:
|
|
494 |
"s = s1 + s2" "s1 ## s2" "?fm1 s1" "?fm2 s2"
|
|
495 |
by (auto dest!:sep_conjD)
|
|
496 |
from pre(3) obtain p1 where pre1:
|
|
497 |
"s1 = (\<Union> i \<in> I1. p1(i))"
|
|
498 |
"(\<forall> i \<in> I1. cpt i (p1 i))"
|
|
499 |
"(\<forall> i \<in> I1. \<forall> j \<in> I1. i \<noteq> j \<longrightarrow> p1(i) ## p1(j))"
|
|
500 |
unfolding fam_conj_def by metis
|
|
501 |
from pre(4) obtain p2 where pre2:
|
|
502 |
"s2 = (\<Union> i \<in> I2. p2(i))"
|
|
503 |
"(\<forall> i \<in> I2. cpt i (p2 i))"
|
|
504 |
"(\<forall> i \<in> I2. \<forall> j \<in> I2. i \<noteq> j \<longrightarrow> p2(i) ## p2(j))"
|
|
505 |
unfolding fam_conj_def by metis
|
|
506 |
let ?p = "\<lambda> i. if i \<in> I1 then p1 i else p2 i"
|
|
507 |
from h2 pre(2)
|
|
508 |
have "s = (\<Union> i \<in> I. ?p(i))"
|
|
509 |
apply (unfold h1 pre(1) pre1(1) pre2(1) set_ins_def, auto split:if_splits)
|
|
510 |
by (metis disjoint_iff_not_equal)
|
|
511 |
moreover from h2 pre1(2) pre2(2) have "(\<forall> i \<in> I. cpt i (?p i))"
|
|
512 |
by (unfold h1 set_ins_def, auto split:if_splits)
|
|
513 |
moreover from pre1(1, 3) pre2(1, 3) pre(2) h2
|
|
514 |
have "(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> ?p(i) ## ?p(j))"
|
|
515 |
apply (unfold h1 set_ins_def, auto split:if_splits)
|
|
516 |
by (metis IntI empty_iff)
|
|
517 |
ultimately show "?fm s" by (unfold fam_conj_def, auto)
|
|
518 |
qed
|
|
519 |
qed
|
|
520 |
|
|
521 |
lemma fam_conj_disj_simp:
|
|
522 |
assumes h: "I1 ## I2"
|
|
523 |
shows "fam_conj (I1 + I2) cpt = (fam_conj I1 cpt \<and>* fam_conj I2 cpt)"
|
|
524 |
proof -
|
|
525 |
from fam_conj_disj_simp_pre[OF _ h, of "I1 + I2"]
|
|
526 |
show ?thesis by simp
|
|
527 |
qed
|
|
528 |
|
|
529 |
lemma fam_conj_elm_simp:
|
|
530 |
assumes h: "i \<in> I"
|
|
531 |
shows "fam_conj I cpt = (cpt(i) \<and>* fam_conj (I - {i}) cpt)"
|
|
532 |
proof
|
|
533 |
fix s
|
|
534 |
show "fam_conj I cpt s = (cpt i \<and>* fam_conj (I - {i}) cpt) s"
|
|
535 |
proof
|
|
536 |
assume pre: "fam_conj I cpt s"
|
|
537 |
show "(cpt i \<and>* fam_conj (I - {i}) cpt) s"
|
|
538 |
proof -
|
|
539 |
from pre obtain p where pres:
|
|
540 |
"s = (\<Union> i \<in> I. p(i))"
|
|
541 |
"(\<forall> i \<in> I. cpt i (p i))"
|
|
542 |
"(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> p(i) ## p(j))"
|
|
543 |
unfolding fam_conj_def by metis
|
|
544 |
let ?s = "(\<Union>i\<in>(I - {i}). p i)"
|
|
545 |
from pres(3) h have disj: "(p i) ## ?s"
|
|
546 |
by (auto simp:set_ins_def)
|
|
547 |
moreover from pres(1) h have eq_s: "s = (p i) + ?s"
|
|
548 |
by (auto simp:set_ins_def)
|
|
549 |
moreover from pres(2) h have "cpt i (p i)" by auto
|
|
550 |
moreover from pres have "(fam_conj (I - {i}) cpt) ?s"
|
|
551 |
by (unfold fam_conj_def, rule_tac x = p in exI, auto)
|
|
552 |
ultimately show ?thesis by (auto intro!:sep_conjI)
|
|
553 |
qed
|
|
554 |
next
|
|
555 |
let ?fam = "fam_conj (I - {i}) cpt"
|
|
556 |
assume "(cpt i \<and>* ?fam) s"
|
|
557 |
then obtain s1 s2 where pre:
|
|
558 |
"s = s1 + s2" "s1 ## s2" "cpt i s1" "?fam s2"
|
|
559 |
by (auto dest!:sep_conjD)
|
|
560 |
from pre(4) obtain p where pres:
|
|
561 |
"s2 = (\<Union> ia \<in> I - {i}. p(ia))"
|
|
562 |
"(\<forall> ia \<in> I - {i}. cpt ia (p ia))"
|
|
563 |
"(\<forall> ia \<in> I - {i}. \<forall> j \<in> I - {i}. ia \<noteq> j \<longrightarrow> p(ia) ## p(j))"
|
|
564 |
unfolding fam_conj_def by metis
|
|
565 |
let ?p = "p(i:=s1)"
|
|
566 |
from h pres(3) pre(2)[unfolded pres(1)]
|
|
567 |
have " \<forall>ia\<in>I. \<forall>j\<in>I. ia \<noteq> j \<longrightarrow> ?p ia ## ?p j" by (auto simp:set_ins_def)
|
|
568 |
moreover from pres(1) pre(1) h pre(2)
|
|
569 |
have "(s = (\<Union> i \<in> I. ?p(i)))" by (auto simp:set_ins_def split:if_splits)
|
|
570 |
moreover from pre(3) pres(2) h
|
|
571 |
have "(\<forall> i \<in> I. cpt i (?p i))" by (auto simp:set_ins_def split:if_splits)
|
|
572 |
ultimately show "fam_conj I cpt s"
|
|
573 |
by (unfold fam_conj_def, auto)
|
|
574 |
qed
|
|
575 |
qed
|
|
576 |
|
|
577 |
lemma fam_conj_insert_simp:
|
|
578 |
assumes h:"i \<notin> I"
|
|
579 |
shows "fam_conj (insert i I) cpt = (cpt(i) \<and>* fam_conj I cpt)"
|
|
580 |
proof -
|
|
581 |
have "i \<in> insert i I" by auto
|
|
582 |
from fam_conj_elm_simp[OF this] and h
|
|
583 |
show ?thesis by auto
|
|
584 |
qed
|
|
585 |
|
|
586 |
lemmas fam_conj_simps = fam_conj_insert_simp fam_conj_zero_simp
|
|
587 |
|
|
588 |
lemma "fam_conj {0, 2, 3} cpt = (cpt 2 \<and>* cpt (0::int) \<and>* cpt 3)"
|
|
589 |
by (simp add:fam_conj_simps sep_conj_ac)
|
|
590 |
|
|
591 |
lemma fam_conj_ext_eq:
|
|
592 |
assumes h: "\<And> i . i \<in> I \<Longrightarrow> f i = g i"
|
|
593 |
shows "fam_conj I f = fam_conj I g"
|
|
594 |
proof
|
|
595 |
fix s
|
|
596 |
show "fam_conj I f s = fam_conj I g s"
|
|
597 |
by (unfold fam_conj_def, auto simp:h)
|
|
598 |
qed
|
|
599 |
|
|
600 |
end
|
|
601 |
|