| 
1465
 | 
     1  | 
(*  Title:      HOL/gfp
  | 
| 
923
 | 
     2  | 
    ID:         $Id$
  | 
| 
1465
 | 
     3  | 
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
  | 
| 
923
 | 
     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  | 
  | 
| 
 | 
    15  | 
val prems = goalw Gfp.thy [gfp_def] "[| X <= f(X) |] ==> X <= gfp(f)";
  | 
| 
 | 
    16  | 
by (rtac (CollectI RS Union_upper) 1);
  | 
| 
 | 
    17  | 
by (resolve_tac prems 1);
  | 
| 
 | 
    18  | 
qed "gfp_upperbound";
  | 
| 
 | 
    19  | 
  | 
| 
 | 
    20  | 
val prems = goalw Gfp.thy [gfp_def]
  | 
| 
 | 
    21  | 
    "[| !!u. u <= f(u) ==> u<=X |] ==> gfp(f) <= X";
  | 
| 
 | 
    22  | 
by (REPEAT (ares_tac ([Union_least]@prems) 1));
  | 
| 
 | 
    23  | 
by (etac CollectD 1);
  | 
| 
 | 
    24  | 
qed "gfp_least";
  | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))";
  | 
| 
 | 
    27  | 
by (EVERY1 [rtac gfp_least, rtac subset_trans, atac,
  | 
| 
1465
 | 
    28  | 
            rtac (mono RS monoD), rtac gfp_upperbound, atac]);
  | 
| 
923
 | 
    29  | 
qed "gfp_lemma2";
  | 
| 
 | 
    30  | 
  | 
| 
 | 
    31  | 
val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)";
  | 
| 
 | 
    32  | 
by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), 
  | 
| 
1465
 | 
    33  | 
            rtac gfp_lemma2, rtac mono]);
  | 
| 
923
 | 
    34  | 
qed "gfp_lemma3";
  | 
| 
 | 
    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  | 
qed "gfp_Tarski";
  | 
| 
 | 
    39  | 
  | 
| 
 | 
    40  | 
(*** Coinduction rules for greatest fixed points ***)
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
(*weak version*)
  | 
| 
 | 
    43  | 
val prems = goal Gfp.thy
  | 
| 
 | 
    44  | 
    "[| a: X;  X <= f(X) |] ==> a : gfp(f)";
  | 
| 
 | 
    45  | 
by (rtac (gfp_upperbound RS subsetD) 1);
  | 
| 
 | 
    46  | 
by (REPEAT (ares_tac prems 1));
  | 
| 
 | 
    47  | 
qed "weak_coinduct";
  | 
| 
 | 
    48  | 
  | 
| 
 | 
    49  | 
val [prem,mono] = goal Gfp.thy
  | 
| 
 | 
    50  | 
    "[| X <= f(X Un gfp(f));  mono(f) |] ==>  \
  | 
| 
 | 
    51  | 
\    X Un gfp(f) <= f(X Un gfp(f))";
  | 
| 
 | 
    52  | 
by (rtac (prem RS Un_least) 1);
  | 
| 
 | 
    53  | 
by (rtac (mono RS gfp_lemma2 RS subset_trans) 1);
  | 
| 
 | 
    54  | 
by (rtac (Un_upper2 RS subset_trans) 1);
  | 
| 
 | 
    55  | 
by (rtac (mono RS mono_Un) 1);
  | 
| 
 | 
    56  | 
qed "coinduct_lemma";
  | 
| 
 | 
    57  | 
  | 
| 
 | 
    58  | 
(*strong version, thanks to Coen & Frost*)
  | 
| 
 | 
    59  | 
goal Gfp.thy
  | 
| 
 | 
    60  | 
    "!!X. [| mono(f);  a: X;  X <= f(X Un gfp(f)) |] ==> a : gfp(f)";
  | 
| 
 | 
    61  | 
by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1);
  | 
| 
 | 
    62  | 
by (REPEAT (ares_tac [UnI1, Un_least] 1));
  | 
| 
 | 
    63  | 
qed "coinduct";
  | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
val [mono,prem] = goal Gfp.thy
  | 
| 
 | 
    66  | 
    "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))";
  | 
| 
1465
 | 
    67  | 
by (rtac (mono RS mono_Un RS subsetD) 1);
  | 
| 
 | 
    68  | 
by (rtac (mono RS gfp_lemma2 RS subsetD RS UnI2) 1);
  | 
| 
923
 | 
    69  | 
by (rtac prem 1);
  | 
| 
 | 
    70  | 
qed "gfp_fun_UnI2";
  | 
| 
 | 
    71  | 
  | 
| 
 | 
    72  | 
(***  Even Stronger version of coinduct  [by Martin Coen]
  | 
| 
 | 
    73  | 
         - instead of the condition  X <= f(X)
  | 
| 
 | 
    74  | 
                           consider  X <= (f(X) Un f(f(X)) ...) Un gfp(X) ***)
  | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un X Un B)";
  | 
| 
 | 
    77  | 
