| author | wenzelm | 
| Mon, 06 Jul 2015 10:54:15 +0200 | |
| changeset 60654 | ca1e07005b8b | 
| parent 58880 | 0baae4311a9f | 
| child 61032 | b57df8eecad6 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/IOA/ABP/Sender.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 | |
| 58880 | 5 | section {* The implementation: sender *}
 | 
| 17244 | 6 | |
| 7 | theory Sender | |
| 8 | imports IOA Action Lemmas | |
| 9 | begin | |
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 10 | |
| 41476 | 11 | type_synonym | 
| 17244 | 12 |   'm sender_state = "'m list  *  bool"  -- {* messages, Alternating Bit *}
 | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 13 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 14 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 15 | sq :: "'m sender_state => 'm list" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 16 | "sq = fst" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 17 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 18 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 19 | sbit :: "'m sender_state => bool" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 20 | "sbit = snd" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 21 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 22 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 23 | sender_asig :: "'m action signature" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 24 |   "sender_asig = ((UN m. {S_msg(m)}) Un (UN b. {R_ack(b)}),
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 25 |                    UN pkt. {S_pkt(pkt)},
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 26 |                    {})"
 | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 27 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 28 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 29 |   sender_trans :: "('m action, 'm sender_state)transition set" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 30 | "sender_trans = | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 31 |    {tr. let s = fst(tr);
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 32 | t = snd(snd(tr)) | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 33 | in case fst(snd(tr)) | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 34 | of | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 35 | Next => if sq(s)=[] then t=s else False | | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 36 | S_msg(m) => sq(t)=sq(s)@[m] & | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 37 | sbit(t)=sbit(s) | | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 38 | R_msg(m) => False | | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 39 | S_pkt(pkt) => sq(s) ~= [] & | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 40 | hdr(pkt) = sbit(s) & | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 41 | msg(pkt) = hd(sq(s)) & | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 42 | sq(t) = sq(s) & | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 43 | sbit(t) = sbit(s) | | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 44 | R_pkt(pkt) => False | | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 45 | S_ack(b) => False | | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 46 | R_ack(b) => if b = sbit(s) then | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 47 | sq(t)=tl(sq(s)) & sbit(t)=(~sbit(s)) else | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 48 | sq(t)=sq(s) & sbit(t)=sbit(s)}" | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 49 | |
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 50 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 51 |   sender_ioa :: "('m action, 'm sender_state)ioa" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 52 | "sender_ioa = | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 53 |    (sender_asig, {([],True)}, sender_trans,{},{})"
 | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 54 | |
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 55 | end |