author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 14981 | e73f8140af78 |
child 17876 | b9c92f384109 |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/Traces.ML |
3275 | 2 |
ID: $Id$ |
12218 | 3 |
Author: Olaf Müller |
3071 | 4 |
|
5 |
Theorems about Executions and Traces of I/O automata in HOLCF. |
|
6 |
*) |
|
7 |
||
4815
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
oheimb
parents:
4559
diff
changeset
|
8 |
(* global changes to simpset() and claset(), see also TLS.ML *) |
3275 | 9 |
Delsimps (ex_simps @ all_simps); |
4536 | 10 |
Delsimps [split_paired_Ex]; |
4815
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
oheimb
parents:
4559
diff
changeset
|
11 |
Addsimps [Let_def]; |
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
oheimb
parents:
4559
diff
changeset
|
12 |
claset_ref() := claset() delSWrapper "split_all_tac"; |
3071 | 13 |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
14 |
val exec_rws = [executions_def,is_exec_frag_def]; |
3071 | 15 |
|
16 |
||
17 |
||
18 |
(* ----------------------------------------------------------------------------------- *) |
|
19 |
||
20 |
section "recursive equations of operators"; |
|
21 |
||
22 |
||
23 |
(* ---------------------------------------------------------------- *) |
|
24 |
(* filter_act *) |
|
25 |
(* ---------------------------------------------------------------- *) |
|
26 |
||
27 |
||
10835 | 28 |
Goal "filter_act$UU = UU"; |
4098 | 29 |
by (simp_tac (simpset() addsimps [filter_act_def]) 1); |
3071 | 30 |
qed"filter_act_UU"; |
31 |
||
10835 | 32 |
Goal "filter_act$nil = nil"; |
4098 | 33 |
by (simp_tac (simpset() addsimps [filter_act_def]) 1); |
3071 | 34 |
qed"filter_act_nil"; |
35 |
||
10835 | 36 |
Goal "filter_act$(x>>xs) = (fst x) >> filter_act$xs"; |
4098 | 37 |
by (simp_tac (simpset() addsimps [filter_act_def]) 1); |
3071 | 38 |
qed"filter_act_cons"; |
39 |
||
40 |
Addsimps [filter_act_UU,filter_act_nil,filter_act_cons]; |
|
41 |
||
42 |
||
43 |
(* ---------------------------------------------------------------- *) |
|
44 |
(* mk_trace *) |
|
45 |
(* ---------------------------------------------------------------- *) |
|
46 |
||
10835 | 47 |
Goal "mk_trace A$UU=UU"; |
4098 | 48 |
by (simp_tac (simpset() addsimps [mk_trace_def]) 1); |
3071 | 49 |
qed"mk_trace_UU"; |
50 |
||
10835 | 51 |
Goal "mk_trace A$nil=nil"; |
4098 | 52 |
by (simp_tac (simpset() addsimps [mk_trace_def]) 1); |
3071 | 53 |
qed"mk_trace_nil"; |
54 |
||
10835 | 55 |
Goal "mk_trace A$(at >> xs) = \ |
3071 | 56 |
\ (if ((fst at):ext A) \ |
10835 | 57 |
\ then (fst at) >> (mk_trace A$xs) \ |
58 |
\ else mk_trace A$xs)"; |
|
3071 | 59 |
|
4098 | 60 |
by (asm_full_simp_tac (simpset() addsimps [mk_trace_def]) 1); |
3071 | 61 |
qed"mk_trace_cons"; |
62 |
||
63 |
Addsimps [mk_trace_UU,mk_trace_nil,mk_trace_cons]; |
|
64 |
||
65 |
(* ---------------------------------------------------------------- *) |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
66 |
(* is_exec_fragC *) |
3071 | 67 |
(* ---------------------------------------------------------------- *) |
68 |
||
69 |
||
5068 | 70 |
Goal "is_exec_fragC A = (LAM ex. (%s. case ex of \ |
3071 | 71 |
\ nil => TT \ |
72 |
\ | x##xs => (flift1 \ |
|
10835 | 73 |
\ (%p. Def ((s,p):trans_of A) andalso (is_exec_fragC A$xs) (snd p)) \ |
74 |
\ $x) \ |
|
3071 | 75 |
\ ))"; |
76 |
by (rtac trans 1); |
|
3457 | 77 |
by (rtac fix_eq2 1); |
78 |
by (rtac is_exec_fragC_def 1); |
|
79 |
by (rtac beta_cfun 1); |
|
4098 | 80 |
by (simp_tac (simpset() addsimps [flift1_def]) 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
81 |
qed"is_exec_fragC_unfold"; |
3071 | 82 |
|
10835 | 83 |
Goal "(is_exec_fragC A$UU) s=UU"; |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
84 |
by (stac is_exec_fragC_unfold 1); |
3071 | 85 |
by (Simp_tac 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
86 |
qed"is_exec_fragC_UU"; |
3071 | 87 |
|
10835 | 88 |
Goal "(is_exec_fragC A$nil) s = TT"; |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
89 |
by (stac is_exec_fragC_unfold 1); |
3071 | 90 |
by (Simp_tac 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
91 |
qed"is_exec_fragC_nil"; |
3071 | 92 |
|
10835 | 93 |
Goal "(is_exec_fragC A$(pr>>xs)) s = \ |
3071 | 94 |
\ (Def ((s,pr):trans_of A) \ |
10835 | 95 |
\ andalso (is_exec_fragC A$xs)(snd pr))"; |
3457 | 96 |
by (rtac trans 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
97 |
by (stac is_exec_fragC_unfold 1); |
7229
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
wenzelm
parents:
5068
diff
changeset
|
98 |
by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1); |
3071 | 99 |
by (Simp_tac 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
100 |
qed"is_exec_fragC_cons"; |
3071 | 101 |
|
102 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
103 |
Addsimps [is_exec_fragC_UU,is_exec_fragC_nil,is_exec_fragC_cons]; |
3071 | 104 |
|
105 |
||
106 |
(* ---------------------------------------------------------------- *) |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
107 |
(* is_exec_frag *) |
3071 | 108 |
(* ---------------------------------------------------------------- *) |
109 |
||
5068 | 110 |
Goal "is_exec_frag A (s, UU)"; |
4098 | 111 |
by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
112 |
qed"is_exec_frag_UU"; |
3071 | 113 |
|
5068 | 114 |
Goal "is_exec_frag A (s, nil)"; |
4098 | 115 |
by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
116 |
qed"is_exec_frag_nil"; |
3071 | 117 |
|
5068 | 118 |
Goal "is_exec_frag A (s, (a,t)>>ex) = \ |
3071 | 119 |
\ (((s,a,t):trans_of A) & \ |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
120 |
\ is_exec_frag A (t, ex))"; |
4098 | 121 |
by (simp_tac (simpset() addsimps [is_exec_frag_def]) 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
122 |
qed"is_exec_frag_cons"; |
3071 | 123 |
|
124 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
125 |
(* Delsimps [is_exec_fragC_UU,is_exec_fragC_nil,is_exec_fragC_cons]; *) |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
126 |
Addsimps [is_exec_frag_UU,is_exec_frag_nil, is_exec_frag_cons]; |
3071 | 127 |
|
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
128 |
(* ---------------------------------------------------------------------------- *) |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
129 |
section "laststate"; |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
130 |
(* ---------------------------------------------------------------------------- *) |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
131 |
|
5068 | 132 |
Goal "laststate (s,UU) = s"; |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
133 |
by (simp_tac (simpset() addsimps [laststate_def]) 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
134 |
qed"laststate_UU"; |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
135 |
|
5068 | 136 |
Goal "laststate (s,nil) = s"; |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
137 |
by (simp_tac (simpset() addsimps [laststate_def]) 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
138 |
qed"laststate_nil"; |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
139 |
|
5068 | 140 |
Goal "!! ex. Finite ex ==> laststate (s,at>>ex) = laststate (snd at,ex)"; |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
141 |
by (simp_tac (simpset() addsimps [laststate_def]) 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
142 |
by (case_tac "ex=nil" 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
143 |
by (Asm_simp_tac 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
144 |
by (Asm_simp_tac 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
145 |
by (dtac (Finite_Last1 RS mp) 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
146 |
by (assume_tac 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
147 |
by (def_tac 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
148 |
qed"laststate_cons"; |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
149 |
|
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
150 |
Addsimps [laststate_UU,laststate_nil,laststate_cons]; |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
151 |
|
5068 | 152 |
Goal "!!ex. Finite ex ==> (! s. ? u. laststate (s,ex)=u)"; |
4559
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
153 |
by (Seq_Finite_induct_tac 1); |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
154 |
qed"exists_laststate"; |
8e604d885b54
added files containing temproal logic and abstraction;
mueller
parents:
4536
diff
changeset
|
155 |
|
3071 | 156 |
|
157 |
(* -------------------------------------------------------------------------------- *) |
|
158 |
||
159 |
section "has_trace, mk_trace"; |
|
160 |
||
161 |
(* alternative definition of has_trace tailored for the refinement proof, as it does not |
|
162 |
take the detour of schedules *) |
|
163 |
||
5068 | 164 |
Goalw [executions_def,mk_trace_def,has_trace_def,schedules_def,has_schedule_def] |
10835 | 165 |
"has_trace A b = (? ex:executions A. b = mk_trace A$(snd ex))"; |
3071 | 166 |
|
167 |
by (safe_tac set_cs); |
|
168 |
(* 1 *) |
|
169 |
by (res_inst_tac[("x","ex")] bexI 1); |
|
170 |
by (stac beta_cfun 1); |
|
171 |
by (cont_tacR 1); |
|
172 |
by (Simp_tac 1); |
|
173 |
by (Asm_simp_tac 1); |
|
174 |
(* 2 *) |
|
10835 | 175 |
by (res_inst_tac[("x","filter_act$(snd ex)")] bexI 1); |
3071 | 176 |
by (stac beta_cfun 1); |
177 |
by (cont_tacR 1); |
|
178 |
by (Simp_tac 1); |
|
179 |
by (safe_tac set_cs); |
|
180 |
by (res_inst_tac[("x","ex")] bexI 1); |
|
181 |
by (REPEAT (Asm_simp_tac 1)); |
|
182 |
qed"has_trace_def2"; |
|
183 |
||
3275 | 184 |
|
185 |
(* -------------------------------------------------------------------------------- *) |
|
186 |
||
187 |
section "signatures and executions, schedules"; |
|
188 |
||
189 |
||
190 |
(* All executions of A have only actions of A. This is only true because of the |
|
191 |
predicate state_trans (part of the predicate IOA): We have no dependent types. |
|
192 |
For executions of parallel automata this assumption is not needed, as in par_def |
|
193 |
this condition is included once more. (see Lemmas 1.1.1c in CompoExecs for example) *) |
|
194 |
||
5068 | 195 |
Goal |
3521 | 196 |
"!! A. is_trans_of A ==> \ |
10835 | 197 |
\ ! s. is_exec_frag A (s,xs) --> Forall (%a. a:act A) (filter_act$xs)"; |
3275 | 198 |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
199 |
by (pair_induct_tac "xs" [is_exec_frag_def,Forall_def,sforall_def] 1); |
3275 | 200 |
(* main case *) |
201 |
ren "ss a t" 1; |
|
202 |
by (safe_tac set_cs); |
|
4098 | 203 |
by (REPEAT (asm_full_simp_tac (simpset() addsimps [is_trans_of_def]) 1)); |
3275 | 204 |
qed"execfrag_in_sig"; |
205 |
||
5068 | 206 |
Goal |
3521 | 207 |
"!! A.[| is_trans_of A; x:executions A |] ==> \ |
10835 | 208 |
\ Forall (%a. a:act A) (filter_act$(snd x))"; |
3275 | 209 |
|
4098 | 210 |
by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1); |
3275 | 211 |
by (pair_tac "x" 1); |
3457 | 212 |
by (rtac (execfrag_in_sig RS spec RS mp) 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4423
diff
changeset
|
213 |
by Auto_tac; |
3275 | 214 |
qed"exec_in_sig"; |
215 |
||
5068 | 216 |
Goalw [schedules_def,has_schedule_def] |
3521 | 217 |
"!! A.[| is_trans_of A; x:schedules A |] ==> \ |
3842 | 218 |
\ Forall (%a. a:act A) x"; |
3275 | 219 |
|
4098 | 220 |
by (fast_tac (claset() addSIs [exec_in_sig]) 1); |
3275 | 221 |
qed"scheds_in_sig"; |
222 |
||
4283 | 223 |
(* |
224 |
||
225 |
is ok but needs ForallQFilterP which has to been proven first (is trivial also) |
|
226 |
||
5068 | 227 |
Goalw [traces_def,has_trace_def] |
4283 | 228 |
"!! A.[| x:traces A |] ==> \ |
229 |
\ Forall (%a. a:act A) x"; |
|
230 |
by (safe_tac set_cs ); |
|
4423 | 231 |
by (rtac ForallQFilterP 1); |
4283 | 232 |
by (fast_tac (!claset addSIs [ext_is_act]) 1); |
233 |
qed"traces_in_sig"; |
|
234 |
*) |
|
235 |
||
3275 | 236 |
|
237 |
(* -------------------------------------------------------------------------------- *) |
|
238 |
||
239 |
section "executions are prefix closed"; |
|
240 |
||
241 |
(* only admissible in y, not if done in x !! *) |
|
5068 | 242 |
Goal "!x s. is_exec_frag A (s,x) & y<<x --> is_exec_frag A (s,y)"; |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
243 |
by (pair_induct_tac "y" [is_exec_frag_def] 1); |
3275 | 244 |
by (strip_tac 1); |
245 |
by (Seq_case_simp_tac "xa" 1); |
|
246 |
by (pair_tac "a" 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4423
diff
changeset
|
247 |
by Auto_tac; |
3275 | 248 |
qed"execfrag_prefixclosed"; |
249 |
||
250 |
bind_thm ("exec_prefixclosed",conjI RS (execfrag_prefixclosed RS spec RS spec RS mp)); |
|
251 |
||
3361 | 252 |
|
253 |
(* second prefix notion for Finite x *) |
|
254 |
||
5068 | 255 |
Goal "! y s. is_exec_frag A (s,x@@y) --> is_exec_frag A (s,x)"; |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3361
diff
changeset
|
256 |
by (pair_induct_tac "x" [is_exec_frag_def] 1); |
3361 | 257 |
by (strip_tac 1); |
258 |
by (Seq_case_simp_tac "s" 1); |
|
259 |
by (pair_tac "a" 1); |
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4423
diff
changeset
|
260 |
by Auto_tac; |
3361 | 261 |
qed_spec_mp"exec_prefix2closed"; |
262 |