| author | haftmann | 
| Fri, 27 Aug 2010 10:55:20 +0200 | |
| changeset 38794 | 2d638e963357 | 
| parent 34233 | 156c42518cfc | 
| child 41842 | d8f76db6a207 | 
| permissions | -rw-r--r-- | 
| 13020 | 1  | 
header {* \section{Examples} *}
 | 
2  | 
||
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
3  | 
theory RG_Examples  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
4  | 
imports RG_Syntax  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
5  | 
begin  | 
| 13020 | 6  | 
|
7  | 
lemmas definitions [simp]= stable_def Pre_def Rely_def Guar_def Post_def Com_def  | 
|
8  | 
||
9  | 
subsection {* Set Elements of an Array to Zero *}
 | 
|
10  | 
||
11  | 
lemma le_less_trans2: "\<lbrakk>(j::nat)<k; i\<le> j\<rbrakk> \<Longrightarrow> i<k"  | 
|
12  | 
by simp  | 
|
13  | 
||
14  | 
lemma add_le_less_mono: "\<lbrakk> (a::nat) < c; b\<le>d \<rbrakk> \<Longrightarrow> a + b < c + d"  | 
|
15  | 
by simp  | 
|
16  | 
||
17  | 
record Example1 =  | 
|
18  | 
A :: "nat list"  | 
|
19  | 
||
20  | 
lemma Example1:  | 
|
21  | 
"\<turnstile> COBEGIN  | 
|
22  | 
SCHEME [0 \<le> i < n]  | 
|
23  | 
(\<acute>A := \<acute>A [i := 0],  | 
|
24  | 
\<lbrace> n < length \<acute>A \<rbrace>,  | 
|
25  | 
\<lbrace> length \<ordmasculine>A = length \<ordfeminine>A \<and> \<ordmasculine>A ! i = \<ordfeminine>A ! i \<rbrace>,  | 
|
26  | 
\<lbrace> length \<ordmasculine>A = length \<ordfeminine>A \<and> (\<forall>j<n. i \<noteq> j \<longrightarrow> \<ordmasculine>A ! j = \<ordfeminine>A ! j) \<rbrace>,  | 
|
27  | 
\<lbrace> \<acute>A ! i = 0 \<rbrace>)  | 
|
28  | 
COEND  | 
|
29  | 
SAT [\<lbrace> n < length \<acute>A \<rbrace>, \<lbrace> \<ordmasculine>A = \<ordfeminine>A \<rbrace>, \<lbrace> True \<rbrace>, \<lbrace> \<forall>i < n. \<acute>A ! i = 0 \<rbrace>]"  | 
|
30  | 
apply(rule Parallel)  | 
|
| 15102 | 31  | 
apply (auto intro!: Basic)  | 
| 13020 | 32  | 
done  | 
33  | 
||
34  | 
lemma Example1_parameterized:  | 
|
35  | 
"k < t \<Longrightarrow>  | 
|
36  | 
\<turnstile> COBEGIN  | 
|
37  | 
SCHEME [k*n\<le>i<(Suc k)*n] (\<acute>A:=\<acute>A[i:=0],  | 
|
38  | 
\<lbrace>t*n < length \<acute>A\<rbrace>,  | 
|
39  | 
\<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> \<ordmasculine>A!i = \<ordfeminine>A!i\<rbrace>,  | 
|
40  | 
\<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> (\<forall>j<length \<ordmasculine>A . i\<noteq>j \<longrightarrow> \<ordmasculine>A!j = \<ordfeminine>A!j)\<rbrace>,  | 
|
41  | 
\<lbrace>\<acute>A!i=0\<rbrace>)  | 
|
42  | 
COEND  | 
|
43  | 
SAT [\<lbrace>t*n < length \<acute>A\<rbrace>,  | 
|
44  | 
\<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> (\<forall>i<n. \<ordmasculine>A!(k*n+i)=\<ordfeminine>A!(k*n+i))\<rbrace>,  | 
|
45  | 
\<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and>  | 
|
46  | 
(\<forall>i<length \<ordmasculine>A . (i<k*n \<longrightarrow> \<ordmasculine>A!i = \<ordfeminine>A!i) \<and> ((Suc k)*n \<le> i\<longrightarrow> \<ordmasculine>A!i = \<ordfeminine>A!i))\<rbrace>,  | 
|
47  | 
\<lbrace>\<forall>i<n. \<acute>A!(k*n+i) = 0\<rbrace>]"  | 
|
48  | 
apply(rule Parallel)  | 
|
| 15102 | 49  | 
apply auto  | 
50  | 
apply(erule_tac x="k*n +i" in allE)  | 
|
51  | 
apply(subgoal_tac "k*n+i <length (A b)")  | 
|
| 13020 | 52  | 
apply force  | 
| 15102 | 53  | 
apply(erule le_less_trans2)  | 
54  | 
apply(case_tac t,simp+)  | 
|
55  | 
apply (simp add:add_commute)  | 
|
56  | 
apply(simp add: add_le_mono)  | 
|
| 13020 | 57  | 
apply(rule Basic)  | 
58  | 
apply simp  | 
|
59  | 
apply clarify  | 
|
60  | 
apply (subgoal_tac "k*n+i< length (A x)")  | 
|
61  | 
apply simp  | 
|
62  | 
apply(erule le_less_trans2)  | 
|
63  | 
apply(case_tac t,simp+)  | 
|
64  | 
apply (simp add:add_commute)  | 
|
| 15102 | 65  | 
apply(rule add_le_mono, auto)  | 
| 13020 | 66  | 
done  | 
67  | 
||
| 15102 | 68  | 
|
| 13020 | 69  | 
subsection {* Increment a Variable in Parallel *}
 | 
