| 
17006
 | 
     1  | 
(*  Title:      HOL/FixedPoint.thy
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
  | 
| 
 | 
     4  | 
    Copyright   1992  University of Cambridge
  | 
| 
 | 
     5  | 
*)
  | 
| 
 | 
     6  | 
  | 
| 
 | 
     7  | 
header{* Fixed Points and the Knaster-Tarski Theorem*}
 | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
theory FixedPoint
  | 
| 
 | 
    10  | 
imports Product_Type
  | 
| 
 | 
    11  | 
begin
  | 
| 
 | 
    12  | 
  | 
| 
 | 
    13  | 
constdefs
  | 
| 
 | 
    14  | 
  lfp :: "['a set \<Rightarrow> 'a set] \<Rightarrow> 'a set"
  | 
| 
 | 
    15  | 
    "lfp(f) == Inter({u. f(u) \<subseteq> u})"    --{*least fixed point*}
 | 
| 
 | 
    16  | 
  | 
| 
 | 
    17  | 
  gfp :: "['a set=>'a set] => 'a set"
  | 
| 
 | 
    18  | 
    "gfp(f) == Union({u. u \<subseteq> f(u)})"
 | 
| 
 | 
    19  | 
  | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
subsection{*Proof of Knaster-Tarski Theorem using @{term lfp}*}
 | 
| 
 | 
    22  | 
  | 
| 
 | 
    23  | 
  | 
| 
 | 
    24  | 
text{*@{term "lfp f"} is the least upper bound of 
 | 
| 
 | 
    25  | 
      the set @{term "{u. f(u) \<subseteq> u}"} *}
 | 
| 
 | 
    26  | 
  | 
| 
 | 
    27  | 
lemma lfp_lowerbound: "f(A) \<subseteq> A ==> lfp(f) \<subseteq> A"
  | 
| 
 | 
    28  | 
by (auto simp add: lfp_def)
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
lemma lfp_greatest: "[| !!u. f(u) \<subseteq> u ==> A\<subseteq>u |] ==> A \<subseteq> lfp(f)"
  | 
| 
 | 
    31  | 
by (auto simp add: lfp_def)
  | 
| 
 | 
    32  | 
  | 
| 
 | 
    33  | 
lemma lfp_lemma2: "mono(f) ==> f(lfp(f)) \<subseteq> lfp(f)"
  | 
| 
17589
 | 
    34  | 
by (iprover intro: lfp_greatest subset_trans monoD lfp_lowerbound)
  | 
| 
17006
 | 
    35  | 
  | 
| 
 | 
    36  | 
lemma lfp_lemma3: "mono(f) ==> lfp(f) \<subseteq> f(lfp(f))"
  | 
| 
17589
 | 
    37  | 
by (iprover intro: lfp_lemma2 monoD lfp_lowerbound)
  | 
| 
17006
 | 
    38  | 
  | 
| 
 | 
    39  | 
lemma lfp_unfold: "mono(f) ==> lfp(f) = f(lfp(f))"
  | 
| 
17589
 | 
    40  | 
by (iprover intro: equalityI lfp_lemma2 lfp_lemma3)
  | 
| 
17006
 | 
    41  | 
  | 
| 
 | 
    42  | 
subsection{*General induction rules for greatest fixed points*}
 | 
| 
 | 
    43  | 
  | 
| 
 | 
    44  | 
lemma lfp_induct: 
  | 
| 
 | 
    45  | 
  assumes lfp: "a: lfp(f)"
  | 
| 
 | 
    46  | 
      and mono: "mono(f)"
  | 
| 
 | 
    47  | 
      and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
 | 
| 
 | 
    48  | 
  shows "P(a)"
  | 
| 
 | 
    49  | 
apply (rule_tac a=a in Int_lower2 [THEN subsetD, THEN CollectD]) 
  | 
| 
 | 
    50  | 
apply (rule lfp [THEN [2] lfp_lowerbound [THEN subsetD]]) 
  | 
| 
 | 
    51  | 
apply (rule Int_greatest)
  | 
| 
 | 
    52  | 
 apply (rule subset_trans [OF Int_lower1 [THEN mono [THEN monoD]]
  | 
| 
 | 
    53  | 
                              mono [THEN lfp_lemma2]]) 
  | 
| 
 | 
    54  | 
