| author | nipkow | 
| Tue, 22 Feb 2005 10:54:30 +0100 | |
| changeset 15543 | 0024472afce7 | 
| parent 15045 | d59f7e2e18d3 | 
| child 15561 | 045a07ac35a7 | 
| permissions | -rw-r--r-- | 
| 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  | 
|
| 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 *})
 | 
|
178  | 
--{* 11 subgoals left *}
 | 
|
179  | 
apply(erule less_SucE)  | 
|
180  | 
apply simp  | 
|
181  | 
apply simp  | 
|
182  | 
--{* 10 subgoals left *}
 | 
|
183  | 
apply force  | 
|
184  | 
apply(case_tac "i=k")  | 
|
185  | 
apply force  | 
|
186  | 
apply simp  | 
|
187  | 
apply(case_tac "i=l")  | 
|
188  | 
apply force  | 
|
189  | 
apply force  | 
|
190  | 
--{* 8 subgoals left *}
 | 
|
191  | 
prefer 8  | 
|
192  | 
apply force  | 
|
193  | 
apply force  | 
|
194  | 
--{* 6 subgoals left *}
 | 
|
195  | 
prefer 6  | 
|
196  | 
apply(erule_tac x=i in allE)  | 
|
| 13601 | 197  | 
apply fastsimp  | 
| 13020 | 198  | 
--{* 5 subgoals left *}
 | 
199  | 
prefer 5  | 
|
200  | 
apply(tactic {* ALLGOALS (case_tac "j=k") *})
 | 
|
| 13187 | 201  | 
--{* 10 subgoals left *}
 | 
| 13020 | 202  | 
apply simp_all  | 
| 13601 | 203  | 
apply(erule_tac x=k in allE)  | 
| 13020 | 204  | 
apply force  | 
| 13187 | 205  | 
--{* 9 subgoals left *}
 | 
| 13020 | 206  | 
apply(case_tac "j=l")  | 
207  | 
apply simp  | 
|
208  | 
apply(erule_tac x=k in allE)  | 
|
209  | 
apply(erule_tac x=k in allE)  | 
|
210  | 
apply(erule_tac x=l in allE)  | 
|
211  | 
apply force  | 
|
212  | 
apply(erule_tac x=k in allE)  | 
|
213  | 
apply(erule_tac x=k in allE)  | 
|
214  | 
apply(erule_tac x=l in allE)  | 
|
215  | 
apply force  | 
|
216  | 
--{* 8 subgoals left *}
 | 
|
217  | 
apply force  | 
|
218  | 
apply(case_tac "j=l")  | 
|
219  | 
apply simp  | 
|
220  | 
apply(erule_tac x=k in allE)  | 
|
221  | 
apply(erule_tac x=l in allE)  | 
|
222  | 
apply force  | 
|
223  | 
apply force  | 
|
224  | 
apply force  | 
|
| 13187 | 225  | 
--{* 5 subgoals left *}
 | 
| 13020 | 226  | 
apply(erule_tac x=k in allE)  | 
227  | 
apply(erule_tac x=l in allE)  | 
|
228  | 
apply(case_tac "j=l")  | 
|
229  | 
apply force  | 
|
230  | 
apply force  | 
|
231  | 
apply force  | 
|
| 13187 | 232  | 
--{* 3 subgoals left *}
 | 
| 13020 | 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  | 
|
| 13187 | 239  | 
--{* 1 subgoals left *}
 | 
| 13020 | 240  | 
apply(erule_tac x=k in allE)  | 
241  | 
apply(erule_tac x=l in allE)  | 
|
242  | 
apply(case_tac "j=l")  | 
|
243  | 
apply force  | 
|
244  | 
apply force  | 
|
245  | 
done  | 
|
246  | 
||
247  | 
subsection{* Parallel Zero Search *}
 | 
|
248  | 
||
249  | 
text {* Synchronized Zero Search. Zero-6 *}
 | 
|
250  | 
||
251  | 
text {*Apt and Olderog. "Verification of sequential and concurrent Programs" page 294: *}
 | 
