src/HOL/Lfp.thy
author wenzelm
Wed, 13 Jul 2005 16:07:35 +0200
changeset 16813 67140ae50e77
parent 15386 06757406d8cf
permissions -rw-r--r--
removed ad-hoc atp_hook, cal_atp; removed depth_of; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1370
diff changeset
     1
(*  Title:      HOL/Lfp.thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1370
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   1992  University of Cambridge
15386
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
     5
*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
15386
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
     7
header{*Least Fixed Points and the Knaster-Tarski Theorem*}
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14169
diff changeset
     9
theory Lfp
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    10
imports Product_Type
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: 11609
diff changeset
    14
  lfp :: "['a set \<Rightarrow> 'a set] \<Rightarrow> 'a set"
15386
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    15
    "lfp(f) == Inter({u. f(u) \<subseteq> u})"    --{*least fixed point*}
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    16
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    17
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    18
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    19
subsection{*Proof of Knaster-Tarski Theorem using @{term lfp}*}
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    20
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    21
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    22
text{*@{term "lfp f"} is the least upper bound of 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    23
      the set @{term "{u. f(u) \<subseteq> u}"} *}
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    24
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    25
lemma lfp_lowerbound: "f(A) \<subseteq> A ==> lfp(f) \<subseteq> A"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    26
by (auto simp add: lfp_def)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    27
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    28
lemma lfp_greatest: "[| !!u. f(u) \<subseteq> u ==> A\<subseteq>u |] ==> A \<subseteq> lfp(f)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    29
by (auto simp add: lfp_def)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    30
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    31
lemma lfp_lemma2: "mono(f) ==> f(lfp(f)) \<subseteq> lfp(f)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    32
by (rules intro: lfp_greatest subset_trans monoD lfp_lowerbound)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    33
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    34
lemma lfp_lemma3: "mono(f) ==> lfp(f) \<subseteq> f(lfp(f))"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    35
by (rules intro: lfp_lemma2 monoD lfp_lowerbound)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    36
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    37
lemma lfp_unfold: "mono(f) ==> lfp(f) = f(lfp(f))"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    38
by (rules intro: equalityI lfp_lemma2 lfp_lemma3)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    39
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    40
subsection{*General induction rules for greatest fixed points*}
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    41
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    42
lemma lfp_induct: 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    43
  assumes lfp: "a: lfp(f)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    44
      and mono: "mono(f)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    45
      and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    46
  shows "P(a)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    47
apply (rule_tac a=a in Int_lower2 [THEN subsetD, THEN CollectD]) 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    48
apply (rule lfp [THEN [2] lfp_lowerbound [THEN subsetD]]) 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    49
apply (rule Int_greatest)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    50
 apply (rule subset_trans [OF Int_lower1 [THEN mono [THEN monoD]]
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    51
                              mono [THEN lfp_lemma2]]) 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    52
apply (blast intro: indhyp) 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    53
done
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    54
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    55
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    56
text{*Version of induction for binary relations*}
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    57
lemmas lfp_induct2 =  lfp_induct [of "(a,b)", split_format (complete)]
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    58
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    59
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    60
lemma lfp_ordinal_induct: 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    61
  assumes mono: "mono f"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    62
  shows "[| !!S. P S ==> P(f S); !!M. !S:M. P S ==> P(Union M) |] 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    63
         ==> P(lfp f)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    64
apply(subgoal_tac "lfp f = Union{S. S \<subseteq> lfp f & P S}")
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    65
 apply (erule ssubst, simp) 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    66
apply(subgoal_tac "Union{S. S \<subseteq> lfp f & P S} \<subseteq> lfp f")
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    67
 prefer 2 apply blast
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    68
apply(rule equalityI)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    69
 prefer 2 apply assumption
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    70
apply(drule mono [THEN monoD])
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    71
apply (cut_tac mono [THEN lfp_unfold], simp)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    72
apply (rule lfp_lowerbound, auto) 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    73
done
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    74
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    75
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    76
text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    77
    to control unfolding*}
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    78
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    79
lemma def_lfp_unfold: "[| h==lfp(f);  mono(f) |] ==> h = f(h)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    80
by (auto intro!: lfp_unfold)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    81
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    82
lemma def_lfp_induct: 
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    83
    "[| A == lfp(f);  mono(f);   a:A;                    
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    84
        !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    85
     |] ==> P(a)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    86
by (blast intro: lfp_induct)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    87
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    88
(*Monotonicity of lfp!*)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    89
lemma lfp_mono: "[| !!Z. f(Z)\<subseteq>g(Z) |] ==> lfp(f) \<subseteq> lfp(g)"
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    90
by (rule lfp_lowerbound [THEN lfp_greatest], blast)
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    91
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    92
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    93
ML
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    94
{*
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    95
val lfp_def = thm "lfp_def";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    96
val lfp_lowerbound = thm "lfp_lowerbound";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    97
val lfp_greatest = thm "lfp_greatest";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    98
val lfp_unfold = thm "lfp_unfold";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
    99
val lfp_induct = thm "lfp_induct";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
   100
val lfp_induct2 = thm "lfp_induct2";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
   101
val lfp_ordinal_induct = thm "lfp_ordinal_induct";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
   102
val def_lfp_unfold = thm "def_lfp_unfold";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
   103
val def_lfp_induct = thm "def_lfp_induct";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
   104
val lfp_mono = thm "lfp_mono";
06757406d8cf converted Lfp to new-style theory
paulson
parents: 15140
diff changeset
   105
*}
1558
9c6ebfab4e05 added constdefs section
clasohm
parents: 1475
diff changeset
   106
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
end