Documentation

Mathlib.Order.ModularLattice

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 #

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

  • covBy_sup_of_inf_covBy_covBy {a b : α} : ab aab ba ab

    a ⊔ b covers a and b if a and b both cover a ⊓ b.

Instances

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

    • inf_covBy_of_covBy_covBy_sup {a b : α} : a abb abab a

      a and b cover a ⊓ b if a ⊔ b covers both a and b

    Instances
      class IsUpperModularLattice (α : Type u_2) [Lattice α] :

      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.

      • covBy_sup_of_inf_covBy {a b : α} : ab ab ab

        a ⊔ b covers a and b if either a or b covers a ⊓ b

      Instances
        class IsLowerModularLattice (α : Type u_2) [Lattice α] :

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

        • inf_covBy_of_covBy_sup {a b : α} : a abab b

          a and b both cover a ⊓ b if a ⊔ b covers either a or b

        Instances
          class IsModularLattice (α : Type u_2) [Lattice α] :

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

          • sup_inf_le_assoc_of_le {x : α} (y : α) {z : α} : x z(xy)z xyz

            Whenever x ≤ z, then for any y, (x ⊔ y) ⊓ z ≤ x ⊔ (y ⊓ z)

          Instances
            theorem covBy_sup_of_inf_covBy_of_inf_covBy_left {α : Type u_1} [Lattice α] [IsWeakUpperModularLattice α] {a b : α} :
            ab aab ba ab
            theorem covBy_sup_of_inf_covBy_of_inf_covBy_right {α : Type u_1} [Lattice α] [IsWeakUpperModularLattice α] {a b : α} :
            ab aab bb ab
            theorem CovBy.sup_of_inf_of_inf_left {α : Type u_1} [Lattice α] [IsWeakUpperModularLattice α] {a b : α} :
            ab aab ba ab

            Alias of covBy_sup_of_inf_covBy_of_inf_covBy_left.

            theorem CovBy.sup_of_inf_of_inf_right {α : Type u_1} [Lattice α] [IsWeakUpperModularLattice α] {a b : α} :
            ab aab bb ab

            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 α] [IsWeakLowerModularLattice α] {a b : α} :
            a abb abab a
            theorem inf_covBy_of_covBy_sup_of_covBy_sup_right {α : Type u_1} [Lattice α] [IsWeakLowerModularLattice α] {a b : α} :
            a abb abab b
            theorem CovBy.inf_of_sup_of_sup_left {α : Type u_1} [Lattice α] [IsWeakLowerModularLattice α] {a b : α} :
            a abb abab a

            Alias of inf_covBy_of_covBy_sup_of_covBy_sup_left.

            theorem CovBy.inf_of_sup_of_sup_right {α : Type u_1} [Lattice α] [IsWeakLowerModularLattice α] {a b : α} :
            a abb abab b

            Alias of inf_covBy_of_covBy_sup_of_covBy_sup_right.

            theorem covBy_sup_of_inf_covBy_left {α : Type u_1} [Lattice α] [IsUpperModularLattice α] {a b : α} :
            ab ab ab
            theorem covBy_sup_of_inf_covBy_right {α : Type u_1} [Lattice α] [IsUpperModularLattice α] {a b : α} :
            ab ba ab
            theorem CovBy.sup_of_inf_left {α : Type u_1} [Lattice α] [IsUpperModularLattice α] {a b : α} :
            ab ab ab

            Alias of covBy_sup_of_inf_covBy_left.

            theorem CovBy.sup_of_inf_right {α : Type u_1} [Lattice α] [IsUpperModularLattice α] {a b : α} :
            ab ba ab

            Alias of covBy_sup_of_inf_covBy_right.

            theorem inf_covBy_of_covBy_sup_left {α : Type u_1} [Lattice α] [IsLowerModularLattice α] {a b : α} :
            a abab b
            theorem inf_covBy_of_covBy_sup_right {α : Type u_1} [Lattice α] [IsLowerModularLattice α] {a b : α} :
            b abab a
            theorem CovBy.inf_of_sup_left {α : Type u_1} [Lattice α] [IsLowerModularLattice α] {a b : α} :
            a abab b

            Alias of inf_covBy_of_covBy_sup_left.

            theorem CovBy.inf_of_sup_right {α : Type u_1} [Lattice α] [IsLowerModularLattice α] {a b : α} :
            b abab a

            Alias of inf_covBy_of_covBy_sup_right.

            theorem sup_inf_assoc_of_le {α : Type u_1} [Lattice α] [IsModularLattice α] {x : α} (y : α) {z : α} (h : x z) :
            (xy)z = xyz
            theorem IsModularLattice.inf_sup_inf_assoc {α : Type u_1} [Lattice α] [IsModularLattice α] {x y z : α} :
            xzyz = (xzy)z
            theorem inf_sup_assoc_of_le {α : Type u_1} [Lattice α] [IsModularLattice α] {x : α} (y : α) {z : α} (h : z x) :
            xyz = x(yz)
            theorem IsModularLattice.sup_inf_sup_assoc {α : Type u_1} [Lattice α] [IsModularLattice α] {x y z : α} :
            (xz)(yz) = (xz)yz
            theorem eq_of_le_of_inf_le_of_le_sup {α : Type u_1} [Lattice α] [IsModularLattice α] {x y z : α} (hxy : x y) (hinf : yz x) (hsup : y xz) :
            x = y
            theorem eq_of_le_of_inf_le_of_sup_le {α : Type u_1} [Lattice α] [IsModularLattice α] {x y z : α} (hxy : x y) (hinf : yz xz) (hsup : yz xz) :
            x = y
            theorem sup_lt_sup_of_lt_of_inf_le_inf {α : Type u_1} [Lattice α] [IsModularLattice α] {x y z : α} (hxy : x < y) (hinf : yz xz) :
            xz < yz
            theorem inf_lt_inf_of_lt_of_sup_le_sup {α : Type u_1} [Lattice α] [IsModularLattice α] {x y z : α} (hxy : x < y) (hinf : yz xz) :
            xz < yz
            theorem strictMono_inf_prod_sup {α : Type u_1} [Lattice α] [IsModularLattice α] {z : α} :
            StrictMono fun (x : α) => (xz, xz)
            theorem wellFounded_lt_exact_sequence {α : Type u_1} [Lattice α] [IsModularLattice α] {β : Type u_2} {γ : Type u_3} [Preorder β] [Preorder γ] [h₁ : WellFoundedLT β] [h₂ : WellFoundedLT γ] (K : α) (f₁ : βα) (f₂ : αβ) (g₁ : γα) (g₂ : αγ) (gci : GaloisCoinsertion f₁ f₂) (gi : GaloisInsertion g₂ g₁) (hf : ∀ (a : α), f₁ (f₂ a) = aK) (hg : ∀ (a : α), g₁ (g₂ a) = aK) :

            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 wellFounded_gt_exact_sequence {α : Type u_1} [Lattice α] [IsModularLattice α] {β : Type u_2} {γ : Type u_3} [Preorder β] [Preorder γ] [WellFoundedGT β] [WellFoundedGT γ] (K : α) (f₁ : βα) (f₂ : αβ) (g₁ : γα) (g₂ : αγ) (gci : GaloisCoinsertion f₁ f₂) (gi : GaloisInsertion g₂ g₁) (hf : ∀ (a : α), f₁ (f₂ a) = aK) (hg : ∀ (a : α), g₁ (g₂ a) = aK) :

            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 infIccOrderIsoIccSup {α : Type u_1} [Lattice α] [IsModularLattice α] (a b : α) :
            (Set.Icc (ab) a) ≃o (Set.Icc b (ab))

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

            Equations
            • infIccOrderIsoIccSup a b = { toFun := fun (x : (Set.Icc (ab) a)) => xb, , invFun := fun (x : (Set.Icc b (ab))) => ax, , left_inv := , right_inv := , map_rel_iff' := }
            @[simp]
            theorem infIccOrderIsoIccSup_symm_apply_coe {α : Type u_1} [Lattice α] [IsModularLattice α] (a b : α) (x : (Set.Icc b (ab))) :
            ((RelIso.symm (infIccOrderIsoIccSup a b)) x) = ax
            @[simp]
            theorem infIccOrderIsoIccSup_apply_coe {α : Type u_1} [Lattice α] [IsModularLattice α] (a b : α) (x : (Set.Icc (ab) a)) :
            ((infIccOrderIsoIccSup a b) x) = xb
            theorem inf_strictMonoOn_Icc_sup {α : Type u_1} [Lattice α] [IsModularLattice α] {a b : α} :
            StrictMonoOn (fun (c : α) => ac) (Set.Icc b (ab))
            theorem sup_strictMonoOn_Icc_inf {α : Type u_1} [Lattice α] [IsModularLattice α] {a b : α} :
            StrictMonoOn (fun (c : α) => cb) (Set.Icc (ab) a)
            def infIooOrderIsoIooSup {α : Type u_1} [Lattice α] [IsModularLattice α] (a b : α) :
            (Set.Ioo (ab) a) ≃o (Set.Ioo b (ab))

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

            Equations
            • infIooOrderIsoIooSup a b = { toFun := fun (c : (Set.Ioo (ab) a)) => cb, , invFun := fun (c : (Set.Ioo b (ab))) => ac, , left_inv := , right_inv := , map_rel_iff' := }
            @[simp]
            theorem infIooOrderIsoIooSup_apply_coe {α : Type u_1} [Lattice α] [IsModularLattice α] (a b : α) (c : (Set.Ioo (ab) a)) :
            ((infIooOrderIsoIooSup a b) c) = cb
            @[simp]
            theorem infIooOrderIsoIooSup_symm_apply_coe {α : Type u_1} [Lattice α] [IsModularLattice α] (a b : α) (c : (Set.Ioo b (ab))) :
            ((RelIso.symm (infIooOrderIsoIooSup a b)) c) = ac
            def IsCompl.IicOrderIsoIci {α : Type u_1} [Lattice α] [BoundedOrder α] [IsModularLattice α] {a b : α} (h : IsCompl a b) :
            (Set.Iic a) ≃o (Set.Ici b)

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

            Equations
            theorem isModularLattice_iff_inf_sup_inf_assoc {α : Type u_1} [Lattice α] :
            IsModularLattice α ∀ (x y z : α), xzyz = (xzy)z
            theorem Disjoint.disjoint_sup_right_of_disjoint_sup_left {α : Type u_1} {a b c : α} [Lattice α] [OrderBot α] [IsModularLattice α] (h : Disjoint a b) (hsup : Disjoint (ab) c) :
            Disjoint a (bc)
            theorem Disjoint.disjoint_sup_left_of_disjoint_sup_right {α : Type u_1} {a b c : α} [Lattice α] [OrderBot α] [IsModularLattice α] (h : Disjoint b c) (hsup : Disjoint a (bc)) :
            Disjoint (ab) c
            theorem Disjoint.isCompl_sup_right_of_isCompl_sup_left {α : Type u_1} {a b c : α} [Lattice α] [BoundedOrder α] [IsModularLattice α] (h : Disjoint a b) (hcomp : IsCompl (ab) c) :
            IsCompl a (bc)
            theorem Disjoint.isCompl_sup_left_of_isCompl_sup_right {α : Type u_1} {a b c : α} [Lattice α] [BoundedOrder α] [IsModularLattice α] (h : Disjoint b c) (hcomp : IsCompl a (bc)) :
            IsCompl (ab) c
            theorem Set.Iic.isCompl_inf_inf_of_isCompl_of_le {α : Type u_1} [Lattice α] [BoundedOrder α] [IsModularLattice α] {a b c : α} (h₁ : IsCompl b c) (h₂ : b a) :
            IsCompl ab, ac,