src/HOL/Gfp.ML
author paulson
Mon, 12 Jan 2004 16:51:45 +0100
changeset 14353 79f9fbef9106
parent 14169 0590de71a016
permissions -rw-r--r--
Added lemmas to Ring_and_Field with slightly modified simplification rules Deleted some little-used integer theorems, replacing them by the generic ones in Ring_and_Field Consolidated integer powers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9422
4b6bc2b347e5 avoid referencing thy value;
wenzelm
parents: 5316
diff changeset
     1
(*  Title:      HOL/Gfp.ML
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5069
diff changeset
     6
The Knaster-Tarski Theorem for greatest fixed points.
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
(*** Proof of Knaster-Tarski Theorem using gfp ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
14169
0590de71a016 Converted to new style theories.
skalberg
parents: 11335
diff changeset
    11
val gfp_def = thm "gfp_def";
0590de71a016 Converted to new style theories.
skalberg
parents: 11335
diff changeset
    12
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
(* gfp(f) is the least upper bound of {u. u <= f(u)} *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5069
diff changeset
    15
Goalw [gfp_def] "[| X <= f(X) |] ==> X <= gfp(f)";
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5069
diff changeset
    16
by (etac (CollectI RS Union_upper) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
qed "gfp_upperbound";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    19
val prems = Goalw [gfp_def]
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
    "[| !!u. u <= f(u) ==> u<=X |] ==> gfp(f) <= X";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
by (REPEAT (ares_tac ([Union_least]@prems) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
by (etac CollectD 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
qed "gfp_least";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    25
Goal "mono(f) ==> gfp(f) <= f(gfp(f))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
by (EVERY1 [rtac gfp_least, rtac subset_trans, atac,
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    27
            etac monoD, rtac gfp_upperbound, atac]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
qed "gfp_lemma2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    30
Goal "mono(f) ==> f(gfp(f)) <= gfp(f)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    31
by (EVERY1 [rtac gfp_upperbound, rtac monoD, assume_tac,
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    32
            etac gfp_lemma2]);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
qed "gfp_lemma3";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    35
Goal "mono(f) ==> gfp(f) = f(gfp(f))";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    36
by (REPEAT (ares_tac [equalityI,gfp_lemma2,gfp_lemma3] 1));
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
    37
qed "gfp_unfold";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
(*** Coinduction rules for greatest fixed points ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
(*weak version*)
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5069
diff changeset
    42
Goal "[| a: X;  X <= f(X) |] ==> a : gfp(f)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
by (rtac (gfp_upperbound RS subsetD) 1);
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5069
diff changeset
    44
by Auto_tac;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    45
qed "weak_coinduct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
11335
c150861633da added weak_coinduct_image
oheimb
parents: 10186
diff changeset
    47
Goal "!!X. [| a : X; g`X <= f (g`X) |] ==> g a : gfp f";
c150861633da added weak_coinduct_image
oheimb
parents: 10186
diff changeset
    48
by (etac (gfp_upperbound RS subsetD) 1);
c150861633da added weak_coinduct_image
oheimb
parents: 10186
diff changeset
    49
by (etac imageI 1);
c150861633da added weak_coinduct_image
oheimb
parents: 10186
diff changeset
    50
qed "weak_coinduct_image";
c150861633da added weak_coinduct_image
oheimb
parents: 10186
diff changeset
    51
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    52
Goal "[| X <= f(X Un gfp(f));  mono(f) |] ==>  \
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
\    X Un gfp(f) <= f(X Un gfp(f))";
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    54
by (blast_tac (claset() addDs [gfp_lemma2, mono_Un]) 1); 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
qed "coinduct_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
(*strong version, thanks to Coen & Frost*)
5148
74919e8f221c More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5069
diff changeset
    58
Goal "[| mono(f);  a: X;  X <= f(X Un gfp(f)) |] ==> a : gfp(f)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
by (rtac (coinduct_lemma RSN (2, weak_coinduct)) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
by (REPEAT (ares_tac [UnI1, Un_least] 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
qed "coinduct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    63
Goal "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))";
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    64
by (blast_tac (claset() addDs [gfp_lemma2, mono_Un]) 1); 
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
qed "gfp_fun_UnI2";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
(***  Even Stronger version of coinduct  [by Martin Coen]
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
         - instead of the condition  X <= f(X)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
                           consider  X <= (f(X) Un f(f(X)) ...) Un gfp(X) ***)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    71
Goal "mono(f) ==> mono(%x. f(x) Un X Un B)";
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    72
by (REPEAT (ares_tac [subset_refl, monoI, Un_mono] 1 ORELSE etac monoD 1));
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
qed "coinduct3_mono_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    75
Goal "[| X <= f(lfp(%x. f(x) Un X Un gfp(f)));  mono(f) |] ==> \
3842
b55686a7b22c fixed dots;
wenzelm
parents: 2036
diff changeset
    76
\    lfp(%x. f(x) Un X Un gfp(f)) <= f(lfp(%x. f(x) Un X Un gfp(f)))";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
by (rtac subset_trans 1);
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    78
by (etac (coinduct3_mono_lemma RS lfp_lemma3) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
by (rtac (Un_least RS Un_least) 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
by (rtac subset_refl 1);
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    81
by (assume_tac 1); 
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
    82
by (rtac (gfp_unfold RS equalityD1 RS subset_trans) 1);
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    83
by (assume_tac 1);
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    84
by (rtac monoD 1 THEN assume_tac 1);
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
    85
by (stac (coinduct3_mono_lemma RS lfp_unfold) 1);
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    86
by Auto_tac;  
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
qed "coinduct3_lemma";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    89
Goal
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    90
  "[| mono(f);  a:X;  X <= f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
by (rtac (coinduct3_lemma RSN (2,weak_coinduct)) 1);
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
    92
by (resolve_tac [coinduct3_mono_lemma RS lfp_unfold RS ssubst] 1);
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
    93
by Auto_tac;
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
qed "coinduct3";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
    97
(** Definition forms of gfp_unfold and coinduct, to control unfolding **)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
    99
Goal "[| A==gfp(f);  mono(f) |] ==> A = f(A)";
10186
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
   100
by (auto_tac (claset() addSIs [gfp_unfold], simpset()));  
499637e8f2c6 *** empty log message ***
nipkow
parents: 10067
diff changeset
   101
qed "def_gfp_unfold";
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
   103
Goal "[| A==gfp(f);  mono(f);  a:X;  X <= f(X Un A) |] ==> a: A";
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
   104
by (auto_tac (claset() addSIs [coinduct], simpset()));  
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   105
qed "def_coinduct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
(*The version used in the induction/coinduction package*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
   108
val prems = Goal
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));  \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
\       a: X;  !!z. z: X ==> P (X Un A) z |] ==> \
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
\    a : A";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   112
by (rtac def_coinduct 1);
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   113
by (REPEAT (ares_tac (prems @ [subsetI,CollectI]) 1));
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   114
qed "def_Collect_coinduct";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   115
10067
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
   116
Goal "[| A==gfp(f); mono(f);  a:X;  X <= f(lfp(%x. f(x) Un X Un A)) |] \
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
   117
\     ==> a: A";
ab03cfd6be3a tidied, removing obsolete "goal" commands
paulson
parents: 9422
diff changeset
   118
by (auto_tac (claset() addSIs [coinduct3], simpset()));  
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   119
qed "def_coinduct3";
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   120
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
(*Monotonicity of gfp!*)
5316
7a8975451a89 even more tidying of Goal commands
paulson
parents: 5148
diff changeset
   122
val [prem] = Goal "[| !!Z. f(Z)<=g(Z) |] ==> gfp(f) <= gfp(g)";
1465
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
   123
by (rtac (gfp_upperbound RS gfp_least) 1);
5d7a7e439cec expanded tabs
clasohm
parents: 923
diff changeset
   124
by (etac (prem RSN (2,subset_trans)) 1);
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   125
qed "gfp_mono";