author | wenzelm |
Thu, 22 Nov 2007 14:51:34 +0100 | |
changeset 25456 | 6f79698f294d |
parent 25135 | 4f8176c940cf |
child 26359 | 6d437bde2f1d |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/CompoScheds.thy |
3275 | 2 |
ID: $Id$ |
12218 | 3 |
Author: Olaf Müller |
17233 | 4 |
*) |
3071 | 5 |
|
17233 | 6 |
header {* Compositionality on Schedule level *} |
3071 | 7 |
|
17233 | 8 |
theory CompoScheds |
9 |
imports CompoExecs |
|
10 |
begin |
|
3071 | 11 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
12 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
13 |
mkex2 :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq -> |
17233 | 14 |
('a,'s)pairs -> ('a,'t)pairs -> |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
15 |
('s => 't => ('a,'s*'t)pairs)" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
16 |
"mkex2 A B = (fix$(LAM h sch exA exB. (%s t. case sch of |
3071 | 17 |
nil => nil |
17233 | 18 |
| x##xs => |
19 |
(case x of |
|
12028 | 20 |
UU => UU |
17233 | 21 |
| Def y => |
22 |
(if y:act A then |
|
23 |
(if y:act B then |
|
10835 | 24 |
(case HD$exA of |
12028 | 25 |
UU => UU |
10835 | 26 |
| Def a => (case HD$exB of |
12028 | 27 |
UU => UU |
17233 | 28 |
| Def b => |
3071 | 29 |
(y,(snd a,snd b))>> |
10835 | 30 |
(h$xs$(TL$exA)$(TL$exB)) (snd a) (snd b))) |
3071 | 31 |
else |
10835 | 32 |
(case HD$exA of |
12028 | 33 |
UU => UU |
17233 | 34 |
| Def a => |
10835 | 35 |
(y,(snd a,t))>>(h$xs$(TL$exA)$exB) (snd a) t) |
3071 | 36 |
) |
17233 | 37 |
else |
38 |
(if y:act B then |
|
10835 | 39 |
(case HD$exB of |
12028 | 40 |
UU => UU |
17233 | 41 |
| Def b => |
10835 | 42 |
(y,(s,snd b))>>(h$xs$exA$(TL$exB)) s (snd b)) |
3071 | 43 |
else |
44 |
UU |
|
45 |
) |
|
46 |
) |
|
47 |
))))" |
|
48 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
49 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
50 |
mkex :: "('a,'s)ioa => ('a,'t)ioa => 'a Seq => |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
51 |
('a,'s)execution => ('a,'t)execution =>('a,'s*'t)execution" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
52 |
"mkex A B sch exA exB = |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
53 |
((fst exA,fst exB), |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
54 |
(mkex2 A B$sch$(snd exA)$(snd exB)) (fst exA) (fst exB))" |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
55 |
|
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
56 |
definition |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
57 |
par_scheds ::"['a schedule_module,'a schedule_module] => 'a schedule_module" where |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
58 |
"par_scheds SchedsA SchedsB = |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
59 |
(let schA = fst SchedsA; sigA = snd SchedsA; |
17233 | 60 |
schB = fst SchedsB; sigB = snd SchedsB |
3521 | 61 |
in |
10835 | 62 |
( {sch. Filter (%a. a:actions sigA)$sch : schA} |
63 |
Int {sch. Filter (%a. a:actions sigB)$sch : schB} |
|
3521 | 64 |
Int {sch. Forall (%x. x:(actions sigA Un actions sigB)) sch}, |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
65 |
asig_comp sigA sigB))" |
3521 | 66 |
|
19741 | 67 |
|
68 |
declare surjective_pairing [symmetric, simp] |
|
69 |
||
70 |
||
71 |
subsection "mkex rewrite rules" |
|
72 |
||
73 |
||
74 |
lemma mkex2_unfold: |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
75 |
"mkex2 A B = (LAM sch exA exB. (%s t. case sch of |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
76 |
nil => nil |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
77 |
| x##xs => |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
78 |
(case x of |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
79 |
UU => UU |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
80 |
| Def y => |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
81 |
(if y:act A then |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
82 |
(if y:act B then |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
83 |
(case HD$exA of |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
84 |
UU => UU |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
85 |
| Def a => (case HD$exB of |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
86 |
UU => UU |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
87 |
| Def b => |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
88 |
(y,(snd a,snd b))>> |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
89 |
(mkex2 A B$xs$(TL$exA)$(TL$exB)) (snd a) (snd b))) |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
90 |
else |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
91 |
(case HD$exA of |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
92 |
UU => UU |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
93 |
| Def a => |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
94 |
(y,(snd a,t))>>(mkex2 A B$xs$(TL$exA)$exB) (snd a) t) |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
95 |
) |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
96 |
else |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
97 |
(if y:act B then |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
98 |
(case HD$exB of |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
99 |
UU => UU |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
100 |
| Def b => |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
101 |
(y,(s,snd b))>>(mkex2 A B$xs$exA$(TL$exB)) s (snd b)) |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
102 |
else |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
103 |
UU |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
104 |
) |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
105 |
) |
19741 | 106 |
)))" |
107 |
apply (rule trans) |
|
108 |
apply (rule fix_eq2) |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
109 |
apply (simp only: mkex2_def) |
19741 | 110 |
apply (rule beta_cfun) |
111 |
apply simp |
|
112 |
done |
|
113 |
||
114 |
lemma mkex2_UU: "(mkex2 A B$UU$exA$exB) s t = UU" |
|
115 |
apply (subst mkex2_unfold) |
|
116 |
apply simp |
|
117 |
done |
|
118 |
||
119 |
lemma mkex2_nil: "(mkex2 A B$nil$exA$exB) s t= nil" |
|
120 |
apply (subst mkex2_unfold) |
|
121 |
apply simp |
|
122 |
done |
|
123 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
124 |
lemma mkex2_cons_1: "[| x:act A; x~:act B; HD$exA=Def a|] |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
125 |
==> (mkex2 A B$(x>>sch)$exA$exB) s t = |
19741 | 126 |
(x,snd a,t) >> (mkex2 A B$sch$(TL$exA)$exB) (snd a) t" |
127 |
apply (rule trans) |
|
128 |
apply (subst mkex2_unfold) |
|
129 |
apply (simp add: Consq_def If_and_if) |
|
130 |
apply (simp add: Consq_def) |
|
131 |
done |
|
132 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
133 |
lemma mkex2_cons_2: "[| x~:act A; x:act B; HD$exB=Def b|] |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
134 |
==> (mkex2 A B$(x>>sch)$exA$exB) s t = |
19741 | 135 |
(x,s,snd b) >> (mkex2 A B$sch$exA$(TL$exB)) s (snd b)" |
136 |
apply (rule trans) |
|
137 |
apply (subst mkex2_unfold) |
|
138 |
apply (simp add: Consq_def If_and_if) |
|
139 |
apply (simp add: Consq_def) |
|
140 |
done |
|
141 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
142 |
lemma mkex2_cons_3: "[| x:act A; x:act B; HD$exA=Def a;HD$exB=Def b|] |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
143 |
==> (mkex2 A B$(x>>sch)$exA$exB) s t = |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
144 |
(x,snd a,snd b) >> |
19741 | 145 |
(mkex2 A B$sch$(TL$exA)$(TL$exB)) (snd a) (snd b)" |
146 |
apply (rule trans) |
|
147 |
apply (subst mkex2_unfold) |
|
148 |
apply (simp add: Consq_def If_and_if) |
|
149 |
apply (simp add: Consq_def) |
|
150 |
done |
|
151 |
||
152 |
declare mkex2_UU [simp] mkex2_nil [simp] mkex2_cons_1 [simp] |
|
153 |
mkex2_cons_2 [simp] mkex2_cons_3 [simp] |
|
154 |
||
155 |
||
156 |
subsection {* mkex *} |
|
157 |
||
158 |
lemma mkex_UU: "mkex A B UU (s,exA) (t,exB) = ((s,t),UU)" |
|
159 |
apply (simp add: mkex_def) |
|
160 |
done |
|
161 |
||
162 |
lemma mkex_nil: "mkex A B nil (s,exA) (t,exB) = ((s,t),nil)" |
|
163 |
apply (simp add: mkex_def) |
|
164 |
done |
|
165 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
166 |
lemma mkex_cons_1: "[| x:act A; x~:act B |] |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
167 |
==> mkex A B (x>>sch) (s,a>>exA) (t,exB) = |
19741 | 168 |
((s,t), (x,snd a,t) >> snd (mkex A B sch (snd a,exA) (t,exB)))" |
169 |
apply (simp (no_asm) add: mkex_def) |
|
170 |
apply (cut_tac exA = "a>>exA" in mkex2_cons_1) |
|
171 |
apply auto |
|
172 |
done |
|
173 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
174 |
lemma mkex_cons_2: "[| x~:act A; x:act B |] |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
175 |
==> mkex A B (x>>sch) (s,exA) (t,b>>exB) = |
19741 | 176 |
((s,t), (x,s,snd b) >> snd (mkex A B sch (s,exA) (snd b,exB)))" |
177 |
apply (simp (no_asm) add: mkex_def) |
|
178 |
apply (cut_tac exB = "b>>exB" in mkex2_cons_2) |
|
179 |
apply auto |
|
180 |
done |
|
181 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
182 |
lemma mkex_cons_3: "[| x:act A; x:act B |] |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
183 |
==> mkex A B (x>>sch) (s,a>>exA) (t,b>>exB) = |
19741 | 184 |
((s,t), (x,snd a,snd b) >> snd (mkex A B sch (snd a,exA) (snd b,exB)))" |
185 |
apply (simp (no_asm) add: mkex_def) |
|
186 |
apply (cut_tac exB = "b>>exB" and exA = "a>>exA" in mkex2_cons_3) |
|
187 |
apply auto |
|
188 |
done |
|
189 |
||
190 |
declare mkex2_UU [simp del] mkex2_nil [simp del] |
|
191 |
mkex2_cons_1 [simp del] mkex2_cons_2 [simp del] mkex2_cons_3 [simp del] |
|
192 |
||
193 |
lemmas composch_simps = mkex_UU mkex_nil mkex_cons_1 mkex_cons_2 mkex_cons_3 |
|
194 |
||
195 |
declare composch_simps [simp] |
|
196 |
||
197 |
||
198 |
subsection {* COMPOSITIONALITY on SCHEDULE Level *} |
|
199 |
||
200 |
subsubsection "Lemmas for ==>" |
|
201 |
||
202 |
(* --------------------------------------------------------------------- *) |
|
203 |
(* Lemma_2_1 : tfilter(ex) and filter_act are commutative *) |
|
204 |
(* --------------------------------------------------------------------- *) |
|
205 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
206 |
lemma lemma_2_1a: |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
207 |
"filter_act$(Filter_ex2 (asig_of A)$xs)= |
19741 | 208 |
Filter (%a. a:act A)$(filter_act$xs)" |
209 |
||
210 |
apply (unfold filter_act_def Filter_ex2_def) |
|
211 |
apply (simp (no_asm) add: MapFilter o_def) |
|
212 |
done |
|
213 |
||
214 |
||
215 |
(* --------------------------------------------------------------------- *) |
|
216 |
(* Lemma_2_2 : State-projections do not affect filter_act *) |
|
217 |
(* --------------------------------------------------------------------- *) |
|
218 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
219 |
lemma lemma_2_1b: |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
220 |
"filter_act$(ProjA2$xs) =filter_act$xs & |
19741 | 221 |
filter_act$(ProjB2$xs) =filter_act$xs" |
222 |
apply (tactic {* pair_induct_tac "xs" [] 1 *}) |
|
223 |
done |
|
224 |
||
225 |
||
226 |
(* --------------------------------------------------------------------- *) |
|
227 |
(* Schedules of A||B have only A- or B-actions *) |
|
228 |
(* --------------------------------------------------------------------- *) |
|
229 |
||
230 |
(* very similar to lemma_1_1c, but it is not checking if every action element of |
|
231 |
an ex is in A or B, but after projecting it onto the action schedule. Of course, this |
|
232 |
is the same proposition, but we cannot change this one, when then rather lemma_1_1c *) |
|
233 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
234 |
lemma sch_actions_in_AorB: "!s. is_exec_frag (A||B) (s,xs) |
19741 | 235 |
--> Forall (%x. x:act (A||B)) (filter_act$xs)" |
236 |
||
237 |
apply (tactic {* pair_induct_tac "xs" [thm "is_exec_frag_def", thm "Forall_def", |
|
238 |
thm "sforall_def"] 1 *}) |
|
239 |
(* main case *) |
|
240 |
apply (tactic "safe_tac set_cs") |
|
241 |
apply (auto simp add: trans_of_defs2 actions_asig_comp asig_of_par) |
|
242 |
done |
|
243 |
||
244 |
||
245 |
subsubsection "Lemmas for <==" |
|
246 |
||
247 |
(*--------------------------------------------------------------------------- |
|
248 |
Filtering actions out of mkex(sch,exA,exB) yields the oracle sch |
|
249 |
structural induction |
|
250 |
--------------------------------------------------------------------------- *) |
|
251 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
252 |
lemma Mapfst_mkex_is_sch: "! exA exB s t. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
253 |
Forall (%x. x:act (A||B)) sch & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
254 |
Filter (%a. a:act A)$sch << filter_act$exA & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
255 |
Filter (%a. a:act B)$sch << filter_act$exB |
19741 | 256 |
--> filter_act$(snd (mkex A B sch (s,exA) (t,exB))) = sch" |
257 |
||
258 |
apply (tactic {* Seq_induct_tac "sch" [thm "Filter_def", thm "Forall_def", |
|
259 |
thm "sforall_def", thm "mkex_def"] 1 *}) |
|
260 |
||
261 |
(* main case *) |
|
262 |
(* splitting into 4 cases according to a:A, a:B *) |
|
263 |
apply (tactic "safe_tac set_cs") |
|
264 |
||
265 |
(* Case y:A, y:B *) |
|
266 |
apply (tactic {* Seq_case_simp_tac "exA" 1 *}) |
|
267 |
(* Case exA=UU, Case exA=nil*) |
|
268 |
(* These UU and nil cases are the only places where the assumption filter A sch<<f_act exA |
|
269 |
is used! --> to generate a contradiction using ~a>>ss<< UU(nil), using theorems |
|
270 |
Cons_not_less_UU and Cons_not_less_nil *) |
|
271 |
apply (tactic {* Seq_case_simp_tac "exB" 1 *}) |
|
272 |
(* Case exA=a>>x, exB=b>>y *) |
|
273 |
(* here it is important that Seq_case_simp_tac uses no !full!_simp_tac for the cons case, |
|
274 |
as otherwise mkex_cons_3 would not be rewritten without use of rotate_tac: then tactic |
|
275 |
would not be generally applicable *) |
|
276 |
apply simp |
|
277 |
||
278 |
(* Case y:A, y~:B *) |
|
279 |
apply (tactic {* Seq_case_simp_tac "exA" 1 *}) |
|
280 |
apply simp |
|
281 |
||
282 |
(* Case y~:A, y:B *) |
|
283 |
apply (tactic {* Seq_case_simp_tac "exB" 1 *}) |
|
284 |
apply simp |
|
285 |
||
286 |
(* Case y~:A, y~:B *) |
|
287 |
apply (simp add: asig_of_par actions_asig_comp) |
|
288 |
done |
|
289 |
||
290 |
||
291 |
(* generalizing the proof above to a tactic *) |
|
292 |
||
293 |
ML {* |
|
294 |
||
295 |
local |
|
296 |
val defs = [thm "Filter_def", thm "Forall_def", thm "sforall_def", thm "mkex_def", |
|
297 |
thm "stutter_def"] |
|
298 |
val asigs = [thm "asig_of_par", thm "actions_asig_comp"] |
|
299 |
in |
|
300 |
||
301 |
fun mkex_induct_tac sch exA exB = |
|
302 |
EVERY1[Seq_induct_tac sch defs, |
|
303 |
SIMPSET' asm_full_simp_tac, |
|
304 |
SELECT_GOAL (safe_tac set_cs), |
|
305 |
Seq_case_simp_tac exA, |
|
306 |
Seq_case_simp_tac exB, |
|
307 |
SIMPSET' asm_full_simp_tac, |
|
308 |
Seq_case_simp_tac exA, |
|
309 |
SIMPSET' asm_full_simp_tac, |
|
310 |
Seq_case_simp_tac exB, |
|
311 |
SIMPSET' asm_full_simp_tac, |
|
312 |
SIMPSET' (fn ss => asm_full_simp_tac (ss addsimps asigs)) |
|
313 |
] |
|
17233 | 314 |
|
3521 | 315 |
end |
19741 | 316 |
*} |
317 |
||
318 |
||
319 |
(*--------------------------------------------------------------------------- |
|
320 |
Projection of mkex(sch,exA,exB) onto A stutters on A |
|
321 |
structural induction |
|
322 |
--------------------------------------------------------------------------- *) |
|
323 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
324 |
lemma stutterA_mkex: "! exA exB s t. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
325 |
Forall (%x. x:act (A||B)) sch & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
326 |
Filter (%a. a:act A)$sch << filter_act$exA & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
327 |
Filter (%a. a:act B)$sch << filter_act$exB |
19741 | 328 |
--> stutter (asig_of A) (s,ProjA2$(snd (mkex A B sch (s,exA) (t,exB))))" |
329 |
||
330 |
apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *}) |
|
331 |
done |
|
332 |
||
333 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
334 |
lemma stutter_mkex_on_A: "[| |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
335 |
Forall (%x. x:act (A||B)) sch ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
336 |
Filter (%a. a:act A)$sch << filter_act$(snd exA) ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
337 |
Filter (%a. a:act B)$sch << filter_act$(snd exB) |] |
19741 | 338 |
==> stutter (asig_of A) (ProjA (mkex A B sch exA exB))" |
339 |
||
340 |
apply (cut_tac stutterA_mkex) |
|
341 |
apply (simp add: stutter_def ProjA_def mkex_def) |
|
342 |
apply (erule allE)+ |
|
343 |
apply (drule mp) |
|
344 |
prefer 2 apply (assumption) |
|
345 |
apply simp |
|
346 |
done |
|
347 |
||
348 |
||
349 |
(*--------------------------------------------------------------------------- |
|
350 |
Projection of mkex(sch,exA,exB) onto B stutters on B |
|
351 |
structural induction |
|
352 |
--------------------------------------------------------------------------- *) |
|
353 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
354 |
lemma stutterB_mkex: "! exA exB s t. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
355 |
Forall (%x. x:act (A||B)) sch & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
356 |
Filter (%a. a:act A)$sch << filter_act$exA & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
357 |
Filter (%a. a:act B)$sch << filter_act$exB |
19741 | 358 |
--> stutter (asig_of B) (t,ProjB2$(snd (mkex A B sch (s,exA) (t,exB))))" |
359 |
apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *}) |
|
360 |
done |
|
361 |
||
362 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
363 |
lemma stutter_mkex_on_B: "[| |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
364 |
Forall (%x. x:act (A||B)) sch ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
365 |
Filter (%a. a:act A)$sch << filter_act$(snd exA) ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
366 |
Filter (%a. a:act B)$sch << filter_act$(snd exB) |] |
19741 | 367 |
==> stutter (asig_of B) (ProjB (mkex A B sch exA exB))" |
368 |
apply (cut_tac stutterB_mkex) |
|
369 |
apply (simp add: stutter_def ProjB_def mkex_def) |
|
370 |
apply (erule allE)+ |
|
371 |
apply (drule mp) |
|
372 |
prefer 2 apply (assumption) |
|
373 |
apply simp |
|
374 |
done |
|
375 |
||
376 |
||
377 |
(*--------------------------------------------------------------------------- |
|
378 |
Filter of mkex(sch,exA,exB) to A after projection onto A is exA |
|
379 |
-- using zip$(proj1$exA)$(proj2$exA) instead of exA -- |
|
380 |
-- because of admissibility problems -- |
|
381 |
structural induction |
|
382 |
--------------------------------------------------------------------------- *) |
|
383 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
384 |
lemma filter_mkex_is_exA_tmp: "! exA exB s t. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
385 |
Forall (%x. x:act (A||B)) sch & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
386 |
Filter (%a. a:act A)$sch << filter_act$exA & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
387 |
Filter (%a. a:act B)$sch << filter_act$exB |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
388 |
--> Filter_ex2 (asig_of A)$(ProjA2$(snd (mkex A B sch (s,exA) (t,exB)))) = |
19741 | 389 |
Zip$(Filter (%a. a:act A)$sch)$(Map snd$exA)" |
390 |
apply (tactic {* mkex_induct_tac "sch" "exB" "exA" *}) |
|
391 |
done |
|
392 |
||
393 |
(*--------------------------------------------------------------------------- |
|
394 |
zip$(proj1$y)$(proj2$y) = y (using the lift operations) |
|
395 |
lemma for admissibility problems |
|
396 |
--------------------------------------------------------------------------- *) |
|
397 |
||
398 |
lemma Zip_Map_fst_snd: "Zip$(Map fst$y)$(Map snd$y) = y" |
|
399 |
apply (tactic {* Seq_induct_tac "y" [] 1 *}) |
|
400 |
done |
|
401 |
||
402 |
||
403 |
(*--------------------------------------------------------------------------- |
|
404 |
filter A$sch = proj1$ex --> zip$(filter A$sch)$(proj2$ex) = ex |
|
405 |
lemma for eliminating non admissible equations in assumptions |
|
406 |
--------------------------------------------------------------------------- *) |
|
407 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
408 |
lemma trick_against_eq_in_ass: "!! sch ex. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
409 |
Filter (%a. a:act AB)$sch = filter_act$ex |
19741 | 410 |
==> ex = Zip$(Filter (%a. a:act AB)$sch)$(Map snd$ex)" |
411 |
apply (simp add: filter_act_def) |
|
412 |
apply (rule Zip_Map_fst_snd [symmetric]) |
|
413 |
done |
|
414 |
||
415 |
(*--------------------------------------------------------------------------- |
|
416 |
Filter of mkex(sch,exA,exB) to A after projection onto A is exA |
|
417 |
using the above trick |
|
418 |
--------------------------------------------------------------------------- *) |
|
419 |
||
420 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
421 |
lemma filter_mkex_is_exA: "!!sch exA exB. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
422 |
[| Forall (%a. a:act (A||B)) sch ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
423 |
Filter (%a. a:act A)$sch = filter_act$(snd exA) ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
424 |
Filter (%a. a:act B)$sch = filter_act$(snd exB) |] |
19741 | 425 |
==> Filter_ex (asig_of A) (ProjA (mkex A B sch exA exB)) = exA" |
426 |
apply (simp add: ProjA_def Filter_ex_def) |
|
427 |
apply (tactic {* pair_tac "exA" 1 *}) |
|
428 |
apply (tactic {* pair_tac "exB" 1 *}) |
|
429 |
apply (rule conjI) |
|
430 |
apply (simp (no_asm) add: mkex_def) |
|
431 |
apply (simplesubst trick_against_eq_in_ass) |
|
432 |
back |
|
433 |
apply assumption |
|
434 |
apply (simp add: filter_mkex_is_exA_tmp) |
|
435 |
done |
|
436 |
||
437 |
||
438 |
(*--------------------------------------------------------------------------- |
|
439 |
Filter of mkex(sch,exA,exB) to B after projection onto B is exB |
|
440 |
-- using zip$(proj1$exB)$(proj2$exB) instead of exB -- |
|
441 |
-- because of admissibility problems -- |
|
442 |
structural induction |
|
443 |
--------------------------------------------------------------------------- *) |
|
444 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
445 |
lemma filter_mkex_is_exB_tmp: "! exA exB s t. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
446 |
Forall (%x. x:act (A||B)) sch & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
447 |
Filter (%a. a:act A)$sch << filter_act$exA & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
448 |
Filter (%a. a:act B)$sch << filter_act$exB |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
449 |
--> Filter_ex2 (asig_of B)$(ProjB2$(snd (mkex A B sch (s,exA) (t,exB)))) = |
19741 | 450 |
Zip$(Filter (%a. a:act B)$sch)$(Map snd$exB)" |
451 |
||
452 |
(* notice necessary change of arguments exA and exB *) |
|
453 |
apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *}) |
|
454 |
done |
|
455 |
||
456 |
||
457 |
(*--------------------------------------------------------------------------- |
|
458 |
Filter of mkex(sch,exA,exB) to A after projection onto B is exB |
|
459 |
using the above trick |
|
460 |
--------------------------------------------------------------------------- *) |
|
461 |
||
462 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
463 |
lemma filter_mkex_is_exB: "!!sch exA exB. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
464 |
[| Forall (%a. a:act (A||B)) sch ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
465 |
Filter (%a. a:act A)$sch = filter_act$(snd exA) ; |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
466 |
Filter (%a. a:act B)$sch = filter_act$(snd exB) |] |
19741 | 467 |
==> Filter_ex (asig_of B) (ProjB (mkex A B sch exA exB)) = exB" |
468 |
apply (simp add: ProjB_def Filter_ex_def) |
|
469 |
apply (tactic {* pair_tac "exA" 1 *}) |
|
470 |
apply (tactic {* pair_tac "exB" 1 *}) |
|
471 |
apply (rule conjI) |
|
472 |
apply (simp (no_asm) add: mkex_def) |
|
473 |
apply (simplesubst trick_against_eq_in_ass) |
|
474 |
back |
|
475 |
apply assumption |
|
476 |
apply (simp add: filter_mkex_is_exB_tmp) |
|
477 |
done |
|
478 |
||
479 |
(* --------------------------------------------------------------------- *) |
|
480 |
(* mkex has only A- or B-actions *) |
|
481 |
(* --------------------------------------------------------------------- *) |
|
482 |
||
483 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
484 |
lemma mkex_actions_in_AorB: "!s t exA exB. |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
485 |
Forall (%x. x : act (A || B)) sch & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
486 |
Filter (%a. a:act A)$sch << filter_act$exA & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
487 |
Filter (%a. a:act B)$sch << filter_act$exB |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
488 |
--> Forall (%x. fst x : act (A ||B)) |
19741 | 489 |
(snd (mkex A B sch (s,exA) (t,exB)))" |
490 |
apply (tactic {* mkex_induct_tac "sch" "exA" "exB" *}) |
|
491 |
done |
|
492 |
||
493 |
||
494 |
(* ------------------------------------------------------------------ *) |
|
495 |
(* COMPOSITIONALITY on SCHEDULE Level *) |
|
496 |
(* Main Theorem *) |
|
497 |
(* ------------------------------------------------------------------ *) |
|
498 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
499 |
lemma compositionality_sch: |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
500 |
"(sch : schedules (A||B)) = |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
501 |
(Filter (%a. a:act A)$sch : schedules A & |
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
502 |
Filter (%a. a:act B)$sch : schedules B & |
19741 | 503 |
Forall (%x. x:act (A||B)) sch)" |
504 |
apply (simp (no_asm) add: schedules_def has_schedule_def) |
|
505 |
apply (tactic "safe_tac set_cs") |
|
506 |
(* ==> *) |
|
507 |
apply (rule_tac x = "Filter_ex (asig_of A) (ProjA ex) " in bexI) |
|
508 |
prefer 2 |
|
509 |
apply (simp add: compositionality_ex) |
|
510 |
apply (simp (no_asm) add: Filter_ex_def ProjA_def lemma_2_1a lemma_2_1b) |
|
511 |
apply (rule_tac x = "Filter_ex (asig_of B) (ProjB ex) " in bexI) |
|
512 |
prefer 2 |
|
513 |
apply (simp add: compositionality_ex) |
|
514 |
apply (simp (no_asm) add: Filter_ex_def ProjB_def lemma_2_1a lemma_2_1b) |
|
515 |
apply (simp add: executions_def) |
|
516 |
apply (tactic {* pair_tac "ex" 1 *}) |
|
517 |
apply (erule conjE) |
|
518 |
apply (simp add: sch_actions_in_AorB) |
|
519 |
||
520 |
(* <== *) |
|
521 |
||
522 |
(* mkex is exactly the construction of exA||B out of exA, exB, and the oracle sch, |
|
523 |
we need here *) |
|
524 |
apply (rename_tac exA exB) |
|
525 |
apply (rule_tac x = "mkex A B sch exA exB" in bexI) |
|
526 |
(* mkex actions are just the oracle *) |
|
527 |
apply (tactic {* pair_tac "exA" 1 *}) |
|
528 |
apply (tactic {* pair_tac "exB" 1 *}) |
|
529 |
apply (simp add: Mapfst_mkex_is_sch) |
|
530 |
||
531 |
(* mkex is an execution -- use compositionality on ex-level *) |
|
532 |
apply (simp add: compositionality_ex) |
|
533 |
apply (simp add: stutter_mkex_on_A stutter_mkex_on_B filter_mkex_is_exB filter_mkex_is_exA) |
|
534 |
apply (tactic {* pair_tac "exA" 1 *}) |
|
535 |
apply (tactic {* pair_tac "exB" 1 *}) |
|
536 |
apply (simp add: mkex_actions_in_AorB) |
|
537 |
done |
|
538 |
||
539 |
||
540 |
subsection {* COMPOSITIONALITY on SCHEDULE Level -- for Modules *} |
|
541 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
19741
diff
changeset
|
542 |
lemma compositionality_sch_modules: |
19741 | 543 |
"Scheds (A||B) = par_scheds (Scheds A) (Scheds B)" |
544 |
||
545 |
apply (unfold Scheds_def par_scheds_def) |
|
546 |
apply (simp add: asig_of_par) |
|
547 |
apply (rule set_ext) |
|
548 |
apply (simp add: compositionality_sch actions_of_par) |
|
549 |
done |
|
550 |
||
551 |
||
552 |
declare compoex_simps [simp del] |
|
553 |
declare composch_simps [simp del] |
|
554 |
||
555 |
end |