|
252  | 
||
253  | 
record Zero_search =  | 
|
254  | 
turn :: nat  | 
|
255  | 
found :: bool  | 
|
256  | 
x :: nat  | 
|
257  | 
y :: nat  | 
|
258  | 
||
259  | 
lemma Zero_search:  | 
|
260  | 
"\<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))  | 
|
261  | 
\<and> (\<not>\<acute>found \<and> a<\<acute> x \<longrightarrow> f(\<acute>x)\<noteq>0) \<guillemotright> ;  | 
|
262  | 
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))  | 
|
263  | 
\<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0) \<guillemotright> \<rbrakk> \<Longrightarrow>  | 
|
264  | 
  \<parallel>- .{\<exists> u. f(u)=0}.  
 | 
|
265  | 
\<acute>turn:=1,, \<acute>found:= False,,  | 
|
266  | 
\<acute>x:=a,, \<acute>y:=a+1 ,,  | 
|
267  | 
  COBEGIN .{\<acute>I1}.  
 | 
|
268  | 
WHILE \<not>\<acute>found  | 
|
269  | 
       INV .{\<acute>I1}.  
 | 
|
270  | 
       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)}.  
 | 
|
271  | 
WAIT \<acute>turn=1 END;;  | 
|
272  | 
          .{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)}.  
 | 
|
273  | 
\<acute>turn:=2;;  | 
|
274  | 
          .{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)}.    
 | 
|
275  | 
\<langle> \<acute>x:=\<acute>x+1,,  | 
|
276  | 
IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  | 
|
277  | 
OD;;  | 
|
278  | 
       .{\<acute>I1  \<and> \<acute>found}.  
 | 
|
279  | 
\<acute>turn:=2  | 
|
280  | 
       .{\<acute>I1 \<and> \<acute>found}.  
 | 
|
281  | 
\<parallel>  | 
|
282  | 
      .{\<acute>I2}.  
 | 
|
283  | 
WHILE \<not>\<acute>found  | 
|
284  | 
       INV .{\<acute>I2}.  
 | 
|
285  | 
       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)}.  
 | 
|
286  | 
WAIT \<acute>turn=2 END;;  | 
|
287  | 
          .{\<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)}.  
 | 
|
288  | 
\<acute>turn:=1;;  | 
|
289  | 
          .{\<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)}.  
 | 
|
290  | 
\<langle> \<acute>y:=(\<acute>y - 1),,  | 
|
291  | 
IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  | 
|
292  | 
OD;;  | 
|
293  | 
       .{\<acute>I2 \<and> \<acute>found}.  
 | 
|
294  | 
\<acute>turn:=1  | 
|
295  | 
       .{\<acute>I2 \<and> \<acute>found}.  
 | 
|
296  | 
COEND  | 
|
297  | 
  .{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
 | 
|
298  | 
apply oghoare  | 
|
299  | 
--{* 98 verification conditions *}
 | 
|
300  | 
apply auto  | 
|
301  | 
--{* auto takes about 3 minutes !! *}
 | 
|
302  | 
apply arith+  | 
|
303  | 
done  | 
|
304  | 
||
305  | 
text {* Easier Version: without AWAIT.  Apt and Olderog. page 256: *}
 | 
|
306  | 
||
307  | 
lemma Zero_Search_2:  | 
|
308  | 
"\<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))  | 
|
309  | 
\<and> (\<not>\<acute>found \<and> a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)\<guillemotright>;  | 
|
310  | 
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))  | 
|
311  | 
\<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)\<guillemotright>\<rbrakk> \<Longrightarrow>  | 
|
312  | 
  \<parallel>- .{\<exists>u. f(u)=0}.  
 | 
|
313  | 
\<acute>found:= False,,  | 
|
314  | 
\<acute>x:=a,, \<acute>y:=a+1,,  | 
|
315  | 
  COBEGIN .{\<acute>I1}.  
 | 
|
316  | 
WHILE \<not>\<acute>found  | 
|
317  | 
       INV .{\<acute>I1}.  
 | 
|
318  | 
       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)}.  
 | 
|
319  | 
\<langle> \<acute>x:=\<acute>x+1,,IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  | 
|
320  | 
OD  | 
|
321  | 
       .{\<acute>I1 \<and> \<acute>found}.  
 | 
