| author | blanchet | 
| Thu, 06 Jun 2013 21:18:39 +0200 | |
| changeset 52324 | 095c88b93e8d | 
| parent 42151 | 4da4fc77664b | 
| child 58880 | 0baae4311a9f | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/IOA/ABP/Receiver.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: receiver *}
 | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
6  | 
|
| 17244 | 7  | 
theory Receiver  | 
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 receiver_state = "'m list * bool"  -- {* messages, mode *}
 | 
| 
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: 
19738 
diff
changeset
 | 
14  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
15  | 
rq :: "'m receiver_state => 'm list" where  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
16  | 
"rq = 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: 
19738 
diff
changeset
 | 
18  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
19  | 
rbit :: "'m receiver_state => bool" where  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
20  | 
"rbit = 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: 
19738 
diff
changeset
 | 
22  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
23  | 
receiver_asig :: "'m action signature" where  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
24  | 
"receiver_asig =  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
25  | 
    (UN pkt. {R_pkt(pkt)},
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
26  | 
    (UN m. {R_msg(m)}) Un (UN b. {S_ack(b)}),
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
27  | 
    {})"
 | 
| 
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: 
19738 
diff
changeset
 | 
29  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
30  | 
  receiver_trans :: "('m action, 'm receiver_state)transition set" where
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
31  | 
"receiver_trans =  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
32  | 
   {tr. let s = fst(tr);
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
33  | 
t = snd(snd(tr))  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
34  | 
in  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
35  | 
case fst(snd(tr))  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
36  | 
of  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
37  | 
Next => False |  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
38  | 
S_msg(m) => False |  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
39  | 
R_msg(m) => (rq(s) ~= []) &  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
40  | 
m = hd(rq(s)) &  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
41  | 
rq(t) = tl(rq(s)) &  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
42  | 
rbit(t)=rbit(s) |  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
43  | 
S_pkt(pkt) => False |  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
44  | 
R_pkt(pkt) => if (hdr(pkt) ~= rbit(s))&rq(s)=[] then  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
45  | 
rq(t) = (rq(s)@[msg(pkt)]) &rbit(t) = (~rbit(s)) else  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
46  | 
rq(t) =rq(s) & rbit(t)=rbit(s) |  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
47  | 
S_ack(b) => b = rbit(s) &  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
48  | 
rq(t) = rq(s) &  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
49  | 
rbit(t)=rbit(s) |  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
50  | 
R_ack(b) => False}"  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
51  | 
|
| 
25131
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
52  | 
definition  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
53  | 
  receiver_ioa :: "('m action, 'm receiver_state)ioa" where
 | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
54  | 
"receiver_ioa =  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
55  | 
   (receiver_asig, {([],False)}, receiver_trans,{},{})"
 | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
57  | 
end  |