# HG changeset patch # User haftmann # Date 1193426537 -7200 # Node ID 9c84ec7217a9d0b44305eccde0ad4e320103b5a6 # Parent b408ceba462736e16ed6e67bbb457e52a28fb2b2 localized monotonicity; tuned syntax diff -r b408ceba4627 -r 9c84ec7217a9 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Fri Oct 26 21:22:16 2007 +0200 +++ b/src/HOL/Lattices.thy Fri Oct 26 21:22:17 2007 +0200 @@ -11,6 +11,11 @@ subsection{* Lattices *} +notation + less_eq (infix "\" 50) +and + less (infix "\" 50) + class lower_semilattice = order + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) assumes inf_le1 [simp]: "x \ y \ x" @@ -61,11 +66,12 @@ lemma le_iff_inf: "(x \ y) = (x \ y = x)" by (blast intro: antisym dest: eq_iff [THEN iffD1]) -end +lemma mono_inf: + fixes f :: "'a \ 'b\lower_semilattice" + shows "mono f \ f (A \ B) \ f A \ f B" + by (auto simp add: mono_def intro: Lattices.inf_greatest) -lemma mono_inf: "mono f \ f (inf A B) \ inf (f A) (f B)" - by (auto simp add: mono_def) - +end context upper_semilattice begin @@ -93,15 +99,16 @@ lemma le_iff_sup: "(x \ y) = (x \ y = y)" by (blast intro: antisym dest: eq_iff [THEN iffD1]) -end +lemma mono_sup: + fixes f :: "'a \ 'b\upper_semilattice" + shows "mono f \ f A \ f B \ f (A \ B)" + by (auto simp add: mono_def intro: Lattices.sup_least) -lemma mono_sup: "mono f \ sup (f A) (f B) \ f (sup A B)" - by (auto simp add: mono_def) +end subsubsection{* Equational laws *} - context lower_semilattice begin @@ -393,12 +400,12 @@ definition top :: 'a where - "top = Inf {}" + "top = \{}" definition bot :: 'a where - "bot = Sup {}" + "bot = \{}" lemma top_greatest [simp]: "x \ top" by (unfold top_def, rule Inf_greatest, simp) @@ -409,12 +416,12 @@ definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where - "SUPR A f == Sup (f ` A)" + "SUPR A f == \ (f ` A)" definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - "INFI A f == Inf (f ` A)" + "INFI A f == \ (f ` A)" end @@ -473,17 +480,17 @@ subsection {* Bool as lattice *} instance bool :: distrib_lattice - inf_bool_eq: "inf P Q \ P \ Q" - sup_bool_eq: "sup P Q \ P \ Q" + inf_bool_eq: "P \ Q \ P \ Q" + sup_bool_eq: "P \ Q \ P \ Q" by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def) instance bool :: complete_lattice - Inf_bool_def: "Inf A \ \x\A. x" - Sup_bool_def: "Sup A \ \x\A. x" + Inf_bool_def: "\A \ \x\A. x" + Sup_bool_def: "\A \ \x\A. x" by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) lemma Inf_empty_bool [simp]: - "Inf {}" + "\{}" unfolding Inf_bool_def by auto lemma not_Sup_empty_bool [simp]: @@ -500,8 +507,8 @@ subsection {* Set as lattice *} instance set :: (type) distrib_lattice - inf_set_eq: "inf A B \ A \ B" - sup_set_eq: "sup A B \ A \ B" + inf_set_eq: "A \ B \ A \ B" + sup_set_eq: "A \ B \ A \ B" by intro_classes (auto simp add: inf_set_eq sup_set_eq) lemmas [code func del] = inf_set_eq sup_set_eq @@ -517,8 +524,8 @@ done instance set :: (type) complete_lattice - Inf_set_def: "Inf S \ \S" - Sup_set_def: "Sup S \ \S" + Inf_set_def: "\S \ \S" + Sup_set_def: "\S \ \S" by intro_classes (auto simp add: Inf_set_def Sup_set_def) lemmas [code func del] = Inf_set_def Sup_set_def @@ -533,8 +540,8 @@ subsection {* Fun as lattice *} instance "fun" :: (type, lattice) lattice - inf_fun_eq: "inf f g \ (\x. inf (f x) (g x))" - sup_fun_eq: "sup f g \ (\x. sup (f x) (g x))" + inf_fun_eq: "f \ g \ (\x. f x \ g x)" + sup_fun_eq: "f \ g \ (\x. f x \ g x)" apply intro_classes unfolding inf_fun_eq sup_fun_eq apply (auto intro: le_funI) @@ -550,8 +557,8 @@ by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) instance "fun" :: (type, complete_lattice) complete_lattice - Inf_fun_def: "Inf A \ (\x. Inf {y. \f\A. y = f x})" - Sup_fun_def: "Sup A \ (\x. Sup {y. \f\A. y = f x})" + Inf_fun_def: "\A \ (\x. \{y. \f\A. y = f x})" + Sup_fun_def: "\A \ (\x. \{y. \f\A. y = f x})" by intro_classes (auto simp add: Inf_fun_def Sup_fun_def le_fun_def intro: Inf_lower Sup_upper Inf_greatest Sup_least) @@ -559,11 +566,11 @@ lemmas [code func del] = Inf_fun_def Sup_fun_def lemma Inf_empty_fun: - "Inf {} = (\_. Inf {})" + "\{} = (\_. \{})" by rule (auto simp add: Inf_fun_def) lemma Sup_empty_fun: - "Sup {} = (\_. Sup {})" + "\{} = (\_. \{})" by rule (auto simp add: Sup_fun_def) lemma top_fun_eq: "top = (\x. top)" @@ -579,15 +586,16 @@ lemmas sup_aci = sup_ACI no_notation - inf (infixl "\" 70) - -no_notation - sup (infixl "\" 65) - -no_notation - Inf ("\_" [900] 900) - -no_notation - Sup ("\_" [900] 900) + less_eq (infix "\" 50) +and + less (infix "\" 50) +and + inf (infixl "\" 70) +and + sup (infixl "\" 65) +and + Inf ("\_" [900] 900) +and + Sup ("\_" [900] 900) end