70  | 
||
71  | 
subsubsection {* Two components *}
 | 
|
72  | 
||
73  | 
record Example2 =  | 
|
74  | 
x :: nat  | 
|
75  | 
c_0 :: nat  | 
|
76  | 
c_1 :: nat  | 
|
77  | 
||
78  | 
lemma Example2:  | 
|
79  | 
"\<turnstile> COBEGIN  | 
|
80  | 
(\<langle> \<acute>x:=\<acute>x+1;; \<acute>c_0:=\<acute>c_0 + 1 \<rangle>,  | 
|
81  | 
\<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_0=0\<rbrace>,  | 
|
82  | 
\<lbrace>\<ordmasculine>c_0 = \<ordfeminine>c_0 \<and>  | 
|
83  | 
(\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1  | 
|
84  | 
\<longrightarrow> \<ordfeminine>x = \<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  | 
|
85  | 
\<lbrace>\<ordmasculine>c_1 = \<ordfeminine>c_1 \<and>  | 
|
86  | 
(\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1  | 
|
87  | 
\<longrightarrow> \<ordfeminine>x =\<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  | 
|
88  | 
\<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_0=1 \<rbrace>)  | 
|
89  | 
\<parallel>  | 
|
90  | 
(\<langle> \<acute>x:=\<acute>x+1;; \<acute>c_1:=\<acute>c_1+1 \<rangle>,  | 
|
91  | 
\<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_1=0 \<rbrace>,  | 
|
92  | 
\<lbrace>\<ordmasculine>c_1 = \<ordfeminine>c_1 \<and>  | 
|
93  | 
(\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1  | 
|
94  | 
\<longrightarrow> \<ordfeminine>x = \<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  | 
|
95  | 
\<lbrace>\<ordmasculine>c_0 = \<ordfeminine>c_0 \<and>  | 
|
96  | 
(\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1  | 
|
97  | 
\<longrightarrow> \<ordfeminine>x =\<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  | 
|
98  | 
\<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_1=1\<rbrace>)  | 
|
99  | 
COEND  | 
|
100  | 
SAT [\<lbrace>\<acute>x=0 \<and> \<acute>c_0=0 \<and> \<acute>c_1=0\<rbrace>,  | 
|
101  | 
\<lbrace>\<ordmasculine>x=\<ordfeminine>x \<and> \<ordmasculine>c_0= \<ordfeminine>c_0 \<and> \<ordmasculine>c_1=\<ordfeminine>c_1\<rbrace>,  | 
|
102  | 
\<lbrace>True\<rbrace>,  | 
|
103  | 
\<lbrace>\<acute>x=2\<rbrace>]"  | 
|
104  | 
apply(rule Parallel)  | 
|
105  | 
apply simp_all  | 
|
106  | 
apply clarify  | 
|
107  | 
apply(case_tac i)  | 
|
108  | 
apply simp  | 
|
| 15102 | 109  | 
apply(rule conjI)  | 
| 13020 | 110  | 
apply clarify  | 
111  | 
apply simp  | 
|
112  | 
apply clarify  | 
|
113  | 
apply simp  | 
|
114  | 
apply(case_tac j,simp)  | 
|
115  | 
apply simp  | 
|
116  | 
apply simp  | 
|
| 15102 | 117  | 
apply(rule conjI)  | 
| 13020 | 118  | 
apply clarify  | 
119  | 
apply simp  | 
|
120  | 
apply clarify  | 
|
121  | 
apply simp  | 
|
| 13187 | 122  | 
apply(subgoal_tac "j=0")  | 
| 
34233
 
156c42518cfc
removed more asm_rl's - unfortunately slowdown of 1 min.
 
nipkow 
parents: 
32621 
diff
changeset
 | 
123  | 
apply (simp)  | 
| 13187 | 124  | 
apply arith  | 
| 13020 | 125  | 
apply clarify  | 
126  | 
apply(case_tac i,simp,simp)  | 
|
| 
34233
 
156c42518cfc
removed more asm_rl's - unfortunately slowdown of 1 min.
 
nipkow 
parents: 
32621 
diff
changeset
 | 
127  | 
apply clarify  | 
| 13020 | 128  | 
apply simp  | 
129  | 
apply(erule_tac x=0 in all_dupE)  | 
|
130  | 
apply(erule_tac x=1 in allE,simp)  | 
|
131  | 
apply clarify  | 
|
132  | 
apply(case_tac i,simp)  | 
|
133  | 
apply(rule Await)  | 
|
134  | 
apply simp_all  | 
|
135  | 
apply(clarify)  | 
|
136  | 
apply(rule Seq)  | 
|
137  | 
prefer 2  | 
|
138  | 
apply(rule Basic)  | 
|
139  | 
apply simp_all  | 
|
140  | 
apply(rule subset_refl)  | 
|
141  | 
apply(rule Basic)  | 
|
142  | 
apply simp_all  | 
|
143  | 
apply clarify  | 
|
144  | 
apply simp  | 
|
145  | 
apply(rule Await)  | 
|
146  | 
apply simp_all  | 
|
147  | 
apply(clarify)  | 
|
148  | 
apply(rule Seq)  | 
|
149  | 
prefer 2  | 
|
150  | 
apply(rule Basic)  | 
|
151  | 
apply simp_all  | 
|
152  | 
apply(rule subset_refl)  | 
|
| 15102 | 153  | 
apply(auto intro!: Basic)  | 
| 13020 | 154  | 
done  | 
155  | 
||
156  | 
subsubsection {* Parameterized *}
 | 
|
157  | 
||
| 15561 | 158  | 
lemma Example2_lemma2_aux: "j<n \<Longrightarrow>  | 
159  | 
(\<Sum>i=0..<n. (b i::nat)) =  | 
|
160  | 
(\<Sum>i=0..<j. b i) + b j + (\<Sum>i=0..<n-(Suc j) . b (Suc j + i))"  | 
|
| 13020 | 161  | 
apply(induct n)  | 
162  | 
apply simp_all  | 
|
163  | 
apply(simp add:less_Suc_eq)  | 
|
164  | 
apply(auto)  | 
|
165  | 
apply(subgoal_tac "n - j = Suc(n- Suc j)")  | 
|
166  | 
apply simp  | 
|
167  | 
apply arith  | 
|
| 15561 | 168  | 
done  | 
| 13020 | 169  | 
|
| 15561 | 170  | 
lemma Example2_lemma2_aux2:  | 
171  | 
"j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)"  | 
|
| 13020 | 172  | 
apply(induct j)  | 
| 
15041
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
173  | 
apply (simp_all cong:setsum_cong)  | 
| 13020 | 174  | 
done  | 
175  | 
||
| 
15041
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
176  | 
lemma Example2_lemma2:  | 
| 15561 | 177  | 
"\<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)"  | 
| 
15041
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
178  | 
apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)  | 
| 15561 | 179  | 
apply(erule_tac  t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
 | 
| 13020 | 180  | 
apply(frule_tac b=b in Example2_lemma2_aux)  | 
| 15561 | 181  | 
apply(erule_tac  t="setsum b {0..<n}" in ssubst)
 | 
182  | 
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)))")
 | 
|
| 
15041
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
183  | 
apply(rotate_tac -1)  | 
| 
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
184  | 
apply(erule ssubst)  | 
| 
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
185  | 
apply(subgoal_tac "j\<le>j")  | 
| 
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
186  | 
apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2)  | 
| 
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
187  | 
apply(rotate_tac -1)  | 
| 
 
