# HG changeset patch # User haftmann # Date 1236237789 -3600 # Node ID 5ffa9d4dbea75931aa58e03e7d746d8fc47af2b9 # Parent 429612400fe9d1368e8e658b938093466bcb401b moved complete_lattice to Set.thy diff -r 429612400fe9 -r 5ffa9d4dbea7 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Thu Mar 05 08:23:08 2009 +0100 +++ b/src/HOL/Lattices.thy Thu Mar 05 08:23:09 2009 +0100 @@ -5,7 +5,7 @@ header {* Abstract lattices *} theory Lattices -imports Fun +imports Orderings begin subsection {* Lattices *} @@ -328,135 +328,6 @@ min_max.le_infI1 min_max.le_infI2 -subsection {* Complete lattices *} - -class complete_lattice = lattice + bot + top + - fixes Inf :: "'a set \ 'a" ("\_" [900] 900) - and Sup :: "'a set \ 'a" ("\_" [900] 900) - assumes Inf_lower: "x \ A \ \A \ x" - and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" - assumes Sup_upper: "x \ A \ x \ \A" - and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" -begin - -lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Inf_Univ: "\UNIV = \{}" - unfolding Sup_Inf by auto - -lemma Sup_Univ: "\UNIV = \{}" - unfolding Inf_Sup by auto - -lemma Inf_insert: "\insert a A = a \ \A" - by (auto intro: antisym Inf_greatest Inf_lower) - -lemma Sup_insert: "\insert a A = a \ \A" - by (auto intro: antisym Sup_least Sup_upper) - -lemma Inf_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Inf_lower Inf_greatest) - -lemma Sup_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Sup_upper Sup_least) - -lemma Inf_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Inf_insert) - -lemma Sup_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Sup_insert) - -lemma Inf_binary: - "\{a, b} = a \ b" - by (simp add: Inf_insert_simp) - -lemma Sup_binary: - "\{a, b} = a \ b" - by (simp add: Sup_insert_simp) - -lemma bot_def: - "bot = \{}" - by (auto intro: antisym Sup_least) - -lemma top_def: - "top = \{}" - by (auto intro: antisym Inf_greatest) - -lemma sup_bot [simp]: - "x \ bot = x" - using bot_least [of x] by (simp add: le_iff_sup sup_commute) - -lemma inf_top [simp]: - "x \ top = x" - using top_greatest [of x] by (simp add: le_iff_inf inf_commute) - -definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where - "SUPR A f == \ (f ` A)" - -definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - "INFI A f == \ (f ` A)" - -end - -syntax - "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) - "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) - "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) - "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) - -translations - "SUP x y. B" == "SUP x. SUP y. B" - "SUP x. B" == "CONST SUPR UNIV (%x. B)" - "SUP x. B" == "SUP x:UNIV. B" - "SUP x:A. B" == "CONST SUPR A (%x. B)" - "INF x y. B" == "INF x. INF y. B" - "INF x. B" == "CONST INFI UNIV (%x. B)" - "INF x. B" == "INF x:UNIV. B" - "INF x:A. B" == "CONST INFI A (%x. B)" - -(* To avoid eta-contraction of body: *) -print_translation {* -let - fun btr' syn (A :: Abs abs :: ts) = - let val (x,t) = atomic_abs_tr' abs - in list_comb (Syntax.const syn $ x $ A $ t, ts) end - val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const -in -[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")] -end -*} - -context complete_lattice -begin - -lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" - by (auto simp add: SUPR_def intro: Sup_upper) - -lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" - by (auto simp add: SUPR_def intro: Sup_least) - -lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" - by (auto simp add: INFI_def intro: Inf_lower) - -lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" - by (auto simp add: INFI_def intro: Inf_greatest) - -lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" - by (auto intro: antisym SUP_leI le_SUPI) - -lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" - by (auto intro: antisym INF_leI le_INFI) - -end - - subsection {* Bool as lattice *} instantiation bool :: distrib_lattice @@ -473,28 +344,6 @@ end -instantiation bool :: complete_lattice -begin - -definition - Inf_bool_def: "\A \ (\x\A. x)" - -definition - Sup_bool_def: "\A \ (\x\A. x)" - -instance - by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) - -end - -lemma Inf_empty_bool [simp]: - "\{}" - unfolding Inf_bool_def by auto - -lemma not_Sup_empty_bool [simp]: - "\ Sup {}" - unfolding Sup_bool_def by auto - subsection {* Fun as lattice *} @@ -522,85 +371,6 @@ instance "fun" :: (type, distrib_lattice) distrib_lattice by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) -instantiation "fun" :: (type, complete_lattice) complete_lattice -begin - -definition - Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -definition - Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -instance - by intro_classes - (auto simp add: Inf_fun_def Sup_fun_def le_fun_def - intro: Inf_lower Sup_upper Inf_greatest Sup_least) - -end - -lemma Inf_empty_fun: - "\{} = (\_. \{})" - by rule (auto simp add: Inf_fun_def) - -lemma Sup_empty_fun: - "\{} = (\_. \{})" - by rule (auto simp add: Sup_fun_def) - - -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 *} @@ -611,8 +381,6 @@ less_eq (infix "\" 50) and less (infix "\" 50) and inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) + sup (infixl "\" 65) end