Up to index of Isabelle/HOLCF/FOCUS
theory Buffer(* Title: HOLCF/FOCUS/Buffer.thy
ID: $Id: Buffer.thy,v 1.7 2005/09/26 06:41:24 haftmann Exp $
Author: David von Oheimb, TU Muenchen
Formalization of section 4 of
@inproceedings {broy_mod94,
author = {Manfred Broy},
title = {{Specification and Refinement of a Buffer of Length One}},
booktitle = {Deductive Program Design},
year = {1994},
editor = {Manfred Broy},
volume = {152},
series = {ASI Series, Series F: Computer and System Sciences},
pages = {273 -- 304},
publisher = {Springer}
}
Slides available from http://ddvo.net/talks/1-Buffer.ps.gz
*)
theory Buffer
imports FOCUS
begin
typedecl D
datatype
M = Md D | Mreq ("•")
datatype
State = Sd D | Snil ("¤")
types
SPF11 = "M fstream -> D fstream"
SPEC11 = "SPF11 set"
SPSF11 = "State => SPF11"
SPECS11 = "SPSF11 set"
consts
BufEq_F :: "SPEC11 => SPEC11"
BufEq :: "SPEC11"
BufEq_alt :: "SPEC11"
BufAC_Asm_F :: " (M fstream set) => (M fstream set)"
BufAC_Asm :: " (M fstream set)"
BufAC_Cmt_F :: "((M fstream * D fstream) set) =>
((M fstream * D fstream) set)"
BufAC_Cmt :: "((M fstream * D fstream) set)"
BufAC :: "SPEC11"
BufSt_F :: "SPECS11 => SPECS11"
BufSt_P :: "SPECS11"
BufSt :: "SPEC11"
defs
BufEq_F_def: "BufEq_F B ≡ {f. ∀d. f·(Md d\<leadsto><>) = <> ∧
(∀x. ∃ff∈B. f·(Md d\<leadsto>•\<leadsto>x) = d\<leadsto>ff·x)}"
BufEq_def: "BufEq ≡ gfp BufEq_F"
BufEq_alt_def: "BufEq_alt ≡ gfp (λB. {f. ∀d. f·(Md d\<leadsto><> ) = <> ∧
(∃ff∈B. (∀x. f·(Md d\<leadsto>•\<leadsto>x) = d\<leadsto>ff·x))})"
BufAC_Asm_F_def: "BufAC_Asm_F A ≡ {s. s = <> ∨
(∃d x. s = Md d\<leadsto>x ∧ (x = <> ∨ (ft·x = Def • ∧ (rt·x)∈A)))}"
BufAC_Asm_def: "BufAC_Asm ≡ gfp BufAC_Asm_F"
BufAC_Cmt_F_def: "BufAC_Cmt_F C ≡ {(s,t). ∀d x.
(s = <> --> t = <> ) ∧
(s = Md d\<leadsto><> --> t = <> ) ∧
(s = Md d\<leadsto>•\<leadsto>x --> (ft·t = Def d ∧ (x,rt·t)∈C))}"
BufAC_Cmt_def: "BufAC_Cmt ≡ gfp BufAC_Cmt_F"
BufAC_def: "BufAC ≡ {f. ∀x. x∈BufAC_Asm --> (x,f·x)∈BufAC_Cmt}"
BufSt_F_def: "BufSt_F H ≡ {h. ∀s . h s ·<> = <> ∧
(∀d x. h ¤ ·(Md d\<leadsto>x) = h (Sd d)·x ∧
(∃hh∈H. h (Sd d)·(• \<leadsto>x) = d\<leadsto>(hh ¤·x)))}"
BufSt_P_def: "BufSt_P ≡ gfp BufSt_F"
BufSt_def: "BufSt ≡ {f. ∃h∈BufSt_P. f = h ¤}"
ML {* use_legacy_bindings (the_context ()) *}
end
theorem BufAC_f_d:
f ∈ BufAC ==> f·(Md d~>UU) = UU
theorem ex_elim_lemma:
(∃ff∈B. ∀x. f·(a~>b~>x) = d~>ff·x) = ((∀x. ft·(f·(a~>b~>x)) = Def d) ∧ (LAM x. rt·(f·(a~>b~>x))) ∈ B)
theorem BufAC_f_d_req:
f ∈ BufAC ==> ∃ff∈BufAC. ∀x. f·(Md d~>•~>x) = d~>ff·x
theorem Buf_AC_imp_Eq:
BufAC ⊆ BufEq
theorem BufAC_Asm_cong_lemma:
[| f ∈ BufEq; ff ∈ BufEq; s ∈ BufAC_Asm |] ==> stream_take n·(f·s) = stream_take n·(ff·s)
theorem BufAC_Asm_cong:
[| f ∈ BufEq; ff ∈ BufEq; s ∈ BufAC_Asm |] ==> f·s = ff·s
theorem Buf_Eq_imp_AC_lemma:
[| f ∈ BufEq; x ∈ BufAC_Asm |] ==> (x, f·x) ∈ BufAC_Cmt
theorem Buf_Eq_imp_AC:
BufEq ⊆ BufAC
theorem Buf_AC_imp_Eq_alt:
BufAC ⊆ BufEq_alt
theorem Buf_Eq_imp_Eq_alt:
BufEq ⊆ BufEq_alt
theorem Buf_Eq_alt_eq:
BufEq_alt = BufEq
theorem Buf_St_imp_Eq:
BufSt ⊆ BufEq
theorem Buf_Eq_imp_St:
BufEq ⊆ BufSt
theorem Buf_Eq_eq_St:
BufEq = BufSt