| author | wenzelm | 
| Tue, 26 Dec 2023 12:46:10 +0100 | |
| changeset 79359 | 5b01b93de062 | 
| parent 62002 | f1599e98c4d0 | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/IOA/ABP/Packet.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  | 
|
| 62002 | 5  | 
section \<open>Packets\<close>  | 
| 17244 | 6  | 
|
7  | 
theory Packet  | 
|
8  | 
imports Main  | 
|
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  | 
'msg packet = "bool * 'msg"  | 
| 
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  | 
hdr :: "'msg packet => bool" where  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
16  | 
"hdr = 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  | 
msg :: "'msg packet => 'msg" where  | 
| 
 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 
wenzelm 
parents: 
19738 
diff
changeset
 | 
20  | 
"msg = snd"  | 
| 
3072
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
a31419014be5
Old ABP files now running under the IOA meta theory based on HOLCF;
 
mueller 
parents:  
diff
changeset
 | 
22  | 
end  |