apply (blast intro: indhyp) 
  | 
| 
 | 
    55  | 
done
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
  | 
| 
 | 
    58  | 
text{*Version of induction for binary relations*}
 | 
| 
 | 
    59  | 
lemmas lfp_induct2 =  lfp_induct [of "(a,b)", split_format (complete)]
  | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
  | 
| 
 | 
    62  | 
lemma lfp_ordinal_induct: 
  | 
| 
 | 
    63  | 
  assumes mono: "mono f"
  | 
| 
 | 
    64  | 
  shows "[| !!S. P S ==> P(f S); !!M. !S:M. P S ==> P(Union M) |] 
  | 
| 
 | 
    65  | 
         ==> P(lfp f)"
  | 
| 
 | 
    66  | 
apply(subgoal_tac "lfp f = Union{S. S \<subseteq> lfp f & P S}")
 | 
| 
 | 
    67  | 
 apply (erule ssubst, simp) 
  | 
| 
 | 
    68  | 
apply(subgoal_tac "Union{S. S \<subseteq> lfp f & P S} \<subseteq> lfp f")
 | 
| 
 | 
    69  | 
 prefer 2 apply blast
  | 
| 
 | 
    70  | 
apply(rule equalityI)
  | 
| 
 | 
    71  | 
 prefer 2 apply assumption
  | 
| 
 | 
    72  | 
apply(drule mono [THEN monoD])
  | 
| 
 | 
    73  | 
apply (cut_tac mono [THEN lfp_unfold], simp)
  | 
| 
 | 
    74  | 
apply (rule lfp_lowerbound, auto) 
  | 
| 
 | 
    75  | 
done
  | 
| 
 | 
    76  | 
  | 
| 
 | 
    77  | 
  | 
| 
 | 
    78  | 
text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
 | 
| 
 | 
    79  | 
    to control unfolding*}
  | 
| 
 | 
    80  | 
  | 
| 
 | 
    81  | 
lemma def_lfp_unfold: "[| h==lfp(f);  mono(f) |] ==> h = f(h)"
  | 
| 
 | 
    82  | 
by (auto intro!: lfp_unfold)
  | 
| 
 | 
    83  | 
  | 
| 
 | 
    84  | 
lemma def_lfp_induct: 
  | 
| 
 | 
    85  | 
    "[| A == lfp(f);  mono(f);   a:A;                    
  | 
| 
 | 
    86  | 
        !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
 | 
| 
 | 
    87  | 
     |] ==> P(a)"
  | 
| 
 | 
    88  | 
by (blast intro: lfp_induct)
  | 
| 
 | 
    89  | 
  | 
| 
 | 
    90  | 
(*Monotonicity of lfp!*)
  | 
| 
 | 
    91  | 
lemma lfp_mono: "[| !!Z. f(Z)\<subseteq>g(Z) |] ==> lfp(f) \<subseteq> lfp(g)"
  | 
| 
 | 
    92  | 
by (rule lfp_lowerbound [THEN lfp_greatest], blast)
  | 
| 
 | 
    93  | 
  | 
| 
 | 
    94  | 
  | 
| 
 | 
    95  | 
subsection{*Proof of Knaster-Tarski Theorem using @{term gfp}*}
 | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
  | 
| 
 | 
    98  | 
text{*@{term "gfp f"} is the greatest lower bound of 
 | 
| 
 | 
    99  | 
      the set @{term "{u. u \<subseteq> f(u)}"} *}
 | 
| 
 | 
   100  | 
  | 
| 
 | 
   101  | 
lemma gfp_upperbound: "[| X \<subseteq> f(X) |] ==> X \<subseteq> gfp(f)"
  | 
| 
 | 
   102  | 
by (auto simp add: gfp_def)
  | 
| 
 | 
   103  | 
  | 
| 
 | 
   104  | 
lemma gfp_least: "[| !!u. u \<subseteq> f(u) ==> u\<subseteq>X |] ==> gfp(f) \<subseteq> X"
  | 
| 
 | 
   105  | 
by (auto simp add: gfp_def)
  | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
lemma gfp_lemma2: "mono(f) ==> gfp(f) \<subseteq> f(gfp(f))"
  | 
| 
17589
 | 
   108  | 
by (iprover intro: gfp_least subset_trans monoD gfp_upperbound)
  | 
| 
17006
 | 
   109  | 
  | 
