| author | wenzelm | 
| Mon, 31 Aug 2015 20:56:24 +0200 | |
| changeset 61068 | 6cb92c2a5ece | 
| parent 58889 | 5b7a9633cfa8 | 
| child 61424 | c3658c18b7bc | 
| permissions | -rw-r--r-- | 
| 6297 | 1  | 
(* Title: HOL/UNITY/Extend.thy  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Copyright 1998 University of Cambridge  | 
|
4  | 
||
| 13798 | 5  | 
Extending of state setsExtending of state sets  | 
| 6297 | 6  | 
function f (forget) maps the extended state to the original state  | 
7  | 
function g (forgotten) maps the extended state to the "extending part"  | 
|
8  | 
*)  | 
|
9  | 
||
| 58889 | 10  | 
section{*Extending State Sets*}
 | 
| 13798 | 11  | 
|
| 16417 | 12  | 
theory Extend imports Guar begin  | 
| 6297 | 13  | 
|
| 36866 | 14  | 
definition  | 
| 
8948
 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 
paulson 
parents: 
8703 
diff
changeset
 | 
15  | 
(*MOVE to Relation.thy?*)  | 
| 
 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 
paulson 
parents: 
8703 
diff
changeset
 | 
16  | 
  Restrict :: "[ 'a set, ('a*'b) set] => ('a*'b) set"
 | 
| 36866 | 17  | 
where "Restrict A r = r \<inter> (A <*> UNIV)"  | 
| 
8948
 
b797cfa3548d
restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
 
paulson 
parents: 
8703 
diff
changeset
 | 
18  | 
|
| 36866 | 19  | 
definition  | 
| 7482 | 20  | 
good_map :: "['a*'b => 'c] => bool"  | 
| 36866 | 21  | 
where "good_map h <-> surj h & (\<forall>x y. fst (inv h (h (x,y))) = x)"  | 
| 7482 | 22  | 
(*Using the locale constant "f", this is f (h (x,y))) = x*)  | 
23  | 
||
| 36866 | 24  | 
definition  | 
| 6297 | 25  | 
extend_set :: "['a*'b => 'c, 'a set] => 'c set"  | 
| 36866 | 26  | 
where "extend_set h A = h ` (A <*> UNIV)"  | 
| 6297 | 27  | 
|
| 36866 | 28  | 
definition  | 
| 7342 | 29  | 
project_set :: "['a*'b => 'c, 'c set] => 'a set"  | 
| 36866 | 30  | 
  where "project_set h C = {x. \<exists>y. h(x,y) \<in> C}"
 | 
| 7342 | 31  | 
|
| 36866 | 32  | 
definition  | 
| 7342 | 33  | 
  extend_act :: "['a*'b => 'c, ('a*'a) set] => ('c*'c) set"
 | 
| 36866 | 34  | 
  where "extend_act h = (%act. \<Union>(s,s') \<in> act. \<Union>y. {(h(s,y), h(s',y))})"
 | 
| 6297 | 35  | 
|
| 36866 | 36  | 
definition  | 
| 
7878
 
43b03d412b82
working version with localTo[C] instead of localTo
 
paulson 
parents: 
7826 
diff
changeset
 | 
37  | 
  project_act :: "['a*'b => 'c, ('c*'c) set] => ('a*'a) set"
 | 
| 36866 | 38  | 
  where "project_act h act = {(x,x'). \<exists>y y'. (h(x,y), h(x',y')) \<in> act}"
 | 
| 7342 | 39  | 
|
| 36866 | 40  | 
definition  | 
| 6297 | 41  | 
extend :: "['a*'b => 'c, 'a program] => 'c program"  | 
| 36866 | 42  | 
where "extend h F = mk_program (extend_set h (Init F),  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
43  | 
extend_act h ` Acts F,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
44  | 
project_act h -` AllowedActs F)"  | 
| 6297 | 45  | 
|
| 36866 | 46  | 
definition  | 
| 
7878
 
43b03d412b82
working version with localTo[C] instead of localTo
 
paulson 
parents: 
7826 
diff
changeset
 | 
47  | 
(*Argument C allows weak safety laws to be projected*)  | 
| 
7880
 
62fb24e28e5e
exchanged the first two args of "project" and "drop_prog"
 
paulson 
parents: 
7878 
diff
changeset
 | 
48  | 
project :: "['a*'b => 'c, 'c set, 'c program] => 'a program"  | 
| 36866 | 49  | 
where "project h C F =  | 
| 
10064
 
1a77667b21ef
added compatibility relation: AllowedActs, Allowed, ok,
 
paulson 
parents: 
8948 
diff
changeset
 | 
