Gfp.ML
author clasohm
Wed, 02 Mar 1994 12:26:55 +0100
changeset 48 21291189b51e
parent 2 befa4e9f7c90
child 128 89669c58e506
permissions -rw-r--r--
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     1
(*  Title: 	HOL/gfp
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     6
For gfp.thy.  The Knaster-Tarski Theorem for greatest fixed points.
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     7
*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     8
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     9
open Gfp;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    10
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    11
(*** Proof of Knaster-Tarski Theorem using gfp ***)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    12
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    13
(* gfp(f) is the least upper bound of {u. u <= f(u)} *)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    16
by (rtac (CollectI RS Union_upper) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    17
by (resolve_tac prems 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    18
val gfp_upperbound = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    19
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    22
by (REPEAT (ares_tac ([Union_least]@prems) 1));
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    23
by (etac CollectD 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    24
val gfp_least = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    25
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    26
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) <= f(gfp(f))";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    27
by (EVERY1 [rtac gfp_least, rtac subset_trans, atac,
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    28
	    rtac (mono RS monoD), rtac gfp_upperbound, atac]);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    29
val gfp_lemma2 = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    30
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    31
val [mono] = goal Gfp.thy "mono(f) ==> f(gfp(f)) <= gfp(f)";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    32
by (EVERY1 [rtac gfp_upperbound, rtac (mono RS monoD), 
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    33
	    rtac gfp_lemma2, rtac mono]);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    34
val gfp_lemma3 = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    35
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    36
val [mono] = goal Gfp.thy "mono(f) ==> gfp(f) = f(gfp(f))";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    37
by (REPEAT (resolve_tac [equalityI,gfp_lemma2,gfp_lemma3,mono] 1));
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    38
val gfp_Tarski = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    39
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    40
(*** Coinduction rules for greatest fixed points ***)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    41
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    42
(*weak version*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    45
by (rtac (gfp_upperbound RS subsetD) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    48
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    71
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    77
by (REPEAT (ares_tac [subset_refl, monoI, Un_mono, prem RS monoD] 1));
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    78
val coinduct3_mono_lemma= result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    79
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    83
by (rtac subset_trans 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    84
by (rtac (mono RS coinduct3_mono_lemma RS lfp_lemma3) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    85
by (rtac (Un_least RS Un_least) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    86
by (rtac subset_refl 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    87
by (rtac prem 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    88
by (rtac (mono RS gfp_Tarski RS equalityD1 RS subset_trans) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    89
by (rtac (mono RS monoD) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    90
by (rtac (mono RS coinduct3_mono_lemma RS lfp_Tarski RS ssubst) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    91
by (rtac Un_upper2 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    92
val coinduct3_lemma = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    93
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    97
by (resolve_tac (prems RL [coinduct3_mono_lemma RS lfp_Tarski RS ssubst]) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    98
by (rtac (UnI2 RS UnI1) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    99
by (REPEAT (resolve_tac prems 1));
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   100
val coinduct3 = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   101
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   102
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   103
(** Definition forms of gfp_Tarski and coinduct, to control unfolding **)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   106
by (rewtac rew);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   107
by (rtac (mono RS gfp_Tarski) 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   108
val def_gfp_Tarski = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   109
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   114
val def_coinduct = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   115
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   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
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   118
by (rewtac rew);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   119
by (REPEAT (ares_tac (map (rewrite_rule [rew]) prems @ [coinduct3]) 1));
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   120
val def_coinduct3 = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   121
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   122
(*Monotonicity of gfp!*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   123
val prems = goal Gfp.thy
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   124
    "[| mono(f);  !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   125
by (rtac gfp_upperbound 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   126
by (rtac subset_trans 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   127
by (rtac gfp_lemma2 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   128
by (resolve_tac prems 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   129
by (resolve_tac prems 1);
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   130
val gfp_mono = result();
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   131
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   132
(*Monotonicity of gfp!*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   133
val [prem] = goal Gfp.thy "[| !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   134
br (gfp_upperbound RS gfp_least) 1;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   135
be (prem RSN (2,subset_trans)) 1;
7949f97df77a Initial revision
clasohm
parents:
diff changeset
   136
val gfp_mono = result();