author | wenzelm |
Fri, 15 Aug 2008 15:50:52 +0200 | |
changeset 27885 | 76b51cd0a37c |
parent 27361 | 24ec32bee347 |
child 35174 | e15040ae75d7 |
permissions | -rw-r--r-- |
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
1 |
(* Title: HOL/IOA/NTP/Abschannel.thy |
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
2 |
ID: $Id$ |
12218 | 3 |
Author: Olaf Müller |
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
4 |
*) |
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
5 |
|
17244 | 6 |
header {* The (faulty) transmission channel (both directions) *} |
7 |
||
8 |
theory Abschannel |
|
9 |
imports IOA Action |
|
10 |
begin |
|
11 |
||
12 |
datatype 'a abs_action = S 'a | R 'a |
|
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
13 |
|
27361 | 14 |
definition |
15 |
ch_asig :: "'a abs_action signature" where |
|
16 |
"ch_asig = (UN b. {S(b)}, UN b. {R(b)}, {})" |
|
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
17 |
|
27361 | 18 |
definition |
19 |
ch_trans :: "('a abs_action, 'a multiset)transition set" where |
|
20 |
"ch_trans = |
|
21 |
{tr. let s = fst(tr); |
|
22 |
t = snd(snd(tr)) |
|
23 |
in |
|
24 |
case fst(snd(tr)) |
|
25 |
of S(b) => t = addm s b | |
|
26 |
R(b) => count s b ~= 0 & t = delm s b}" |
|
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
27 |
|
27361 | 28 |
definition |
29 |
ch_ioa :: "('a abs_action, 'a multiset)ioa" where |
|
30 |
"ch_ioa = (ch_asig, {{|}}, ch_trans,{},{})" |
|
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
31 |
|
27361 | 32 |
definition |
33 |
rsch_actions :: "'m action => bool abs_action option" where |
|
34 |
"rsch_actions (akt) = |
|
35 |
(case akt of |
|
17244 | 36 |
S_msg(m) => None | |
37 |
R_msg(m) => None | |
|
38 |
S_pkt(packet) => None | |
|
39 |
R_pkt(packet) => None | |
|
40 |
S_ack(b) => Some(S(b)) | |
|
41 |
R_ack(b) => Some(R(b)) | |
|
42 |
C_m_s => None | |
|
43 |
C_m_r => None | |
|
44 |
C_r_s => None | |
|
27361 | 45 |
C_r_r(m) => None)" |
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
46 |
|
27361 | 47 |
definition |
48 |
srch_actions :: "'m action =>(bool * 'm) abs_action option" where |
|
49 |
"srch_actions (akt) = |
|
50 |
(case akt of |
|
17244 | 51 |
S_msg(m) => None | |
52 |
R_msg(m) => None | |
|
53 |
S_pkt(p) => Some(S(p)) | |
|
54 |
R_pkt(p) => Some(R(p)) | |
|
55 |
S_ack(b) => None | |
|
56 |
R_ack(b) => None | |
|
57 |
C_m_s => None | |
|
58 |
C_m_r => None | |
|
59 |
C_r_s => None | |
|
27361 | 60 |
C_r_r(m) => None)" |
61 |
||
62 |
definition |
|
63 |
srch_ioa :: "('m action, 'm packet multiset)ioa" where |
|
64 |
"srch_ioa = rename ch_ioa srch_actions" |
|
65 |
||
66 |
definition |
|
67 |
rsch_ioa :: "('m action, bool multiset)ioa" where |
|
68 |
"rsch_ioa = rename ch_ioa rsch_actions" |
|
69 |
||
70 |
definition |
|
71 |
srch_asig :: "'m action signature" where |
|
72 |
"srch_asig = asig_of(srch_ioa)" |
|
73 |
||
74 |
definition |
|
75 |
rsch_asig :: "'m action signature" where |
|
76 |
"rsch_asig = asig_of(rsch_ioa)" |
|
77 |
||
78 |
definition |
|
79 |
srch_wfair :: "('m action)set set" where |
|
80 |
"srch_wfair = wfair_of(srch_ioa)" |
|
81 |
definition |
|
82 |
srch_sfair :: "('m action)set set" where |
|
83 |
"srch_sfair = sfair_of(srch_ioa)" |
|
84 |
definition |
|
85 |
rsch_sfair :: "('m action)set set" where |
|
86 |
"rsch_sfair = sfair_of(rsch_ioa)" |
|
87 |
definition |
|
88 |
rsch_wfair :: "('m action)set set" where |
|
89 |
"rsch_wfair = wfair_of(rsch_ioa)" |
|
90 |
||
91 |
definition |
|
92 |
srch_trans :: "('m action, 'm packet multiset)transition set" where |
|
93 |
"srch_trans = trans_of(srch_ioa)" |
|
94 |
definition |
|
95 |
rsch_trans :: "('m action, bool multiset)transition set" where |
|
96 |
"rsch_trans = trans_of(rsch_ioa)" |
|
97 |
||
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
98 |
|
19739 | 99 |
lemmas unfold_renaming = |
100 |
srch_asig_def rsch_asig_def rsch_ioa_def srch_ioa_def ch_ioa_def |
|
101 |
ch_asig_def srch_actions_def rsch_actions_def rename_def rename_set_def asig_of_def |
|
102 |
actions_def srch_trans_def rsch_trans_def ch_trans_def starts_of_def |
|
103 |
trans_of_def asig_projections |
|
104 |
||
105 |
lemma in_srch_asig: |
|
106 |
"S_msg(m) ~: actions(srch_asig) & |
|
107 |
R_msg(m) ~: actions(srch_asig) & |
|
108 |
S_pkt(pkt) : actions(srch_asig) & |
|
109 |
R_pkt(pkt) : actions(srch_asig) & |
|
110 |
S_ack(b) ~: actions(srch_asig) & |
|
111 |
R_ack(b) ~: actions(srch_asig) & |
|
112 |
C_m_s ~: actions(srch_asig) & |
|
113 |
C_m_r ~: actions(srch_asig) & |
|
114 |
C_r_s ~: actions(srch_asig) & C_r_r(m) ~: actions(srch_asig)" |
|
115 |
by (simp add: unfold_renaming) |
|
116 |
||
117 |
lemma in_rsch_asig: |
|
118 |
"S_msg(m) ~: actions(rsch_asig) & |
|
119 |
R_msg(m) ~: actions(rsch_asig) & |
|
120 |
S_pkt(pkt) ~: actions(rsch_asig) & |
|
121 |
R_pkt(pkt) ~: actions(rsch_asig) & |
|
122 |
S_ack(b) : actions(rsch_asig) & |
|
123 |
R_ack(b) : actions(rsch_asig) & |
|
124 |
C_m_s ~: actions(rsch_asig) & |
|
125 |
C_m_r ~: actions(rsch_asig) & |
|
126 |
C_r_s ~: actions(rsch_asig) & |
|
127 |
C_r_r(m) ~: actions(rsch_asig)" |
|
128 |
by (simp add: unfold_renaming) |
|
129 |
||
130 |
lemma srch_ioa_thm: "srch_ioa = |
|
131 |
(srch_asig, {{|}}, srch_trans,srch_wfair,srch_sfair)" |
|
132 |
apply (simp (no_asm) add: srch_asig_def srch_trans_def asig_of_def trans_of_def wfair_of_def sfair_of_def srch_wfair_def srch_sfair_def) |
|
133 |
apply (simp (no_asm) add: unfold_renaming) |
|
134 |
done |
|
135 |
||
136 |
lemma rsch_ioa_thm: "rsch_ioa = |
|
137 |
(rsch_asig, {{|}}, rsch_trans,rsch_wfair,rsch_sfair)" |
|
138 |
apply (simp (no_asm) add: rsch_asig_def rsch_trans_def asig_of_def trans_of_def wfair_of_def sfair_of_def rsch_wfair_def rsch_sfair_def) |
|
139 |
apply (simp (no_asm) add: unfold_renaming) |
|
140 |
done |
|
3073
88366253a09a
Old NTP files now running under the IOA meta theory based on HOLCF;
mueller
parents:
diff
changeset
|
141 |
|
17244 | 142 |
end |