mathlib documentation

order.modular_lattice

Modular Lattices #

This file defines (semi)modular lattices, a kind of lattice useful in algebra. For examples, look to the subobject lattices of abelian groups, submodules, and ideals, or consider any distributive lattice.

Typeclasses #

We define (semi)modularity typeclasses as Prop-valued mixins.

Main Definitions #

Main Results #

References #

TODO #

@[class]
structure is_weak_upper_modular_lattice (α : Type u_2) [lattice α] :
Prop
  • covby_sup_of_inf_covby_covby : ∀ {a b : α}, a b aa b ba a b

A weakly upper modular lattice is a lattice where a ⊔ b covers a and b if a and b both cover a ⊓ b.

Instances of this typeclass
@[class]
structure is_weak_lower_modular_lattice (α : Type u_2) [lattice α] :
Prop
  • inf_covby_of_covby_covby_sup : ∀ {a b : α}, a a bb a ba b a

A weakly lower modular lattice is a lattice where a and b cover a ⊓ b if a ⊔ b covers both a and b.

Instances of this typeclass
@[class]
structure is_upper_modular_lattice (α : Type u_2) [lattice α] :
Prop
  • covby_sup_of_inf_covby : ∀ {a b : α}, a b ab a b

An upper modular lattice, aka semimodular lattice, is a lattice where a ⊔ b covers a and b if either a or b covers a ⊓ b.

Instances of this typeclass
@[class]
structure is_lower_modular_lattice (α : Type u_2) [lattice α] :
Prop
  • inf_covby_of_covby_sup : ∀ {a b : α}, a a ba b b

A lower modular lattice is a lattice where a and b both cover a ⊓ b if a ⊔ b covers either a or b.

Instances of this typeclass
@[class]
structure is_modular_lattice (α : Type u_2) [lattice α] :
Prop
  • sup_inf_le_assoc_of_le : ∀ {x : α} (y : α) {z : α}, x z(x y) z x y z

A modular lattice is one with a limited associativity between and .

Instances of this typeclass
theorem covby_sup_of_inf_covby_of_inf_covby_left {α : Type u_1} [lattice α] [is_weak_upper_modular_lattice α] {a b : α} :
a b aa b ba a b
theorem covby_sup_of_inf_covby_of_inf_covby_right {α : Type u_1} [lattice α] [is_weak_upper_modular_lattice α] {a b : α} :
a b aa b bb a b
theorem covby.sup_of_inf_of_inf_left {α : Type u_1} [lattice α] [is_weak_upper_modular_lattice α] {a b : α} :
a b aa b ba a b

Alias of covby_sup_of_inf_covby_of_inf_covby_left.

theorem covby.sup_of_inf_of_inf_right {α : Type u_1} [lattice α] [is_weak_upper_modular_lattice α] {a b : α} :
a b aa b bb a b

Alias of covby_sup_of_inf_covby_of_inf_covby_right.

theorem inf_covby_of_covby_sup_of_covby_sup_left {α : Type u_1} [lattice α] [is_weak_lower_modular_lattice α] {a b : α} :
a a bb a ba b a
theorem inf_covby_of_covby_sup_of_covby_sup_right {α : Type u_1} [lattice α] [is_weak_lower_modular_lattice α] {a b : α} :
a a bb a ba b b
theorem covby.inf_of_sup_of_sup_left {α : Type u_1} [lattice α] [is_weak_lower_modular_lattice α] {a b : α} :
a a bb a ba b a

Alias of inf_covby_of_covby_sup_of_covby_sup_left.

theorem covby.inf_of_sup_of_sup_right {α : Type u_1} [lattice α] [is_weak_lower_modular_lattice α] {a b : α} :
a a bb a ba b b

Alias of inf_covby_of_covby_sup_of_covby_sup_right.

theorem covby_sup_of_inf_covby_left {α : Type u_1} [lattice α] [is_upper_modular_lattice α] {a b : α} :
a b ab a b
theorem covby_sup_of_inf_covby_right {α : Type u_1} [lattice α] [is_upper_modular_lattice α] {a b : α} :
a b ba a b
theorem covby.sup_of_inf_left {α : Type u_1} [lattice α] [is_upper_modular_lattice α] {a b : α} :
a b ab a b

Alias of covby_sup_of_inf_covby_left.

theorem covby.sup_of_inf_right {α : Type u_1} [lattice α] [is_upper_modular_lattice α] {a b : α} :
a b ba a b

Alias of covby_sup_of_inf_covby_right.

theorem inf_covby_of_covby_sup_left {α : Type u_1} [lattice α] [is_lower_modular_lattice α] {a b : α} :
a a ba b b
theorem inf_covby_of_covby_sup_right {α : Type u_1} [lattice α] [is_lower_modular_lattice α] {a b : α} :
b a ba b a
theorem covby.inf_of_sup_left {α : Type u_1} [lattice α] [is_lower_modular_lattice α] {a b : α} :
a a ba b b

