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