author | wenzelm |
Sat, 13 May 2006 02:51:48 +0200 | |
changeset 19635 | f7aa7d174343 |
parent 16733 | 236dfafbeb63 |
child 20217 | 25b068a99d2b |
permissions | -rw-r--r-- |
13020 | 1 |
|
2 |
header {* \section{Examples} *} |
|
3 |
||
16417 | 4 |
theory OG_Examples imports OG_Syntax begin |
13020 | 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 |
|
13187 | 139 |
--{* 20 vc *} |
13020 | 140 |
apply auto |
141 |
done |
|
142 |
||
143 |
subsubsection{* The Ticket Algorithm *} |
|
144 |
||
145 |
record Ticket_mutex = |
|
146 |
num :: nat |
|
147 |
nextv :: nat |
|
148 |
turn :: "nat list" |
|
149 |
index :: nat |
|
150 |
||
151 |
lemma Ticket_mutex: |
|
152 |
"\<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 |
|
153 |
\<longrightarrow> \<acute>turn!k < \<acute>num \<and> (\<acute>turn!k =0 \<or> \<acute>turn!k\<noteq>\<acute>turn!l))\<guillemotright> \<rbrakk> |
|
154 |
\<Longrightarrow> \<parallel>- .{n=length \<acute>turn}. |
|
155 |
\<acute>index:= 0,, |
|
156 |
WHILE \<acute>index < n INV .{n=length \<acute>turn \<and> (\<forall>i<\<acute>index. \<acute>turn!i=0)}. |
|
157 |
DO \<acute>turn:= \<acute>turn[\<acute>index:=0],, \<acute>index:=\<acute>index +1 OD,, |
|
158 |
\<acute>num:=1 ,, \<acute>nextv:=1 ,, |
|
159 |
COBEGIN |
|
160 |
SCHEME [0\<le> i< n] |
|
161 |
.{\<acute>I}. |
|
162 |
WHILE True INV .{\<acute>I}. |
|
163 |
DO .{\<acute>I}. \<langle> \<acute>turn :=\<acute>turn[i:=\<acute>num],, \<acute>num:=\<acute>num+1 \<rangle>;; |
|
164 |
.{\<acute>I}. WAIT \<acute>turn!i=\<acute>nextv END;; |
|
165 |
.{\<acute>I \<and> \<acute>turn!i=\<acute>nextv}. \<acute>nextv:=\<acute>nextv+1 |
|
166 |
OD |
|
167 |
.{False}. |
|
168 |
COEND |
|
169 |
.{False}." |
|
170 |
apply oghoare |
|
171 |
--{* 35 vc *} |
|
172 |
apply simp_all |
|
13187 | 173 |
--{* 21 vc *} |
13020 | 174 |
apply(tactic {* ALLGOALS Clarify_tac *}) |
13187 | 175 |
--{* 11 vc *} |
13020 | 176 |
apply simp_all |
177 |
apply(tactic {* ALLGOALS Clarify_tac *}) |
|
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
178 |
--{* 10 subgoals left *} |
13020 | 179 |
apply(erule less_SucE) |
180 |
apply simp |
|
181 |
apply simp |
|
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
182 |
--{* 9 subgoals left *} |
13020 | 183 |
apply(case_tac "i=k") |
184 |
apply force |
|
185 |
apply simp |
|
186 |
apply(case_tac "i=l") |
|
187 |
apply force |
|
188 |
apply force |
|
189 |
--{* 8 subgoals left *} |
|
190 |
prefer 8 |
|
191 |
apply force |
|
192 |
apply force |
|
193 |
--{* 6 subgoals left *} |
|
194 |
prefer 6 |
|
195 |
apply(erule_tac x=i in allE) |
|
13601 | 196 |
apply fastsimp |
13020 | 197 |
--{* 5 subgoals left *} |
198 |
prefer 5 |
|
199 |
apply(tactic {* ALLGOALS (case_tac "j=k") *}) |
|
13187 | 200 |
--{* 10 subgoals left *} |
13020 | 201 |
apply simp_all |
13601 | 202 |
apply(erule_tac x=k in allE) |
13020 | 203 |
apply force |
13187 | 204 |
--{* 9 subgoals left *} |
13020 | 205 |
apply(case_tac "j=l") |
206 |
apply simp |
|
207 |
apply(erule_tac x=k in allE) |
|
208 |
apply(erule_tac x=k in allE) |
|
209 |
apply(erule_tac x=l in allE) |
|
210 |
apply force |
|
211 |
apply(erule_tac x=k in allE) |
|
212 |
apply(erule_tac x=k in allE) |
|
213 |
apply(erule_tac x=l in allE) |
|
214 |
apply force |
|
215 |
--{* 8 subgoals left *} |
|
216 |
apply force |
|
217 |
apply(case_tac "j=l") |
|
218 |
apply simp |
|
219 |
apply(erule_tac x=k in allE) |
|
220 |
apply(erule_tac x=l in allE) |
|
221 |
apply force |
|
222 |
apply force |
|
223 |
apply force |
|
13187 | 224 |
--{* 5 subgoals left *} |
13020 | 225 |
apply(erule_tac x=k in allE) |
226 |
apply(erule_tac x=l in allE) |
|
227 |
apply(case_tac "j=l") |
|
228 |
apply force |
|
229 |
apply force |
|
230 |
apply force |
|
13187 | 231 |
--{* 3 subgoals left *} |
13020 | 232 |
apply(erule_tac x=k in allE) |
233 |
apply(erule_tac x=l in allE) |
|
234 |
apply(case_tac "j=l") |
|
235 |
apply force |
|
236 |
apply force |
|
237 |
apply force |
|
13187 | 238 |
--{* 1 subgoals left *} |
13020 | 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 |
||
13517 | 342 |
lemma nat_lemma2: "\<lbrakk> b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \<rbrakk> \<Longrightarrow> m \<le> s" |
13020 | 343 |
proof - |
13517 | 344 |
assume "b = m*(n::nat) + t" "a = s*n + u" "t=u" |
13020 | 345 |
hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib) |
346 |
also assume "\<dots> < n" |
|
347 |
finally have "m - s < 1" by simp |
|
348 |
thus ?thesis by arith |
|
349 |
qed |
|
350 |
||
351 |
lemma mod_lemma: "\<lbrakk> (c::nat) \<le> a; a < b; b - c < n \<rbrakk> \<Longrightarrow> b mod n \<noteq> a mod n" |
|
352 |
apply(subgoal_tac "b=b div n*n + b mod n" ) |
|
353 |
prefer 2 apply (simp add: mod_div_equality [symmetric]) |
|
354 |
apply(subgoal_tac "a=a div n*n + a mod n") |
|
355 |
prefer 2 |
|
356 |
apply(simp add: mod_div_equality [symmetric]) |
|
13517 | 357 |
apply(subgoal_tac "b - a \<le> b - c") |
358 |
prefer 2 apply arith |
|
13020 | 359 |
apply(drule le_less_trans) |
360 |
back |
|
361 |
apply assumption |
|
362 |
apply(frule less_not_refl2) |
|
363 |
apply(drule less_imp_le) |
|
13517 | 364 |
apply (drule_tac m = "a" and k = n in div_le_mono) |
13020 | 365 |
apply(safe) |
13517 | 366 |
apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2, assumption, assumption) |
13020 | 367 |
apply assumption |
13517 | 368 |
apply(drule order_antisym, assumption) |
369 |
apply(rotate_tac -3) |
|
13020 | 370 |
apply(simp) |
371 |
done |
|
372 |
||
13517 | 373 |
|
13020 | 374 |
subsubsection {* Producer/Consumer Algorithm *} |
375 |
||
376 |
record Producer_consumer = |
|
377 |
ins :: nat |
|
378 |
outs :: nat |
|
379 |
li :: nat |
|
380 |
lj :: nat |
|
381 |
vx :: nat |
|
382 |
vy :: nat |
|
383 |
buffer :: "nat list" |
|
384 |
b :: "nat list" |
|
385 |
||
386 |
text {* The whole proof takes aprox. 4 minutes. *} |
|
387 |
||
388 |
lemma Producer_consumer: |
|
389 |
"\<lbrakk>INIT= \<guillemotleft>0<length a \<and> 0<length \<acute>buffer \<and> length \<acute>b=length a\<guillemotright> ; |
|
390 |
I= \<guillemotleft>(\<forall>k<\<acute>ins. \<acute>outs\<le>k \<longrightarrow> (a ! k) = \<acute>buffer ! (k mod (length \<acute>buffer))) \<and> |
|
391 |
\<acute>outs\<le>\<acute>ins \<and> \<acute>ins-\<acute>outs\<le>length \<acute>buffer\<guillemotright> ; |
|
392 |
I1= \<guillemotleft>\<acute>I \<and> \<acute>li\<le>length a\<guillemotright> ; |
|
393 |
p1= \<guillemotleft>\<acute>I1 \<and> \<acute>li=\<acute>ins\<guillemotright> ; |
|
394 |
I2 = \<guillemotleft>\<acute>I \<and> (\<forall>k<\<acute>lj. (a ! k)=(\<acute>b ! k)) \<and> \<acute>lj\<le>length a\<guillemotright> ; |
|
395 |
p2 = \<guillemotleft>\<acute>I2 \<and> \<acute>lj=\<acute>outs\<guillemotright> \<rbrakk> \<Longrightarrow> |
|
396 |
\<parallel>- .{\<acute>INIT}. |
|
397 |
\<acute>ins:=0,, \<acute>outs:=0,, \<acute>li:=0,, \<acute>lj:=0,, |
|
398 |
COBEGIN .{\<acute>p1 \<and> \<acute>INIT}. |
|
399 |
WHILE \<acute>li <length a |
|
400 |
INV .{\<acute>p1 \<and> \<acute>INIT}. |
|
401 |
DO .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a}. |
|
402 |
\<acute>vx:= (a ! \<acute>li);; |
|
403 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)}. |
|
404 |
WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;; |
|
405 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li) |
|
406 |
\<and> \<acute>ins-\<acute>outs < length \<acute>buffer}. |
|
407 |
\<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);; |
|
408 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a |
|
409 |
\<and> (a ! \<acute>li)=(\<acute>buffer ! (\<acute>ins mod (length \<acute>buffer))) |
|
410 |
\<and> \<acute>ins-\<acute>outs <length \<acute>buffer}. |
|
411 |
\<acute>ins:=\<acute>ins+1;; |
|
412 |
.{\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a}. |
|
413 |
\<acute>li:=\<acute>li+1 |
|
414 |
OD |
|
415 |
.{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a}. |
|
416 |
\<parallel> |
|
417 |
.{\<acute>p2 \<and> \<acute>INIT}. |
|
418 |
WHILE \<acute>lj < length a |
|
419 |
INV .{\<acute>p2 \<and> \<acute>INIT}. |
|
420 |
DO .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT}. |
|
421 |
WAIT \<acute>outs<\<acute>ins END;; |
|
422 |
.{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT}. |
|
423 |
\<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));; |
|
424 |
.{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}. |
|
425 |
\<acute>outs:=\<acute>outs+1;; |
|
426 |
.{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}. |
|
427 |
\<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);; |
|
428 |
.{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT}. |
|
429 |
\<acute>lj:=\<acute>lj+1 |
|
430 |
OD |
|
431 |
.{\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT}. |
|
432 |
COEND |
|
433 |
.{ \<forall>k<length a. (a ! k)=(\<acute>b ! k)}." |
|
434 |
apply oghoare |
|
435 |
--{* 138 vc *} |
|
436 |
apply(tactic {* ALLGOALS Clarify_tac *}) |
|
437 |
--{* 112 subgoals left *} |
|
438 |
apply(simp_all (no_asm)) |
|
439 |
apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *}) |
|
13601 | 440 |
--{* 930 subgoals left *} |
13020 | 441 |
apply(tactic {* ALLGOALS Clarify_tac *}) |
13601 | 442 |
apply(simp_all (asm_lr) only:length_0_conv [THEN sym]) |
443 |
--{* 44 subgoals left *} |
|
444 |
apply (simp_all (asm_lr) del:length_0_conv add: nth_list_update mod_less_divisor mod_lemma) |
|
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
445 |
--{* 32 subgoals left *} |
13020 | 446 |
apply(tactic {* ALLGOALS Clarify_tac *}) |
14757
556ce89b7d41
tactic call changed from TRYALL arith_tac to TRYALL simple_arith_tac preventing a call to presburger.
chaieb
parents:
13601
diff
changeset
|
447 |
|
556ce89b7d41
tactic call changed from TRYALL arith_tac to TRYALL simple_arith_tac preventing a call to presburger.
chaieb
parents:
13601
diff
changeset
|
448 |
apply(tactic {* TRYALL simple_arith_tac *}) |
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
449 |
--{* 9 subgoals left *} |
13020 | 450 |
apply (force simp add:less_Suc_eq) |
451 |
apply(drule sym) |
|
452 |
apply (force simp add:less_Suc_eq)+ |
|
453 |
done |
|
454 |
||
455 |
subsection {* Parameterized Examples *} |
|
456 |
||
457 |
subsubsection {* Set Elements of an Array to Zero *} |
|
458 |
||
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
459 |
record Example1 = |
13020 | 460 |
a :: "nat \<Rightarrow> nat" |
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
461 |
|
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
462 |
lemma Example1: |
13020 | 463 |
"\<parallel>- .{True}. |
464 |
COBEGIN SCHEME [0\<le>i<n] .{True}. \<acute>a:=\<acute>a (i:=0) .{\<acute>a i=0}. COEND |
|
465 |
.{\<forall>i < n. \<acute>a i = 0}." |
|
466 |
apply oghoare |
|
467 |
apply simp_all |
|
468 |
done |
|
469 |
||
470 |
text {* Same example with lists as auxiliary variables. *} |
|
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
471 |
record Example1_list = |
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
472 |
A :: "nat list" |
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
473 |
lemma Example1_list: |
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
474 |
"\<parallel>- .{n < length \<acute>A}. |
13020 | 475 |
COBEGIN |
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
476 |
SCHEME [0\<le>i<n] .{n < length \<acute>A}. \<acute>A:=\<acute>A[i:=0] .{\<acute>A!i=0}. |
13020 | 477 |
COEND |
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
478 |
.{\<forall>i < n. \<acute>A!i = 0}." |
13020 | 479 |
apply oghoare |
13187 | 480 |
apply force+ |
13020 | 481 |
done |
482 |
||
483 |
subsubsection {* Increment a Variable in Parallel *} |
|
484 |
||
15561 | 485 |
text {* First some lemmas about summation properties. *} |
486 |
(* |
|
15043 | 487 |
lemma Example2_lemma1: "!!b. j<n \<Longrightarrow> (\<Sum>i::nat<n. b i) = (0::nat) \<Longrightarrow> b j = 0 " |
13020 | 488 |
apply(induct n) |
489 |
apply simp_all |
|
490 |
apply(force simp add: less_Suc_eq) |
|
491 |
done |
|
15561 | 492 |
*) |
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
493 |
lemma Example2_lemma2_aux: "!!b. j<n \<Longrightarrow> |
15561 | 494 |
(\<Sum>i=0..<n. (b i::nat)) = |
495 |
(\<Sum>i=0..<j. b i) + b j + (\<Sum>i=0..<n-(Suc j) . b (Suc j + i))" |
|
13020 | 496 |
apply(induct n) |
497 |
apply simp_all |
|
498 |
apply(simp add:less_Suc_eq) |
|
499 |
apply(auto) |
|
500 |
apply(subgoal_tac "n - j = Suc(n- Suc j)") |
|
501 |
apply simp |
|
502 |
apply arith |
|
13187 | 503 |
done |
13020 | 504 |
|
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
505 |
lemma Example2_lemma2_aux2: |
15561 | 506 |
"!!b. j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)" |
13020 | 507 |
apply(induct j) |
15041
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
nipkow
parents:
14757
diff
changeset
|
508 |
apply (simp_all cong:setsum_cong) |
13020 | 509 |
done |
510 |
||
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
511 |
lemma Example2_lemma2: |
15561 | 512 |
"!!b. \<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)" |
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
513 |
apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux) |
15561 | 514 |
apply(erule_tac t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst) |
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
515 |
apply(frule_tac b=b in Example2_lemma2_aux) |
15561 | 516 |
apply(erule_tac t="setsum b {0..<n}" in ssubst) |
517 |
apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))") |
|
13020 | 518 |
apply(rotate_tac -1) |
519 |
apply(erule ssubst) |
|
520 |
apply(subgoal_tac "j\<le>j") |
|
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
521 |
apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2) |
13020 | 522 |
apply(rotate_tac -1) |
523 |
apply(erule ssubst) |
|
524 |
apply simp_all |
|
525 |
done |
|
526 |
||
527 |
||
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
528 |
record Example2 = |
13020 | 529 |
c :: "nat \<Rightarrow> nat" |
530 |
x :: nat |
|
13022
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
531 |
|
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
prensani
parents:
13020
diff
changeset
|
532 |
lemma Example_2: "0<n \<Longrightarrow> |
15561 | 533 |
\<parallel>- .{\<acute>x=0 \<and> (\<Sum>i=0..<n. \<acute>c i)=0}. |
13020 | 534 |
COBEGIN |
535 |
SCHEME [0\<le>i<n] |
|
15561 | 536 |
.{\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=0}. |
13020 | 537 |
\<langle> \<acute>x:=\<acute>x+(Suc 0),, \<acute>c:=\<acute>c (i:=(Suc 0)) \<rangle> |
15561 | 538 |
.{\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=(Suc 0)}. |
13020 | 539 |
COEND |
540 |
.{\<acute>x=n}." |
|
541 |
apply oghoare |
|
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
542 |
apply (simp_all cong del: strong_setsum_cong) |
13020 | 543 |
apply (tactic {* ALLGOALS Clarify_tac *}) |
16733
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
544 |
apply (simp_all cong del: strong_setsum_cong) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
545 |
apply(erule (1) Example2_lemma2) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
546 |
apply(erule (1) Example2_lemma2) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
547 |
apply(erule (1) Example2_lemma2) |
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
nipkow
parents:
16417
diff
changeset
|
548 |
apply(simp) |
13020 | 549 |
done |
550 |
||
13187 | 551 |
end |