Alias of inf_covby_of_covby_sup_left.

theorem covby.inf_of_sup_right {α : Type u_1} [lattice α] [is_lower_modular_lattice α] {a b : α} :
b a ba b a

Alias of inf_covby_of_covby_sup_right.

theorem sup_inf_assoc_of_le {α : Type u_1} [lattice α] [is_modular_lattice α] {x : α} (y : α) {z : α} (h : x z) :
(x y) z = x y z
theorem is_modular_lattice.inf_sup_inf_assoc {α : Type u_1} [lattice α] [is_modular_lattice α] {x y z : α} :
x z y z = (x z y) z
theorem inf_sup_assoc_of_le {α : Type u_1} [lattice α] [is_modular_lattice α] {x : α} (y : α) {z : α} (h : z x) :
x y z = x (y z)
@[protected, instance]
theorem is_modular_lattice.sup_inf_sup_assoc {α : Type u_1} [lattice α] [is_modular_lattice α] {x y z : α} :
(x z) (y z) = (x z) y z
theorem eq_of_le_of_inf_le_of_sup_le {α : Type u_1} [lattice α] [is_modular_lattice α] {x y z : α} (hxy : x y) (hinf : y z x z) (hsup : y z x z) :
x = y
theorem sup_lt_sup_of_lt_of_inf_le_inf {α : Type u_1} [lattice α] [is_modular_lattice α] {x y z : α} (hxy : x < y) (hinf : y z x z) :
x z < y z
theorem inf_lt_inf_of_lt_of_sup_le_sup {α : Type u_1} [lattice α] [is_modular_lattice α] {x y z : α} (hxy : x < y) (hinf : y z x z) :
x z < y z
theorem well_founded_lt_exact_sequence {α : Type u_1} [lattice α] [is_modular_lattice α] {β : Type u_2} {γ : Type u_3} [partial_order β] [preorder γ] (h₁ : well_founded has_lt.lt) (h₂ : well_founded has_lt.lt) (K : α) (f₁ : β → α) (f₂ : α → β) (g₁ : γ → α) (g₂ : α → γ) (gci : galois_coinsertion f₁ f₂) (gi : galois_insertion g₂ g₁) (hf : ∀ (a : α), f₁ (f₂ a) = a K) (hg : ∀ (a : α), g₁ (g₂ a) = a K) :

A generalization of the theorem that if N is a submodule of M and N and M / N are both Artinian, then M is Artinian.

theorem well_founded_gt_exact_sequence {α : Type u_1} [lattice α] [is_modular_lattice α] {β : Type u_2} {γ : Type u_3} [preorder β] [partial_order γ] (h₁ : well_founded gt) (h₂ : well_founded gt) (K : α) (f₁ : β → α) (f₂ : α → β) (g₁ : γ → α) (g₂ : α → γ) (gci : galois_coinsertion f₁ f₂) (gi : galois_insertion g₂ g₁) (hf : ∀ (a : α), f₁ (f₂ a) = a K) (hg : ∀ (a : α), g₁ (g₂ a) = a K) :

A generalization of the theorem that if N is a submodule of M and N and M / N are both Noetherian, then M is Noetherian.

def inf_Icc_order_iso_Icc_sup {α : Type u_1} [lattice α] [is_modular_lattice α] (a b : α) :
(set.Icc (a b) a) ≃o (set.Icc b (a b))

The diamond isomorphism between the intervals [a ⊓ b, a] and [b, a ⊔ b]

Equations
def is_compl.Iic_order_iso_Ici {α : Type u_1} [lattice α] [bounded_order α] [is_modular_lattice α] {a b : α} (h : is_compl a b) :

The diamond isomorphism between the intervals set.Iic a and set.Ici b.

Equations
theorem is_modular_lattice_iff_inf_sup_inf_assoc {α : Type u_1} [lattice α] :
is_modular_lattice α ∀ (x y z : α), x z y z = (x z y) z
@[protected, instance]
theorem disjoint.disjoint_sup_right_of_disjoint_sup_left {α : Type u_1} [lattice α] [order_bot α] [is_modular_lattice α] {a b c : α} (h : disjoint a b) (hsup : disjoint (a b) c) :
disjoint a (b c)
theorem disjoint.disjoint_sup_left_of_disjoint_sup_right {α : Type u_1} [lattice α] [order_bot α] [is_modular_lattice α] {a b c : α} (h : disjoint b c) (hsup : disjoint a (b c)) :
disjoint (a b) c
@[protected, instance]
@[protected, instance]