| author | wenzelm | 
| Mon, 20 Jun 2005 22:14:19 +0200 | |
| changeset 16504 | 7c1cb7ce24eb | 
| parent 14981 | e73f8140af78 | 
| child 17233 | 41eee2e7b465 | 
| permissions | -rw-r--r-- | 
| 3071 | 1  | 
(* Title: HOLCF/IOA/meta_theory/RefCorrectness.ML  | 
| 3275 | 2  | 
ID: $Id$  | 
| 12218 | 3  | 
Author: Olaf Müller  | 
| 3071 | 4  | 
|
| 12218 | 5  | 
Correctness of Refinement Mappings in HOLCF/IOA.  | 
| 3071 | 6  | 
*)  | 
7  | 
||
8  | 
||
9  | 
||
10  | 
(* -------------------------------------------------------------------------------- *)  | 
|
11  | 
||
12  | 
section "corresp_ex";  | 
|
13  | 
||
14  | 
||
15  | 
(* ---------------------------------------------------------------- *)  | 
|
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
16  | 
(* corresp_exC *)  | 
| 3071 | 17  | 
(* ---------------------------------------------------------------- *)  | 
18  | 
||
19  | 
||
| 5068 | 20  | 
Goal "corresp_exC A f = (LAM ex. (%s. case ex of \  | 
| 3071 | 21  | 
\ nil => nil \  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
22  | 
\ | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr))) \  | 
| 10835 | 23  | 
\ @@ ((corresp_exC A f $xs) (snd pr))) \  | 
24  | 
\ $x) ))";  | 
|
| 3071 | 25  | 
by (rtac trans 1);  | 
| 3457 | 26  | 
by (rtac fix_eq2 1);  | 
27  | 
by (rtac corresp_exC_def 1);  | 
|
28  | 
by (rtac beta_cfun 1);  | 
|
| 4098 | 29  | 
by (simp_tac (simpset() addsimps [flift1_def]) 1);  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
30  | 
qed"corresp_exC_unfold";  | 
| 3071 | 31  | 
|
| 10835 | 32  | 
Goal "(corresp_exC A f$UU) s=UU";  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
33  | 
by (stac corresp_exC_unfold 1);  | 
| 3071 | 34  | 
by (Simp_tac 1);  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
35  | 
qed"corresp_exC_UU";  | 
| 3071 | 36  | 
|
| 10835 | 37  | 
Goal "(corresp_exC A f$nil) s = nil";  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
38  | 
by (stac corresp_exC_unfold 1);  | 
| 3071 | 39  | 
by (Simp_tac 1);  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
40  | 
qed"corresp_exC_nil";  | 
| 3071 | 41  | 
|
| 10835 | 42  | 
Goal "(corresp_exC A f$(at>>xs)) s = \  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
43  | 
\ (@cex. move A cex (f s) (fst at) (f (snd at))) \  | 
| 10835 | 44  | 
\ @@ ((corresp_exC A f$xs) (snd at))";  | 
| 3457 | 45  | 
by (rtac trans 1);  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
46  | 
by (stac corresp_exC_unfold 1);  | 
| 
7229
 
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
 
wenzelm 
parents: 
6161 
diff
changeset
 | 
