author | haftmann |
Fri, 03 Nov 2006 14:22:44 +0100 | |
changeset 21158 | b379fdc3a3bd |
parent 19763 | ec18656a2c10 |
child 21404 | eb85850d3eb7 |
permissions | -rw-r--r-- |
17293 | 1 |
(* Title: HOLCF/FOCUS/Buffer.thy |
11355 | 2 |
ID: $Id$ |
17293 | 3 |
Author: David von Oheimb, TU Muenchen |
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
4 |
|
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
5 |
Formalization of section 4 of |
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
6 |
|
15038 | 7 |
@inproceedings {broy_mod94, |
8 |
author = {Manfred Broy}, |
|
9 |
title = {{Specification and Refinement of a Buffer of Length One}}, |
|
10 |
booktitle = {Deductive Program Design}, |
|
11 |
year = {1994}, |
|
12 |
editor = {Manfred Broy}, |
|
13 |
volume = {152}, |
|
14 |
series = {ASI Series, Series F: Computer and System Sciences}, |
|
15 |
pages = {273 -- 304}, |
|
16 |
publisher = {Springer} |
|
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
17 |
} |
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
18 |
|
17646 | 19 |
Slides available from http://ddvo.net/talks/1-Buffer.ps.gz |
15038 | 20 |
|
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
21 |
*) |
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
22 |
|
17293 | 23 |
theory Buffer |
24 |
imports FOCUS |
|
25 |
begin |
|
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
26 |
|
17293 | 27 |
typedecl D |
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
28 |
|
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
29 |
datatype |
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
30 |
|
17293 | 31 |
M = Md D | Mreq ("\<bullet>") |
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
32 |
|
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
33 |
datatype |
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
34 |
|
17293 | 35 |
State = Sd D | Snil ("\<currency>") |
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
36 |
|
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
37 |
types |
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
38 |
|
17293 | 39 |
SPF11 = "M fstream \<rightarrow> D fstream" |
40 |
SPEC11 = "SPF11 set" |
|
41 |
SPSF11 = "State \<Rightarrow> SPF11" |
|
42 |
SPECS11 = "SPSF11 set" |
|
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
43 |
|
19763 | 44 |
definition |
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
45 |
|
17293 | 46 |
BufEq_F :: "SPEC11 \<Rightarrow> SPEC11" |
19763 | 47 |
"BufEq_F B = {f. \<forall>d. f\<cdot>(Md d\<leadsto><>) = <> \<and> |
48 |
(\<forall>x. \<exists>ff\<in>B. f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>ff\<cdot>x)}" |
|
49 |
||
17293 | 50 |
BufEq :: "SPEC11" |
19763 | 51 |
"BufEq = gfp BufEq_F" |
52 |
||
17293 | 53 |
BufEq_alt :: "SPEC11" |
19763 | 54 |
"BufEq_alt = gfp (\<lambda>B. {f. \<forall>d. f\<cdot>(Md d\<leadsto><> ) = <> \<and> |
55 |
(\<exists>ff\<in>B. (\<forall>x. f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>ff\<cdot>x))})" |
|
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
56 |
|
17293 | 57 |
BufAC_Asm_F :: " (M fstream set) \<Rightarrow> (M fstream set)" |
19763 | 58 |
"BufAC_Asm_F A = {s. s = <> \<or> |
59 |
(\<exists>d x. s = Md d\<leadsto>x \<and> (x = <> \<or> (ft\<cdot>x = Def \<bullet> \<and> (rt\<cdot>x)\<in>A)))}" |
|
60 |
||
17293 | 61 |
BufAC_Asm :: " (M fstream set)" |
19763 | 62 |
"BufAC_Asm = gfp BufAC_Asm_F" |
63 |
||
17293 | 64 |
BufAC_Cmt_F :: "((M fstream * D fstream) set) \<Rightarrow> |
65 |
((M fstream * D fstream) set)" |
|
19763 | 66 |
"BufAC_Cmt_F C = {(s,t). \<forall>d x. |
17293 | 67 |
(s = <> \<longrightarrow> t = <> ) \<and> |
68 |
(s = Md d\<leadsto><> \<longrightarrow> t = <> ) \<and> |
|
69 |
(s = Md d\<leadsto>\<bullet>\<leadsto>x \<longrightarrow> (ft\<cdot>t = Def d \<and> (x,rt\<cdot>t)\<in>C))}" |
|
19763 | 70 |
|
71 |
BufAC_Cmt :: "((M fstream * D fstream) set)" |
|
72 |
"BufAC_Cmt = gfp BufAC_Cmt_F" |
|
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
73 |
|
19763 | 74 |
BufAC :: "SPEC11" |
75 |
"BufAC = {f. \<forall>x. x\<in>BufAC_Asm \<longrightarrow> (x,f\<cdot>x)\<in>BufAC_Cmt}" |
|
76 |
||
77 |
BufSt_F :: "SPECS11 \<Rightarrow> SPECS11" |
|
78 |
"BufSt_F H = {h. \<forall>s . h s \<cdot><> = <> \<and> |
|
17293 | 79 |
(\<forall>d x. h \<currency> \<cdot>(Md d\<leadsto>x) = h (Sd d)\<cdot>x \<and> |
80 |
(\<exists>hh\<in>H. h (Sd d)\<cdot>(\<bullet> \<leadsto>x) = d\<leadsto>(hh \<currency>\<cdot>x)))}" |
|
19763 | 81 |
|
82 |
BufSt_P :: "SPECS11" |
|
83 |
"BufSt_P = gfp BufSt_F" |
|
17293 | 84 |
|
19763 | 85 |
BufSt :: "SPEC11" |
86 |
"BufSt = {f. \<exists>h\<in>BufSt_P. f = h \<currency>}" |
|
87 |
||
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
88 |
|
19759 | 89 |
lemma set_cong: "!!X. A = B ==> (x:A) = (x:B)" |
90 |
by (erule subst, rule refl) |
|
91 |
||
92 |
ML {* |
|
93 |
fun B_prover s tac eqs = prove_goal (the_context ()) s (fn prems => [cut_facts_tac prems 1, |
|
94 |
tac 1, auto_tac (claset(), simpset() addsimps eqs)]); |
|
95 |
||
96 |
fun prove_forw s thm = B_prover s (dtac (thm RS iffD1)) []; |
|
97 |
fun prove_backw s thm eqs = B_prover s (rtac (thm RS iffD2)) eqs; |
|
98 |
*} |
|
99 |
||
100 |
(**** BufEq *******************************************************************) |
|
101 |
||
102 |
lemma mono_BufEq_F: "mono BufEq_F" |
|
103 |
by (unfold mono_def BufEq_F_def, fast) |
|
104 |
||
19763 | 105 |
lemmas BufEq_fix = mono_BufEq_F [THEN BufEq_def [THEN eq_reflection, THEN def_gfp_unfold]] |
19759 | 106 |
|
107 |
lemma BufEq_unfold: "(f:BufEq) = (!d. f\<cdot>(Md d\<leadsto><>) = <> & |
|
108 |
(!x. ? ff:BufEq. f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>(ff\<cdot>x)))" |
|
109 |
apply (subst BufEq_fix [THEN set_cong]) |
|
110 |
apply (unfold BufEq_F_def) |
|
111 |
apply (simp) |
|
112 |
done |
|
113 |
||
114 |
lemma Buf_f_empty: "f:BufEq \<Longrightarrow> f\<cdot><> = <>" |
|
115 |
by (drule BufEq_unfold [THEN iffD1], auto) |
|
116 |
||
117 |
lemma Buf_f_d: "f:BufEq \<Longrightarrow> f\<cdot>(Md d\<leadsto><>) = <>" |
|
118 |
by (drule BufEq_unfold [THEN iffD1], auto) |
|
119 |
||
120 |
lemma Buf_f_d_req: |
|
121 |
"f:BufEq \<Longrightarrow> \<exists>ff. ff:BufEq \<and> f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>ff\<cdot>x" |
|
122 |
by (drule BufEq_unfold [THEN iffD1], auto) |
|
123 |
||
124 |
||
125 |
(**** BufAC_Asm ***************************************************************) |
|
126 |
||
127 |
lemma mono_BufAC_Asm_F: "mono BufAC_Asm_F" |
|
128 |
by (unfold mono_def BufAC_Asm_F_def, fast) |
|
129 |
||
19763 | 130 |
lemmas BufAC_Asm_fix = |
131 |
mono_BufAC_Asm_F [THEN BufAC_Asm_def [THEN eq_reflection, THEN def_gfp_unfold]] |
|
19759 | 132 |
|
133 |
lemma BufAC_Asm_unfold: "(s:BufAC_Asm) = (s = <> | (? d x. |
|
134 |
s = Md d\<leadsto>x & (x = <> | (ft\<cdot>x = Def \<bullet> & (rt\<cdot>x):BufAC_Asm))))" |
|
135 |
apply (subst BufAC_Asm_fix [THEN set_cong]) |
|
136 |
apply (unfold BufAC_Asm_F_def) |
|
137 |
apply (simp) |
|
138 |
done |
|
139 |
||
140 |
lemma BufAC_Asm_empty: "<> :BufAC_Asm" |
|
141 |
by (rule BufAC_Asm_unfold [THEN iffD2], auto) |
|
142 |
||
143 |
lemma BufAC_Asm_d: "Md d\<leadsto><>:BufAC_Asm" |
|
144 |
by (rule BufAC_Asm_unfold [THEN iffD2], auto) |
|
145 |
lemma BufAC_Asm_d_req: "x:BufAC_Asm ==> Md d\<leadsto>\<bullet>\<leadsto>x:BufAC_Asm" |
|
146 |
by (rule BufAC_Asm_unfold [THEN iffD2], auto) |
|
147 |
lemma BufAC_Asm_prefix2: "a\<leadsto>b\<leadsto>s:BufAC_Asm ==> s:BufAC_Asm" |
|
148 |
by (drule BufAC_Asm_unfold [THEN iffD1], auto) |
|
149 |
||
150 |
||
151 |
(**** BBufAC_Cmt **************************************************************) |
|
152 |
||
153 |
lemma mono_BufAC_Cmt_F: "mono BufAC_Cmt_F" |
|
154 |
by (unfold mono_def BufAC_Cmt_F_def, fast) |
|
155 |
||
19763 | 156 |
lemmas BufAC_Cmt_fix = |
157 |
mono_BufAC_Cmt_F [THEN BufAC_Cmt_def [THEN eq_reflection, THEN def_gfp_unfold]] |
|
19759 | 158 |
|
159 |
lemma BufAC_Cmt_unfold: "((s,t):BufAC_Cmt) = (!d x. |
|
160 |
(s = <> --> t = <>) & |
|
161 |
(s = Md d\<leadsto><> --> t = <>) & |
|
162 |
(s = Md d\<leadsto>\<bullet>\<leadsto>x --> ft\<cdot>t = Def d & (x, rt\<cdot>t):BufAC_Cmt))" |
|
163 |
apply (subst BufAC_Cmt_fix [THEN set_cong]) |
|
164 |
apply (unfold BufAC_Cmt_F_def) |
|
165 |
apply (simp) |
|
166 |
done |
|
167 |
||
168 |
lemma BufAC_Cmt_empty: "f:BufEq ==> (<>, f\<cdot><>):BufAC_Cmt" |
|
169 |
by (rule BufAC_Cmt_unfold [THEN iffD2], auto simp add: Buf_f_empty) |
|
170 |
||
171 |
lemma BufAC_Cmt_d: "f:BufEq ==> (a\<leadsto>\<bottom>, f\<cdot>(a\<leadsto>\<bottom>)):BufAC_Cmt" |
|
172 |
by (rule BufAC_Cmt_unfold [THEN iffD2], auto simp add: Buf_f_d) |
|
173 |
||
174 |
lemma BufAC_Cmt_d2: |
|
175 |
"(Md d\<leadsto>\<bottom>, f\<cdot>(Md d\<leadsto>\<bottom>)):BufAC_Cmt ==> f\<cdot>(Md d\<leadsto>\<bottom>) = \<bottom>" |
|
176 |
by (drule BufAC_Cmt_unfold [THEN iffD1], auto) |
|
177 |
||
178 |
lemma BufAC_Cmt_d3: |
|
179 |
"(Md d\<leadsto>\<bullet>\<leadsto>x, f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x)):BufAC_Cmt ==> (x, rt\<cdot>(f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x))):BufAC_Cmt" |
|
180 |
by (drule BufAC_Cmt_unfold [THEN iffD1], auto) |
|
181 |
||
182 |
lemma BufAC_Cmt_d32: |
|
183 |
"(Md d\<leadsto>\<bullet>\<leadsto>x, f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x)):BufAC_Cmt ==> ft\<cdot>(f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x)) = Def d" |
|
184 |
by (drule BufAC_Cmt_unfold [THEN iffD1], auto) |
|
185 |
||
186 |
(**** BufAC *******************************************************************) |
|
187 |
||
188 |
lemma BufAC_f_d: "f \<in> BufAC \<Longrightarrow> f\<cdot>(Md d\<leadsto>\<bottom>) = \<bottom>" |
|
189 |
apply (unfold BufAC_def) |
|
190 |
apply (fast intro: BufAC_Cmt_d2 BufAC_Asm_d) |
|
191 |
done |
|
192 |
||
193 |
lemma ex_elim_lemma: "(? ff:B. (!x. f\<cdot>(a\<leadsto>b\<leadsto>x) = d\<leadsto>ff\<cdot>x)) = |
|
194 |
((!x. ft\<cdot>(f\<cdot>(a\<leadsto>b\<leadsto>x)) = Def d) & (LAM x. rt\<cdot>(f\<cdot>(a\<leadsto>b\<leadsto>x))):B)" |
|
195 |
(* this is an instance (though unification cannot handle this) of |
|
196 |
lemma "(? ff:B. (!x. f\<cdot>x = d\<leadsto>ff\<cdot>x)) = \ |
|
197 |
\((!x. ft\<cdot>(f\<cdot>x) = Def d) & (LAM x. rt\<cdot>(f\<cdot>x)):B)"*) |
|
198 |
apply safe |
|
199 |
apply ( rule_tac [2] P="(%x. x:B)" in ssubst) |
|
200 |
prefer 3 |
|
201 |
apply ( assumption) |
|
202 |
apply ( rule_tac [2] ext_cfun) |
|
203 |
apply ( drule_tac [2] spec) |
|
204 |
apply ( drule_tac [2] f="rt" in cfun_arg_cong) |
|
205 |
prefer 2 |
|
206 |
apply ( simp) |
|
207 |
prefer 2 |
|
208 |
apply ( simp) |
|
209 |
apply (rule_tac bexI) |
|
210 |
apply auto |
|
211 |
apply (drule spec) |
|
212 |
apply (erule exE) |
|
213 |
apply (erule ssubst) |
|
214 |
apply (simp) |
|
215 |
done |
|
216 |
||
217 |
lemma BufAC_f_d_req: "f\<in>BufAC \<Longrightarrow> \<exists>ff\<in>BufAC. \<forall>x. f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>ff\<cdot>x" |
|
218 |
apply (unfold BufAC_def) |
|
219 |
apply (rule ex_elim_lemma [THEN iffD2]) |
|
220 |
apply safe |
|
221 |
apply (fast intro: BufAC_Cmt_d32 [THEN Def_maximal] |
|
222 |
monofun_cfun_arg BufAC_Asm_empty [THEN BufAC_Asm_d_req]) |
|
223 |
apply (auto intro: BufAC_Cmt_d3 BufAC_Asm_d_req) |
|
224 |
done |
|
225 |
||
226 |
||
227 |
(**** BufSt *******************************************************************) |
|
228 |
||
229 |
lemma mono_BufSt_F: "mono BufSt_F" |
|
230 |
by (unfold mono_def BufSt_F_def, fast) |
|
231 |
||
19763 | 232 |
lemmas BufSt_P_fix = |
233 |
mono_BufSt_F [THEN BufSt_P_def [THEN eq_reflection, THEN def_gfp_unfold]] |
|
19759 | 234 |
|
235 |
lemma BufSt_P_unfold: "(h:BufSt_P) = (!s. h s\<cdot><> = <> & |
|
236 |
(!d x. h \<currency> \<cdot>(Md d\<leadsto>x) = h (Sd d)\<cdot>x & |
|
237 |
(? hh:BufSt_P. h (Sd d)\<cdot>(\<bullet>\<leadsto>x) = d\<leadsto>(hh \<currency> \<cdot>x))))" |
|
238 |
apply (subst BufSt_P_fix [THEN set_cong]) |
|
239 |
apply (unfold BufSt_F_def) |
|
240 |
apply (simp) |
|
241 |
done |
|
242 |
||
243 |
lemma BufSt_P_empty: "h:BufSt_P ==> h s \<cdot> <> = <>" |
|
244 |
by (drule BufSt_P_unfold [THEN iffD1], auto) |
|
245 |
lemma BufSt_P_d: "h:BufSt_P ==> h \<currency> \<cdot>(Md d\<leadsto>x) = h (Sd d)\<cdot>x" |
|
246 |
by (drule BufSt_P_unfold [THEN iffD1], auto) |
|
247 |
lemma BufSt_P_d_req: "h:BufSt_P ==> \<exists>hh\<in>BufSt_P. |
|
248 |
h (Sd d)\<cdot>(\<bullet> \<leadsto>x) = d\<leadsto>(hh \<currency> \<cdot>x)" |
|
249 |
by (drule BufSt_P_unfold [THEN iffD1], auto) |
|
250 |
||
251 |
||
252 |
(**** Buf_AC_imp_Eq ***********************************************************) |
|
253 |
||
254 |
lemma Buf_AC_imp_Eq: "BufAC \<subseteq> BufEq" |
|
255 |
apply (unfold BufEq_def) |
|
256 |
apply (rule gfp_upperbound) |
|
257 |
apply (unfold BufEq_F_def) |
|
258 |
apply safe |
|
259 |
apply (erule BufAC_f_d) |
|
260 |
apply (drule BufAC_f_d_req) |
|
261 |
apply (fast) |
|
262 |
done |
|
263 |
||
264 |
||
265 |
(**** Buf_Eq_imp_AC by coinduction ********************************************) |
|
266 |
||
267 |
lemma BufAC_Asm_cong_lemma [rule_format]: "\<forall>s f ff. f\<in>BufEq \<longrightarrow> ff\<in>BufEq \<longrightarrow> |
|
268 |
s\<in>BufAC_Asm \<longrightarrow> stream_take n\<cdot>(f\<cdot>s) = stream_take n\<cdot>(ff\<cdot>s)" |
|
269 |
apply (induct_tac "n") |
|
270 |
apply (simp) |
|
271 |
apply (intro strip) |
|
272 |
apply (drule BufAC_Asm_unfold [THEN iffD1]) |
|
273 |
apply safe |
|
274 |
apply (simp add: Buf_f_empty) |
|
275 |
apply (simp add: Buf_f_d) |
|
276 |
apply (drule ft_eq [THEN iffD1]) |
|
277 |
apply (clarsimp) |
|
278 |
apply (drule Buf_f_d_req)+ |
|
279 |
apply safe |
|
280 |
apply (erule ssubst)+ |
|
281 |
apply (simp (no_asm)) |
|
282 |
apply (fast) |
|
283 |
done |
|
284 |
||
285 |
lemma BufAC_Asm_cong: "\<lbrakk>f \<in> BufEq; ff \<in> BufEq; s \<in> BufAC_Asm\<rbrakk> \<Longrightarrow> f\<cdot>s = ff\<cdot>s" |
|
286 |
apply (rule stream.take_lemmas) |
|
287 |
apply (erule (2) BufAC_Asm_cong_lemma) |
|
288 |
done |
|
289 |
||
290 |
lemma Buf_Eq_imp_AC_lemma: "\<lbrakk>f \<in> BufEq; x \<in> BufAC_Asm\<rbrakk> \<Longrightarrow> (x, f\<cdot>x) \<in> BufAC_Cmt" |
|
291 |
apply (unfold BufAC_Cmt_def) |
|
292 |
apply (rotate_tac) |
|
293 |
apply (erule weak_coinduct_image) |
|
294 |
apply (unfold BufAC_Cmt_F_def) |
|
295 |
apply safe |
|
296 |
apply (erule Buf_f_empty) |
|
297 |
apply (erule Buf_f_d) |
|
298 |
apply (drule Buf_f_d_req) |
|
299 |
apply (clarsimp) |
|
300 |
apply (erule exI) |
|
301 |
apply (drule BufAC_Asm_prefix2) |
|
302 |
apply (frule Buf_f_d_req) |
|
303 |
apply (clarsimp) |
|
304 |
apply (erule ssubst) |
|
305 |
apply (simp) |
|
306 |
apply (drule (2) BufAC_Asm_cong) |
|
307 |
apply (erule subst) |
|
308 |
apply (erule imageI) |
|
309 |
done |
|
310 |
lemma Buf_Eq_imp_AC: "BufEq \<subseteq> BufAC" |
|
311 |
apply (unfold BufAC_def) |
|
312 |
apply (clarify) |
|
313 |
apply (erule (1) Buf_Eq_imp_AC_lemma) |
|
314 |
done |
|
315 |
||
316 |
(**** Buf_Eq_eq_AC ************************************************************) |
|
317 |
||
318 |
lemmas Buf_Eq_eq_AC = Buf_AC_imp_Eq [THEN Buf_Eq_imp_AC [THEN subset_antisym]] |
|
319 |
||
320 |
||
321 |
(**** alternative (not strictly) stronger version of Buf_Eq *******************) |
|
322 |
||
323 |
lemma Buf_Eq_alt_imp_Eq: "BufEq_alt \<subseteq> BufEq" |
|
324 |
apply (unfold BufEq_def BufEq_alt_def) |
|
325 |
apply (rule gfp_mono) |
|
326 |
apply (unfold BufEq_F_def) |
|
327 |
apply (fast) |
|
328 |
done |
|
329 |
||
330 |
(* direct proof of "BufEq \<subseteq> BufEq_alt" seems impossible *) |
|
331 |
||
332 |
||
333 |
lemma Buf_AC_imp_Eq_alt: "BufAC <= BufEq_alt" |
|
334 |
apply (unfold BufEq_alt_def) |
|
335 |
apply (rule gfp_upperbound) |
|
336 |
apply (fast elim: BufAC_f_d BufAC_f_d_req) |
|
337 |
done |
|
338 |
||
339 |
lemmas Buf_Eq_imp_Eq_alt = subset_trans [OF Buf_Eq_imp_AC Buf_AC_imp_Eq_alt] |
|
340 |
||
341 |
lemmas Buf_Eq_alt_eq = subset_antisym [OF Buf_Eq_alt_imp_Eq Buf_Eq_imp_Eq_alt] |
|
342 |
||
343 |
||
344 |
(**** Buf_Eq_eq_St ************************************************************) |
|
345 |
||
346 |
lemma Buf_St_imp_Eq: "BufSt <= BufEq" |
|
347 |
apply (unfold BufSt_def BufEq_def) |
|
348 |
apply (rule gfp_upperbound) |
|
349 |
apply (unfold BufEq_F_def) |
|
350 |
apply safe |
|
351 |
apply ( simp add: BufSt_P_d BufSt_P_empty) |
|
352 |
apply (simp add: BufSt_P_d) |
|
353 |
apply (drule BufSt_P_d_req) |
|
354 |
apply (force) |
|
355 |
done |
|
356 |
||
357 |
lemma Buf_Eq_imp_St: "BufEq <= BufSt" |
|
358 |
apply (unfold BufSt_def BufSt_P_def) |
|
359 |
apply safe |
|
360 |
apply (rename_tac f) |
|
361 |
apply (rule_tac x="\<lambda>s. case s of Sd d => \<Lambda> x. f\<cdot>(Md d\<leadsto>x)| \<currency> => f" in bexI) |
|
362 |
apply ( simp) |
|
363 |
apply (erule weak_coinduct_image) |
|
364 |
apply (unfold BufSt_F_def) |
|
365 |
apply (simp) |
|
366 |
apply safe |
|
367 |
apply ( rename_tac "s") |
|
368 |
apply ( induct_tac "s") |
|
369 |
apply ( simp add: Buf_f_d) |
|
370 |
apply ( simp add: Buf_f_empty) |
|
371 |
apply ( simp) |
|
372 |
apply (simp) |
|
373 |
apply (rename_tac f d x) |
|
374 |
apply (drule_tac d="d" and x="x" in Buf_f_d_req) |
|
375 |
apply auto |
|
376 |
done |
|
377 |
||
378 |
lemmas Buf_Eq_eq_St = Buf_St_imp_Eq [THEN Buf_Eq_imp_St [THEN subset_antisym]] |
|
379 |
||
11350
4c55b020d6ee
added FOCUS including the One-Element Buffer by Manfred Broy
oheimb
parents:
diff
changeset
|
380 |
end |