author | nipkow |
Fri, 17 Oct 1997 15:23:14 +0200 | |
changeset 3918 | 94e0fdcb7b91 |
parent 3847 | d5905b98291f |
child 4034 | 5bb30bedbdc2 |
permissions | -rw-r--r-- |
3071 | 1 |
(* Title: HOLCF/IOA/meta_theory/RefCorrectness.ML |
3275 | 2 |
ID: $Id$ |
3071 | 3 |
Author: Olaf Mueller |
4 |
Copyright 1996 TU Muenchen |
|
5 |
||
6 |
Correctness of Refinement Mappings in HOLCF/IOA |
|
7 |
*) |
|
8 |
||
9 |
||
10 |
||
11 |
(* -------------------------------------------------------------------------------- *) |
|
12 |
||
13 |
section "corresp_ex"; |
|
14 |
||
15 |
||
16 |
(* ---------------------------------------------------------------- *) |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
17 |
(* corresp_exC *) |
3071 | 18 |
(* ---------------------------------------------------------------- *) |
19 |
||
20 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
21 |
goal thy "corresp_exC A f = (LAM ex. (%s. case ex of \ |
3071 | 22 |
\ nil => nil \ |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
23 |
\ | x##xs => (flift1 (%pr. (@cex. move A cex (f s) (fst pr) (f (snd pr))) \ |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
24 |
\ @@ ((corresp_exC A f `xs) (snd pr))) \ |
3071 | 25 |
\ `x) ))"; |
26 |
by (rtac trans 1); |
|
3457 | 27 |
by (rtac fix_eq2 1); |
28 |
by (rtac corresp_exC_def 1); |
|
29 |
by (rtac beta_cfun 1); |
|
3071 | 30 |
by (simp_tac (!simpset addsimps [flift1_def]) 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
31 |
qed"corresp_exC_unfold"; |
3071 | 32 |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
33 |
goal thy "(corresp_exC A f`UU) s=UU"; |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
34 |
by (stac corresp_exC_unfold 1); |
3071 | 35 |
by (Simp_tac 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
36 |
qed"corresp_exC_UU"; |
3071 | 37 |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
38 |
goal thy "(corresp_exC A f`nil) s = nil"; |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
39 |
by (stac corresp_exC_unfold 1); |
3071 | 40 |
by (Simp_tac 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
41 |
qed"corresp_exC_nil"; |
3071 | 42 |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
43 |
goal thy "(corresp_exC A f`(at>>xs)) s = \ |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
44 |
\ (@cex. move A cex (f s) (fst at) (f (snd at))) \ |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
45 |
\ @@ ((corresp_exC A f`xs) (snd at))"; |
3457 | 46 |
by (rtac trans 1); |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
47 |
by (stac corresp_exC_unfold 1); |
3071 | 48 |
by (asm_full_simp_tac (!simpset addsimps [Cons_def,flift1_def]) 1); |
49 |
by (Simp_tac 1); |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
50 |
qed"corresp_exC_cons"; |
3071 | 51 |
|
52 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
53 |
Addsimps [corresp_exC_UU,corresp_exC_nil,corresp_exC_cons]; |
3071 | 54 |
|
55 |
||
56 |
||
57 |
(* ------------------------------------------------------------------ *) |
|
58 |
(* The following lemmata describe the definition *) |
|
59 |
(* of move in more detail *) |
|
60 |
(* ------------------------------------------------------------------ *) |
|
61 |
||
62 |
section"properties of move"; |
|
63 |
||
64 |
goalw thy [is_ref_map_def] |
|
65 |
"!!f. [|is_ref_map f C A; reachable C s; (s,a,t):trans_of C|] ==>\ |
|
66 |
\ move A (@x. move A x (f s) a (f t)) (f s) a (f t)"; |
|
67 |
||
3847 | 68 |
by (subgoal_tac "? ex. move A ex (f s) a (f t)" 1); |
3071 | 69 |
by (Asm_full_simp_tac 2); |
70 |
by (etac exE 1); |
|
71 |
by (rtac selectI 1); |
|
72 |
by (assume_tac 1); |
|
73 |
qed"move_is_move"; |
|
74 |
||
75 |
goal thy |
|
76 |
"!!f. [|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
|
77 |
\ is_exec_frag A (f s,@x. move A x (f s) a (f t))"; |
3071 | 78 |
by (cut_inst_tac [] move_is_move 1); |
79 |
by (REPEAT (assume_tac 1)); |
|
80 |
by (asm_full_simp_tac (!simpset addsimps [move_def]) 1); |
|
81 |
qed"move_subprop1"; |
|
82 |
||
83 |
goal thy |
|
84 |
"!!f. [|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
|
85 |
\ Finite ((@x. move A x (f s) a (f t)))"; |
3071 | 86 |
by (cut_inst_tac [] move_is_move 1); |
87 |
by (REPEAT (assume_tac 1)); |
|
88 |
by (asm_full_simp_tac (!simpset addsimps [move_def]) 1); |
|
89 |
qed"move_subprop2"; |
|
90 |
||
91 |
goal thy |
|
92 |
"!!f. [|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
|
93 |
\ laststate (f s,@x. move A x (f s) a (f t)) = (f t)"; |
3071 | 94 |
by (cut_inst_tac [] move_is_move 1); |
95 |
by (REPEAT (assume_tac 1)); |
|
96 |
by (asm_full_simp_tac (!simpset addsimps [move_def]) 1); |
|
97 |
qed"move_subprop3"; |
|
98 |
||
99 |
goal thy |
|
100 |
"!!f. [|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
|
101 |
\ mk_trace A`((@x. move A x (f s) a (f t))) = \ |
3071 | 102 |
\ (if a:ext A then a>>nil else nil)"; |
103 |
||
104 |
by (cut_inst_tac [] move_is_move 1); |
|
105 |
by (REPEAT (assume_tac 1)); |
|
106 |
by (asm_full_simp_tac (!simpset addsimps [move_def]) 1); |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
107 |
qed"move_subprop4"; |
3071 | 108 |
|
109 |
||
110 |
(* ------------------------------------------------------------------ *) |
|
111 |
(* The following lemmata contribute to *) |
|
112 |
(* TRACE INCLUSION Part 1: Traces coincide *) |
|
113 |
(* ------------------------------------------------------------------ *) |
|
114 |
||
115 |
section "Lemmata for <=="; |
|
116 |
||
117 |
(* --------------------------------------------------- *) |
|
118 |
(* Lemma 1.1: Distribution of mk_trace and @@ *) |
|
119 |
(* --------------------------------------------------- *) |
|
120 |
||
121 |
||
122 |
goal thy "mk_trace C`(ex1 @@ ex2)= (mk_trace C`ex1) @@ (mk_trace C`ex2)"; |
|
123 |
by (simp_tac (!simpset addsimps [mk_trace_def,filter_act_def, |
|
124 |
FilterConc,MapConc]) 1); |
|
125 |
qed"mk_traceConc"; |
|
126 |
||
127 |
||
128 |
||
129 |
(* ------------------------------------------------------ |
|
130 |
Lemma 1 :Traces coincide |
|
131 |
------------------------------------------------------- *) |
|
132 |
||
133 |
goalw thy [corresp_ex_def] |
|
134 |
"!!f.[|is_ref_map f C A; ext C = ext A|] ==> \ |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
135 |
\ !s. reachable C s & is_exec_frag C (s,xs) --> \ |
3071 | 136 |
\ mk_trace C`xs = mk_trace A`(snd (corresp_ex A f (s,xs)))"; |
137 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
138 |
by (pair_induct_tac "xs" [is_exec_frag_def] 1); |
3071 | 139 |
(* cons case *) |
140 |
by (safe_tac set_cs); |
|
141 |
by (asm_full_simp_tac (!simpset addsimps [mk_traceConc]) 1); |
|
142 |
by (forward_tac [reachable.reachable_n] 1); |
|
3457 | 143 |
by (assume_tac 1); |
3071 | 144 |
by (eres_inst_tac [("x","y")] allE 1); |
145 |
by (Asm_full_simp_tac 1); |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
146 |
by (asm_full_simp_tac (!simpset addsimps [move_subprop4] |
3071 | 147 |
setloop split_tac [expand_if] ) 1); |
148 |
qed"lemma_1"; |
|
149 |
||
150 |
||
151 |
(* ------------------------------------------------------------------ *) |
|
152 |
(* The following lemmata contribute to *) |
|
153 |
(* TRACE INCLUSION Part 2: corresp_ex is execution *) |
|
154 |
(* ------------------------------------------------------------------ *) |
|
155 |
||
156 |
section "Lemmata for ==>"; |
|
157 |
||
158 |
(* -------------------------------------------------- *) |
|
159 |
(* Lemma 2.1 *) |
|
160 |
(* -------------------------------------------------- *) |
|
161 |
||
162 |
goal thy |
|
163 |
"Finite xs --> \ |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
164 |
\(!s .is_exec_frag A (s,xs) & is_exec_frag A (t,ys) & \ |
3071 | 165 |
\ t = laststate (s,xs) \ |
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
166 |
\ --> is_exec_frag A (s,xs @@ ys))"; |
3071 | 167 |
|
3457 | 168 |
by (rtac impI 1); |
3071 | 169 |
by (Seq_Finite_induct_tac 1); |
170 |
(* base_case *) |
|
171 |
by (fast_tac HOL_cs 1); |
|
172 |
(* main case *) |
|
173 |
by (safe_tac set_cs); |
|
174 |
by (pair_tac "a" 1); |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
175 |
qed_spec_mp"lemma_2_1"; |
3071 | 176 |
|
177 |
||
178 |
(* ----------------------------------------------------------- *) |
|
179 |
(* Lemma 2 : corresp_ex is execution *) |
|
180 |
(* ----------------------------------------------------------- *) |
|
181 |
||
182 |
||
183 |
||
184 |
goalw thy [corresp_ex_def] |
|
185 |
"!!f.[| is_ref_map f C A |] ==>\ |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
186 |
\ !s. reachable C s & is_exec_frag C (s,xs) \ |
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
187 |
\ --> is_exec_frag A (corresp_ex A f (s,xs))"; |
3071 | 188 |
|
189 |
by (Asm_full_simp_tac 1); |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
190 |
by (pair_induct_tac "xs" [is_exec_frag_def] 1); |
3071 | 191 |
(* main case *) |
192 |
by (safe_tac set_cs); |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
193 |
by (res_inst_tac [("t","f y")] lemma_2_1 1); |
3071 | 194 |
|
195 |
(* Finite *) |
|
3457 | 196 |
by (etac move_subprop2 1); |
3071 | 197 |
by (REPEAT (atac 1)); |
198 |
by (rtac conjI 1); |
|
199 |
||
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
200 |
(* is_exec_frag *) |
3457 | 201 |
by (etac move_subprop1 1); |
3071 | 202 |
by (REPEAT (atac 1)); |
203 |
by (rtac conjI 1); |
|
204 |
||
205 |
(* Induction hypothesis *) |
|
206 |
(* reachable_n looping, therefore apply it manually *) |
|
207 |
by (eres_inst_tac [("x","y")] allE 1); |
|
208 |
by (Asm_full_simp_tac 1); |
|
209 |
by (forward_tac [reachable.reachable_n] 1); |
|
3457 | 210 |
by (assume_tac 1); |
3071 | 211 |
by (Asm_full_simp_tac 1); |
212 |
(* laststate *) |
|
3457 | 213 |
by (etac (move_subprop3 RS sym) 1); |
3071 | 214 |
by (REPEAT (atac 1)); |
215 |
qed"lemma_2"; |
|
216 |
||
217 |
||
218 |
(* -------------------------------------------------------------------------------- *) |
|
219 |
||
220 |
section "Main Theorem: T R A C E - I N C L U S I O N"; |
|
221 |
||
222 |
(* -------------------------------------------------------------------------------- *) |
|
223 |
||
224 |
||
225 |
goalw thy [traces_def] |
|
3433
2de17c994071
added deadlock freedom, polished definitions and proofs
mueller
parents:
3275
diff
changeset
|
226 |
"!!f. [| ext C = ext A; is_ref_map f C A |] \ |
3071 | 227 |
\ ==> traces C <= traces A"; |
228 |
||
229 |
by (simp_tac(!simpset addsimps [has_trace_def2])1); |
|
230 |
by (safe_tac set_cs); |
|
231 |
||
232 |
(* give execution of abstract automata *) |
|
233 |
by (res_inst_tac[("x","corresp_ex A f ex")] bexI 1); |
|
234 |
||
235 |
(* Traces coincide, Lemma 1 *) |
|
236 |
by (pair_tac "ex" 1); |
|
237 |
by (etac (lemma_1 RS spec RS mp) 1); |
|
238 |
by (REPEAT (atac 1)); |
|
239 |
by (asm_full_simp_tac (!simpset addsimps [executions_def,reachable.reachable_0]) 1); |
|
240 |
||
241 |
(* corresp_ex is execution, Lemma 2 *) |
|
242 |
by (pair_tac "ex" 1); |
|
243 |
by (asm_full_simp_tac (!simpset addsimps [executions_def]) 1); |
|
244 |
(* start state *) |
|
245 |
by (rtac conjI 1); |
|
246 |
by (asm_full_simp_tac (!simpset addsimps [is_ref_map_def,corresp_ex_def]) 1); |
|
247 |
(* is-execution-fragment *) |
|
248 |
by (etac (lemma_2 RS spec RS mp) 1); |
|
249 |
by (asm_full_simp_tac (!simpset addsimps [reachable.reachable_0]) 1); |
|
250 |
qed"trace_inclusion"; |
|
251 |
||
252 |
||
253 |
||
254 |
||
255 |