|
322  | 
\<parallel>  | 
|
323  | 
      .{\<acute>I2}.  
 | 
|
324  | 
WHILE \<not>\<acute>found  | 
|
325  | 
       INV .{\<acute>I2}.  
 | 
|
326  | 
       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)}.  
 | 
|
327  | 
\<langle> \<acute>y:=(\<acute>y - 1),,IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  | 
|
328  | 
OD  | 
|
329  | 
       .{\<acute>I2 \<and> \<acute>found}.  
 | 
|
330  | 
COEND  | 
|
331  | 
  .{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
 | 
|
332  | 
apply oghoare  | 
|
333  | 
--{* 20 vc *}
 | 
|
334  | 
apply auto  | 
|
335  | 
--{* auto takes aprox. 2 minutes. *}
 | 
|
336  | 
apply arith+  | 
|
337  | 
done  | 
|
338  | 
||
339  | 
subsection {* Producer/Consumer *}
 | 
|
340  | 
||
341  | 
subsubsection {* Previous lemmas *}
 | 
|
342  | 
||
| 13517 | 343  | 
lemma nat_lemma2: "\<lbrakk> b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \<rbrakk> \<Longrightarrow> m \<le> s"  | 
| 13020 | 344  | 
proof -  | 
| 13517 | 345  | 
assume "b = m*(n::nat) + t" "a = s*n + u" "t=u"  | 
| 13020 | 346  | 
hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib)  | 
347  | 
also assume "\<dots> < n"  | 
|
348  | 
finally have "m - s < 1" by simp  | 
|
349  | 
thus ?thesis by arith  | 
|
350  | 
qed  | 
|
351  | 
||
352  | 
lemma mod_lemma: "\<lbrakk> (c::nat) \<le> a; a < b; b - c < n \<rbrakk> \<Longrightarrow> b mod n \<noteq> a mod n"  | 
|
353  | 
apply(subgoal_tac "b=b div n*n + b mod n" )  | 
|
354  | 
prefer 2 apply (simp add: mod_div_equality [symmetric])  | 
|
355  | 
apply(subgoal_tac "a=a div n*n + a mod n")  | 
|
356  | 
prefer 2  | 
|
357  | 
apply(simp add: mod_div_equality [symmetric])  | 
|
| 13517 | 358  | 
apply(subgoal_tac "b - a \<le> b - c")  | 
359  | 
prefer 2 apply arith  | 
|
| 13020 | 360  | 
apply(drule le_less_trans)  | 
361  | 
back  | 
|
362  | 
apply assumption  | 
|
363  | 
apply(frule less_not_refl2)  | 
|
364  | 
apply(drule less_imp_le)  | 
|
| 13517 | 365  | 
apply (drule_tac m = "a" and k = n in div_le_mono)  | 
| 13020 | 366  | 
apply(safe)  | 
| 13517 | 367  | 
apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2, assumption, assumption)  | 
| 13020 | 368  | 
apply assumption  | 
| 13517 | 369  | 
apply(drule order_antisym, assumption)  | 
370  | 
apply(rotate_tac -3)  | 
|
| 13020 | 371  | 
apply(simp)  | 
372  | 
done  | 
|
373  | 
||
| 13517 | 374  | 
|
| 13020 | 375  | 
subsubsection {* Producer/Consumer Algorithm *}
 | 
376  | 
||
377  | 
record Producer_consumer =  | 
|
378  | 
ins :: nat  | 
|
379  | 
outs :: nat  | 
|
380  | 
li :: nat  | 
|
381  | 
lj :: nat  | 
|
382  | 
vx :: nat  | 
|
383  | 
vy :: nat  | 
|
384  | 
buffer :: "nat list"  | 
|
385  | 
b :: "nat list"  | 
|
386  | 
||
387  | 
text {* The whole proof takes aprox. 4 minutes. *}
 | 
