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