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