--- a/src/HOL/Algebra/Group.thy Sat May 01 22:04:14 2004 +0200
+++ b/src/HOL/Algebra/Group.thy Sat May 01 22:05:05 2004 +0200
@@ -42,10 +42,10 @@
pow :: "[('a, 'm) monoid_scheme, 'a, 'b::number] => 'a" (infixr "'(^')\<index>" 75)
defs (overloaded)
- nat_pow_def: "pow G a n == nat_rec (one G) (%u b. mult G b a) n"
+ nat_pow_def: "pow G a n == nat_rec \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a) n"
int_pow_def: "pow G a z ==
- let p = nat_rec (one G) (%u b. mult G b a)
- in if neg z then m_inv G (p (nat (-z))) else p (nat z)"
+ let p = nat_rec \<one>\<^bsub>G\<^esub> (%u b. b \<otimes>\<^bsub>G\<^esub> a)
+ in if neg z then inv\<^bsub>G\<^esub> (p (nat (-z))) else p (nat z)"
locale magma = struct G +
assumes m_closed [intro, simp]:
@@ -62,14 +62,15 @@
and r_one [simp]: "x \<in> carrier G ==> x \<otimes> \<one> = x"
lemma monoidI:
+ includes struct G
assumes m_closed:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y \<in> carrier G"
- and one_closed: "one G \<in> carrier G"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G"
+ and one_closed: "\<one> \<in> carrier G"
and m_assoc:
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==>
- mult G (mult G x y) z = mult G x (mult G y z)"
- and l_one: "!!x. x \<in> carrier G ==> mult G (one G) x = x"
- and r_one: "!!x. x \<in> carrier G ==> mult G x (one G) = x"
+ (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
+ and l_one: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x"
+ and r_one: "!!x. x \<in> carrier G ==> x \<otimes> \<one> = x"
shows "monoid G"
by (fast intro!: monoid.intro magma.intro semigroup_axioms.intro
semigroup.intro monoid_axioms.intro
@@ -80,8 +81,8 @@
by (unfold Units_def) fast
lemma (in monoid) inv_unique:
- assumes eq: "y \<otimes> x = \<one>" "x \<otimes> y' = \<one>"
- and G: "x \<in> carrier G" "y \<in> carrier G" "y' \<in> carrier G"
+ assumes eq: "y \<otimes> x = \<one>" "x \<otimes> y' = \<one>"
+ and G: "x \<in> carrier G" "y \<in> carrier G" "y' \<in> carrier G"
shows "y = y'"
proof -
from G eq have "y = y \<otimes> (x \<otimes> y')" by simp
@@ -129,13 +130,13 @@
(x \<otimes> y = x \<otimes> z) = (y = z)"
proof
assume eq: "x \<otimes> y = x \<otimes> z"
- and G: "x \<in> Units G" "y \<in> carrier G" "z \<in> carrier G"
+ and G: "x \<in> Units G" "y \<in> carrier G" "z \<in> carrier G"
then have "(inv x \<otimes> x) \<otimes> y = (inv x \<otimes> x) \<otimes> z"
by (simp add: m_assoc Units_closed)
with G show "y = z" by (simp add: Units_l_inv)
next
assume eq: "y = z"
- and G: "x \<in> Units G" "y \<in> carrier G" "z \<in> carrier G"
+ and G: "x \<in> Units G" "y \<in> carrier G" "z \<in> carrier G"
then show "x \<otimes> y = x \<otimes> z" by simp
qed
@@ -152,14 +153,14 @@
"inj_on (m_inv G) (Units G)"
proof (rule inj_onI)
fix x y
- assume G: "x \<in> Units G" "y \<in> Units G" and eq: "inv x = inv y"
+ assume G: "x \<in> Units G" "y \<in> Units G" and eq: "inv x = inv y"
then have "inv (inv x) = inv (inv y)" by simp
with G show "x = y" by simp
qed
lemma (in monoid) Units_inv_comm:
assumes inv: "x \<otimes> y = \<one>"
- and G: "x \<in> Units G" "y \<in> Units G"
+ and G: "x \<in> Units G" "y \<in> Units G"
shows "y \<otimes> x = \<one>"
proof -
from G have "x \<otimes> y \<otimes> x = x \<otimes> \<one>" by (auto simp add: inv Units_closed)
@@ -200,59 +201,58 @@
assumes Units: "carrier G <= Units G"
theorem groupI:
+ includes struct G
assumes m_closed [simp]:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y \<in> carrier G"
- and one_closed [simp]: "one G \<in> carrier G"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G"
+ and one_closed [simp]: "\<one> \<in> carrier G"
and m_assoc:
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==>
- mult G (mult G x y) z = mult G x (mult G y z)"
- and l_one [simp]: "!!x. x \<in> carrier G ==> mult G (one G) x = x"
- and l_inv_ex: "!!x. x \<in> carrier G ==> EX y : carrier G. mult G y x = one G"
+ (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
+ and l_one [simp]: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x"
+ and l_inv_ex: "!!x. x \<in> carrier G ==> EX y : carrier G. y \<otimes> x = \<one>"
shows "group G"
proof -
have l_cancel [simp]:
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==>
- (mult G x y = mult G x z) = (y = z)"
+ (x \<otimes> y = x \<otimes> z) = (y = z)"
proof
fix x y z
- assume eq: "mult G x y = mult G x z"
- and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
+ assume eq: "x \<otimes> y = x \<otimes> z"
+ and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
with l_inv_ex obtain x_inv where xG: "x_inv \<in> carrier G"
- and l_inv: "mult G x_inv x = one G" by fast
- from G eq xG have "mult G (mult G x_inv x) y = mult G (mult G x_inv x) z"
+ and l_inv: "x_inv \<otimes> x = \<one>" by fast
+ from G eq xG have "(x_inv \<otimes> x) \<otimes> y = (x_inv \<otimes> x) \<otimes> z"
by (simp add: m_assoc)
with G show "y = z" by (simp add: l_inv)
next
fix x y z
assume eq: "y = z"
- and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
- then show "mult G x y = mult G x z" by simp
+ and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
+ then show "x \<otimes> y = x \<otimes> z" by simp
qed
have r_one:
- "!!x. x \<in> carrier G ==> mult G x (one G) = x"
+ "!!x. x \<in> carrier G ==> x \<otimes> \<one> = x"
proof -
fix x
assume x: "x \<in> carrier G"
with l_inv_ex obtain x_inv where xG: "x_inv \<in> carrier G"
- and l_inv: "mult G x_inv x = one G" by fast
- from x xG have "mult G x_inv (mult G x (one G)) = mult G x_inv x"
+ and l_inv: "x_inv \<otimes> x = \<one>" by fast
+ from x xG have "x_inv \<otimes> (x \<otimes> \<one>) = x_inv \<otimes> x"
by (simp add: m_assoc [symmetric] l_inv)
- with x xG show "mult G x (one G) = x" by simp
+ with x xG show "x \<otimes> \<one> = x" by simp
qed
have inv_ex:
- "!!x. x \<in> carrier G ==> EX y : carrier G. mult G y x = one G &
- mult G x y = one G"
+ "!!x. x \<in> carrier G ==> EX y : carrier G. y \<otimes> x = \<one> & x \<otimes> y = \<one>"
proof -
fix x
assume x: "x \<in> carrier G"
with l_inv_ex obtain y where y: "y \<in> carrier G"
- and l_inv: "mult G y x = one G" by fast
- from x y have "mult G y (mult G x y) = mult G y (one G)"
+ and l_inv: "y \<otimes> x = \<one>" by fast
+ from x y have "y \<otimes> (x \<otimes> y) = y \<otimes> \<one>"
by (simp add: m_assoc [symmetric] l_inv r_one)
- with x y have r_inv: "mult G x y = one G"
+ with x y have r_inv: "x \<otimes> y = \<one>"
by simp
- from x y show "EX y : carrier G. mult G y x = one G &
- mult G x y = one G"
+ from x y show "EX y : carrier G. y \<otimes> x = \<one> & x \<otimes> y = \<one>"
by (fast intro: l_inv r_inv)
qed
then have carrier_subset_Units: "carrier G <= Units G"
@@ -265,7 +265,7 @@
lemma (in monoid) monoid_groupI:
assumes l_inv_ex:
- "!!x. x \<in> carrier G ==> EX y : carrier G. mult G y x = one G"
+ "!!x. x \<in> carrier G ==> EX y : carrier G. y \<otimes> x = \<one>"
shows "group G"
by (rule groupI) (auto intro: m_assoc l_inv_ex)
@@ -306,13 +306,13 @@
(y \<otimes> x = z \<otimes> x) = (y = z)"
proof
assume eq: "y \<otimes> x = z \<otimes> x"
- and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
+ and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
then have "y \<otimes> (x \<otimes> inv x) = z \<otimes> (x \<otimes> inv x)"
by (simp add: m_assoc [symmetric])
with G show "y = z" by (simp add: r_inv)
next
assume eq: "y = z"
- and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
+ and G: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
then show "y \<otimes> x = z \<otimes> x" by simp
qed
@@ -335,7 +335,7 @@
lemma (in group) inv_mult_group:
"[| x \<in> carrier G; y \<in> carrier G |] ==> inv (x \<otimes> y) = inv y \<otimes> inv x"
proof -
- assume G: "x \<in> carrier G" "y \<in> carrier G"
+ assume G: "x \<in> carrier G" "y \<in> carrier G"
then have "inv (x \<otimes> y) \<otimes> (x \<otimes> y) = (inv y \<otimes> inv x) \<otimes> (x \<otimes> y)"
by (simp add: m_assoc l_inv) (simp add: m_assoc [symmetric] l_inv)
with G show ?thesis by simp
@@ -343,14 +343,14 @@
lemma (in group) inv_comm:
"[| x \<otimes> y = \<one>; x \<in> carrier G; y \<in> carrier G |] ==> y \<otimes> x = \<one>"
- by (rule Units_inv_comm) auto
+ by (rule Units_inv_comm) auto
lemma (in group) inv_equality:
"[|y \<otimes> x = \<one>; x \<in> carrier G; y \<in> carrier G|] ==> inv x = y"
apply (simp add: m_inv_def)
apply (rule the_equality)
- apply (simp add: inv_comm [of y x])
-apply (rule r_cancel [THEN iffD1], auto)
+ apply (simp add: inv_comm [of y x])
+apply (rule r_cancel [THEN iffD1], auto)
done
text {* Power *}
@@ -493,10 +493,10 @@
constdefs (structure G and H)
DirProdSemigroup :: "_ => _ => ('a \<times> 'b) semigroup" (infixr "\<times>\<^sub>s" 80)
"G \<times>\<^sub>s H == (| carrier = carrier G \<times> carrier H,
- mult = (%(xg, xh) (yg, yh). (xg \<otimes> yg, xh \<otimes>\<^sub>2 yh)) |)"
+ mult = (%(g, h) (g', h'). (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')) |)"
DirProdGroup :: "_ => _ => ('a \<times> 'b) monoid" (infixr "\<times>\<^sub>g" 80)
- "G \<times>\<^sub>g H == semigroup.extend (G \<times>\<^sub>s H) (| one = (\<one>, \<one>\<^sub>2) |)"
+ "G \<times>\<^sub>g H == semigroup.extend (G \<times>\<^sub>s H) (| one = (\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>) |)"
lemma DirProdSemigroup_magma:
includes magma G + magma H
@@ -550,18 +550,18 @@
by (simp add: DirProdGroup_def DirProdSemigroup_def semigroup.defs)
lemma one_DirProdGroup [simp]:
- "one (G \<times>\<^sub>g H) = (one G, one H)"
+ "\<one>\<^bsub>(G \<times>\<^sub>g H)\<^esub> = (\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>)"
by (simp add: DirProdGroup_def DirProdSemigroup_def semigroup.defs)
lemma mult_DirProdGroup [simp]:
- "mult (G \<times>\<^sub>g H) (g, h) (g', h') = (mult G g g', mult H h h')"
+ "(g, h) \<otimes>\<^bsub>(G \<times>\<^sub>g H)\<^esub> (g', h') = (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')"
by (simp add: DirProdGroup_def DirProdSemigroup_def semigroup.defs)
lemma inv_DirProdGroup [simp]:
includes group G + group H
assumes g: "g \<in> carrier G"
and h: "h \<in> carrier H"
- shows "m_inv (G \<times>\<^sub>g H) (g, h) = (m_inv G g, m_inv H h)"
+ shows "m_inv (G \<times>\<^sub>g H) (g, h) = (inv\<^bsub>G\<^esub> g, inv\<^bsub>H\<^esub> h)"
apply (rule group.inv_equality [OF DirProdGroup_group])
apply (simp_all add: prems group_def group.l_inv)
done
@@ -572,7 +572,7 @@
hom :: "_ => _ => ('a => 'b) set"
"hom G H ==
{h. h \<in> carrier G -> carrier H &
- (\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes> y) = h x \<otimes>\<^sub>2 h y)}"
+ (\<forall>x \<in> carrier G. \<forall>y \<in> carrier G. h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y)}"
lemma (in semigroup) hom:
includes semigroup G
@@ -586,9 +586,9 @@
qed
lemma hom_mult:
- "[| h \<in> hom G H; x \<in> carrier G; y \<in> carrier G |]
- ==> h (mult G x y) = mult H (h x) (h y)"
- by (simp add: hom_def)
+ "[| h \<in> hom G H; x \<in> carrier G; y \<in> carrier G |]
+ ==> h (x \<otimes>\<^bsub>G\<^esub> y) = h x \<otimes>\<^bsub>H\<^esub> h y"
+ by (simp add: hom_def)
lemma hom_closed:
"[| h \<in> hom G H; x \<in> carrier G |] ==> h x \<in> carrier H"
@@ -598,9 +598,9 @@
"[|group G; h \<in> hom G G; h' \<in> hom G G; h' \<in> carrier G -> carrier G|]
==> compose (carrier G) h h' \<in> hom G G"
apply (simp (no_asm_simp) add: hom_def)
-apply (intro conjI)
+apply (intro conjI)
apply (force simp add: funcset_compose hom_def)
-apply (simp add: compose_def group.axioms hom_mult funcset_mem)
+apply (simp add: compose_def group.axioms hom_mult funcset_mem)
done
locale group_hom = group G + group H + var h +
@@ -613,9 +613,9 @@
by simp
lemma (in group_hom) hom_one [simp]:
- "h \<one> = \<one>\<^sub>2"
+ "h \<one> = \<one>\<^bsub>H\<^esub>"
proof -
- have "h \<one> \<otimes>\<^sub>2 \<one>\<^sub>2 = h \<one> \<otimes>\<^sub>2 h \<one>"
+ have "h \<one> \<otimes>\<^bsub>H\<^esub> \<one>\<^bsub>H\<^esub> = h \<one> \<otimes>\<^sub>2 h \<one>"
by (simp add: hom_mult [symmetric] del: hom_mult)
then show ?thesis by (simp del: r_one)
qed
@@ -625,14 +625,14 @@
by simp
lemma (in group_hom) hom_inv [simp]:
- "x \<in> carrier G ==> h (inv x) = inv\<^sub>2 (h x)"
+ "x \<in> carrier G ==> h (inv x) = inv\<^bsub>H\<^esub> (h x)"
proof -
assume x: "x \<in> carrier G"
- then have "h x \<otimes>\<^sub>2 h (inv x) = \<one>\<^sub>2"
+ then have "h x \<otimes>\<^bsub>H\<^esub> h (inv x) = \<one>\<^bsub>H\<^esub>"
by (simp add: hom_mult [symmetric] G.r_inv del: hom_mult)
- also from x have "... = h x \<otimes>\<^sub>2 inv\<^sub>2 (h x)"
+ also from x have "... = h x \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)"
by (simp add: hom_mult [symmetric] H.r_inv del: hom_mult)
- finally have "h x \<otimes>\<^sub>2 h (inv x) = h x \<otimes>\<^sub>2 inv\<^sub>2 (h x)" .
+ finally have "h x \<otimes>\<^bsub>H\<^esub> h (inv x) = h x \<otimes>\<^bsub>H\<^esub> inv\<^bsub>H\<^esub> (h x)" .
with x show ?thesis by simp
qed
@@ -653,7 +653,7 @@
"[| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==>
x \<otimes> (y \<otimes> z) = y \<otimes> (x \<otimes> z)"
proof -
- assume xyz: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
+ assume xyz: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
from xyz have "x \<otimes> (y \<otimes> z) = (x \<otimes> y) \<otimes> z" by (simp add: m_assoc)
also from xyz have "... = (y \<otimes> x) \<otimes> z" by (simp add: m_comm)
also from xyz have "... = y \<otimes> (x \<otimes> z)" by (simp add: m_assoc)
@@ -665,15 +665,16 @@
locale comm_monoid = comm_semigroup + monoid
lemma comm_monoidI:
+ includes struct G
assumes m_closed:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y \<in> carrier G"
- and one_closed: "one G \<in> carrier G"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G"
+ and one_closed: "\<one> \<in> carrier G"
and m_assoc:
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==>
- mult G (mult G x y) z = mult G x (mult G y z)"
- and l_one: "!!x. x \<in> carrier G ==> mult G (one G) x = x"
+ (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
+ and l_one: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x"
and m_comm:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y = mult G y x"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x"
shows "comm_monoid G"
using l_one
by (auto intro!: comm_monoid.intro magma.intro semigroup_axioms.intro
@@ -682,20 +683,17 @@
lemma (in monoid) monoid_comm_monoidI:
assumes m_comm:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y = mult G y x"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x"
shows "comm_monoid G"
by (rule comm_monoidI) (auto intro: m_assoc m_comm)
-(*
-lemma (in comm_monoid) r_one [simp]:
+(*lemma (in comm_monoid) r_one [simp]:
"x \<in> carrier G ==> x \<otimes> \<one> = x"
proof -
assume G: "x \<in> carrier G"
then have "x \<otimes> \<one> = \<one> \<otimes> x" by (simp add: m_comm)
also from G have "... = x" by simp
finally show ?thesis .
-qed
-*)
-
+qed*)
lemma (in comm_monoid) nat_pow_distr:
"[| x \<in> carrier G; y \<in> carrier G |] ==>
(x \<otimes> y) (^) (n::nat) = x (^) n \<otimes> y (^) n"
@@ -705,22 +703,23 @@
lemma (in group) group_comm_groupI:
assumes m_comm: "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==>
- mult G x y = mult G y x"
+ x \<otimes> y = y \<otimes> x"
shows "comm_group G"
by (fast intro: comm_group.intro comm_semigroup_axioms.intro
group.axioms prems)
lemma comm_groupI:
+ includes struct G
assumes m_closed:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y \<in> carrier G"
- and one_closed: "one G \<in> carrier G"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y \<in> carrier G"
+ and one_closed: "\<one> \<in> carrier G"
and m_assoc:
"!!x y z. [| x \<in> carrier G; y \<in> carrier G; z \<in> carrier G |] ==>
- mult G (mult G x y) z = mult G x (mult G y z)"
+ (x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
and m_comm:
- "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> mult G x y = mult G y x"
- and l_one: "!!x. x \<in> carrier G ==> mult G (one G) x = x"
- and l_inv_ex: "!!x. x \<in> carrier G ==> EX y : carrier G. mult G y x = one G"
+ "!!x y. [| x \<in> carrier G; y \<in> carrier G |] ==> x \<otimes> y = y \<otimes> x"
+ and l_one: "!!x. x \<in> carrier G ==> \<one> \<otimes> x = x"
+ and l_inv_ex: "!!x. x \<in> carrier G ==> EX y : carrier G. y \<otimes> x = \<one>"
shows "comm_group G"
by (fast intro: group.group_comm_groupI groupI prems)
--- a/src/HOL/Algebra/Lattice.thy Sat May 01 22:04:14 2004 +0200
+++ b/src/HOL/Algebra/Lattice.thy Sat May 01 22:05:05 2004 +0200
@@ -14,9 +14,7 @@
record 'a order = "'a partial_object" +
le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50)
-locale order_syntax = struct L
-
-locale partial_order = order_syntax +
+locale partial_order = struct L +
assumes refl [intro, simp]:
"x \<in> carrier L ==> x \<sqsubseteq> x"
and anti_sym [intro]:
@@ -31,19 +29,19 @@
-- {* Upper and lower bounds of a set. *}
Upper :: "[_, 'a set] => 'a set"
- "Upper L A == {u. (ALL x. x \<in> A \<inter> carrier L --> le L x u)} \<inter>
+ "Upper L A == {u. (ALL x. x \<in> A \<inter> carrier L --> x \<sqsubseteq> u)} \<inter>
carrier L"
Lower :: "[_, 'a set] => 'a set"
- "Lower L A == {l. (ALL x. x \<in> A \<inter> carrier L --> le L l x)} \<inter>
+ "Lower L A == {l. (ALL x. x \<in> A \<inter> carrier L --> l \<sqsubseteq> x)} \<inter>
carrier L"
-- {* Least and greatest, as predicate. *}
least :: "[_, 'a, 'a set] => bool"
- "least L l A == A \<subseteq> carrier L & l \<in> A & (ALL x : A. le L l x)"
+ "least L l A == A \<subseteq> carrier L & l \<in> A & (ALL x : A. l \<sqsubseteq> x)"
greatest :: "[_, 'a, 'a set] => bool"
- "greatest L g A == A \<subseteq> carrier L & g \<in> A & (ALL x : A. le L x g)"
+ "greatest L g A == A \<subseteq> carrier L & g \<in> A & (ALL x : A. x \<sqsubseteq> g)"
-- {* Supremum and infimum *}
sup :: "[_, 'a set] => 'a" ("\<Squnion>\<index>_" [90] 90)
@@ -66,14 +64,14 @@
by (unfold Upper_def) clarify
lemma UpperD [dest]:
- includes order_syntax
+ includes struct L
shows "[| u \<in> Upper L A; x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> u"
- by (unfold Upper_def) blast
+ by (unfold Upper_def) blast
lemma Upper_memI:
- includes order_syntax
+ includes struct L
shows "[| !! y. y \<in> A ==> y \<sqsubseteq> x; x \<in> carrier L |] ==> x \<in> Upper L A"
- by (unfold Upper_def) blast
+ by (unfold Upper_def) blast
lemma Upper_antimono:
"A \<subseteq> B ==> Upper L B \<subseteq> Upper L A"
@@ -87,14 +85,14 @@
by (unfold Lower_def) clarify
lemma LowerD [dest]:
- includes order_syntax
+ includes struct L
shows "[| l \<in> Lower L A; x \<in> A; A \<subseteq> carrier L |] ==> l \<sqsubseteq> x"
- by (unfold Lower_def) blast
+ by (unfold Lower_def) blast
lemma Lower_memI:
- includes order_syntax
+ includes struct L
shows "[| !! y. y \<in> A ==> x \<sqsubseteq> y; x \<in> carrier L |] ==> x \<in> Lower L A"
- by (unfold Lower_def) blast
+ by (unfold Lower_def) blast
lemma Lower_antimono:
"A \<subseteq> B ==> Lower L B \<subseteq> Lower L A"
@@ -116,22 +114,21 @@
by (unfold least_def) blast
lemma least_le:
- includes order_syntax
+ includes struct L
shows "[| least L x A; a \<in> A |] ==> x \<sqsubseteq> a"
by (unfold least_def) fast
lemma least_UpperI:
- includes order_syntax
+ includes struct L
assumes above: "!! x. x \<in> A ==> x \<sqsubseteq> s"
and below: "!! y. y \<in> Upper L A ==> s \<sqsubseteq> y"
- and L: "A \<subseteq> carrier L" "s \<in> carrier L"
+ and L: "A \<subseteq> carrier L" "s \<in> carrier L"
shows "least L s (Upper L A)"
-proof (unfold least_def, intro conjI)
- show "Upper L A \<subseteq> carrier L" by simp
-next
- from above L show "s \<in> Upper L A" by (simp add: Upper_def)
-next
- from below show "ALL x : Upper L A. s \<sqsubseteq> x" by fast
+proof -
+ have "Upper L A \<subseteq> carrier L" by simp
+ moreover from above L have "s \<in> Upper L A" by (simp add: Upper_def)
+ moreover from below have "ALL x : Upper L A. s \<sqsubseteq> x" by fast
+ ultimately show ?thesis by (simp add: least_def)
qed
@@ -150,24 +147,24 @@
by (unfold greatest_def) blast
lemma greatest_le:
- includes order_syntax
+ includes struct L
shows "[| greatest L x A; a \<in> A |] ==> a \<sqsubseteq> x"
by (unfold greatest_def) fast
lemma greatest_LowerI:
- includes order_syntax
+ includes struct L
assumes below: "!! x. x \<in> A ==> i \<sqsubseteq> x"
and above: "!! y. y \<in> Lower L A ==> y \<sqsubseteq> i"
- and L: "A \<subseteq> carrier L" "i \<in> carrier L"
+ and L: "A \<subseteq> carrier L" "i \<in> carrier L"
shows "greatest L i (Lower L A)"
-proof (unfold greatest_def, intro conjI)
- show "Lower L A \<subseteq> carrier L" by simp
-next
- from below L show "i \<in> Lower L A" by (simp add: Lower_def)
-next
- from above show "ALL x : Lower L A. x \<sqsubseteq> i" by fast
+proof -
+ have "Lower L A \<subseteq> carrier L" by simp
+ moreover from below L have "i \<in> Lower L A" by (simp add: Lower_def)
+ moreover from above have "ALL x : Lower L A. x \<sqsubseteq> i" by fast
+ ultimately show ?thesis by (simp add: greatest_def)
qed
+
subsection {* Lattices *}
locale lattice = partial_order +
@@ -177,12 +174,12 @@
"[| x \<in> carrier L; y \<in> carrier L |] ==> EX s. greatest L s (Lower L {x, y})"
lemma least_Upper_above:
- includes order_syntax
+ includes struct L
shows "[| least L s (Upper L A); x \<in> A; A \<subseteq> carrier L |] ==> x \<sqsubseteq> s"
by (unfold least_def) blast
lemma greatest_Lower_above:
- includes order_syntax
+ includes struct L
shows "[| greatest L i (Lower L A); x \<in> A; A \<subseteq> carrier L |] ==> i \<sqsubseteq> x"
by (unfold greatest_def) blast
@@ -193,11 +190,11 @@
"[| !!l. least L l (Upper L {x, y}) ==> P l; x \<in> carrier L; y \<in> carrier L |]
==> P (x \<squnion> y)"
proof (unfold join_def sup_def)
- assume L: "x \<in> carrier L" "y \<in> carrier L"
+ assume L: "x \<in> carrier L" "y \<in> 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)
+ by (fast intro: theI2 least_unique P)
qed
lemma (in lattice) join_closed [simp]:
@@ -209,8 +206,8 @@
by (rule least_UpperI) fast+
lemma (in partial_order) sup_of_singleton [simp]:
- includes order_syntax
- shows "x \<in> carrier L ==> \<Squnion> {x} = x"
+ includes struct L
+ shows "x \<in> carrier L ==> \<Squnion>{x} = x"
by (unfold sup_def) (blast intro: least_unique least_UpperI sup_of_singletonI)
@@ -219,129 +216,104 @@
lemma (in lattice) sup_insertI:
"[| !!s. least L s (Upper L (insert x A)) ==> P s;
least L a (Upper L A); x \<in> carrier L; A \<subseteq> carrier L |]
- ==> P (\<Squnion> (insert x A))"
+ ==> P (\<Squnion>(insert x A))"
proof (unfold sup_def)
- assume L: "x \<in> carrier L" "A \<subseteq> carrier L"
+ assume L: "x \<in> carrier L" "A \<subseteq> 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 \<in> carrier L" by simp
from L sup_of_two_exists least_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 [where a = s])
+ proof (rule theI2)
show "least L s (Upper L (insert x A))"
proof (rule least_UpperI)
fix z
- assume xA: "z \<in> insert x A"
- show "z \<sqsubseteq> s"
- proof -
- {
- assume "z = x" then have ?thesis
- by (simp add: least_Upper_above [OF least_s] L La)
- }
- moreover
- {
- assume "z \<in> A"
- with L least_s least_a have ?thesis
- by (rule_tac trans [where y = a]) (auto dest: least_Upper_above)
- }
- moreover note xA
- ultimately show ?thesis by blast
+ assume "z \<in> insert x A"
+ then show "z \<sqsubseteq> s"
+ proof
+ assume "z = x" then show ?thesis
+ by (simp add: least_Upper_above [OF least_s] L La)
+ next
+ assume "z \<in> A"
+ with L least_s least_a show ?thesis
+ by (rule_tac trans [where y = a]) (auto dest: least_Upper_above)
+ qed
+ next
+ fix y
+ assume y: "y \<in> Upper L (insert x A)"
+ show "s \<sqsubseteq> y"
+ proof (rule least_le [OF least_s], rule Upper_memI)
+ fix z
+ assume z: "z \<in> {a, x}"
+ then show "z \<sqsubseteq> y"
+ proof
+ have y': "y \<in> 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"
+ with y' least_a show ?thesis by (fast dest: least_le)
+ next
+ assume "z \<in> {x}" (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
+ with y L show ?thesis by blast
+ qed
+ qed (rule Upper_closed [THEN subsetD])
+ next
+ from L show "insert x A \<subseteq> carrier L" by simp
+ from least_s show "s \<in> carrier L" by simp
qed
next
- fix y
- assume y: "y \<in> Upper L (insert x A)"
- show "s \<sqsubseteq> y"
- proof (rule least_le [OF least_s], rule Upper_memI)
- fix z
- assume z: "z \<in> {a, x}"
- show "z \<sqsubseteq> y"
- proof -
- {
- have y': "y \<in> 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"
- with y' least_a have ?thesis by (fast dest: least_le)
- }
- moreover
- {
- assume "z = x"
- with y L have ?thesis by blast
- }
- moreover note z
- ultimately show ?thesis by blast
- qed
- qed (rule Upper_closed [THEN subsetD])
- next
- from L show "insert x A \<subseteq> carrier L" by simp
- next
- from least_s show "s \<in> carrier L" by simp
- qed
-next
fix l
assume least_l: "least L l (Upper L (insert x A))"
show "l = s"
proof (rule least_unique)
show "least L s (Upper L (insert x A))"
proof (rule least_UpperI)
- fix z
- assume xA: "z \<in> insert x A"
- show "z \<sqsubseteq> s"
- proof -
- {
- assume "z = x" then have ?thesis
- by (simp add: least_Upper_above [OF least_s] L La)
- }
- moreover
- {
- assume "z \<in> A"
- with L least_s least_a have ?thesis
- by (rule_tac trans [where y = a]) (auto dest: least_Upper_above)
- }
- moreover note xA
- ultimately show ?thesis by blast
+ fix z
+ assume "z \<in> insert x A"
+ then show "z \<sqsubseteq> s"
+ proof
+ assume "z = x" then show ?thesis
+ by (simp add: least_Upper_above [OF least_s] L La)
+ next
+ assume "z \<in> A"
+ with L least_s least_a show ?thesis
+ by (rule_tac trans [where y = a]) (auto dest: least_Upper_above)
qed
next
- fix y
- assume y: "y \<in> Upper L (insert x A)"
- show "s \<sqsubseteq> y"
- proof (rule least_le [OF least_s], rule Upper_memI)
- fix z
- assume z: "z \<in> {a, x}"
- show "z \<sqsubseteq> y"
- proof -
- {
- have y': "y \<in> 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"
- with y' least_a have ?thesis by (fast dest: least_le)
- }
- moreover
- {
- assume "z = x"
- with y L have ?thesis by blast
- }
- moreover note z
- ultimately show ?thesis by blast
- qed
- qed (rule Upper_closed [THEN subsetD])
+ fix y
+ assume y: "y \<in> Upper L (insert x A)"
+ show "s \<sqsubseteq> y"
+ proof (rule least_le [OF least_s], rule Upper_memI)
+ fix z
+ assume z: "z \<in> {a, x}"
+ then show "z \<sqsubseteq> y"
+ proof
+ have y': "y \<in> 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"
+ with y' least_a show ?thesis by (fast dest: least_le)
+ next
+ assume "z \<in> {x}"
+ with y L show ?thesis by blast
+ qed
+ qed (rule Upper_closed [THEN subsetD])
next
- from L show "insert x A \<subseteq> carrier L" by simp
- next
- from least_s show "s \<in> carrier L" by simp
+ from L show "insert x A \<subseteq> carrier L" by simp
+ from least_s show "s \<in> carrier L" by simp
qed
qed
qed
qed
lemma (in lattice) finite_sup_least:
- "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> least L (\<Squnion> A) (Upper L A)"
+ "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> least L (\<Squnion>A) (Upper L A)"
proof (induct set: Finites)
- case empty then show ?case by simp
+ case empty
+ then show ?case by simp
next
case (insert A x)
show ?case
@@ -350,16 +322,15 @@
with insert show ?thesis by (simp add: sup_of_singletonI)
next
case False
- from insert show ?thesis
- proof (rule_tac sup_insertI)
- from False insert show "least L (\<Squnion> A) (Upper L A)" by simp
- qed simp_all
+ with insert have "least L (\<Squnion>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 l (Upper L (insert x A)) ==> P l"
- and xA: "finite A" "x \<in> carrier L" "A \<subseteq> carrier L"
+ and xA: "finite A" "x \<in> carrier L" "A \<subseteq> carrier L"
shows "P (\<Squnion> (insert x A))"
proof (cases "A = {}")
case True with P and xA show ?thesis
@@ -370,44 +341,44 @@
qed
lemma (in lattice) finite_sup_closed:
- "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Squnion> A \<in> carrier L"
+ "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Squnion>A \<in> carrier L"
proof (induct set: Finites)
case empty then show ?case by simp
next
case (insert A x) then show ?case
- by (rule_tac finite_sup_insertI) (simp_all)
+ by - (rule finite_sup_insertI, simp_all)
qed
lemma (in lattice) join_left:
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqsubseteq> x \<squnion> y"
- by (rule joinI [folded join_def]) (blast dest: least_mem )
+ by (rule joinI [folded join_def]) (blast dest: least_mem)
lemma (in lattice) join_right:
"[| x \<in> carrier L; y \<in> carrier L |] ==> y \<sqsubseteq> x \<squnion> y"
- by (rule joinI [folded join_def]) (blast dest: least_mem )
+ by (rule joinI [folded join_def]) (blast dest: least_mem)
lemma (in lattice) sup_of_two_least:
- "[| x \<in> carrier L; y \<in> carrier L |] ==> least L (\<Squnion> {x, y}) (Upper L {x, y})"
+ "[| x \<in> carrier L; y \<in> carrier L |] ==> least L (\<Squnion>{x, y}) (Upper L {x, y})"
proof (unfold sup_def)
- assume L: "x \<in> carrier L" "y \<in> carrier L"
+ assume L: "x \<in> carrier L" "y \<in> 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 \<sqsubseteq> z" "y \<sqsubseteq> z"
- and L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
+ assumes sub: "x \<sqsubseteq> z" "y \<sqsubseteq> z"
+ and L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
shows "x \<squnion> y \<sqsubseteq> z"
proof (rule joinI)
fix s
assume "least L s (Upper L {x, y})"
with sub L show "s \<sqsubseteq> z" by (fast elim: least_le intro: Upper_memI)
qed
-
+
lemma (in lattice) join_assoc_lemma:
- assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
- shows "x \<squnion> (y \<squnion> z) = \<Squnion> {x, y, z}"
+ assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
+ shows "x \<squnion> (y \<squnion> z) = \<Squnion>{x, y, z}"
proof (rule finite_sup_insertI)
-- {* The textbook argument in Jacobson I, p 457 *}
fix s
@@ -424,21 +395,22 @@
qed (simp_all add: L)
lemma join_comm:
- includes order_syntax
+ includes struct L
shows "x \<squnion> y = y \<squnion> x"
by (unfold join_def) (simp add: insert_commute)
lemma (in lattice) join_assoc:
- assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
+ assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
shows "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)"
proof -
have "(x \<squnion> y) \<squnion> z = z \<squnion> (x \<squnion> y)" by (simp only: join_comm)
- also from L have "... = \<Squnion> {z, x, y}" by (simp add: join_assoc_lemma)
- also from L have "... = \<Squnion> {x, y, z}" by (simp add: insert_commute)
+ also from L have "... = \<Squnion>{z, x, y}" by (simp add: join_assoc_lemma)
+ also from L have "... = \<Squnion>{x, y, z}" by (simp add: insert_commute)
also from L have "... = x \<squnion> (y \<squnion> z)" by (simp add: join_assoc_lemma)
finally show ?thesis .
qed
+
subsubsection {* Infimum *}
lemma (in lattice) meetI:
@@ -446,7 +418,7 @@
x \<in> carrier L; y \<in> carrier L |]
==> P (x \<sqinter> y)"
proof (unfold meet_def inf_def)
- assume L: "x \<in> carrier L" "y \<in> carrier L"
+ assume L: "x \<in> carrier L" "y \<in> 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}))"
@@ -462,7 +434,7 @@
by (rule greatest_LowerI) fast+
lemma (in partial_order) inf_of_singleton [simp]:
- includes order_syntax
+ includes struct L
shows "x \<in> carrier L ==> \<Sqinter> {x} = x"
by (unfold inf_def) (blast intro: greatest_unique greatest_LowerI inf_of_singletonI)
@@ -471,127 +443,101 @@
lemma (in lattice) inf_insertI:
"[| !!i. greatest L i (Lower L (insert x A)) ==> P i;
greatest L a (Lower L A); x \<in> carrier L; A \<subseteq> carrier L |]
- ==> P (\<Sqinter> (insert x A))"
+ ==> P (\<Sqinter>(insert x A))"
proof (unfold inf_def)
- assume L: "x \<in> carrier L" "A \<subseteq> carrier L"
+ assume L: "x \<in> carrier L" "A \<subseteq> 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 \<in> carrier L" by simp
from L inf_of_two_exists greatest_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 [where a = i])
+ proof (rule theI2)
show "greatest L i (Lower L (insert x A))"
proof (rule greatest_LowerI)
fix z
- assume xA: "z \<in> insert x A"
- show "i \<sqsubseteq> z"
- proof -
- {
- assume "z = x" then have ?thesis
- by (simp add: greatest_Lower_above [OF greatest_i] L La)
- }
- moreover
- {
- assume "z \<in> A"
- with L greatest_i greatest_a have ?thesis
- by (rule_tac trans [where y = a]) (auto dest: greatest_Lower_above)
- }
- moreover note xA
- ultimately show ?thesis by blast
+ assume "z \<in> insert x A"
+ then show "i \<sqsubseteq> z"
+ proof
+ assume "z = x" then show ?thesis
+ by (simp add: greatest_Lower_above [OF greatest_i] L La)
+ next
+ assume "z \<in> A"
+ with L greatest_i greatest_a show ?thesis
+ by (rule_tac trans [where y = a]) (auto dest: greatest_Lower_above)
+ qed
+ next
+ fix y
+ assume y: "y \<in> Lower L (insert x A)"
+ show "y \<sqsubseteq> i"
+ proof (rule greatest_le [OF greatest_i], rule Lower_memI)
+ fix z
+ assume z: "z \<in> {a, x}"
+ then show "y \<sqsubseteq> z"
+ proof
+ have y': "y \<in> 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"
+ with y' greatest_a show ?thesis by (fast dest: greatest_le)
+ next
+ assume "z \<in> {x}"
+ with y L show ?thesis by blast
+ qed
+ qed (rule Lower_closed [THEN subsetD])
+ next
+ from L show "insert x A \<subseteq> carrier L" by simp
+ from greatest_i show "i \<in> carrier L" by simp
qed
next
- fix y
- assume y: "y \<in> Lower L (insert x A)"
- show "y \<sqsubseteq> i"
- proof (rule greatest_le [OF greatest_i], rule Lower_memI)
- fix z
- assume z: "z \<in> {a, x}"
- show "y \<sqsubseteq> z"
- proof -
- {
- have y': "y \<in> 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"
- with y' greatest_a have ?thesis by (fast dest: greatest_le)
- }
- moreover
- {
- assume "z = x"
- with y L have ?thesis by blast
- }
- moreover note z
- ultimately show ?thesis by blast
- qed
- qed (rule Lower_closed [THEN subsetD])
- next
- from L show "insert x A \<subseteq> carrier L" by simp
- next
- from greatest_i show "i \<in> carrier L" by simp
- qed
-next
fix g
assume greatest_g: "greatest L g (Lower L (insert x A))"
show "g = i"
proof (rule greatest_unique)
show "greatest L i (Lower L (insert x A))"
proof (rule greatest_LowerI)
- fix z
- assume xA: "z \<in> insert x A"
- show "i \<sqsubseteq> z"
- proof -
- {
- assume "z = x" then have ?thesis
- by (simp add: greatest_Lower_above [OF greatest_i] L La)
- }
- moreover
- {
- assume "z \<in> A"
- with L greatest_i greatest_a have ?thesis
- by (rule_tac trans [where y = a]) (auto dest: greatest_Lower_above)
- }
- moreover note xA
- ultimately show ?thesis by blast
- qed
+ fix z
+ assume "z \<in> insert x A"
+ then show "i \<sqsubseteq> z"
+ proof
+ assume "z = x" then show ?thesis
+ by (simp add: greatest_Lower_above [OF greatest_i] L La)
+ next
+ assume "z \<in> A"
+ with L greatest_i greatest_a show ?thesis
+ by (rule_tac trans [where y = a]) (auto dest: greatest_Lower_above)
+ qed
next
- fix y
- assume y: "y \<in> Lower L (insert x A)"
- show "y \<sqsubseteq> i"
- proof (rule greatest_le [OF greatest_i], rule Lower_memI)
- fix z
- assume z: "z \<in> {a, x}"
- show "y \<sqsubseteq> z"
- proof -
- {
- have y': "y \<in> 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"
- with y' greatest_a have ?thesis by (fast dest: greatest_le)
- }
- moreover
- {
- assume "z = x"
- with y L have ?thesis by blast
- }
- moreover note z
- ultimately show ?thesis by blast
+ fix y
+ assume y: "y \<in> Lower L (insert x A)"
+ show "y \<sqsubseteq> i"
+ proof (rule greatest_le [OF greatest_i], rule Lower_memI)
+ fix z
+ assume z: "z \<in> {a, x}"
+ then show "y \<sqsubseteq> z"
+ proof
+ have y': "y \<in> 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"
+ with y' greatest_a show ?thesis by (fast dest: greatest_le)
+ next
+ assume "z \<in> {x}"
+ with y L show ?thesis by blast
qed
- qed (rule Lower_closed [THEN subsetD])
+ qed (rule Lower_closed [THEN subsetD])
next
- from L show "insert x A \<subseteq> carrier L" by simp
- next
- from greatest_i show "i \<in> carrier L" by simp
+ from L show "insert x A \<subseteq> carrier L" by simp
+ from greatest_i show "i \<in> carrier L" by simp
qed
qed
qed
qed
lemma (in lattice) finite_inf_greatest:
- "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> greatest L (\<Sqinter> A) (Lower L A)"
+ "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> greatest L (\<Sqinter>A) (Lower L A)"
proof (induct set: Finites)
case empty then show ?case by simp
next
@@ -604,14 +550,14 @@
case False
from insert show ?thesis
proof (rule_tac inf_insertI)
- from False insert show "greatest L (\<Sqinter> A) (Lower L A)" by simp
+ from False insert show "greatest L (\<Sqinter>A) (Lower L A)" by simp
qed simp_all
qed
qed
lemma (in lattice) finite_inf_insertI:
assumes P: "!!i. greatest L i (Lower L (insert x A)) ==> P i"
- and xA: "finite A" "x \<in> carrier L" "A \<subseteq> carrier L"
+ and xA: "finite A" "x \<in> carrier L" "A \<subseteq> carrier L"
shows "P (\<Sqinter> (insert x A))"
proof (cases "A = {}")
case True with P and xA show ?thesis
@@ -622,7 +568,7 @@
qed
lemma (in lattice) finite_inf_closed:
- "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Sqinter> A \<in> carrier L"
+ "[| finite A; A \<subseteq> carrier L; A ~= {} |] ==> \<Sqinter>A \<in> carrier L"
proof (induct set: Finites)
case empty then show ?case by simp
next
@@ -632,17 +578,17 @@
lemma (in lattice) meet_left:
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<sqsubseteq> x"
- by (rule meetI [folded meet_def]) (blast dest: greatest_mem )
+ by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
lemma (in lattice) meet_right:
"[| x \<in> carrier L; y \<in> carrier L |] ==> x \<sqinter> y \<sqsubseteq> y"
- by (rule meetI [folded meet_def]) (blast dest: greatest_mem )
+ by (rule meetI [folded meet_def]) (blast dest: greatest_mem)
lemma (in lattice) inf_of_two_greatest:
"[| x \<in> carrier L; y \<in> carrier L |] ==>
greatest L (\<Sqinter> {x, y}) (Lower L {x, y})"
proof (unfold inf_def)
- assume L: "x \<in> carrier L" "y \<in> carrier L"
+ assume L: "x \<in> carrier L" "y \<in> carrier L"
with inf_of_two_exists obtain s where "greatest L s (Lower L {x, y})" by fast
with L
show "greatest L (THE xa. greatest L xa (Lower L {x, y})) (Lower L {x, y})"
@@ -650,18 +596,18 @@
qed
lemma (in lattice) meet_le:
- assumes sub: "z \<sqsubseteq> x" "z \<sqsubseteq> y"
- and L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
+ assumes sub: "z \<sqsubseteq> x" "z \<sqsubseteq> y"
+ and L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
shows "z \<sqsubseteq> x \<sqinter> y"
proof (rule meetI)
fix i
assume "greatest L i (Lower L {x, y})"
with sub L show "z \<sqsubseteq> i" by (fast elim: greatest_le intro: Lower_memI)
qed
-
+
lemma (in lattice) meet_assoc_lemma:
- assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
- shows "x \<sqinter> (y \<sqinter> z) = \<Sqinter> {x, y, z}"
+ assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
+ shows "x \<sqinter> (y \<sqinter> z) = \<Sqinter>{x, y, z}"
proof (rule finite_inf_insertI)
txt {* The textbook argument in Jacobson I, p 457 *}
fix i
@@ -678,12 +624,12 @@
qed (simp_all add: L)
lemma meet_comm:
- includes order_syntax
+ includes struct L
shows "x \<sqinter> y = y \<sqinter> x"
by (unfold meet_def) (simp add: insert_commute)
lemma (in lattice) meet_assoc:
- assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
+ assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
shows "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)"
proof -
have "(x \<sqinter> y) \<sqinter> z = z \<sqinter> (x \<sqinter> y)" by (simp only: meet_comm)
@@ -693,6 +639,7 @@
finally show ?thesis .
qed
+
subsection {* Total Orders *}
locale total_order = lattice +
@@ -707,47 +654,48 @@
show "lattice_axioms L"
proof (rule lattice_axioms.intro)
fix x y
- assume L: "x \<in> carrier L" "y \<in> carrier L"
+ assume L: "x \<in> carrier L" "y \<in> carrier L"
show "EX s. least L s (Upper L {x, y})"
proof -
note total L
moreover
{
- assume "x \<sqsubseteq> y"
+ assume "x \<sqsubseteq> y"
with L have "least L y (Upper L {x, y})"
- by (rule_tac least_UpperI) auto
+ by (rule_tac least_UpperI) auto
}
moreover
{
- assume "y \<sqsubseteq> x"
+ assume "y \<sqsubseteq> x"
with L have "least L x (Upper L {x, y})"
- by (rule_tac least_UpperI) auto
+ by (rule_tac least_UpperI) auto
}
ultimately show ?thesis by blast
qed
next
fix x y
- assume L: "x \<in> carrier L" "y \<in> carrier L"
+ assume L: "x \<in> carrier L" "y \<in> carrier L"
show "EX i. greatest L i (Lower L {x, y})"
proof -
note total L
moreover
{
- assume "y \<sqsubseteq> x"
+ assume "y \<sqsubseteq> x"
with L have "greatest L y (Lower L {x, y})"
- by (rule_tac greatest_LowerI) auto
+ by (rule_tac greatest_LowerI) auto
}
moreover
{
- assume "x \<sqsubseteq> y"
+ assume "x \<sqsubseteq> y"
with L have "greatest L x (Lower L {x, y})"
- by (rule_tac greatest_LowerI) auto
+ by (rule_tac greatest_LowerI) auto
}
ultimately show ?thesis by blast
qed
qed
qed (assumption | rule total_order_axioms.intro)+
+
subsection {* Complete lattices *}
locale complete_lattice = lattice +
@@ -766,7 +714,7 @@
shows "complete_lattice L"
proof (rule complete_lattice.intro)
show "lattice_axioms L"
- by (rule lattice_axioms.intro) (blast intro: sup_exists inf_exists)+
+ by (rule lattice_axioms.intro) (blast intro: sup_exists inf_exists)+
qed (assumption | rule complete_lattice_axioms.intro)+
constdefs (structure L)
@@ -789,7 +737,7 @@
qed
lemma (in complete_lattice) sup_closed [simp]:
- "A \<subseteq> carrier L ==> \<Squnion> A \<in> carrier L"
+ "A \<subseteq> carrier L ==> \<Squnion>A \<in> carrier L"
by (rule supI) simp_all
lemma (in complete_lattice) top_closed [simp, intro]:
@@ -798,7 +746,7 @@
lemma (in complete_lattice) infI:
"[| !!i. greatest L i (Lower L A) ==> P i; A \<subseteq> carrier L |]
- ==> P (\<Sqinter> A)"
+ ==> P (\<Sqinter>A)"
proof (unfold inf_def)
assume L: "A \<subseteq> carrier L"
and P: "!!l. greatest L l (Lower L A) ==> P l"
@@ -808,7 +756,7 @@
qed
lemma (in complete_lattice) inf_closed [simp]:
- "A \<subseteq> carrier L ==> \<Sqinter> A \<in> carrier L"
+ "A \<subseteq> carrier L ==> \<Sqinter>A \<in> carrier L"
by (rule infI) simp_all
lemma (in complete_lattice) bottom_closed [simp, intro]:
@@ -842,7 +790,7 @@
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 L ?B"])
+ apply (rule greatest_le [where A = "Lower L ?B"])
apply (rule b_inf_B)
apply (rule Lower_memI)
apply (erule UpperD)
@@ -942,7 +890,7 @@
with subgr [THEN subgroup.m_inv_closed]
show "inv x \<in> \<Inter>A" by blast
next
- fix x y assume "x \<in> \<Inter>A" "y \<in> \<Inter>A"
+ fix x y assume "x \<in> \<Inter>A" "y \<in> \<Inter>A"
with subgr [THEN subgroup.m_closed]
show "x \<otimes> y \<in> \<Inter>A" by blast
qed
@@ -986,4 +934,4 @@
then show "EX I. greatest ?L I (Lower ?L A)" ..
qed
-end
\ No newline at end of file
+end