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