|
388  | 
||
389  | 
lemma Producer_consumer:  | 
|
390  | 
"\<lbrakk>INIT= \<guillemotleft>0<length a \<and> 0<length \<acute>buffer \<and> length \<acute>b=length a\<guillemotright> ;  | 
|
391  | 
I= \<guillemotleft>(\<forall>k<\<acute>ins. \<acute>outs\<le>k \<longrightarrow> (a ! k) = \<acute>buffer ! (k mod (length \<acute>buffer))) \<and>  | 
|
392  | 
\<acute>outs\<le>\<acute>ins \<and> \<acute>ins-\<acute>outs\<le>length \<acute>buffer\<guillemotright> ;  | 
|
393  | 
I1= \<guillemotleft>\<acute>I \<and> \<acute>li\<le>length a\<guillemotright> ;  | 
|
394  | 
p1= \<guillemotleft>\<acute>I1 \<and> \<acute>li=\<acute>ins\<guillemotright> ;  | 
|
395  | 
I2 = \<guillemotleft>\<acute>I \<and> (\<forall>k<\<acute>lj. (a ! k)=(\<acute>b ! k)) \<and> \<acute>lj\<le>length a\<guillemotright> ;  | 
|
396  | 
p2 = \<guillemotleft>\<acute>I2 \<and> \<acute>lj=\<acute>outs\<guillemotright> \<rbrakk> \<Longrightarrow>  | 
|
397  | 
  \<parallel>- .{\<acute>INIT}.  
 | 
|
398  | 
\<acute>ins:=0,, \<acute>outs:=0,, \<acute>li:=0,, \<acute>lj:=0,,  | 
|
399  | 
 COBEGIN .{\<acute>p1 \<and> \<acute>INIT}. 
 | 
|
400  | 
WHILE \<acute>li <length a  | 
|
401  | 
     INV .{\<acute>p1 \<and> \<acute>INIT}.   
 | 
|
402  | 
   DO .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a}.  
 | 
|
403  | 
\<acute>vx:= (a ! \<acute>li);;  | 
|
404  | 
      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)}. 
 | 
|
405  | 
WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;;  | 
|
406  | 
      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li) 
 | 
|
407  | 
\<and> \<acute>ins-\<acute>outs < length \<acute>buffer}.  | 
|
408  | 
\<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);;  | 
|
409  | 
      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a 
 | 
|
410  | 
\<and> (a ! \<acute>li)=(\<acute>buffer ! (\<acute>ins mod (length \<acute>buffer)))  | 
|
411  | 
\<and> \<acute>ins-\<acute>outs <length \<acute>buffer}.  | 
|
412  | 
\<acute>ins:=\<acute>ins+1;;  | 
|
413  | 
      .{\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a}.  
 | 
|
414  | 
\<acute>li:=\<acute>li+1  | 
|
415  | 
OD  | 
|
416  | 
  .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a}.  
 | 
|
417  | 
\<parallel>  | 
|
418  | 
  .{\<acute>p2 \<and> \<acute>INIT}.  
 | 
|
419  | 
WHILE \<acute>lj < length a  | 
|
420  | 
     INV .{\<acute>p2 \<and> \<acute>INIT}.  
 | 
|
421  | 
   DO .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT}.  
 | 
|
422  | 
WAIT \<acute>outs<\<acute>ins END;;  | 
|
423  | 
      .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT}.  
 | 
|
424  | 
\<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));;  | 
|
425  | 
      .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.  
 | 
|
426  | 
\<acute>outs:=\<acute>outs+1;;  | 
|
427  | 
      .{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.  
 | 
|
428  | 
\<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);;  | 
|
429  | 
      .{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT}.  
 | 
|
430  | 
\<acute>lj:=\<acute>lj+1  | 
|
431  | 
OD  | 
|
432  | 
  .{\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT}.  
 | 
|
433  | 
COEND  | 
|
434  | 
 .{ \<forall>k<length a. (a ! k)=(\<acute>b ! k)}."
 | 
|
435  | 
apply oghoare  | 
|
436  | 
--{* 138 vc  *}
 | 
|
437  | 
apply(tactic {* ALLGOALS Clarify_tac *})
 | 
|
438  | 
--{* 112 subgoals left *}
 | 
|
439  | 
apply(simp_all (no_asm))  | 
|
440  | 
apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *})
 | 
|
| 13601 | 441  | 
--{* 930 subgoals left *}
 | 
| 13020 | 442  | 
apply(tactic {* ALLGOALS Clarify_tac *})
 | 
