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