# HG changeset patch # User ballarin # Date 1168612641 -3600 # Node ID 71742560919279baeaa6eade105eb574ef2c9762 # Parent f4cfc4101c8f9040d7f4816e53b1ca60b29a2f42 Reverted to structure representation with records. diff -r f4cfc4101c8f -r 717425609192 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Fri Jan 12 09:58:31 2007 +0100 +++ b/src/HOL/Algebra/Group.thy Fri Jan 12 15:37:21 2007 +0100 @@ -685,7 +685,7 @@ text_raw {* \label{sec:subgroup-lattice} *} theorem (in group) subgroups_partial_order: - "partial_order {H. subgroup H G} (op \)" + "partial_order (| carrier = {H. subgroup H G}, le = op \ |)" by (rule partial_order.intro) simp_all lemma (in group) subgroup_self: @@ -730,23 +730,23 @@ qed theorem (in group) subgroups_complete_lattice: - "complete_lattice {H. subgroup H G} (op \)" - (is "complete_lattice ?car ?le") + "complete_lattice (| carrier = {H. subgroup H G}, le = op \ |)" + (is "complete_lattice ?L") proof (rule partial_order.complete_lattice_criterion1) - show "partial_order ?car ?le" by (rule subgroups_partial_order) + show "partial_order ?L" by (rule subgroups_partial_order) next - have "order_syntax.greatest ?car ?le (carrier G) ?car" - by (unfold order_syntax.greatest_def) + have "greatest ?L (carrier G) (carrier ?L)" + by (unfold greatest_def) (simp add: subgroup.subset subgroup_self) - then show "\G. order_syntax.greatest ?car ?le G ?car" .. + then show "\G. greatest ?L G (carrier ?L)" .. next fix A - assume L: "A \ ?car" and non_empty: "A ~= {}" + assume L: "A \ carrier ?L" and non_empty: "A ~= {}" then have Int_subgroup: "subgroup (\A) G" by (fastsimp intro: subgroups_Inter) - have "order_syntax.greatest ?car ?le (\A) (order_syntax.Lower ?car ?le A)" - (is "order_syntax.greatest _ _ ?Int _") - proof (rule order_syntax.greatest_LowerI) + have "greatest ?L (\A) (Lower ?L A)" + (is "greatest _ ?Int _") + proof (rule greatest_LowerI) fix H assume H: "H \ A" with L have subgroupH: "subgroup H G" by auto @@ -755,18 +755,18 @@ from groupH have monoidH: "monoid ?H" by (rule group.is_monoid) from H have Int_subset: "?Int \ H" by fastsimp - then show "?le ?Int H" by simp + then show "le ?L ?Int H" by simp next fix H - assume H: "H \ order_syntax.Lower ?car ?le A" - with L Int_subgroup show "?le H ?Int" - by (fastsimp simp: order_syntax.Lower_def intro: Inter_greatest) + assume H: "H \ Lower ?L A" + with L Int_subgroup show "le ?L H ?Int" + by (fastsimp simp: Lower_def intro: Inter_greatest) next - show "A \ ?car" by (rule L) + show "A \ carrier ?L" by (rule L) next - show "?Int \ ?car" by simp (rule Int_subgroup) + show "?Int \ carrier ?L" by simp (rule Int_subgroup) qed - then show "\I. order_syntax.greatest ?car ?le I (order_syntax.Lower ?car ?le A)" .. + then show "\I. greatest ?L I (Lower ?L A)" .. qed end diff -r f4cfc4101c8f -r 717425609192 src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Fri Jan 12 09:58:31 2007 +0100 +++ b/src/HOL/Algebra/IntRing.thy Fri Jan 12 15:37:21 2007 +0100 @@ -96,54 +96,60 @@ interpretation "domain" ["\"] by (rule int_is_domain) lemma int_le_total_order: - "total_order (UNIV::int set) (op \)" -apply (rule partial_order.total_orderI) - apply (rule partial_order.intro, simp+) -apply clarsimp -done + "total_order (| carrier = UNIV::int set, le = op \ |)" + apply (rule partial_order.total_orderI) + apply (rule partial_order.intro, simp+) + apply clarsimp + done lemma less_int: - "order_syntax.less (op \::[int, int] => bool) = (op <)" - by (auto simp add: expand_fun_eq order_syntax.less_def) + "lless (| carrier = UNIV::int set, le = op \ |) = (op <)" + by (auto simp add: expand_fun_eq lless_def) lemma join_int: - "order_syntax.join (UNIV::int set) (op \) = max" + "join (| carrier = UNIV::int set, le = op \ |) = max" apply (simp add: expand_fun_eq max_def) apply (rule+) apply (rule lattice.joinI) apply (simp add: int_le_total_order total_order.axioms) - apply (simp add: order_syntax.least_def order_syntax.Upper_def) + apply (simp add: least_def Upper_def) apply arith apply simp apply simp apply (rule lattice.joinI) apply (simp add: int_le_total_order total_order.axioms) - apply (simp add: order_syntax.least_def order_syntax.Upper_def) + apply (simp add: least_def Upper_def) apply arith apply simp apply simp done lemma meet_int: - "order_syntax.meet (UNIV::int set) (op \) = min" + "meet (| carrier = UNIV::int set, le = op \ |) = min" apply (simp add: expand_fun_eq min_def) apply (rule+) apply (rule lattice.meetI) apply (simp add: int_le_total_order total_order.axioms) - apply (simp add: order_syntax.greatest_def order_syntax.Lower_def) + apply (simp add: greatest_def Lower_def) apply arith apply simp apply simp apply (rule lattice.meetI) apply (simp add: int_le_total_order total_order.axioms) - apply (simp add: order_syntax.greatest_def order_syntax.Lower_def) + apply (simp add: greatest_def Lower_def) apply arith apply simp apply simp done -text {* Interpretation unfolding @{text less}, @{text join} and @{text +lemma carrier_int: + "carrier (| carrier = UNIV::int set, le = op \ |) = UNIV" + apply simp + done + +text {* Interpretation unfolding @{text lless}, @{text join} and @{text meet} since they have natural representations in @{typ int}. *} interpretation - int [unfolded less_int join_int meet_int]: - total_order ["UNIV::int set" "op \"] by (rule int_le_total_order) + int [unfolded less_int join_int meet_int carrier_int]: + total_order ["(| carrier = UNIV::int set, le = op \ |)"] + by (rule int_le_total_order) subsubsection {* Generated Ideals of @{text "\"} *} diff -r f4cfc4101c8f -r 717425609192 src/HOL/Algebra/Lattice.thy --- a/src/HOL/Algebra/Lattice.thy Fri Jan 12 09:58:31 2007 +0100 +++ b/src/HOL/Algebra/Lattice.thy Fri Jan 12 15:37:21 2007 +0100 @@ -18,163 +18,157 @@ subsection {* Partial Orders *} -text {* Locale @{text order_syntax} is required since we want to refer - to definitions (and their derived theorems) outside of @{text partial_order}. - *} - -locale order_syntax = - fixes L :: "'a set" and le :: "['a, 'a] => bool" (infix "\" 50) - -text {* Note that the type constraints above are necessary, because the - definition command cannot specialise the types. *} - -definition (in order_syntax) - less (infixl "\" 50) where "x \ y == x \ y & x ~= y" - -text {* Upper and lower bounds of a set. *} - -definition (in order_syntax) - Upper :: "'a set => 'a set" where - "Upper A == {u. (ALL x. x \ A \ L --> x \ u)} \ L" - -definition (in order_syntax) - Lower :: "'a set => 'a set" where - "Lower A == {l. (ALL x. x \ A \ L --> l \ x)} \ L" - -text {* Least and greatest, as predicate. *} +record 'a order = "'a partial_object" + + le :: "['a, 'a] => bool" (infixl "\\" 50) -definition (in order_syntax) - least :: "['a, 'a set] => bool" where - "least l A == A \ L & l \ A & (ALL x : A. l \ x)" - -definition (in order_syntax) - greatest :: "['a, 'a set] => bool" where - "greatest g A == A \ L & g \ A & (ALL x : A. x \ g)" - -text {* Supremum and infimum *} - -definition (in order_syntax) - sup :: "'a set => 'a" ("\_" [90] 90) where - "\A == THE x. least x (Upper A)" - -definition (in order_syntax) - inf :: "'a set => 'a" ("\_" [90] 90) where - "\A == THE x. greatest x (Lower A)" - -definition (in order_syntax) - join :: "['a, 'a] => 'a" (infixl "\" 65) where - "x \ y == sup {x, y}" - -definition (in order_syntax) - meet :: "['a, 'a] => 'a" (infixl "\" 70) where - "x \ y == inf {x, y}" - -locale partial_order = order_syntax + +locale partial_order = + fixes L (structure) assumes refl [intro, simp]: - "x \ L ==> x \ x" + "x \ carrier L ==> x \ x" and anti_sym [intro]: - "[| x \ y; y \ x; x \ L; y \ L |] ==> x = y" + "[| x \ y; y \ x; x \ carrier L; y \ carrier L |] ==> x = y" and trans [trans]: "[| x \ y; y \ z; - x \ L; y \ L; z \ L |] ==> x \ z" + x \ carrier L; y \ carrier L; z \ carrier L |] ==> x \ z" + +constdefs (structure L) + lless :: "[_, 'a, 'a] => bool" (infixl "\\" 50) + "x \ y == x \ y & x ~= y" + + -- {* Upper and lower bounds of a set. *} + Upper :: "[_, 'a set] => 'a set" + "Upper L A == {u. (ALL x. x \ A \ carrier L --> x \ u)} \ + carrier L" + + Lower :: "[_, 'a set] => 'a set" + "Lower L A == {l. (ALL x. x \ A \ carrier L --> l \ x)} \ + carrier L" + + -- {* Least and greatest, as predicate. *} + least :: "[_, 'a, 'a set] => bool" + "least L l A == A \ carrier L & l \ A & (ALL x : A. l \ x)" + + greatest :: "[_, 'a, 'a set] => bool" + "greatest L g A == A \ carrier L & g \ A & (ALL x : A. x \ g)" + + -- {* Supremum and infimum *} + sup :: "[_, 'a set] => 'a" ("\\_" [90] 90) + "\A == THE x. least L x (Upper L A)" + + inf :: "[_, 'a set] => 'a" ("\\_" [90] 90) + "\A == THE x. greatest L x (Lower L A)" + + join :: "[_, 'a, 'a] => 'a" (infixl "\\" 65) + "x \ y == sup L {x, y}" + + meet :: "[_, 'a, 'a] => 'a" (infixl "\\" 70) + "x \ y == inf L {x, y}" subsubsection {* Upper *} -lemma (in order_syntax) Upper_closed [intro, simp]: - "Upper A \ L" +lemma Upper_closed [intro, simp]: + "Upper L A \ carrier L" by (unfold Upper_def) clarify -lemma (in order_syntax) UpperD [dest]: - "[| u \ Upper A; x \ A; A \ L |] ==> x \ u" +lemma UpperD [dest]: + fixes L (structure) + shows "[| u \ Upper L A; x \ A; A \ carrier L |] ==> x \ u" by (unfold Upper_def) blast -lemma (in order_syntax) Upper_memI: - "[| !! y. y \ A ==> y \ x; x \ L |] ==> x \ Upper A" +lemma Upper_memI: + fixes L (structure) + shows "[| !! y. y \ A ==> y \ x; x \ carrier L |] ==> x \ Upper L A" by (unfold Upper_def) blast -lemma (in order_syntax) Upper_antimono: - "A \ B ==> Upper B \ Upper A" +lemma Upper_antimono: + "A \ B ==> Upper L B \ Upper L A" by (unfold Upper_def) blast subsubsection {* Lower *} -lemma (in order_syntax) Lower_closed [intro, simp]: - "Lower A \ L" +lemma Lower_closed [intro, simp]: + "Lower L A \ carrier L" by (unfold Lower_def) clarify -lemma (in order_syntax) LowerD [dest]: - "[| l \ Lower A; x \ A; A \ L |] ==> l \ x" +lemma LowerD [dest]: + fixes L (structure) + shows "[| l \ Lower L A; x \ A; A \ carrier L |] ==> l \ x" by (unfold Lower_def) blast -lemma (in order_syntax) Lower_memI: - "[| !! y. y \ A ==> x \ y; x \ L |] ==> x \ Lower A" +lemma Lower_memI: + fixes L (structure) + shows "[| !! y. y \ A ==> x \ y; x \ carrier L |] ==> x \ Lower L A" by (unfold Lower_def) blast -lemma (in order_syntax) Lower_antimono: - "A \ B ==> Lower B \ Lower A" +lemma Lower_antimono: + "A \ B ==> Lower L B \ Lower L A" by (unfold Lower_def) blast subsubsection {* least *} -lemma (in order_syntax) least_closed [intro, simp]: - "least l A ==> l \ L" +lemma least_carrier [intro, simp]: + shows "least L l A ==> l \ carrier L" by (unfold least_def) fast -lemma (in order_syntax) least_mem: - "least l A ==> l \ A" +lemma least_mem: + "least L l A ==> l \ A" by (unfold least_def) fast lemma (in partial_order) least_unique: - "[| least x A; least y A |] ==> x = y" + "[| least L x A; least L y A |] ==> x = y" by (unfold least_def) blast -lemma (in order_syntax) least_le: - "[| least x A; a \ A |] ==> x \ a" +lemma least_le: + fixes L (structure) + shows "[| least L x A; a \ A |] ==> x \ a" by (unfold least_def) fast -lemma (in order_syntax) least_UpperI: +lemma least_UpperI: + fixes L (structure) assumes above: "!! x. x \ A ==> x \ s" - and below: "!! y. y \ Upper A ==> s \ y" - and L: "A \ L" "s \ L" - shows "least s (Upper A)" + and below: "!! y. y \ Upper L A ==> s \ y" + and L: "A \ carrier L" "s \ carrier L" + shows "least L s (Upper L A)" proof - - have "Upper A \ L" by simp - moreover from above L have "s \ Upper A" by (simp add: Upper_def) - moreover from below have "ALL x : Upper A. s \ x" by fast + have "Upper L A \ carrier L" by simp + moreover from above L have "s \ Upper L A" by (simp add: Upper_def) + moreover from below have "ALL x : Upper L A. s \ x" by fast ultimately show ?thesis by (simp add: least_def) qed subsubsection {* greatest *} -lemma (in order_syntax) greatest_closed [intro, simp]: - "greatest l A ==> l \ L" +lemma greatest_carrier [intro, simp]: + shows "greatest L l A ==> l \ carrier L" by (unfold greatest_def) fast -lemma (in order_syntax) greatest_mem: - "greatest l A ==> l \ A" +lemma greatest_mem: + "greatest L l A ==> l \ A" by (unfold greatest_def) fast lemma (in partial_order) greatest_unique: - "[| greatest x A; greatest y A |] ==> x = y" + "[| greatest L x A; greatest L y A |] ==> x = y" by (unfold greatest_def) blast -lemma (in order_syntax) greatest_le: - "[| greatest x A; a \ A |] ==> a \ x" +lemma greatest_le: + fixes L (structure) + shows "[| greatest L x A; a \ A |] ==> a \ x" by (unfold greatest_def) fast -lemma (in order_syntax) greatest_LowerI: +lemma greatest_LowerI: + fixes L (structure) assumes below: "!! x. x \ A ==> i \ x" - and above: "!! y. y \ Lower A ==> y \ i" - and L: "A \ L" "i \ L" - shows "greatest i (Lower A)" + and above: "!! y. y \ Lower L A ==> y \ i" + and L: "A \ carrier L" "i \ carrier L" + shows "greatest L i (Lower L A)" proof - - have "Lower A \ L" by simp - moreover from below L have "i \ Lower A" by (simp add: Lower_def) - moreover from above have "ALL x : Lower A. x \ i" by fast + have "Lower L A \ carrier L" by simp + moreover from below L have "i \ Lower L A" by (simp add: Lower_def) + moreover from above have "ALL x : Lower L A. x \ i" by fast ultimately show ?thesis by (simp add: greatest_def) qed @@ -183,61 +177,63 @@ locale lattice = partial_order + assumes sup_of_two_exists: - "[| x \ L; y \ L |] ==> EX s. order_syntax.least L le s (order_syntax.Upper L le {x, y})" + "[| x \ carrier L; y \ carrier L |] ==> EX s. least L s (Upper L {x, y})" and inf_of_two_exists: - "[| x \ L; y \ L |] ==> EX s. order_syntax.greatest L le s (order_syntax.Lower L le {x, y})" + "[| x \ carrier L; y \ carrier L |] ==> EX s. greatest L s (Lower L {x, y})" -lemma (in order_syntax) least_Upper_above: - "[| least s (Upper A); x \ A; A \ L |] ==> x \ s" +lemma least_Upper_above: + fixes L (structure) + shows "[| least L s (Upper L A); x \ A; A \ carrier L |] ==> x \ s" by (unfold least_def) blast -lemma (in order_syntax) greatest_Lower_above: - "[| greatest i (Lower A); x \ A; A \ L |] ==> i \ x" +lemma greatest_Lower_above: + fixes L (structure) + shows "[| greatest L i (Lower L A); x \ A; A \ carrier L |] ==> i \ x" by (unfold greatest_def) blast subsubsection {* Supremum *} lemma (in lattice) joinI: - "[| !!l. least l (Upper {x, y}) ==> P l; x \ L; y \ L |] + "[| !!l. least L l (Upper L {x, y}) ==> P l; x \ carrier L; y \ carrier L |] ==> P (x \ y)" proof (unfold join_def sup_def) - assume L: "x \ L" "y \ L" - and P: "!!l. least l (Upper {x, y}) ==> P l" - with sup_of_two_exists obtain s where "least s (Upper {x, y})" by fast - with L show "P (THE l. least l (Upper {x, y}))" + assume L: "x \ carrier L" "y \ carrier L" + and P: "!!l. least L l (Upper L {x, y}) ==> P l" + with sup_of_two_exists obtain s where "least L s (Upper L {x, y})" by fast + with L show "P (THE l. least L l (Upper L {x, y}))" by (fast intro: theI2 least_unique P) qed lemma (in lattice) join_closed [simp]: - "[| x \ L; y \ L |] ==> x \ y \ L" - by (rule joinI) (rule least_closed) + "[| x \ carrier L; y \ carrier L |] ==> x \ y \ carrier L" + by (rule joinI) (rule least_carrier) -lemma (in partial_order) sup_of_singletonI: (* only reflexivity needed ? *) - "x \ L ==> least x (Upper {x})" +lemma (in partial_order) sup_of_singletonI: (* only reflexivity needed ? *) + "x \ carrier L ==> least L x (Upper L {x})" by (rule least_UpperI) fast+ lemma (in partial_order) sup_of_singleton [simp]: - "x \ L ==> \{x} = x" + "x \ carrier L ==> \{x} = x" by (unfold sup_def) (blast intro: least_unique least_UpperI sup_of_singletonI) text {* Condition on @{text A}: supremum exists. *} lemma (in lattice) sup_insertI: - "[| !!s. least s (Upper (insert x A)) ==> P s; - least a (Upper A); x \ L; A \ L |] + "[| !!s. least L s (Upper L (insert x A)) ==> P s; + least L a (Upper L A); x \ carrier L; A \ carrier L |] ==> P (\(insert x A))" proof (unfold sup_def) - assume L: "x \ L" "A \ L" - and P: "!!l. least l (Upper (insert x A)) ==> P l" - and least_a: "least a (Upper A)" - from least_a have La: "a \ L" by simp + assume L: "x \ carrier L" "A \ carrier L" + and P: "!!l. least L l (Upper L (insert x A)) ==> P l" + and least_a: "least L a (Upper L A)" + from L least_a have La: "a \ carrier L" by simp from L sup_of_two_exists least_a - obtain s where least_s: "least s (Upper {a, x})" by blast - show "P (THE l. least l (Upper (insert x A)))" + obtain s where least_s: "least L s (Upper L {a, x})" by blast + show "P (THE l. least L l (Upper L (insert x A)))" proof (rule theI2) - show "least s (Upper (insert x A))" + show "least L s (Upper L (insert x A))" proof (rule least_UpperI) fix z assume "z \ insert x A" @@ -252,15 +248,15 @@ qed next fix y - assume y: "y \ Upper (insert x A)" + assume y: "y \ Upper L (insert x A)" show "s \ y" proof (rule least_le [OF least_s], rule Upper_memI) fix z assume z: "z \ {a, x}" then show "z \ y" proof - have y': "y \ Upper A" - apply (rule subsetD [where A = "Upper (insert x A)"]) + have y': "y \ Upper L A" + apply (rule subsetD [where A = "Upper L (insert x A)"]) apply (rule Upper_antimono) apply clarify apply assumption done assume "z = a" @@ -271,15 +267,15 @@ qed qed (rule Upper_closed [THEN subsetD]) next - from L show "insert x A \ L" by simp - from least_s show "s \ L" by simp + from L show "insert x A \ carrier L" by simp + from least_s show "s \ carrier L" by simp qed next fix l - assume least_l: "least l (Upper (insert x A))" + assume least_l: "least L l (Upper L (insert x A))" show "l = s" proof (rule least_unique) - show "least s (Upper (insert x A))" + show "least L s (Upper L (insert x A))" proof (rule least_UpperI) fix z assume "z \ insert x A" @@ -294,15 +290,15 @@ qed next fix y - assume y: "y \ Upper (insert x A)" + assume y: "y \ Upper L (insert x A)" show "s \ y" proof (rule least_le [OF least_s], rule Upper_memI) fix z assume z: "z \ {a, x}" then show "z \ y" proof - have y': "y \ Upper A" - apply (rule subsetD [where A = "Upper (insert x A)"]) + have y': "y \ Upper L A" + apply (rule subsetD [where A = "Upper L (insert x A)"]) apply (rule Upper_antimono) apply clarify apply assumption done assume "z = a" @@ -313,15 +309,15 @@ qed qed (rule Upper_closed [THEN subsetD]) next - from L show "insert x A \ L" by simp - from least_s show "s \ L" by simp + from L show "insert x A \ carrier L" by simp + from least_s show "s \ carrier L" by simp qed qed qed qed lemma (in lattice) finite_sup_least: - "[| finite A; A \ L; A ~= {} |] ==> least (\A) (Upper A)" + "[| finite A; A \ carrier L; A ~= {} |] ==> least L (\A) (Upper L A)" proof (induct set: Finites) case empty then show ?case by simp @@ -333,15 +329,15 @@ with insert show ?thesis by (simp add: sup_of_singletonI) next case False - with insert have "least (\A) (Upper A)" by simp + with insert have "least L (\A) (Upper L A)" by simp with _ show ?thesis by (rule sup_insertI) (simp_all add: insert [simplified]) qed qed lemma (in lattice) finite_sup_insertI: - assumes P: "!!l. least l (Upper (insert x A)) ==> P l" - and xA: "finite A" "x \ L" "A \ L" + assumes P: "!!l. least L l (Upper L (insert x A)) ==> P l" + and xA: "finite A" "x \ carrier L" "A \ carrier L" shows "P (\ (insert x A))" proof (cases "A = {}") case True with P and xA show ?thesis @@ -352,7 +348,7 @@ qed lemma (in lattice) finite_sup_closed: - "[| finite A; A \ L; A ~= {} |] ==> \A \ L" + "[| finite A; A \ carrier L; A ~= {} |] ==> \A \ carrier L" proof (induct set: Finites) case empty then show ?case by simp next @@ -361,39 +357,39 @@ qed lemma (in lattice) join_left: - "[| x \ L; y \ L |] ==> x \ x \ y" + "[| x \ carrier L; y \ carrier L |] ==> x \ x \ y" by (rule joinI [folded join_def]) (blast dest: least_mem) lemma (in lattice) join_right: - "[| x \ L; y \ L |] ==> y \ x \ y" + "[| x \ carrier L; y \ carrier L |] ==> y \ x \ y" by (rule joinI [folded join_def]) (blast dest: least_mem) lemma (in lattice) sup_of_two_least: - "[| x \ L; y \ L |] ==> least (\{x, y}) (Upper {x, y})" + "[| x \ carrier L; y \ carrier L |] ==> least L (\{x, y}) (Upper L {x, y})" proof (unfold sup_def) - assume L: "x \ L" "y \ L" - with sup_of_two_exists obtain s where "least s (Upper {x, y})" by fast - with L show "least (THE xa. least xa (Upper {x, y})) (Upper {x, y})" + assume L: "x \ carrier L" "y \ carrier L" + with sup_of_two_exists obtain s where "least L s (Upper L {x, y})" by fast + with L show "least L (THE xa. least L xa (Upper L {x, y})) (Upper L {x, y})" by (fast intro: theI2 least_unique) (* blast fails *) qed lemma (in lattice) join_le: assumes sub: "x \ z" "y \ z" - and L: "x \ L" "y \ L" "z \ L" + and L: "x \ carrier L" "y \ carrier L" "z \ carrier L" shows "x \ y \ z" proof (rule joinI) fix s - assume "least s (Upper {x, y})" + assume "least L s (Upper L {x, y})" with sub L show "s \ z" by (fast elim: least_le intro: Upper_memI) qed lemma (in lattice) join_assoc_lemma: - assumes L: "x \ L" "y \ L" "z \ L" + assumes L: "x \ carrier L" "y \ carrier L" "z \ carrier L" shows "x \ (y \ z) = \{x, y, z}" proof (rule finite_sup_insertI) -- {* The textbook argument in Jacobson I, p 457 *} fix s - assume sup: "least s (Upper {x, y, z})" + assume sup: "least L s (Upper L {x, y, z})" show "x \ (y \ z) = s" proof (rule anti_sym) from sup L show "x \ (y \ z) \ s" @@ -402,15 +398,16 @@ from sup L show "s \ x \ (y \ z)" by (erule_tac least_le) (blast intro!: Upper_memI intro: trans join_left join_right join_closed) - qed (simp_all add: L least_closed [OF sup]) + qed (simp_all add: L least_carrier [OF sup]) qed (simp_all add: L) -lemma (in order_syntax) join_comm: - "x \ y = y \ x" +lemma join_comm: + fixes L (structure) + shows "x \ y = y \ x" by (unfold join_def) (simp add: insert_commute) lemma (in lattice) join_assoc: - assumes L: "x \ L" "y \ L" "z \ L" + assumes L: "x \ carrier L" "y \ carrier L" "z \ carrier L" shows "(x \ y) \ z = x \ (y \ z)" proof - have "(x \ y) \ z = z \ (x \ y)" by (simp only: join_comm) @@ -424,44 +421,45 @@ subsubsection {* Infimum *} lemma (in lattice) meetI: - "[| !!i. greatest i (Lower {x, y}) ==> P i; x \ L; y \ L |] + "[| !!i. greatest L i (Lower L {x, y}) ==> P i; + x \ carrier L; y \ carrier L |] ==> P (x \ y)" proof (unfold meet_def inf_def) - assume L: "x \ L" "y \ L" - and P: "!!g. greatest g (Lower {x, y}) ==> P g" - with inf_of_two_exists obtain i where "greatest i (Lower {x, y})" by fast - with L show "P (THE g. greatest g (Lower {x, y}))" + assume L: "x \ carrier L" "y \ carrier L" + and P: "!!g. greatest L g (Lower L {x, y}) ==> P g" + with inf_of_two_exists obtain i where "greatest L i (Lower L {x, y})" by fast + with L show "P (THE g. greatest L g (Lower L {x, y}))" by (fast intro: theI2 greatest_unique P) qed lemma (in lattice) meet_closed [simp]: - "[| x \ L; y \ L |] ==> x \ y \ L" - by (rule meetI) (rule greatest_closed) + "[| x \ carrier L; y \ carrier L |] ==> x \ y \ carrier L" + by (rule meetI) (rule greatest_carrier) lemma (in partial_order) inf_of_singletonI: (* only reflexivity needed ? *) - "x \ L ==> greatest x (Lower {x})" + "x \ carrier L ==> greatest L x (Lower L {x})" by (rule greatest_LowerI) fast+ lemma (in partial_order) inf_of_singleton [simp]: - "x \ L ==> \ {x} = x" + "x \ carrier L ==> \ {x} = x" by (unfold inf_def) (blast intro: greatest_unique greatest_LowerI inf_of_singletonI) text {* Condition on A: infimum exists. *} lemma (in lattice) inf_insertI: - "[| !!i. greatest i (Lower (insert x A)) ==> P i; - greatest a (Lower A); x \ L; A \ L |] + "[| !!i. greatest L i (Lower L (insert x A)) ==> P i; + greatest L a (Lower L A); x \ carrier L; A \ carrier L |] ==> P (\(insert x A))" proof (unfold inf_def) - assume L: "x \ L" "A \ L" - and P: "!!g. greatest g (Lower (insert x A)) ==> P g" - and greatest_a: "greatest a (Lower A)" - from greatest_a have La: "a \ L" by simp + assume L: "x \ carrier L" "A \ carrier L" + and P: "!!g. greatest L g (Lower L (insert x A)) ==> P g" + and greatest_a: "greatest L a (Lower L A)" + from L greatest_a have La: "a \ carrier L" by simp from L inf_of_two_exists greatest_a - obtain i where greatest_i: "greatest i (Lower {a, x})" by blast - show "P (THE g. greatest g (Lower (insert x A)))" + obtain i where greatest_i: "greatest L i (Lower L {a, x})" by blast + show "P (THE g. greatest L g (Lower L (insert x A)))" proof (rule theI2) - show "greatest i (Lower (insert x A))" + show "greatest L i (Lower L (insert x A))" proof (rule greatest_LowerI) fix z assume "z \ insert x A" @@ -476,15 +474,15 @@ qed next fix y - assume y: "y \ Lower (insert x A)" + assume y: "y \ Lower L (insert x A)" show "y \ i" proof (rule greatest_le [OF greatest_i], rule Lower_memI) fix z assume z: "z \ {a, x}" then show "y \ z" proof - have y': "y \ Lower A" - apply (rule subsetD [where A = "Lower (insert x A)"]) + have y': "y \ Lower L A" + apply (rule subsetD [where A = "Lower L (insert x A)"]) apply (rule Lower_antimono) apply clarify apply assumption done assume "z = a" @@ -495,15 +493,15 @@ qed qed (rule Lower_closed [THEN subsetD]) next - from L show "insert x A \ L" by simp - from greatest_i show "i \ L" by simp + from L show "insert x A \ carrier L" by simp + from greatest_i show "i \ carrier L" by simp qed next fix g - assume greatest_g: "greatest g (Lower (insert x A))" + assume greatest_g: "greatest L g (Lower L (insert x A))" show "g = i" proof (rule greatest_unique) - show "greatest i (Lower (insert x A))" + show "greatest L i (Lower L (insert x A))" proof (rule greatest_LowerI) fix z assume "z \ insert x A" @@ -518,15 +516,15 @@ qed next fix y - assume y: "y \ Lower (insert x A)" + assume y: "y \ Lower L (insert x A)" show "y \ i" proof (rule greatest_le [OF greatest_i], rule Lower_memI) fix z assume z: "z \ {a, x}" then show "y \ z" proof - have y': "y \ Lower A" - apply (rule subsetD [where A = "Lower (insert x A)"]) + have y': "y \ Lower L A" + apply (rule subsetD [where A = "Lower L (insert x A)"]) apply (rule Lower_antimono) apply clarify apply assumption done assume "z = a" @@ -537,15 +535,15 @@ qed qed (rule Lower_closed [THEN subsetD]) next - from L show "insert x A \ L" by simp - from greatest_i show "i \ L" by simp + from L show "insert x A \ carrier L" by simp + from greatest_i show "i \ carrier L" by simp qed qed qed qed lemma (in lattice) finite_inf_greatest: - "[| finite A; A \ L; A ~= {} |] ==> greatest (\A) (Lower A)" + "[| finite A; A \ carrier L; A ~= {} |] ==> greatest L (\A) (Lower L A)" proof (induct set: Finites) case empty then show ?case by simp next @@ -558,14 +556,14 @@ case False from insert show ?thesis proof (rule_tac inf_insertI) - from False insert show "greatest (\A) (Lower A)" by simp + from False insert show "greatest L (\A) (Lower L A)" by simp qed simp_all qed qed lemma (in lattice) finite_inf_insertI: - assumes P: "!!i. greatest i (Lower (insert x A)) ==> P i" - and xA: "finite A" "x \ L" "A \ L" + assumes P: "!!i. greatest L i (Lower L (insert x A)) ==> P i" + and xA: "finite A" "x \ carrier L" "A \ carrier L" shows "P (\ (insert x A))" proof (cases "A = {}") case True with P and xA show ?thesis @@ -576,7 +574,7 @@ qed lemma (in lattice) finite_inf_closed: - "[| finite A; A \ L; A ~= {} |] ==> \A \ L" + "[| finite A; A \ carrier L; A ~= {} |] ==> \A \ carrier L" proof (induct set: Finites) case empty then show ?case by simp next @@ -585,40 +583,41 @@ qed lemma (in lattice) meet_left: - "[| x \ L; y \ L |] ==> x \ y \ x" + "[| x \ carrier L; y \ carrier L |] ==> x \ y \ x" by (rule meetI [folded meet_def]) (blast dest: greatest_mem) lemma (in lattice) meet_right: - "[| x \ L; y \ L |] ==> x \ y \ y" + "[| x \ carrier L; y \ carrier L |] ==> x \ y \ y" by (rule meetI [folded meet_def]) (blast dest: greatest_mem) lemma (in lattice) inf_of_two_greatest: - "[| x \ L; y \ L |] ==> greatest (\ {x, y}) (Lower {x, y})" + "[| x \ carrier L; y \ carrier L |] ==> + greatest L (\ {x, y}) (Lower L {x, y})" proof (unfold inf_def) - assume L: "x \ L" "y \ L" - with inf_of_two_exists obtain s where "greatest s (Lower {x, y})" by fast + assume L: "x \ carrier L" "y \ carrier L" + with inf_of_two_exists obtain s where "greatest L s (Lower L {x, y})" by fast with L - show "greatest (THE xa. greatest xa (Lower {x, y})) (Lower {x, y})" + show "greatest L (THE xa. greatest L xa (Lower L {x, y})) (Lower L {x, y})" by (fast intro: theI2 greatest_unique) (* blast fails *) qed lemma (in lattice) meet_le: assumes sub: "z \ x" "z \ y" - and L: "x \ L" "y \ L" "z \ L" + and L: "x \ carrier L" "y \ carrier L" "z \ carrier L" shows "z \ x \ y" proof (rule meetI) fix i - assume "greatest i (Lower {x, y})" + assume "greatest L i (Lower L {x, y})" with sub L show "z \ i" by (fast elim: greatest_le intro: Lower_memI) qed lemma (in lattice) meet_assoc_lemma: - assumes L: "x \ L" "y \ L" "z \ L" + assumes L: "x \ carrier L" "y \ carrier L" "z \ carrier L" shows "x \ (y \ z) = \{x, y, z}" proof (rule finite_inf_insertI) txt {* The textbook argument in Jacobson I, p 457 *} fix i - assume inf: "greatest i (Lower {x, y, z})" + assume inf: "greatest L i (Lower L {x, y, z})" show "x \ (y \ z) = i" proof (rule anti_sym) from inf L show "i \ x \ (y \ z)" @@ -627,15 +626,16 @@ from inf L show "x \ (y \ z) \ i" by (erule_tac greatest_le) (blast intro!: Lower_memI intro: trans meet_left meet_right meet_closed) - qed (simp_all add: L greatest_closed [OF inf]) + qed (simp_all add: L greatest_carrier [OF inf]) qed (simp_all add: L) -lemma (in order_syntax) meet_comm: - "x \ y = y \ x" +lemma meet_comm: + fixes L (structure) + shows "x \ y = y \ x" by (unfold meet_def) (simp add: insert_commute) lemma (in lattice) meet_assoc: - assumes L: "x \ L" "y \ L" "z \ L" + assumes L: "x \ carrier L" "y \ carrier L" "z \ carrier L" shows "(x \ y) \ z = x \ (y \ z)" proof - have "(x \ y) \ z = z \ (x \ y)" by (simp only: meet_comm) @@ -649,52 +649,51 @@ subsection {* Total Orders *} locale total_order = lattice + - assumes total: "[| x \ L; y \ L |] ==> x \ y | y \ x" - + assumes total: "[| x \ carrier L; y \ carrier L |] ==> x \ y | y \ x" text {* Introduction rule: the usual definition of total order *} lemma (in partial_order) total_orderI: - assumes total: "!!x y. [| x \ L; y \ L |] ==> x \ y | y \ x" - shows "total_order L le" + assumes total: "!!x y. [| x \ carrier L; y \ carrier L |] ==> x \ y | y \ x" + shows "total_order L" proof intro_locales - show "lattice_axioms L le" + show "lattice_axioms L" proof (rule lattice_axioms.intro) fix x y - assume L: "x \ L" "y \ L" - show "EX s. least s (Upper {x, y})" + assume L: "x \ carrier L" "y \ carrier L" + show "EX s. least L s (Upper L {x, y})" proof - note total L moreover { assume "x \ y" - with L have "least y (Upper {x, y})" + with L have "least L y (Upper L {x, y})" by (rule_tac least_UpperI) auto } moreover { assume "y \ x" - with L have "least x (Upper {x, y})" + with L have "least L x (Upper L {x, y})" by (rule_tac least_UpperI) auto } ultimately show ?thesis by blast qed next fix x y - assume L: "x \ L" "y \ L" - show "EX i. greatest i (Lower {x, y})" + assume L: "x \ carrier L" "y \ carrier L" + show "EX i. greatest L i (Lower L {x, y})" proof - note total L moreover { assume "y \ x" - with L have "greatest y (Lower {x, y})" + with L have "greatest L y (Lower L {x, y})" by (rule_tac greatest_LowerI) auto } moreover { assume "x \ y" - with L have "greatest x (Lower {x, y})" + with L have "greatest L x (Lower L {x, y})" by (rule_tac greatest_LowerI) auto } ultimately show ?thesis by blast @@ -707,98 +706,97 @@ locale complete_lattice = lattice + assumes sup_exists: - "[| A \ L |] ==> EX s. order_syntax.least L le s (order_syntax.Upper L le A)" + "[| A \ carrier L |] ==> EX s. least L s (Upper L A)" and inf_exists: - "[| A \ L |] ==> EX i. order_syntax.greatest L le i (order_syntax.Lower L le A)" + "[| A \ carrier L |] ==> EX i. greatest L i (Lower L A)" text {* Introduction rule: the usual definition of complete lattice *} lemma (in partial_order) complete_latticeI: assumes sup_exists: - "!!A. [| A \ L |] ==> EX s. least s (Upper A)" + "!!A. [| A \ carrier L |] ==> EX s. least L s (Upper L A)" and inf_exists: - "!!A. [| A \ L |] ==> EX i. greatest i (Lower A)" - shows "complete_lattice L le" + "!!A. [| A \ carrier L |] ==> EX i. greatest L i (Lower L A)" + shows "complete_lattice L" proof intro_locales - show "lattice_axioms L le" + show "lattice_axioms L" by (rule lattice_axioms.intro) (blast intro: sup_exists inf_exists)+ qed (assumption | rule complete_lattice_axioms.intro)+ -definition (in order_syntax) - top ("\") where - "\ == sup L" +constdefs (structure L) + top :: "_ => 'a" ("\\") + "\ == sup L (carrier L)" -definition (in order_syntax) - bottom ("\") where - "\ == inf L" + bottom :: "_ => 'a" ("\\") + "\ == inf L (carrier L)" lemma (in complete_lattice) supI: - "[| !!l. least l (Upper A) ==> P l; A \ L |] + "[| !!l. least L l (Upper L A) ==> P l; A \ carrier L |] ==> P (\A)" proof (unfold sup_def) - assume L: "A \ L" - and P: "!!l. least l (Upper A) ==> P l" - with sup_exists obtain s where "least s (Upper A)" by blast - with L show "P (THE l. least l (Upper A))" + assume L: "A \ carrier L" + and P: "!!l. least L l (Upper L A) ==> P l" + with sup_exists obtain s where "least L s (Upper L A)" by blast + with L show "P (THE l. least L l (Upper L A))" by (fast intro: theI2 least_unique P) qed lemma (in complete_lattice) sup_closed [simp]: - "A \ L ==> \A \ L" + "A \ carrier L ==> \A \ carrier L" by (rule supI) simp_all lemma (in complete_lattice) top_closed [simp, intro]: - "\ \ L" + "\ \ carrier L" by (unfold top_def) simp lemma (in complete_lattice) infI: - "[| !!i. greatest i (Lower A) ==> P i; A \ L |] + "[| !!i. greatest L i (Lower L A) ==> P i; A \ carrier L |] ==> P (\A)" proof (unfold inf_def) - assume L: "A \ L" - and P: "!!l. greatest l (Lower A) ==> P l" - with inf_exists obtain s where "greatest s (Lower A)" by blast - with L show "P (THE l. greatest l (Lower A))" + assume L: "A \ carrier L" + and P: "!!l. greatest L l (Lower L A) ==> P l" + with inf_exists obtain s where "greatest L s (Lower L A)" by blast + with L show "P (THE l. greatest L l (Lower L A))" by (fast intro: theI2 greatest_unique P) qed lemma (in complete_lattice) inf_closed [simp]: - "A \ L ==> \A \ L" + "A \ carrier L ==> \A \ carrier L" by (rule infI) simp_all lemma (in complete_lattice) bottom_closed [simp, intro]: - "\ \ L" + "\ \ carrier L" by (unfold bottom_def) simp text {* Jacobson: Theorem 8.1 *} -lemma (in order_syntax) Lower_empty [simp]: - "Lower {} = L" +lemma Lower_empty [simp]: + "Lower L {} = carrier L" by (unfold Lower_def) simp -lemma (in order_syntax) Upper_empty [simp]: - "Upper {} = L" +lemma Upper_empty [simp]: + "Upper L {} = carrier L" by (unfold Upper_def) simp theorem (in partial_order) complete_lattice_criterion1: - assumes top_exists: "EX g. greatest g L" + assumes top_exists: "EX g. greatest L g (carrier L)" and inf_exists: - "!!A. [| A \ L; A ~= {} |] ==> EX i. greatest i (Lower A)" - shows "complete_lattice L le" + "!!A. [| A \ carrier L; A ~= {} |] ==> EX i. greatest L i (Lower L A)" + shows "complete_lattice L" proof (rule complete_latticeI) - from top_exists obtain top where top: "greatest top L" .. + from top_exists obtain top where top: "greatest L top (carrier L)" .. fix A - assume L: "A \ L" - let ?B = "Upper A" + assume L: "A \ carrier L" + let ?B = "Upper L A" from L top have "top \ ?B" by (fast intro!: Upper_memI intro: greatest_le) then have B_non_empty: "?B ~= {}" by fast - have B_L: "?B \ L" by simp + have B_L: "?B \ carrier L" by simp from inf_exists [OF B_L B_non_empty] - obtain b where b_inf_B: "greatest b (Lower ?B)" .. - have "least b (Upper A)" + obtain b where b_inf_B: "greatest L b (Lower L ?B)" .. + have "least L b (Upper L A)" apply (rule least_UpperI) - apply (rule greatest_le [where A = "Lower ?B"]) + apply (rule greatest_le [where A = "Lower L ?B"]) apply (rule b_inf_B) apply (rule Lower_memI) apply (erule UpperD) @@ -808,13 +806,13 @@ apply (erule greatest_Lower_above [OF b_inf_B]) apply simp apply (rule L) -apply (rule greatest_closed [OF b_inf_B]) (* rename rule: _closed *) +apply (rule greatest_carrier [OF b_inf_B]) (* rename rule: _closed *) done - then show "EX s. least s (Upper A)" .. + then show "EX s. least L s (Upper L A)" .. next fix A - assume L: "A \ L" - show "EX i. greatest i (Lower A)" + assume L: "A \ carrier L" + show "EX i. greatest L i (Lower L A)" proof (cases "A = {}") case True then show ?thesis by (simp add: top_exists) @@ -832,25 +830,25 @@ subsubsection {* Powerset of a Set is a Complete Lattice *} theorem powerset_is_complete_lattice: - "complete_lattice (Pow A) (op \)" - (is "complete_lattice ?L ?le") + "complete_lattice (| carrier = Pow A, le = op \ |)" + (is "complete_lattice ?L") proof (rule partial_order.complete_latticeI) - show "partial_order ?L ?le" + show "partial_order ?L" by (rule partial_order.intro) auto next fix B - assume "B \ ?L" - then have "order_syntax.least ?L ?le (\ B) (order_syntax.Upper ?L ?le B)" - by (fastsimp intro!: order_syntax.least_UpperI simp: order_syntax.Upper_def) - then show "EX s. order_syntax.least ?L ?le s (order_syntax.Upper ?L ?le B)" .. + assume "B \ carrier ?L" + then have "least ?L (\ B) (Upper ?L B)" + by (fastsimp intro!: least_UpperI simp: Upper_def) + then show "EX s. least ?L s (Upper ?L B)" .. next fix B - assume "B \ ?L" - then have "order_syntax.greatest ?L ?le (\ B \ A) (order_syntax.Lower ?L ?le B)" + assume "B \ carrier ?L" + then have "greatest ?L (\ B \ A) (Lower ?L B)" txt {* @{term "\ B"} is not the infimum of @{term B}: @{term "\ {} = UNIV"} which is in general bigger than @{term "A"}! *} - by (fastsimp intro!: order_syntax.greatest_LowerI simp: order_syntax.Lower_def) - then show "EX i. order_syntax.greatest ?L ?le i (order_syntax.Lower ?L ?le B)" .. + by (fastsimp intro!: greatest_LowerI simp: Lower_def) + then show "EX i. greatest ?L i (Lower ?L B)" .. qed text {* An other example, that of the lattice of subgroups of a group,