| 13601 | 443  | 
apply(simp_all (asm_lr) only:length_0_conv [THEN sym])  | 
444  | 
--{* 44 subgoals left *}
 | 
|
445  | 
apply (simp_all (asm_lr) del:length_0_conv add: nth_list_update mod_less_divisor mod_lemma)  | 
|
446  | 
--{* 33 subgoals left *}
 | 
|
| 13020 | 447  | 
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
 | 
448  | 
|
| 
 
556ce89b7d41
tactic call changed from TRYALL arith_tac to TRYALL simple_arith_tac preventing a call to presburger.
 
chaieb 
parents: 
13601 
diff
changeset
 | 
449  | 
ML "set Presburger.trace"  | 
| 
 
556ce89b7d41
tactic call changed from TRYALL arith_tac to TRYALL simple_arith_tac preventing a call to presburger.
 
chaieb 
parents: 
13601 
diff
changeset
 | 
450  | 
apply(tactic {* TRYALL simple_arith_tac *})
 | 
| 13601 | 451  | 
--{* 10 subgoals left *}
 | 
| 13020 | 452  | 
apply (force simp add:less_Suc_eq)  | 
453  | 
apply(drule sym)  | 
|
454  | 
apply (force simp add:less_Suc_eq)+  | 
|
455  | 
done  | 
|
456  | 
||
457  | 
subsection {* Parameterized Examples *}
 | 
|
458  | 
||
459  | 
subsubsection {* Set Elements of an Array to Zero *}
 | 
|
460  | 
||
| 
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  | 
record Example1 =  | 
| 13020 | 462  | 
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
 | 
463  | 
|
| 
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
464  | 
lemma Example1:  | 
| 13020 | 465  | 
 "\<parallel>- .{True}.
 | 
466  | 
   COBEGIN SCHEME [0\<le>i<n] .{True}. \<acute>a:=\<acute>a (i:=0) .{\<acute>a i=0}. COEND 
 | 
|
467  | 
  .{\<forall>i < n. \<acute>a i = 0}."
 | 
|
468  | 
apply oghoare  | 
|
469  | 
apply simp_all  | 
|
470  | 
done  | 
|
471  | 
||
472  | 
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
 | 
473  | 
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
 | 
474  | 
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
 | 
475  | 
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
 | 
476  | 
 "\<parallel>- .{n < length \<acute>A}. 
 | 
| 13020 | 477  | 
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
 | 
478  | 
     SCHEME [0\<le>i<n] .{n < length \<acute>A}. \<acute>A:=\<acute>A[i:=0] .{\<acute>A!i=0}. 
 | 
| 13020 | 479  | 
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
 | 
480  | 
    .{\<forall>i < n. \<acute>A!i = 0}."
 | 
| 13020 | 481  | 
apply oghoare  | 
| 13187 | 482  | 
apply force+  | 
| 13020 | 483  | 
done  | 
484  | 
||
485  | 
subsubsection {* Increment a Variable in Parallel *}
 | 
|
486  | 
||
487  | 
text {* First some lemmas about summation properties. Summation is
 | 
|
488  | 
defined in PreList. *}  | 
|
489  | 
||
| 15043 | 490  | 
lemma Example2_lemma1: "!!b. j<n \<Longrightarrow> (\<Sum>i::nat<n. b i) = (0::nat) \<Longrightarrow> b j = 0 "  | 
| 13020 | 491  | 
apply(induct n)  | 
492  | 
apply simp_all  | 
|
493  | 
apply(force simp add: less_Suc_eq)  | 
|
494  | 
done  | 
|
495  | 
||
| 
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
 | 
496  | 
lemma Example2_lemma2_aux: "!!b. j<n \<Longrightarrow>  | 
| 13020 | 497  | 
(\<Sum>i<n. (b i::nat)) = (\<Sum>i<j. b i) + b j + (\<Sum>i<n-(Suc j) . b (Suc j + i))"  | 
498  | 
apply(induct n)  | 
|
499  | 
apply simp_all  | 
|
500  | 
apply(simp add:less_Suc_eq)  | 
|
501  | 
apply(auto)  | 
|
502  | 
apply(subgoal_tac "n - j = Suc(n- Suc j)")  | 
|
503  | 
apply simp  | 
|
504  | 
apply arith  | 
|
| 13187 | 505  | 
done  | 
| 13020 | 506  | 
|
| 
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
 | 