47  | 
by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1);  | 
| 3071 | 48  | 
by (Simp_tac 1);  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
49  | 
qed"corresp_exC_cons";  | 
| 3071 | 50  | 
|
51  | 
||
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
52  | 
Addsimps [corresp_exC_UU,corresp_exC_nil,corresp_exC_cons];  | 
| 3071 | 53  | 
|
54  | 
||
55  | 
||
56  | 
(* ------------------------------------------------------------------ *)  | 
|
57  | 
(* The following lemmata describe the definition *)  | 
|
58  | 
(* of move in more detail *)  | 
|
59  | 
(* ------------------------------------------------------------------ *)  | 
|
60  | 
||
61  | 
section"properties of move";  | 
|
62  | 
||
| 5068 | 63  | 
Goalw [is_ref_map_def]  | 
| 6161 | 64  | 
"[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\  | 
| 3071 | 65  | 
\ move A (@x. move A x (f s) a (f t)) (f s) a (f t)";  | 
66  | 
||
| 3847 | 67  | 
by (subgoal_tac "? ex. move A ex (f s) a (f t)" 1);  | 
| 3071 | 68  | 
by (Asm_full_simp_tac 2);  | 
69  | 
by (etac exE 1);  | 
|
| 9970 | 70  | 
by (rtac someI 1);  | 
| 3071 | 71  | 
by (assume_tac 1);  | 
72  | 
qed"move_is_move";  | 
|
73  | 
||
| 5068 | 74  | 
Goal  | 
| 6161 | 75  | 
"[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
76  | 
\ is_exec_frag A (f s,@x. move A x (f s) a (f t))";  | 
| 3071 | 77  | 
by (cut_inst_tac [] move_is_move 1);  | 
78  | 
by (REPEAT (assume_tac 1));  | 
|
| 4098 | 79  | 
by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);  | 
| 3071 | 80  | 
qed"move_subprop1";  | 
81  | 
||
| 5068 | 82  | 
Goal  | 
| 6161 | 83  | 
"[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
84  | 
\ Finite ((@x. move A x (f s) a (f t)))";  | 
| 3071 | 85  | 
by (cut_inst_tac [] move_is_move 1);  | 
86  | 
by (REPEAT (assume_tac 1));  | 
|
| 4098 | 87  | 
by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);  | 
| 3071 | 88  | 
qed"move_subprop2";  | 
89  | 
||
| 5068 | 90  | 
Goal  | 
| 6161 | 91  | 
"[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
92  | 
\ laststate (f s,@x. move A x (f s) a (f t)) = (f t)";  | 
| 3071 | 93  | 
by (cut_inst_tac [] move_is_move 1);  | 
94  | 
by (REPEAT (assume_tac 1));  | 
|
| 4098 | 95  | 
by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);  | 
| 3071 | 96  | 
qed"move_subprop3";  | 
97  | 
||
| 5068 | 98  | 
Goal  | 
| 6161 | 99  | 
"[|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\  | 
| 10835 | 100  | 
\ mk_trace A$((@x. move A x (f s) a (f t))) = \  | 
| 3071 | 101  | 
\ (if a:ext A then a>>nil else nil)";  | 
102  | 
||
103  | 
by (cut_inst_tac [] move_is_move 1);  | 
|
104  | 
by (REPEAT (assume_tac 1));  | 
|
| 4098 | 105  | 
by (asm_full_simp_tac (simpset() addsimps [move_def]) 1);  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
106  | 
qed"move_subprop4";  | 
| 3071 | 107  | 
|
108  | 
||
109  | 
(* ------------------------------------------------------------------ *)  | 
|
110  | 
(* The following lemmata contribute to *)  | 
|
111  | 
(* TRACE INCLUSION Part 1: Traces coincide *)  | 
|
112  | 
(* ------------------------------------------------------------------ *)  | 
|
113  | 
||
114  | 
section "Lemmata for <==";  | 
|
115  | 
||
116  | 
(* --------------------------------------------------- *)  | 
|
117  | 
(* Lemma 1.1: Distribution of mk_trace and @@ *)  | 
|
118  | 
(* --------------------------------------------------- *)  | 
|
119  | 
||
120  | 
||
| 10835 | 121  | 
Goal "mk_trace C$(ex1 @@ ex2)= (mk_trace C$ex1) @@ (mk_trace C$ex2)";  | 
| 4098 | 122  | 
by (simp_tac (simpset() addsimps [mk_trace_def,filter_act_def,  | 
| 3071 | 123  | 
FilterConc,MapConc]) 1);  | 
124  | 
qed"mk_traceConc";  | 
|
125  | 
||
126  | 
||
127  | 
||
128  | 
(* ------------------------------------------------------  | 
|
129  | 
Lemma 1 :Traces coincide  | 
|
130  | 
------------------------------------------------------- *)  | 
|
| 4833 | 131  | 
Delsplits[split_if];  | 
| 5068 | 132  | 
Goalw [corresp_ex_def]  | 
| 6161 | 133  | 
"[|is_ref_map f C A; ext C = ext A|] ==> \  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
134  | 
\ !s. reachable C s & is_exec_frag C (s,xs) --> \  | 
| 10835 | 135  | 
\ mk_trace C$xs = mk_trace A$(snd (corresp_ex A f (s,xs)))";  | 
| 3071 | 136  | 
|
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
137  | 
by (pair_induct_tac "xs" [is_exec_frag_def] 1);  | 
| 3071 | 138  | 
(* cons case *)  | 
139  | 
by (safe_tac set_cs);  | 
|
| 4098 | 140  | 
by (asm_full_simp_tac (simpset() addsimps [mk_traceConc]) 1);  | 
| 3071 | 141  | 
by (forward_tac [reachable.reachable_n] 1);  | 
| 3457 | 142  | 
by (assume_tac 1);  | 
| 3071 | 143  | 
by (eres_inst_tac [("x","y")] allE 1);
 | 
