author | wenzelm |
Wed, 04 Nov 2015 23:27:00 +0100 | |
changeset 61578 | 6623c81cb15a |
parent 60773 | d09c66a0ea10 |
child 61824 | dcbe9f756ae0 |
permissions | -rw-r--r-- |
37936 | 1 |
(* Title: HOL/UNITY/SubstAx.thy |
4776 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1998 University of Cambridge |
|
4 |
||
6536 | 5 |
Weak LeadsTo relation (restricted to the set of reachable states) |
4776 | 6 |
*) |
7 |
||
58889 | 8 |
section{*Weak Progress*} |
13798 | 9 |
|
16417 | 10 |
theory SubstAx imports WFair Constrains begin |
4776 | 11 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
19769
diff
changeset
|
12 |
definition Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) where |
13805 | 13 |
"A Ensures B == {F. F \<in> (reachable F \<inter> A) ensures B}" |
8122
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
paulson
parents:
8041
diff
changeset
|
14 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
19769
diff
changeset
|
15 |
definition LeadsTo :: "['a set, 'a set] => 'a program set" (infixl "LeadsTo" 60) where |
13805 | 16 |
"A LeadsTo B == {F. F \<in> (reachable F \<inter> A) leadsTo B}" |
4776 | 17 |
|
60773 | 18 |
notation LeadsTo (infixl "\<longmapsto>w" 60) |
13796 | 19 |
|
20 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
21 |
text{*Resembles the previous definition of LeadsTo*} |
13796 | 22 |
lemma LeadsTo_eq_leadsTo: |
13805 | 23 |
"A LeadsTo B = {F. F \<in> (reachable F \<inter> A) leadsTo (reachable F \<inter> B)}" |
13796 | 24 |
apply (unfold LeadsTo_def) |
25 |
apply (blast dest: psp_stable2 intro: leadsTo_weaken) |
|
26 |
done |
|
27 |
||
28 |
||
13798 | 29 |
subsection{*Specialized laws for handling invariants*} |
13796 | 30 |
|
31 |
(** Conjoining an Always property **) |
|
32 |
||
33 |
lemma Always_LeadsTo_pre: |
|
13805 | 34 |
"F \<in> Always INV ==> (F \<in> (INV \<inter> A) LeadsTo A') = (F \<in> A LeadsTo A')" |
35 |
by (simp add: LeadsTo_def Always_eq_includes_reachable Int_absorb2 |
|
36 |
Int_assoc [symmetric]) |
|
13796 | 37 |
|
38 |
lemma Always_LeadsTo_post: |
|
13805 | 39 |
"F \<in> Always INV ==> (F \<in> A LeadsTo (INV \<inter> A')) = (F \<in> A LeadsTo A')" |
40 |
by (simp add: LeadsTo_eq_leadsTo Always_eq_includes_reachable Int_absorb2 |
|
41 |
Int_assoc [symmetric]) |
|
13796 | 42 |
|
13805 | 43 |
(* [| F \<in> Always C; F \<in> (C \<inter> A) LeadsTo A' |] ==> F \<in> A LeadsTo A' *) |
45605 | 44 |
lemmas Always_LeadsToI = Always_LeadsTo_pre [THEN iffD1] |
13796 | 45 |
|
13805 | 46 |
(* [| F \<in> Always INV; F \<in> A LeadsTo A' |] ==> F \<in> A LeadsTo (INV \<inter> A') *) |
45605 | 47 |
lemmas Always_LeadsToD = Always_LeadsTo_post [THEN iffD2] |
13796 | 48 |
|
49 |
||
13798 | 50 |
subsection{*Introduction rules: Basis, Trans, Union*} |
13796 | 51 |
|
13805 | 52 |
lemma leadsTo_imp_LeadsTo: "F \<in> A leadsTo B ==> F \<in> A LeadsTo B" |
13796 | 53 |
apply (simp add: LeadsTo_def) |
54 |
apply (blast intro: leadsTo_weaken_L) |
|
55 |
done |
|
56 |
||
57 |
lemma LeadsTo_Trans: |
|
13805 | 58 |
"[| F \<in> A LeadsTo B; F \<in> B LeadsTo C |] ==> F \<in> A LeadsTo C" |
13796 | 59 |
apply (simp add: LeadsTo_eq_leadsTo) |
60 |
apply (blast intro: leadsTo_Trans) |
|
61 |
done |
|
62 |
||
63 |
lemma LeadsTo_Union: |
|
13805 | 64 |
"(!!A. A \<in> S ==> F \<in> A LeadsTo B) ==> F \<in> (Union S) LeadsTo B" |
13796 | 65 |
apply (simp add: LeadsTo_def) |
66 |
apply (subst Int_Union) |
|
67 |
apply (blast intro: leadsTo_UN) |
|
68 |
done |
|
69 |
||
70 |
||
13798 | 71 |
subsection{*Derived rules*} |
13796 | 72 |
|
13805 | 73 |
lemma LeadsTo_UNIV [simp]: "F \<in> A LeadsTo UNIV" |
13796 | 74 |
by (simp add: LeadsTo_def) |
75 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
76 |
text{*Useful with cancellation, disjunction*} |
13796 | 77 |
lemma LeadsTo_Un_duplicate: |
13805 | 78 |
"F \<in> A LeadsTo (A' \<union> A') ==> F \<in> A LeadsTo A'" |
13796 | 79 |
by (simp add: Un_ac) |
80 |
||
81 |
lemma LeadsTo_Un_duplicate2: |
|
13805 | 82 |
"F \<in> A LeadsTo (A' \<union> C \<union> C) ==> F \<in> A LeadsTo (A' \<union> C)" |
13796 | 83 |
by (simp add: Un_ac) |
84 |
||
85 |
lemma LeadsTo_UN: |
|
13805 | 86 |
"(!!i. i \<in> I ==> F \<in> (A i) LeadsTo B) ==> F \<in> (\<Union>i \<in> I. A i) LeadsTo B" |
44106 | 87 |
apply (unfold SUP_def) |
13796 | 88 |
apply (blast intro: LeadsTo_Union) |
89 |
done |
|
90 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
91 |
text{*Binary union introduction rule*} |
13796 | 92 |
lemma LeadsTo_Un: |
13805 | 93 |
"[| F \<in> A LeadsTo C; F \<in> B LeadsTo C |] ==> F \<in> (A \<union> B) LeadsTo C" |
44106 | 94 |
using LeadsTo_UN [of "{A, B}" F id C] by auto |
13796 | 95 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
96 |
text{*Lets us look at the starting state*} |
13796 | 97 |
lemma single_LeadsTo_I: |
13805 | 98 |
"(!!s. s \<in> A ==> F \<in> {s} LeadsTo B) ==> F \<in> A LeadsTo B" |
13796 | 99 |
by (subst UN_singleton [symmetric], rule LeadsTo_UN, blast) |
100 |
||
13805 | 101 |
lemma subset_imp_LeadsTo: "A \<subseteq> B ==> F \<in> A LeadsTo B" |
13796 | 102 |
apply (simp add: LeadsTo_def) |
103 |
apply (blast intro: subset_imp_leadsTo) |
|
104 |
done |
|
105 |
||
45605 | 106 |
lemmas empty_LeadsTo = empty_subsetI [THEN subset_imp_LeadsTo, simp] |
13796 | 107 |
|
45477 | 108 |
lemma LeadsTo_weaken_R: |
13805 | 109 |
"[| F \<in> A LeadsTo A'; A' \<subseteq> B' |] ==> F \<in> A LeadsTo B'" |
110 |
apply (simp add: LeadsTo_def) |
|
13796 | 111 |
apply (blast intro: leadsTo_weaken_R) |
112 |
done |
|
113 |
||
45477 | 114 |
lemma LeadsTo_weaken_L: |
13805 | 115 |
"[| F \<in> A LeadsTo A'; B \<subseteq> A |] |
116 |
==> F \<in> B LeadsTo A'" |
|
117 |
apply (simp add: LeadsTo_def) |
|
13796 | 118 |
apply (blast intro: leadsTo_weaken_L) |
119 |
done |
|
120 |
||
121 |
lemma LeadsTo_weaken: |
|
13805 | 122 |
"[| F \<in> A LeadsTo A'; |
123 |
B \<subseteq> A; A' \<subseteq> B' |] |
|
124 |
==> F \<in> B LeadsTo B'" |
|
13796 | 125 |
by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans) |
126 |
||
127 |
lemma Always_LeadsTo_weaken: |
|
13805 | 128 |
"[| F \<in> Always C; F \<in> A LeadsTo A'; |
129 |
C \<inter> B \<subseteq> A; C \<inter> A' \<subseteq> B' |] |
|
130 |
==> F \<in> B LeadsTo B'" |
|
13796 | 131 |
by (blast dest: Always_LeadsToI intro: LeadsTo_weaken intro: Always_LeadsToD) |
132 |
||
133 |
(** Two theorems for "proof lattices" **) |
|
134 |
||
13805 | 135 |
lemma LeadsTo_Un_post: "F \<in> A LeadsTo B ==> F \<in> (A \<union> B) LeadsTo B" |
13796 | 136 |
by (blast intro: LeadsTo_Un subset_imp_LeadsTo) |
137 |
||
138 |
lemma LeadsTo_Trans_Un: |
|
13805 | 139 |
"[| F \<in> A LeadsTo B; F \<in> B LeadsTo C |] |
140 |
==> F \<in> (A \<union> B) LeadsTo C" |
|
13796 | 141 |
by (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans) |
142 |
||
143 |
||
144 |
(** Distributive laws **) |
|
145 |
||
146 |
lemma LeadsTo_Un_distrib: |
|
13805 | 147 |
"(F \<in> (A \<union> B) LeadsTo C) = (F \<in> A LeadsTo C & F \<in> B LeadsTo C)" |
13796 | 148 |
by (blast intro: LeadsTo_Un LeadsTo_weaken_L) |
149 |
||
150 |
lemma LeadsTo_UN_distrib: |
|
13805 | 151 |
"(F \<in> (\<Union>i \<in> I. A i) LeadsTo B) = (\<forall>i \<in> I. F \<in> (A i) LeadsTo B)" |
13796 | 152 |
by (blast intro: LeadsTo_UN LeadsTo_weaken_L) |
153 |
||
154 |
lemma LeadsTo_Union_distrib: |
|
13805 | 155 |
"(F \<in> (Union S) LeadsTo B) = (\<forall>A \<in> S. F \<in> A LeadsTo B)" |
13796 | 156 |
by (blast intro: LeadsTo_Union LeadsTo_weaken_L) |
157 |
||
158 |
||
159 |
(** More rules using the premise "Always INV" **) |
|
160 |
||
13805 | 161 |
lemma LeadsTo_Basis: "F \<in> A Ensures B ==> F \<in> A LeadsTo B" |
13796 | 162 |
by (simp add: Ensures_def LeadsTo_def leadsTo_Basis) |
163 |
||
164 |
lemma EnsuresI: |
|
13805 | 165 |
"[| F \<in> (A-B) Co (A \<union> B); F \<in> transient (A-B) |] |
166 |
==> F \<in> A Ensures B" |
|
13796 | 167 |
apply (simp add: Ensures_def Constrains_eq_constrains) |
168 |
apply (blast intro: ensuresI constrains_weaken transient_strengthen) |
|
169 |
done |
|
170 |
||
171 |
lemma Always_LeadsTo_Basis: |
|
13805 | 172 |
"[| F \<in> Always INV; |
173 |
F \<in> (INV \<inter> (A-A')) Co (A \<union> A'); |
|
174 |
F \<in> transient (INV \<inter> (A-A')) |] |
|
175 |
==> F \<in> A LeadsTo A'" |
|
13796 | 176 |
apply (rule Always_LeadsToI, assumption) |
177 |
apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen) |
|
178 |
done |
|
179 |
||
14150 | 180 |
text{*Set difference: maybe combine with @{text leadsTo_weaken_L}?? |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
181 |
This is the most useful form of the "disjunction" rule*} |
13796 | 182 |
lemma LeadsTo_Diff: |
13805 | 183 |
"[| F \<in> (A-B) LeadsTo C; F \<in> (A \<inter> B) LeadsTo C |] |
184 |
==> F \<in> A LeadsTo C" |
|
13796 | 185 |
by (blast intro: LeadsTo_Un LeadsTo_weaken) |
186 |
||
187 |
||
188 |
lemma LeadsTo_UN_UN: |
|
13805 | 189 |
"(!! i. i \<in> I ==> F \<in> (A i) LeadsTo (A' i)) |
190 |
==> F \<in> (\<Union>i \<in> I. A i) LeadsTo (\<Union>i \<in> I. A' i)" |
|
13796 | 191 |
apply (simp only: Union_image_eq [symmetric]) |
192 |
apply (blast intro: LeadsTo_Union LeadsTo_weaken_R) |
|
193 |
done |
|
194 |
||
195 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
196 |
text{*Version with no index set*} |
13796 | 197 |
lemma LeadsTo_UN_UN_noindex: |
13805 | 198 |
"(!!i. F \<in> (A i) LeadsTo (A' i)) ==> F \<in> (\<Union>i. A i) LeadsTo (\<Union>i. A' i)" |
13796 | 199 |
by (blast intro: LeadsTo_UN_UN) |
200 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
201 |
text{*Version with no index set*} |
13796 | 202 |
lemma all_LeadsTo_UN_UN: |
13805 | 203 |
"\<forall>i. F \<in> (A i) LeadsTo (A' i) |
204 |
==> F \<in> (\<Union>i. A i) LeadsTo (\<Union>i. A' i)" |
|
13796 | 205 |
by (blast intro: LeadsTo_UN_UN) |
206 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
207 |
text{*Binary union version*} |
13796 | 208 |
lemma LeadsTo_Un_Un: |
13805 | 209 |
"[| F \<in> A LeadsTo A'; F \<in> B LeadsTo B' |] |
210 |
==> F \<in> (A \<union> B) LeadsTo (A' \<union> B')" |
|
13796 | 211 |
by (blast intro: LeadsTo_Un LeadsTo_weaken_R) |
212 |
||
213 |
||
214 |
(** The cancellation law **) |
|
215 |
||
216 |
lemma LeadsTo_cancel2: |
|
13805 | 217 |
"[| F \<in> A LeadsTo (A' \<union> B); F \<in> B LeadsTo B' |] |
218 |
==> F \<in> A LeadsTo (A' \<union> B')" |
|
13796 | 219 |
by (blast intro: LeadsTo_Un_Un subset_imp_LeadsTo LeadsTo_Trans) |
220 |
||
221 |
lemma LeadsTo_cancel_Diff2: |
|
13805 | 222 |
"[| F \<in> A LeadsTo (A' \<union> B); F \<in> (B-A') LeadsTo B' |] |
223 |
==> F \<in> A LeadsTo (A' \<union> B')" |
|
13796 | 224 |
apply (rule LeadsTo_cancel2) |
225 |
prefer 2 apply assumption |
|
226 |
apply (simp_all (no_asm_simp)) |
|
227 |
done |
|
228 |
||
229 |
lemma LeadsTo_cancel1: |
|
13805 | 230 |
"[| F \<in> A LeadsTo (B \<union> A'); F \<in> B LeadsTo B' |] |
231 |
==> F \<in> A LeadsTo (B' \<union> A')" |
|
13796 | 232 |
apply (simp add: Un_commute) |
233 |
apply (blast intro!: LeadsTo_cancel2) |
|
234 |
done |
|
235 |
||
236 |
lemma LeadsTo_cancel_Diff1: |
|
13805 | 237 |
"[| F \<in> A LeadsTo (B \<union> A'); F \<in> (B-A') LeadsTo B' |] |
238 |
==> F \<in> A LeadsTo (B' \<union> A')" |
|
13796 | 239 |
apply (rule LeadsTo_cancel1) |
240 |
prefer 2 apply assumption |
|
241 |
apply (simp_all (no_asm_simp)) |
|
242 |
done |
|
243 |
||
244 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
245 |
text{*The impossibility law*} |
13796 | 246 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
247 |
text{*The set "A" may be non-empty, but it contains no reachable states*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
248 |
lemma LeadsTo_empty: "[|F \<in> A LeadsTo {}; all_total F|] ==> F \<in> Always (-A)" |
13805 | 249 |
apply (simp add: LeadsTo_def Always_eq_includes_reachable) |
13796 | 250 |
apply (drule leadsTo_empty, auto) |
251 |
done |
|
252 |
||
253 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
254 |
subsection{*PSP: Progress-Safety-Progress*} |
13796 | 255 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
256 |
text{*Special case of PSP: Misra's "stable conjunction"*} |
13796 | 257 |
lemma PSP_Stable: |
13805 | 258 |
"[| F \<in> A LeadsTo A'; F \<in> Stable B |] |
259 |
==> F \<in> (A \<inter> B) LeadsTo (A' \<inter> B)" |
|
260 |
apply (simp add: LeadsTo_eq_leadsTo Stable_eq_stable) |
|
13796 | 261 |
apply (drule psp_stable, assumption) |
262 |
apply (simp add: Int_ac) |
|
263 |
done |
|
264 |
||
265 |
lemma PSP_Stable2: |
|
13805 | 266 |
"[| F \<in> A LeadsTo A'; F \<in> Stable B |] |
267 |
==> F \<in> (B \<inter> A) LeadsTo (B \<inter> A')" |
|
13796 | 268 |
by (simp add: PSP_Stable Int_ac) |
269 |
||
270 |
lemma PSP: |
|
13805 | 271 |
"[| F \<in> A LeadsTo A'; F \<in> B Co B' |] |
272 |
==> F \<in> (A \<inter> B') LeadsTo ((A' \<inter> B) \<union> (B' - B))" |
|
273 |
apply (simp add: LeadsTo_def Constrains_eq_constrains) |
|
13796 | 274 |
apply (blast dest: psp intro: leadsTo_weaken) |
275 |
done |
|
276 |
||
277 |
lemma PSP2: |
|
13805 | 278 |
"[| F \<in> A LeadsTo A'; F \<in> B Co B' |] |
279 |
==> F \<in> (B' \<inter> A) LeadsTo ((B \<inter> A') \<union> (B' - B))" |
|
13796 | 280 |
by (simp add: PSP Int_ac) |
281 |
||
282 |
lemma PSP_Unless: |
|
13805 | 283 |
"[| F \<in> A LeadsTo A'; F \<in> B Unless B' |] |
284 |
==> F \<in> (A \<inter> B) LeadsTo ((A' \<inter> B) \<union> B')" |
|
13796 | 285 |
apply (unfold Unless_def) |
286 |
apply (drule PSP, assumption) |
|
287 |
apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo) |
|
288 |
done |
|
289 |
||
290 |
||
291 |
lemma Stable_transient_Always_LeadsTo: |
|
13805 | 292 |
"[| F \<in> Stable A; F \<in> transient C; |
293 |
F \<in> Always (-A \<union> B \<union> C) |] ==> F \<in> A LeadsTo B" |
|
13796 | 294 |
apply (erule Always_LeadsTo_weaken) |
295 |
apply (rule LeadsTo_Diff) |
|
296 |
prefer 2 |
|
297 |
apply (erule |
|
298 |
transient_imp_leadsTo [THEN leadsTo_imp_LeadsTo, THEN PSP_Stable2]) |
|
299 |
apply (blast intro: subset_imp_LeadsTo)+ |
|
300 |
done |
|
301 |
||
302 |
||
13798 | 303 |
subsection{*Induction rules*} |
13796 | 304 |
|
305 |
(** Meta or object quantifier ????? **) |
|
306 |
lemma LeadsTo_wf_induct: |
|
307 |
"[| wf r; |
|
13805 | 308 |
\<forall>m. F \<in> (A \<inter> f-`{m}) LeadsTo |
309 |
((A \<inter> f-`(r^-1 `` {m})) \<union> B) |] |
|
310 |
==> F \<in> A LeadsTo B" |
|
311 |
apply (simp add: LeadsTo_eq_leadsTo) |
|
13796 | 312 |
apply (erule leadsTo_wf_induct) |
313 |
apply (blast intro: leadsTo_weaken) |
|
314 |
done |
|
315 |
||
316 |
||
317 |
lemma Bounded_induct: |
|
318 |
"[| wf r; |
|
13805 | 319 |
\<forall>m \<in> I. F \<in> (A \<inter> f-`{m}) LeadsTo |
320 |
((A \<inter> f-`(r^-1 `` {m})) \<union> B) |] |
|
321 |
==> F \<in> A LeadsTo ((A - (f-`I)) \<union> B)" |
|
13796 | 322 |
apply (erule LeadsTo_wf_induct, safe) |
13805 | 323 |
apply (case_tac "m \<in> I") |
13796 | 324 |
apply (blast intro: LeadsTo_weaken) |
325 |
apply (blast intro: subset_imp_LeadsTo) |
|
326 |
done |
|
327 |
||
328 |
||
329 |
lemma LessThan_induct: |
|
13805 | 330 |
"(!!m::nat. F \<in> (A \<inter> f-`{m}) LeadsTo ((A \<inter> f-`(lessThan m)) \<union> B)) |
331 |
==> F \<in> A LeadsTo B" |
|
332 |
by (rule wf_less_than [THEN LeadsTo_wf_induct], auto) |
|
13796 | 333 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
334 |
text{*Integer version. Could generalize from 0 to any lower bound*} |
13796 | 335 |
lemma integ_0_le_induct: |
13805 | 336 |
"[| F \<in> Always {s. (0::int) \<le> f s}; |
337 |
!! z. F \<in> (A \<inter> {s. f s = z}) LeadsTo |
|
338 |
((A \<inter> {s. f s < z}) \<union> B) |] |
|
339 |
==> F \<in> A LeadsTo B" |
|
13796 | 340 |
apply (rule_tac f = "nat o f" in LessThan_induct) |
341 |
apply (simp add: vimage_def) |
|
342 |
apply (rule Always_LeadsTo_weaken, assumption+) |
|
343 |
apply (auto simp add: nat_eq_iff nat_less_iff) |
|
344 |
done |
|
345 |
||
346 |
lemma LessThan_bounded_induct: |
|
13805 | 347 |
"!!l::nat. \<forall>m \<in> greaterThan l. |
348 |
F \<in> (A \<inter> f-`{m}) LeadsTo ((A \<inter> f-`(lessThan m)) \<union> B) |
|
349 |
==> F \<in> A LeadsTo ((A \<inter> (f-`(atMost l))) \<union> B)" |
|
350 |
apply (simp only: Diff_eq [symmetric] vimage_Compl |
|
351 |
Compl_greaterThan [symmetric]) |
|
352 |
apply (rule wf_less_than [THEN Bounded_induct], simp) |
|
13796 | 353 |
done |
354 |
||
355 |
lemma GreaterThan_bounded_induct: |
|
13805 | 356 |
"!!l::nat. \<forall>m \<in> lessThan l. |
357 |
F \<in> (A \<inter> f-`{m}) LeadsTo ((A \<inter> f-`(greaterThan m)) \<union> B) |
|
358 |
==> F \<in> A LeadsTo ((A \<inter> (f-`(atLeast l))) \<union> B)" |
|
13796 | 359 |
apply (rule_tac f = f and f1 = "%k. l - k" |
360 |
in wf_less_than [THEN wf_inv_image, THEN LeadsTo_wf_induct]) |
|
19769
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents:
16417
diff
changeset
|
361 |
apply (simp add: Image_singleton, clarify) |
13796 | 362 |
apply (case_tac "m<l") |
13805 | 363 |
apply (blast intro: LeadsTo_weaken_R diff_less_mono2) |
364 |
apply (blast intro: not_leE subset_imp_LeadsTo) |
|
13796 | 365 |
done |
366 |
||
367 |
||
13798 | 368 |
subsection{*Completion: Binary and General Finite versions*} |
13796 | 369 |
|
370 |
lemma Completion: |
|
13805 | 371 |
"[| F \<in> A LeadsTo (A' \<union> C); F \<in> A' Co (A' \<union> C); |
372 |
F \<in> B LeadsTo (B' \<union> C); F \<in> B' Co (B' \<union> C) |] |
|
373 |
==> F \<in> (A \<inter> B) LeadsTo ((A' \<inter> B') \<union> C)" |
|
374 |
apply (simp add: LeadsTo_eq_leadsTo Constrains_eq_constrains Int_Un_distrib) |
|
13796 | 375 |
apply (blast intro: completion leadsTo_weaken) |
376 |
done |
|
377 |
||
378 |
lemma Finite_completion_lemma: |
|
379 |
"finite I |
|
13805 | 380 |
==> (\<forall>i \<in> I. F \<in> (A i) LeadsTo (A' i \<union> C)) --> |
381 |
(\<forall>i \<in> I. F \<in> (A' i) Co (A' i \<union> C)) --> |
|
382 |
F \<in> (\<Inter>i \<in> I. A i) LeadsTo ((\<Inter>i \<in> I. A' i) \<union> C)" |
|
13796 | 383 |
apply (erule finite_induct, auto) |
384 |
apply (rule Completion) |
|
385 |
prefer 4 |
|
386 |
apply (simp only: INT_simps [symmetric]) |
|
387 |
apply (rule Constrains_INT, auto) |
|
388 |
done |
|
389 |
||
390 |
lemma Finite_completion: |
|
391 |
"[| finite I; |
|
13805 | 392 |
!!i. i \<in> I ==> F \<in> (A i) LeadsTo (A' i \<union> C); |
393 |
!!i. i \<in> I ==> F \<in> (A' i) Co (A' i \<union> C) |] |
|
394 |
==> F \<in> (\<Inter>i \<in> I. A i) LeadsTo ((\<Inter>i \<in> I. A' i) \<union> C)" |
|
13796 | 395 |
by (blast intro: Finite_completion_lemma [THEN mp, THEN mp]) |
396 |
||
397 |
lemma Stable_completion: |
|
13805 | 398 |
"[| F \<in> A LeadsTo A'; F \<in> Stable A'; |
399 |
F \<in> B LeadsTo B'; F \<in> Stable B' |] |
|
400 |
==> F \<in> (A \<inter> B) LeadsTo (A' \<inter> B')" |
|
13796 | 401 |
apply (unfold Stable_def) |
402 |
apply (rule_tac C1 = "{}" in Completion [THEN LeadsTo_weaken_R]) |
|
403 |
apply (force+) |
|
404 |
done |
|
405 |
||
406 |
lemma Finite_stable_completion: |
|
407 |
"[| finite I; |
|
13805 | 408 |
!!i. i \<in> I ==> F \<in> (A i) LeadsTo (A' i); |
409 |
!!i. i \<in> I ==> F \<in> Stable (A' i) |] |
|
410 |
==> F \<in> (\<Inter>i \<in> I. A i) LeadsTo (\<Inter>i \<in> I. A' i)" |
|
13796 | 411 |
apply (unfold Stable_def) |
412 |
apply (rule_tac C1 = "{}" in Finite_completion [THEN LeadsTo_weaken_R]) |
|
13805 | 413 |
apply (simp_all, blast+) |
13796 | 414 |
done |
415 |
||
4776 | 416 |
end |