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