author | paulson |
Fri, 24 Sep 1999 16:33:57 +0200 | |
changeset 7594 | 8a188ef6545e |
parent 7524 | 15e4a6db638a |
child 7826 | c6a8b73b6c2a |
permissions | -rw-r--r-- |
4776 | 1 |
(* Title: HOL/UNITY/WFair |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Weak Fairness versions of transient, ensures, leadsTo. |
|
7 |
||
8 |
From Misra, "A Logic for Concurrent Programming", 1994 |
|
9 |
*) |
|
10 |
||
11 |
||
5648 | 12 |
overload_1st_set "WFair.transient"; |
13 |
overload_1st_set "WFair.ensures"; |
|
6536 | 14 |
overload_1st_set "WFair.op leadsTo"; |
5340 | 15 |
|
4776 | 16 |
(*** transient ***) |
17 |
||
5069 | 18 |
Goalw [stable_def, constrains_def, transient_def] |
5648 | 19 |
"[| F : stable A; F : transient A |] ==> A = {}"; |
4776 | 20 |
by (Blast_tac 1); |
21 |
qed "stable_transient_empty"; |
|
22 |
||
5069 | 23 |
Goalw [transient_def] |
5648 | 24 |
"[| F : transient A; B<=A |] ==> F : transient B"; |
4776 | 25 |
by (Clarify_tac 1); |
6012
1894bfc4aee9
Addition of the States component; parts of Comp not working
paulson
parents:
5971
diff
changeset
|
26 |
by (blast_tac (claset() addSIs [rev_bexI]) 1); |
4776 | 27 |
qed "transient_strengthen"; |
28 |
||
5069 | 29 |
Goalw [transient_def] |
5648 | 30 |
"[| act: Acts F; A <= Domain act; act^^A <= -A |] ==> F : transient A"; |
4776 | 31 |
by (Blast_tac 1); |
32 |
qed "transient_mem"; |
|
33 |
||
34 |
||
35 |
(*** ensures ***) |
|
36 |
||
5069 | 37 |
Goalw [ensures_def] |
7524 | 38 |
"[| F : (A-B) co (A Un B); F : transient (A-B) |] ==> F : A ensures B"; |
4776 | 39 |
by (Blast_tac 1); |
40 |
qed "ensuresI"; |
|
41 |
||
5069 | 42 |
Goalw [ensures_def] |
6536 | 43 |
"F : A ensures B ==> F : (A-B) co (A Un B) & F : transient (A-B)"; |
4776 | 44 |
by (Blast_tac 1); |
45 |
qed "ensuresD"; |
|
46 |
||
47 |
(*The L-version (precondition strengthening) doesn't hold for ENSURES*) |
|
5069 | 48 |
Goalw [ensures_def] |
6536 | 49 |
"[| F : A ensures A'; A'<=B' |] ==> F : A ensures B'"; |
4776 | 50 |
by (blast_tac (claset() addIs [constrains_weaken, transient_strengthen]) 1); |
51 |
qed "ensures_weaken_R"; |
|
52 |
||
5069 | 53 |
Goalw [ensures_def, constrains_def, transient_def] |
6536 | 54 |
"F : A ensures UNIV"; |
5340 | 55 |
by Auto_tac; |
4776 | 56 |
qed "ensures_UNIV"; |
57 |
||
5069 | 58 |
Goalw [ensures_def] |
5648 | 59 |
"[| F : stable C; \ |
6536 | 60 |
\ F : (C Int (A - A')) co (A Un A'); \ |
5648 | 61 |
\ F : transient (C Int (A-A')) |] \ |
6536 | 62 |
\ ==> F : (C Int A) ensures (C Int A')"; |
4776 | 63 |
by (asm_simp_tac (simpset() addsimps [Int_Un_distrib RS sym, |
64 |
Diff_Int_distrib RS sym, |
|
65 |
stable_constrains_Int]) 1); |
|
66 |
qed "stable_ensures_Int"; |
|
67 |
||
6536 | 68 |
Goal "[| F : stable A; F : transient C; A <= B Un C |] ==> F : A ensures B"; |
5640 | 69 |
by (asm_full_simp_tac (simpset() addsimps [ensures_def, stable_def]) 1); |
70 |
by (blast_tac (claset() addIs [constrains_weaken, transient_strengthen]) 1); |
|
71 |
qed "stable_transient_ensures"; |
|
72 |
||
4776 | 73 |
|
74 |
(*** leadsTo ***) |
|
75 |
||
6536 | 76 |
Goalw [leadsTo_def] "F : A ensures B ==> F : A leadsTo B"; |
6801
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
paulson
parents:
6714
diff
changeset
|
77 |
by (blast_tac (claset() addIs [leads.Basis]) 1); |
5648 | 78 |
qed "leadsTo_Basis"; |
4776 | 79 |
|
5648 | 80 |
Goalw [leadsTo_def] |
6536 | 81 |
"[| F : A leadsTo B; F : B leadsTo C |] ==> F : A leadsTo C"; |
6801
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
paulson
parents:
6714
diff
changeset
|
82 |
by (blast_tac (claset() addIs [leads.Trans]) 1); |
5648 | 83 |
qed "leadsTo_Trans"; |
84 |
||
6536 | 85 |
Goal "F : transient A ==> F : A leadsTo (-A)"; |
5640 | 86 |
by (asm_simp_tac |
87 |
(simpset() addsimps [leadsTo_Basis, ensuresI, Compl_partition]) 1); |
|
88 |
qed "transient_imp_leadsTo"; |
|
89 |
||
6536 | 90 |
Goal "F : A leadsTo UNIV"; |
4776 | 91 |
by (blast_tac (claset() addIs [ensures_UNIV RS leadsTo_Basis]) 1); |
92 |
qed "leadsTo_UNIV"; |
|
93 |
Addsimps [leadsTo_UNIV]; |
|
94 |
||
95 |
(*Useful with cancellation, disjunction*) |
|
6536 | 96 |
Goal "F : A leadsTo (A' Un A') ==> F : A leadsTo A'"; |
4776 | 97 |
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); |
98 |
qed "leadsTo_Un_duplicate"; |
|
99 |
||
6536 | 100 |
Goal "F : A leadsTo (A' Un C Un C) ==> F : A leadsTo (A' Un C)"; |
4776 | 101 |
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1); |
102 |
qed "leadsTo_Un_duplicate2"; |
|
103 |
||
104 |
(*The Union introduction rule as we should have liked to state it*) |
|
5648 | 105 |
val prems = Goalw [leadsTo_def] |
6536 | 106 |
"(!!A. A : S ==> F : A leadsTo B) ==> F : (Union S) leadsTo B"; |
6801
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
paulson
parents:
6714
diff
changeset
|
107 |
by (blast_tac (claset() addIs [leads.Union] addDs prems) 1); |
4776 | 108 |
qed "leadsTo_Union"; |
109 |
||
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
110 |
val prems = Goalw [leadsTo_def] |
7524 | 111 |
"(!!A. A : S ==> F : (A Int C) leadsTo B) ==> F : (Union S Int C) leadsTo B"; |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
112 |
by (simp_tac (HOL_ss addsimps [Int_Union_Union]) 1); |
6801
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
paulson
parents:
6714
diff
changeset
|
113 |
by (blast_tac (claset() addIs [leads.Union] addDs prems) 1); |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
114 |
qed "leadsTo_Union_Int"; |
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
115 |
|
5648 | 116 |
val prems = Goal |
6536 | 117 |
"(!!i. i : I ==> F : (A i) leadsTo B) ==> F : (UN i:I. A i) leadsTo B"; |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
118 |
by (stac (Union_image_eq RS sym) 1); |
5648 | 119 |
by (blast_tac (claset() addIs leadsTo_Union::prems) 1); |
4776 | 120 |
qed "leadsTo_UN"; |
121 |
||
122 |
(*Binary union introduction rule*) |
|
6536 | 123 |
Goal "[| F : A leadsTo C; F : B leadsTo C |] ==> F : (A Un B) leadsTo C"; |
4776 | 124 |
by (stac Un_eq_Union 1); |
125 |
by (blast_tac (claset() addIs [leadsTo_Union]) 1); |
|
126 |
qed "leadsTo_Un"; |
|
127 |
||
6714 | 128 |
val prems = |
129 |
Goal "(!!x. x : A ==> F : {x} leadsTo B) ==> F : A leadsTo B"; |
|
130 |
by (stac (UN_singleton RS sym) 1 THEN rtac leadsTo_UN 1); |
|
131 |
by (blast_tac (claset() addIs prems) 1); |
|
132 |
qed "single_leadsTo_I"; |
|
133 |
||
4776 | 134 |
|
135 |
(*The INDUCTION rule as we should have liked to state it*) |
|
5648 | 136 |
val major::prems = Goalw [leadsTo_def] |
6536 | 137 |
"[| F : za leadsTo zb; \ |
138 |
\ !!A B. F : A ensures B ==> P A B; \ |
|
139 |
\ !!A B C. [| F : A leadsTo B; P A B; F : B leadsTo C; P B C |] \ |
|
4776 | 140 |
\ ==> P A C; \ |
6536 | 141 |
\ !!B S. ALL A:S. F : A leadsTo B & P A B ==> P (Union S) B \ |
4776 | 142 |
\ |] ==> P za zb"; |
6801
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
paulson
parents:
6714
diff
changeset
|
143 |
by (rtac (major RS CollectD RS leads.induct) 1); |
4776 | 144 |
by (REPEAT (blast_tac (claset() addIs prems) 1)); |
145 |
qed "leadsTo_induct"; |
|
146 |
||
147 |
||
7594
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7524
diff
changeset
|
148 |
Goal "A<=B ==> F : A ensures B"; |
4776 | 149 |
by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]); |
150 |
by (Blast_tac 1); |
|
7594
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7524
diff
changeset
|
151 |
qed "subset_imp_ensures"; |
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7524
diff
changeset
|
152 |
|
8a188ef6545e
working version with co-guarantees-leadsto results
paulson
parents:
7524
diff
changeset
|
153 |
bind_thm ("subset_imp_leadsTo", subset_imp_ensures RS leadsTo_Basis); |
4776 | 154 |
|
155 |
bind_thm ("empty_leadsTo", empty_subsetI RS subset_imp_leadsTo); |
|
156 |
Addsimps [empty_leadsTo]; |
|
157 |
||
158 |
||
6536 | 159 |
Goal "[| F : A leadsTo A'; A'<=B' |] ==> F : A leadsTo B'"; |
5648 | 160 |
by (blast_tac (claset() addIs [subset_imp_leadsTo, leadsTo_Trans]) 1); |
161 |
qed "leadsTo_weaken_R"; |
|
4776 | 162 |
|
6536 | 163 |
Goal "[| F : A leadsTo A'; B<=A |] ==> F : B leadsTo A'"; |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
164 |
by (blast_tac (claset() addIs [leadsTo_Trans, subset_imp_leadsTo]) 1); |
4776 | 165 |
qed_spec_mp "leadsTo_weaken_L"; |
166 |
||
167 |
(*Distributes over binary unions*) |
|
6536 | 168 |
Goal "F : (A Un B) leadsTo C = (F : A leadsTo C & F : B leadsTo C)"; |
4776 | 169 |
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken_L]) 1); |
170 |
qed "leadsTo_Un_distrib"; |
|
171 |
||
6536 | 172 |
Goal "F : (UN i:I. A i) leadsTo B = (ALL i : I. F : (A i) leadsTo B)"; |
4776 | 173 |
by (blast_tac (claset() addIs [leadsTo_UN, leadsTo_weaken_L]) 1); |
174 |
qed "leadsTo_UN_distrib"; |
|
175 |
||
6536 | 176 |
Goal "F : (Union S) leadsTo B = (ALL A : S. F : A leadsTo B)"; |
4776 | 177 |
by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_L]) 1); |
178 |
qed "leadsTo_Union_distrib"; |
|
179 |
||
180 |
||
6536 | 181 |
Goal "[| F : A leadsTo A'; B<=A; A'<=B' |] ==> F : B leadsTo B'"; |
5340 | 182 |
by (blast_tac (claset() addIs [leadsTo_weaken_R, leadsTo_weaken_L, |
183 |
leadsTo_Trans]) 1); |
|
4776 | 184 |
qed "leadsTo_weaken"; |
185 |
||
186 |
||
187 |
(*Set difference: maybe combine with leadsTo_weaken_L??*) |
|
6536 | 188 |
Goal "[| F : (A-B) leadsTo C; F : B leadsTo C |] ==> F : A leadsTo C"; |
4776 | 189 |
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken]) 1); |
190 |
qed "leadsTo_Diff"; |
|
191 |
||
192 |
||
193 |
(** Meta or object quantifier ??? |
|
194 |
see ball_constrains_UN in UNITY.ML***) |
|
195 |
||
196 |
val prems = goal thy |
|
6536 | 197 |
"(!! i. i:I ==> F : (A i) leadsTo (A' i)) \ |
198 |
\ ==> F : (UN i:I. A i) leadsTo (UN i:I. A' i)"; |
|
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
199 |
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1); |
4776 | 200 |
by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_R] |
201 |
addIs prems) 1); |
|
202 |
qed "leadsTo_UN_UN"; |
|
203 |
||
204 |
||
205 |
(*Version with no index set*) |
|
206 |
val prems = goal thy |
|
6536 | 207 |
"(!! i. F : (A i) leadsTo (A' i)) \ |
208 |
\ ==> F : (UN i. A i) leadsTo (UN i. A' i)"; |
|
4776 | 209 |
by (blast_tac (claset() addIs [leadsTo_UN_UN] |
210 |
addIs prems) 1); |
|
211 |
qed "leadsTo_UN_UN_noindex"; |
|
212 |
||
213 |
(*Version with no index set*) |
|
6536 | 214 |
Goal "ALL i. F : (A i) leadsTo (A' i) \ |
215 |
\ ==> F : (UN i. A i) leadsTo (UN i. A' i)"; |
|
4776 | 216 |
by (blast_tac (claset() addIs [leadsTo_UN_UN]) 1); |
217 |
qed "all_leadsTo_UN_UN"; |
|
218 |
||
219 |
||
220 |
(*Binary union version*) |
|
6714 | 221 |
Goal "[| F : A leadsTo A'; F : B leadsTo B' |] \ |
222 |
\ ==> F : (A Un B) leadsTo (A' Un B')"; |
|
4776 | 223 |
by (blast_tac (claset() addIs [leadsTo_Un, |
224 |
leadsTo_weaken_R]) 1); |
|
225 |
qed "leadsTo_Un_Un"; |
|
226 |
||
227 |
||
228 |
(** The cancellation law **) |
|
229 |
||
6536 | 230 |
Goal "[| F : A leadsTo (A' Un B); F : B leadsTo B' |] \ |
6714 | 231 |
\ ==> F : A leadsTo (A' Un B')"; |
4776 | 232 |
by (blast_tac (claset() addIs [leadsTo_Un_Un, |
233 |
subset_imp_leadsTo, leadsTo_Trans]) 1); |
|
234 |
qed "leadsTo_cancel2"; |
|
235 |
||
6536 | 236 |
Goal "[| F : A leadsTo (A' Un B); F : (B-A') leadsTo B' |] \ |
6714 | 237 |
\ ==> F : A leadsTo (A' Un B')"; |
4776 | 238 |
by (rtac leadsTo_cancel2 1); |
239 |
by (assume_tac 2); |
|
240 |
by (ALLGOALS Asm_simp_tac); |
|
241 |
qed "leadsTo_cancel_Diff2"; |
|
242 |
||
6536 | 243 |
Goal "[| F : A leadsTo (B Un A'); F : B leadsTo B' |] \ |
244 |
\ ==> F : A leadsTo (B' Un A')"; |
|
4776 | 245 |
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1); |
246 |
by (blast_tac (claset() addSIs [leadsTo_cancel2]) 1); |
|
247 |
qed "leadsTo_cancel1"; |
|
248 |
||
6536 | 249 |
Goal "[| F : A leadsTo (B Un A'); F : (B-A') leadsTo B' |] \ |
250 |
\ ==> F : A leadsTo (B' Un A')"; |
|
4776 | 251 |
by (rtac leadsTo_cancel1 1); |
252 |
by (assume_tac 2); |
|
253 |
by (ALLGOALS Asm_simp_tac); |
|
254 |
qed "leadsTo_cancel_Diff1"; |
|
255 |
||
256 |
||
257 |
||
258 |
(** The impossibility law **) |
|
259 |
||
6536 | 260 |
Goal "F : A leadsTo B ==> B={} --> A={}"; |
4776 | 261 |
by (etac leadsTo_induct 1); |
262 |
by (ALLGOALS Asm_simp_tac); |
|
263 |
by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]); |
|
264 |
by (Blast_tac 1); |
|
265 |
val lemma = result() RS mp; |
|
266 |
||
6536 | 267 |
Goal "F : A leadsTo {} ==> A={}"; |
4776 | 268 |
by (blast_tac (claset() addSIs [lemma]) 1); |
269 |
qed "leadsTo_empty"; |
|
270 |
||
271 |
||
272 |
(** PSP: Progress-Safety-Progress **) |
|
273 |
||
5640 | 274 |
(*Special case of PSP: Misra's "stable conjunction"*) |
5069 | 275 |
Goalw [stable_def] |
6536 | 276 |
"[| F : A leadsTo A'; F : stable B |] \ |
277 |
\ ==> F : (A Int B) leadsTo (A' Int B)"; |
|
4776 | 278 |
by (etac leadsTo_induct 1); |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
279 |
by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3); |
4776 | 280 |
by (blast_tac (claset() addIs [leadsTo_Trans]) 2); |
281 |
by (rtac leadsTo_Basis 1); |
|
282 |
by (asm_full_simp_tac |
|
283 |
(simpset() addsimps [ensures_def, |
|
284 |
Diff_Int_distrib2 RS sym, Int_Un_distrib2 RS sym]) 1); |
|
285 |
by (blast_tac (claset() addIs [transient_strengthen, constrains_Int]) 1); |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
286 |
qed "psp_stable"; |
4776 | 287 |
|
7524 | 288 |
Goal |
289 |
"[| F : A leadsTo A'; F : stable B |] ==> F : (B Int A) leadsTo (B Int A')"; |
|
5536 | 290 |
by (asm_simp_tac (simpset() addsimps psp_stable::Int_ac) 1); |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
291 |
qed "psp_stable2"; |
4776 | 292 |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
293 |
Goalw [ensures_def, constrains_def] |
6536 | 294 |
"[| F : A ensures A'; F : B co B' |] \ |
6714 | 295 |
\ ==> F : (A Int B') ensures ((A' Int B) Un (B' - B))"; |
296 |
by (Clarify_tac 1); (*speeds up the proof*) |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
297 |
by (blast_tac (claset() addIs [transient_strengthen]) 1); |
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
298 |
qed "psp_ensures"; |
4776 | 299 |
|
6536 | 300 |
Goal "[| F : A leadsTo A'; F : B co B' |] \ |
6714 | 301 |
\ ==> F : (A Int B') leadsTo ((A' Int B) Un (B' - B))"; |
4776 | 302 |
by (etac leadsTo_induct 1); |
6295
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
paulson
parents:
6012
diff
changeset
|
303 |
by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3); |
4776 | 304 |
(*Transitivity case has a delicate argument involving "cancellation"*) |
305 |
by (rtac leadsTo_Un_duplicate2 2); |
|
306 |
by (etac leadsTo_cancel_Diff1 2); |
|
307 |
by (asm_full_simp_tac (simpset() addsimps [Int_Diff, Diff_triv]) 2); |
|
6714 | 308 |
by (blast_tac (claset() addIs [leadsTo_weaken_L] |
309 |
addDs [constrains_imp_subset]) 2); |
|
4776 | 310 |
(*Basis case*) |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
311 |
by (blast_tac (claset() addIs [leadsTo_Basis, psp_ensures]) 1); |
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
312 |
qed "psp"; |
4776 | 313 |
|
6536 | 314 |
Goal "[| F : A leadsTo A'; F : B co B' |] \ |
6714 | 315 |
\ ==> F : (B' Int A) leadsTo ((B Int A') Un (B' - B))"; |
5536 | 316 |
by (asm_simp_tac (simpset() addsimps psp::Int_ac) 1); |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
317 |
qed "psp2"; |
4776 | 318 |
|
319 |
||
5069 | 320 |
Goalw [unless_def] |
6536 | 321 |
"[| F : A leadsTo A'; F : B unless B' |] \ |
322 |
\ ==> F : (A Int B) leadsTo ((A' Int B) Un B')"; |
|
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
323 |
by (dtac psp 1); |
4776 | 324 |
by (assume_tac 1); |
6714 | 325 |
by (blast_tac (claset() addIs [leadsTo_weaken]) 1); |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
326 |
qed "psp_unless"; |
4776 | 327 |
|
328 |
||
329 |
(*** Proving the induction rules ***) |
|
330 |
||
5257 | 331 |
(** The most general rule: r is any wf relation; f is any variant function **) |
332 |
||
5239 | 333 |
Goal "[| wf r; \ |
6536 | 334 |
\ ALL m. F : (A Int f-``{m}) leadsTo \ |
7524 | 335 |
\ ((A Int f-``(r^-1 ^^ {m})) Un B) |] \ |
6536 | 336 |
\ ==> F : (A Int f-``{m}) leadsTo B"; |
4776 | 337 |
by (eres_inst_tac [("a","m")] wf_induct 1); |
6536 | 338 |
by (subgoal_tac "F : (A Int (f -`` (r^-1 ^^ {x}))) leadsTo B" 1); |
4776 | 339 |
by (stac vimage_eq_UN 2); |
340 |
by (asm_simp_tac (HOL_ss addsimps (UN_simps RL [sym])) 2); |
|
341 |
by (blast_tac (claset() addIs [leadsTo_UN]) 2); |
|
342 |
by (blast_tac (claset() addIs [leadsTo_cancel1, leadsTo_Un_duplicate]) 1); |
|
343 |
val lemma = result(); |
|
344 |
||
345 |
||
346 |
(** Meta or object quantifier ????? **) |
|
5239 | 347 |
Goal "[| wf r; \ |
6536 | 348 |
\ ALL m. F : (A Int f-``{m}) leadsTo \ |
7524 | 349 |
\ ((A Int f-``(r^-1 ^^ {m})) Un B) |] \ |
6536 | 350 |
\ ==> F : A leadsTo B"; |
4776 | 351 |
by (res_inst_tac [("t", "A")] subst 1); |
352 |
by (rtac leadsTo_UN 2); |
|
353 |
by (etac lemma 2); |
|
354 |
by (REPEAT (assume_tac 2)); |
|
355 |
by (Fast_tac 1); (*Blast_tac: Function unknown's argument not a parameter*) |
|
356 |
qed "leadsTo_wf_induct"; |
|
357 |
||
358 |
||
5239 | 359 |
Goal "[| wf r; \ |
6536 | 360 |
\ ALL m:I. F : (A Int f-``{m}) leadsTo \ |
7524 | 361 |
\ ((A Int f-``(r^-1 ^^ {m})) Un B) |] \ |
6536 | 362 |
\ ==> F : A leadsTo ((A - (f-``I)) Un B)"; |
4776 | 363 |
by (etac leadsTo_wf_induct 1); |
364 |
by Safe_tac; |
|
365 |
by (case_tac "m:I" 1); |
|
366 |
by (blast_tac (claset() addIs [leadsTo_weaken]) 1); |
|
367 |
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); |
|
368 |
qed "bounded_induct"; |
|
369 |
||
370 |
||
6536 | 371 |
(*Alternative proof is via the lemma F : (A Int f-``(lessThan m)) leadsTo B*) |
372 |
Goal "[| ALL m. F : (A Int f-``{m}) leadsTo \ |
|
7524 | 373 |
\ ((A Int f-``(lessThan m)) Un B) |] \ |
6536 | 374 |
\ ==> F : A leadsTo B"; |
4776 | 375 |
by (rtac (wf_less_than RS leadsTo_wf_induct) 1); |
376 |
by (Asm_simp_tac 1); |
|
377 |
qed "lessThan_induct"; |
|
378 |
||
7524 | 379 |
Goal "[| ALL m:(greaterThan l). \ |
380 |
\ F : (A Int f-``{m}) leadsTo ((A Int f-``(lessThan m)) Un B) |] \ |
|
6536 | 381 |
\ ==> F : A leadsTo ((A Int (f-``(atMost l))) Un B)"; |
5648 | 382 |
by (simp_tac (HOL_ss addsimps [Diff_eq RS sym, vimage_Compl, |
383 |
Compl_greaterThan RS sym]) 1); |
|
4776 | 384 |
by (rtac (wf_less_than RS bounded_induct) 1); |
385 |
by (Asm_simp_tac 1); |
|
386 |
qed "lessThan_bounded_induct"; |
|
387 |
||
7524 | 388 |
Goal "[| ALL m:(lessThan l). \ |
389 |
\ F : (A Int f-``{m}) leadsTo ((A Int f-``(greaterThan m)) Un B) |] \ |
|
6536 | 390 |
\ ==> F : A leadsTo ((A Int (f-``(atLeast l))) Un B)"; |
4776 | 391 |
by (res_inst_tac [("f","f"),("f1", "%k. l - k")] |
392 |
(wf_less_than RS wf_inv_image RS leadsTo_wf_induct) 1); |
|
393 |
by (simp_tac (simpset() addsimps [inv_image_def, Image_singleton]) 1); |
|
394 |
by (Clarify_tac 1); |
|
395 |
by (case_tac "m<l" 1); |
|
396 |
by (blast_tac (claset() addIs [not_leE, subset_imp_leadsTo]) 2); |
|
397 |
by (blast_tac (claset() addIs [leadsTo_weaken_R, diff_less_mono2]) 1); |
|
398 |
qed "greaterThan_bounded_induct"; |
|
399 |
||
400 |
||
401 |
(*** wlt ****) |
|
402 |
||
403 |
(*Misra's property W3*) |
|
6536 | 404 |
Goalw [wlt_def] "F : (wlt F B) leadsTo B"; |
4776 | 405 |
by (blast_tac (claset() addSIs [leadsTo_Union]) 1); |
406 |
qed "wlt_leadsTo"; |
|
407 |
||
6536 | 408 |
Goalw [wlt_def] "F : A leadsTo B ==> A <= wlt F B"; |
4776 | 409 |
by (blast_tac (claset() addSIs [leadsTo_Union]) 1); |
410 |
qed "leadsTo_subset"; |
|
411 |
||
412 |
(*Misra's property W2*) |
|
6536 | 413 |
Goal "F : A leadsTo B = (A <= wlt F B)"; |
4776 | 414 |
by (blast_tac (claset() addSIs [leadsTo_subset, |
415 |
wlt_leadsTo RS leadsTo_weaken_L]) 1); |
|
416 |
qed "leadsTo_eq_subset_wlt"; |
|
417 |
||
418 |
(*Misra's property W4*) |
|
5648 | 419 |
Goal "B <= wlt F B"; |
4776 | 420 |
by (asm_simp_tac (simpset() addsimps [leadsTo_eq_subset_wlt RS sym, |
421 |
subset_imp_leadsTo]) 1); |
|
422 |
qed "wlt_increasing"; |
|
423 |
||
424 |
||
425 |
(*Used in the Trans case below*) |
|
5069 | 426 |
Goalw [constrains_def] |
5111 | 427 |
"[| B <= A2; \ |
6536 | 428 |
\ F : (A1 - B) co (A1 Un B); \ |
429 |
\ F : (A2 - C) co (A2 Un C) |] \ |
|
430 |
\ ==> F : (A1 Un A2 - C) co (A1 Un A2 Un C)"; |
|
5669 | 431 |
by (Clarify_tac 1); |
5620 | 432 |
by (Blast_tac 1); |
4776 | 433 |
val lemma1 = result(); |
434 |
||
435 |
||
436 |
(*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*) |
|
6536 | 437 |
Goal "F : A leadsTo A' ==> \ |
438 |
\ EX B. A<=B & F : B leadsTo A' & F : (B-A') co (B Un A')"; |
|
4776 | 439 |
by (etac leadsTo_induct 1); |
440 |
(*Basis*) |
|
441 |
by (blast_tac (claset() addIs [leadsTo_Basis] |
|
442 |
addDs [ensuresD]) 1); |
|
443 |
(*Trans*) |
|
444 |
by (Clarify_tac 1); |
|
445 |
by (res_inst_tac [("x", "Ba Un Bb")] exI 1); |
|
446 |
by (blast_tac (claset() addIs [lemma1, leadsTo_Un_Un, leadsTo_cancel1, |
|
447 |
leadsTo_Un_duplicate]) 1); |
|
448 |
(*Union*) |
|
449 |
by (clarify_tac (claset() addSDs [ball_conj_distrib RS iffD1, |
|
450 |
bchoice, ball_constrains_UN]) 1);; |
|
451 |
by (res_inst_tac [("x", "UN A:S. f A")] exI 1); |
|
452 |
by (blast_tac (claset() addIs [leadsTo_UN, constrains_weaken]) 1); |
|
453 |
qed "leadsTo_123"; |
|
454 |
||
455 |
||
456 |
(*Misra's property W5*) |
|
6536 | 457 |
Goal "F : (wlt F B - B) co (wlt F B)"; |
5648 | 458 |
by (cut_inst_tac [("F","F")] (wlt_leadsTo RS leadsTo_123) 1); |
4776 | 459 |
by (Clarify_tac 1); |
5648 | 460 |
by (subgoal_tac "Ba = wlt F B" 1); |
461 |
by (blast_tac (claset() addDs [leadsTo_eq_subset_wlt RS iffD1]) 2); |
|
4776 | 462 |
by (Clarify_tac 1); |
463 |
by (asm_full_simp_tac (simpset() addsimps [wlt_increasing, Un_absorb2]) 1); |
|
464 |
qed "wlt_constrains_wlt"; |
|
465 |
||
466 |
||
467 |
(*** Completion: Binary and General Finite versions ***) |
|
468 |
||
6536 | 469 |
Goal "[| F : A leadsTo A'; F : stable A'; \ |
470 |
\ F : B leadsTo B'; F : stable B' |] \ |
|
471 |
\ ==> F : (A Int B) leadsTo (A' Int B')"; |
|
5648 | 472 |
by (subgoal_tac "F : stable (wlt F B')" 1); |
4776 | 473 |
by (asm_full_simp_tac (simpset() addsimps [stable_def]) 2); |
474 |
by (EVERY [etac (constrains_Un RS constrains_weaken) 2, |
|
5648 | 475 |
rtac wlt_constrains_wlt 2, |
4776 | 476 |
fast_tac (claset() addEs [wlt_increasing RSN (2,rev_subsetD)]) 3, |
477 |
Blast_tac 2]); |
|
6536 | 478 |
by (subgoal_tac "F : (A Int wlt F B') leadsTo (A' Int wlt F B')" 1); |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
479 |
by (blast_tac (claset() addIs [psp_stable]) 2); |
6536 | 480 |
by (subgoal_tac "F : (A' Int wlt F B') leadsTo (A' Int B')" 1); |
5277
e4297d03e5d2
A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents:
5257
diff
changeset
|
481 |
by (blast_tac (claset() addIs [wlt_leadsTo, psp_stable2]) 2); |
6536 | 482 |
by (subgoal_tac "F : (A Int B) leadsTo (A Int wlt F B')" 1); |
4776 | 483 |
by (blast_tac (claset() addIs [leadsTo_subset RS subsetD, |
484 |
subset_imp_leadsTo]) 2); |
|
5479 | 485 |
by (blast_tac (claset() addIs [leadsTo_Trans]) 1); |
4776 | 486 |
qed "stable_completion"; |
487 |
||
488 |
||
6536 | 489 |
Goal "finite I ==> (ALL i:I. F : (A i) leadsTo (A' i)) --> \ |
5648 | 490 |
\ (ALL i:I. F : stable (A' i)) --> \ |
6536 | 491 |
\ F : (INT i:I. A i) leadsTo (INT i:I. A' i)"; |
4776 | 492 |
by (etac finite_induct 1); |
493 |
by (Asm_simp_tac 1); |
|
494 |
by (asm_simp_tac |
|
495 |
(simpset() addsimps [stable_completion, stable_def, |
|
496 |
ball_constrains_INT]) 1); |
|
497 |
qed_spec_mp "finite_stable_completion"; |
|
498 |
||
499 |
||
5648 | 500 |
Goal "[| W = wlt F (B' Un C); \ |
6536 | 501 |
\ F : A leadsTo (A' Un C); F : A' co (A' Un C); \ |
502 |
\ F : B leadsTo (B' Un C); F : B' co (B' Un C) |] \ |
|
503 |
\ ==> F : (A Int B) leadsTo ((A' Int B') Un C)"; |
|
504 |
by (subgoal_tac "F : (W-C) co (W Un B' Un C)" 1); |
|
4776 | 505 |
by (blast_tac (claset() addIs [[asm_rl, wlt_constrains_wlt] |
506 |
MRS constrains_Un RS constrains_weaken]) 2); |
|
6536 | 507 |
by (subgoal_tac "F : (W-C) co W" 1); |
4776 | 508 |
by (asm_full_simp_tac |
509 |
(simpset() addsimps [wlt_increasing, Un_assoc, Un_absorb2]) 2); |
|
6536 | 510 |
by (subgoal_tac "F : (A Int W - C) leadsTo (A' Int W Un C)" 1); |
4776 | 511 |
by (simp_tac (simpset() addsimps [Int_Diff]) 2); |
6714 | 512 |
by (blast_tac (claset() addIs [wlt_leadsTo, psp RS leadsTo_weaken]) 2); |
5456 | 513 |
(** LEVEL 7 **) |
6536 | 514 |
by (subgoal_tac "F : (A' Int W Un C) leadsTo (A' Int B' Un C)" 1); |
6714 | 515 |
by (rtac leadsTo_Un_duplicate2 2); |
516 |
by (blast_tac (claset() addIs [leadsTo_Un_Un, |
|
517 |
wlt_leadsTo RS psp2 RS leadsTo_weaken, |
|
518 |
subset_refl RS subset_imp_leadsTo]) 2); |
|
4776 | 519 |
by (dtac leadsTo_Diff 1); |
520 |
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); |
|
521 |
by (subgoal_tac "A Int B <= A Int W" 1); |
|
5456 | 522 |
by (blast_tac (claset() addSDs [leadsTo_subset] |
523 |
addSIs [subset_refl RS Int_mono]) 2); |
|
4776 | 524 |
by (blast_tac (claset() addIs [leadsTo_Trans, subset_imp_leadsTo]) 1); |
525 |
bind_thm("completion", refl RS result()); |
|
526 |
||
527 |
||
6536 | 528 |
Goal "finite I ==> (ALL i:I. F : (A i) leadsTo (A' i Un C)) --> \ |
529 |
\ (ALL i:I. F : (A' i) co (A' i Un C)) --> \ |
|
530 |
\ F : (INT i:I. A i) leadsTo ((INT i:I. A' i) Un C)"; |
|
4776 | 531 |
by (etac finite_induct 1); |
532 |
by (ALLGOALS Asm_simp_tac); |
|
533 |
by (Clarify_tac 1); |
|
534 |
by (dtac ball_constrains_INT 1); |
|
535 |
by (asm_full_simp_tac (simpset() addsimps [completion]) 1); |
|
6564 | 536 |
qed_spec_mp "finite_completion"; |
4776 | 537 |