diff -r 9d4d5852eb47 -r 6d3987f3aad9 src/HOL/Lattice/CompleteLattice.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Lattice/CompleteLattice.thy Fri Oct 06 01:04:56 2000 +0200 @@ -0,0 +1,409 @@ +(* Title: HOL/Lattice/CompleteLattice.thy + ID: $Id$ + Author: Markus Wenzel, TU Muenchen +*) + +header {* Complete lattices *} + +theory CompleteLattice = Lattice: + +subsection {* Complete lattice operations *} + +text {* + A \emph{complete lattice} is a partial order with general + (infinitary) infimum of any set of elements. General supremum + exists as well, as a consequence of the connection of infinitary + bounds (see \S\ref{sec:connect-bounds}). +*} + +axclass complete_lattice < partial_order + ex_Inf: "\inf. is_Inf A inf" + +theorem ex_Sup: "\sup::'a::complete_lattice. is_Sup A sup" +proof - + from ex_Inf obtain sup where "is_Inf {b. \a\A. a \ b} sup" by blast + hence "is_Sup A sup" by (rule Inf_Sup) + thus ?thesis .. +qed + +text {* + The general @{text \} (meet) and @{text \} (join) operations select + such infimum and supremum elements. +*} + +consts + Meet :: "'a::complete_lattice set \ 'a" + Join :: "'a::complete_lattice set \ 'a" +syntax (symbols) + Meet :: "'a::complete_lattice set \ 'a" ("\_" [90] 90) + Join :: "'a::complete_lattice set \ 'a" ("\_" [90] 90) +defs + Meet_def: "\A \ SOME inf. is_Inf A inf" + Join_def: "\A \ SOME sup. is_Sup A sup" + +text {* + Due to unique existence of bounds, the complete lattice operations + may be exhibited as follows. +*} + +lemma Meet_equality [elim?]: "is_Inf A inf \ \A = inf" +proof (unfold Meet_def) + assume "is_Inf A inf" + thus "(SOME inf. is_Inf A inf) = inf" + by (rule some_equality) (rule is_Inf_uniq) +qed + +lemma MeetI [intro?]: + "(\a. a \ A \ inf \ a) \ + (\b. \a \ A. b \ a \ b \ inf) \ + \A = inf" + by (rule Meet_equality, rule is_InfI) blast+ + +lemma Join_equality [elim?]: "is_Sup A sup \ \A = sup" +proof (unfold Join_def) + assume "is_Sup A sup" + thus "(SOME sup. is_Sup A sup) = sup" + by (rule some_equality) (rule is_Sup_uniq) +qed + +lemma JoinI [intro?]: + "(\a. a \ A \ a \ sup) \ + (\b. \a \ A. a \ b \ sup \ b) \ + \A = sup" + by (rule Join_equality, rule is_SupI) blast+ + + +text {* + \medskip The @{text \} and @{text \} operations indeed determine + bounds on a complete lattice structure. +*} + +lemma is_Inf_Meet [intro?]: "is_Inf A (\A)" +proof (unfold Meet_def) + from ex_Inf show "is_Inf A (SOME inf. is_Inf A inf)" + by (rule ex_someI) +qed + +lemma Meet_greatest [intro?]: "(\a. a \ A \ x \ a) \ x \ \A" + by (rule is_Inf_greatest, rule is_Inf_Meet) blast + +lemma Meet_lower [intro?]: "a \ A \ \A \ a" + by (rule is_Inf_lower) (rule is_Inf_Meet) + + +lemma is_Sup_Join [intro?]: "is_Sup A (\A)" +proof (unfold Join_def) + from ex_Sup show "is_Sup A (SOME sup. is_Sup A sup)" + by (rule ex_someI) +qed + +lemma Join_least [intro?]: "(\a. a \ A \ a \ x) \ \A \ x" + by (rule is_Sup_least, rule is_Sup_Join) blast +lemma Join_lower [intro?]: "a \ A \ a \ \A" + by (rule is_Sup_upper) (rule is_Sup_Join) + + +subsection {* The Knaster-Tarski Theorem *} + +text {* + The Knaster-Tarski Theorem (in its simplest formulation) states that + any monotone function on a complete lattice has a least fixed-point + (see \cite[pages 93--94]{Davey-Priestley:1990} for example). This + is a consequence of the basic boundary properties of the complete + lattice operations. +*} + +theorem Knaster_Tarski: + "(\x y. x \ y \ f x \ f y) \ \a::'a::complete_lattice. f a = a" +proof + assume mono: "\x y. x \ y \ f x \ f y" + let ?H = "{u. f u \ u}" let ?a = "\?H" + have ge: "f ?a \ ?a" + proof + fix x assume x: "x \ ?H" + hence "?a \ x" .. + hence "f ?a \ f x" by (rule mono) + also from x have "... \ x" .. + finally show "f ?a \ x" . + qed + also have "?a \ f ?a" + proof + from ge have "f (f ?a) \ f ?a" by (rule mono) + thus "f ?a : ?H" .. + qed + finally show "f ?a = ?a" . +qed + + +subsection {* Bottom and top elements *} + +text {* + With general bounds available, complete lattices also have least and + greatest elements. +*} + +constdefs + bottom :: "'a::complete_lattice" ("\") + "\ \ \UNIV" + top :: "'a::complete_lattice" ("\") + "\ \ \UNIV" + +lemma bottom_least [intro?]: "\ \ x" +proof (unfold bottom_def) + have "x \ UNIV" .. + thus "\UNIV \ x" .. +qed + +lemma bottomI [intro?]: "(\a. x \ a) \ \ = x" +proof (unfold bottom_def) + assume "\a. x \ a" + show "\UNIV = x" + proof + fix a show "x \ a" . + next + fix b :: "'a::complete_lattice" + assume b: "\a \ UNIV. b \ a" + have "x \ UNIV" .. + with b show "b \ x" .. + qed +qed + +lemma top_greatest [intro?]: "x \ \" +proof (unfold top_def) + have "x \ UNIV" .. + thus "x \ \UNIV" .. +qed + +lemma topI [intro?]: "(\a. a \ x) \ \ = x" +proof (unfold top_def) + assume "\a. a \ x" + show "\UNIV = x" + proof + fix a show "a \ x" . + next + fix b :: "'a::complete_lattice" + assume b: "\a \ UNIV. a \ b" + have "x \ UNIV" .. + with b show "x \ b" .. + qed +qed + + +subsection {* Duality *} + +text {* + The class of complete lattices is closed under formation of dual + structures. +*} + +instance dual :: (complete_lattice) complete_lattice +proof intro_classes + fix A' :: "'a::complete_lattice dual set" + show "\inf'. is_Inf A' inf'" + proof - + have "\sup. is_Sup (undual `` A') sup" by (rule ex_Sup) + hence "\sup. is_Inf (dual `` undual `` A') (dual sup)" by (simp only: dual_Inf) + thus ?thesis by (simp add: dual_ex [symmetric] image_compose [symmetric]) + qed +qed + +text {* + Apparently, the @{text \} and @{text \} operations are dual to each + other. +*} + +theorem dual_Meet [intro?]: "dual (\A) = \(dual `` A)" +proof - + from is_Inf_Meet have "is_Sup (dual `` A) (dual (\A))" .. + hence "\(dual `` A) = dual (\A)" .. + thus ?thesis .. +qed + +theorem dual_Join [intro?]: "dual (\A) = \(dual `` A)" +proof - + from is_Sup_Join have "is_Inf (dual `` A) (dual (\A))" .. + hence "\(dual `` A) = dual (\A)" .. + thus ?thesis .. +qed + +text {* + Likewise are @{text \} and @{text \} duals of each other. +*} + +theorem dual_bottom [intro?]: "dual \ = \" +proof - + have "\ = dual \" + proof + fix a' have "\ \ undual a'" .. + hence "dual (undual a') \ dual \" .. + thus "a' \ dual \" by simp + qed + thus ?thesis .. +qed + +theorem dual_top [intro?]: "dual \ = \" +proof - + have "\ = dual \" + proof + fix a' have "undual a' \ \" .. + hence "dual \ \ dual (undual a')" .. + thus "dual \ \ a'" by simp + qed + thus ?thesis .. +qed + + +subsection {* Complete lattices are lattices *} + +text {* + Complete lattices (with general bounds) available are indeed plain + lattices as well. This holds due to the connection of general + versus binary bounds that has been formally established in + \S\ref{sec:gen-bin-bounds}. +*} + +lemma is_inf_binary: "is_inf x y (\{x, y})" +proof - + have "is_Inf {x, y} (\{x, y})" .. + thus ?thesis by (simp only: is_Inf_binary) +qed + +lemma is_sup_binary: "is_sup x y (\{x, y})" +proof - + have "is_Sup {x, y} (\{x, y})" .. + thus ?thesis by (simp only: is_Sup_binary) +qed + +instance complete_lattice < lattice +proof intro_classes + fix x y :: "'a::complete_lattice" + from is_inf_binary show "\inf. is_inf x y inf" .. + from is_sup_binary show "\sup. is_sup x y sup" .. +qed + +theorem meet_binary: "x \ y = \{x, y}" + by (rule meet_equality) (rule is_inf_binary) + +theorem join_binary: "x \ y = \{x, y}" + by (rule join_equality) (rule is_sup_binary) + + +subsection {* Complete lattices and set-theory operations *} + +text {* + The complete lattice operations are (anti) monotone wrt.\ set + inclusion. +*} + +theorem Meet_subset_antimono: "A \ B \ \B \ \A" +proof (rule Meet_greatest) + fix a assume "a \ A" + also assume "A \ B" + finally have "a \ B" . + thus "\B \ a" .. +qed + +theorem Join_subset_mono: "A \ B \ \A \ \B" +proof - + assume "A \ B" + hence "dual `` A \ dual `` B" by blast + hence "\(dual `` B) \ \(dual `` A)" by (rule Meet_subset_antimono) + hence "dual (\B) \ dual (\A)" by (simp only: dual_Join) + thus ?thesis by (simp only: dual_leq) +qed + +text {* + Bounds over unions of sets may be obtained separately. +*} + +theorem Meet_Un: "\(A \ B) = \A \ \B" +proof + fix a assume "a \ A \ B" + thus "\A \ \B \ a" + proof + assume a: "a \ A" + have "\A \ \B \ \A" .. + also from a have "\ \ a" .. + finally show ?thesis . + next + assume a: "a \ B" + have "\A \ \B \ \B" .. + also from a have "\ \ a" .. + finally show ?thesis . + qed +next + fix b assume b: "\a \ A \ B. b \ a" + show "b \ \A \ \B" + proof + show "b \ \A" + proof + fix a assume "a \ A" + hence "a \ A \ B" .. + with b show "b \ a" .. + qed + show "b \ \B" + proof + fix a assume "a \ B" + hence "a \ A \ B" .. + with b show "b \ a" .. + qed + qed +qed + +theorem Join_Un: "\(A \ B) = \A \ \B" +proof - + have "dual (\(A \ B)) = \(dual `` A \ dual `` B)" + by (simp only: dual_Join image_Un) + also have "\ = \(dual `` A) \ \(dual `` B)" + by (rule Meet_Un) + also have "\ = dual (\A \ \B)" + by (simp only: dual_join dual_Join) + finally show ?thesis .. +qed + +text {* + Bounds over singleton sets are trivial. +*} + +theorem Meet_singleton: "\{x} = x" +proof + fix a assume "a \ {x}" + hence "a = x" by simp + thus "x \ a" by (simp only: leq_refl) +next + fix b assume "\a \ {x}. b \ a" + thus "b \ x" by simp +qed + +theorem Join_singleton: "\{x} = x" +proof - + have "dual (\{x}) = \{dual x}" by (simp add: dual_Join) + also have "\ = dual x" by (rule Meet_singleton) + finally show ?thesis .. +qed + +text {* + Bounds over the empty and universal set correspond to each other. +*} + +theorem Meet_empty: "\{} = \UNIV" +proof + fix a :: "'a::complete_lattice" + assume "a \ {}" + hence False by simp + thus "\UNIV \ a" .. +next + fix b :: "'a::complete_lattice" + have "b \ UNIV" .. + thus "b \ \UNIV" .. +qed + +theorem Join_empty: "\{} = \UNIV" +proof - + have "dual (\{}) = \{}" by (simp add: dual_Join) + also have "\ = \UNIV" by (rule Meet_empty) + also have "\ = dual (\UNIV)" by (simp add: dual_Meet) + finally show ?thesis .. +qed + +end