a6b1f0cef7b3
Got rid of Summation and made it a translation into setsum instead.
 
nipkow 
parents: 
14174 
diff
changeset
 | 
188  | 
apply(erule ssubst)  | 
| 13020 | 189  | 
apply simp_all  | 
190  | 
done  | 
|
191  | 
||
| 15561 | 192  | 
lemma Example2_lemma2_Suc0: "\<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow>  | 
193  | 
Suc (\<Sum>i::nat=0..< n. b i)=(\<Sum>i=0..< n. (b (j:=Suc 0)) i)"  | 
|
| 13020 | 194  | 
by(simp add:Example2_lemma2)  | 
195  | 
||
196  | 
record Example2_parameterized =  | 
|
197  | 
C :: "nat \<Rightarrow> nat"  | 
|
198  | 
y :: nat  | 
|
199  | 
||
200  | 
lemma Example2_parameterized: "0<n \<Longrightarrow>  | 
|
201  | 
\<turnstile> COBEGIN SCHEME [0\<le>i<n]  | 
|
202  | 
(\<langle> \<acute>y:=\<acute>y+1;; \<acute>C:=\<acute>C (i:=1) \<rangle>,  | 
|
| 15561 | 203  | 
\<lbrace>\<acute>y=(\<Sum>i=0..<n. \<acute>C i) \<and> \<acute>C i=0\<rbrace>,  | 
| 13020 | 204  | 
\<lbrace>\<ordmasculine>C i = \<ordfeminine>C i \<and>  | 
| 15561 | 205  | 
(\<ordmasculine>y=(\<Sum>i=0..<n. \<ordmasculine>C i) \<longrightarrow> \<ordfeminine>y =(\<Sum>i=0..<n. \<ordfeminine>C i))\<rbrace>,  | 
| 13020 | 206  | 
\<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>C j = \<ordfeminine>C j) \<and>  | 
| 15561 | 207  | 
(\<ordmasculine>y=(\<Sum>i=0..<n. \<ordmasculine>C i) \<longrightarrow> \<ordfeminine>y =(\<Sum>i=0..<n. \<ordfeminine>C i))\<rbrace>,  | 
208  | 
\<lbrace>\<acute>y=(\<Sum>i=0..<n. \<acute>C i) \<and> \<acute>C i=1\<rbrace>)  | 
|
| 13020 | 209  | 
COEND  | 
| 15561 | 210  | 
SAT [\<lbrace>\<acute>y=0 \<and> (\<Sum>i=0..<n. \<acute>C i)=0 \<rbrace>, \<lbrace>\<ordmasculine>C=\<ordfeminine>C \<and> \<ordmasculine>y=\<ordfeminine>y\<rbrace>, \<lbrace>True\<rbrace>, \<lbrace>\<acute>y=n\<rbrace>]"  | 
| 13020 | 211  | 
apply(rule Parallel)  | 
212  | 
apply force  | 
|
213  | 
apply force  | 
|
| 15561 | 214  | 
apply(force)  | 
| 13020 | 215  | 
apply clarify  | 
216  | 
apply simp  | 
|
| 15561 | 217  | 
apply(simp cong:setsum_ivl_cong)  | 
| 13020 | 218  | 
apply clarify  | 
219  | 
apply simp  | 
|
220  | 
apply(rule Await)  | 
|
221  | 
apply simp_all  | 
|
222  | 
apply clarify  | 
|
223  | 
apply(rule Seq)  | 
|
224  | 
prefer 2  | 
|
225  | 
apply(rule Basic)  | 
|
226  | 
apply(rule subset_refl)  | 
|
227  | 
apply simp+  | 
|
228  | 
apply(rule Basic)  | 
|
229  | 
apply simp  | 
|
230  | 
apply clarify  | 
|
231  | 
apply simp  | 
|
| 
16733
 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 