| 
 | 
   110  | 
lemma gfp_lemma3: "mono(f) ==> f(gfp(f)) \<subseteq> gfp(f)"
  | 
| 
17589
 | 
   111  | 
by (iprover intro: gfp_lemma2 monoD gfp_upperbound)
  | 
| 
17006
 | 
   112  | 
  | 
| 
 | 
   113  | 
lemma gfp_unfold: "mono(f) ==> gfp(f) = f(gfp(f))"
  | 
| 
17589
 | 
   114  | 
by (iprover intro: equalityI gfp_lemma2 gfp_lemma3)
  | 
| 
17006
 | 
   115  | 
  | 
| 
 | 
   116  | 
subsection{*Coinduction rules for greatest fixed points*}
 | 
| 
 | 
   117  | 
  | 
| 
 | 
   118  | 
text{*weak version*}
 | 
| 
 | 
   119  | 
lemma weak_coinduct: "[| a: X;  X \<subseteq> f(X) |] ==> a : gfp(f)"
  | 
| 
 | 
   120  | 
by (rule gfp_upperbound [THEN subsetD], auto)
  | 
| 
 | 
   121  | 
  | 
| 
 | 
   122  | 
lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f"
  | 
| 
 | 
   123  | 
apply (erule gfp_upperbound [THEN subsetD])
  | 
| 
 | 
   124  | 
apply (erule imageI)
  | 
| 
 | 
   125  | 
done
  | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
lemma coinduct_lemma:
  | 
| 
 | 
   128  | 
     "[| X \<subseteq> f(X Un gfp(f));  mono(f) |] ==> X Un gfp(f) \<subseteq> f(X Un gfp(f))"
  | 
| 
 | 
   129  | 
by (blast dest: gfp_lemma2 mono_Un)
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
text{*strong version, thanks to Coen and Frost*}
 | 
| 
 | 
   132  | 
lemma coinduct: "[| mono(f);  a: X;  X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)"
  | 
| 
 | 
   133  | 
by (blast intro: weak_coinduct [OF _ coinduct_lemma])
  | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
lemma gfp_fun_UnI2: "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))"
  | 
| 
 | 
   136  | 
by (blast dest: gfp_lemma2 mono_Un)
  | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
subsection{*Even Stronger Coinduction Rule, by Martin Coen*}
 | 
| 
 | 
   139  | 
  | 
| 
 | 
   140  | 
text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
 | 
| 
 | 
   141  | 
  @{term lfp} and @{term gfp}*}
 | 
| 
 | 
   142  | 
  | 
| 
 | 
   143  | 
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)"
  | 
| 
17589
 | 
   144  | 
by (iprover intro: subset_refl monoI Un_mono monoD)
  | 
| 
17006
 | 
   145  | 
  | 
| 
 | 
   146  | 
lemma coinduct3_lemma:
  | 
| 
 | 
   147  | 
     "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)));  mono(f) |]
  | 
| 
 | 
   148  | 
      ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))"
  | 
| 
 | 
   149  | 
apply (rule subset_trans)
  | 
| 
 | 
   150  | 
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3])
  | 
| 
 | 
   151  | 
apply (rule Un_least [THEN Un_least])
  | 
| 
 | 
   152  | 
apply (rule subset_refl, assumption)
  | 
| 
 | 
   153  | 
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption)
  | 
| 
 | 
   154  | 
apply (rule monoD, assumption)
  | 
| 
 | 
   155  | 
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto)
  | 
| 
 | 
   156  | 
done
  | 
| 
 | 
   157  | 
  | 
| 
 | 
   158  | 
lemma coinduct3: 
  | 
| 
 | 
   159  | 
  "[| mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"
  | 
| 
 | 
   160  | 
apply (rule coinduct3_lemma [THEN [2] weak_coinduct])
  | 
| 
 | 
   161  | 
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto)
  | 
| 
 | 
   162  | 
done
  | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
  | 
| 
 | 
   165  | 
text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
 | 
| 
 | 
   166  | 
    to control unfolding*}
  | 
| 
 | 
   167  | 
  | 
| 
 | 
   168  | 
lemma def_gfp_unfold: "[| A==gfp(f);  mono(f) |] ==> A = f(A)"
  | 
| 
 | 
   169  | 
by (auto intro!: gfp_unfold)
  | 
