author | wenzelm |
Mon, 22 May 2000 11:56:55 +0200 | |
changeset 8907 | 813fabceec00 |
parent 7499 | 23e090051cb8 |
child 9877 | b2a62260f8ac |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/CompoTraces.ML |
3080 | 2 |
ID: $Id$ |
3071 | 3 |
Author: Olaf M"uller |
4 |
Copyright 1996 TU Muenchen |
|
3275 | 5 |
|
3071 | 6 |
Compositionality on Trace level. |
7 |
*) |
|
8 |
||
4678 | 9 |
simpset_ref () := simpset() setmksym (K None); |
3275 | 10 |
|
3071 | 11 |
(* ---------------------------------------------------------------- *) |
12 |
section "mksch rewrite rules"; |
|
13 |
(* ---------------------------------------------------------------- *) |
|
14 |
||
15 |
||
16 |
||
17 |
||
18 |
bind_thm ("mksch_unfold", fix_prover2 thy mksch_def |
|
19 |
"mksch A B = (LAM tr schA schB. case tr of \ |
|
20 |
\ nil => nil\ |
|
21 |
\ | x##xs => \ |
|
22 |
\ (case x of \ |
|
23 |
\ Undef => UU \ |
|
24 |
\ | Def y => \ |
|
25 |
\ (if y:act A then \ |
|
26 |
\ (if y:act B then \ |
|
3842 | 27 |
\ ((Takewhile (%a. a:int A)`schA) \ |
28 |
\ @@(Takewhile (%a. a:int B)`schB) \ |
|
3071 | 29 |
\ @@(y>>(mksch A B`xs \ |
3842 | 30 |
\ `(TL`(Dropwhile (%a. a:int A)`schA)) \ |
31 |
\ `(TL`(Dropwhile (%a. a:int B)`schB)) \ |
|
3071 | 32 |
\ ))) \ |
33 |
\ else \ |
|
3842 | 34 |
\ ((Takewhile (%a. a:int A)`schA) \ |
3071 | 35 |
\ @@ (y>>(mksch A B`xs \ |
3842 | 36 |
\ `(TL`(Dropwhile (%a. a:int A)`schA)) \ |
3071 | 37 |
\ `schB))) \ |
38 |
\ ) \ |
|
39 |
\ else \ |
|
40 |
\ (if y:act B then \ |
|
3842 | 41 |
\ ((Takewhile (%a. a:int B)`schB) \ |
3071 | 42 |
\ @@ (y>>(mksch A B`xs \ |
43 |
\ `schA \ |
|
3842 | 44 |
\ `(TL`(Dropwhile (%a. a:int B)`schB)) \ |
3071 | 45 |
\ ))) \ |
46 |
\ else \ |
|
47 |
\ UU \ |
|
48 |
\ ) \ |
|
49 |
\ ) \ |
|
50 |
\ ))"); |
|
51 |
||
52 |
||
5068 | 53 |
Goal "mksch A B`UU`schA`schB = UU"; |
3071 | 54 |
by (stac mksch_unfold 1); |
55 |
by (Simp_tac 1); |
|
56 |
qed"mksch_UU"; |
|
57 |
||
5068 | 58 |
Goal "mksch A B`nil`schA`schB = nil"; |
3071 | 59 |
by (stac mksch_unfold 1); |
60 |
by (Simp_tac 1); |
|
61 |
qed"mksch_nil"; |
|
62 |
||
6161 | 63 |
Goal "[|x:act A;x~:act B|] \ |
3071 | 64 |
\ ==> mksch A B`(x>>tr)`schA`schB = \ |
3842 | 65 |
\ (Takewhile (%a. a:int A)`schA) \ |
66 |
\ @@ (x>>(mksch A B`tr`(TL`(Dropwhile (%a. a:int A)`schA)) \ |
|
3071 | 67 |
\ `schB))"; |
3457 | 68 |
by (rtac trans 1); |
3071 | 69 |
by (stac mksch_unfold 1); |
7229
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,If_and_if]) 1); |
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
6161
diff
changeset
|
71 |
by (simp_tac (simpset() addsimps [Consq_def]) 1); |
3071 | 72 |
qed"mksch_cons1"; |
73 |
||
6161 | 74 |
Goal "[|x~:act A;x:act B|] \ |
3071 | 75 |
\ ==> mksch A B`(x>>tr)`schA`schB = \ |
3842 | 76 |
\ (Takewhile (%a. a:int B)`schB) \ |
77 |
\ @@ (x>>(mksch A B`tr`schA`(TL`(Dropwhile (%a. a:int B)`schB)) \ |
|
3071 | 78 |
\ ))"; |
3457 | 79 |
by (rtac trans 1); |
3071 | 80 |
by (stac mksch_unfold 1); |
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
6161
diff
changeset
|
81 |
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
|
82 |
by (simp_tac (simpset() addsimps [Consq_def]) 1); |
3071 | 83 |
qed"mksch_cons2"; |
84 |
||
6161 | 85 |
Goal "[|x:act A;x:act B|] \ |
3071 | 86 |
\ ==> mksch A B`(x>>tr)`schA`schB = \ |
3842 | 87 |
\ (Takewhile (%a. a:int A)`schA) \ |
88 |
\ @@ ((Takewhile (%a. a:int B)`schB) \ |
|
89 |
\ @@ (x>>(mksch A B`tr`(TL`(Dropwhile (%a. a:int A)`schA)) \ |
|
90 |
\ `(TL`(Dropwhile (%a. a:int B)`schB)))) \ |
|
3071 | 91 |
\ )"; |
3457 | 92 |
by (rtac trans 1); |
3071 | 93 |
by (stac mksch_unfold 1); |
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
6161
diff
changeset
|
94 |
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
|
95 |
by (simp_tac (simpset() addsimps [Consq_def]) 1); |
3071 | 96 |
qed"mksch_cons3"; |
97 |
||
98 |
val compotr_simps =[mksch_UU,mksch_nil, mksch_cons1,mksch_cons2,mksch_cons3]; |
|
99 |
||
100 |
Addsimps compotr_simps; |
|
101 |
||
102 |
||
103 |
(* ------------------------------------------------------------------ *) |
|
104 |
(* The following lemmata aim for *) |
|
105 |
(* COMPOSITIONALITY on TRACE Level *) |
|
106 |
(* ------------------------------------------------------------------ *) |
|
107 |
||
108 |
||
109 |
(* ---------------------------------------------------------------------------- *) |
|
3275 | 110 |
section"Lemmata for ==>"; |
3071 | 111 |
(* ---------------------------------------------------------------------------- *) |
112 |
||
113 |
(* Consequence out of ext1_ext2_is_not_act1(2), which in turn are consequences out of |
|
114 |
the compatibility of IOA, in particular out of the condition that internals are |
|
115 |
really hidden. *) |
|
116 |
||
5068 | 117 |
Goal "(eB & ~eA --> ~A) --> \ |
3071 | 118 |
\ (A & (eA | eB)) = (eA & A)"; |
119 |
by (Fast_tac 1); |
|
120 |
qed"compatibility_consequence1"; |
|
121 |
||
122 |
||
123 |
(* very similar to above, only the commutativity of | is used to make a slight change *) |
|
124 |
||
5068 | 125 |
Goal "(eB & ~eA --> ~A) --> \ |
3071 | 126 |
\ (A & (eB | eA)) = (eA & A)"; |
127 |
by (Fast_tac 1); |
|
128 |
qed"compatibility_consequence2"; |
|
129 |
||
130 |
||
3275 | 131 |
(* ---------------------------------------------------------------------------- *) |
132 |
section"Lemmata for <=="; |
|
133 |
(* ---------------------------------------------------------------------------- *) |
|
3071 | 134 |
|
135 |
||
3275 | 136 |
(* Lemma for substitution of looping assumption in another specific assumption *) |
137 |
val [p1,p2] = goal thy "[| f << (g x) ; x=(h x) |] ==> f << g (h x)"; |
|
138 |
by (cut_facts_tac [p1] 1); |
|
3457 | 139 |
by (etac (p2 RS subst) 1); |
3275 | 140 |
qed"subst_lemma1"; |
3071 | 141 |
|
3275 | 142 |
(* Lemma for substitution of looping assumption in another specific assumption *) |
143 |
val [p1,p2] = goal thy "[| (f x) = y >> g; x=(h x) |] ==> (f (h x)) = y >> g"; |
|
144 |
by (cut_facts_tac [p1] 1); |
|
3457 | 145 |
by (etac (p2 RS subst) 1); |
3275 | 146 |
qed"subst_lemma2"; |
3071 | 147 |
|
148 |
||
5068 | 149 |
Goal "!!A B. compatible A B ==> \ |
3071 | 150 |
\ ! schA schB. Forall (%x. x:act (A||B)) tr \ |
151 |
\ --> Forall (%x. x:act (A||B)) (mksch A B`tr`schA`schB)"; |
|
152 |
by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1); |
|
153 |
by (safe_tac set_cs); |
|
4098 | 154 |
by (asm_full_simp_tac (simpset() addsimps [actions_of_par]) 1); |
3071 | 155 |
by (case_tac "a:act A" 1); |
156 |
by (case_tac "a:act B" 1); |
|
157 |
(* a:A, a:B *) |
|
158 |
by (Asm_full_simp_tac 1); |
|
3457 | 159 |
by (rtac (Forall_Conc_impl RS mp) 1); |
4098 | 160 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1); |
3457 | 161 |
by (rtac (Forall_Conc_impl RS mp) 1); |
4098 | 162 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1); |
3071 | 163 |
(* a:A,a~:B *) |
164 |
by (Asm_full_simp_tac 1); |
|
3457 | 165 |
by (rtac (Forall_Conc_impl RS mp) 1); |
4098 | 166 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1); |
3071 | 167 |
by (case_tac "a:act B" 1); |
168 |
(* a~:A, a:B *) |
|
169 |
by (Asm_full_simp_tac 1); |
|
3457 | 170 |
by (rtac (Forall_Conc_impl RS mp) 1); |
4098 | 171 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ,intA_is_not_actB,int_is_act]) 1); |
3071 | 172 |
(* a~:A,a~:B *) |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
173 |
by Auto_tac; |
4678 | 174 |
qed_spec_mp"ForallAorB_mksch"; |
3071 | 175 |
|
5068 | 176 |
Goal "!!A B. compatible B A ==> \ |
3071 | 177 |
\ ! schA schB. (Forall (%x. x:act B & x~:act A) tr \ |
178 |
\ --> Forall (%x. x:act B & x~:act A) (mksch A B`tr`schA`schB))"; |
|
179 |
||
180 |
by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1); |
|
181 |
by (safe_tac set_cs); |
|
3457 | 182 |
by (rtac (Forall_Conc_impl RS mp) 1); |
4098 | 183 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, |
3275 | 184 |
intA_is_not_actB,int_is_act]) 1); |
4678 | 185 |
qed_spec_mp "ForallBnAmksch"; |
3275 | 186 |
|
5068 | 187 |
Goal "!!A B. compatible A B ==> \ |
3071 | 188 |
\ ! schA schB. (Forall (%x. x:act A & x~:act B) tr \ |
189 |
\ --> Forall (%x. x:act A & x~:act B) (mksch A B`tr`schA`schB))"; |
|
190 |
||
191 |
by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1); |
|
192 |
by (safe_tac set_cs); |
|
193 |
by (Asm_full_simp_tac 1); |
|
3457 | 194 |
by (rtac (Forall_Conc_impl RS mp) 1); |
4098 | 195 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, |
3275 | 196 |
intA_is_not_actB,int_is_act]) 1); |
4678 | 197 |
qed_spec_mp"ForallAnBmksch"; |
3071 | 198 |
|
3275 | 199 |
(* safe-tac makes too many case distinctions with this lemma in the next proof *) |
200 |
Delsimps [FiniteConc]; |
|
3071 | 201 |
|
6161 | 202 |
Goal "[| Finite tr; is_asig(asig_of A); is_asig(asig_of B) |] ==> \ |
3842 | 203 |
\ ! x y. Forall (%x. x:act A) x & Forall (%x. x:act B) y & \ |
3275 | 204 |
\ Filter (%a. a:ext A)`x = Filter (%a. a:act A)`tr & \ |
3071 | 205 |
\ Filter (%a. a:ext B)`y = Filter (%a. a:act B)`tr &\ |
206 |
\ Forall (%x. x:ext (A||B)) tr \ |
|
207 |
\ --> Finite (mksch A B`tr`x`y)"; |
|
208 |
||
3457 | 209 |
by (etac Seq_Finite_ind 1); |
3071 | 210 |
by (Asm_full_simp_tac 1); |
211 |
(* main case *) |
|
4833 | 212 |
by (Asm_full_simp_tac 1); |
3071 | 213 |
by (safe_tac set_cs); |
3275 | 214 |
|
215 |
(* a: act A; a: act B *) |
|
216 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
217 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
218 |
back(); |
|
219 |
by (REPEAT (etac conjE 1)); |
|
220 |
(* Finite (tw iA x) and Finite (tw iB y) *) |
|
4098 | 221 |
by (asm_full_simp_tac (simpset() addsimps |
3275 | 222 |
[not_ext_is_int_or_not_act,FiniteConc]) 1); |
223 |
(* now for conclusion IH applicable, but assumptions have to be transformed *) |
|
224 |
by (dres_inst_tac [("x","x"), |
|
225 |
("g","Filter (%a. a:act A)`s")] subst_lemma2 1); |
|
3457 | 226 |
by (assume_tac 1); |
3275 | 227 |
by (dres_inst_tac [("x","y"), |
228 |
("g","Filter (%a. a:act B)`s")] subst_lemma2 1); |
|
3457 | 229 |
by (assume_tac 1); |
3275 | 230 |
(* IH *) |
4678 | 231 |
by (asm_full_simp_tac (simpset() |
232 |
addsimps [not_ext_is_int_or_not_act,ForallTL,ForallDropwhile]) 1); |
|
3275 | 233 |
|
234 |
(* a: act B; a~: act A *) |
|
235 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
3071 | 236 |
|
3275 | 237 |
by (REPEAT (etac conjE 1)); |
238 |
(* Finite (tw iB y) *) |
|
4098 | 239 |
by (asm_full_simp_tac (simpset() addsimps |
3275 | 240 |
[not_ext_is_int_or_not_act,FiniteConc]) 1); |
241 |
(* now for conclusion IH applicable, but assumptions have to be transformed *) |
|
242 |
by (dres_inst_tac [("x","y"), |
|
243 |
("g","Filter (%a. a:act B)`s")] subst_lemma2 1); |
|
3457 | 244 |
by (assume_tac 1); |
3275 | 245 |
(* IH *) |
4678 | 246 |
by (asm_full_simp_tac (simpset() |
247 |
addsimps [not_ext_is_int_or_not_act,ForallTL,ForallDropwhile]) 1); |
|
3071 | 248 |
|
3275 | 249 |
(* a~: act B; a: act A *) |
250 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
3071 | 251 |
|
3275 | 252 |
by (REPEAT (etac conjE 1)); |
253 |
(* Finite (tw iA x) *) |
|
4098 | 254 |
by (asm_full_simp_tac (simpset() addsimps |
3275 | 255 |
[not_ext_is_int_or_not_act,FiniteConc]) 1); |
256 |
(* now for conclusion IH applicable, but assumptions have to be transformed *) |
|
257 |
by (dres_inst_tac [("x","x"), |
|
258 |
("g","Filter (%a. a:act A)`s")] subst_lemma2 1); |
|
3457 | 259 |
by (assume_tac 1); |
3275 | 260 |
(* IH *) |
4678 | 261 |
by (asm_full_simp_tac (simpset() |
262 |
addsimps [not_ext_is_int_or_not_act,ForallTL,ForallDropwhile]) 1); |
|
3275 | 263 |
|
264 |
(* a~: act B; a~: act A *) |
|
4098 | 265 |
by (fast_tac (claset() addSIs [ext_is_act] |
266 |
addss (simpset() addsimps [externals_of_par]) ) 1); |
|
4678 | 267 |
qed_spec_mp"FiniteL_mksch"; |
3071 | 268 |
|
3275 | 269 |
Addsimps [FiniteConc]; |
270 |
||
271 |
||
272 |
(* otherwise subst_lemma2 does not fit exactly, just to avoid a subst_lemma3 *) |
|
273 |
Delsimps [FilterConc]; |
|
274 |
||
6161 | 275 |
Goal " [| Finite bs; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==> \ |
3842 | 276 |
\! y. Forall (%x. x:act B) y & Forall (%x. x:act B & x~:act A) bs &\ |
3275 | 277 |
\ Filter (%a. a:ext B)`y = Filter (%a. a:act B)`(bs @@ z) \ |
278 |
\ --> (? y1 y2. (mksch A B`(bs @@ z)`x`y) = (y1 @@ (mksch A B`z`x`y2)) & \ |
|
279 |
\ Forall (%x. x:act B & x~:act A) y1 & \ |
|
280 |
\ Finite y1 & y = (y1 @@ y2) & \ |
|
281 |
\ Filter (%a. a:ext B)`y1 = bs)"; |
|
4473 | 282 |
by (forw_inst_tac [("A1","A")] (compat_commute RS iffD1) 1); |
3457 | 283 |
by (etac Seq_Finite_ind 1); |
3275 | 284 |
by (REPEAT (rtac allI 1)); |
3071 | 285 |
by (rtac impI 1); |
286 |
by (res_inst_tac [("x","nil")] exI 1); |
|
287 |
by (res_inst_tac [("x","y")] exI 1); |
|
288 |
by (Asm_full_simp_tac 1); |
|
289 |
(* main case *) |
|
3275 | 290 |
by (REPEAT (rtac allI 1)); |
3071 | 291 |
by (rtac impI 1); |
292 |
by (Asm_full_simp_tac 1); |
|
293 |
by (REPEAT (etac conjE 1)); |
|
3275 | 294 |
by (Asm_full_simp_tac 1); |
295 |
(* divide_Seq on s *) |
|
296 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
297 |
by (REPEAT (etac conjE 1)); |
|
298 |
(* transform assumption f eB y = f B (s@z) *) |
|
299 |
by (dres_inst_tac [("x","y"), |
|
300 |
("g","Filter (%a. a:act B)`(s@@z)")] subst_lemma2 1); |
|
3457 | 301 |
by (assume_tac 1); |
3275 | 302 |
Addsimps [FilterConc]; |
4678 | 303 |
by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act]) 1); |
3275 | 304 |
(* apply IH *) |
3842 | 305 |
by (eres_inst_tac [("x","TL`(Dropwhile (%a. a:int B)`y)")] allE 1); |
4098 | 306 |
by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile])1); |
3275 | 307 |
by (REPEAT (etac exE 1)); |
308 |
by (REPEAT (etac conjE 1)); |
|
309 |
by (Asm_full_simp_tac 1); |
|
310 |
(* for replacing IH in conclusion *) |
|
311 |
by (rotate_tac ~2 1); |
|
312 |
by (Asm_full_simp_tac 1); |
|
313 |
(* instantiate y1a and y2a *) |
|
3842 | 314 |
by (res_inst_tac [("x","Takewhile (%a. a:int B)`y @@ a>>y1")] exI 1); |
3275 | 315 |
by (res_inst_tac [("x","y2")] exI 1); |
316 |
(* elminate all obligations up to two depending on Conc_assoc *) |
|
4098 | 317 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, intA_is_not_actB, |
4678 | 318 |
int_is_act,int_is_not_ext]) 1); |
4098 | 319 |
by (simp_tac (simpset() addsimps [Conc_assoc]) 1); |
4678 | 320 |
qed_spec_mp "reduceA_mksch1"; |
3275 | 321 |
|
4678 | 322 |
bind_thm("reduceA_mksch",conjI RSN (6,conjI RSN (5,reduceA_mksch1))); |
3275 | 323 |
|
324 |
||
325 |
(* otherwise subst_lemma2 does not fit exactly, just to avoid a subst_lemma3 *) |
|
326 |
Delsimps [FilterConc]; |
|
327 |
||
328 |
||
6161 | 329 |
Goal " [| Finite as; is_asig(asig_of A); is_asig(asig_of B);compatible A B|] ==> \ |
3842 | 330 |
\! x. Forall (%x. x:act A) x & Forall (%x. x:act A & x~:act B) as &\ |
3275 | 331 |
\ Filter (%a. a:ext A)`x = Filter (%a. a:act A)`(as @@ z) \ |
332 |
\ --> (? x1 x2. (mksch A B`(as @@ z)`x`y) = (x1 @@ (mksch A B`z`x2`y)) & \ |
|
333 |
\ Forall (%x. x:act A & x~:act B) x1 & \ |
|
334 |
\ Finite x1 & x = (x1 @@ x2) & \ |
|
335 |
\ Filter (%a. a:ext A)`x1 = as)"; |
|
4473 | 336 |
by (forw_inst_tac [("A1","A")] (compat_commute RS iffD1) 1); |
3457 | 337 |
by (etac Seq_Finite_ind 1); |
3275 | 338 |
by (REPEAT (rtac allI 1)); |
339 |
by (rtac impI 1); |
|
340 |
by (res_inst_tac [("x","nil")] exI 1); |
|
341 |
by (res_inst_tac [("x","x")] exI 1); |
|
342 |
by (Asm_full_simp_tac 1); |
|
343 |
(* main case *) |
|
344 |
by (REPEAT (rtac allI 1)); |
|
345 |
by (rtac impI 1); |
|
346 |
by (Asm_full_simp_tac 1); |
|
347 |
by (REPEAT (etac conjE 1)); |
|
348 |
by (Asm_full_simp_tac 1); |
|
349 |
(* divide_Seq on s *) |
|
350 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
351 |
by (REPEAT (etac conjE 1)); |
|
352 |
(* transform assumption f eA x = f A (s@z) *) |
|
353 |
by (dres_inst_tac [("x","x"), |
|
354 |
("g","Filter (%a. a:act A)`(s@@z)")] subst_lemma2 1); |
|
3457 | 355 |
by (assume_tac 1); |
3275 | 356 |
Addsimps [FilterConc]; |
4678 | 357 |
by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act]) 1); |
3275 | 358 |
(* apply IH *) |
3842 | 359 |
by (eres_inst_tac [("x","TL`(Dropwhile (%a. a:int A)`x)")] allE 1); |
4098 | 360 |
by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile])1); |
3275 | 361 |
by (REPEAT (etac exE 1)); |
362 |
by (REPEAT (etac conjE 1)); |
|
363 |
by (Asm_full_simp_tac 1); |
|
364 |
(* for replacing IH in conclusion *) |
|
365 |
by (rotate_tac ~2 1); |
|
366 |
by (Asm_full_simp_tac 1); |
|
367 |
(* instantiate y1a and y2a *) |
|
3842 | 368 |
by (res_inst_tac [("x","Takewhile (%a. a:int A)`x @@ a>>x1")] exI 1); |
3275 | 369 |
by (res_inst_tac [("x","x2")] exI 1); |
370 |
(* elminate all obligations up to two depending on Conc_assoc *) |
|
4098 | 371 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, intA_is_not_actB, |
4678 | 372 |
int_is_act,int_is_not_ext]) 1); |
4098 | 373 |
by (simp_tac (simpset() addsimps [Conc_assoc]) 1); |
4678 | 374 |
qed_spec_mp"reduceB_mksch1"; |
3275 | 375 |
|
4678 | 376 |
bind_thm("reduceB_mksch",conjI RSN (6,conjI RSN (5,reduceB_mksch1))); |
3071 | 377 |
|
378 |
||
3314 | 379 |
(* |
3071 | 380 |
|
381 |
||
6161 | 382 |
Goal "Finite y ==> ! z tr. Forall (%a.a:ext (A||B)) tr & \ |
3314 | 383 |
\ y = z @@ mksch A B`tr`a`b\ |
384 |
\ --> Finite tr"; |
|
3457 | 385 |
by (etac Seq_Finite_ind 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
386 |
by Auto_tac; |
3314 | 387 |
by (Seq_case_simp_tac "tr" 1); |
388 |
(* tr = UU *) |
|
4098 | 389 |
by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc]) 1); |
3314 | 390 |
(* tr = nil *) |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
391 |
by Auto_tac; |
3314 | 392 |
(* tr = Conc *) |
393 |
ren "s ss" 1; |
|
394 |
||
395 |
by (case_tac "s:act A" 1); |
|
396 |
by (case_tac "s:act B" 1); |
|
397 |
by (rotate_tac ~2 1); |
|
398 |
by (rotate_tac ~2 2); |
|
4098 | 399 |
by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc,nil_is_Conc2]) 1); |
400 |
by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc,nil_is_Conc2]) 1); |
|
3314 | 401 |
by (case_tac "s:act B" 1); |
402 |
by (rotate_tac ~2 1); |
|
4098 | 403 |
by (asm_full_simp_tac (simpset() addsimps [nil_is_Conc,nil_is_Conc2]) 1); |
404 |
by (fast_tac (claset() addSIs [ext_is_act] |
|
405 |
addss (simpset() addsimps [externals_of_par]) ) 1); |
|
3314 | 406 |
(* main case *) |
407 |
by (Seq_case_simp_tac "tr" 1); |
|
408 |
(* tr = UU *) |
|
4098 | 409 |
by (asm_full_simp_tac (simpset() addsimps [Conc_Conc_eq]) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
410 |
by Auto_tac; |
3314 | 411 |
(* tr = nil ok *) |
412 |
(* tr = Conc *) |
|
413 |
by (Seq_case_simp_tac "z" 1); |
|
414 |
(* z = UU ok *) |
|
415 |
(* z = nil *) |
|
416 |
||
417 |
(* z= Cons *) |
|
418 |
||
419 |
||
420 |
by (case_tac "aaa:act A" 2); |
|
421 |
by (case_tac "aaa:act B" 2); |
|
422 |
by (rotate_tac ~2 2); |
|
423 |
by (rotate_tac ~2 3); |
|
424 |
by (asm_full_simp_tac (HOL_basic_ss addsimps [mksch_cons3]) 2); |
|
3842 | 425 |
by (eres_inst_tac [("x","sb@@Takewhile (%a. a: int A)`a @@ Takewhile (%a. a:int B)`b@@(aaa>>nil)")] allE 2); |
3314 | 426 |
by (eres_inst_tac [("x","sa")] allE 2); |
4098 | 427 |
by (asm_full_simp_tac (simpset() addsimps [Conc_assoc])2); |
3314 | 428 |
|
429 |
||
430 |
||
431 |
by (eres_inst_tac [("x","sa")] allE 1); |
|
432 |
by (Asm_full_simp_tac 1); |
|
433 |
by (case_tac "aaa:act A" 1); |
|
434 |
by (case_tac "aaa:act B" 1); |
|
435 |
by (rotate_tac ~2 1); |
|
436 |
by (rotate_tac ~2 2); |
|
4098 | 437 |
by (asm_full_simp_tac (simpset() addsimps [Conc_Conc_eq]) 1); |
3314 | 438 |
|
439 |
||
440 |
||
5068 | 441 |
Goal "(x>>xs = y @@ z) = ((y=nil & x>>xs=z) | (? ys. y=x>>ys & xs=ys@@z))"; |
3314 | 442 |
by (Seq_case_simp_tac "y" 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
443 |
by Auto_tac; |
3314 | 444 |
qed"Conc_Conc_eq"; |
445 |
||
5068 | 446 |
Goal "!! (y::'a Seq).Finite y ==> ~ y= x@@UU"; |
3457 | 447 |
by (etac Seq_Finite_ind 1); |
3314 | 448 |
by (Seq_case_simp_tac "x" 1); |
449 |
by (Seq_case_simp_tac "x" 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
450 |
by Auto_tac; |
3314 | 451 |
qed"FiniteConcUU1"; |
452 |
||
5068 | 453 |
Goal "~ Finite ((x::'a Seq)@@UU)"; |
3457 | 454 |
by (rtac (FiniteConcUU1 COMP rev_contrapos) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4473
diff
changeset
|
455 |
by Auto_tac; |
3314 | 456 |
qed"FiniteConcUU"; |
457 |
||
458 |
finiteR_mksch |
|
459 |
"Finite (mksch A B`tr`x`y) --> Finite tr" |
|
460 |
*) |
|
461 |
||
3071 | 462 |
|
463 |
||
3275 | 464 |
(*------------------------------------------------------------------------------------- *) |
465 |
section"Filtering external actions out of mksch(tr,schA,schB) yields the oracle tr"; |
|
466 |
(* structural induction |
|
467 |
------------------------------------------------------------------------------------- *) |
|
3071 | 468 |
|
5068 | 469 |
Goal |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
470 |
"!! A B. [| compatible A B; compatible B A;\ |
3275 | 471 |
\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \ |
3842 | 472 |
\ ! schA schB. Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \ |
473 |
\ Forall (%x. x:ext (A||B)) tr & \ |
|
474 |
\ Filter (%a. a:act A)`tr << Filter (%a. a:ext A)`schA &\ |
|
475 |
\ Filter (%a. a:act B)`tr << Filter (%a. a:ext B)`schB \ |
|
476 |
\ --> Filter (%a. a:ext (A||B))`(mksch A B`tr`schA`schB) = tr"; |
|
3071 | 477 |
|
478 |
by (Seq_induct_tac "tr" [Forall_def,sforall_def,mksch_def] 1); |
|
479 |
||
480 |
(* main case *) |
|
481 |
(* splitting into 4 cases according to a:A, a:B *) |
|
4833 | 482 |
by (Asm_full_simp_tac 1); |
3071 | 483 |
by (safe_tac set_cs); |
484 |
||
485 |
(* Case a:A, a:B *) |
|
7499 | 486 |
by (ftac divide_Seq 1); |
487 |
by (ftac divide_Seq 1); |
|
3071 | 488 |
back(); |
489 |
by (REPEAT (etac conjE 1)); |
|
3275 | 490 |
(* filtering internals of A in schA and of B in schB is nil *) |
4098 | 491 |
by (asm_full_simp_tac (simpset() addsimps |
4678 | 492 |
[not_ext_is_int_or_not_act, |
3275 | 493 |
externals_of_par, intA_is_not_extB,int_is_not_ext]) 1); |
3071 | 494 |
(* conclusion of IH ok, but assumptions of IH have to be transformed *) |
495 |
by (dres_inst_tac [("x","schA")] subst_lemma1 1); |
|
3457 | 496 |
by (assume_tac 1); |
3071 | 497 |
by (dres_inst_tac [("x","schB")] subst_lemma1 1); |
3457 | 498 |
by (assume_tac 1); |
3275 | 499 |
(* IH *) |
4098 | 500 |
by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act, |
4678 | 501 |
ForallTL,ForallDropwhile]) 1); |
4520 | 502 |
|
503 |
(* Case a:A, a~:B *) |
|
7499 | 504 |
by (ftac divide_Seq 1); |
4520 | 505 |
by (REPEAT (etac conjE 1)); |
506 |
(* filtering internals of A is nil *) |
|
507 |
by (asm_full_simp_tac (simpset() addsimps |
|
4678 | 508 |
[not_ext_is_int_or_not_act, |
4520 | 509 |
externals_of_par, intA_is_not_extB,int_is_not_ext]) 1); |
510 |
by (dres_inst_tac [("x","schA")] subst_lemma1 1); |
|
511 |
by (assume_tac 1); |
|
512 |
by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act, |
|
4678 | 513 |
ForallTL,ForallDropwhile]) 1); |
4520 | 514 |
|
3071 | 515 |
(* Case a:B, a~:A *) |
7499 | 516 |
by (ftac divide_Seq 1); |
3071 | 517 |
by (REPEAT (etac conjE 1)); |
518 |
(* filtering internals of A is nil *) |
|
4098 | 519 |
by (asm_full_simp_tac (simpset() addsimps |
4678 | 520 |
[not_ext_is_int_or_not_act, |
3275 | 521 |
externals_of_par, intA_is_not_extB,int_is_not_ext]) 1); |
3071 | 522 |
by (dres_inst_tac [("x","schB")] subst_lemma1 1); |
523 |
back(); |
|
3457 | 524 |
by (assume_tac 1); |
4098 | 525 |
by (asm_full_simp_tac (simpset() addsimps [not_ext_is_int_or_not_act, |
4678 | 526 |
ForallTL,ForallDropwhile]) 1); |
3071 | 527 |
|
528 |
(* Case a~:A, a~:B *) |
|
4098 | 529 |
by (fast_tac (claset() addSIs [ext_is_act] |
530 |
addss (simpset() addsimps [externals_of_par]) ) 1); |
|
3275 | 531 |
qed"FilterA_mksch_is_tr"; |
3071 | 532 |
|
533 |
||
534 |
||
3314 | 535 |
(* |
536 |
||
537 |
***************************************************************8 |
|
538 |
With uncomplete take lemma rule should be reused afterwards !!!!!!!!!!!!!!!!! |
|
539 |
********************************************************************** |
|
540 |
||
541 |
(*--------------------------------------------------------------------------- |
|
542 |
Filter of mksch(tr,schA,schB) to A is schA |
|
543 |
take lemma |
|
544 |
--------------------------------------------------------------------------- *) |
|
545 |
||
5068 | 546 |
Goal "!! A B. [| compatible A B; compatible B A; \ |
3314 | 547 |
\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \ |
3842 | 548 |
\ Forall (%x. x:ext (A||B)) tr & \ |
549 |
\ Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \ |
|
550 |
\ Filter (%a. a:ext A)`schA = Filter (%a. a:act A)`tr &\ |
|
551 |
\ Filter (%a. a:ext B)`schB = Filter (%a. a:act B)`tr &\ |
|
3314 | 552 |
\ LastActExtsch A schA & LastActExtsch B schB \ |
3842 | 553 |
\ --> Filter (%a. a:act A)`(mksch A B`tr`schA`schB) = schA"; |
3314 | 554 |
|
555 |
by (res_inst_tac [("Q","%x. x:act B & x~:act A"),("x","tr")] take_lemma_less_induct 1); |
|
556 |
by (REPEAT (etac conjE 1)); |
|
557 |
||
558 |
by (case_tac "Finite s" 1); |
|
559 |
||
560 |
(* both sides of this equation are nil *) |
|
561 |
by (subgoal_tac "schA=nil" 1); |
|
562 |
by (Asm_simp_tac 1); |
|
563 |
(* first side: mksch = nil *) |
|
4098 | 564 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil,ForallBnAmksch,FiniteL_mksch], |
565 |
simpset())) 1); |
|
3314 | 566 |
(* second side: schA = nil *) |
567 |
by (eres_inst_tac [("A","A")] LastActExtimplnil 1); |
|
568 |
by (Asm_simp_tac 1); |
|
4098 | 569 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil], |
570 |
simpset())) 1); |
|
3314 | 571 |
|
572 |
(* case ~ Finite s *) |
|
573 |
||
574 |
(* both sides of this equation are UU *) |
|
575 |
by (subgoal_tac "schA=UU" 1); |
|
576 |
by (Asm_simp_tac 1); |
|
577 |
(* first side: mksch = UU *) |
|
4098 | 578 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU, |
3314 | 579 |
(finiteR_mksch RS mp COMP rev_contrapos), |
580 |
ForallBnAmksch], |
|
4098 | 581 |
simpset())) 1); |
3314 | 582 |
(* schA = UU *) |
583 |
by (eres_inst_tac [("A","A")] LastActExtimplUU 1); |
|
584 |
by (Asm_simp_tac 1); |
|
4098 | 585 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU], |
586 |
simpset())) 1); |
|
3314 | 587 |
|
588 |
(* case" ~ Forall (%x.x:act B & x~:act A) s" *) |
|
589 |
||
590 |
by (REPEAT (etac conjE 1)); |
|
591 |
||
592 |
(* bring in lemma reduceA_mksch *) |
|
593 |
by (forw_inst_tac [("y","schB"),("x","schA")] reduceA_mksch 1); |
|
594 |
by (REPEAT (atac 1)); |
|
595 |
by (REPEAT (etac exE 1)); |
|
596 |
by (REPEAT (etac conjE 1)); |
|
597 |
||
598 |
(* use reduceA_mksch to rewrite conclusion *) |
|
599 |
by (hyp_subst_tac 1); |
|
600 |
by (Asm_full_simp_tac 1); |
|
601 |
||
602 |
(* eliminate the B-only prefix *) |
|
603 |
||
3842 | 604 |
by (subgoal_tac "(Filter (%a. a :act A)`y1) = nil" 1); |
3457 | 605 |
by (etac ForallQFilterPnil 2); |
606 |
by (assume_tac 2); |
|
3314 | 607 |
by (Fast_tac 2); |
608 |
||
609 |
(* Now real recursive step follows (in y) *) |
|
610 |
||
611 |
by (Asm_full_simp_tac 1); |
|
612 |
by (case_tac "y:act A" 1); |
|
613 |
by (case_tac "y~:act B" 1); |
|
614 |
by (rotate_tac ~2 1); |
|
615 |
by (Asm_full_simp_tac 1); |
|
616 |
||
617 |
by (subgoal_tac "Filter (%a. a:act A & a:ext B)`y1=nil" 1); |
|
618 |
by (rotate_tac ~1 1); |
|
619 |
by (Asm_full_simp_tac 1); |
|
620 |
(* eliminate introduced subgoal 2 *) |
|
3457 | 621 |
by (etac ForallQFilterPnil 2); |
622 |
by (assume_tac 2); |
|
3314 | 623 |
by (Fast_tac 2); |
624 |
||
625 |
(* bring in divide Seq for s *) |
|
626 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
627 |
by (REPEAT (etac conjE 1)); |
|
628 |
||
629 |
(* subst divide_Seq in conclusion, but only at the righest occurence *) |
|
630 |
by (res_inst_tac [("t","schA")] ssubst 1); |
|
631 |
back(); |
|
632 |
back(); |
|
633 |
back(); |
|
3457 | 634 |
by (assume_tac 1); |
3314 | 635 |
|
636 |
(* reduce trace_takes from n to strictly smaller k *) |
|
3457 | 637 |
by (rtac take_reduction 1); |
3314 | 638 |
|
639 |
(* f A (tw iA) = tw ~eA *) |
|
4098 | 640 |
by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act, |
3314 | 641 |
not_ext_is_int_or_not_act]) 1); |
3457 | 642 |
by (rtac refl 1); |
3314 | 643 |
|
644 |
(* now conclusion fulfills induction hypothesis, but assumptions are not ready *) |
|
645 |
(* |
|
646 |
||
647 |
nacessary anymore ???????????????? |
|
648 |
by (rotate_tac ~10 1); |
|
649 |
||
650 |
*) |
|
651 |
(* assumption schB *) |
|
4098 | 652 |
by (asm_full_simp_tac (simpset() addsimps [ext_and_act]) 1); |
3314 | 653 |
(* assumption schA *) |
654 |
by (dres_inst_tac [("x","schA"), |
|
655 |
("g","Filter (%a. a:act A)`s2")] subst_lemma2 1); |
|
3457 | 656 |
by (assume_tac 1); |
4678 | 657 |
by (Asm_full_simp_tac 1); |
3314 | 658 |
(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *) |
659 |
by (dres_inst_tac [("sch","schA"),("P","%a. a:int A")] LastActExtsmall1 1); |
|
660 |
by (dres_inst_tac [("sch1.0","y1")] LastActExtsmall2 1); |
|
3457 | 661 |
by (assume_tac 1); |
3314 | 662 |
|
663 |
FIX: problem: schA and schB are not quantified in the new take lemma version !!! |
|
664 |
||
665 |
by (Asm_full_simp_tac 1); |
|
666 |
||
667 |
********************************************************************************************** |
|
668 |
*) |
|
669 |
||
670 |
||
671 |
||
3275 | 672 |
(*--------------------------------------------------------------------------- *) |
3071 | 673 |
|
3275 | 674 |
section" Filter of mksch(tr,schA,schB) to A is schA -- take lemma proof"; |
675 |
||
676 |
(* --------------------------------------------------------------------------- *) |
|
3071 | 677 |
|
678 |
||
679 |
||
5068 | 680 |
Goal "!! A B. [| compatible A B; compatible B A; \ |
3275 | 681 |
\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \ |
3842 | 682 |
\ Forall (%x. x:ext (A||B)) tr & \ |
683 |
\ Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \ |
|
684 |
\ Filter (%a. a:ext A)`schA = Filter (%a. a:act A)`tr &\ |
|
685 |
\ Filter (%a. a:ext B)`schB = Filter (%a. a:act B)`tr &\ |
|
3275 | 686 |
\ LastActExtsch A schA & LastActExtsch B schB \ |
3842 | 687 |
\ --> Filter (%a. a:act A)`(mksch A B`tr`schA`schB) = schA"; |
3071 | 688 |
|
3275 | 689 |
by (strip_tac 1); |
4042 | 690 |
by (resolve_tac seq.take_lemmas 1); |
3457 | 691 |
by (rtac mp 1); |
692 |
by (assume_tac 2); |
|
3275 | 693 |
back();back();back();back(); |
694 |
by (res_inst_tac [("x","schA")] spec 1); |
|
695 |
by (res_inst_tac [("x","schB")] spec 1); |
|
696 |
by (res_inst_tac [("x","tr")] spec 1); |
|
697 |
by (thin_tac' 5 1); |
|
3457 | 698 |
by (rtac less_induct 1); |
3275 | 699 |
by (REPEAT (rtac allI 1)); |
700 |
ren "tr schB schA" 1; |
|
701 |
by (strip_tac 1); |
|
702 |
by (REPEAT (etac conjE 1)); |
|
3071 | 703 |
|
3275 | 704 |
by (case_tac "Forall (%x. x:act B & x~:act A) tr" 1); |
705 |
||
3457 | 706 |
by (rtac (seq_take_lemma RS iffD2 RS spec) 1); |
3275 | 707 |
by (thin_tac' 5 1); |
3071 | 708 |
|
709 |
||
3275 | 710 |
by (case_tac "Finite tr" 1); |
3071 | 711 |
|
3275 | 712 |
(* both sides of this equation are nil *) |
713 |
by (subgoal_tac "schA=nil" 1); |
|
3071 | 714 |
by (Asm_simp_tac 1); |
3275 | 715 |
(* first side: mksch = nil *) |
4678 | 716 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil,ForallBnAmksch,FiniteL_mksch], |
4098 | 717 |
simpset())) 1); |
3275 | 718 |
(* second side: schA = nil *) |
3071 | 719 |
by (eres_inst_tac [("A","A")] LastActExtimplnil 1); |
720 |
by (Asm_simp_tac 1); |
|
3842 | 721 |
by (eres_inst_tac [("Q","%x. x:act B & x~:act A")] ForallQFilterPnil 1); |
3457 | 722 |
by (assume_tac 1); |
3314 | 723 |
by (Fast_tac 1); |
3275 | 724 |
|
725 |
(* case ~ Finite s *) |
|
3071 | 726 |
|
3275 | 727 |
(* both sides of this equation are UU *) |
728 |
by (subgoal_tac "schA=UU" 1); |
|
729 |
by (Asm_simp_tac 1); |
|
730 |
(* first side: mksch = UU *) |
|
4098 | 731 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU, |
4813 | 732 |
(finiteR_mksch RS mp COMP rev_contrapos),ForallBnAmksch],simpset())) 1); |
3071 | 733 |
(* schA = UU *) |
734 |
by (eres_inst_tac [("A","A")] LastActExtimplUU 1); |
|
735 |
by (Asm_simp_tac 1); |
|
3842 | 736 |
by (eres_inst_tac [("Q","%x. x:act B & x~:act A")] ForallQFilterPUU 1); |
3457 | 737 |
by (assume_tac 1); |
3314 | 738 |
by (Fast_tac 1); |
3275 | 739 |
|
740 |
(* case" ~ Forall (%x.x:act B & x~:act A) s" *) |
|
741 |
||
3457 | 742 |
by (dtac divide_Seq3 1); |
3071 | 743 |
|
3275 | 744 |
by (REPEAT (etac exE 1)); |
745 |
by (REPEAT (etac conjE 1)); |
|
746 |
by (hyp_subst_tac 1); |
|
3071 | 747 |
|
3275 | 748 |
(* bring in lemma reduceA_mksch *) |
4678 | 749 |
by (forw_inst_tac [("x","schA"),("y","schB"),("A","A"),("B","B")] reduceA_mksch 1); |
3071 | 750 |
by (REPEAT (atac 1)); |
751 |
by (REPEAT (etac exE 1)); |
|
752 |
by (REPEAT (etac conjE 1)); |
|
753 |
||
3275 | 754 |
(* use reduceA_mksch to rewrite conclusion *) |
3071 | 755 |
by (hyp_subst_tac 1); |
756 |
by (Asm_full_simp_tac 1); |
|
757 |
||
758 |
(* eliminate the B-only prefix *) |
|
3275 | 759 |
|
3842 | 760 |
by (subgoal_tac "(Filter (%a. a :act A)`y1) = nil" 1); |
3457 | 761 |
by (etac ForallQFilterPnil 2); |
762 |
by (assume_tac 2); |
|
3275 | 763 |
by (Fast_tac 2); |
764 |
||
765 |
(* Now real recursive step follows (in y) *) |
|
3071 | 766 |
|
3275 | 767 |
by (Asm_full_simp_tac 1); |
768 |
by (case_tac "x:act A" 1); |
|
769 |
by (case_tac "x~:act B" 1); |
|
3071 | 770 |
by (rotate_tac ~2 1); |
3275 | 771 |
by (Asm_full_simp_tac 1); |
3071 | 772 |
|
3275 | 773 |
by (subgoal_tac "Filter (%a. a:act A & a:ext B)`y1=nil" 1); |
3071 | 774 |
by (rotate_tac ~1 1); |
775 |
by (Asm_full_simp_tac 1); |
|
776 |
(* eliminate introduced subgoal 2 *) |
|
3457 | 777 |
by (etac ForallQFilterPnil 2); |
778 |
by (assume_tac 2); |
|
3071 | 779 |
by (Fast_tac 2); |
780 |
||
3275 | 781 |
(* bring in divide Seq for s *) |
782 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
783 |
by (REPEAT (etac conjE 1)); |
|
784 |
||
785 |
(* subst divide_Seq in conclusion, but only at the righest occurence *) |
|
786 |
by (res_inst_tac [("t","schA")] ssubst 1); |
|
787 |
back(); |
|
788 |
back(); |
|
789 |
back(); |
|
3457 | 790 |
by (assume_tac 1); |
3275 | 791 |
|
792 |
(* reduce trace_takes from n to strictly smaller k *) |
|
3457 | 793 |
by (rtac take_reduction 1); |
3275 | 794 |
|
795 |
(* f A (tw iA) = tw ~eA *) |
|
4098 | 796 |
by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act, |
3275 | 797 |
not_ext_is_int_or_not_act]) 1); |
3457 | 798 |
by (rtac refl 1); |
4098 | 799 |
by (asm_full_simp_tac (simpset() addsimps [int_is_act, |
3275 | 800 |
not_ext_is_int_or_not_act]) 1); |
801 |
by (rotate_tac ~11 1); |
|
802 |
||
803 |
(* now conclusion fulfills induction hypothesis, but assumptions are not ready *) |
|
3071 | 804 |
|
3275 | 805 |
(* assumption Forall tr *) |
4098 | 806 |
by (asm_full_simp_tac (simpset() addsimps [Forall_Conc]) 1); |
3275 | 807 |
(* assumption schB *) |
4098 | 808 |
by (asm_full_simp_tac (simpset() addsimps [Forall_Conc,ext_and_act]) 1); |
3275 | 809 |
by (REPEAT (etac conjE 1)); |
810 |
(* assumption schA *) |
|
811 |
by (dres_inst_tac [("x","schA"), |
|
812 |
("g","Filter (%a. a:act A)`rs")] subst_lemma2 1); |
|
3457 | 813 |
by (assume_tac 1); |
4678 | 814 |
by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1); |
3275 | 815 |
(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *) |
816 |
by (dres_inst_tac [("sch","schA"),("P","%a. a:int A")] LastActExtsmall1 1); |
|
817 |
by (forw_inst_tac [("sch1.0","y1")] LastActExtsmall2 1); |
|
3457 | 818 |
by (assume_tac 1); |
3275 | 819 |
|
820 |
(* assumption Forall schA *) |
|
821 |
by (dres_inst_tac [("s","schA"), |
|
3842 | 822 |
("P","Forall (%x. x:act A)")] subst 1); |
3457 | 823 |
by (assume_tac 1); |
4098 | 824 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, int_is_act]) 1); |
3275 | 825 |
|
826 |
(* case x:actions(asig_of A) & x: actions(asig_of B) *) |
|
827 |
||
828 |
||
829 |
by (rotate_tac ~2 1); |
|
830 |
by (Asm_full_simp_tac 1); |
|
831 |
||
832 |
by (subgoal_tac "Filter (%a. a:act A & a:ext B)`y1=nil" 1); |
|
833 |
by (rotate_tac ~1 1); |
|
834 |
by (Asm_full_simp_tac 1); |
|
835 |
(* eliminate introduced subgoal 2 *) |
|
3457 | 836 |
by (etac ForallQFilterPnil 2); |
837 |
by (assume_tac 2); |
|
3275 | 838 |
by (Fast_tac 2); |
4678 | 839 |
|
3275 | 840 |
(* bring in divide Seq for s *) |
841 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
842 |
by (REPEAT (etac conjE 1)); |
|
843 |
||
844 |
(* subst divide_Seq in conclusion, but only at the righest occurence *) |
|
3071 | 845 |
by (res_inst_tac [("t","schA")] ssubst 1); |
846 |
back(); |
|
847 |
back(); |
|
848 |
back(); |
|
3457 | 849 |
by (assume_tac 1); |
3275 | 850 |
|
851 |
(* f A (tw iA) = tw ~eA *) |
|
4098 | 852 |
by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act, |
3275 | 853 |
not_ext_is_int_or_not_act]) 1); |
3071 | 854 |
|
3275 | 855 |
(* rewrite assumption forall and schB *) |
856 |
by (rotate_tac 13 1); |
|
4098 | 857 |
by (asm_full_simp_tac (simpset() addsimps [ext_and_act]) 1); |
3275 | 858 |
|
859 |
(* divide_Seq for schB2 *) |
|
860 |
by (forw_inst_tac [("y","y2")] (sym RS antisym_less_inverse RS conjunct1 RS divide_Seq) 1); |
|
3071 | 861 |
by (REPEAT (etac conjE 1)); |
862 |
(* assumption schA *) |
|
863 |
by (dres_inst_tac [("x","schA"), |
|
3275 | 864 |
("g","Filter (%a. a:act A)`rs")] subst_lemma2 1); |
3457 | 865 |
by (assume_tac 1); |
4678 | 866 |
by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1); |
3275 | 867 |
|
868 |
(* f A (tw iB schB2) = nil *) |
|
4098 | 869 |
by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext,not_ext_is_int_or_not_act, |
4678 | 870 |
intA_is_not_actB]) 1); |
3275 | 871 |
|
872 |
||
873 |
(* reduce trace_takes from n to strictly smaller k *) |
|
3457 | 874 |
by (rtac take_reduction 1); |
875 |
by (rtac refl 1); |
|
876 |
by (rtac refl 1); |
|
3275 | 877 |
|
878 |
(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *) |
|
879 |
||
880 |
(* assumption schB *) |
|
881 |
by (dres_inst_tac [("x","y2"), |
|
882 |
("g","Filter (%a. a:act B)`rs")] subst_lemma2 1); |
|
3457 | 883 |
by (assume_tac 1); |
4678 | 884 |
by (asm_full_simp_tac (simpset() addsimps [intA_is_not_actB,int_is_not_ext]) 1); |
3275 | 885 |
|
886 |
(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *) |
|
887 |
by (dres_inst_tac [("sch","schA"),("P","%a. a:int A")] LastActExtsmall1 1); |
|
888 |
by (forw_inst_tac [("sch1.0","y1")] LastActExtsmall2 1); |
|
3457 | 889 |
by (assume_tac 1); |
3275 | 890 |
by (dres_inst_tac [("sch","y2"),("P","%a. a:int B")] LastActExtsmall1 1); |
891 |
||
892 |
(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *) |
|
4098 | 893 |
by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile]) 1); |
3275 | 894 |
|
895 |
(* case x~:A & x:B *) |
|
896 |
(* cannot occur, as just this case has been scheduled out before as the B-only prefix *) |
|
897 |
by (case_tac "x:act B" 1); |
|
898 |
by (Fast_tac 1); |
|
899 |
||
900 |
(* case x~:A & x~:B *) |
|
901 |
(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *) |
|
902 |
by (rotate_tac ~9 1); |
|
903 |
(* reduce forall assumption from tr to (x>>rs) *) |
|
4098 | 904 |
by (asm_full_simp_tac (simpset() addsimps [externals_of_par]) 1); |
3071 | 905 |
by (REPEAT (etac conjE 1)); |
4098 | 906 |
by (fast_tac (claset() addSIs [ext_is_act]) 1); |
3275 | 907 |
|
908 |
qed"FilterAmksch_is_schA"; |
|
909 |
||
910 |
||
911 |
||
912 |
(*--------------------------------------------------------------------------- *) |
|
913 |
||
914 |
section" Filter of mksch(tr,schA,schB) to B is schB -- take lemma proof"; |
|
915 |
||
916 |
(* --------------------------------------------------------------------------- *) |
|
917 |
||
918 |
||
919 |
||
5068 | 920 |
Goal "!! A B. [| compatible A B; compatible B A; \ |
3275 | 921 |
\ is_asig(asig_of A); is_asig(asig_of B) |] ==> \ |
3842 | 922 |
\ Forall (%x. x:ext (A||B)) tr & \ |
923 |
\ Forall (%x. x:act A) schA & Forall (%x. x:act B) schB & \ |
|
924 |
\ Filter (%a. a:ext A)`schA = Filter (%a. a:act A)`tr &\ |
|
925 |
\ Filter (%a. a:ext B)`schB = Filter (%a. a:act B)`tr &\ |
|
3275 | 926 |
\ LastActExtsch A schA & LastActExtsch B schB \ |
3842 | 927 |
\ --> Filter (%a. a:act B)`(mksch A B`tr`schA`schB) = schB"; |
3275 | 928 |
|
929 |
by (strip_tac 1); |
|
4042 | 930 |
by (resolve_tac seq.take_lemmas 1); |
3457 | 931 |
by (rtac mp 1); |
932 |
by (assume_tac 2); |
|
3275 | 933 |
back();back();back();back(); |
934 |
by (res_inst_tac [("x","schA")] spec 1); |
|
935 |
by (res_inst_tac [("x","schB")] spec 1); |
|
936 |
by (res_inst_tac [("x","tr")] spec 1); |
|
937 |
by (thin_tac' 5 1); |
|
3457 | 938 |
by (rtac less_induct 1); |
3275 | 939 |
by (REPEAT (rtac allI 1)); |
940 |
ren "tr schB schA" 1; |
|
941 |
by (strip_tac 1); |
|
942 |
by (REPEAT (etac conjE 1)); |
|
943 |
||
944 |
by (case_tac "Forall (%x. x:act A & x~:act B) tr" 1); |
|
945 |
||
3457 | 946 |
by (rtac (seq_take_lemma RS iffD2 RS spec) 1); |
3275 | 947 |
by (thin_tac' 5 1); |
948 |
||
949 |
||
950 |
by (case_tac "Finite tr" 1); |
|
951 |
||
952 |
(* both sides of this equation are nil *) |
|
953 |
by (subgoal_tac "schB=nil" 1); |
|
954 |
by (Asm_simp_tac 1); |
|
955 |
(* first side: mksch = nil *) |
|
4678 | 956 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPnil,ForallAnBmksch,FiniteL_mksch], |
4098 | 957 |
simpset())) 1); |
3275 | 958 |
(* second side: schA = nil *) |
959 |
by (eres_inst_tac [("A","B")] LastActExtimplnil 1); |
|
960 |
by (Asm_simp_tac 1); |
|
3842 | 961 |
by (eres_inst_tac [("Q","%x. x:act A & x~:act B")] ForallQFilterPnil 1); |
3457 | 962 |
by (assume_tac 1); |
3314 | 963 |
by (Fast_tac 1); |
3275 | 964 |
|
965 |
(* case ~ Finite tr *) |
|
966 |
||
967 |
(* both sides of this equation are UU *) |
|
968 |
by (subgoal_tac "schB=UU" 1); |
|
969 |
by (Asm_simp_tac 1); |
|
970 |
(* first side: mksch = UU *) |
|
4098 | 971 |
by (SELECT_GOAL (auto_tac (claset() addSIs [ForallQFilterPUU, |
3275 | 972 |
(finiteR_mksch RS mp COMP rev_contrapos), |
973 |
ForallAnBmksch], |
|
4098 | 974 |
simpset())) 1); |
3275 | 975 |
(* schA = UU *) |
976 |
by (eres_inst_tac [("A","B")] LastActExtimplUU 1); |
|
977 |
by (Asm_simp_tac 1); |
|
3842 | 978 |
by (eres_inst_tac [("Q","%x. x:act A & x~:act B")] ForallQFilterPUU 1); |
3457 | 979 |
by (assume_tac 1); |
3314 | 980 |
by (Fast_tac 1); |
3275 | 981 |
|
982 |
(* case" ~ Forall (%x.x:act B & x~:act A) s" *) |
|
983 |
||
3457 | 984 |
by (dtac divide_Seq3 1); |
3275 | 985 |
|
986 |
by (REPEAT (etac exE 1)); |
|
987 |
by (REPEAT (etac conjE 1)); |
|
988 |
by (hyp_subst_tac 1); |
|
989 |
||
990 |
(* bring in lemma reduceB_mksch *) |
|
991 |
by (forw_inst_tac [("y","schB"),("x","schA"),("A","A"),("B","B")] reduceB_mksch 1); |
|
992 |
by (REPEAT (atac 1)); |
|
993 |
by (REPEAT (etac exE 1)); |
|
994 |
by (REPEAT (etac conjE 1)); |
|
995 |
||
996 |
(* use reduceB_mksch to rewrite conclusion *) |
|
997 |
by (hyp_subst_tac 1); |
|
998 |
by (Asm_full_simp_tac 1); |
|
999 |
||
1000 |
(* eliminate the A-only prefix *) |
|
1001 |
||
3842 | 1002 |
by (subgoal_tac "(Filter (%a. a :act B)`x1) = nil" 1); |
3457 | 1003 |
by (etac ForallQFilterPnil 2); |
1004 |
by (assume_tac 2); |
|
3275 | 1005 |
by (Fast_tac 2); |
1006 |
||
1007 |
(* Now real recursive step follows (in x) *) |
|
1008 |
||
1009 |
by (Asm_full_simp_tac 1); |
|
1010 |
by (case_tac "x:act B" 1); |
|
1011 |
by (case_tac "x~:act A" 1); |
|
1012 |
by (rotate_tac ~2 1); |
|
3071 | 1013 |
by (Asm_full_simp_tac 1); |
1014 |
||
3275 | 1015 |
by (subgoal_tac "Filter (%a. a:act B & a:ext A)`x1=nil" 1); |
3071 | 1016 |
by (rotate_tac ~1 1); |
1017 |
by (Asm_full_simp_tac 1); |
|
1018 |
(* eliminate introduced subgoal 2 *) |
|
3457 | 1019 |
by (etac ForallQFilterPnil 2); |
1020 |
by (assume_tac 2); |
|
3071 | 1021 |
by (Fast_tac 2); |
1022 |
||
3275 | 1023 |
(* bring in divide Seq for s *) |
1024 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
1025 |
by (REPEAT (etac conjE 1)); |
|
3071 | 1026 |
|
3275 | 1027 |
(* subst divide_Seq in conclusion, but only at the righest occurence *) |
1028 |
by (res_inst_tac [("t","schB")] ssubst 1); |
|
3071 | 1029 |
back(); |
1030 |
back(); |
|
1031 |
back(); |
|
3457 | 1032 |
by (assume_tac 1); |
3275 | 1033 |
|
1034 |
(* reduce trace_takes from n to strictly smaller k *) |
|
3457 | 1035 |
by (rtac take_reduction 1); |
3275 | 1036 |
|
1037 |
(* f B (tw iB) = tw ~eB *) |
|
4098 | 1038 |
by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act, |
3275 | 1039 |
not_ext_is_int_or_not_act]) 1); |
3457 | 1040 |
by (rtac refl 1); |
4098 | 1041 |
by (asm_full_simp_tac (simpset() addsimps [int_is_act, |
3275 | 1042 |
not_ext_is_int_or_not_act]) 1); |
1043 |
by (rotate_tac ~11 1); |
|
1044 |
||
1045 |
(* now conclusion fulfills induction hypothesis, but assumptions are not ready *) |
|
1046 |
||
1047 |
(* assumption Forall tr *) |
|
4098 | 1048 |
by (asm_full_simp_tac (simpset() addsimps [Forall_Conc]) 1); |
3275 | 1049 |
(* assumption schA *) |
4098 | 1050 |
by (asm_full_simp_tac (simpset() addsimps [Forall_Conc,ext_and_act]) 1); |
3071 | 1051 |
by (REPEAT (etac conjE 1)); |
3275 | 1052 |
(* assumption schB *) |
1053 |
by (dres_inst_tac [("x","schB"), |
|
1054 |
("g","Filter (%a. a:act B)`rs")] subst_lemma2 1); |
|
3457 | 1055 |
by (assume_tac 1); |
4678 | 1056 |
by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1); |
3275 | 1057 |
(* assumptions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *) |
1058 |
by (dres_inst_tac [("sch","schB"),("P","%a. a:int B")] LastActExtsmall1 1); |
|
1059 |
by (forw_inst_tac [("sch1.0","x1")] LastActExtsmall2 1); |
|
3457 | 1060 |
by (assume_tac 1); |
3275 | 1061 |
|
1062 |
(* assumption Forall schB *) |
|
1063 |
by (dres_inst_tac [("s","schB"), |
|
3842 | 1064 |
("P","Forall (%x. x:act B)")] subst 1); |
3457 | 1065 |
by (assume_tac 1); |
4098 | 1066 |
by (asm_full_simp_tac (simpset() addsimps [ForallPTakewhileQ, int_is_act]) 1); |
3275 | 1067 |
|
1068 |
(* case x:actions(asig_of A) & x: actions(asig_of B) *) |
|
1069 |
||
1070 |
||
1071 |
by (rotate_tac ~2 1); |
|
1072 |
by (Asm_full_simp_tac 1); |
|
1073 |
||
1074 |
by (subgoal_tac "Filter (%a. a:act B & a:ext A)`x1=nil" 1); |
|
1075 |
by (rotate_tac ~1 1); |
|
1076 |
by (Asm_full_simp_tac 1); |
|
1077 |
(* eliminate introduced subgoal 2 *) |
|
3457 | 1078 |
by (etac ForallQFilterPnil 2); |
1079 |
by (assume_tac 2); |
|
3275 | 1080 |
by (Fast_tac 2); |
1081 |
||
1082 |
(* bring in divide Seq for s *) |
|
1083 |
by (forward_tac [sym RS antisym_less_inverse RS conjunct1 RS divide_Seq] 1); |
|
3071 | 1084 |
by (REPEAT (etac conjE 1)); |
1085 |
||
3275 | 1086 |
(* subst divide_Seq in conclusion, but only at the righest occurence *) |
1087 |
by (res_inst_tac [("t","schB")] ssubst 1); |
|
1088 |
back(); |
|
1089 |
back(); |
|
1090 |
back(); |
|
3457 | 1091 |
by (assume_tac 1); |
3275 | 1092 |
|
1093 |
(* f B (tw iB) = tw ~eB *) |
|
4098 | 1094 |
by (asm_full_simp_tac (simpset() addsimps [FilterPTakewhileQid,int_is_act, |
3275 | 1095 |
not_ext_is_int_or_not_act]) 1); |
1096 |
||
1097 |
(* rewrite assumption forall and schB *) |
|
1098 |
by (rotate_tac 13 1); |
|
4098 | 1099 |
by (asm_full_simp_tac (simpset() addsimps [ext_and_act]) 1); |
3275 | 1100 |
|
1101 |
(* divide_Seq for schB2 *) |
|
1102 |
by (forw_inst_tac [("y","x2")] (sym RS antisym_less_inverse RS conjunct1 RS divide_Seq) 1); |
|
1103 |
by (REPEAT (etac conjE 1)); |
|
1104 |
(* assumption schA *) |
|
1105 |
by (dres_inst_tac [("x","schB"), |
|
1106 |
("g","Filter (%a. a:act B)`rs")] subst_lemma2 1); |
|
3457 | 1107 |
by (assume_tac 1); |
4678 | 1108 |
by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext]) 1); |
3275 | 1109 |
|
1110 |
(* f B (tw iA schA2) = nil *) |
|
4098 | 1111 |
by (asm_full_simp_tac (simpset() addsimps [int_is_not_ext,not_ext_is_int_or_not_act, |
4678 | 1112 |
intA_is_not_actB]) 1); |
3275 | 1113 |
|
1114 |
||
1115 |
(* reduce trace_takes from n to strictly smaller k *) |
|
3457 | 1116 |
by (rtac take_reduction 1); |
1117 |
by (rtac refl 1); |
|
1118 |
by (rtac refl 1); |
|
3275 | 1119 |
|
1120 |
(* now conclusion fulfills induction hypothesis, but assumptions are not all ready *) |
|
1121 |
||
1122 |
(* assumption schA *) |
|
1123 |
by (dres_inst_tac [("x","x2"), |
|
1124 |
("g","Filter (%a. a:act A)`rs")] subst_lemma2 1); |
|
3457 | 1125 |
by (assume_tac 1); |
4678 | 1126 |
by (asm_full_simp_tac (simpset() addsimps [intA_is_not_actB,int_is_not_ext]) 1); |
3275 | 1127 |
|
1128 |
(* conclusions concerning LastActExtsch, cannot be rewritten, as LastActExtsmalli are looping *) |
|
1129 |
by (dres_inst_tac [("sch","schB"),("P","%a. a:int B")] LastActExtsmall1 1); |
|
1130 |
by (forw_inst_tac [("sch1.0","x1")] LastActExtsmall2 1); |
|
3457 | 1131 |
by (assume_tac 1); |
3275 | 1132 |
by (dres_inst_tac [("sch","x2"),("P","%a. a:int A")] LastActExtsmall1 1); |
1133 |
||
1134 |
(* assumption Forall schA, and Forall schA are performed by ForallTL,ForallDropwhile *) |
|
4098 | 1135 |
by (asm_full_simp_tac (simpset() addsimps [ForallTL,ForallDropwhile]) 1); |
3275 | 1136 |
|
1137 |
(* case x~:B & x:A *) |
|
3071 | 1138 |
(* cannot occur, as just this case has been scheduled out before as the B-only prefix *) |
3275 | 1139 |
by (case_tac "x:act A" 1); |
3071 | 1140 |
by (Fast_tac 1); |
1141 |
||
3275 | 1142 |
(* case x~:B & x~:A *) |
1143 |
(* cannot occur because of assumption: Forall (a:ext A | a:ext B) *) |
|
1144 |
by (rotate_tac ~9 1); |
|
1145 |
(* reduce forall assumption from tr to (x>>rs) *) |
|
4098 | 1146 |
by (asm_full_simp_tac (simpset() addsimps [externals_of_par]) 1); |
3071 | 1147 |
by (REPEAT (etac conjE 1)); |
4098 | 1148 |
by (fast_tac (claset() addSIs [ext_is_act]) 1); |
3071 | 1149 |
|
3275 | 1150 |
qed"FilterBmksch_is_schB"; |
3071 | 1151 |
|
1152 |
||
1153 |
||
1154 |
(* ------------------------------------------------------------------ *) |
|
3275 | 1155 |
section"COMPOSITIONALITY on TRACE Level -- Main Theorem"; |
3071 | 1156 |
(* ------------------------------------------------------------------ *) |
1157 |
||
5068 | 1158 |
Goal |
3521 | 1159 |
"!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A; \ |
3071 | 1160 |
\ is_asig(asig_of A); is_asig(asig_of B)|] \ |
3275 | 1161 |
\ ==> tr: traces(A||B) = \ |
3842 | 1162 |
\ (Filter (%a. a:act A)`tr : traces A &\ |
1163 |
\ Filter (%a. a:act B)`tr : traces B &\ |
|
3275 | 1164 |
\ Forall (%x. x:ext(A||B)) tr)"; |
3071 | 1165 |
|
4098 | 1166 |
by (simp_tac (simpset() addsimps [traces_def,has_trace_def]) 1); |
3071 | 1167 |
by (safe_tac set_cs); |
1168 |
||
1169 |
(* ==> *) |
|
1170 |
(* There is a schedule of A *) |
|
3842 | 1171 |
by (res_inst_tac [("x","Filter (%a. a:act A)`sch")] bexI 1); |
4098 | 1172 |
by (asm_full_simp_tac (simpset() addsimps [compositionality_sch]) 2); |
1173 |
by (asm_full_simp_tac (simpset() addsimps [compatibility_consequence1, |
|
3071 | 1174 |
externals_of_par,ext1_ext2_is_not_act1]) 1); |
1175 |
(* There is a schedule of B *) |
|
3842 | 1176 |
by (res_inst_tac [("x","Filter (%a. a:act B)`sch")] bexI 1); |
4098 | 1177 |
by (asm_full_simp_tac (simpset() addsimps [compositionality_sch]) 2); |
1178 |
by (asm_full_simp_tac (simpset() addsimps [compatibility_consequence2, |
|
3071 | 1179 |
externals_of_par,ext1_ext2_is_not_act2]) 1); |
1180 |
(* Traces of A||B have only external actions from A or B *) |
|
3457 | 1181 |
by (rtac ForallPFilterP 1); |
3071 | 1182 |
|
1183 |
(* <== *) |
|
1184 |
||
3275 | 1185 |
(* replace schA and schB by Cut(schA) and Cut(schB) *) |
3071 | 1186 |
by (dtac exists_LastActExtsch 1); |
3457 | 1187 |
by (assume_tac 1); |
3071 | 1188 |
by (dtac exists_LastActExtsch 1); |
3457 | 1189 |
by (assume_tac 1); |
3071 | 1190 |
by (REPEAT (etac exE 1)); |
1191 |
by (REPEAT (etac conjE 1)); |
|
3275 | 1192 |
(* Schedules of A(B) have only actions of A(B) *) |
3457 | 1193 |
by (dtac scheds_in_sig 1); |
1194 |
by (assume_tac 1); |
|
1195 |
by (dtac scheds_in_sig 1); |
|
1196 |
by (assume_tac 1); |
|
3071 | 1197 |
|
3275 | 1198 |
ren "h1 h2 schA schB" 1; |
3071 | 1199 |
(* mksch is exactly the construction of trA||B out of schA, schB, and the oracle tr, |
1200 |
we need here *) |
|
3275 | 1201 |
by (res_inst_tac [("x","mksch A B`tr`schA`schB")] bexI 1); |
3071 | 1202 |
|
1203 |
(* External actions of mksch are just the oracle *) |
|
4098 | 1204 |
by (asm_full_simp_tac (simpset() addsimps [FilterA_mksch_is_tr RS spec RS spec RS mp]) 1); |
3071 | 1205 |
|
1206 |
(* mksch is a schedule -- use compositionality on sch-level *) |
|
4098 | 1207 |
by (asm_full_simp_tac (simpset() addsimps [compositionality_sch]) 1); |
1208 |
by (asm_full_simp_tac (simpset() addsimps [FilterAmksch_is_schA,FilterBmksch_is_schB]) 1); |
|
3457 | 1209 |
by (etac ForallAorB_mksch 1); |
1210 |
by (etac ForallPForallQ 1); |
|
1211 |
by (etac ext_is_act 1); |
|
3275 | 1212 |
qed"compositionality_tr"; |
3071 | 1213 |
|
1214 |
||
1215 |
||
3521 | 1216 |
(* ------------------------------------------------------------------ *) |
1217 |
(* COMPOSITIONALITY on TRACE Level *) |
|
1218 |
(* For Modules *) |
|
1219 |
(* ------------------------------------------------------------------ *) |
|
1220 |
||
5068 | 1221 |
Goalw [Traces_def,par_traces_def] |
3521 | 1222 |
|
1223 |
"!! A B. [| is_trans_of A; is_trans_of B; compatible A B; compatible B A; \ |
|
1224 |
\ is_asig(asig_of A); is_asig(asig_of B)|] \ |
|
1225 |
\==> Traces (A||B) = par_traces (Traces A) (Traces B)"; |
|
1226 |
||
4098 | 1227 |
by (asm_full_simp_tac (simpset() addsimps [asig_of_par]) 1); |
4423 | 1228 |
by (rtac set_ext 1); |
4098 | 1229 |
by (asm_full_simp_tac (simpset() addsimps [compositionality_tr,externals_of_par]) 1); |
3521 | 1230 |
qed"compositionality_tr_modules"; |
3071 | 1231 |
|
1232 |
||
4678 | 1233 |
simpset_ref () := simpset() setmksym (Some o symmetric_fun); |