# HG changeset patch # User berghofe # Date 1210150596 -7200 # Node ID 354c3844dfde9e5bcb196934b87680d644ee99b9 # Parent e36a92ff543e291a9909d61841a275092f5ec055 - Now imports Fun rather than Orderings - Moved "Set as lattice" section behind "Fun as lattice" section, since sets are just functions. - The instantiations instantiation set :: (type) distrib_lattice instantiation set :: (type) complete_lattice are no longer needed, and the former definitions inf_set_eq, sup_set_eq, Inf_set_def, and Sup_set_def can now be derived from abstract properties of sup, inf, etc. diff -r e36a92ff543e -r 354c3844dfde src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Wed May 07 10:56:35 2008 +0200 +++ b/src/HOL/Lattices.thy Wed May 07 10:56:36 2008 +0200 @@ -6,7 +6,7 @@ header {* Abstract lattices *} theory Lattices -imports Orderings +imports Fun begin subsection{* Lattices *} @@ -506,53 +506,6 @@ by (iprover intro!: order_antisym le_boolI bot_least) -subsection {* Set as lattice *} - -instantiation set :: (type) distrib_lattice -begin - -definition - inf_set_eq [code func del]: "A \ B = A \ B" - -definition - sup_set_eq [code func del]: "A \ B = A \ B" - -instance - by intro_classes (auto simp add: inf_set_eq sup_set_eq) - -end - -lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" - apply (fold inf_set_eq sup_set_eq) - apply (erule mono_inf) - done - -lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" - apply (fold inf_set_eq sup_set_eq) - apply (erule mono_sup) - done - -instantiation set :: (type) complete_lattice -begin - -definition - Inf_set_def [code func del]: "\S \ \S" - -definition - Sup_set_def [code func del]: "\S \ \S" - -instance - by intro_classes (auto simp add: Inf_set_def Sup_set_def) - -end - -lemma top_set_eq: "top = UNIV" - by (iprover intro!: subset_antisym subset_UNIV top_greatest) - -lemma bot_set_eq: "bot = {}" - by (iprover intro!: subset_antisym empty_subsetI bot_least) - - subsection {* Fun as lattice *} instantiation "fun" :: (type, lattice) lattice @@ -610,6 +563,61 @@ by (iprover intro!: order_antisym le_funI bot_least) +subsection {* Set as lattice *} + +lemma inf_set_eq: "A \ B = A \ B" + apply (rule subset_antisym) + apply (rule Int_greatest) + apply (rule inf_le1) + apply (rule inf_le2) + apply (rule inf_greatest) + apply (rule Int_lower1) + apply (rule Int_lower2) + done + +lemma sup_set_eq: "A \ B = A \ B" + apply (rule subset_antisym) + apply (rule sup_least) + apply (rule Un_upper1) + apply (rule Un_upper2) + apply (rule Un_least) + apply (rule sup_ge1) + apply (rule sup_ge2) + done + +lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" + apply (fold inf_set_eq sup_set_eq) + apply (erule mono_inf) + done + +lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" + apply (fold inf_set_eq sup_set_eq) + apply (erule mono_sup) + done + +lemma Inf_set_eq: "\S = \S" + apply (rule subset_antisym) + apply (rule Inter_greatest) + apply (erule Inf_lower) + apply (rule Inf_greatest) + apply (erule Inter_lower) + done + +lemma Sup_set_eq: "\S = \S" + apply (rule subset_antisym) + apply (rule Sup_least) + apply (erule Union_upper) + apply (rule Union_least) + apply (erule Sup_upper) + done + +lemma top_set_eq: "top = UNIV" + by (iprover intro!: subset_antisym subset_UNIV top_greatest) + +lemma bot_set_eq: "bot = {}" + by (iprover intro!: subset_antisym empty_subsetI bot_least) + + text {* redundant bindings *} lemmas inf_aci = inf_ACI