144  | 
by (Asm_full_simp_tac 1);  | 
|
| 4098 | 145  | 
by (asm_full_simp_tac (simpset() addsimps [move_subprop4]  | 
| 4833 | 146  | 
addsplits [split_if]) 1);  | 
| 3071 | 147  | 
qed"lemma_1";  | 
| 4833 | 148  | 
Addsplits[split_if];  | 
| 3071 | 149  | 
|
150  | 
(* ------------------------------------------------------------------ *)  | 
|
151  | 
(* The following lemmata contribute to *)  | 
|
152  | 
(* TRACE INCLUSION Part 2: corresp_ex is execution *)  | 
|
153  | 
(* ------------------------------------------------------------------ *)  | 
|
154  | 
||
155  | 
section "Lemmata for ==>";  | 
|
156  | 
||
157  | 
(* -------------------------------------------------- *)  | 
|
158  | 
(* Lemma 2.1 *)  | 
|
159  | 
(* -------------------------------------------------- *)  | 
|
160  | 
||
| 5068 | 161  | 
Goal  | 
| 3071 | 162  | 
"Finite xs --> \  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
163  | 
\(!s .is_exec_frag A (s,xs) & is_exec_frag A (t,ys) & \  | 
| 3071 | 164  | 
\ t = laststate (s,xs) \  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
165  | 
\ --> is_exec_frag A (s,xs @@ ys))";  | 
| 3071 | 166  | 
|
| 3457 | 167  | 
by (rtac impI 1);  | 
| 3071 | 168  | 
by (Seq_Finite_induct_tac 1);  | 
169  | 
(* main case *)  | 
|
170  | 
by (safe_tac set_cs);  | 
|
171  | 
by (pair_tac "a" 1);  | 
|
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
172  | 
qed_spec_mp"lemma_2_1";  | 
| 3071 | 173  | 
|
174  | 
||
175  | 
(* ----------------------------------------------------------- *)  | 
|
176  | 
(* Lemma 2 : corresp_ex is execution *)  | 
|
177  | 
(* ----------------------------------------------------------- *)  | 
|
178  | 
||
179  | 
||
180  | 
||
| 5068 | 181  | 
Goalw [corresp_ex_def]  | 
| 6161 | 182  | 
"[| is_ref_map f C A |] ==>\  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
183  | 
\ !s. reachable C s & is_exec_frag C (s,xs) \  | 
| 
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
184  | 
\ --> is_exec_frag A (corresp_ex A f (s,xs))";  | 
| 3071 | 185  | 
|
186  | 
by (Asm_full_simp_tac 1);  | 
|
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
187  | 
by (pair_induct_tac "xs" [is_exec_frag_def] 1);  | 
| 3071 | 188  | 
(* main case *)  | 
189  | 
by (safe_tac set_cs);  | 
|
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
190  | 
by (res_inst_tac [("t","f y")]  lemma_2_1 1);
 | 
