src/HOL/FixedPoint.thy
author wenzelm
Fri, 21 Sep 2007 22:51:08 +0200
changeset 24669 4579eac2c997
parent 24390 9b5073c79a0b
child 24915 fc90277c0dd7
permissions -rw-r--r--
proper signature constraint; minor tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     1
(*  Title:      HOL/FixedPoint.thy
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     2
    ID:         $Id$
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
     4
    Author:     Stefan Berghofer, TU Muenchen
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     5
    Copyright   1992  University of Cambridge
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     6
*)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     7
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 22439
diff changeset
     8
header {* Fixed Points and the Knaster-Tarski Theorem*}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
     9
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    10
theory FixedPoint
23878
bd651ecd4b8a simplified HOL bootstrap
haftmann
parents: 23737
diff changeset
    11
imports Lattices
22918
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
    12
begin
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
    13
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    14
subsection {* Least and greatest fixed points *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    15
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    16
definition
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 22439
diff changeset
    17
  lfp :: "('a\<Colon>complete_lattice \<Rightarrow> 'a) \<Rightarrow> 'a" where
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    18
  "lfp f = Inf {u. f u \<le> u}"    --{*least fixed point*}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    19
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    20
definition
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 22439
diff changeset
    21
  gfp :: "('a\<Colon>complete_lattice \<Rightarrow> 'a) \<Rightarrow> 'a" where
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    22
  "gfp f = Sup {u. u \<le> f u}"    --{*greatest fixed point*}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    23
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    24
22918
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
    25
subsection{* Proof of Knaster-Tarski Theorem using @{term lfp} *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    26
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    27
text{*@{term "lfp f"} is the least upper bound of 
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    28
      the set @{term "{u. f(u) \<le> u}"} *}
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    29
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    30
lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A"
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    31
  by (auto simp add: lfp_def intro: Inf_lower)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    32
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    33
lemma lfp_greatest: "(!!u. f u \<le> u ==> A \<le> u) ==> A \<le> lfp f"
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    34
  by (auto simp add: lfp_def intro: Inf_greatest)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    35
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    36
lemma lfp_lemma2: "mono f ==> f (lfp f) \<le> lfp f"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    37
  by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    38
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    39
lemma lfp_lemma3: "mono f ==> lfp f \<le> f (lfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    40
  by (iprover intro: lfp_lemma2 monoD lfp_lowerbound)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    41
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    42
lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    43
  by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    44
22356
fe054a72d9ce Added lemma lfp_const: "lfp (%x. t) = t
krauss
parents: 22276
diff changeset
    45
lemma lfp_const: "lfp (\<lambda>x. t) = t"
fe054a72d9ce Added lemma lfp_const: "lfp (%x. t) = t
krauss
parents: 22276
diff changeset
    46
  by (rule lfp_unfold) (simp add:mono_def)
fe054a72d9ce Added lemma lfp_const: "lfp (%x. t) = t
krauss
parents: 22276
diff changeset
    47
22918
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
    48
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
    49
subsection {* General induction rules for least fixed points *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    50
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    51
theorem lfp_induct:
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    52
  assumes mono: "mono f" and ind: "f (inf (lfp f) P) <= P"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    53
  shows "lfp f <= P"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    54
proof -
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    55
  have "inf (lfp f) P <= lfp f" by (rule inf_le1)
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    56
  with mono have "f (inf (lfp f) P) <= f (lfp f)" ..
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    57
  also from mono have "f (lfp f) = lfp f" by (rule lfp_unfold [symmetric])
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    58
  finally have "f (inf (lfp f) P) <= lfp f" .
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    59
  from this and ind have "f (inf (lfp f) P) <= inf (lfp f) P" by (rule le_infI)
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    60
  hence "lfp f <= inf (lfp f) P" by (rule lfp_lowerbound)
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    61
  also have "inf (lfp f) P <= P" by (rule inf_le2)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    62
  finally show ?thesis .
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    63
qed
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    64
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    65
lemma lfp_induct_set:
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    66
  assumes lfp: "a: lfp(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    67
      and mono: "mono(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    68
      and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    69
  shows "P(a)"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
    70
  by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp])
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
    71
    (auto simp: inf_set_eq intro: indhyp)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    72
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    73
lemma lfp_ordinal_induct: 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    74
  assumes mono: "mono f"
24390
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    75
  and P_f: "!!S. P S ==> P(f S)"
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    76
  and P_Union: "!!M. !S:M. P S ==> P(Union M)"
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    77
  shows "P(lfp f)"
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    78
proof -
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    79
  let ?M = "{S. S \<subseteq> lfp f & P S}"
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    80
  have "P (Union ?M)" using P_Union by simp
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    81
  also have "Union ?M = lfp f"
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    82
  proof
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    83
    show "Union ?M \<subseteq> lfp f" by blast
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    84
    hence "f (Union ?M) \<subseteq> f (lfp f)" by (rule mono [THEN monoD])
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    85
    hence "f (Union ?M) \<subseteq> lfp f" using mono [THEN lfp_unfold] by simp
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    86
    hence "f (Union ?M) \<in> ?M" using P_f P_Union by simp
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    87
    hence "f (Union ?M) \<subseteq> Union ?M" by (rule Union_upper)
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    88
    thus "lfp f \<subseteq> Union ?M" by (rule lfp_lowerbound)
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    89
  qed
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    90
  finally show ?thesis .
9b5073c79a0b Isar-style proof for lfp_ordinal_induct
huffman
parents: 23878
diff changeset
    91
qed
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    92
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    93
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    94
text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    95
    to control unfolding*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    96
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    97
lemma def_lfp_unfold: "[| h==lfp(f);  mono(f) |] ==> h = f(h)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    98
by (auto intro!: lfp_unfold)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
    99
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   100
lemma def_lfp_induct: 
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   101
    "[| A == lfp(f); mono(f);
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   102
        f (inf A P) \<le> P
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   103
     |] ==> A \<le> P"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   104
  by (blast intro: lfp_induct)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   105
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   106
lemma def_lfp_induct_set: 
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   107
    "[| A == lfp(f);  mono(f);   a:A;                    
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   108
        !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   109
     |] ==> P(a)"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   110
  by (blast intro: lfp_induct_set)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   111
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   112
(*Monotonicity of lfp!*)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   113
lemma lfp_mono: "(!!Z. f Z \<le> g Z) ==> lfp f \<le> lfp g"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   114
  by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   115
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   116
22918
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
   117