50  | 
mk_program (project_set h (Init F),  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
51  | 
project_act h ` Restrict C ` Acts F,  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
52  | 
                   {act. Restrict (project_set h C) act :
 | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
53  | 
project_act h ` Restrict C ` AllowedActs F})"  | 
| 7342 | 54  | 
|
| 6297 | 55  | 
locale Extend =  | 
| 13790 | 56  | 
fixes f :: "'c => 'a"  | 
57  | 
and g :: "'c => 'b"  | 
|
58  | 
and h :: "'a*'b => 'c" (*isomorphism between 'a * 'b and 'c *)  | 
|
59  | 
and slice :: "['c set, 'b] => 'a set"  | 
|
60  | 
assumes  | 
|
61  | 
good_h: "good_map h"  | 
|
62  | 
defines f_def: "f z == fst (inv h z)"  | 
|
63  | 
and g_def: "g z == snd (inv h z)"  | 
|
| 13805 | 64  | 
      and slice_def: "slice Z y == {x. h(x,y) \<in> Z}"
 | 
| 13790 | 65  | 
|
66  | 
||
67  | 
(** These we prove OUTSIDE the locale. **)  | 
|
68  | 
||
69  | 
||
| 13798 | 70  | 
subsection{*Restrict*}
 | 
71  | 
(*MOVE to Relation.thy?*)  | 
|
| 13790 | 72  | 
|
| 13805 | 73  | 
lemma Restrict_iff [iff]: "((x,y): Restrict A r) = ((x,y): r & x \<in> A)"  | 
| 13790 | 74  | 
by (unfold Restrict_def, blast)  | 
75  | 
||
76  | 
lemma Restrict_UNIV [simp]: "Restrict UNIV = id"  | 
|
77  | 
apply (rule ext)  | 
|
78  | 
apply (auto simp add: Restrict_def)  | 
|
79  | 
done  | 
|
80  | 
||
81  | 
lemma Restrict_empty [simp]: "Restrict {} r = {}"
 | 
|
82  | 
by (auto simp add: Restrict_def)  | 
|
83  | 
||
| 13805 | 84  | 
lemma Restrict_Int [simp]: "Restrict A (Restrict B r) = Restrict (A \<inter> B) r"  | 
| 13790 | 85  | 
by (unfold Restrict_def, blast)  | 
86  | 
||
| 13805 | 87  | 
lemma Restrict_triv: "Domain r \<subseteq> A ==> Restrict A r = r"  | 
| 13790 | 88  | 
by (unfold Restrict_def, auto)  | 
89  | 
||
| 13805 | 90  | 
lemma Restrict_subset: "Restrict A r \<subseteq> r"  | 
| 13790 | 91  | 
by (unfold Restrict_def, auto)  | 
92  | 
||
93  | 
lemma Restrict_eq_mono:  | 
|
| 13805 | 94  | 
"[| A \<subseteq> B; Restrict B r = Restrict B s |]  | 
| 13790 | 95  | 
==> Restrict A r = Restrict A s"  | 
96  | 
by (unfold Restrict_def, blast)  | 
|
97  | 
||
98  | 
lemma Restrict_imageI:  | 
|
| 13805 | 99  | 
"[| s \<in> RR; Restrict A r = Restrict A s |]  | 
100  | 
==> Restrict A r \<in> Restrict A ` RR"  | 
|
| 13790 | 101  | 
by (unfold Restrict_def image_def, auto)  | 
102  | 
||
| 13805 | 103  | 
lemma Domain_Restrict [simp]: "Domain (Restrict A r) = A \<inter> Domain r"  | 
| 13790 | 104  | 
by blast  | 
105  | 
||
| 13805 | 106  | 
lemma Image_Restrict [simp]: "(Restrict A r) `` B = r `` (A \<inter> B)"  | 
| 13790 | 107  | 
by blast  | 
108  | 
||
109  | 
(*Possibly easier than reasoning about "inv h"*)  | 
|
110  | 
lemma good_mapI:  | 
|
111  | 
assumes surj_h: "surj h"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
112  | 
and prem: "!! x x' y y'. h(x,y) = h(x',y') ==> x=x'"  | 
| 13790 | 113  | 
shows "good_map h"  | 
114  | 
apply (simp add: good_map_def)  | 
|
115  | 
apply (safe intro!: surj_h)  | 
|
116  | 
apply (rule prem)  | 
|
117  | 
apply (subst surjective_pairing [symmetric])  | 
|
118  | 
apply (subst surj_h [THEN surj_f_inv_f])  | 
|
119  | 
apply (rule refl)  | 
|
120  | 
done  | 
|
121  | 
||
122  | 
lemma good_map_is_surj: "good_map h ==> surj h"  | 
|
123  | 
by (unfold good_map_def, auto)  | 
|
124  | 
||
125  | 
(*A convenient way of finding a closed form for inv h*)  | 
|
126  | 
lemma fst_inv_equalityI:  | 
|
127  | 
assumes surj_h: "surj h"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
128  | 
and prem: "!! x y. g (h(x,y)) = x"  | 
| 13790 | 129  | 
shows "fst (inv h z) = g z"  | 
| 40702 | 130  | 
by (metis UNIV_I f_inv_into_f pair_collapse prem surj_h)  | 
| 13790 | 131  | 
|
132  | 
||
| 13798 | 133  | 
subsection{*Trivial properties of f, g, h*}
 | 
| 13790 | 134  | 
|
| 46912 | 135  | 
context Extend  | 
136  | 
begin  | 
|
137  | 
||
138  | 
lemma f_h_eq [simp]: "f(h(x,y)) = x"  | 
|
| 13790 | 139  | 
by (simp add: f_def good_h [unfolded good_map_def, THEN conjunct2])  | 
140  | 
||
| 46912 | 141  | 
lemma h_inject1 [dest]: "h(x,y) = h(x',y') ==> x=x'"  | 
| 13790 | 142  | 
apply (drule_tac f = f in arg_cong)  | 
143  | 
apply (simp add: f_def good_h [unfolded good_map_def, THEN conjunct2])  | 
|
144  | 
done  | 
|
145  | 
||
| 46912 | 146  | 
lemma h_f_g_equiv: "h(f z, g z) == z"  | 
| 13790 | 147  | 
by (simp add: f_def g_def  | 
148  | 
good_h [unfolded good_map_def, THEN conjunct1, THEN surj_f_inv_f])  | 
|
149  | 
||
| 46912 | 150  | 
lemma h_f_g_eq: "h(f z, g z) = z"  | 
| 13790 | 151  | 
by (simp add: h_f_g_equiv)  | 
152  | 
||
153  | 
||
| 46912 | 154  | 
lemma split_extended_all:  | 
| 13790 | 155  | 
"(!!z. PROP P z) == (!!u y. PROP P (h (u, y)))"  | 
156  | 
proof  | 
|
157  | 
assume allP: "\<And>z. PROP P z"  | 
|
158  | 
fix u y  | 
|
159  | 
show "PROP P (h (u, y))" by (rule allP)  | 
|
160  | 
next  | 
|
161  | 
assume allPh: "\<And>u y. PROP P (h(u,y))"  | 
|
162  | 
fix z  | 
|
163  | 
have Phfgz: "PROP P (h (f z, g z))" by (rule allPh)  | 
|
164  | 
show "PROP P z" by (rule Phfgz [unfolded h_f_g_equiv])  | 
|
165  | 
qed  | 
|
166  | 
||
| 46912 | 167  | 
end  | 
| 13790 | 168  | 
|
169  | 
||
| 13798 | 170  | 
subsection{*@{term extend_set}: basic properties*}
 | 
| 13790 | 171  | 
|
172  | 
lemma project_set_iff [iff]:  | 
|
| 13805 | 173  | 
"(x \<in> project_set h C) = (\<exists>y. h(x,y) \<in> C)"  | 
| 13790 | 174  | 
by (simp add: project_set_def)  | 
175  | 
||
| 13805 | 176  | 
lemma extend_set_mono: "A \<subseteq> B ==> extend_set h A \<subseteq> extend_set h B"  | 
| 13790 | 177  | 
by (unfold extend_set_def, blast)  | 
178  | 
||
| 46912 | 179  | 
context Extend  | 
180  | 
begin  | 
|
181  | 
||
182  | 
lemma mem_extend_set_iff [iff]: "z \<in> extend_set h A = (f z \<in> A)"  | 
|
| 13790 | 183  | 
apply (unfold extend_set_def)  | 
184  | 
apply (force intro: h_f_g_eq [symmetric])  | 
|
185  | 
done  | 
|
186  | 
||
| 46912 | 187  | 
lemma extend_set_strict_mono [iff]:  | 
| 13805 | 188  | 
"(extend_set h A \<subseteq> extend_set h B) = (A \<subseteq> B)"  | 
| 13790 | 189  | 
by (unfold extend_set_def, force)  | 
190  | 
||
| 46912 | 191  | 
lemma (in -) extend_set_empty [simp]: "extend_set h {} = {}"
 | 
| 13790 | 192  | 
by (unfold extend_set_def, auto)  | 
193  | 
||
| 46912 | 194  | 
lemma extend_set_eq_Collect: "extend_set h {s. P s} = {s. P(f s)}"
 | 
| 13790 | 195  | 
by auto  | 
196  | 
||
| 46912 | 197  | 
lemma extend_set_sing: "extend_set h {x} = {s. f s = x}"
 | 
| 13790 | 198  | 
by auto  | 
199  | 
||
| 46912 | 200  | 
lemma extend_set_inverse [simp]: "project_set h (extend_set h C) = C"  | 
| 13790 | 201  | 
by (unfold extend_set_def, auto)  | 
202  | 
||
| 46912 | 203  | 
lemma extend_set_project_set: "C \<subseteq> extend_set h (project_set h C)"  | 
| 13790 | 204  | 
apply (unfold extend_set_def)  | 
205  | 
apply (auto simp add: split_extended_all, blast)  | 
|
206  | 
done  | 
|
207  | 
||
| 46912 | 208  | 
lemma inj_extend_set: "inj (extend_set h)"  | 
| 13790 | 209  | 
apply (rule inj_on_inverseI)  | 
210  | 
apply (rule extend_set_inverse)  | 
|
211  | 
done  | 
|
212  | 
||
| 46912 | 213  | 
lemma extend_set_UNIV_eq [simp]: "extend_set h UNIV = UNIV"  | 
| 13790 | 214  | 
apply (unfold extend_set_def)  | 
215  | 
apply (auto simp add: split_extended_all)  | 
|
216  | 
done  | 
|
217  | 
||
| 13798 | 218  | 
subsection{*@{term project_set}: basic properties*}
 | 
| 13790 | 219  | 
|
220  | 
(*project_set is simply image!*)  | 
|
| 46912 | 221  | 
lemma project_set_eq: "project_set h C = f ` C"  | 
| 13790 | 222  | 
by (auto intro: f_h_eq [symmetric] simp add: split_extended_all)  | 
223  | 
||
224  | 
(*Converse appears to fail*)  | 
|
| 46912 | 225  | 
lemma project_set_I: "!!z. z \<in> C ==> f z \<in> project_set h C"  | 
| 13790 | 226  | 
by (auto simp add: split_extended_all)  | 
227  | 
||
228  | 
||
| 13798 | 229  | 
subsection{*More laws*}
 | 
| 13790 | 230  | 
|
231  | 
(*Because A and B could differ on the "other" part of the state,  | 
|
232  | 
cannot generalize to  | 
|
| 13805 | 233  | 
project_set h (A \<inter> B) = project_set h A \<inter> project_set h B  | 
| 13790 | 234  | 
*)  | 
| 46912 | 235  | 
lemma project_set_extend_set_Int: "project_set h ((extend_set h A) \<inter> B) = A \<inter> (project_set h B)"  | 
236  | 
by auto  | 
|
| 13790 | 237  | 
|
238  | 
(*Unused, but interesting?*)  | 
|
| 46912 | 239  | 
lemma project_set_extend_set_Un: "project_set h ((extend_set h A) \<union> B) = A \<union> (project_set h B)"  | 
240  | 
by auto  | 
|
| 13790 | 241  | 
|
| 46912 | 242  | 
lemma (in -) project_set_Int_subset:  | 
243  | 
"project_set h (A \<inter> B) \<subseteq> (project_set h A) \<inter> (project_set h B)"  | 
|
244  | 
by auto  | 
|
| 13790 | 245  | 
|
| 46912 | 246  | 
lemma extend_set_Un_distrib: "extend_set h (A \<union> B) = extend_set h A \<union> extend_set h B"  | 
247  | 
by auto  | 
|
| 13790 | 248  | 
|
| 46912 | 249  | 
lemma extend_set_Int_distrib: "extend_set h (A \<inter> B) = extend_set h A \<inter> extend_set h B"  | 
250  | 
by auto  | 
|
| 13790 | 251  | 
|
| 46912 | 252  | 
lemma extend_set_INT_distrib: "extend_set h (INTER A B) = (\<Inter>x \<in> A. extend_set h (B x))"  | 
253  | 
by auto  | 
|
| 13790 | 254  | 
|
| 46912 | 255  | 
lemma extend_set_Diff_distrib: "extend_set h (A - B) = extend_set h A - extend_set h B"  | 
256  | 
by auto  | 
|
| 13790 | 257  | 
|
| 46912 | 258  | 
lemma extend_set_Union: "extend_set h (Union A) = (\<Union>X \<in> A. extend_set h X)"  | 
259  | 
by blast  | 
|
260  | 
||
261  | 
lemma extend_set_subset_Compl_eq: "(extend_set h A \<subseteq> - extend_set h B) = (A \<subseteq> - B)"  | 
|
262  | 
by (auto simp: extend_set_def)  | 
|
| 13790 | 263  | 
|
264  | 
||
| 13798 | 265  | 
subsection{*@{term extend_act}*}
 | 
| 13790 | 266  | 
|
267  | 
(*Can't strengthen it to  | 
|
| 13805 | 268  | 
((h(s,y), h(s',y')) \<in> extend_act h act) = ((s, s') \<in> act & y=y')  | 
| 13790 | 269  | 
because h doesn't have to be injective in the 2nd argument*)  | 
| 46912 | 270  | 
lemma mem_extend_act_iff [iff]: "((h(s,y), h(s',y)) \<in> extend_act h act) = ((s, s') \<in> act)"  | 
271  | 
by (auto simp: extend_act_def)  | 
|
| 13790 | 272  | 
|
273  | 
(*Converse fails: (z,z') would include actions that changed the g-part*)  | 
|
| 46912 | 274  | 
lemma extend_act_D: "(z, z') \<in> extend_act h act ==> (f z, f z') \<in> act"  | 
275  | 
by (auto simp: extend_act_def)  | 
|
| 13790 | 276  | 
|
| 46912 | 277  | 
lemma extend_act_inverse [simp]: "project_act h (extend_act h act) = act"  | 
278  | 
unfolding extend_act_def project_act_def by blast  | 
|
| 13790 | 279  | 
|
| 46912 | 280  | 
lemma project_act_extend_act_restrict [simp]:  | 
| 13790 | 281  | 
"project_act h (Restrict C (extend_act h act)) =  | 
282  | 
Restrict (project_set h C) act"  | 
|
| 46912 | 283  | 
unfolding extend_act_def project_act_def by blast  | 
| 13790 | 284  | 
|
| 46912 | 285  | 
lemma subset_extend_act_D: "act' \<subseteq> extend_act h act ==> project_act h act' \<subseteq> act"  | 
286  | 
unfolding extend_act_def project_act_def by force  | 
|
| 13790 | 287  | 
|
| 46912 | 288  | 
lemma inj_extend_act: "inj (extend_act h)"  | 
| 13790 | 289  | 
apply (rule inj_on_inverseI)  | 
290  | 
apply (rule extend_act_inverse)  | 
|
291  | 
done  | 
|
292  | 
||
| 46912 | 293  | 
lemma extend_act_Image [simp]:  | 
| 13790 | 294  | 
"extend_act h act `` (extend_set h A) = extend_set h (act `` A)"  | 
| 46912 | 295  | 
unfolding extend_set_def extend_act_def by force  | 
| 13790 | 296  | 
|
| 46912 | 297  | 
lemma extend_act_strict_mono [iff]:  | 
| 13805 | 298  | 
"(extend_act h act' \<subseteq> extend_act h act) = (act'<=act)"  | 
| 46912 | 299  | 
by (auto simp: extend_act_def)  | 
| 13790 | 300  | 
|
| 46912 | 301  | 
lemma [iff]: "(extend_act h act = extend_act h act') = (act = act')"  | 
302  | 
by (rule inj_extend_act [THEN inj_eq])  | 
|
| 13790 | 303  | 
|
| 46912 | 304  | 
lemma (in -) Domain_extend_act:  | 
305  | 
"Domain (extend_act h act) = extend_set h (Domain act)"  | 
|
306  | 
unfolding extend_set_def extend_act_def by force  | 
|
307  | 
||
308  | 
lemma extend_act_Id [simp]: "extend_act h Id = Id"  | 
|
309  | 
unfolding extend_act_def by (force intro: h_f_g_eq [symmetric])  | 
|
| 13790 | 310  | 
|
| 46912 | 311  | 
lemma project_act_I: "!!z z'. (z, z') \<in> act ==> (f z, f z') \<in> project_act h act"  | 
312  | 
unfolding project_act_def by (force simp add: split_extended_all)  | 
|
| 13790 | 313  | 
|
| 46912 | 314  | 
lemma project_act_Id [simp]: "project_act h Id = Id"  | 
315  | 
unfolding project_act_def by force  | 
|
| 13790 | 316  | 
|
| 46912 | 317  | 
lemma Domain_project_act: "Domain (project_act h act) = project_set h (Domain act)"  | 
318  | 
unfolding project_act_def by (force simp add: split_extended_all)  | 
|
| 13790 | 319  | 
|
320  | 
||
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
321  | 
subsection{*extend*}
 | 
| 13790 | 322  | 
|
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
323  | 
text{*Basic properties*}
 | 
| 13790 | 324  | 
|
| 46912 | 325  | 
lemma (in -) Init_extend [simp]:  | 
| 13790 | 326  | 
"Init (extend h F) = extend_set h (Init F)"  | 
| 46912 | 327  | 
by (auto simp: extend_def)  | 
| 13790 | 328  | 
|
| 46912 | 329  | 
lemma (in -) Init_project [simp]:  | 
| 13790 | 330  | 
"Init (project h C F) = project_set h (Init F)"  | 
| 46912 | 331  | 
by (auto simp: project_def)  | 
| 13790 | 332  | 
|
| 46912 | 333  | 
lemma Acts_extend [simp]: "Acts (extend h F) = (extend_act h ` Acts F)"  | 
334  | 
by (simp add: extend_def insert_Id_image_Acts)  | 
|
| 13790 | 335  | 
|
| 46912 | 336  | 
lemma AllowedActs_extend [simp]:  | 
| 13790 | 337  | 
"AllowedActs (extend h F) = project_act h -` AllowedActs F"  | 
| 46912 | 338  | 
by (simp add: extend_def insert_absorb)  | 
| 13790 | 339  | 
|
| 46912 | 340  | 
lemma (in -) Acts_project [simp]:  | 
| 13790 | 341  | 
"Acts(project h C F) = insert Id (project_act h ` Restrict C ` Acts F)"  | 
| 46912 | 342  | 
by (auto simp add: project_def image_iff)  | 
| 13790 | 343  | 
|
| 46912 | 344  | 
lemma AllowedActs_project [simp]:  | 
| 13790 | 345  | 
"AllowedActs(project h C F) =  | 
346  | 
        {act. Restrict (project_set h C) act  
 | 
|
| 13805 | 347  | 
\<in> project_act h ` Restrict C ` AllowedActs F}"  | 
| 13790 | 348  | 
apply (simp (no_asm) add: project_def image_iff)  | 
349  | 
apply (subst insert_absorb)  | 
|
350  | 
apply (auto intro!: bexI [of _ Id] simp add: project_act_def)  | 
|
351  | 
done  | 
|
352  | 
||
| 46912 | 353  | 
lemma Allowed_extend: "Allowed (extend h F) = project h UNIV -` Allowed F"  | 
354  | 
by (auto simp add: Allowed_def)  | 
|
| 13790 | 355  | 
|
| 46912 | 356  | 
lemma extend_SKIP [simp]: "extend h SKIP = SKIP"  | 
| 13790 | 357  | 
apply (unfold SKIP_def)  | 
358  | 
apply (rule program_equalityI, auto)  | 
|
359  | 
done  | 
|
360  | 
||
| 46912 | 361  | 
lemma (in -) project_set_UNIV [simp]: "project_set h UNIV = UNIV"  | 
362  | 
by auto  | 
|
| 13790 | 363  | 
|
| 46912 | 364  | 
lemma (in -) project_set_Union: "project_set h (Union A) = (\<Union>X \<in> A. project_set h X)"  | 
365  | 
by blast  | 
|
| 13790 | 366  | 
|
| 6297 | 367  | 
|
| 13790 | 368  | 
(*Converse FAILS: the extended state contributing to project_set h C  | 
369  | 
may not coincide with the one contributing to project_act h act*)  | 
|
| 46912 | 370  | 
lemma (in -) project_act_Restrict_subset:  | 
371  | 
"project_act h (Restrict C act) \<subseteq> Restrict (project_set h C) (project_act h act)"  | 
|
372  | 
by (auto simp add: project_act_def)  | 
|
| 13790 | 373  | 
|
| 46912 | 374  | 
lemma project_act_Restrict_Id_eq: "project_act h (Restrict C Id) = Restrict (project_set h C) Id"  | 
375  | 
by (auto simp add: project_act_def)  | 
|
| 13790 | 376  | 
|
| 46912 | 377  | 
lemma project_extend_eq:  | 
| 13790 | 378  | 
"project h C (extend h F) =  | 
379  | 
mk_program (Init F, Restrict (project_set h C) ` Acts F,  | 
|
380  | 
                  {act. Restrict (project_set h C) act 
 | 
|
| 13805 | 381  | 
\<in> project_act h ` Restrict C `  | 
| 13790 | 382  | 
(project_act h -` AllowedActs F)})"  | 
383  | 
apply (rule program_equalityI)  | 
|
384  | 
apply simp  | 
|
385  | 
apply (simp add: image_eq_UN)  | 
|
386  | 
apply (simp add: project_def)  | 
|
387  | 
done  | 
|
388  | 
||
| 46912 | 389  | 
lemma extend_inverse [simp]:  | 
| 13790 | 390  | 
"project h UNIV (extend h F) = F"  | 
391  | 
apply (simp (no_asm_simp) add: project_extend_eq image_eq_UN  | 
|
392  | 
subset_UNIV [THEN subset_trans, THEN Restrict_triv])  | 
|
393  | 
apply (rule program_equalityI)  | 
|
394  | 
apply (simp_all (no_asm))  | 
|
395  | 
apply (subst insert_absorb)  | 
|
396  | 
apply (simp (no_asm) add: bexI [of _ Id])  | 
|
397  | 
apply auto  | 
|
398  | 
apply (rename_tac "act")  | 
|
399  | 
apply (rule_tac x = "extend_act h act" in bexI, auto)  | 
|
400  | 
done  | 
|
401  | 
||
| 46912 | 402  | 
lemma inj_extend: "inj (extend h)"  | 
| 13790 | 403  | 
apply (rule inj_on_inverseI)  | 
404  | 
apply (rule extend_inverse)  | 
|
405  | 
done  | 
|
406  | 
||
| 46912 | 407  | 
lemma extend_Join [simp]: "extend h (F\<squnion>G) = extend h F\<squnion>extend h G"  | 
| 13790 | 408  | 
apply (rule program_equalityI)  | 
409  | 
apply (simp (no_asm) add: extend_set_Int_distrib)  | 
|
410  | 
apply (simp add: image_Un, auto)  | 
|
411  | 
done  | 
|
412  | 
||
| 46912 | 413  | 
lemma extend_JN [simp]: "extend h (JOIN I F) = (\<Squnion>i \<in> I. extend h (F i))"  | 
| 13790 | 414  | 
apply (rule program_equalityI)  | 
415  | 
apply (simp (no_asm) add: extend_set_INT_distrib)  | 
|
416  | 
apply (simp add: image_UN, auto)  | 
|
417  | 
done  | 
|
418  | 
||
419  | 
(** These monotonicity results look natural but are UNUSED **)  | 
|
420  | 
||
| 46912 | 421  | 
lemma extend_mono: "F \<le> G ==> extend h F \<le> extend h G"  | 
422  | 
by (force simp add: component_eq_subset)  | 
|
| 13790 | 423  | 
|
| 46912 | 424  | 
lemma project_mono: "F \<le> G ==> project h C F \<le> project h C G"  | 
425  | 
by (simp add: component_eq_subset, blast)  | 
|
| 13790 | 426  | 
|
| 46912 | 427  | 
lemma all_total_extend: "all_total F ==> all_total (extend h F)"  | 
428  | 
by (simp add: all_total_def Domain_extend_act)  | 
|
| 13790 | 429  | 
|
| 13798 | 430  | 
subsection{*Safety: co, stable*}
 | 
| 13790 | 431  | 
|
| 46912 | 432  | 
lemma extend_constrains:  | 
| 13805 | 433  | 
"(extend h F \<in> (extend_set h A) co (extend_set h B)) =  | 
434  | 
(F \<in> A co B)"  | 
|
| 46912 | 435  | 
by (simp add: constrains_def)  | 
| 13790 | 436  | 
|
| 46912 | 437  | 
lemma extend_stable:  | 
| 13805 | 438  | 
"(extend h F \<in> stable (extend_set h A)) = (F \<in> stable A)"  | 
| 46912 | 439  | 
by (simp add: stable_def extend_constrains)  | 
| 13790 | 440  | 
|
| 46912 | 441  | 
lemma extend_invariant:  | 
| 13805 | 442  | 
"(extend h F \<in> invariant (extend_set h A)) = (F \<in> invariant A)"  | 
| 46912 | 443  | 
by (simp add: invariant_def extend_stable)  | 
| 13790 | 444  | 
|
445  | 
(*Projects the state predicates in the property satisfied by extend h F.  | 
|
446  | 
Converse fails: A and B may differ in their extra variables*)  | 
|
| 46912 | 447  | 
lemma extend_constrains_project_set:  | 
| 13805 | 448  | 
"extend h F \<in> A co B ==> F \<in> (project_set h A) co (project_set h B)"  | 
| 46912 | 449  | 
by (auto simp add: constrains_def, force)  | 
| 13790 | 450  | 
|
| 46912 | 451  | 
lemma extend_stable_project_set:  | 
| 13805 | 452  | 
"extend h F \<in> stable A ==> F \<in> stable (project_set h A)"  | 
| 46912 | 453  | 
by (simp add: stable_def extend_constrains_project_set)  | 
| 13790 | 454  | 
|
455  | 
||
| 13798 | 456  | 
subsection{*Weak safety primitives: Co, Stable*}
 | 
| 13790 | 457  | 
|
| 46912 | 458  | 
lemma reachable_extend_f: "p \<in> reachable (extend h F) ==> f p \<in> reachable F"  | 
459  | 
by (induct set: reachable) (auto intro: reachable.intros simp add: extend_act_def image_iff)  | 
|
| 13790 | 460  | 
|
| 46912 | 461  | 
lemma h_reachable_extend: "h(s,y) \<in> reachable (extend h F) ==> s \<in> reachable F"  | 
462  | 
by (force dest!: reachable_extend_f)  | 
|
| 13790 | 463  | 
|
| 46912 | 464  | 
lemma reachable_extend_eq: "reachable (extend h F) = extend_set h (reachable F)"  | 
| 13790 | 465  | 
apply (unfold extend_set_def)  | 
466  | 
apply (rule equalityI)  | 
|
467  | 
apply (force intro: h_f_g_eq [symmetric] dest!: reachable_extend_f, clarify)  | 
|
468  | 
apply (erule reachable.induct)  | 
|
469  | 
apply (force intro: reachable.intros)+  | 
|
470  | 
done  | 
|
471  | 
||
| 46912 | 472  | 
lemma extend_Constrains:  | 
| 13805 | 473  | 
"(extend h F \<in> (extend_set h A) Co (extend_set h B)) =  | 
474  | 
(F \<in> A Co B)"  | 
|
| 46912 | 475  | 
by (simp add: Constrains_def reachable_extend_eq extend_constrains  | 
| 13790 | 476  | 
extend_set_Int_distrib [symmetric])  | 
477  | 
||
| 46912 | 478  | 
lemma extend_Stable: "(extend h F \<in> Stable (extend_set h A)) = (F \<in> Stable A)"  | 
479  | 
by (simp add: Stable_def extend_Constrains)  | 
|
| 13790 | 480  | 
|
| 46912 | 481  | 
lemma extend_Always: "(extend h F \<in> Always (extend_set h A)) = (F \<in> Always A)"  | 
482  | 
by (simp add: Always_def extend_Stable)  | 
|
| 13790 | 483  | 
|
484  | 
||
485  | 
(** Safety and "project" **)  | 
|
486  | 
||
487  | 
(** projection: monotonicity for safety **)  | 
|
488  | 
||
| 46912 | 489  | 
lemma (in -) project_act_mono:  | 
| 13805 | 490  | 
"D \<subseteq> C ==>  | 
491  | 
project_act h (Restrict D act) \<subseteq> project_act h (Restrict C act)"  | 
|
| 46912 | 492  | 
by (auto simp add: project_act_def)  | 
| 13790 | 493  | 
|
| 46912 | 494  | 
lemma project_constrains_mono:  | 
| 13805 | 495  | 
"[| D \<subseteq> C; project h C F \<in> A co B |] ==> project h D F \<in> A co B"  | 
| 13790 | 496  | 
apply (auto simp add: constrains_def)  | 
497  | 
apply (drule project_act_mono, blast)  | 
|
498  | 
done  | 
|
499  | 
||
| 46912 | 500  | 
lemma project_stable_mono:  | 
| 13805 | 501  | 
"[| D \<subseteq> C; project h C F \<in> stable A |] ==> project h D F \<in> stable A"  | 
| 46912 | 502  | 
by (simp add: stable_def project_constrains_mono)  | 
| 13790 | 503  | 
|
504  | 
(*Key lemma used in several proofs about project and co*)  | 
|
| 46912 | 505  | 
lemma project_constrains:  | 
| 13805 | 506  | 
"(project h C F \<in> A co B) =  | 
507  | 
(F \<in> (C \<inter> extend_set h A) co (extend_set h B) & A \<subseteq> B)"  | 
|
| 13790 | 508  | 
apply (unfold constrains_def)  | 
509  | 
apply (auto intro!: project_act_I simp add: ball_Un)  | 
|
510  | 
apply (force intro!: project_act_I dest!: subsetD)  | 
|
511  | 
(*the <== direction*)  | 
|
512  | 
apply (unfold project_act_def)  | 
|
513  | 
apply (force dest!: subsetD)  | 
|
514  | 
done  | 
|
515  | 
||
| 46912 | 516  | 
lemma project_stable: "(project h UNIV F \<in> stable A) = (F \<in> stable (extend_set h A))"  | 
517  | 
by (simp add: stable_def project_constrains)  | 
|
| 13790 | 518  | 
|
| 46912 | 519  | 
lemma project_stable_I: "F \<in> stable (extend_set h A) ==> project h C F \<in> stable A"  | 
| 13790 | 520  | 
apply (drule project_stable [THEN iffD2])  | 
521  | 
apply (blast intro: project_stable_mono)  | 
|
522  | 
done  | 
|
523  | 
||
| 46912 | 524  | 
lemma Int_extend_set_lemma:  | 
| 13805 | 525  | 
"A \<inter> extend_set h ((project_set h A) \<inter> B) = A \<inter> extend_set h B"  | 
| 46912 | 526  | 
by (auto simp add: split_extended_all)  | 
| 13790 | 527  | 
|
528  | 
(*Strange (look at occurrences of C) but used in leadsETo proofs*)  | 
|
529  | 
lemma project_constrains_project_set:  | 
|
| 13805 | 530  | 
"G \<in> C co B ==> project h C G \<in> project_set h C co project_set h B"  | 
| 46912 | 531  | 
by (simp add: constrains_def project_def project_act_def, blast)  | 
| 13790 | 532  | 
|
533  | 
lemma project_stable_project_set:  | 
|
| 13805 | 534  | 
"G \<in> stable C ==> project h C G \<in> stable (project_set h C)"  | 
| 46912 | 535  | 
by (simp add: stable_def project_constrains_project_set)  | 
| 13790 | 536  | 
|
537  | 
||
| 13798 | 538  | 
subsection{*Progress: transient, ensures*}
 | 
| 13790 | 539  | 
|
| 46912 | 540  | 
lemma extend_transient:  | 
| 13805 | 541  | 
"(extend h F \<in> transient (extend_set h A)) = (F \<in> transient A)"  | 
| 46912 | 542  | 
by (auto simp add: transient_def extend_set_subset_Compl_eq Domain_extend_act)  | 
| 13790 | 543  | 
|
| 46912 | 544  | 
lemma extend_ensures:  | 
| 13805 | 545  | 
"(extend h F \<in> (extend_set h A) ensures (extend_set h B)) =  | 
546  | 
(F \<in> A ensures B)"  | 
|
| 46912 | 547  | 
by (simp add: ensures_def extend_constrains extend_transient  | 
| 13790 | 548  | 
extend_set_Un_distrib [symmetric] extend_set_Diff_distrib [symmetric])  | 
549  | 
||
| 46912 | 550  | 
lemma leadsTo_imp_extend_leadsTo:  | 
| 13805 | 551  | 
"F \<in> A leadsTo B  | 
552  | 
==> extend h F \<in> (extend_set h A) leadsTo (extend_set h B)"  | 
|
| 13790 | 553  | 
apply (erule leadsTo_induct)  | 
554  | 
apply (simp add: leadsTo_Basis extend_ensures)  | 
|
555  | 
apply (blast intro: leadsTo_Trans)  | 
|
556  | 
apply (simp add: leadsTo_UN extend_set_Union)  | 
|
557  | 
done  | 
|
558  | 
||
| 13798 | 559  | 
subsection{*Proving the converse takes some doing!*}
 | 
| 13790 | 560  | 
|
| 46912 | 561  | 
lemma slice_iff [iff]: "(x \<in> slice C y) = (h(x,y) \<in> C)"  | 
562  | 
by (simp add: slice_def)  | 
|
| 13790 | 563  | 
|
| 46912 | 564  | 
lemma slice_Union: "slice (Union S) y = (\<Union>x \<in> S. slice x y)"  | 
565  | 
by auto  | 
|
| 13790 | 566  | 
|
| 46912 | 567  | 
lemma slice_extend_set: "slice (extend_set h A) y = A"  | 
568  | 
by auto  | 
|
| 13790 | 569  | 
|
| 46912 | 570  | 
lemma project_set_is_UN_slice: "project_set h A = (\<Union>y. slice A y)"  | 
571  | 
by auto  | 
|
| 13790 | 572  | 
|
| 46912 | 573  | 
lemma extend_transient_slice:  | 
| 13805 | 574  | 
"extend h F \<in> transient A ==> F \<in> transient (slice A y)"  | 
| 46912 | 575  | 
by (auto simp: transient_def)  | 
| 13790 | 576  | 
|
577  | 
(*Converse?*)  | 
|
| 46912 | 578  | 
lemma extend_constrains_slice:  | 
| 13805 | 579  | 
"extend h F \<in> A co B ==> F \<in> (slice A y) co (slice B y)"  | 
| 46912 | 580  | 
by (auto simp add: constrains_def)  | 
| 13790 | 581  | 
|
| 46912 | 582  | 
lemma extend_ensures_slice:  | 
| 13805 | 583  | 
"extend h F \<in> A ensures B ==> F \<in> (slice A y) ensures (project_set h B)"  | 
| 13790 | 584  | 
apply (auto simp add: ensures_def extend_constrains extend_transient)  | 
585  | 
apply (erule_tac [2] extend_transient_slice [THEN transient_strengthen])  | 
|
586  | 
apply (erule extend_constrains_slice [THEN constrains_weaken], auto)  | 
|
587  | 
done  | 
|
588  | 
||
| 46912 | 589  | 
lemma leadsTo_slice_project_set:  | 
| 13805 | 590  | 
"\<forall>y. F \<in> (slice B y) leadsTo CU ==> F \<in> (project_set h B) leadsTo CU"  | 
| 46912 | 591  | 
apply (simp add: project_set_is_UN_slice)  | 
| 13790 | 592  | 
apply (blast intro: leadsTo_UN)  | 
593  | 
done  | 
|
594  | 
||
| 46912 | 595  | 
lemma extend_leadsTo_slice [rule_format]:  | 
| 13805 | 596  | 
"extend h F \<in> AU leadsTo BU  | 
597  | 
==> \<forall>y. F \<in> (slice AU y) leadsTo (project_set h BU)"  | 
|
| 13790 | 598  | 
apply (erule leadsTo_induct)  | 
| 46577 | 599  | 
apply (blast intro: extend_ensures_slice)  | 
| 13790 | 600  | 
apply (blast intro: leadsTo_slice_project_set leadsTo_Trans)  | 
601  | 
apply (simp add: leadsTo_UN slice_Union)  | 
|
602  | 
done  | 
|
603  | 
||
| 46912 | 604  | 
lemma extend_leadsTo:  | 
| 13805 | 605  | 
"(extend h F \<in> (extend_set h A) leadsTo (extend_set h B)) =  | 
606  | 
(F \<in> A leadsTo B)"  | 
|
| 13790 | 607  | 
apply safe  | 
608  | 
apply (erule_tac [2] leadsTo_imp_extend_leadsTo)  | 
|
609  | 
apply (drule extend_leadsTo_slice)  | 
|
610  | 
apply (simp add: slice_extend_set)  | 
|
611  | 
done  | 
|
612  | 
||
| 46912 | 613  | 
lemma extend_LeadsTo:  | 
| 13805 | 614  | 
"(extend h F \<in> (extend_set h A) LeadsTo (extend_set h B)) =  | 
615  | 
(F \<in> A LeadsTo B)"  | 
|
| 46912 | 616  | 
by (simp add: LeadsTo_def reachable_extend_eq extend_leadsTo  | 
| 13790 | 617  | 
extend_set_Int_distrib [symmetric])  | 
618  | 
||
619  | 
||
| 13798 | 620  | 
subsection{*preserves*}
 | 
| 13790 | 621  | 
|
| 46912 | 622  | 
lemma project_preserves_I:  | 
| 13805 | 623  | 
"G \<in> preserves (v o f) ==> project h C G \<in> preserves v"  | 
| 46912 | 624  | 
by (auto simp add: preserves_def project_stable_I extend_set_eq_Collect)  | 
| 13790 | 625  | 
|
626  | 
(*to preserve f is to preserve the whole original state*)  | 
|
| 46912 | 627  | 
lemma project_preserves_id_I:  | 
| 13805 | 628  | 
"G \<in> preserves f ==> project h C G \<in> preserves id"  | 
| 46912 | 629  | 
by (simp add: project_preserves_I)  | 
| 13790 | 630  | 
|
| 46912 | 631  | 
lemma extend_preserves:  | 
| 13805 | 632  | 
"(extend h G \<in> preserves (v o f)) = (G \<in> preserves v)"  | 
| 46912 | 633  | 
by (auto simp add: preserves_def extend_stable [symmetric]  | 
| 13790 | 634  | 
extend_set_eq_Collect)  | 
635  | 
||
| 46912 | 636  | 
lemma inj_extend_preserves: "inj h ==> (extend h G \<in> preserves g)"  | 
637  | 
by (auto simp add: preserves_def extend_def extend_act_def stable_def  | 
|
| 13790 | 638  | 
constrains_def g_def)  | 
639  | 
||
640  | 
||
| 13798 | 641  | 
subsection{*Guarantees*}
 | 
| 13790 | 642  | 
|
| 46912 | 643  | 
lemma project_extend_Join: "project h UNIV ((extend h F)\<squnion>G) = F\<squnion>(project h UNIV G)"  | 
| 13790 | 644  | 
apply (rule program_equalityI)  | 
645  | 
apply (simp add: project_set_extend_set_Int)  | 
|
| 46577 | 646  | 
apply (auto simp add: image_eq_UN)  | 
| 13790 | 647  | 
done  | 
648  | 
||
| 46912 | 649  | 
lemma extend_Join_eq_extend_D:  | 
| 13819 | 650  | 
"(extend h F)\<squnion>G = extend h H ==> H = F\<squnion>(project h UNIV G)"  | 
| 13790 | 651  | 
apply (drule_tac f = "project h UNIV" in arg_cong)  | 
652  | 
apply (simp add: project_extend_Join)  | 
|
653  | 
done  | 
|
654  | 
||
655  | 
(** Strong precondition and postcondition; only useful when  | 
|
656  | 
the old and new state sets are in bijection **)  | 
|
657  | 
||
658  | 
||
| 46912 | 659  | 
lemma ok_extend_imp_ok_project: "extend h F ok G ==> F ok project h UNIV G"  | 
| 13790 | 660  | 
apply (auto simp add: ok_def)  | 
661  | 
apply (drule subsetD)  | 
|
662  | 
apply (auto intro!: rev_image_eqI)  | 
|
663  | 
done  | 
|
664  | 
||
| 46912 | 665  | 
lemma ok_extend_iff: "(extend h F ok extend h G) = (F ok G)"  | 
| 13790 | 666  | 
apply (simp add: ok_def, safe)  | 
| 46912 | 667  | 
apply force+  | 
| 13790 | 668  | 
done  | 
669  | 
||
| 46912 | 670  | 
lemma OK_extend_iff: "OK I (%i. extend h (F i)) = (OK I F)"  | 
| 13790 | 671  | 
apply (unfold OK_def, safe)  | 
672  | 
apply (drule_tac x = i in bspec)  | 
|
673  | 
apply (drule_tac [2] x = j in bspec)  | 
|
| 46912 | 674  | 
apply force+  | 
| 13790 | 675  | 
done  | 
676  | 
||
| 46912 | 677  | 
lemma guarantees_imp_extend_guarantees:  | 
| 13805 | 678  | 
"F \<in> X guarantees Y ==>  | 
679  | 
extend h F \<in> (extend h ` X) guarantees (extend h ` Y)"  | 
|
| 13790 | 680  | 
apply (rule guaranteesI, clarify)  | 
681  | 
apply (blast dest: ok_extend_imp_ok_project extend_Join_eq_extend_D  | 
|
682  | 
guaranteesD)  | 
|
683  | 
done  | 
|
684  | 
||
| 46912 | 685  | 
lemma extend_guarantees_imp_guarantees:  | 
| 13805 | 686  | 
"extend h F \<in> (extend h ` X) guarantees (extend h ` Y)  | 
687  | 
==> F \<in> X guarantees Y"  | 
|
| 13790 | 688  | 
apply (auto simp add: guar_def)  | 
689  | 
apply (drule_tac x = "extend h G" in spec)  | 
|
690  | 
apply (simp del: extend_Join  | 
|
691  | 
add: extend_Join [symmetric] ok_extend_iff  | 
|
692  | 
inj_extend [THEN inj_image_mem_iff])  | 
|
693  | 
done  | 
|
694  | 
||
| 46912 | 695  | 
lemma extend_guarantees_eq:  | 
| 13805 | 696  | 
"(extend h F \<in> (extend h ` X) guarantees (extend h ` Y)) =  | 
697  | 
(F \<in> X guarantees Y)"  | 
|
| 46912 | 698  | 
by (blast intro: guarantees_imp_extend_guarantees  | 
| 13790 | 699  | 
extend_guarantees_imp_guarantees)  | 
| 6297 | 700  | 
|
701  | 
end  | 
|
| 46912 | 702  | 
|
703  | 
end  |