| author | blanchet | 
| Thu, 19 Dec 2013 17:24:17 +0100 | |
| changeset 54823 | a510fc40559c | 
| parent 42151 | 4da4fc77664b | 
| child 58880 | 0baae4311a9f | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/ABP/Spec.thy | 
| 40945 | 2 | Author: Olaf Müller | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 3 | *) | 
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 4 | |
| 17244 | 5 | header {* The specification of reliable transmission *}
 | 
| 6 | ||
| 7 | theory Spec | |
| 8 | imports IOA Action | |
| 9 | begin | |
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 10 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 11 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 12 | spec_sig :: "'m action signature" where | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 13 |   sig_def: "spec_sig = (UN m.{S_msg(m)},
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 14 |                        UN m.{R_msg(m)} Un {Next},
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 15 |                        {})"
 | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 16 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 17 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 18 |   spec_trans :: "('m action, 'm list)transition set" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 19 | trans_def: "spec_trans = | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 20 |    {tr. let s = fst(tr);
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 21 | t = snd(snd(tr)) | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 22 | in | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 23 | case fst(snd(tr)) | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 24 | of | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 25 | Next => t=s | (* Note that there is condition as in Sender *) | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 26 | S_msg(m) => t = s@[m] | | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 27 | R_msg(m) => s = (m#t) | | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 28 | S_pkt(pkt) => False | | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 29 | R_pkt(pkt) => False | | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 30 | S_ack(b) => False | | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 31 | R_ack(b) => False}" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 32 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 33 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 34 |   spec_ioa :: "('m action, 'm list)ioa" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
17244diff
changeset | 35 |   ioa_def: "spec_ioa = (spec_sig, {[]}, spec_trans)"
 | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 36 | |
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 37 | end |