| 3071 | 191  | 
|
192  | 
(* Finite *)  | 
|
| 3457 | 193  | 
by (etac move_subprop2 1);  | 
| 3071 | 194  | 
by (REPEAT (atac 1));  | 
195  | 
by (rtac conjI 1);  | 
|
196  | 
||
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents: 
3275 
diff
changeset
 | 
197  | 
(* is_exec_frag *)  | 
| 3457 | 198  | 
by (etac move_subprop1 1);  | 
| 3071 | 199  | 
by (REPEAT (atac 1));  | 
200  | 
by (rtac conjI 1);  | 
|
201  | 
||
202  | 
(* Induction hypothesis *)  | 
|
203  | 
(* reachable_n looping, therefore apply it manually *)  | 
|
204  | 
by (eres_inst_tac [("x","y")] allE 1);
 | 
|
205  | 
by (Asm_full_simp_tac 1);  | 
|
206  | 
by (forward_tac [reachable.reachable_n] 1);  | 
|
| 3457 | 207  | 
by (assume_tac 1);  | 
| 3071 | 208  | 
by (Asm_full_simp_tac 1);  | 
209  | 
(* laststate *)  | 
|
| 3457 | 210  | 
by (etac (move_subprop3 RS sym) 1);  | 
| 3071 | 211  | 
by (REPEAT (atac 1));  | 
212  | 
qed"lemma_2";  | 
|
213  | 
||
214  | 
||
215  | 
(* -------------------------------------------------------------------------------- *)  | 
|
216  | 
||
217  | 
section "Main Theorem: T R A C E - I N C L U S I O N";  | 
|
218  | 
||
219  | 
(* -------------------------------------------------------------------------------- *)  | 
|
220  | 
||
221  | 
||
| 5068 | 222  | 
Goalw [traces_def]  | 
| 6161 | 223  | 
"[| ext C = ext A; is_ref_map f C A |] \  | 
| 3071 | 224  | 
\ ==> traces C <= traces A";  | 
225  | 
||
| 4098 | 226  | 
by (simp_tac(simpset() addsimps [has_trace_def2])1);  | 
| 3071 | 227  | 
by (safe_tac set_cs);  | 
228  | 
||
229  | 
(* give execution of abstract automata *)  | 
|
230  | 
  by (res_inst_tac[("x","corresp_ex A f ex")] bexI 1);
 | 