subsection {* Proof of Knaster-Tarski Theorem using @{term gfp} *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   118
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   119
text{*@{term "gfp f"} is the greatest lower bound of 
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   120
      the set @{term "{u. u \<le> f(u)}"} *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   121
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   122
lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   123
  by (auto simp add: gfp_def intro: Sup_upper)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   124
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   125
lemma gfp_least: "(!!u. u \<le> f u ==> u \<le> X) ==> gfp f \<le> X"
21312
1d39091a3208 started reorgnization of lattice theories
nipkow
parents: 21017
diff changeset
   126
  by (auto simp add: gfp_def intro: Sup_least)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   127
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   128
lemma gfp_lemma2: "mono f ==> gfp f \<le> f (gfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   129
  by (iprover intro: gfp_least order_trans monoD gfp_upperbound)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   130
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   131
lemma gfp_lemma3: "mono f ==> f (gfp f) \<le> gfp f"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   132
  by (iprover intro: gfp_lemma2 monoD gfp_upperbound)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   133
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   134
lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   135
  by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   136
22918
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
   137
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
   138
subsection {* Coinduction rules for greatest fixed points *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   139
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   140
text{*weak version*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   141
lemma weak_coinduct: "[| a: X;  X \<subseteq> f(X) |] ==> a : gfp(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   142
by (rule gfp_upperbound [THEN subsetD], auto)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   143
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   144
lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   145
apply (erule gfp_upperbound [THEN subsetD])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   146
apply (erule imageI)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   147
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   148
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   149
lemma coinduct_lemma:
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   150
     "[| X \<le> f (sup X (gfp f));  mono f |] ==> sup X (gfp f) \<le> f (sup X (gfp f))"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   151
  apply (frule gfp_lemma2)
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   152
  apply (drule mono_sup)
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   153
  apply (rule le_supI)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   154
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   155
  apply (rule order_trans)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   156
  apply (rule order_trans)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   157
  apply assumption
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   158
  apply (rule sup_ge2)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   159
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   160
  done
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   161
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   162
text{*strong version, thanks to Coen and Frost*}
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   163
lemma coinduct_set: "[| mono(f);  a: X;  X \<subseteq> f(X Un gfp(f)) |] ==> a : gfp(f)"
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   164
by (blast intro: weak_coinduct [OF _ coinduct_lemma, simplified sup_set_eq])
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   165
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   166
lemma coinduct: "[| mono(f); X \<le> f (sup X (gfp f)) |] ==> X \<le> gfp(f)"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   167
  apply (rule order_trans)
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   168
  apply (rule sup_ge1)
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   169
  apply (erule gfp_upperbound [OF coinduct_lemma])
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   170
  apply assumption
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   171
  done
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   172
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   173
lemma gfp_fun_UnI2: "[| mono(f);  a: gfp(f) |] ==> a: f(X Un gfp(f))"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   174
by (blast dest: gfp_lemma2 mono_Un)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   175
22918
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
   176
b8b4d53ccd24 localized Sup/Inf
haftmann
parents: 22845
diff changeset
   177
subsection {* Even Stronger Coinduction Rule, by Martin Coen *}
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   178
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   179
text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   180
  @{term lfp} and @{term gfp}*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   181
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   182
lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)"
17589
58eeffd73be1 renamed rules to iprover
nipkow
parents: 17006
diff changeset
   183
by (iprover intro: subset_refl monoI Un_mono monoD)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   184
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   185
lemma coinduct3_lemma:
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   186
     "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)));  mono(f) |]
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   187
      ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   188
apply (rule subset_trans)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   189
apply (erule coinduct3_mono_lemma [THEN lfp_lemma3])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   190
apply (rule Un_least [THEN Un_least])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   191
apply (rule subset_refl, assumption)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   192
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   193
apply (rule monoD, assumption)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   194
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   195
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   196
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   197
lemma coinduct3: 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   198
  "[| mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   199
apply (rule coinduct3_lemma [THEN [2] weak_coinduct])
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   200
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   201
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   202
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   203
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   204
text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   205
    to control unfolding*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   206
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   207
lemma def_gfp_unfold: "[| A==gfp(f);  mono(f) |] ==> A = f(A)"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   208
by (auto intro!: gfp_unfold)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   209
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   210
lemma def_coinduct:
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   211
     "[| A==gfp(f);  mono(f);  X \<le> f(sup X A) |] ==> X \<le> A"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   212
by (iprover intro!: coinduct)
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   213
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   214
lemma def_coinduct_set:
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   215
     "[| A==gfp(f);  mono(f);  a:X;  X \<subseteq> f(X Un A) |] ==> a: A"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   216
by (auto intro!: coinduct_set)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   217
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   218
(*The version used in the induction/coinduction package*)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   219
lemma def_Collect_coinduct:
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   220
    "[| A == gfp(%w. Collect(P(w)));  mono(%w. Collect(P(w)));   
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   221
        a: X;  !!z. z: X ==> P (X Un A) z |] ==>  
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   222
     a : A"
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   223
apply (erule def_coinduct_set, auto) 
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   224
done
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   225
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   226
lemma def_coinduct3:
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   227
    "[| A==gfp(f); mono(f);  a:X;  X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A"
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   228
by (auto intro!: coinduct3)
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   229
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   230
text{*Monotonicity of @{term gfp}!*}
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   231
lemma gfp_mono: "(!!Z. f Z \<le> g Z) ==> gfp f \<le> gfp g"
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   232
  by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans)
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   233
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   234
ML
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   235
{*
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   236
val lfp_def = thm "lfp_def";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   237
val lfp_lowerbound = thm "lfp_lowerbound";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   238
val lfp_greatest = thm "lfp_greatest";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   239
val lfp_unfold = thm "lfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   240
val lfp_induct = thm "lfp_induct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   241
val lfp_ordinal_induct = thm "lfp_ordinal_induct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   242
val def_lfp_unfold = thm "def_lfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   243
val def_lfp_induct = thm "def_lfp_induct";
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   244
val def_lfp_induct_set = thm "def_lfp_induct_set";
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   245
val lfp_mono = thm "lfp_mono";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   246
val gfp_def = thm "gfp_def";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   247
val gfp_upperbound = thm "gfp_upperbound";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   248
val gfp_least = thm "gfp_least";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   249
val gfp_unfold = thm "gfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   250
val weak_coinduct = thm "weak_coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   251
val weak_coinduct_image = thm "weak_coinduct_image";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   252
val coinduct = thm "coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   253
val gfp_fun_UnI2 = thm "gfp_fun_UnI2";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   254
val coinduct3 = thm "coinduct3";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   255
val def_gfp_unfold = thm "def_gfp_unfold";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   256
val def_coinduct = thm "def_coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   257
val def_Collect_coinduct = thm "def_Collect_coinduct";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   258
val def_coinduct3 = thm "def_coinduct3";
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   259
val gfp_mono = thm "gfp_mono";
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   260
val le_funI = thm "le_funI";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   261
val le_boolI = thm "le_boolI";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   262
val le_boolI' = thm "le_boolI'";
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   263
val inf_fun_eq = thm "inf_fun_eq";
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 22390
diff changeset
   264
val inf_bool_eq = thm "inf_bool_eq";
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   265
val le_funE = thm "le_funE";
22276
96a4db55a0b3 Introduction and elimination rules for <= on predicates
berghofe
parents: 21547
diff changeset
   266
val le_funD = thm "le_funD";
21017
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   267
val le_boolE = thm "le_boolE";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   268
val le_boolD = thm "le_boolD";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   269
val le_bool_def = thm "le_bool_def";
5693e4471c2b Generalized gfp and lfp to arbitrary complete lattices.
berghofe
parents: 17589
diff changeset
   270
val le_fun_def = thm "le_fun_def";
17006
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   271
*}
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   272
cffca870816a combined Lfp and Gfp to FixedPoint
avigad
parents:
diff changeset
   273
end