507  | 
lemma Example2_lemma2_aux2:  | 
| 15043 | 508  | 
"!!b. j\<le> s \<Longrightarrow> (\<Sum>i::nat<j. (b (s:=t)) i) = (\<Sum>i<j. b i)"  | 
| 13020 | 509  | 
apply(induct j)  | 
| 
15041
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14757 
diff
changeset
 | 
510  | 
apply (simp_all cong:setsum_cong)  | 
| 13020 | 511  | 
done  | 
512  | 
||
| 
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  | 
lemma Example2_lemma2:  | 
| 15043 | 514  | 
"!!b. \<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat< n. b i)=(\<Sum>i< 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
 | 
515  | 
apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)  | 
| 15045 | 516  | 
apply(erule_tac  t="setsum (b(j := (Suc 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
 | 
517  | 
apply(frule_tac b=b in Example2_lemma2_aux)  | 
| 15045 | 518  | 
apply(erule_tac  t="setsum b {..<n}" in ssubst)
 | 
519  | 
apply(subgoal_tac "Suc (setsum b {..<j} + b j + (\<Sum>i<n - Suc j. b (Suc j + i)))=(setsum b {..<j} + Suc (b j) + (\<Sum>i<n - Suc j. b (Suc j + i)))")
 | 
|
| 13020 | 520  | 
apply(rotate_tac -1)  | 
521  | 
apply(erule ssubst)  | 
|
522  | 
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
 | 
523  | 
apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2)  | 
| 13020 | 524  | 
apply(rotate_tac -1)  | 
525  | 
apply(erule ssubst)  | 
|
526  | 
apply simp_all  | 
|
527  | 
done  | 
|
528  | 
||
| 
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
 | 
529  | 
lemma Example2_lemma3: "!!b. \<forall>i< n. b i = (Suc 0) \<Longrightarrow> (\<Sum>i<n. b i)= n"  | 
| 13020 | 530  | 
apply (induct n)  | 
531  | 
apply auto  | 
|
532  | 
done  | 
|
533  | 
||
| 
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
 | 
534  | 
record Example2 =  | 
| 13020 | 535  | 
c :: "nat \<Rightarrow> nat"  | 
536  | 
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
 | 
537  | 
|
| 
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
538  | 
lemma Example_2: "0<n \<Longrightarrow>  | 
| 13020 | 539  | 
 \<parallel>- .{\<acute>x=0 \<and> (\<Sum>i< n. \<acute>c i)=0}.  
 | 
540  | 
COBEGIN  | 
|
541  | 
SCHEME [0\<le>i<n]  | 
|
542  | 
  .{\<acute>x=(\<Sum>i< n. \<acute>c i) \<and> \<acute>c i=0}. 
 | 
|
543  | 
\<langle> \<acute>x:=\<acute>x+(Suc 0),, \<acute>c:=\<acute>c (i:=(Suc 0)) \<rangle>  | 
|
544  | 
  .{\<acute>x=(\<Sum>i< n. \<acute>c i) \<and> \<acute>c i=(Suc 0)}.
 | 
|
545  | 
COEND  | 
|
546  | 
 .{\<acute>x=n}."
 | 
|
547  | 
apply oghoare  | 
|
548  | 
apply simp_all  | 
|
549  | 
apply (tactic {* ALLGOALS Clarify_tac *})
 | 
|
550  | 
apply simp_all  | 
|
| 
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
 | 
551  | 
apply(erule Example2_lemma2)  | 
| 13020 | 552  | 
apply simp  | 
| 
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
 | 
553  | 
apply(erule Example2_lemma2)  | 
| 13020 | 554  | 
apply simp  | 
| 
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
 | 
555  | 
apply(erule Example2_lemma2)  | 
| 13020 | 556  | 
apply simp  | 
| 
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
 | 
557  | 
apply(force intro: Example2_lemma3)  | 
| 13020 | 558  | 
done  | 
559  | 
||
| 13187 | 560  | 
end  |