author | paulson |
Thu, 12 Jul 2001 16:36:26 +0200 | |
changeset 11412 | 54dd65d0ae87 |
parent 10216 | e928bdf62014 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/fixedpt.ML |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
5321 | 6 |
Least and greatest fixed points; the Knaster-Tarski Theorem |
0 | 7 |
|
8 |
Proved in the lattice of subsets of D, namely Pow(D), with Inter as glb |
|
9 |
*) |
|
10 |
||
11 |
(*** Monotone operators ***) |
|
12 |
||
5321 | 13 |
val prems = Goalw [bnd_mono_def] |
0 | 14 |
"[| h(D)<=D; \ |
15 |
\ !!W X. [| W<=D; X<=D; W<=X |] ==> h(W) <= h(X) \ |
|
16 |
\ |] ==> bnd_mono(D,h)"; |
|
17 |
by (REPEAT (ares_tac (prems@[conjI,allI,impI]) 1 |
|
18 |
ORELSE etac subset_trans 1)); |
|
760 | 19 |
qed "bnd_monoI"; |
0 | 20 |
|
5321 | 21 |
Goalw [bnd_mono_def] "bnd_mono(D,h) ==> h(D) <= D"; |
22 |
by (etac conjunct1 1); |
|
760 | 23 |
qed "bnd_monoD1"; |
0 | 24 |
|
5321 | 25 |
Goalw [bnd_mono_def] "[| bnd_mono(D,h); W<=X; X<=D |] ==> h(W) <= h(X)"; |
26 |
by (Blast_tac 1); |
|
760 | 27 |
qed "bnd_monoD2"; |
0 | 28 |
|
9907 | 29 |
val [major,minor] = goal (the_context ()) |
0 | 30 |
"[| bnd_mono(D,h); X<=D |] ==> h(X) <= D"; |
31 |
by (rtac (major RS bnd_monoD2 RS subset_trans) 1); |
|
32 |
by (rtac (major RS bnd_monoD1) 3); |
|
33 |
by (rtac minor 1); |
|
34 |
by (rtac subset_refl 1); |
|
760 | 35 |
qed "bnd_mono_subset"; |
0 | 36 |
|
5137 | 37 |
Goal "[| bnd_mono(D,h); A <= D; B <= D |] ==> \ |
0 | 38 |
\ h(A) Un h(B) <= h(A Un B)"; |
39 |
by (REPEAT (ares_tac [Un_upper1, Un_upper2, Un_least] 1 |
|
40 |
ORELSE etac bnd_monoD2 1)); |
|
760 | 41 |
qed "bnd_mono_Un"; |
0 | 42 |
|
43 |
(*Useful??*) |
|
5137 | 44 |
Goal "[| bnd_mono(D,h); A <= D; B <= D |] ==> \ |
0 | 45 |
\ h(A Int B) <= h(A) Int h(B)"; |
46 |
by (REPEAT (ares_tac [Int_lower1, Int_lower2, Int_greatest] 1 |
|
47 |
ORELSE etac bnd_monoD2 1)); |
|
760 | 48 |
qed "bnd_mono_Int"; |
0 | 49 |
|
50 |
(**** Proof of Knaster-Tarski Theorem for the lfp ****) |
|
51 |
||
52 |
(*lfp is contained in each pre-fixedpoint*) |
|
5067 | 53 |
Goalw [lfp_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
54 |
"[| h(A) <= A; A<=D |] ==> lfp(D,h) <= A"; |
3016 | 55 |
by (Blast_tac 1); |
744 | 56 |
(*or by (rtac (PowI RS CollectI RS Inter_lower) 1); *) |
760 | 57 |
qed "lfp_lowerbound"; |
0 | 58 |
|
59 |
(*Unfolding the defn of Inter dispenses with the premise bnd_mono(D,h)!*) |
|
5067 | 60 |
Goalw [lfp_def,Inter_def] "lfp(D,h) <= D"; |
3016 | 61 |
by (Blast_tac 1); |
760 | 62 |
qed "lfp_subset"; |
0 | 63 |
|
64 |
(*Used in datatype package*) |
|
9907 | 65 |
val [rew] = goal (the_context ()) "A==lfp(D,h) ==> A <= D"; |
0 | 66 |
by (rewtac rew); |
67 |
by (rtac lfp_subset 1); |
|
760 | 68 |
qed "def_lfp_subset"; |
0 | 69 |
|
5321 | 70 |
val prems = Goalw [lfp_def] |
0 | 71 |
"[| h(D) <= D; !!X. [| h(X) <= X; X<=D |] ==> A<=X |] ==> \ |
72 |
\ A <= lfp(D,h)"; |
|
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
0
diff
changeset
|
73 |
by (rtac (Pow_top RS CollectI RS Inter_greatest) 1); |
0 | 74 |
by (REPEAT (ares_tac prems 1 ORELSE eresolve_tac [CollectE,PowD] 1)); |
760 | 75 |
qed "lfp_greatest"; |
0 | 76 |
|
9907 | 77 |
val hmono::prems = goal (the_context ()) |
0 | 78 |
"[| bnd_mono(D,h); h(A)<=A; A<=D |] ==> h(lfp(D,h)) <= A"; |
79 |
by (rtac (hmono RS bnd_monoD2 RS subset_trans) 1); |
|
80 |
by (rtac lfp_lowerbound 1); |
|
81 |
by (REPEAT (resolve_tac prems 1)); |
|
760 | 82 |
qed "lfp_lemma1"; |
0 | 83 |
|
5321 | 84 |
Goal "bnd_mono(D,h) ==> h(lfp(D,h)) <= lfp(D,h)"; |
0 | 85 |
by (rtac (bnd_monoD1 RS lfp_greatest) 1); |
86 |
by (rtac lfp_lemma1 2); |
|
5321 | 87 |
by (REPEAT (assume_tac 1)); |
760 | 88 |
qed "lfp_lemma2"; |
0 | 89 |
|
9907 | 90 |
val [hmono] = goal (the_context ()) |
0 | 91 |
"bnd_mono(D,h) ==> lfp(D,h) <= h(lfp(D,h))"; |
92 |
by (rtac lfp_lowerbound 1); |
|
93 |
by (rtac (hmono RS bnd_monoD2) 1); |
|
94 |
by (rtac (hmono RS lfp_lemma2) 1); |
|
95 |
by (rtac (hmono RS bnd_mono_subset) 2); |
|
96 |
by (REPEAT (rtac lfp_subset 1)); |
|
760 | 97 |
qed "lfp_lemma3"; |
0 | 98 |
|
5321 | 99 |
Goal "bnd_mono(D,h) ==> lfp(D,h) = h(lfp(D,h))"; |
100 |
by (REPEAT (ares_tac [equalityI,lfp_lemma2,lfp_lemma3] 1)); |
|
10216 | 101 |
qed "lfp_unfold"; |
0 | 102 |
|
103 |
(*Definition form, to control unfolding*) |
|
9907 | 104 |
val [rew,mono] = goal (the_context ()) |
0 | 105 |
"[| A==lfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"; |
106 |
by (rewtac rew); |
|
10216 | 107 |
by (rtac (mono RS lfp_unfold) 1); |
108 |
qed "def_lfp_unfold"; |
|
0 | 109 |
|
110 |
(*** General induction rule for least fixedpoints ***) |
|
111 |
||
5321 | 112 |
val [hmono,indstep] = Goal |
0 | 113 |
"[| bnd_mono(D,h); !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ |
114 |
\ |] ==> h(Collect(lfp(D,h),P)) <= Collect(lfp(D,h),P)"; |
|
115 |
by (rtac subsetI 1); |
|
116 |
by (rtac CollectI 1); |
|
117 |
by (etac indstep 2); |
|
118 |
by (rtac (hmono RS lfp_lemma2 RS subsetD) 1); |
|
119 |
by (rtac (hmono RS bnd_monoD2 RS subsetD) 1); |
|
120 |
by (REPEAT (ares_tac [Collect_subset, lfp_subset] 1)); |
|
760 | 121 |
qed "Collect_is_pre_fixedpt"; |
0 | 122 |
|
123 |
(*This rule yields an induction hypothesis in which the components of a |
|
124 |
data structure may be assumed to be elements of lfp(D,h)*) |
|
5321 | 125 |
val prems = Goal |
1461 | 126 |
"[| bnd_mono(D,h); a : lfp(D,h); \ |
127 |
\ !!x. x : h(Collect(lfp(D,h),P)) ==> P(x) \ |
|
0 | 128 |
\ |] ==> P(a)"; |
129 |
by (rtac (Collect_is_pre_fixedpt RS lfp_lowerbound RS subsetD RS CollectD2) 1); |
|
130 |
by (rtac (lfp_subset RS (Collect_subset RS subset_trans)) 3); |
|
131 |
by (REPEAT (ares_tac prems 1)); |
|
760 | 132 |
qed "induct"; |
0 | 133 |
|
134 |
(*Definition form, to control unfolding*) |
|
5321 | 135 |
val rew::prems = Goal |
0 | 136 |
"[| A == lfp(D,h); bnd_mono(D,h); a:A; \ |
137 |
\ !!x. x : h(Collect(A,P)) ==> P(x) \ |
|
138 |
\ |] ==> P(a)"; |
|
139 |
by (rtac induct 1); |
|
140 |
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1)); |
|
760 | 141 |
qed "def_induct"; |
0 | 142 |
|
143 |
(*This version is useful when "A" is not a subset of D; |
|
144 |
second premise could simply be h(D Int A) <= D or !!X. X<=D ==> h(X)<=D *) |
|
9907 | 145 |
val [hsub,hmono] = goal (the_context ()) |
0 | 146 |
"[| h(D Int A) <= A; bnd_mono(D,h) |] ==> lfp(D,h) <= A"; |
147 |
by (rtac (lfp_lowerbound RS subset_trans) 1); |
|
148 |
by (rtac (hmono RS bnd_mono_subset RS Int_greatest) 1); |
|
149 |
by (REPEAT (resolve_tac [hsub,Int_lower1,Int_lower2] 1)); |
|
760 | 150 |
qed "lfp_Int_lowerbound"; |
0 | 151 |
|
152 |
(*Monotonicity of lfp, where h precedes i under a domain-like partial order |
|
153 |
monotonicity of h is not strictly necessary; h must be bounded by D*) |
|
5321 | 154 |
val [hmono,imono,subhi] = Goal |
1461 | 155 |
"[| bnd_mono(D,h); bnd_mono(E,i); \ |
0 | 156 |
\ !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(E,i)"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
0
diff
changeset
|
157 |
by (rtac (bnd_monoD1 RS lfp_greatest) 1); |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
0
diff
changeset
|
158 |
by (rtac imono 1); |
0 | 159 |
by (rtac (hmono RSN (2, lfp_Int_lowerbound)) 1); |
160 |
by (rtac (Int_lower1 RS subhi RS subset_trans) 1); |
|
161 |
by (rtac (imono RS bnd_monoD2 RS subset_trans) 1); |
|
162 |
by (REPEAT (ares_tac [Int_lower2] 1)); |
|
760 | 163 |
qed "lfp_mono"; |
0 | 164 |
|
165 |
(*This (unused) version illustrates that monotonicity is not really needed, |
|
166 |
but both lfp's must be over the SAME set D; Inter is anti-monotonic!*) |
|
5321 | 167 |
val [isubD,subhi] = Goal |
0 | 168 |
"[| i(D) <= D; !!X. X<=D ==> h(X) <= i(X) |] ==> lfp(D,h) <= lfp(D,i)"; |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
0
diff
changeset
|
169 |
by (rtac lfp_greatest 1); |
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
0
diff
changeset
|
170 |
by (rtac isubD 1); |
0 | 171 |
by (rtac lfp_lowerbound 1); |
14
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents:
0
diff
changeset
|
172 |
by (etac (subhi RS subset_trans) 1); |
0 | 173 |
by (REPEAT (assume_tac 1)); |
760 | 174 |
qed "lfp_mono2"; |
0 | 175 |
|
176 |
||
177 |
(**** Proof of Knaster-Tarski Theorem for the gfp ****) |
|
178 |
||
179 |
(*gfp contains each post-fixedpoint that is contained in D*) |
|
5321 | 180 |
Goalw [gfp_def] "[| A <= h(A); A<=D |] ==> A <= gfp(D,h)"; |
0 | 181 |
by (rtac (PowI RS CollectI RS Union_upper) 1); |
5321 | 182 |
by (REPEAT (assume_tac 1)); |
760 | 183 |
qed "gfp_upperbound"; |
0 | 184 |
|
5067 | 185 |
Goalw [gfp_def] "gfp(D,h) <= D"; |
3016 | 186 |
by (Blast_tac 1); |
760 | 187 |
qed "gfp_subset"; |
0 | 188 |
|
189 |
(*Used in datatype package*) |
|
9907 | 190 |
val [rew] = goal (the_context ()) "A==gfp(D,h) ==> A <= D"; |
0 | 191 |
by (rewtac rew); |
192 |
by (rtac gfp_subset 1); |
|
760 | 193 |
qed "def_gfp_subset"; |
0 | 194 |
|
5321 | 195 |
val hmono::prems = Goalw [gfp_def] |
0 | 196 |
"[| bnd_mono(D,h); !!X. [| X <= h(X); X<=D |] ==> X<=A |] ==> \ |
197 |
\ gfp(D,h) <= A"; |
|
198 |
by (fast_tac (subset_cs addIs ((hmono RS bnd_monoD1)::prems)) 1); |
|
760 | 199 |
qed "gfp_least"; |
0 | 200 |
|
9907 | 201 |
val hmono::prems = goal (the_context ()) |
0 | 202 |
"[| bnd_mono(D,h); A<=h(A); A<=D |] ==> A <= h(gfp(D,h))"; |
203 |
by (rtac (hmono RS bnd_monoD2 RSN (2,subset_trans)) 1); |
|
204 |
by (rtac gfp_subset 3); |
|
205 |
by (rtac gfp_upperbound 2); |
|
206 |
by (REPEAT (resolve_tac prems 1)); |
|
760 | 207 |
qed "gfp_lemma1"; |
0 | 208 |
|
5321 | 209 |
Goal "bnd_mono(D,h) ==> gfp(D,h) <= h(gfp(D,h))"; |
0 | 210 |
by (rtac gfp_least 1); |
211 |
by (rtac gfp_lemma1 2); |
|
5321 | 212 |
by (REPEAT (assume_tac 1)); |
760 | 213 |
qed "gfp_lemma2"; |
0 | 214 |
|
9907 | 215 |
val [hmono] = goal (the_context ()) |
0 | 216 |
"bnd_mono(D,h) ==> h(gfp(D,h)) <= gfp(D,h)"; |
217 |
by (rtac gfp_upperbound 1); |
|
218 |
by (rtac (hmono RS bnd_monoD2) 1); |
|
219 |
by (rtac (hmono RS gfp_lemma2) 1); |
|
220 |
by (REPEAT (rtac ([hmono, gfp_subset] MRS bnd_mono_subset) 1)); |
|
760 | 221 |
qed "gfp_lemma3"; |
0 | 222 |
|
5321 | 223 |
Goal "bnd_mono(D,h) ==> gfp(D,h) = h(gfp(D,h))"; |
224 |
by (REPEAT (ares_tac [equalityI,gfp_lemma2,gfp_lemma3] 1)); |
|
10216 | 225 |
qed "gfp_unfold"; |
0 | 226 |
|
227 |
(*Definition form, to control unfolding*) |
|
9907 | 228 |
val [rew,mono] = goal (the_context ()) |
0 | 229 |
"[| A==gfp(D,h); bnd_mono(D,h) |] ==> A = h(A)"; |
230 |
by (rewtac rew); |
|
10216 | 231 |
by (rtac (mono RS gfp_unfold) 1); |
232 |
qed "def_gfp_unfold"; |
|
0 | 233 |
|
234 |
||
235 |
(*** Coinduction rules for greatest fixed points ***) |
|
236 |
||
237 |
(*weak version*) |
|
5137 | 238 |
Goal "[| a: X; X <= h(X); X <= D |] ==> a : gfp(D,h)"; |
0 | 239 |
by (REPEAT (ares_tac [gfp_upperbound RS subsetD] 1)); |
760 | 240 |
qed "weak_coinduct"; |
0 | 241 |
|
9907 | 242 |
val [subs_h,subs_D,mono] = goal (the_context ()) |
0 | 243 |
"[| X <= h(X Un gfp(D,h)); X <= D; bnd_mono(D,h) |] ==> \ |
244 |
\ X Un gfp(D,h) <= h(X Un gfp(D,h))"; |
|
245 |
by (rtac (subs_h RS Un_least) 1); |
|
246 |
by (rtac (mono RS gfp_lemma2 RS subset_trans) 1); |
|
247 |
by (rtac (Un_upper2 RS subset_trans) 1); |
|
248 |
by (rtac ([mono, subs_D, gfp_subset] MRS bnd_mono_Un) 1); |
|
760 | 249 |
qed "coinduct_lemma"; |
0 | 250 |
|
251 |
(*strong version*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
252 |
Goal "[| bnd_mono(D,h); a: X; X <= h(X Un gfp(D,h)); X <= D |] ==> \ |
0 | 253 |
\ a : gfp(D,h)"; |
647
fb7345cccddc
ZF/Fixedpt/coinduct: modified proof to suppress deep unification
lcp
parents:
484
diff
changeset
|
254 |
by (rtac weak_coinduct 1); |
fb7345cccddc
ZF/Fixedpt/coinduct: modified proof to suppress deep unification
lcp
parents:
484
diff
changeset
|
255 |
by (etac coinduct_lemma 2); |
0 | 256 |
by (REPEAT (ares_tac [gfp_subset, UnI1, Un_least] 1)); |
760 | 257 |
qed "coinduct"; |
0 | 258 |
|
259 |
(*Definition form, to control unfolding*) |
|
9907 | 260 |
val rew::prems = goal (the_context ()) |
0 | 261 |
"[| A == gfp(D,h); bnd_mono(D,h); a: X; X <= h(X Un A); X <= D |] ==> \ |
262 |
\ a : A"; |
|
263 |
by (rewtac rew); |
|
264 |
by (rtac coinduct 1); |
|
265 |
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems) 1)); |
|
760 | 266 |
qed "def_coinduct"; |
0 | 267 |
|
268 |
(*Lemma used immediately below!*) |
|
5321 | 269 |
val [subsA,XimpP] = Goal |
0 | 270 |
"[| X <= A; !!z. z:X ==> P(z) |] ==> X <= Collect(A,P)"; |
271 |
by (rtac (subsA RS subsetD RS CollectI RS subsetI) 1); |
|
272 |
by (assume_tac 1); |
|
273 |
by (etac XimpP 1); |
|
760 | 274 |
qed "subset_Collect"; |
0 | 275 |
|
276 |
(*The version used in the induction/coinduction package*) |
|
5321 | 277 |
val prems = Goal |
0 | 278 |
"[| A == gfp(D, %w. Collect(D,P(w))); bnd_mono(D, %w. Collect(D,P(w))); \ |
279 |
\ a: X; X <= D; !!z. z: X ==> P(X Un A, z) |] ==> \ |
|
280 |
\ a : A"; |
|
281 |
by (rtac def_coinduct 1); |
|
282 |
by (REPEAT (ares_tac (subset_Collect::prems) 1)); |
|
760 | 283 |
qed "def_Collect_coinduct"; |
0 | 284 |
|
285 |
(*Monotonicity of gfp!*) |
|
5321 | 286 |
val [hmono,subde,subhi] = Goal |
1461 | 287 |
"[| bnd_mono(D,h); D <= E; \ |
0 | 288 |
\ !!X. X<=D ==> h(X) <= i(X) |] ==> gfp(D,h) <= gfp(E,i)"; |
289 |
by (rtac gfp_upperbound 1); |
|
290 |
by (rtac (hmono RS gfp_lemma2 RS subset_trans) 1); |
|
291 |
by (rtac (gfp_subset RS subhi) 1); |
|
292 |
by (rtac ([gfp_subset, subde] MRS subset_trans) 1); |
|
760 | 293 |
qed "gfp_mono"; |
0 | 294 |