nipkow 
parents: 
16417 
diff
changeset
 | 
232  | 
apply(simp add:Example2_lemma2_Suc0 cong:if_cong)  | 
| 13020 | 233  | 
apply simp+  | 
234  | 
done  | 
|
235  | 
||
236  | 
subsection {* Find Least Element *}
 | 
|
237  | 
||
238  | 
text {* A previous lemma: *}
 | 
|
239  | 
||
240  | 
lemma mod_aux :"\<lbrakk>i < (n::nat); a mod n = i; j < a + n; j mod n = i; a < j\<rbrakk> \<Longrightarrow> False"  | 
|
241  | 
apply(subgoal_tac "a=a div n*n + a mod n" )  | 
|
| 13517 | 242  | 
prefer 2 apply (simp (no_asm_use))  | 
| 13020 | 243  | 
apply(subgoal_tac "j=j div n*n + j mod n")  | 
| 13517 | 244  | 
prefer 2 apply (simp (no_asm_use))  | 
| 13020 | 245  | 
apply simp  | 
246  | 
apply(subgoal_tac "a div n*n < j div n*n")  | 
|
247  | 
prefer 2 apply arith  | 
|
248  | 
apply(subgoal_tac "j div n*n < (a div n + 1)*n")  | 
|
| 13517 | 249  | 
prefer 2 apply simp  | 
| 13020 | 250  | 
apply (simp only:mult_less_cancel2)  | 
251  | 
apply arith  | 
|
252  | 
done  | 
|
253  | 
||
254  | 
record Example3 =  | 
|
255  | 
X :: "nat \<Rightarrow> nat"  | 
|
256  | 
Y :: "nat \<Rightarrow> nat"  | 
|
257  | 
||
258  | 
lemma Example3: "m mod n=0 \<Longrightarrow>  | 
|
259  | 
\<turnstile> COBEGIN  | 
|
260  | 
SCHEME [0\<le>i<n]  | 
|
261  | 
(WHILE (\<forall>j<n. \<acute>X i < \<acute>Y j) DO  | 
|
262  | 
IF P(B!(\<acute>X i)) THEN \<acute>Y:=\<acute>Y (i:=\<acute>X i)  | 
|
263  | 
ELSE \<acute>X:= \<acute>X (i:=(\<acute>X i)+ n) FI  | 
|
264  | 
OD,  | 
|
265  | 
\<lbrace>(\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i)\<rbrace>,  | 
|
266  | 
\<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordfeminine>Y j \<le> \<ordmasculine>Y j) \<and> \<ordmasculine>X i = \<ordfeminine>X i \<and>  | 
|
267  | 
\<ordmasculine>Y i = \<ordfeminine>Y i\<rbrace>,  | 
|
268  | 
\<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>X j = \<ordfeminine>X j \<and> \<ordmasculine>Y j = \<ordfeminine>Y j) \<and>  | 
|
269  | 
\<ordfeminine>Y i \<le> \<ordmasculine>Y i\<rbrace>,  | 
|
270  | 
\<lbrace>(\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i) \<rbrace>)  | 
|
271  | 
COEND  | 
|
272  | 
SAT [\<lbrace> \<forall>i<n. \<acute>X i=i \<and> \<acute>Y i=m+i \<rbrace>,\<lbrace>\<ordmasculine>X=\<ordfeminine>X \<and> \<ordmasculine>Y=\<ordfeminine>Y\<rbrace>,\<lbrace>True\<rbrace>,  | 
|
273  | 
\<lbrace>\<forall>i<n. (\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and>  | 
|
274  | 
(\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i)\<rbrace>]"  | 
|
275  | 
apply(rule Parallel)  | 
|
| 13099 | 276  | 
--{*5 subgoals left *}
 | 
| 13020 | 277  | 
apply force+  | 
278  | 
apply clarify  | 
|
279  | 
apply simp  | 
|
280  | 
apply(rule While)  | 
|
281  | 
apply force  | 
|
282  | 
apply force  | 
|
283  | 
apply force  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
13601 
diff
changeset
 | 
284  | 
apply(rule_tac pre'="\<lbrace> \<acute>X i mod n = i \<and> (\<forall>j. j<\<acute>X i \<longrightarrow> j mod n = i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i < n * q \<longrightarrow> P (B!(\<acute>Y i))) \<and> \<acute>X i<\<acute>Y i\<rbrace>" in Conseq)  | 
| 13020 | 285  | 
apply force  | 
286  | 
apply(rule subset_refl)+  | 
|
287  | 
apply(rule Cond)  | 
|
288  | 
apply force  | 
|
289  | 
apply(rule Basic)  | 
|
290  | 
apply force  | 
|
| 13187 | 291  | 
apply fastsimp  | 
| 13020 | 292  | 
apply force  | 
293  | 
apply force  | 
|
294  | 
apply(rule Basic)  | 
|
| 27676 | 295  | 
apply simp  | 
| 13020 | 296  | 
apply clarify  | 
297  | 
apply simp  | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
298  | 
apply (case_tac "X x (j mod n) \<le> j")  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
299  | 
apply (drule le_imp_less_or_eq)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
300  | 
apply (erule disjE)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
301  | 
apply (drule_tac j=j and n=n and i="j mod n" and a="X x (j mod n)" in mod_aux)  | 
| 
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
302  | 
apply auto  | 
| 13020 | 303  | 
done  | 
304  | 
||
305  | 
text {* Same but with a list as auxiliary variable: *}
 | 
|
306  | 
||
307  | 
record Example3_list =  | 
|
308  | 
X :: "nat list"  | 
|
309  | 
Y :: "nat list"  | 
|
310  | 
||
311  | 
lemma Example3_list: "m mod n=0 \<Longrightarrow> \<turnstile> (COBEGIN SCHEME [0\<le>i<n]  | 
|
312  | 
(WHILE (\<forall>j<n. \<acute>X!i < \<acute>Y!j) DO  | 
|
313  | 
IF P(B!(\<acute>X!i)) THEN \<acute>Y:=\<acute>Y[i:=\<acute>X!i] ELSE \<acute>X:= \<acute>X[i:=(\<acute>X!i)+ n] FI  | 
|
314  | 
OD,  | 
|
315  | 
\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> (\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i)\<rbrace>,  | 
|
316  | 
\<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordfeminine>Y!j \<le> \<ordmasculine>Y!j) \<and> \<ordmasculine>X!i = \<ordfeminine>X!i \<and>  | 
|
317  | 
\<ordmasculine>Y!i = \<ordfeminine>Y!i \<and> length \<ordmasculine>X = length \<ordfeminine>X \<and> length \<ordmasculine>Y = length \<ordfeminine>Y\<rbrace>,  | 
|
318  | 
\<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>X!j = \<ordfeminine>X!j \<and> \<ordmasculine>Y!j = \<ordfeminine>Y!j) \<and>  | 
|
319  | 
\<ordfeminine>Y!i \<le> \<ordmasculine>Y!i \<and> length \<ordmasculine>X = length \<ordfeminine>X \<and> length \<ordmasculine>Y = length \<ordfeminine>Y\<rbrace>,  | 
|
320  | 
\<lbrace>(\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y!j \<le> \<acute>X!i) \<rbrace>) COEND)  | 
|
321  | 
SAT [\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> (\<forall>i<n. \<acute>X!i=i \<and> \<acute>Y!i=m+i) \<rbrace>,  | 
|
322  | 
\<lbrace>\<ordmasculine>X=\<ordfeminine>X \<and> \<ordmasculine>Y=\<ordfeminine>Y\<rbrace>,  | 
|
323  | 
\<lbrace>True\<rbrace>,  | 
|
324  | 
\<lbrace>\<forall>i<n. (\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and>  | 
|
325  | 
(\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y!j \<le> \<acute>X!i)\<rbrace>]"  | 
|
326  | 
apply(rule Parallel)  | 
|
| 13099 | 327  | 
--{* 5 subgoals left *}
 | 
| 13020 | 328  | 
apply force+  | 
329  | 
apply clarify  | 
|
330  | 
apply simp  | 
|
331  | 
apply(rule While)  | 
|
332  | 
apply force  | 
|
333  | 
apply force  | 
|
334  | 
apply force  | 
|
| 
14174
 
f3cafd2929d5
Methods rule_tac etc support static (Isar) contexts.
 
ballarin 
parents: 
13601 
diff
changeset
 | 
335  | 
apply(rule_tac pre'="\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> \<acute>X ! i mod n = i \<and> (\<forall>j. j < \<acute>X ! i \<longrightarrow> j mod n = i \<longrightarrow> \<not> P (B ! j)) \<and> (\<acute>Y ! i < n * q \<longrightarrow> P (B ! (\<acute>Y ! i))) \<and> \<acute>X!i<\<acute>Y!i\<rbrace>" in Conseq)  | 
| 13020 | 336  | 
apply force  | 
337  | 
apply(rule subset_refl)+  | 
|
338  | 
apply(rule Cond)  | 
|
339  | 
apply force  | 
|
340  | 
apply(rule Basic)  | 
|
341  | 
apply force  | 
|
342  | 
apply force  | 
|
343  | 
apply force  | 
|
344  | 
apply force  | 
|
345  | 
apply(rule Basic)  | 
|
346  | 
apply simp  | 
|
347  | 
apply clarify  | 
|
| 27676 | 348  | 
apply simp  | 
| 13020 | 349  | 
apply(rule allI)  | 
350  | 
apply(rule impI)+  | 
|
351  | 
apply(case_tac "X x ! i\<le> j")  | 
|
352  | 
apply(drule le_imp_less_or_eq)  | 
|
353  | 
apply(erule disjE)  | 
|
354  | 
apply(drule_tac j=j and n=n and i=i and a="X x ! i" in mod_aux)  | 
|
| 
27651
 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 
haftmann 
parents: 
16733 
diff
changeset
 | 
355  | 
apply auto  | 
| 13020 | 356  | 
done  | 
357  | 
||
| 13187 | 358  | 
end  |