13020
|
1 |
|
|
2 |
header {* \section{Examples} *}
|
|
3 |
|
|
4 |
theory OG_Examples = OG_Syntax:
|
|
5 |
|
|
6 |
subsection {* Mutual Exclusion *}
|
|
7 |
|
|
8 |
subsubsection {* Peterson's Algorithm I*}
|
|
9 |
|
|
10 |
text {* Eike Best. "Semantics of Sequential and Parallel Programs", page 217. *}
|
|
11 |
|
|
12 |
record Petersons_mutex_1 =
|
|
13 |
pr1 :: nat
|
|
14 |
pr2 :: nat
|
|
15 |
in1 :: bool
|
|
16 |
in2 :: bool
|
|
17 |
hold :: nat
|
|
18 |
|
|
19 |
lemma Petersons_mutex_1:
|
|
20 |
"\<parallel>- .{\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2 }.
|
|
21 |
COBEGIN .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.
|
|
22 |
WHILE True INV .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.
|
|
23 |
DO
|
|
24 |
.{\<acute>pr1=0 \<and> \<not>\<acute>in1}. \<langle> \<acute>in1:=True,,\<acute>pr1:=1 \<rangle>;;
|
|
25 |
.{\<acute>pr1=1 \<and> \<acute>in1}. \<langle> \<acute>hold:=1,,\<acute>pr1:=2 \<rangle>;;
|
|
26 |
.{\<acute>pr1=2 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)}.
|
|
27 |
AWAIT (\<not>\<acute>in2 \<or> \<not>(\<acute>hold=1)) THEN \<acute>pr1:=3 END;;
|
|
28 |
.{\<acute>pr1=3 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)}.
|
|
29 |
\<langle>\<acute>in1:=False,,\<acute>pr1:=0\<rangle>
|
|
30 |
OD .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.
|
|
31 |
\<parallel>
|
|
32 |
.{\<acute>pr2=0 \<and> \<not>\<acute>in2}.
|
|
33 |
WHILE True INV .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.
|
|
34 |
DO
|
|
35 |
.{\<acute>pr2=0 \<and> \<not>\<acute>in2}. \<langle> \<acute>in2:=True,,\<acute>pr2:=1 \<rangle>;;
|
|
36 |
.{\<acute>pr2=1 \<and> \<acute>in2}. \<langle> \<acute>hold:=2,,\<acute>pr2:=2 \<rangle>;;
|
|
37 |
.{\<acute>pr2=2 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))}.
|
|
38 |
AWAIT (\<not>\<acute>in1 \<or> \<not>(\<acute>hold=2)) THEN \<acute>pr2:=3 END;;
|
|
39 |
.{\<acute>pr2=3 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))}.
|
|
40 |
\<langle>\<acute>in2:=False,,\<acute>pr2:=0\<rangle>
|
|
41 |
OD .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.
|
|
42 |
COEND
|
|
43 |
.{\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2}."
|
|
44 |
apply oghoare
|
|
45 |
--{* 104 verification conditions. *}
|
|
46 |
apply auto
|
|
47 |
done
|
|
48 |
|
|
49 |
subsubsection {*Peterson's Algorithm II: A Busy Wait Solution *}
|
|
50 |
|
|
51 |
text {* Apt and Olderog. "Verification of sequential and concurrent Programs", page 282. *}
|
|
52 |
|
|
53 |
record Busy_wait_mutex =
|
|
54 |
flag1 :: bool
|
|
55 |
flag2 :: bool
|
|
56 |
turn :: nat
|
|
57 |
after1 :: bool
|
|
58 |
after2 :: bool
|
|
59 |
|
|
60 |
lemma Busy_wait_mutex:
|
|
61 |
"\<parallel>- .{True}.
|
|
62 |
\<acute>flag1:=False,, \<acute>flag2:=False,,
|
|
63 |
COBEGIN .{\<not>\<acute>flag1}.
|
|
64 |
WHILE True
|
|
65 |
INV .{\<not>\<acute>flag1}.
|
|
66 |
DO .{\<not>\<acute>flag1}. \<langle> \<acute>flag1:=True,,\<acute>after1:=False \<rangle>;;
|
|
67 |
.{\<acute>flag1 \<and> \<not>\<acute>after1}. \<langle> \<acute>turn:=1,,\<acute>after1:=True \<rangle>;;
|
|
68 |
.{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.
|
|
69 |
WHILE \<not>(\<acute>flag2 \<longrightarrow> \<acute>turn=2)
|
|
70 |
INV .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.
|
|
71 |
DO .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}. SKIP OD;;
|
|
72 |
.{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>flag2 \<and> \<acute>after2 \<longrightarrow> \<acute>turn=2)}.
|
|
73 |
\<acute>flag1:=False
|
|
74 |
OD
|
|
75 |
.{False}.
|
|
76 |
\<parallel>
|
|
77 |
.{\<not>\<acute>flag2}.
|
|
78 |
WHILE True
|
|
79 |
INV .{\<not>\<acute>flag2}.
|
|
80 |
DO .{\<not>\<acute>flag2}. \<langle> \<acute>flag2:=True,,\<acute>after2:=False \<rangle>;;
|
|
81 |
.{\<acute>flag2 \<and> \<not>\<acute>after2}. \<langle> \<acute>turn:=2,,\<acute>after2:=True \<rangle>;;
|
|
82 |
.{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.
|
|
83 |
WHILE \<not>(\<acute>flag1 \<longrightarrow> \<acute>turn=1)
|
|
84 |
INV .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.
|
|
85 |
DO .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}. SKIP OD;;
|
|
86 |
.{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>flag1 \<and> \<acute>after1 \<longrightarrow> \<acute>turn=1)}.
|
|
87 |
\<acute>flag2:=False
|
|
88 |
OD
|
|
89 |
.{False}.
|
|
90 |
COEND
|
|
91 |
.{False}."
|
|
92 |
apply oghoare
|
|
93 |
--{* 122 vc *}
|
|
94 |
apply auto
|
|
95 |
done
|
|
96 |
|
|
97 |
subsubsection {* Peterson's Algorithm III: A Solution using Semaphores *}
|
|
98 |
|
|
99 |
record Semaphores_mutex =
|
|
100 |
out :: bool
|
|
101 |
who :: nat
|
|
102 |
|
|
103 |
lemma Semaphores_mutex:
|
|
104 |
"\<parallel>- .{i\<noteq>j}.
|
|
105 |
\<acute>out:=True ,,
|
|
106 |
COBEGIN .{i\<noteq>j}.
|
|
107 |
WHILE True INV .{i\<noteq>j}.
|
|
108 |
DO .{i\<noteq>j}. AWAIT \<acute>out THEN \<acute>out:=False,, \<acute>who:=i END;;
|
|
109 |
.{\<not>\<acute>out \<and> \<acute>who=i \<and> i\<noteq>j}. \<acute>out:=True OD
|
|
110 |
.{False}.
|
|
111 |
\<parallel>
|
|
112 |
.{i\<noteq>j}.
|
|
113 |
WHILE True INV .{i\<noteq>j}.
|
|
114 |
DO .{i\<noteq>j}. AWAIT \<acute>out THEN \<acute>out:=False,,\<acute>who:=j END;;
|
|
115 |
.{\<not>\<acute>out \<and> \<acute>who=j \<and> i\<noteq>j}. \<acute>out:=True OD
|
|
116 |
.{False}.
|
|
117 |
COEND
|
|
118 |
.{False}."
|
|
119 |
apply oghoare
|
|
120 |
--{* 38 vc *}
|
|
121 |
apply auto
|
|
122 |
done
|
|
123 |
|
|
124 |
subsubsection {* Peterson's Algorithm III: Parameterized version: *}
|
|
125 |
|
|
126 |
lemma Semaphores_parameterized_mutex:
|
|
127 |
"0<n \<Longrightarrow> \<parallel>- .{True}.
|
|
128 |
\<acute>out:=True ,,
|
|
129 |
COBEGIN
|
|
130 |
SCHEME [0\<le> i< n]
|
|
131 |
.{True}.
|
|
132 |
WHILE True INV .{True}.
|
|
133 |
DO .{True}. AWAIT \<acute>out THEN \<acute>out:=False,, \<acute>who:=i END;;
|
|
134 |
.{\<not>\<acute>out \<and> \<acute>who=i}. \<acute>out:=True OD
|
|
135 |
.{False}.
|
|
136 |
COEND
|
|
137 |
.{False}."
|
|
138 |
apply oghoare
|
|
139 |
apply auto
|
|
140 |
done
|
|
141 |
|
|
142 |
subsubsection{* The Ticket Algorithm *}
|
|
143 |
|
|
144 |
record Ticket_mutex =
|
|
145 |
num :: nat
|
|
146 |
nextv :: nat
|
|
147 |
turn :: "nat list"
|
|
148 |
index :: nat
|
|
149 |
|
|
150 |
lemma Ticket_mutex:
|
|
151 |
"\<lbrakk> 0<n; I=\<guillemotleft>n=length \<acute>turn \<and> 0<\<acute>nextv \<and> (\<forall>k l. k<n \<and> l<n \<and> k\<noteq>l
|
|
152 |
\<longrightarrow> \<acute>turn!k < \<acute>num \<and> (\<acute>turn!k =0 \<or> \<acute>turn!k\<noteq>\<acute>turn!l))\<guillemotright> \<rbrakk>
|
|
153 |
\<Longrightarrow> \<parallel>- .{n=length \<acute>turn}.
|
|
154 |
\<acute>index:= 0,,
|
|
155 |
WHILE \<acute>index < n INV .{n=length \<acute>turn \<and> (\<forall>i<\<acute>index. \<acute>turn!i=0)}.
|
|
156 |
DO \<acute>turn:= \<acute>turn[\<acute>index:=0],, \<acute>index:=\<acute>index +1 OD,,
|
|
157 |
\<acute>num:=1 ,, \<acute>nextv:=1 ,,
|
|
158 |
COBEGIN
|
|
159 |
SCHEME [0\<le> i< n]
|
|
160 |
.{\<acute>I}.
|
|
161 |
WHILE True INV .{\<acute>I}.
|
|
162 |
DO .{\<acute>I}. \<langle> \<acute>turn :=\<acute>turn[i:=\<acute>num],, \<acute>num:=\<acute>num+1 \<rangle>;;
|
|
163 |
.{\<acute>I}. WAIT \<acute>turn!i=\<acute>nextv END;;
|
|
164 |
.{\<acute>I \<and> \<acute>turn!i=\<acute>nextv}. \<acute>nextv:=\<acute>nextv+1
|
|
165 |
OD
|
|
166 |
.{False}.
|
|
167 |
COEND
|
|
168 |
.{False}."
|
|
169 |
apply oghoare
|
|
170 |
--{* 35 vc *}
|
|
171 |
apply simp_all
|
|
172 |
apply(tactic {* ALLGOALS Clarify_tac *})
|
|
173 |
apply simp_all
|
|
174 |
apply(tactic {* ALLGOALS Clarify_tac *})
|
|
175 |
--{* 11 subgoals left *}
|
|
176 |
apply(erule less_SucE)
|
|
177 |
apply simp
|
|
178 |
apply simp
|
|
179 |
--{* 10 subgoals left *}
|
|
180 |
apply force
|
|
181 |
apply(case_tac "i=k")
|
|
182 |
apply force
|
|
183 |
apply simp
|
|
184 |
apply(case_tac "i=l")
|
|
185 |
apply force
|
|
186 |
apply force
|
|
187 |
--{* 8 subgoals left *}
|
|
188 |
prefer 8
|
|
189 |
apply force
|
|
190 |
apply force
|
|
191 |
--{* 6 subgoals left *}
|
|
192 |
prefer 6
|
|
193 |
apply(erule_tac x=i in allE)
|
|
194 |
apply force
|
|
195 |
--{* 5 subgoals left *}
|
|
196 |
prefer 5
|
|
197 |
apply(rule conjI)
|
|
198 |
apply clarify
|
|
199 |
prefer 2
|
|
200 |
apply(case_tac "j=i")
|
|
201 |
apply simp
|
|
202 |
apply simp
|
|
203 |
--{* 4 subgoals left *}
|
|
204 |
apply(tactic {* ALLGOALS (case_tac "j=k") *})
|
|
205 |
apply simp_all
|
|
206 |
apply(erule_tac x=i in allE)
|
|
207 |
apply force
|
|
208 |
apply(case_tac "j=l")
|
|
209 |
apply simp
|
|
210 |
apply(erule_tac x=k in allE)
|
|
211 |
apply(erule_tac x=k in allE)
|
|
212 |
apply(erule_tac x=l in allE)
|
|
213 |
apply force
|
|
214 |
apply(erule_tac x=k in allE)
|
|
215 |
apply(erule_tac x=k in allE)
|
|
216 |
apply(erule_tac x=l in allE)
|
|
217 |
apply force
|
|
218 |
--{* 8 subgoals left *}
|
|
219 |
apply force
|
|
220 |
apply(case_tac "j=l")
|
|
221 |
apply simp
|
|
222 |
apply(erule_tac x=k in allE)
|
|
223 |
apply(erule_tac x=l in allE)
|
|
224 |
apply force
|
|
225 |
apply force
|
|
226 |
apply force
|
|
227 |
apply(erule_tac x=k in allE)
|
|
228 |
apply(erule_tac x=l in allE)
|
|
229 |
apply(case_tac "j=l")
|
|
230 |
apply force
|
|
231 |
apply force
|
|
232 |
apply force
|
|
233 |
apply(erule_tac x=k in allE)
|
|
234 |
apply(erule_tac x=l in allE)
|
|
235 |
apply(case_tac "j=l")
|
|
236 |
apply force
|
|
237 |
apply force
|
|
238 |
apply force
|
|
239 |
apply(erule_tac x=k in allE)
|
|
240 |
apply(erule_tac x=l in allE)
|
|
241 |
apply(case_tac "j=l")
|
|
242 |
apply force
|
|
243 |
apply force
|
|
244 |
done
|
|
245 |
|
|
246 |
subsection{* Parallel Zero Search *}
|
|
247 |
|
|
248 |
text {* Synchronized Zero Search. Zero-6 *}
|
|
249 |
|
|
250 |
text {*Apt and Olderog. "Verification of sequential and concurrent Programs" page 294: *}
|
|
251 |
|
|
252 |
record Zero_search =
|
|
253 |
turn :: nat
|
|
254 |
found :: bool
|
|
255 |
x :: nat
|
|
256 |
y :: nat
|
|
257 |
|
|
258 |
lemma Zero_search:
|
|
259 |
"\<lbrakk>I1= \<guillemotleft> a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0))
|
|
260 |
\<and> (\<not>\<acute>found \<and> a<\<acute> x \<longrightarrow> f(\<acute>x)\<noteq>0) \<guillemotright> ;
|
|
261 |
I2= \<guillemotleft>\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0))
|
|
262 |
\<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0) \<guillemotright> \<rbrakk> \<Longrightarrow>
|
|
263 |
\<parallel>- .{\<exists> u. f(u)=0}.
|
|
264 |
\<acute>turn:=1,, \<acute>found:= False,,
|
|
265 |
\<acute>x:=a,, \<acute>y:=a+1 ,,
|
|
266 |
COBEGIN .{\<acute>I1}.
|
|
267 |
WHILE \<not>\<acute>found
|
|
268 |
INV .{\<acute>I1}.
|
|
269 |
DO .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.
|
|
270 |
WAIT \<acute>turn=1 END;;
|
|
271 |
.{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.
|
|
272 |
\<acute>turn:=2;;
|
|
273 |
.{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.
|
|
274 |
\<langle> \<acute>x:=\<acute>x+1,,
|
|
275 |
IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
|
|
276 |
OD;;
|
|
277 |
.{\<acute>I1 \<and> \<acute>found}.
|
|
278 |
\<acute>turn:=2
|
|
279 |
.{\<acute>I1 \<and> \<acute>found}.
|
|
280 |
\<parallel>
|
|
281 |
.{\<acute>I2}.
|
|
282 |
WHILE \<not>\<acute>found
|
|
283 |
INV .{\<acute>I2}.
|
|
284 |
DO .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.
|
|
285 |
WAIT \<acute>turn=2 END;;
|
|
286 |
.{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.
|
|
287 |
\<acute>turn:=1;;
|
|
288 |
.{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.
|
|
289 |
\<langle> \<acute>y:=(\<acute>y - 1),,
|
|
290 |
IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
|
|
291 |
OD;;
|
|
292 |
.{\<acute>I2 \<and> \<acute>found}.
|
|
293 |
\<acute>turn:=1
|
|
294 |
.{\<acute>I2 \<and> \<acute>found}.
|
|
295 |
COEND
|
|
296 |
.{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
|
|
297 |
apply oghoare
|
|
298 |
--{* 98 verification conditions *}
|
|
299 |
apply auto
|
|
300 |
--{* auto takes about 3 minutes !! *}
|
|
301 |
apply arith+
|
|
302 |
done
|
|
303 |
|
|
304 |
text {* Easier Version: without AWAIT. Apt and Olderog. page 256: *}
|
|
305 |
|
|
306 |
lemma Zero_Search_2:
|
|
307 |
"\<lbrakk>I1=\<guillemotleft> a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0))
|
|
308 |
\<and> (\<not>\<acute>found \<and> a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)\<guillemotright>;
|
|
309 |
I2= \<guillemotleft>\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0))
|
|
310 |
\<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)\<guillemotright>\<rbrakk> \<Longrightarrow>
|
|
311 |
\<parallel>- .{\<exists>u. f(u)=0}.
|
|
312 |
\<acute>found:= False,,
|
|
313 |
\<acute>x:=a,, \<acute>y:=a+1,,
|
|
314 |
COBEGIN .{\<acute>I1}.
|
|
315 |
WHILE \<not>\<acute>found
|
|
316 |
INV .{\<acute>I1}.
|
|
317 |
DO .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.
|
|
318 |
\<langle> \<acute>x:=\<acute>x+1,,IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
|
|
319 |
OD
|
|
320 |
.{\<acute>I1 \<and> \<acute>found}.
|
|
321 |
\<parallel>
|
|
322 |
.{\<acute>I2}.
|
|
323 |
WHILE \<not>\<acute>found
|
|
324 |
INV .{\<acute>I2}.
|
|
325 |
DO .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.
|
|
326 |
\<langle> \<acute>y:=(\<acute>y - 1),,IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
|
|
327 |
OD
|
|
328 |
.{\<acute>I2 \<and> \<acute>found}.
|
|
329 |
COEND
|
|
330 |
.{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
|
|
331 |
apply oghoare
|
|
332 |
--{* 20 vc *}
|
|
333 |
apply auto
|
|
334 |
--{* auto takes aprox. 2 minutes. *}
|
|
335 |
apply arith+
|
|
336 |
done
|
|
337 |
|
|
338 |
subsection {* Producer/Consumer *}
|
|
339 |
|
|
340 |
subsubsection {* Previous lemmas *}
|
|
341 |
|
|
342 |
lemma nat_lemma1: "\<lbrakk>(c::nat) \<le> a ;a < b\<rbrakk> \<Longrightarrow> b - a \<le> b - c"
|
|
343 |
by (simp split: nat_diff_split)
|
|
344 |
|
|
345 |
lemma nat_lemma2: "\<lbrakk> b = m*(n::nat) + t; a = s*n + t ; b - a < n \<rbrakk> \<Longrightarrow> m - s = 0"
|
|
346 |
proof -
|
|
347 |
assume "b = m*(n::nat) + t" and "a = s*n + t"
|
|
348 |
hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib)
|
|
349 |
also assume "\<dots> < n"
|
|
350 |
finally have "m - s < 1" by simp
|
|
351 |
thus ?thesis by arith
|
|
352 |
qed
|
|
353 |
|
|
354 |
lemma less_imp_diff_is_0: "m < Suc(n) \<Longrightarrow> m-n = 0"
|
|
355 |
by arith
|
|
356 |
lemma mod_lemma: "\<lbrakk> (c::nat) \<le> a; a < b; b - c < n \<rbrakk> \<Longrightarrow> b mod n \<noteq> a mod n"
|
|
357 |
apply(subgoal_tac "b=b div n*n + b mod n" )
|
|
358 |
prefer 2 apply (simp add: mod_div_equality [symmetric])
|
|
359 |
apply(subgoal_tac "a=a div n*n + a mod n")
|
|
360 |
prefer 2
|
|
361 |
apply(simp add: mod_div_equality [symmetric])
|
|
362 |
apply(frule nat_lemma1 , assumption)
|
|
363 |
apply(drule le_less_trans)
|
|
364 |
back
|
|
365 |
apply assumption
|
|
366 |
apply(frule less_not_refl2)
|
|
367 |
apply(drule less_imp_le)
|
|
368 |
apply (drule_tac m = "a" in div_le_mono)
|
|
369 |
apply(drule_tac m = "a div n" in le_imp_less_Suc)
|
|
370 |
apply(drule less_imp_diff_is_0)
|
|
371 |
apply(safe)
|
|
372 |
apply(simp)
|
|
373 |
apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2 , assumption)
|
|
374 |
apply assumption
|
|
375 |
apply(drule diffs0_imp_equal)
|
|
376 |
apply(simp)
|
|
377 |
apply(simp)
|
|
378 |
done
|
|
379 |
|
|
380 |
subsubsection {* Producer/Consumer Algorithm *}
|
|
381 |
|
|
382 |
record Producer_consumer =
|
|
383 |
ins :: nat
|
|
384 |
outs :: nat
|
|
385 |
li :: nat
|
|
386 |
lj :: nat
|
|
387 |
vx :: nat
|
|
388 |
vy :: nat
|
|
389 |
buffer :: "nat list"
|
|
390 |
b :: "nat list"
|
|
391 |
|
|
392 |
text {* The whole proof takes aprox. 4 minutes. *}
|
|
393 |
|
|
394 |
lemma Producer_consumer:
|
|
395 |
"\<lbrakk>INIT= \<guillemotleft>0<length a \<and> 0<length \<acute>buffer \<and> length \<acute>b=length a\<guillemotright> ;
|
|
396 |
I= \<guillemotleft>(\<forall>k<\<acute>ins. \<acute>outs\<le>k \<longrightarrow> (a ! k) = \<acute>buffer ! (k mod (length \<acute>buffer))) \<and>
|
|
397 |
\<acute>outs\<le>\<acute>ins \<and> \<acute>ins-\<acute>outs\<le>length \<acute>buffer\<guillemotright> ;
|
|
398 |
I1= \<guillemotleft>\<acute>I \<and> \<acute>li\<le>length a\<guillemotright> ;
|
|
399 |
p1= \<guillemotleft>\<acute>I1 \<and> \<acute>li=\<acute>ins\<guillemotright> ;
|
|
400 |
I2 = \<guillemotleft>\<acute>I \<and> (\<forall>k<\<acute>lj. (a ! k)=(\<acute>b ! k)) \<and> \<acute>lj\<le>length a\<guillemotright> ;
|
|
401 |
p2 = \<guillemotleft>\<acute>I2 \<and> \<acute>lj=\<acute>outs\<guillemotright> \<rbrakk> \<Longrightarrow>
|
|
402 |
\<parallel>- .{\<acute>INIT}.
|
|
403 |
\<acute>ins:=0,, \<acute>outs:=0,, \<acute>li:=0,, \<acute>lj:=0,,
|
|
404 |
COBEGIN .{\<acute>p1 \<and> \<acute>INIT}.
|
|
405 |
WHILE \<acute>li <length a
|
|
406 |
INV .{\<acute>p1 \<and> \<acute>INIT}.
|
|
407 |
DO .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a}.
|
|
408 |
\<acute>vx:= (a ! \<acute>li);;
|
|
409 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)}.
|
|
410 |
WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;;
|
|
411 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)
|
|
412 |
\<and> \<acute>ins-\<acute>outs < length \<acute>buffer}.
|
|
413 |
\<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);;
|
|
414 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a
|
|
415 |
\<and> (a ! \<acute>li)=(\<acute>buffer ! (\<acute>ins mod (length \<acute>buffer)))
|
|
416 |
\<and> \<acute>ins-\<acute>outs <length \<acute>buffer}.
|
|
417 |
\<acute>ins:=\<acute>ins+1;;
|
|
418 |
.{\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a}.
|
|
419 |
\<acute>li:=\<acute>li+1
|
|
420 |
OD
|
|
421 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a}.
|
|
422 |
\<parallel>
|
|
423 |
.{\<acute>p2 \<and> \<acute>INIT}.
|
|
424 |
WHILE \<acute>lj < length a
|
|
425 |
INV .{\<acute>p2 \<and> \<acute>INIT}.
|
|
426 |
DO .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT}.
|
|
427 |
WAIT \<acute>outs<\<acute>ins END;;
|
|
428 |
.{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT}.
|
|
429 |
\<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));;
|
|
430 |
.{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.
|
|
431 |
\<acute>outs:=\<acute>outs+1;;
|
|
432 |
.{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.
|
|
433 |
\<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);;
|
|
434 |
.{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT}.
|
|
435 |
\<acute>lj:=\<acute>lj+1
|
|
436 |
OD
|
|
437 |
.{\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT}.
|
|
438 |
COEND
|
|
439 |
.{ \<forall>k<length a. (a ! k)=(\<acute>b ! k)}."
|
|
440 |
apply oghoare
|
|
441 |
--{* 138 vc *}
|
|
442 |
apply(tactic {* ALLGOALS Clarify_tac *})
|
|
443 |
--{* 112 subgoals left *}
|
|
444 |
apply(simp_all (no_asm))
|
|
445 |
apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *})
|
|
446 |
--{* 860 subgoals left *}
|
|
447 |
apply(tactic {* ALLGOALS Clarify_tac *})
|
|
448 |
apply(simp_all only:length_0_conv [THEN sym])
|
|
449 |
--{* 36 subgoals left *}
|
|
450 |
apply (simp_all del:length_0_conv add: nth_list_update mod_less_divisor mod_lemma)
|
|
451 |
--{* 32 subgoals left *}
|
|
452 |
apply(tactic {* ALLGOALS Clarify_tac *})
|
|
453 |
apply(tactic {* TRYALL arith_tac *})
|
|
454 |
--{* 9 subgoals left *}
|
|
455 |
apply (force simp add:less_Suc_eq)
|
|
456 |
apply(drule sym)
|
|
457 |
apply (force simp add:less_Suc_eq)+
|
|
458 |
done
|
|
459 |
|
|
460 |
subsection {* Parameterized Examples *}
|
|
461 |
|
|
462 |
subsubsection {* Set Elements of an Array to Zero *}
|
|
463 |
|
|
464 |
record scheme1_vars =
|
|
465 |
a :: "nat \<Rightarrow> nat"
|
|
466 |
lemma scheme1:
|
|
467 |
"\<parallel>- .{True}.
|
|
468 |
COBEGIN SCHEME [0\<le>i<n] .{True}. \<acute>a:=\<acute>a (i:=0) .{\<acute>a i=0}. COEND
|
|
469 |
.{\<forall>i < n. \<acute>a i = 0}."
|
|
470 |
apply oghoare
|
|
471 |
apply simp_all
|
|
472 |
done
|
|
473 |
|
|
474 |
text {* Same example with lists as auxiliary variables. *}
|
|
475 |
record scheme1_list_vars =
|
|
476 |
a :: "nat list"
|
|
477 |
lemma scheme1_list:
|
|
478 |
"\<parallel>- .{n < length \<acute>a}.
|
|
479 |
COBEGIN
|
|
480 |
SCHEME [0\<le>i<n] .{n < length \<acute>a}. \<acute>a:=\<acute>a[i:=0] .{\<acute>a!i=0}.
|
|
481 |
COEND
|
|
482 |
.{\<forall>i < n. \<acute>a!i = 0}."
|
|
483 |
apply oghoare
|
|
484 |
apply simp_all
|
|
485 |
apply force+
|
|
486 |
apply clarify
|
|
487 |
apply (simp add:nth_list_update)
|
|
488 |
done
|
|
489 |
|
|
490 |
subsubsection {* Increment a Variable in Parallel *}
|
|
491 |
|
|
492 |
text {* First some lemmas about summation properties. Summation is
|
|
493 |
defined in PreList. *}
|
|
494 |
|
|
495 |
lemma scheme2_lemma1: "!!b. j<n \<Longrightarrow> (\<Sum>i<n. b i) = (0::nat) \<Longrightarrow> b j = 0 "
|
|
496 |
apply(induct n)
|
|
497 |
apply simp_all
|
|
498 |
apply(force simp add: less_Suc_eq)
|
|
499 |
done
|
|
500 |
|
|
501 |
lemma scheme2_lemma2_aux: "!!b. j<n \<Longrightarrow>
|
|
502 |
(\<Sum>i<n. (b i::nat)) = (\<Sum>i<j. b i) + b j + (\<Sum>i<n-(Suc j) . b (Suc j + i))"
|
|
503 |
apply(induct n)
|
|
504 |
apply simp_all
|
|
505 |
apply(simp add:less_Suc_eq)
|
|
506 |
apply(auto)
|
|
507 |
apply(subgoal_tac "n - j = Suc(n- Suc j)")
|
|
508 |
apply simp
|
|
509 |
apply arith
|
|
510 |
done
|
|
511 |
|
|
512 |
lemma scheme2_lemma2_aux2: "!!b. j\<le> s \<Longrightarrow> (\<Sum>i<j. (b (s:=t)) i) = (\<Sum>i<j. b i)"
|
|
513 |
apply(induct j)
|
|
514 |
apply simp_all
|
|
515 |
done
|
|
516 |
|
|
517 |
lemma scheme2_lemma2 [rule_format]:
|
|
518 |
"!!b. \<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i< n. b i)=(\<Sum>i< n. (b (j := Suc 0)) i)"
|
|
519 |
apply(frule_tac b="(b (j:=(Suc 0)))" in scheme2_lemma2_aux)
|
|
520 |
apply(erule_tac t="Summation (b(j := (Suc 0))) n" in ssubst)
|
|
521 |
apply(frule_tac b=b in scheme2_lemma2_aux)
|
|
522 |
apply(erule_tac t="Summation b n" in ssubst)
|
|
523 |
apply(subgoal_tac "Suc (Summation b j + b j + (\<Sum>i<n - Suc j. b (Suc j + i)))=(Summation b j + Suc (b j) + (\<Sum>i<n - Suc j. b (Suc j + i)))")
|
|
524 |
apply(rotate_tac -1)
|
|
525 |
apply(erule ssubst)
|
|
526 |
apply(subgoal_tac "j\<le>j")
|
|
527 |
apply(drule_tac b="b" and t="(Suc 0)" in scheme2_lemma2_aux2)
|
|
528 |
apply(rotate_tac -1)
|
|
529 |
apply(erule ssubst)
|
|
530 |
apply simp_all
|
|
531 |
done
|
|
532 |
|
|
533 |
lemma scheme2_lemma3: "!!b. \<forall>i< n. b i = (Suc 0) \<Longrightarrow> (\<Sum>i<n. b i)= n"
|
|
534 |
apply (induct n)
|
|
535 |
apply auto
|
|
536 |
done
|
|
537 |
|
|
538 |
record scheme2_vars =
|
|
539 |
c :: "nat \<Rightarrow> nat"
|
|
540 |
x :: nat
|
|
541 |
lemma scheme_2: "0<n \<Longrightarrow>
|
|
542 |
\<parallel>- .{\<acute>x=0 \<and> (\<Sum>i< n. \<acute>c i)=0}.
|
|
543 |
COBEGIN
|
|
544 |
SCHEME [0\<le>i<n]
|
|
545 |
.{\<acute>x=(\<Sum>i< n. \<acute>c i) \<and> \<acute>c i=0}.
|
|
546 |
\<langle> \<acute>x:=\<acute>x+(Suc 0),, \<acute>c:=\<acute>c (i:=(Suc 0)) \<rangle>
|
|
547 |
.{\<acute>x=(\<Sum>i< n. \<acute>c i) \<and> \<acute>c i=(Suc 0)}.
|
|
548 |
COEND
|
|
549 |
.{\<acute>x=n}."
|
|
550 |
apply oghoare
|
|
551 |
apply simp_all
|
|
552 |
apply (tactic {* ALLGOALS Clarify_tac *})
|
|
553 |
apply simp_all
|
|
554 |
apply(force elim:scheme2_lemma1)
|
|
555 |
apply(erule scheme2_lemma2)
|
|
556 |
apply simp
|
|
557 |
apply(erule scheme2_lemma2)
|
|
558 |
apply simp
|
|
559 |
apply(erule scheme2_lemma2)
|
|
560 |
apply simp
|
|
561 |
apply(force intro: scheme2_lemma3)
|
|
562 |
done
|
|
563 |
|
|
564 |
end |