author | clasohm |
Wed, 02 Mar 1994 12:26:55 +0100 | |
changeset 48 | 21291189b51e |
parent 2 | befa4e9f7c90 |
child 128 | 89669c58e506 |
permissions | -rw-r--r-- |
0 | 1 |
(* Title: HOL/gfp |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1993 University of Cambridge |
|
5 |
||
6 |
For gfp.thy. The Knaster-Tarski Theorem for greatest fixed points. |
|
7 |
*) |
|
8 |
||
9 |
open Gfp; |
|
10 |
||
11 |
(*** Proof of Knaster-Tarski Theorem using gfp ***) |
|
12 |
||
13 |
(* gfp(f) is the least upper bound of {u. u <= f(u)} *) |
|
14 |
||
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
15 |
val prems = goalw Gfp.thy [gfp_def] "[| X <= f(X) |] ==> X <= gfp(f)"; |
0 | 16 |
by (rtac (CollectI RS Union_upper) 1); |
17 |
by (resolve_tac prems 1); |
|
18 |
val gfp_upperbound = result(); |
|
19 |
||
20 |
val prems = goalw Gfp.thy [gfp_def] |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
21 |
"[| !!u. u <= f(u) ==> u<=X |] ==> gfp(f) <= X"; |
0 | 22 |
by (REPEAT (ares_tac ([Union_least]@prems) 1)); |
23 |
by (etac CollectD 1); |
|
24 |
val gfp_least = result(); |
|
25 |
||
26 |
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))"; |
|
27 |
by (EVERY1 [rtac gfp_least, rtac subset_trans, atac, |
|
28 |
rtac (mono RS monoD), rtac gfp_upperbound, atac]); |
|
29 |
val gfp_lemma2 = result(); |
|
30 |
||
31 |
val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)"; |
|
32 |
by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), |
|
33 |
rtac gfp_lemma2, rtac mono]); |
|
34 |
val gfp_lemma3 = result(); |
|
35 |
||
36 |
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))"; |
|
37 |
by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1)); |
|
38 |
val gfp_Tarski = result(); |
|
39 |
||
40 |
(*** Coinduction rules for greatest fixed points ***) |
|
41 |
||
42 |
(*weak version*) |
|
43 |
val prems = goal Gfp.thy |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
44 |
"[| a: X; X <= f(X) |] ==> a : gfp(f)"; |
0 | 45 |
by (rtac (gfp_upperbound RS subsetD) 1); |
46 |
by (REPEAT (ares_tac prems 1)); |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
47 |
val weak_coinduct = result(); |
0 | 48 |
|
49 |
val [prem,mono] = goal Gfp.thy |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
50 |
"[| X <= f(X Un gfp(f)); mono(f) |] ==> \ |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
51 |
\ X Un gfp(f) <= f(X Un gfp(f))"; |
0 | 52 |
by (rtac (prem RS Un_least) 1); |
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
53 |
by (rtac (mono RS gfp_lemma2 RS subset_trans) 1); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
54 |
by (rtac (Un_upper2 RS subset_trans) 1); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
55 |
by (rtac (mono RS mono_Un) 1); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
56 |
val coinduct_lemma = result(); |
0 | 57 |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
58 |
(*strong version, thanks to Coen & Frost*) |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
59 |
goal Gfp.thy |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
60 |
"!!X. [| mono(f); a: X; X <= f(X Un gfp(f)) |] ==> a : gfp(f)"; |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
61 |
by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
62 |
by (REPEAT (ares_tac [UnI1, Un_least] 1)); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
63 |
val coinduct = result(); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
64 |
|
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
65 |
val [mono,prem] = goal Gfp.thy |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
66 |
"[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))"; |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
67 |
br (mono RS mono_Un RS subsetD) 1; |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
68 |
br (mono RS gfp_lemma2 RS subsetD RS UnI2) 1; |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
69 |
by (rtac prem 1); |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
70 |
val gfp_fun_UnI2 = result(); |
0 | 71 |
|
72 |
(*** Even Stronger version of coinduct [by Martin Coen] |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
73 |
- instead of the condition X <= f(X) |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
74 |
consider X <= (f(X) Un f(f(X)) ...) Un gfp(X) ***) |
0 | 75 |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
76 |
val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un X Un B)"; |
0 | 77 |
by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1)); |
78 |
val coinduct3_mono_lemma= result(); |
|
79 |
||
80 |
val [prem,mono] = goal Gfp.thy |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
81 |
"[| X <= f(lfp(%x.f(x) Un X Un gfp(f))); mono(f) |] ==> \ |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
82 |
\ lfp(%x.f(x) Un X Un gfp(f)) <= f(lfp(%x.f(x) Un X Un gfp(f)))"; |
0 | 83 |
by (rtac subset_trans 1); |
84 |
by (rtac (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1); |
|
85 |
by (rtac (Un_least RS Un_least) 1); |
|
86 |
by (rtac subset_refl 1); |
|
87 |
by (rtac prem 1); |
|
88 |
by (rtac (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1); |
|
89 |
by (rtac (mono RS monoD) 1); |
|
90 |
by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1); |
|
91 |
by (rtac Un_upper2 1); |
|
92 |
val coinduct3_lemma = result(); |
|
93 |
||
94 |
val prems = goal Gfp.thy |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
95 |
"[| mono(f); a:X; X <= f(lfp(%x.f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"; |
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
96 |
by (rtac (coinduct3_lemma RSN (2,weak_coinduct)) 1); |
0 | 97 |
by (resolve_tac (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1); |
98 |
by (rtac (UnI2 RS UnI1) 1); |
|
99 |
by (REPEAT (resolve_tac prems 1)); |
|
100 |
val coinduct3 = result(); |
|
101 |
||
102 |
||
103 |
(** Definition forms of gfp_Tarski and coinduct, to control unfolding **) |
|
104 |
||
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
105 |
val [rew,mono] = goal Gfp.thy "[| A==gfp(f); mono(f) |] ==> A = f(A)"; |
0 | 106 |
by (rewtac rew); |
107 |
by (rtac (mono RS gfp_Tarski) 1); |
|
108 |
val def_gfp_Tarski = result(); |
|
109 |
||
110 |
val rew::prems = goal Gfp.thy |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
111 |
"[| A==gfp(f); mono(f); a:X; X <= f(X Un A) |] ==> a: A"; |
0 | 112 |
by (rewtac rew); |
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
113 |
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct]) 1)); |
0 | 114 |
val def_coinduct = result(); |
115 |
||
116 |
val rew::prems = goal Gfp.thy |
|
2
befa4e9f7c90
Added weak congruence rules to HOL: if_weak_cong, case_weak_cong,
lcp
parents:
0
diff
changeset
|
117 |
"[| A==gfp(f); mono(f); a:X; X <= f(lfp(%x.f(x) Un X Un A)) |] ==> a: A"; |
0 | 118 |
by (rewtac rew); |
119 |
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1)); |
|
120 |
val def_coinduct3 = result(); |
|
121 |
||
122 |
(*Monotonicity of gfp!*) |
|
123 |
val prems = goal Gfp.thy |
|
124 |
"[| mono(f); !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)"; |
|
125 |
by (rtac gfp_upperbound 1); |
|
126 |
by (rtac subset_trans 1); |
|
127 |
by (rtac gfp_lemma2 1); |
|
128 |
by (resolve_tac prems 1); |
|
129 |
by (resolve_tac prems 1); |
|
130 |
val gfp_mono = result(); |
|
131 |
||
132 |
(*Monotonicity of gfp!*) |
|
133 |
val [prem] = goal Gfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)"; |
|
134 |
br (gfp_upperbound RS gfp_least) 1; |
|
135 |
be (prem RSN (2,subset_trans)) 1; |
|
136 |
val gfp_mono = result(); |