Gfp.ML
changeset 0 7949f97df77a
child 2 befa4e9f7c90
equal deleted inserted replaced
-1:000000000000 0:7949f97df77a
       
     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 
       
    15 val prems = goalw Gfp.thy [gfp_def] "[| A <= f(A) |] ==> A <= gfp(f)";
       
    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]
       
    21     "[| !!u. u <= f(u) ==> u<=A |] ==> gfp(f) <= A";
       
    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
       
    44     "[| a: A;  A <= f(A) |] ==> a : gfp(f)";
       
    45 by (rtac (gfp_upperbound RS subsetD) 1);
       
    46 by (REPEAT (ares_tac prems 1));
       
    47 val coinduct = result();
       
    48 
       
    49 val [prem,mono] = goal Gfp.thy
       
    50     "[| A <= f(A) Un gfp(f);  mono(f) |] ==>  \
       
    51 \    A Un gfp(f) <= f(A Un gfp(f))";
       
    52 by (rtac subset_trans 1);
       
    53 by (rtac (mono RS mono_Un) 2);
       
    54 by (rtac (mono RS gfp_Tarski RS subst) 1);
       
    55 by (rtac (prem RS Un_least) 1);
       
    56 by (rtac Un_upper2 1);
       
    57 val coinduct2_lemma = result();
       
    58 
       
    59 (*strong version, thanks to Martin Coen*)
       
    60 val prems = goal Gfp.thy
       
    61     "[| a: A;  A <= f(A) Un gfp(f);  mono(f) |] ==> a : gfp(f)";
       
    62 by (rtac (coinduct2_lemma RSN (2,coinduct)) 1);
       
    63 by (REPEAT (resolve_tac (prems@[UnI1]) 1));
       
    64 val coinduct2 = result();
       
    65 
       
    66 (***  Even Stronger version of coinduct  [by Martin Coen]
       
    67          - instead of the condition  A <= f(A)
       
    68                            consider  A <= (f(A) Un f(f(A)) ...) Un gfp(A) ***)
       
    69 
       
    70 val [prem] = goal Gfp.thy "mono(f) ==> mono(%x.f(x) Un A Un B)";
       
    71 by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1));
       
    72 val coinduct3_mono_lemma= result();
       
    73 
       
    74 val [prem,mono] = goal Gfp.thy
       
    75     "[| A <= f(lfp(%x.f(x) Un A Un gfp(f)));  mono(f) |] ==> \
       
    76 \    lfp(%x.f(x) Un A Un gfp(f)) <= f(lfp(%x.f(x) Un A Un gfp(f)))";
       
    77 by (rtac subset_trans 1);
       
    78 by (rtac (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1);
       
    79 by (rtac (Un_least RS Un_least) 1);
       
    80 by (rtac subset_refl 1);
       
    81 by (rtac prem 1);
       
    82 by (rtac (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1);
       
    83 by (rtac (mono RS monoD) 1);
       
    84 by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1);
       
    85 by (rtac Un_upper2 1);
       
    86 val coinduct3_lemma = result();
       
    87 
       
    88 val prems = goal Gfp.thy
       
    89     "[| a:A;  A <= f(lfp(%x.f(x) Un A Un gfp(f))); mono(f) |] ==> a : gfp(f)";
       
    90 by (rtac (coinduct3_lemma RSN (2,coinduct)) 1);
       
    91 by (resolve_tac (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1);
       
    92 by (rtac (UnI2 RS UnI1) 1);
       
    93 by (REPEAT (resolve_tac prems 1));
       
    94 val coinduct3 = result();
       
    95 
       
    96 
       
    97 (** Definition forms of gfp_Tarski and coinduct, to control unfolding **)
       
    98 
       
    99 val [rew,mono] = goal Gfp.thy "[| h==gfp(f);  mono(f) |] ==> h = f(h)";
       
   100 by (rewtac rew);
       
   101 by (rtac (mono RS gfp_Tarski) 1);
       
   102 val def_gfp_Tarski = result();
       
   103 
       
   104 val rew::prems = goal Gfp.thy
       
   105     "[| h==gfp(f);  a:A;  A <= f(A) |] ==> a: h";
       
   106 by (rewtac rew);
       
   107 by (REPEAT (ares_tac (prems @ [coinduct]) 1));
       
   108 val def_coinduct = result();
       
   109 
       
   110 val rew::prems = goal Gfp.thy
       
   111     "[| h==gfp(f);  a:A;  A <= f(A) Un h; mono(f) |] ==> a: h";
       
   112 by (rewtac rew);
       
   113 by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct2]) 1));
       
   114 val def_coinduct2 = result();
       
   115 
       
   116 val rew::prems = goal Gfp.thy
       
   117     "[| h==gfp(f);  a:A;  A <= f(lfp(%x.f(x) Un A Un h)); mono(f) |] ==> a: h";
       
   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();