| author | wenzelm | 
| Thu, 10 May 2007 00:39:51 +0200 | |
| changeset 22905 | dab6a898b47c | 
| parent 22845 | 5f9138bcb3d7 | 
| child 22918 | b8b4d53ccd24 | 
| permissions | -rw-r--r-- | 
| 17006 | 1 | (* Title: HOL/FixedPoint.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 4 | Author: Stefan Berghofer, TU Muenchen | 
| 17006 | 5 | Copyright 1992 University of Cambridge | 
| 6 | *) | |
| 7 | ||
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 8 | header {* Fixed Points and the Knaster-Tarski Theorem*}
 | 
| 17006 | 9 | |
| 10 | theory FixedPoint | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 11 | imports Product_Type | 
| 17006 | 12 | begin | 
| 13 | ||
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 14 | subsection {* Complete lattices *}
 | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 15 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 16 | class complete_lattice = lattice + | 
| 22845 | 17 |   fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [90] 90)
 | 
| 18 | assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x" | |
| 19 | assumes Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A" | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 20 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 21 | definition | 
| 22845 | 22 | Sup :: "'a\<Colon>complete_lattice set \<Rightarrow> 'a" | 
| 23 | where | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 24 |   "Sup A = Inf {b. \<forall>a \<in> A. a \<le> b}"
 | 
| 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 25 | |
| 22845 | 26 | theorem Sup_upper: "(x::'a::complete_lattice) \<in> A \<Longrightarrow> x \<le> Sup A" | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 27 | by (auto simp: Sup_def intro: Inf_greatest) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 28 | |
| 22845 | 29 | theorem Sup_least: "(\<And>x::'a::complete_lattice. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup A \<le> z" | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 30 | by (auto simp: Sup_def intro: Inf_lower) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 31 | |
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 32 | definition | 
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 33 |   SUPR :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b" where
 | 
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 34 | "SUPR A f == Sup (f ` A)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 35 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 36 | definition | 
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 37 |   INFI :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b" where
 | 
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 38 | "INFI A f == Inf (f ` A)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 39 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 40 | syntax | 
| 22439 | 41 |   "_SUP1"     :: "pttrns => 'b => 'b"           ("(3SUP _./ _)" [0, 10] 10)
 | 
| 42 |   "_SUP"      :: "pttrn => 'a set => 'b => 'b"  ("(3SUP _:_./ _)" [0, 10] 10)
 | |
| 43 |   "_INF1"     :: "pttrns => 'b => 'b"           ("(3INF _./ _)" [0, 10] 10)
 | |