by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1));
  | 
| 
 | 
    78  | 
qed "coinduct3_mono_lemma";
  | 
| 
 | 
    79  | 
  | 
| 
 | 
    80  | 
val [prem,mono] = goal Gfp.thy
  | 
| 
 | 
    81  | 
    "[| X <= f(lfp(%x.f(x) Un X Un gfp(f)));  mono(f) |] ==> \
  | 
| 
 | 
    82  | 
\    lfp(%x.f(x) Un X Un gfp(f)) <= f(lfp(%x.f(x) Un X Un gfp(f)))";
  | 
| 
 | 
    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);
  | 
| 
2036
 | 
    90  | 
by (stac (mono RS coinduct3_mono_lemma RS lfp_Tarski) 1);
  | 
| 
923
 | 
    91  | 
by (rtac Un_upper2 1);
  | 
| 
 | 
    92  | 
qed "coinduct3_lemma";
  | 
| 
 | 
    93  | 
  | 
| 
 | 
    94  | 
val prems = goal Gfp.thy
  | 
| 
 | 
    95  | 
    "[| mono(f);  a:X;  X <= f(lfp(%x.f(x) Un X Un gfp(f))) |] ==> a : gfp(f)";
  | 
| 
 | 
    96  | 
by (rtac (coinduct3_lemma RSN (2,weak_coinduct)) 1);
  | 
| 
 | 
    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  | 
qed "coinduct3";
  | 
| 
 | 
   101  | 
  | 
| 
 | 
   102  | 
  | 
| 
 | 
   103  | 
(** Definition forms of gfp_Tarski and coinduct, to control unfolding **)
  | 
| 
 | 
   104  | 
  | 
| 
 | 
   105  | 
val [rew,mono] = goal Gfp.thy "[| A==gfp(f);  mono(f) |] ==> A = f(A)";
  | 
| 
 | 
   106  | 
by (rewtac rew);
  | 
| 
 | 
   107  | 
by (rtac (mono RS gfp_Tarski) 1);
  | 
| 
 | 
   108  | 
qed "def_gfp_Tarski";
  | 
| 
 | 
   109  | 
  | 
| 
 | 
   110  | 
val rew::prems = goal Gfp.thy
  | 
| 
 | 
   111  | 
    "[| A==gfp(f);  mono(f);  a:X;  X <= f(X Un A) |] ==> a: A";
  | 
| 
 | 
   112  | 
by (rewtac rew);
  | 
| 
 | 
   113  | 
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct]) 1));
  | 
| 
 | 
   114  | 
qed "def_coinduct";
  | 
| 
 | 
   115  | 
  | 
| 
 | 
   116  | 
(*The version used in the induction/coinduction package*)
  | 
| 
 | 
   117  | 
val prems = goal Gfp.thy
  | 
| 
 | 
   118  | 
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));  \
  | 
| 
 | 
   119  | 
\       a: X;  !!z. z: X ==> P (X Un A) z |] ==> \
  | 
| 
 | 
   120  | 
\    a : A";
  | 
| 
 | 
   121  | 
by (rtac def_coinduct 1);
  | 
| 
 | 
   122  | 
by (REPEAT (ares_tac (prems @ [subsetI,CollectI]) 1));
  | 
| 
 | 
   123  | 
qed "def_Collect_coinduct";
  | 
| 
 | 
   124  | 
  | 
| 
 | 
   125  | 
val rew::prems = goal Gfp.thy
  | 
| 
 | 
   126  | 
    "[| A==gfp(f); mono(f);  a:X;  X <= f(lfp(%x.f(x) Un X Un A)) |] ==> a: A";
  | 
| 
 | 
   127  | 
by (rewtac rew);
  | 
| 
 | 
   128  | 
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1));
  | 
| 
 | 
   129  | 
qed "def_coinduct3";
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
(*Monotonicity of gfp!*)
  | 
| 
 | 
   132  | 
val prems = goal Gfp.thy
  | 
| 
 | 
   133  | 
    "[| mono(f);  !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
  | 
| 
 | 
   134  | 
by (rtac gfp_upperbound 1);
  | 
| 
 | 
   135  | 
by (rtac subset_trans 1);
  | 
| 
 | 
   136  | 
by (rtac gfp_lemma2 1);
  | 
| 
 | 
   137  | 
by (resolve_tac prems 1);
  | 
| 
 | 
   138  | 
by (resolve_tac prems 1);
  | 
| 
 | 
   139  | 
val gfp_mono = result();
  | 
| 
 | 
   140  | 
  | 
| 
 | 
   141  | 
(*Monotonicity of gfp!*)
  | 
| 
 | 
   142  | 
val [prem] = goal Gfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
  | 
| 
1465
 | 
   143  | 
by (rtac (gfp_upperbound RS gfp_least) 1);
  | 
| 
 | 
   144  | 
by (etac (prem RSN (2,subset_trans)) 1);
  | 
| 
923
 | 
   145  | 
qed "gfp_mono";
  |