|
231  | 
||
232  | 
(* Traces coincide, Lemma 1 *)  | 
|
233  | 
by (pair_tac "ex" 1);  | 
|
234  | 
by (etac (lemma_1 RS spec RS mp) 1);  | 
|
235  | 
by (REPEAT (atac 1));  | 
|
| 4098 | 236  | 
by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);  | 
| 3071 | 237  | 
|
238  | 
(* corresp_ex is execution, Lemma 2 *)  | 
|
239  | 
by (pair_tac "ex" 1);  | 
|
| 4098 | 240  | 
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);  | 
| 3071 | 241  | 
(* start state *)  | 
242  | 
by (rtac conjI 1);  | 
|
| 4098 | 243  | 
by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);  | 
| 3071 | 244  | 
(* is-execution-fragment *)  | 
245  | 
by (etac (lemma_2 RS spec RS mp) 1);  | 
|
| 4098 | 246  | 
by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);  | 
| 3071 | 247  | 
qed"trace_inclusion";  | 
248  | 
||
249  | 
||
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
250  | 
(* -------------------------------------------------------------------------------- *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
251  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
252  | 
section "Corollary: F A I R T R A C E - I N C L U S I O N";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
253  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
254  | 
(* -------------------------------------------------------------------------------- *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
255  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
256  | 
|
| 5068 | 257  | 
Goalw [fin_often_def] "(~inf_often P s) = fin_often P s";  | 
| 5132 | 258  | 
by Auto_tac;  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
259  | 
qed"fininf";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
260  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
261  | 
|
| 6161 | 262  | 
Goal "is_wfair A W (s,ex) = \  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
263  | 
\ (fin_often (%x. ~Enabled A W (snd x)) ex --> inf_often (%x. fst x :W) ex)";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
264  | 
by (asm_full_simp_tac (simpset() addsimps [is_wfair_def,fin_often_def])1);  | 
| 5132 | 265  | 
by Auto_tac;  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
266  | 
qed"WF_alt";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
267  | 
|
| 6161 | 268  | 
Goal "[|is_wfair A W (s,ex); inf_often (%x. Enabled A W (snd x)) ex; \  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
269  | 
\ en_persistent A W|] \  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
270  | 
\ ==> inf_often (%x. fst x :W) ex";  | 
| 5132 | 271  | 
by (dtac persistent 1);  | 
272  | 
by (assume_tac 1);  | 
|
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
273  | 
by (asm_full_simp_tac (simpset() addsimps [WF_alt])1);  | 
| 5132 | 274  | 
by Auto_tac;  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
275  | 
qed"WF_persistent";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
276  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
277  | 
|
| 5068 | 278  | 
Goal "!! C A. \  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
279  | 
\ [| is_ref_map f C A; ext C = ext A; \  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
280  | 
\ !! ex. [| ex:executions C; fair_ex C ex|] ==> fair_ex A (corresp_ex A f ex) |] \  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
281  | 
\ ==> fairtraces C <= fairtraces A";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
282  | 
by (simp_tac (simpset() addsimps [fairtraces_def,  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
283  | 
fairexecutions_def]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
284  | 
by (safe_tac set_cs);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
285  | 
by (res_inst_tac[("x","corresp_ex A f ex")] exI 1);
 | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
286  | 
by (safe_tac set_cs);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
287  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
288  | 
(* Traces coincide, Lemma 1 *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
289  | 
by (pair_tac "ex" 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
290  | 
by (etac (lemma_1 RS spec RS mp) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
291  | 
by (REPEAT (atac 1));  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
292  | 
by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
293  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
294  | 
(* corresp_ex is execution, Lemma 2 *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
295  | 
by (pair_tac "ex" 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
296  | 
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
297  | 
(* start state *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
298  | 
by (rtac conjI 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
299  | 
by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
300  | 
(* is-execution-fragment *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
301  | 
by (etac (lemma_2 RS spec RS mp) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
302  | 
by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
303  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
304  | 
(* Fairness *)  | 
| 5132 | 305  | 
by Auto_tac;  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
306  | 
qed"fair_trace_inclusion";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
307  | 
|
| 5068 | 308  | 
Goal "!! C A. \  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
309  | 
\ [| inp(C) = inp(A); out(C)=out(A); \  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
310  | 
\ is_fair_ref_map f C A |] \  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
311  | 
\ ==> fair_implements C A";  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
312  | 
by (asm_full_simp_tac (simpset() addsimps [is_fair_ref_map_def, fair_implements_def,  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
313  | 
fairtraces_def, fairexecutions_def]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
314  | 
by (safe_tac set_cs);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
315  | 
by (res_inst_tac[("x","corresp_ex A f ex")] exI 1);
 | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
316  | 
by (safe_tac set_cs);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
317  | 
(* Traces coincide, Lemma 1 *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
318  | 
by (pair_tac "ex" 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
319  | 
by (etac (lemma_1 RS spec RS mp) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
320  | 
by (simp_tac (simpset() addsimps [externals_def])1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
321  | 
by (SELECT_GOAL (auto_tac (claset(),simpset()))1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
322  | 
by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
323  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
324  | 
(* corresp_ex is execution, Lemma 2 *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
325  | 
by (pair_tac "ex" 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
326  | 
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
327  | 
(* start state *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
328  | 
by (rtac conjI 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
329  | 
by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
330  | 
(* is-execution-fragment *)  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
331  | 
by (etac (lemma_2 RS spec RS mp) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
332  | 
by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);  | 
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
333  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
334  | 
(* Fairness *)  | 
| 5132 | 335  | 
by Auto_tac;  | 
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
336  | 
qed"fair_trace_inclusion2";  | 
| 3071 | 337  | 
|
338  | 
||
339  | 
||
| 
4559
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
340  | 
|
| 
 
8e604d885b54
added files containing temproal logic and abstraction;
 
mueller 
parents: 
4098 
diff
changeset
 | 
341  |