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