| author | krauss | 
| Sun, 12 Dec 2010 21:40:59 +0100 | |
| changeset 41113 | b223fa19af3c | 
| parent 40945 | b8703f63bfb2 | 
| child 41476 | 0fa9629aa399 | 
| permissions | -rw-r--r-- | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 1 | (* Title: HOLCF/IOA/ABP/Impl.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 implementation *}
 | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 6 | |
| 17244 | 7 | theory Impl_finite | 
| 8 | imports Sender Receiver Abschannel_finite | |
| 9 | begin | |
| 10 | ||
| 11 | types | |
| 12 | 'm impl_fin_state | |
| 13 | = "'m sender_state * 'm receiver_state * 'm packet list * bool list" | |
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 14 | (* sender_state * receiver_state * srch_state * rsch_state *) | 
| 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 15 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 16 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 17 |   impl_fin_ioa :: "('m action, 'm impl_fin_state)ioa" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 18 | "impl_fin_ioa = (sender_ioa || receiver_ioa || srch_fin_ioa || | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 19 | rsch_fin_ioa)" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 20 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 21 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 22 | sen_fin :: "'m impl_fin_state => 'm sender_state" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 23 | "sen_fin = fst" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 24 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 25 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 26 | rec_fin :: "'m impl_fin_state => 'm receiver_state" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 27 | "rec_fin = fst o snd" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 28 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 29 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 30 | srch_fin :: "'m impl_fin_state => 'm packet list" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 31 | "srch_fin = fst o snd o snd" | 
| 3072 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 mueller parents: diff
changeset | 32 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 33 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 34 | rsch_fin :: "'m impl_fin_state => bool list" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19738diff
changeset | 35 | "rsch_fin = snd o snd o snd" | 
| 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 |