| 
 | 
   170  | 
  | 
| 
 | 
   171  | 
lemma def_coinduct:
  | 
| 
 | 
   172  | 
     "[| A==gfp(f);  mono(f);  a:X;  X \<subseteq> f(X Un A) |] ==> a: A"
  | 
| 
 | 
   173  | 
by (auto intro!: coinduct)
  | 
| 
 | 
   174  | 
  | 
| 
 | 
   175  | 
(*The version used in the induction/coinduction package*)
  | 
| 
 | 
   176  | 
lemma def_Collect_coinduct:
  | 
| 
 | 
   177  | 
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));   
  | 
| 
 | 
   178  | 
        a: X;  !!z. z: X ==> P (X Un A) z |] ==>  
  | 
| 
 | 
   179  | 
     a : A"
  | 
| 
 | 
   180  | 
apply (erule def_coinduct, auto) 
  | 
| 
 | 
   181  | 
done
  | 
| 
 | 
   182  | 
  | 
| 
 | 
   183  | 
lemma def_coinduct3:
  | 
| 
 | 
   184  | 
    "[| A==gfp(f); mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A"
  | 
| 
 | 
   185  | 
by (auto intro!: coinduct3)
  | 
| 
 | 
   186  | 
  | 
| 
 | 
   187  | 
text{*Monotonicity of @{term gfp}!*}
 | 
| 
 | 
   188  | 
lemma gfp_mono: "[| !!Z. f(Z)\<subseteq>g(Z) |] ==> gfp(f) \<subseteq> gfp(g)"
  | 
| 
 | 
   189  | 
by (rule gfp_upperbound [THEN gfp_least], blast)
  | 
| 
 | 
   190  | 
  | 
| 
 | 
   191  | 
  | 
| 
 | 
   192  | 
ML
  | 
| 
 | 
   193  | 
{*
 | 
| 
 | 
   194  | 
val lfp_def = thm "lfp_def";
  | 
| 
 | 
   195  | 
val lfp_lowerbound = thm "lfp_lowerbound";
  | 
| 
 | 
   196  | 
val lfp_greatest = thm "lfp_greatest";
  | 
| 
 | 
   197  | 
val lfp_unfold = thm "lfp_unfold";
  | 
| 
 | 
   198  | 
val lfp_induct = thm "lfp_induct";
  | 
| 
 | 
   199  | 
val lfp_induct2 = thm "lfp_induct2";
  | 
| 
 | 
   200  | 
val lfp_ordinal_induct = thm "lfp_ordinal_induct";
  | 
| 
 | 
   201  | 
val def_lfp_unfold = thm "def_lfp_unfold";
  | 
| 
 | 
   202  | 
val def_lfp_induct = thm "def_lfp_induct";
  | 
| 
 | 
   203  | 
val lfp_mono = thm "lfp_mono";
  | 
| 
 | 
   204  | 
val gfp_def = thm "gfp_def";
  | 
| 
 | 
   205  | 
val gfp_upperbound = thm "gfp_upperbound";
  | 
| 
 | 
   206  | 
val gfp_least = thm "gfp_least";
  | 
| 
 | 
   207  | 
val gfp_unfold = thm "gfp_unfold";
  | 
| 
 | 
   208  | 
val weak_coinduct = thm "weak_coinduct";
  | 
| 
 | 
   209  | 
val weak_coinduct_image = thm "weak_coinduct_image";
  | 
| 
 | 
   210  | 
val coinduct = thm "coinduct";
  | 
| 
 | 
   211  | 
val gfp_fun_UnI2 = thm "gfp_fun_UnI2";
  | 
| 
 | 
   212  | 
val coinduct3 = thm "coinduct3";
  | 
| 
 | 
   213  | 
val def_gfp_unfold = thm "def_gfp_unfold";
  | 
| 
 | 
   214  | 
val def_coinduct = thm "def_coinduct";
  | 
| 
 | 
   215  | 
val def_Collect_coinduct = thm "def_Collect_coinduct";
  | 
| 
 | 
   216  | 
val def_coinduct3 = thm "def_coinduct3";
  | 
| 
 | 
   217  | 
val gfp_mono = thm "gfp_mono";
  | 
| 
 | 
   218  | 
*}
  | 
| 
 | 
   219  | 
  | 
| 
 | 
   220  | 
end
  |