author | oheimb |
Thu, 18 Dec 1997 12:50:58 +0100 | |
changeset 4435 | 41a7e4f0e957 |
parent 3842 | b55686a7b22c |
child 5098 | 48e70d9fe05f |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/lfp.ML |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
For lfp.thy. The Knaster-Tarski Theorem |
|
7 |
*) |
|
8 |
||
9 |
open Lfp; |
|
10 |
||
11 |
(*** Proof of Knaster-Tarski Theorem ***) |
|
12 |
||
13 |
(* lfp(f) is the greatest lower bound of {u. f(u) <= u} *) |
|
14 |
||
15 |
val prems = goalw Lfp.thy [lfp_def] "[| f(A) <= A |] ==> lfp(f) <= A"; |
|
16 |
by (rtac (CollectI RS Inter_lower) 1); |
|
17 |
by (resolve_tac prems 1); |
|
18 |
qed "lfp_lowerbound"; |
|
19 |
||
20 |
val prems = goalw Lfp.thy [lfp_def] |
|
21 |
"[| !!u. f(u) <= u ==> A<=u |] ==> A <= lfp(f)"; |
|
22 |
by (REPEAT (ares_tac ([Inter_greatest]@prems) 1)); |
|
23 |
by (etac CollectD 1); |
|
24 |
qed "lfp_greatest"; |
|
25 |
||
26 |
val [mono] = goal Lfp.thy "mono(f) ==> f(lfp(f)) <= lfp(f)"; |
|
27 |
by (EVERY1 [rtac lfp_greatest, rtac subset_trans, |
|
1465 | 28 |
rtac (mono RS monoD), rtac lfp_lowerbound, atac, atac]); |
923 | 29 |
qed "lfp_lemma2"; |
30 |
||
31 |
val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) <= f(lfp(f))"; |
|
32 |
by (EVERY1 [rtac lfp_lowerbound, rtac (mono RS monoD), |
|
1465 | 33 |
rtac lfp_lemma2, rtac mono]); |
923 | 34 |
qed "lfp_lemma3"; |
35 |
||
36 |
val [mono] = goal Lfp.thy "mono(f) ==> lfp(f) = f(lfp(f))"; |
|
37 |
by (REPEAT (resolve_tac [equalityI,lfp_lemma2,lfp_lemma3,mono] 1)); |
|
38 |
qed "lfp_Tarski"; |
|
39 |
||
40 |
(*** General induction rule for least fixed points ***) |
|
41 |
||
42 |
val [lfp,mono,indhyp] = goal Lfp.thy |
|
1465 | 43 |
"[| a: lfp(f); mono(f); \ |
3842 | 44 |
\ !!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x) \ |
923 | 45 |
\ |] ==> P(a)"; |
46 |
by (res_inst_tac [("a","a")] (Int_lower2 RS subsetD RS CollectD) 1); |
|
47 |
by (rtac (lfp RSN (2, lfp_lowerbound RS subsetD)) 1); |
|
48 |
by (EVERY1 [rtac Int_greatest, rtac subset_trans, |
|
1465 | 49 |
rtac (Int_lower1 RS (mono RS monoD)), |
50 |
rtac (mono RS lfp_lemma2), |
|
51 |
rtac (CollectI RS subsetI), rtac indhyp, atac]); |
|
923 | 52 |
qed "induct"; |
53 |
||
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1552
diff
changeset
|
54 |
bind_thm |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1552
diff
changeset
|
55 |
("induct2", |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1552
diff
changeset
|
56 |
Prod_Syntax.split_rule |
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1552
diff
changeset
|
57 |
(read_instantiate [("a","(a,b)")] induct)); |
1114 | 58 |
|
1125 | 59 |
|
923 | 60 |
(** Definition forms of lfp_Tarski and induct, to control unfolding **) |
61 |
||
62 |
val [rew,mono] = goal Lfp.thy "[| h==lfp(f); mono(f) |] ==> h = f(h)"; |
|
63 |
by (rewtac rew); |
|
64 |
by (rtac (mono RS lfp_Tarski) 1); |
|
65 |
qed "def_lfp_Tarski"; |
|
66 |
||
67 |
val rew::prems = goal Lfp.thy |
|
1465 | 68 |
"[| A == lfp(f); mono(f); a:A; \ |
3842 | 69 |
\ !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x) \ |
923 | 70 |
\ |] ==> P(a)"; |
1465 | 71 |
by (EVERY1 [rtac induct, (*backtracking to force correct induction*) |
72 |
REPEAT1 o (ares_tac (map (rewrite_rule [rew]) prems))]); |
|
923 | 73 |
qed "def_induct"; |
74 |
||
75 |
(*Monotonicity of lfp!*) |
|
76 |
val [prem] = goal Lfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> lfp(f) <= lfp(g)"; |
|
1465 | 77 |
by (rtac (lfp_lowerbound RS lfp_greatest) 1); |
78 |
by (etac (prem RS subset_trans) 1); |
|
923 | 79 |
qed "lfp_mono"; |