| 44 |   "_INF"      :: "pttrn => 'a set => 'b => 'b"  ("(3INF _:_./ _)" [0, 10] 10)
 | |
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 45 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 46 | translations | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 47 | "SUP x y. B" == "SUP x. SUP y. B" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 48 | "SUP x. B" == "CONST SUPR UNIV (%x. B)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 49 | "SUP x. B" == "SUP x:UNIV. B" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 50 | "SUP x:A. B" == "CONST SUPR A (%x. B)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 51 | "INF x y. B" == "INF x. INF y. B" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 52 | "INF x. B" == "CONST INFI UNIV (%x. B)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 53 | "INF x. B" == "INF x:UNIV. B" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 54 | "INF x:A. B" == "CONST INFI A (%x. B)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 55 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 56 | (* To avoid eta-contraction of body: *) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 57 | print_translation {*
 | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 58 | let | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 59 | fun btr' syn (A :: Abs abs :: ts) = | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 60 | let val (x,t) = atomic_abs_tr' abs | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 61 | in list_comb (Syntax.const syn $ x $ A $ t, ts) end | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 62 |   val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const
 | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 63 | in | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 64 | [(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")]
 | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 65 | end | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 66 | *} | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 67 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 68 | lemma le_SUPI: "i : A \<Longrightarrow> M i \<le> (SUP i:A. M i)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 69 | by (auto simp add: SUPR_def intro: Sup_upper) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 70 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 71 | lemma SUP_leI: "(\<And>i. i : A \<Longrightarrow> M i \<le> u) \<Longrightarrow> (SUP i:A. M i) \<le> u" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 72 | by (auto simp add: SUPR_def intro: Sup_least) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 73 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 74 | lemma INF_leI: "i : A \<Longrightarrow> (INF i:A. M i) \<le> M i" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 75 | by (auto simp add: INFI_def intro: Inf_lower) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 76 | |
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 77 | lemma le_INFI: "(\<And>i. i : A \<Longrightarrow> u \<le> M i) \<Longrightarrow> u \<le> (INF i:A. M i)" | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 78 | by (auto simp add: INFI_def intro: Inf_greatest) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 79 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 80 | text {* A complete lattice is a lattice *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 81 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 82 | |
| 21312 | 83 | subsubsection {* Properties *}
 | 
| 84 | ||
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 85 | lemma mono_inf: "mono f \<Longrightarrow> f (inf A B) <= inf (f A) (f B)" | 
| 21312 | 86 | by (auto simp add: mono_def) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 87 | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 88 | lemma mono_sup: "mono f \<Longrightarrow> sup (f A) (f B) <= f (sup A B)" | 
| 21312 | 89 | by (auto simp add: mono_def) | 
| 90 | ||
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 91 | lemma Sup_insert[simp]: "Sup (insert (a::'a::complete_lattice) A) = sup a (Sup A)" | 
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 92 | apply (rule order_antisym) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 93 | apply (rule Sup_least) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 94 | apply (erule insertE) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 95 | apply (rule le_supI1) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 96 | apply simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 97 | apply (rule le_supI2) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 98 | apply (erule Sup_upper) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 99 | apply (rule le_supI) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 100 | apply (rule Sup_upper) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 101 | apply simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 102 | apply (rule Sup_least) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 103 | apply (rule Sup_upper) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 104 | apply simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 105 | done | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 106 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 107 | lemma Inf_insert[simp]: "Inf (insert (a::'a::complete_lattice) A) = inf a (Inf A)" | 
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 108 | apply (rule order_antisym) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 109 | apply (rule le_infI) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 110 | apply (rule Inf_lower) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 111 | apply simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 112 | apply (rule Inf_greatest) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 113 | apply (rule Inf_lower) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 114 | apply simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 115 | apply (rule Inf_greatest) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 116 | apply (erule insertE) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 117 | apply (rule le_infI1) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 118 | apply simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 119 | apply (rule le_infI2) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 120 | apply (erule Inf_lower) | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 121 | done | 
| 21312 | 122 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 123 | lemma bot_least[simp]: "Sup{} \<le> (x::'a::complete_lattice)"
 | 
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 124 | by (rule Sup_least) simp | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 125 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 126 | lemma top_greatest[simp]: "(x::'a::complete_lattice) \<le> Inf{}"
 | 
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 127 | by (rule Inf_greatest) simp | 
| 21312 | 128 | |
| 22477 | 129 | lemma inf_Inf_empty: | 
| 130 |   "inf a (Inf {}) = a"
 | |
| 131 | proof - | |
| 132 |   have "a \<le> Inf {}" by (rule top_greatest)
 | |
| 133 | then show ?thesis by (rule inf_absorb1) | |
| 134 | qed | |
| 135 | ||
| 136 | lemma inf_binary: | |
| 137 |   "Inf {a, b} = inf a b"
 | |
| 138 | unfolding Inf_insert inf_Inf_empty .. | |
| 139 | ||
| 140 | lemma sup_Sup_empty: | |
| 141 |   "sup a (Sup {}) = a"
 | |
| 142 | proof - | |
| 143 |   have "Sup {} \<le> a" by (rule bot_least)
 | |
| 144 | then show ?thesis by (rule sup_absorb1) | |
| 145 | qed | |
| 146 | ||
| 147 | lemma sup_binary: | |
| 148 |   "Sup {a, b} = sup a b"
 | |
| 149 | unfolding Sup_insert sup_Sup_empty .. | |
| 150 | ||
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 151 | lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M"
 | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 152 | by (auto intro: order_antisym SUP_leI le_SUPI) | 
| 21312 | 153 | |
| 22430 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 154 | lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
 | 
| 
6a56bf1b3a64
Generalized version of SUP and INF (with index set).
 berghofe parents: 
22422diff
changeset | 155 | by (auto intro: order_antisym INF_leI le_INFI) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 156 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 157 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 158 | subsection {* Some instances of the type class of complete lattices *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 159 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 160 | subsubsection {* Booleans *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 161 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 162 | instance bool :: complete_lattice | 
| 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 163 | Inf_bool_def: "Inf A \<equiv> \<forall>x\<in>A. x" | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 164 | apply intro_classes | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 165 | apply (unfold Inf_bool_def) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 166 | apply (iprover intro!: le_boolI elim: ballE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 167 | apply (iprover intro!: ballI le_boolI elim: ballE le_boolE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 168 | done | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 169 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 170 | theorem Sup_bool_eq: "Sup A \<longleftrightarrow> (\<exists>x\<in>A. x)" | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 171 | apply (rule order_antisym) | 
| 21312 | 172 | apply (rule Sup_least) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 173 | apply (rule le_boolI) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 174 | apply (erule bexI, assumption) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 175 | apply (rule le_boolI) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 176 | apply (erule bexE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 177 | apply (rule le_boolE) | 
| 21312 | 178 | apply (rule Sup_upper) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 179 | apply assumption+ | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 180 | done | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 181 | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 182 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 183 | subsubsection {* Functions *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 184 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 185 | instance "fun" :: (type, complete_lattice) complete_lattice | 
| 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 186 |   Inf_fun_def: "Inf A \<equiv> (\<lambda>x. Inf {y. \<exists>f\<in>A. y = f x})"
 | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 187 | apply intro_classes | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 188 | apply (unfold Inf_fun_def) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 189 | apply (rule le_funI) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 190 | apply (rule Inf_lower) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 191 | apply (rule CollectI) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 192 | apply (rule bexI) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 193 | apply (rule refl) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 194 | apply assumption | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 195 | apply (rule le_funI) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 196 | apply (rule Inf_greatest) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 197 | apply (erule CollectE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 198 | apply (erule bexE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 199 | apply (iprover elim: le_funE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 200 | done | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 201 | |
| 22845 | 202 | lemmas [code func del] = Inf_fun_def | 
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22477diff
changeset | 203 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 204 | theorem Sup_fun_eq: "Sup A = (\<lambda>x. Sup {y. \<exists>f\<in>A. y = f x})"
 | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 205 | apply (rule order_antisym) | 
| 21312 | 206 | apply (rule Sup_least) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 207 | apply (rule le_funI) | 
| 21312 | 208 | apply (rule Sup_upper) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 209 | apply fast | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 210 | apply (rule le_funI) | 
| 21312 | 211 | apply (rule Sup_least) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 212 | apply (erule CollectE) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 213 | apply (erule bexE) | 
| 21312 | 214 | apply (drule le_funD [OF Sup_upper]) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 215 | apply simp | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 216 | done | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 217 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 218 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 219 | subsubsection {* Sets *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 220 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 221 | instance set :: (type) complete_lattice | 
| 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 222 | Inf_set_def: "Inf S \<equiv> \<Inter>S" | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 223 | by intro_classes (auto simp add: Inf_set_def) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 224 | |
| 22845 | 225 | lemmas [code func del] = Inf_set_def | 
| 22744 
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
 haftmann parents: 
22477diff
changeset | 226 | |
| 21312 | 227 | theorem Sup_set_eq: "Sup S = \<Union>S" | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 228 | apply (rule subset_antisym) | 
| 21312 | 229 | apply (rule Sup_least) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 230 | apply (erule Union_upper) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 231 | apply (rule Union_least) | 
| 21312 | 232 | apply (erule Sup_upper) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 233 | done | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 234 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 235 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 236 | subsection {* Least and greatest fixed points *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 237 | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 238 | definition | 
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 239 |   lfp :: "('a\<Colon>complete_lattice \<Rightarrow> 'a) \<Rightarrow> 'a" where
 | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 240 |   "lfp f = Inf {u. f u \<le> u}"    --{*least fixed point*}
 | 
| 17006 | 241 | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 242 | definition | 
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 243 |   gfp :: "('a\<Colon>complete_lattice \<Rightarrow> 'a) \<Rightarrow> 'a" where
 | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 244 |   "gfp f = Sup {u. u \<le> f u}"    --{*greatest fixed point*}
 | 
| 17006 | 245 | |
| 246 | ||
| 247 | subsection{*Proof of Knaster-Tarski Theorem using @{term lfp}*}
 | |
| 248 | ||
| 249 | text{*@{term "lfp f"} is the least upper bound of 
 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 250 |       the set @{term "{u. f(u) \<le> u}"} *}
 | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 251 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 252 | lemma lfp_lowerbound: "f A \<le> A ==> lfp f \<le> A" | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 253 | by (auto simp add: lfp_def intro: Inf_lower) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 254 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 255 | 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: 
22390diff
changeset | 256 | by (auto simp add: lfp_def intro: Inf_greatest) | 
| 17006 | 257 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 258 | lemma lfp_lemma2: "mono f ==> f (lfp f) \<le> lfp f" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 259 | by (iprover intro: lfp_greatest order_trans monoD lfp_lowerbound) | 
| 17006 | 260 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 261 | lemma lfp_lemma3: "mono f ==> lfp f \<le> f (lfp f)" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 262 | by (iprover intro: lfp_lemma2 monoD lfp_lowerbound) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 263 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 264 | lemma lfp_unfold: "mono f ==> lfp f = f (lfp f)" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 265 | by (iprover intro: order_antisym lfp_lemma2 lfp_lemma3) | 
| 17006 | 266 | |
| 22356 | 267 | lemma lfp_const: "lfp (\<lambda>x. t) = t" | 
| 268 | by (rule lfp_unfold) (simp add:mono_def) | |
| 269 | ||
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 270 | subsection{*General induction rules for least fixed points*}
 | 
| 17006 | 271 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 272 | theorem lfp_induct: | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 273 | assumes mono: "mono f" and ind: "f (inf (lfp f) P) <= P" | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 274 | shows "lfp f <= P" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 275 | proof - | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 276 | have "inf (lfp f) P <= lfp f" by (rule inf_le1) | 
| 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 277 | with mono have "f (inf (lfp f) P) <= f (lfp f)" .. | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 278 | 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: 
22390diff
changeset | 279 | finally have "f (inf (lfp f) P) <= lfp f" . | 
| 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 280 | 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: 
22390diff
changeset | 281 | hence "lfp f <= inf (lfp f) P" by (rule lfp_lowerbound) | 
| 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 282 | also have "inf (lfp f) P <= P" by (rule inf_le2) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 283 | finally show ?thesis . | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 284 | qed | 
| 17006 | 285 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 286 | lemma lfp_induct_set: | 
| 17006 | 287 | assumes lfp: "a: lfp(f)" | 
| 288 | and mono: "mono(f)" | |
| 289 |       and indhyp: "!!x. [| x: f(lfp(f) Int {x. P(x)}) |] ==> P(x)"
 | |
| 290 | shows "P(a)" | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 291 | by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 292 | (auto simp: inf_set_eq intro: indhyp) | 
| 17006 | 293 | |
| 22452 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 294 | text {* Version of induction for binary relations *}
 | 
| 
8a86fd2a1bf0
adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
 haftmann parents: 
22439diff
changeset | 295 | lemmas lfp_induct2 = lfp_induct_set [of "(a, b)", split_format (complete)] | 
| 17006 | 296 | |
| 297 | lemma lfp_ordinal_induct: | |
| 298 | assumes mono: "mono f" | |
| 299 | shows "[| !!S. P S ==> P(f S); !!M. !S:M. P S ==> P(Union M) |] | |
| 300 | ==> P(lfp f)" | |
| 301 | apply(subgoal_tac "lfp f = Union{S. S \<subseteq> lfp f & P S}")
 | |
| 302 | apply (erule ssubst, simp) | |
| 303 | apply(subgoal_tac "Union{S. S \<subseteq> lfp f & P S} \<subseteq> lfp f")
 | |
| 304 | prefer 2 apply blast | |
| 305 | apply(rule equalityI) | |
| 306 | prefer 2 apply assumption | |
| 307 | apply(drule mono [THEN monoD]) | |
| 308 | apply (cut_tac mono [THEN lfp_unfold], simp) | |
| 309 | apply (rule lfp_lowerbound, auto) | |
| 310 | done | |
| 311 | ||
| 312 | ||
| 313 | text{*Definition forms of @{text lfp_unfold} and @{text lfp_induct}, 
 | |
| 314 | to control unfolding*} | |
| 315 | ||
| 316 | lemma def_lfp_unfold: "[| h==lfp(f); mono(f) |] ==> h = f(h)" | |
| 317 | by (auto intro!: lfp_unfold) | |
| 318 | ||
| 319 | lemma def_lfp_induct: | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 320 | "[| A == lfp(f); mono(f); | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 321 | f (inf A P) \<le> P | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 322 | |] ==> A \<le> P" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 323 | by (blast intro: lfp_induct) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 324 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 325 | lemma def_lfp_induct_set: | 
| 17006 | 326 | "[| A == lfp(f); mono(f); a:A; | 
| 327 |         !!x. [| x: f(A Int {x. P(x)}) |] ==> P(x)         
 | |
| 328 | |] ==> P(a)" | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 329 | by (blast intro: lfp_induct_set) | 
| 17006 | 330 | |
| 331 | (*Monotonicity of lfp!*) | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 332 | 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: 
17589diff
changeset | 333 | by (rule lfp_lowerbound [THEN lfp_greatest], blast intro: order_trans) | 
| 17006 | 334 | |
| 335 | ||
| 336 | subsection{*Proof of Knaster-Tarski Theorem using @{term gfp}*}
 | |
| 337 | ||
| 338 | ||
| 339 | text{*@{term "gfp f"} is the greatest lower bound of 
 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 340 |       the set @{term "{u. u \<le> f(u)}"} *}
 | 
| 17006 | 341 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 342 | lemma gfp_upperbound: "X \<le> f X ==> X \<le> gfp f" | 
| 21312 | 343 | by (auto simp add: gfp_def intro: Sup_upper) | 
| 17006 | 344 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 345 | lemma gfp_least: "(!!u. u \<le> f u ==> u \<le> X) ==> gfp f \<le> X" | 
| 21312 | 346 | by (auto simp add: gfp_def intro: Sup_least) | 
| 17006 | 347 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 348 | lemma gfp_lemma2: "mono f ==> gfp f \<le> f (gfp f)" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 349 | by (iprover intro: gfp_least order_trans monoD gfp_upperbound) | 
| 17006 | 350 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 351 | lemma gfp_lemma3: "mono f ==> f (gfp f) \<le> gfp f" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 352 | by (iprover intro: gfp_lemma2 monoD gfp_upperbound) | 
| 17006 | 353 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 354 | lemma gfp_unfold: "mono f ==> gfp f = f (gfp f)" | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 355 | by (iprover intro: order_antisym gfp_lemma2 gfp_lemma3) | 
| 17006 | 356 | |
| 357 | subsection{*Coinduction rules for greatest fixed points*}
 | |
| 358 | ||
| 359 | text{*weak version*}
 | |
| 360 | lemma weak_coinduct: "[| a: X; X \<subseteq> f(X) |] ==> a : gfp(f)" | |
| 361 | by (rule gfp_upperbound [THEN subsetD], auto) | |
| 362 | ||
| 363 | lemma weak_coinduct_image: "!!X. [| a : X; g`X \<subseteq> f (g`X) |] ==> g a : gfp f" | |
| 364 | apply (erule gfp_upperbound [THEN subsetD]) | |
| 365 | apply (erule imageI) | |
| 366 | done | |
| 367 | ||
| 368 | lemma coinduct_lemma: | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 369 | "[| 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: 
17589diff
changeset | 370 | apply (frule gfp_lemma2) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 371 | apply (drule mono_sup) | 
| 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 372 | apply (rule le_supI) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 373 | apply assumption | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 374 | apply (rule order_trans) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 375 | apply (rule order_trans) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 376 | apply assumption | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 377 | apply (rule sup_ge2) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 378 | apply assumption | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 379 | done | 
| 17006 | 380 | |
| 381 | text{*strong version, thanks to Coen and Frost*}
 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 382 | 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: 
22390diff
changeset | 383 | by (blast intro: weak_coinduct [OF _ coinduct_lemma, simplified sup_set_eq]) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 384 | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 385 | 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: 
17589diff
changeset | 386 | apply (rule order_trans) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 387 | apply (rule sup_ge1) | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 388 | apply (erule gfp_upperbound [OF coinduct_lemma]) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 389 | apply assumption | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 390 | done | 
| 17006 | 391 | |
| 392 | lemma gfp_fun_UnI2: "[| mono(f); a: gfp(f) |] ==> a: f(X Un gfp(f))" | |
| 393 | by (blast dest: gfp_lemma2 mono_Un) | |
| 394 | ||
| 395 | subsection{*Even Stronger Coinduction Rule, by Martin Coen*}
 | |
| 396 | ||
| 397 | text{* Weakens the condition @{term "X \<subseteq> f(X)"} to one expressed using both
 | |
| 398 |   @{term lfp} and @{term gfp}*}
 | |
| 399 | ||
| 400 | lemma coinduct3_mono_lemma: "mono(f) ==> mono(%x. f(x) Un X Un B)" | |
| 17589 | 401 | by (iprover intro: subset_refl monoI Un_mono monoD) | 
| 17006 | 402 | |
| 403 | lemma coinduct3_lemma: | |
| 404 | "[| X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))); mono(f) |] | |
| 405 | ==> lfp(%x. f(x) Un X Un gfp(f)) \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f)))" | |
| 406 | apply (rule subset_trans) | |
| 407 | apply (erule coinduct3_mono_lemma [THEN lfp_lemma3]) | |
| 408 | apply (rule Un_least [THEN Un_least]) | |
| 409 | apply (rule subset_refl, assumption) | |
| 410 | apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) | |
| 411 | apply (rule monoD, assumption) | |
| 412 | apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) | |
| 413 | done | |
| 414 | ||
| 415 | lemma coinduct3: | |
| 416 | "[| mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)" | |
| 417 | apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) | |
| 418 | apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto) | |
| 419 | done | |
| 420 | ||
| 421 | ||
| 422 | text{*Definition forms of @{text gfp_unfold} and @{text coinduct}, 
 | |
| 423 | to control unfolding*} | |
| 424 | ||
| 425 | lemma def_gfp_unfold: "[| A==gfp(f); mono(f) |] ==> A = f(A)" | |
| 426 | by (auto intro!: gfp_unfold) | |
| 427 | ||
| 428 | lemma def_coinduct: | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 429 | "[| 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: 
17589diff
changeset | 430 | by (iprover intro!: coinduct) | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 431 | |
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 432 | lemma def_coinduct_set: | 
| 17006 | 433 | "[| 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: 
17589diff
changeset | 434 | by (auto intro!: coinduct_set) | 
| 17006 | 435 | |
| 436 | (*The version used in the induction/coinduction package*) | |
| 437 | lemma def_Collect_coinduct: | |
| 438 | "[| A == gfp(%w. Collect(P(w))); mono(%w. Collect(P(w))); | |
| 439 | a: X; !!z. z: X ==> P (X Un A) z |] ==> | |
| 440 | a : A" | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 441 | apply (erule def_coinduct_set, auto) | 
| 17006 | 442 | done | 
| 443 | ||
| 444 | lemma def_coinduct3: | |
| 445 | "[| A==gfp(f); mono(f); a:X; X \<subseteq> f(lfp(%x. f(x) Un X Un A)) |] ==> a: A" | |
| 446 | by (auto intro!: coinduct3) | |
| 447 | ||
| 448 | text{*Monotonicity of @{term gfp}!*}
 | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 449 | 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: 
17589diff
changeset | 450 | by (rule gfp_upperbound [THEN gfp_least], blast intro: order_trans) | 
| 17006 | 451 | |
| 452 | ML | |
| 453 | {*
 | |
| 454 | val lfp_def = thm "lfp_def"; | |
| 455 | val lfp_lowerbound = thm "lfp_lowerbound"; | |
| 456 | val lfp_greatest = thm "lfp_greatest"; | |
| 457 | val lfp_unfold = thm "lfp_unfold"; | |
| 458 | val lfp_induct = thm "lfp_induct"; | |
| 459 | val lfp_induct2 = thm "lfp_induct2"; | |
| 460 | val lfp_ordinal_induct = thm "lfp_ordinal_induct"; | |
| 461 | val def_lfp_unfold = thm "def_lfp_unfold"; | |
| 462 | val def_lfp_induct = thm "def_lfp_induct"; | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 463 | val def_lfp_induct_set = thm "def_lfp_induct_set"; | 
| 17006 | 464 | val lfp_mono = thm "lfp_mono"; | 
| 465 | val gfp_def = thm "gfp_def"; | |
| 466 | val gfp_upperbound = thm "gfp_upperbound"; | |
| 467 | val gfp_least = thm "gfp_least"; | |
| 468 | val gfp_unfold = thm "gfp_unfold"; | |
| 469 | val weak_coinduct = thm "weak_coinduct"; | |
| 470 | val weak_coinduct_image = thm "weak_coinduct_image"; | |
| 471 | val coinduct = thm "coinduct"; | |
| 472 | val gfp_fun_UnI2 = thm "gfp_fun_UnI2"; | |
| 473 | val coinduct3 = thm "coinduct3"; | |
| 474 | val def_gfp_unfold = thm "def_gfp_unfold"; | |
| 475 | val def_coinduct = thm "def_coinduct"; | |
| 476 | val def_Collect_coinduct = thm "def_Collect_coinduct"; | |
| 477 | val def_coinduct3 = thm "def_coinduct3"; | |
| 478 | val gfp_mono = thm "gfp_mono"; | |
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 479 | val le_funI = thm "le_funI"; | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 480 | val le_boolI = thm "le_boolI"; | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 481 | val le_boolI' = thm "le_boolI'"; | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 482 | val inf_fun_eq = thm "inf_fun_eq"; | 
| 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22390diff
changeset | 483 | val inf_bool_eq = thm "inf_bool_eq"; | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 484 | val le_funE = thm "le_funE"; | 
| 22276 
96a4db55a0b3
Introduction and elimination rules for <= on predicates
 berghofe parents: 
21547diff
changeset | 485 | val le_funD = thm "le_funD"; | 
| 21017 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 486 | val le_boolE = thm "le_boolE"; | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 487 | val le_boolD = thm "le_boolD"; | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 488 | val le_bool_def = thm "le_bool_def"; | 
| 
5693e4471c2b
Generalized gfp and lfp to arbitrary complete lattices.
 berghofe parents: 
17589diff
changeset | 489 | val le_fun_def = thm "le_fun_def"; | 
| 17006 | 490 | *} | 
| 491 | ||
| 492 | end |