author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 20 Apr 2010 14:56:20 +0200 | |
changeset 36215 | 88ff48884d26 |
parent 35422 | e74b6f3b950c |
child 37936 | 1e4c5015a72e |
permissions | -rw-r--r-- |
4776 | 1 |
(* Title: HOL/UNITY/WFair |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
Copyright 1998 University of Cambridge |
|
4 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
5 |
Conditional Fairness versions of transient, ensures, leadsTo. |
4776 | 6 |
|
7 |
From Misra, "A Logic for Concurrent Programming", 1994 |
|
8 |
*) |
|
9 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
10 |
header{*Progress*} |
13798 | 11 |
|
16417 | 12 |
theory WFair imports UNITY begin |
4776 | 13 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
14 |
text{*The original version of this theory was based on weak fairness. (Thus, |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
15 |
the entire UNITY development embodied this assumption, until February 2003.) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
16 |
Weak fairness states that if a command is enabled continuously, then it is |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
17 |
eventually executed. Ernie Cohen suggested that I instead adopt unconditional |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
18 |
fairness: every command is executed infinitely often. |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
19 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
20 |
In fact, Misra's paper on "Progress" seems to be ambiguous about the correct |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
21 |
interpretation, and says that the two forms of fairness are equivalent. They |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
22 |
differ only on their treatment of partial transitions, which under |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
23 |
unconditional fairness behave magically. That is because if there are partial |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
24 |
transitions then there may be no fair executions, making all leads-to |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
25 |
properties hold vacuously. |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
26 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
27 |
Unconditional fairness has some great advantages. By distinguishing partial |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
28 |
transitions from total ones that are the identity on part of their domain, it |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
29 |
is more expressive. Also, by simplifying the definition of the transient |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
30 |
property, it simplifies many proofs. A drawback is that some laws only hold |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
31 |
under the assumption that all transitions are total. The best-known of these |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
32 |
is the impossibility law for leads-to. |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
33 |
*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
34 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32693
diff
changeset
|
35 |
definition |
4776 | 36 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
37 |
--{*This definition specifies conditional fairness. The rest of the theory |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
38 |
is generic to all forms of fairness. To get weak fairness, conjoin |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
39 |
the inclusion below with @{term "A \<subseteq> Domain act"}, which specifies |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
40 |
that the action is enabled over all of @{term A}.*} |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32693
diff
changeset
|
41 |
transient :: "'a set => 'a program set" where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
42 |
"transient A == {F. \<exists>act\<in>Acts F. act``A \<subseteq> -A}" |
4776 | 43 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32693
diff
changeset
|
44 |
definition |
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32693
diff
changeset
|
45 |
ensures :: "['a set, 'a set] => 'a program set" (infixl "ensures" 60) where |
13805 | 46 |
"A ensures B == (A-B co A \<union> B) \<inter> transient (A-B)" |
8006 | 47 |
|
6536 | 48 |
|
23767 | 49 |
inductive_set |
6801
9e0037839d63
renamed the underlying relation of leadsTo from "leadsto"
paulson
parents:
6536
diff
changeset
|
50 |
leads :: "'a program => ('a set * 'a set) set" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
51 |
--{*LEADS-TO constant for the inductive definition*} |
23767 | 52 |
for F :: "'a program" |
53 |
where |
|
4776 | 54 |
|
13805 | 55 |
Basis: "F \<in> A ensures B ==> (A,B) \<in> leads F" |
4776 | 56 |
|
23767 | 57 |
| Trans: "[| (A,B) \<in> leads F; (B,C) \<in> leads F |] ==> (A,C) \<in> leads F" |
4776 | 58 |
|
23767 | 59 |
| Union: "\<forall>A \<in> S. (A,B) \<in> leads F ==> (Union S, B) \<in> leads F" |
4776 | 60 |
|
5155 | 61 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32693
diff
changeset
|
62 |
definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
63 |
--{*visible version of the LEADS-TO relation*} |
13805 | 64 |
"A leadsTo B == {F. (A,B) \<in> leads F}" |
5648 | 65 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32693
diff
changeset
|
66 |
definition wlt :: "['a program, 'a set] => 'a set" where |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
67 |
--{*predicate transformer: the largest set that leads to @{term B}*} |
13805 | 68 |
"wlt F B == Union {A. F \<in> A leadsTo B}" |
4776 | 69 |
|
35355
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
32693
diff
changeset
|
70 |
notation (xsymbols) |
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
wenzelm
parents:
32693
diff
changeset
|
71 |
leadsTo (infixl "\<longmapsto>" 60) |
13797 | 72 |
|
73 |
||
13798 | 74 |
subsection{*transient*} |
13797 | 75 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
76 |
lemma stable_transient: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
77 |
"[| F \<in> stable A; F \<in> transient A |] ==> \<exists>act\<in>Acts F. A \<subseteq> - (Domain act)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
78 |
apply (simp add: stable_def constrains_def transient_def, clarify) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
79 |
apply (rule rev_bexI, auto) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
80 |
done |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
81 |
|
13797 | 82 |
lemma stable_transient_empty: |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
83 |
"[| F \<in> stable A; F \<in> transient A; all_total F |] ==> A = {}" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
84 |
apply (drule stable_transient, assumption) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
85 |
apply (simp add: all_total_def) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
86 |
done |
13797 | 87 |
|
88 |
lemma transient_strengthen: |
|
13805 | 89 |
"[| F \<in> transient A; B \<subseteq> A |] ==> F \<in> transient B" |
13797 | 90 |
apply (unfold transient_def, clarify) |
91 |
apply (blast intro!: rev_bexI) |
|
92 |
done |
|
93 |
||
94 |
lemma transientI: |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
95 |
"[| act: Acts F; act``A \<subseteq> -A |] ==> F \<in> transient A" |
13797 | 96 |
by (unfold transient_def, blast) |
97 |
||
98 |
lemma transientE: |
|
13805 | 99 |
"[| F \<in> transient A; |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
100 |
!!act. [| act: Acts F; act``A \<subseteq> -A |] ==> P |] |
13797 | 101 |
==> P" |
102 |
by (unfold transient_def, blast) |
|
103 |
||
104 |
lemma transient_empty [simp]: "transient {} = UNIV" |
|
105 |
by (unfold transient_def, auto) |
|
106 |
||
107 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
108 |
text{*This equation recovers the notion of weak fairness. A totalized |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
109 |
program satisfies a transient assertion just if the original program |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
110 |
contains a suitable action that is also enabled.*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
111 |
lemma totalize_transient_iff: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
112 |
"(totalize F \<in> transient A) = (\<exists>act\<in>Acts F. A \<subseteq> Domain act & act``A \<subseteq> -A)" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
113 |
apply (simp add: totalize_def totalize_act_def transient_def |
32693 | 114 |
Un_Image, safe) |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
115 |
apply (blast intro!: rev_bexI)+ |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
116 |
done |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
117 |
|
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
118 |
lemma totalize_transientI: |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
119 |
"[| act: Acts F; A \<subseteq> Domain act; act``A \<subseteq> -A |] |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
120 |
==> totalize F \<in> transient A" |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
121 |
by (simp add: totalize_transient_iff, blast) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
122 |
|
13798 | 123 |
subsection{*ensures*} |
13797 | 124 |
|
125 |
lemma ensuresI: |
|
13805 | 126 |
"[| F \<in> (A-B) co (A \<union> B); F \<in> transient (A-B) |] ==> F \<in> A ensures B" |
13797 | 127 |
by (unfold ensures_def, blast) |
128 |
||
129 |
lemma ensuresD: |
|
13805 | 130 |
"F \<in> A ensures B ==> F \<in> (A-B) co (A \<union> B) & F \<in> transient (A-B)" |
13797 | 131 |
by (unfold ensures_def, blast) |
132 |
||
133 |
lemma ensures_weaken_R: |
|
13805 | 134 |
"[| F \<in> A ensures A'; A'<=B' |] ==> F \<in> A ensures B'" |
13797 | 135 |
apply (unfold ensures_def) |
136 |
apply (blast intro: constrains_weaken transient_strengthen) |
|
137 |
done |
|
138 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
139 |
text{*The L-version (precondition strengthening) fails, but we have this*} |
13797 | 140 |
lemma stable_ensures_Int: |
13805 | 141 |
"[| F \<in> stable C; F \<in> A ensures B |] |
142 |
==> F \<in> (C \<inter> A) ensures (C \<inter> B)" |
|
13797 | 143 |
apply (unfold ensures_def) |
144 |
apply (auto simp add: ensures_def Int_Un_distrib [symmetric] Diff_Int_distrib [symmetric]) |
|
145 |
prefer 2 apply (blast intro: transient_strengthen) |
|
146 |
apply (blast intro: stable_constrains_Int constrains_weaken) |
|
147 |
done |
|
148 |
||
149 |
lemma stable_transient_ensures: |
|
13805 | 150 |
"[| F \<in> stable A; F \<in> transient C; A \<subseteq> B \<union> C |] ==> F \<in> A ensures B" |
13797 | 151 |
apply (simp add: ensures_def stable_def) |
152 |
apply (blast intro: constrains_weaken transient_strengthen) |
|
153 |
done |
|
154 |
||
13805 | 155 |
lemma ensures_eq: "(A ensures B) = (A unless B) \<inter> transient (A-B)" |
13797 | 156 |
by (simp (no_asm) add: ensures_def unless_def) |
157 |
||
158 |
||
13798 | 159 |
subsection{*leadsTo*} |
13797 | 160 |
|
13805 | 161 |
lemma leadsTo_Basis [intro]: "F \<in> A ensures B ==> F \<in> A leadsTo B" |
13797 | 162 |
apply (unfold leadsTo_def) |
163 |
apply (blast intro: leads.Basis) |
|
164 |
done |
|
165 |
||
166 |
lemma leadsTo_Trans: |
|
13805 | 167 |
"[| F \<in> A leadsTo B; F \<in> B leadsTo C |] ==> F \<in> A leadsTo C" |
13797 | 168 |
apply (unfold leadsTo_def) |
169 |
apply (blast intro: leads.Trans) |
|
170 |
done |
|
171 |
||
14112 | 172 |
lemma leadsTo_Basis': |
173 |
"[| F \<in> A co A \<union> B; F \<in> transient A |] ==> F \<in> A leadsTo B" |
|
174 |
apply (drule_tac B = "A-B" in constrains_weaken_L) |
|
175 |
apply (drule_tac [2] B = "A-B" in transient_strengthen) |
|
176 |
apply (rule_tac [3] ensuresI [THEN leadsTo_Basis]) |
|
177 |
apply (blast+) |
|
178 |
done |
|
179 |
||
13805 | 180 |
lemma transient_imp_leadsTo: "F \<in> transient A ==> F \<in> A leadsTo (-A)" |
13797 | 181 |
by (simp (no_asm_simp) add: leadsTo_Basis ensuresI Compl_partition) |
182 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
183 |
text{*Useful with cancellation, disjunction*} |
13805 | 184 |
lemma leadsTo_Un_duplicate: "F \<in> A leadsTo (A' \<union> A') ==> F \<in> A leadsTo A'" |
13797 | 185 |
by (simp add: Un_ac) |
186 |
||
187 |
lemma leadsTo_Un_duplicate2: |
|
13805 | 188 |
"F \<in> A leadsTo (A' \<union> C \<union> C) ==> F \<in> A leadsTo (A' \<union> C)" |
13797 | 189 |
by (simp add: Un_ac) |
190 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
191 |
text{*The Union introduction rule as we should have liked to state it*} |
13797 | 192 |
lemma leadsTo_Union: |
13805 | 193 |
"(!!A. A \<in> S ==> F \<in> A leadsTo B) ==> F \<in> (Union S) leadsTo B" |
13797 | 194 |
apply (unfold leadsTo_def) |
195 |
apply (blast intro: leads.Union) |
|
196 |
done |
|
197 |
||
198 |
lemma leadsTo_Union_Int: |
|
13805 | 199 |
"(!!A. A \<in> S ==> F \<in> (A \<inter> C) leadsTo B) ==> F \<in> (Union S \<inter> C) leadsTo B" |
13797 | 200 |
apply (unfold leadsTo_def) |
201 |
apply (simp only: Int_Union_Union) |
|
202 |
apply (blast intro: leads.Union) |
|
203 |
done |
|
204 |
||
205 |
lemma leadsTo_UN: |
|
13805 | 206 |
"(!!i. i \<in> I ==> F \<in> (A i) leadsTo B) ==> F \<in> (\<Union>i \<in> I. A i) leadsTo B" |
13797 | 207 |
apply (subst Union_image_eq [symmetric]) |
208 |
apply (blast intro: leadsTo_Union) |
|
209 |
done |
|
210 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
211 |
text{*Binary union introduction rule*} |
13797 | 212 |
lemma leadsTo_Un: |
13805 | 213 |
"[| F \<in> A leadsTo C; F \<in> B leadsTo C |] ==> F \<in> (A \<union> B) leadsTo C" |
13797 | 214 |
apply (subst Un_eq_Union) |
215 |
apply (blast intro: leadsTo_Union) |
|
216 |
done |
|
217 |
||
218 |
lemma single_leadsTo_I: |
|
13805 | 219 |
"(!!x. x \<in> A ==> F \<in> {x} leadsTo B) ==> F \<in> A leadsTo B" |
13797 | 220 |
by (subst UN_singleton [symmetric], rule leadsTo_UN, blast) |
221 |
||
222 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
223 |
text{*The INDUCTION rule as we should have liked to state it*} |
13797 | 224 |
lemma leadsTo_induct: |
13805 | 225 |
"[| F \<in> za leadsTo zb; |
226 |
!!A B. F \<in> A ensures B ==> P A B; |
|
227 |
!!A B C. [| F \<in> A leadsTo B; P A B; F \<in> B leadsTo C; P B C |] |
|
13797 | 228 |
==> P A C; |
13805 | 229 |
!!B S. \<forall>A \<in> S. F \<in> A leadsTo B & P A B ==> P (Union S) B |
13797 | 230 |
|] ==> P za zb" |
231 |
apply (unfold leadsTo_def) |
|
232 |
apply (drule CollectD, erule leads.induct) |
|
233 |
apply (blast+) |
|
234 |
done |
|
235 |
||
236 |
||
13805 | 237 |
lemma subset_imp_ensures: "A \<subseteq> B ==> F \<in> A ensures B" |
13797 | 238 |
by (unfold ensures_def constrains_def transient_def, blast) |
239 |
||
240 |
lemmas subset_imp_leadsTo = subset_imp_ensures [THEN leadsTo_Basis, standard] |
|
241 |
||
242 |
lemmas leadsTo_refl = subset_refl [THEN subset_imp_leadsTo, standard] |
|
243 |
||
244 |
lemmas empty_leadsTo = empty_subsetI [THEN subset_imp_leadsTo, standard, simp] |
|
245 |
||
246 |
lemmas leadsTo_UNIV = subset_UNIV [THEN subset_imp_leadsTo, standard, simp] |
|
247 |
||
248 |
||
249 |
||
250 |
(** Variant induction rule: on the preconditions for B **) |
|
251 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
252 |
text{*Lemma is the weak version: can't see how to do it in one step*} |
13797 | 253 |
lemma leadsTo_induct_pre_lemma: |
13805 | 254 |
"[| F \<in> za leadsTo zb; |
13797 | 255 |
P zb; |
13805 | 256 |
!!A B. [| F \<in> A ensures B; P B |] ==> P A; |
257 |
!!S. \<forall>A \<in> S. P A ==> P (Union S) |
|
13797 | 258 |
|] ==> P za" |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
259 |
txt{*by induction on this formula*} |
13797 | 260 |
apply (subgoal_tac "P zb --> P za") |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
261 |
txt{*now solve first subgoal: this formula is sufficient*} |
13797 | 262 |
apply (blast intro: leadsTo_refl) |
263 |
apply (erule leadsTo_induct) |
|
264 |
apply (blast+) |
|
265 |
done |
|
266 |
||
267 |
lemma leadsTo_induct_pre: |
|
13805 | 268 |
"[| F \<in> za leadsTo zb; |
13797 | 269 |
P zb; |
13805 | 270 |
!!A B. [| F \<in> A ensures B; F \<in> B leadsTo zb; P B |] ==> P A; |
271 |
!!S. \<forall>A \<in> S. F \<in> A leadsTo zb & P A ==> P (Union S) |
|
13797 | 272 |
|] ==> P za" |
13805 | 273 |
apply (subgoal_tac "F \<in> za leadsTo zb & P za") |
13797 | 274 |
apply (erule conjunct2) |
275 |
apply (erule leadsTo_induct_pre_lemma) |
|
276 |
prefer 3 apply (blast intro: leadsTo_Union) |
|
277 |
prefer 2 apply (blast intro: leadsTo_Trans) |
|
278 |
apply (blast intro: leadsTo_refl) |
|
279 |
done |
|
280 |
||
281 |
||
13805 | 282 |
lemma leadsTo_weaken_R: "[| F \<in> A leadsTo A'; A'<=B' |] ==> F \<in> A leadsTo B'" |
13797 | 283 |
by (blast intro: subset_imp_leadsTo leadsTo_Trans) |
284 |
||
13798 | 285 |
lemma leadsTo_weaken_L [rule_format]: |
13805 | 286 |
"[| F \<in> A leadsTo A'; B \<subseteq> A |] ==> F \<in> B leadsTo A'" |
13797 | 287 |
by (blast intro: leadsTo_Trans subset_imp_leadsTo) |
288 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
289 |
text{*Distributes over binary unions*} |
13797 | 290 |
lemma leadsTo_Un_distrib: |
13805 | 291 |
"F \<in> (A \<union> B) leadsTo C = (F \<in> A leadsTo C & F \<in> B leadsTo C)" |
13797 | 292 |
by (blast intro: leadsTo_Un leadsTo_weaken_L) |
293 |
||
294 |
lemma leadsTo_UN_distrib: |
|
13805 | 295 |
"F \<in> (\<Union>i \<in> I. A i) leadsTo B = (\<forall>i \<in> I. F \<in> (A i) leadsTo B)" |
13797 | 296 |
by (blast intro: leadsTo_UN leadsTo_weaken_L) |
297 |
||
298 |
lemma leadsTo_Union_distrib: |
|
13805 | 299 |
"F \<in> (Union S) leadsTo B = (\<forall>A \<in> S. F \<in> A leadsTo B)" |
13797 | 300 |
by (blast intro: leadsTo_Union leadsTo_weaken_L) |
301 |
||
302 |
||
303 |
lemma leadsTo_weaken: |
|
13805 | 304 |
"[| F \<in> A leadsTo A'; B \<subseteq> A; A'<=B' |] ==> F \<in> B leadsTo B'" |
13797 | 305 |
by (blast intro: leadsTo_weaken_R leadsTo_weaken_L leadsTo_Trans) |
306 |
||
307 |
||
14150 | 308 |
text{*Set difference: maybe combine with @{text leadsTo_weaken_L}??*} |
13797 | 309 |
lemma leadsTo_Diff: |
13805 | 310 |
"[| F \<in> (A-B) leadsTo C; F \<in> B leadsTo C |] ==> F \<in> A leadsTo C" |
13797 | 311 |
by (blast intro: leadsTo_Un leadsTo_weaken) |
312 |
||
313 |
lemma leadsTo_UN_UN: |
|
13805 | 314 |
"(!! i. i \<in> I ==> F \<in> (A i) leadsTo (A' i)) |
315 |
==> F \<in> (\<Union>i \<in> I. A i) leadsTo (\<Union>i \<in> I. A' i)" |
|
13797 | 316 |
apply (simp only: Union_image_eq [symmetric]) |
317 |
apply (blast intro: leadsTo_Union leadsTo_weaken_R) |
|
318 |
done |
|
319 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
320 |
text{*Binary union version*} |
13797 | 321 |
lemma leadsTo_Un_Un: |
13805 | 322 |
"[| F \<in> A leadsTo A'; F \<in> B leadsTo B' |] |
323 |
==> F \<in> (A \<union> B) leadsTo (A' \<union> B')" |
|
13797 | 324 |
by (blast intro: leadsTo_Un leadsTo_weaken_R) |
325 |
||
326 |
||
327 |
(** The cancellation law **) |
|
328 |
||
329 |
lemma leadsTo_cancel2: |
|
13805 | 330 |
"[| F \<in> A leadsTo (A' \<union> B); F \<in> B leadsTo B' |] |
331 |
==> F \<in> A leadsTo (A' \<union> B')" |
|
13797 | 332 |
by (blast intro: leadsTo_Un_Un subset_imp_leadsTo leadsTo_Trans) |
333 |
||
334 |
lemma leadsTo_cancel_Diff2: |
|
13805 | 335 |
"[| F \<in> A leadsTo (A' \<union> B); F \<in> (B-A') leadsTo B' |] |
336 |
==> F \<in> A leadsTo (A' \<union> B')" |
|
13797 | 337 |
apply (rule leadsTo_cancel2) |
338 |
prefer 2 apply assumption |
|
339 |
apply (simp_all (no_asm_simp)) |
|
340 |
done |
|
341 |
||
342 |
lemma leadsTo_cancel1: |
|
13805 | 343 |
"[| F \<in> A leadsTo (B \<union> A'); F \<in> B leadsTo B' |] |
344 |
==> F \<in> A leadsTo (B' \<union> A')" |
|
13797 | 345 |
apply (simp add: Un_commute) |
346 |
apply (blast intro!: leadsTo_cancel2) |
|
347 |
done |
|
348 |
||
349 |
lemma leadsTo_cancel_Diff1: |
|
13805 | 350 |
"[| F \<in> A leadsTo (B \<union> A'); F \<in> (B-A') leadsTo B' |] |
351 |
==> F \<in> A leadsTo (B' \<union> A')" |
|
13797 | 352 |
apply (rule leadsTo_cancel1) |
353 |
prefer 2 apply assumption |
|
354 |
apply (simp_all (no_asm_simp)) |
|
355 |
done |
|
356 |
||
357 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
358 |
text{*The impossibility law*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
359 |
lemma leadsTo_empty: "[|F \<in> A leadsTo {}; all_total F|] ==> A={}" |
13797 | 360 |
apply (erule leadsTo_induct_pre) |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
361 |
apply (simp_all add: ensures_def constrains_def transient_def all_total_def, clarify) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
362 |
apply (drule bspec, assumption)+ |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
363 |
apply blast |
13797 | 364 |
done |
365 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
366 |
subsection{*PSP: Progress-Safety-Progress*} |
13797 | 367 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
368 |
text{*Special case of PSP: Misra's "stable conjunction"*} |
13797 | 369 |
lemma psp_stable: |
13805 | 370 |
"[| F \<in> A leadsTo A'; F \<in> stable B |] |
371 |
==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B)" |
|
13797 | 372 |
apply (unfold stable_def) |
373 |
apply (erule leadsTo_induct) |
|
374 |
prefer 3 apply (blast intro: leadsTo_Union_Int) |
|
375 |
prefer 2 apply (blast intro: leadsTo_Trans) |
|
376 |
apply (rule leadsTo_Basis) |
|
377 |
apply (simp add: ensures_def Diff_Int_distrib2 [symmetric] Int_Un_distrib2 [symmetric]) |
|
378 |
apply (blast intro: transient_strengthen constrains_Int) |
|
379 |
done |
|
380 |
||
381 |
lemma psp_stable2: |
|
13805 | 382 |
"[| F \<in> A leadsTo A'; F \<in> stable B |] ==> F \<in> (B \<inter> A) leadsTo (B \<inter> A')" |
13797 | 383 |
by (simp add: psp_stable Int_ac) |
384 |
||
385 |
lemma psp_ensures: |
|
13805 | 386 |
"[| F \<in> A ensures A'; F \<in> B co B' |] |
387 |
==> F \<in> (A \<inter> B') ensures ((A' \<inter> B) \<union> (B' - B))" |
|
13797 | 388 |
apply (unfold ensures_def constrains_def, clarify) (*speeds up the proof*) |
389 |
apply (blast intro: transient_strengthen) |
|
390 |
done |
|
391 |
||
392 |
lemma psp: |
|
13805 | 393 |
"[| F \<in> A leadsTo A'; F \<in> B co B' |] |
394 |
==> F \<in> (A \<inter> B') leadsTo ((A' \<inter> B) \<union> (B' - B))" |
|
13797 | 395 |
apply (erule leadsTo_induct) |
396 |
prefer 3 apply (blast intro: leadsTo_Union_Int) |
|
397 |
txt{*Basis case*} |
|
398 |
apply (blast intro: psp_ensures) |
|
399 |
txt{*Transitivity case has a delicate argument involving "cancellation"*} |
|
400 |
apply (rule leadsTo_Un_duplicate2) |
|
401 |
apply (erule leadsTo_cancel_Diff1) |
|
402 |
apply (simp add: Int_Diff Diff_triv) |
|
403 |
apply (blast intro: leadsTo_weaken_L dest: constrains_imp_subset) |
|
404 |
done |
|
405 |
||
406 |
lemma psp2: |
|
13805 | 407 |
"[| F \<in> A leadsTo A'; F \<in> B co B' |] |
408 |
==> F \<in> (B' \<inter> A) leadsTo ((B \<inter> A') \<union> (B' - B))" |
|
13797 | 409 |
by (simp (no_asm_simp) add: psp Int_ac) |
410 |
||
411 |
lemma psp_unless: |
|
13805 | 412 |
"[| F \<in> A leadsTo A'; F \<in> B unless B' |] |
413 |
==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B) \<union> B')" |
|
13797 | 414 |
|
415 |
apply (unfold unless_def) |
|
416 |
apply (drule psp, assumption) |
|
417 |
apply (blast intro: leadsTo_weaken) |
|
418 |
done |
|
419 |
||
420 |
||
13798 | 421 |
subsection{*Proving the induction rules*} |
13797 | 422 |
|
423 |
(** The most general rule: r is any wf relation; f is any variant function **) |
|
424 |
||
425 |
lemma leadsTo_wf_induct_lemma: |
|
426 |
"[| wf r; |
|
13805 | 427 |
\<forall>m. F \<in> (A \<inter> f-`{m}) leadsTo |
428 |
((A \<inter> f-`(r^-1 `` {m})) \<union> B) |] |
|
429 |
==> F \<in> (A \<inter> f-`{m}) leadsTo B" |
|
13797 | 430 |
apply (erule_tac a = m in wf_induct) |
13805 | 431 |
apply (subgoal_tac "F \<in> (A \<inter> (f -` (r^-1 `` {x}))) leadsTo B") |
13797 | 432 |
apply (blast intro: leadsTo_cancel1 leadsTo_Un_duplicate) |
433 |
apply (subst vimage_eq_UN) |
|
434 |
apply (simp only: UN_simps [symmetric]) |
|
435 |
apply (blast intro: leadsTo_UN) |
|
436 |
done |
|
437 |
||
438 |
||
439 |
(** Meta or object quantifier ? **) |
|
440 |
lemma leadsTo_wf_induct: |
|
441 |
"[| wf r; |
|
13805 | 442 |
\<forall>m. F \<in> (A \<inter> f-`{m}) leadsTo |
443 |
((A \<inter> f-`(r^-1 `` {m})) \<union> B) |] |
|
444 |
==> F \<in> A leadsTo B" |
|
13797 | 445 |
apply (rule_tac t = A in subst) |
446 |
defer 1 |
|
447 |
apply (rule leadsTo_UN) |
|
448 |
apply (erule leadsTo_wf_induct_lemma) |
|
449 |
apply assumption |
|
450 |
apply fast (*Blast_tac: Function unknown's argument not a parameter*) |
|
451 |
done |
|
452 |
||
453 |
||
454 |
lemma bounded_induct: |
|
455 |
"[| wf r; |
|
13805 | 456 |
\<forall>m \<in> I. F \<in> (A \<inter> f-`{m}) leadsTo |
457 |
((A \<inter> f-`(r^-1 `` {m})) \<union> B) |] |
|
458 |
==> F \<in> A leadsTo ((A - (f-`I)) \<union> B)" |
|
13797 | 459 |
apply (erule leadsTo_wf_induct, safe) |
13805 | 460 |
apply (case_tac "m \<in> I") |
13797 | 461 |
apply (blast intro: leadsTo_weaken) |
462 |
apply (blast intro: subset_imp_leadsTo) |
|
463 |
done |
|
464 |
||
465 |
||
13805 | 466 |
(*Alternative proof is via the lemma F \<in> (A \<inter> f-`(lessThan m)) leadsTo B*) |
13797 | 467 |
lemma lessThan_induct: |
15045 | 468 |
"[| !!m::nat. F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`{..<m}) \<union> B) |] |
13805 | 469 |
==> F \<in> A leadsTo B" |
13797 | 470 |
apply (rule wf_less_than [THEN leadsTo_wf_induct]) |
471 |
apply (simp (no_asm_simp)) |
|
472 |
apply blast |
|
473 |
done |
|
474 |
||
475 |
lemma lessThan_bounded_induct: |
|
13805 | 476 |
"!!l::nat. [| \<forall>m \<in> greaterThan l. |
477 |
F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`(lessThan m)) \<union> B) |] |
|
478 |
==> F \<in> A leadsTo ((A \<inter> (f-`(atMost l))) \<union> B)" |
|
13797 | 479 |
apply (simp only: Diff_eq [symmetric] vimage_Compl Compl_greaterThan [symmetric]) |
480 |
apply (rule wf_less_than [THEN bounded_induct]) |
|
481 |
apply (simp (no_asm_simp)) |
|
482 |
done |
|
483 |
||
484 |
lemma greaterThan_bounded_induct: |
|
13805 | 485 |
"(!!l::nat. \<forall>m \<in> lessThan l. |
486 |
F \<in> (A \<inter> f-`{m}) leadsTo ((A \<inter> f-`(greaterThan m)) \<union> B)) |
|
487 |
==> F \<in> A leadsTo ((A \<inter> (f-`(atLeast l))) \<union> B)" |
|
13797 | 488 |
apply (rule_tac f = f and f1 = "%k. l - k" |
489 |
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
|
490 |
apply (simp (no_asm) add:Image_singleton) |
13797 | 491 |
apply clarify |
492 |
apply (case_tac "m<l") |
|
13805 | 493 |
apply (blast intro: leadsTo_weaken_R diff_less_mono2) |
494 |
apply (blast intro: not_leE subset_imp_leadsTo) |
|
13797 | 495 |
done |
496 |
||
497 |
||
13798 | 498 |
subsection{*wlt*} |
13797 | 499 |
|
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
500 |
text{*Misra's property W3*} |
13805 | 501 |
lemma wlt_leadsTo: "F \<in> (wlt F B) leadsTo B" |
13797 | 502 |
apply (unfold wlt_def) |
503 |
apply (blast intro!: leadsTo_Union) |
|
504 |
done |
|
505 |
||
13805 | 506 |
lemma leadsTo_subset: "F \<in> A leadsTo B ==> A \<subseteq> wlt F B" |
13797 | 507 |
apply (unfold wlt_def) |
508 |
apply (blast intro!: leadsTo_Union) |
|
509 |
done |
|
510 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
511 |
text{*Misra's property W2*} |
13805 | 512 |
lemma leadsTo_eq_subset_wlt: "F \<in> A leadsTo B = (A \<subseteq> wlt F B)" |
13797 | 513 |
by (blast intro!: leadsTo_subset wlt_leadsTo [THEN leadsTo_weaken_L]) |
514 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
515 |
text{*Misra's property W4*} |
13805 | 516 |
lemma wlt_increasing: "B \<subseteq> wlt F B" |
13797 | 517 |
apply (simp (no_asm_simp) add: leadsTo_eq_subset_wlt [symmetric] subset_imp_leadsTo) |
518 |
done |
|
519 |
||
520 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
521 |
text{*Used in the Trans case below*} |
13797 | 522 |
lemma lemma1: |
13805 | 523 |
"[| B \<subseteq> A2; |
524 |
F \<in> (A1 - B) co (A1 \<union> B); |
|
525 |
F \<in> (A2 - C) co (A2 \<union> C) |] |
|
526 |
==> F \<in> (A1 \<union> A2 - C) co (A1 \<union> A2 \<union> C)" |
|
13797 | 527 |
by (unfold constrains_def, clarify, blast) |
528 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
529 |
text{*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*} |
13797 | 530 |
lemma leadsTo_123: |
13805 | 531 |
"F \<in> A leadsTo A' |
532 |
==> \<exists>B. A \<subseteq> B & F \<in> B leadsTo A' & F \<in> (B-A') co (B \<union> A')" |
|
13797 | 533 |
apply (erule leadsTo_induct) |
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
534 |
txt{*Basis*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
535 |
apply (blast dest: ensuresD) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
536 |
txt{*Trans*} |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
537 |
apply clarify |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
538 |
apply (rule_tac x = "Ba \<union> Bb" in exI) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
539 |
apply (blast intro: lemma1 leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate) |
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
540 |
txt{*Union*} |
13797 | 541 |
apply (clarify dest!: ball_conj_distrib [THEN iffD1] bchoice) |
13805 | 542 |
apply (rule_tac x = "\<Union>A \<in> S. f A" in exI) |
13797 | 543 |
apply (auto intro: leadsTo_UN) |
544 |
(*Blast_tac says PROOF FAILED*) |
|
13805 | 545 |
apply (rule_tac I1=S and A1="%i. f i - B" and A'1="%i. f i \<union> B" |
13798 | 546 |
in constrains_UN [THEN constrains_weaken], auto) |
13797 | 547 |
done |
548 |
||
549 |
||
13812
91713a1915ee
converting HOL/UNITY to use unconditional fairness
paulson
parents:
13805
diff
changeset
|
550 |
text{*Misra's property W5*} |
13805 | 551 |
lemma wlt_constrains_wlt: "F \<in> (wlt F B - B) co (wlt F B)" |
13798 | 552 |
proof - |
553 |
from wlt_leadsTo [of F B, THEN leadsTo_123] |
|
554 |
show ?thesis |
|
555 |
proof (elim exE conjE) |
|
556 |
(* assumes have to be in exactly the form as in the goal displayed at |
|
557 |
this point. Isar doesn't give you any automation. *) |
|
558 |
fix C |
|
559 |
assume wlt: "wlt F B \<subseteq> C" |
|
560 |
and lt: "F \<in> C leadsTo B" |
|
561 |
and co: "F \<in> C - B co C \<union> B" |
|
562 |
have eq: "C = wlt F B" |
|
563 |
proof - |
|
564 |
from lt and wlt show ?thesis |
|
565 |
by (blast dest: leadsTo_eq_subset_wlt [THEN iffD1]) |
|
566 |
qed |
|
567 |
from co show ?thesis by (simp add: eq wlt_increasing Un_absorb2) |
|
568 |
qed |
|
569 |
qed |
|
13797 | 570 |
|
571 |
||
13798 | 572 |
subsection{*Completion: Binary and General Finite versions*} |
13797 | 573 |
|
574 |
lemma completion_lemma : |
|
13805 | 575 |
"[| W = wlt F (B' \<union> C); |
576 |
F \<in> A leadsTo (A' \<union> C); F \<in> A' co (A' \<union> C); |
|
577 |
F \<in> B leadsTo (B' \<union> C); F \<in> B' co (B' \<union> C) |] |
|
578 |
==> F \<in> (A \<inter> B) leadsTo ((A' \<inter> B') \<union> C)" |
|
579 |
apply (subgoal_tac "F \<in> (W-C) co (W \<union> B' \<union> C) ") |
|
13797 | 580 |
prefer 2 |
581 |
apply (blast intro: wlt_constrains_wlt [THEN [2] constrains_Un, |
|
582 |
THEN constrains_weaken]) |
|
13805 | 583 |
apply (subgoal_tac "F \<in> (W-C) co W") |
13797 | 584 |
prefer 2 |
585 |
apply (simp add: wlt_increasing Un_assoc Un_absorb2) |
|
13805 | 586 |
apply (subgoal_tac "F \<in> (A \<inter> W - C) leadsTo (A' \<inter> W \<union> C) ") |
13797 | 587 |
prefer 2 apply (blast intro: wlt_leadsTo psp [THEN leadsTo_weaken]) |
588 |
(** LEVEL 6 **) |
|
13805 | 589 |
apply (subgoal_tac "F \<in> (A' \<inter> W \<union> C) leadsTo (A' \<inter> B' \<union> C) ") |
13797 | 590 |
prefer 2 |
591 |
apply (rule leadsTo_Un_duplicate2) |
|
592 |
apply (blast intro: leadsTo_Un_Un wlt_leadsTo |
|
593 |
[THEN psp2, THEN leadsTo_weaken] leadsTo_refl) |
|
594 |
apply (drule leadsTo_Diff) |
|
595 |
apply (blast intro: subset_imp_leadsTo) |
|
13805 | 596 |
apply (subgoal_tac "A \<inter> B \<subseteq> A \<inter> W") |
13797 | 597 |
prefer 2 |
598 |
apply (blast dest!: leadsTo_subset intro!: subset_refl [THEN Int_mono]) |
|
599 |
apply (blast intro: leadsTo_Trans subset_imp_leadsTo) |
|
600 |
done |
|
601 |
||
602 |
lemmas completion = completion_lemma [OF refl] |
|
603 |
||
604 |
lemma finite_completion_lemma: |
|
13805 | 605 |
"finite I ==> (\<forall>i \<in> I. F \<in> (A i) leadsTo (A' i \<union> C)) --> |
606 |
(\<forall>i \<in> I. F \<in> (A' i) co (A' i \<union> C)) --> |
|
607 |
F \<in> (\<Inter>i \<in> I. A i) leadsTo ((\<Inter>i \<in> I. A' i) \<union> C)" |
|
13797 | 608 |
apply (erule finite_induct, auto) |
609 |
apply (rule completion) |
|
610 |
prefer 4 |
|
611 |
apply (simp only: INT_simps [symmetric]) |
|
612 |
apply (rule constrains_INT, auto) |
|
613 |
done |
|
614 |
||
615 |
lemma finite_completion: |
|
616 |
"[| finite I; |
|
13805 | 617 |
!!i. i \<in> I ==> F \<in> (A i) leadsTo (A' i \<union> C); |
618 |
!!i. i \<in> I ==> F \<in> (A' i) co (A' i \<union> C) |] |
|
619 |
==> F \<in> (\<Inter>i \<in> I. A i) leadsTo ((\<Inter>i \<in> I. A' i) \<union> C)" |
|
13797 | 620 |
by (blast intro: finite_completion_lemma [THEN mp, THEN mp]) |
621 |
||
622 |
lemma stable_completion: |
|
13805 | 623 |
"[| F \<in> A leadsTo A'; F \<in> stable A'; |
624 |
F \<in> B leadsTo B'; F \<in> stable B' |] |
|
625 |
==> F \<in> (A \<inter> B) leadsTo (A' \<inter> B')" |
|
13797 | 626 |
apply (unfold stable_def) |
627 |
apply (rule_tac C1 = "{}" in completion [THEN leadsTo_weaken_R]) |
|
628 |
apply (force+) |
|
629 |
done |
|
630 |
||
631 |
lemma finite_stable_completion: |
|
632 |
"[| finite I; |
|
13805 | 633 |
!!i. i \<in> I ==> F \<in> (A i) leadsTo (A' i); |
634 |
!!i. i \<in> I ==> F \<in> stable (A' i) |] |
|
635 |
==> F \<in> (\<Inter>i \<in> I. A i) leadsTo (\<Inter>i \<in> I. A' i)" |
|
13797 | 636 |
apply (unfold stable_def) |
637 |
apply (rule_tac C1 = "{}" in finite_completion [THEN leadsTo_weaken_R]) |
|
638 |
apply (simp_all (no_asm_simp)) |
|
639 |
apply blast+ |
|
640 |
done |
|
9685 | 641 |
|
35422 | 642 |
end |