src/HOL/Gfp.thy
author paulson
Tue, 07 Dec 2004 16:15:44 +0100
changeset 15381 780ea4c697f2
parent 15140 322485b816ac
child 15386 06757406d8cf
permissions -rw-r--r--
converted Gfp to new-style theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15381
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
     1
(*  ID:         $Id$
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1370
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Copyright   1994  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
15381
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
     7
header {*Greatest Fixed Points*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
     8
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
     9
theory Gfp
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    10
imports Lfp
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14169
diff changeset
    11
begin
1558
9c6ebfab4e05 added constdefs section
clasohm
parents: 1475
diff changeset
    12
9c6ebfab4e05 added constdefs section
clasohm
parents: 1475
diff changeset
    13
constdefs
14169
0590de71a016 Converted to new style theories.
skalberg
parents: 8882
diff changeset
    14
  gfp :: "['a set=>'a set] => 'a set"
15381
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    15
    "gfp(f) == Union({u. u \<subseteq> f(u)})"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    16
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    17
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    18
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    19
subsection{*Proof of Knaster-Tarski Theorem using gfp*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    20
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    21
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    22
text{*@{term "gfp f"} is the least upper bound of 
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    23
      the set @{term "{u. u \<subseteq> f(u)}"} *}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    24
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    25
lemma gfp_upperbound: "[| X \<subseteq> f(X) |] ==> X \<subseteq> gfp(f)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    26
by (auto simp add: gfp_def)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    27
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    28
lemma gfp_least: "[| !!u. u \<subseteq> f(u) ==> u\<subseteq>X |] ==> gfp(f) \<subseteq> X"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    29
by (auto simp add: gfp_def)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    30
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    31
lemma gfp_lemma2: "mono(f) ==> gfp(f) \<subseteq> f(gfp(f))"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    32
by (rules intro: gfp_least subset_trans monoD gfp_upperbound)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    33
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    34
lemma gfp_lemma3: "mono(f) ==> f(gfp(f)) \<subseteq> gfp(f)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    35
by (rules intro: gfp_lemma2 monoD gfp_upperbound)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    36
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    37
lemma gfp_unfold: "mono(f) ==> gfp(f) = f(gfp(f))"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    38
by (rules intro: equalityI gfp_lemma2 gfp_lemma3)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    39
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    40
subsection{*Coinduction rules for greatest fixed points*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    41
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    42
text{*weak version*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    43
lemma weak_coinduct: "[| a: X;  X \<subseteq> f(X) |] ==> a : gfp(f)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    44
by (rule gfp_upperbound [THEN subsetD], auto)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    45
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    46
lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    47
apply (erule gfp_upperbound [THEN subsetD])
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    48
apply (erule imageI)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    49
done
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    50
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    51
lemma coinduct_lemma:
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    52
     "[| X \<subseteq> f(X Un gfp(f));  mono(f) |] ==> X Un gfp(f) \<subseteq> f(X Un gfp(f))"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    53
by (blast dest: gfp_lemma2 mono_Un)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    54
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    55
text{*strong version, thanks to Coen and Frost*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    56
lemma coinduct: "[| mono(f);  a: X;  X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    57
by (blast intro: weak_coinduct [OF _ coinduct_lemma])
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    58
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    59
lemma gfp_fun_UnI2: "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    60
by (blast dest: gfp_lemma2 mono_Un)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    61
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    62
subsection{*Even Stronger Coinduction Rule, by Martin Coen*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    63
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    64
text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    65
  @{term lfp} and @{term gfp}*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    66
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    67
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    68
by (rules intro: subset_refl monoI Un_mono monoD)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    69
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    70
lemma coinduct3_lemma:
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    71
     "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)));  mono(f) |]
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    72
      ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    73
apply (rule subset_trans)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    74
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3])
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    75
apply (rule Un_least [THEN Un_least])
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    76
apply (rule subset_refl, assumption)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    77
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    78
apply (rule monoD, assumption)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    79
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    80
done
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    81
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    82
lemma coinduct3: 
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    83
  "[| mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    84
apply (rule coinduct3_lemma [THEN [2] weak_coinduct])
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    85
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    86
done
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    87
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    88
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    89
text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    90
    to control unfolding*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    91
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    92
lemma def_gfp_unfold: "[| A==gfp(f);  mono(f) |] ==> A = f(A)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    93
by (auto intro!: gfp_unfold)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    94
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    95
lemma def_coinduct:
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    96
     "[| A==gfp(f);  mono(f);  a:X;  X \<subseteq> f(X Un A) |] ==> a: A"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    97
by (auto intro!: coinduct)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    98
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
    99
(*The version used in the induction/coinduction package*)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   100
lemma def_Collect_coinduct:
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   101
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));   
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   102
        a: X;  !!z. z: X ==> P (X Un A) z |] ==>  
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   103
     a : A"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   104
apply (erule def_coinduct, auto) 
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   105
done
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   106
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   107
lemma def_coinduct3:
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   108
    "[| A==gfp(f); mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   109
by (auto intro!: coinduct3)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   110
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   111
text{*Monotonicity of @{term gfp}!*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   112
lemma gfp_mono: "[| !!Z. f(Z)\<subseteq>g(Z) |] ==> gfp(f) \<subseteq> gfp(g)"
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   113
by (rule gfp_upperbound [THEN gfp_least], blast)
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   114
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   115
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   116
ML
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   117
{*
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   118
val gfp_def = thm "gfp_def";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   119
val gfp_upperbound = thm "gfp_upperbound";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   120
val gfp_least = thm "gfp_least";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   121
val gfp_lemma2 = thm "gfp_lemma2";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   122
val gfp_lemma3 = thm "gfp_lemma3";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   123
val gfp_unfold = thm "gfp_unfold";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   124
val weak_coinduct = thm "weak_coinduct";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   125
val weak_coinduct_image = thm "weak_coinduct_image";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   126
val coinduct_lemma = thm "coinduct_lemma";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   127
val coinduct = thm "coinduct";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   128
val gfp_fun_UnI2 = thm "gfp_fun_UnI2";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   129
val coinduct3_mono_lemma = thm "coinduct3_mono_lemma";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   130
val coinduct3_lemma = thm "coinduct3_lemma";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   131
val coinduct3 = thm "coinduct3";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   132
val def_gfp_unfold = thm "def_gfp_unfold";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   133
val def_coinduct = thm "def_coinduct";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   134
val def_Collect_coinduct = thm "def_Collect_coinduct";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   135
val def_coinduct3 = thm "def_coinduct3";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   136
val gfp_mono = thm "gfp_mono";
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   137
*}
780ea4c697f2 converted Gfp to new-style theory
paulson
parents: 15140
diff changeset
   138
1558
9c6ebfab4e05 added constdefs section
clasohm
parents: 1475
diff changeset
   139
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   140
end