Documentation

Mathlib.Order.Closure

Closure operators between preorders #

We define (bundled) closure operators on a preorder as monotone (increasing), extensive (inflationary) and idempotent functions. We define closed elements for the operator as elements which are fixed by it.

Lower adjoints to a function between preorders u : β → α allow to generalise closure operators to situations where the closure operator we are dealing with naturally decomposes as u ∘ l where l is a worthy function to have on its own. Typical examples include l : Set G → Subgroup G := Subgroup.closure, u : Subgroup G → Set G := (↑), where G is a group. This shows there is a close connection between closure operators, lower adjoints and Galois connections/insertions: every Galois connection induces a lower adjoint which itself induces a closure operator by composition (see GaloisConnection.lowerAdjoint and LowerAdjoint.closureOperator), and every closure operator on a partial order induces a Galois insertion from the set of closed elements to the underlying type (see ClosureOperator.gi).

Main definitions #

Implementation details #

Although LowerAdjoint is technically a generalisation of ClosureOperator (by defining toFun := id), it is desirable to have both as otherwise ids would be carried all over the place when using concrete closure operators such as ConvexHull.

LowerAdjoint really is a semibundled structure version of GaloisConnection.

References #

Closure operator #

structure ClosureOperator (α : Type u_1) [Preorder α] extends OrderHom :
Type u_1

A closure operator on the preorder α is a monotone function which is extensive (every x is less than its closure) and idempotent.

Instances For
    Equations
    @[simp]
    theorem ClosureOperator.id_apply (α : Type u_1) [PartialOrder α] (a : α) :
    ↑(ClosureOperator.id α) a = a

    The identity function as a closure operator.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      theorem ClosureOperator.ext {α : Type u_1} [PartialOrder α] (c₁ : ClosureOperator α) (c₂ : ClosureOperator α) :
      c₁ = c₂c₁ = c₂
      @[simp]
      theorem ClosureOperator.mk'_apply {α : Type u_1} [PartialOrder α] (f : αα) (hf₁ : Monotone f) (hf₂ : ∀ (x : α), x f x) (hf₃ : ∀ (x : α), f (f x) f x) :
      ∀ (a : α), ↑(ClosureOperator.mk' f hf₁ hf₂ hf₃) a = f a
      def ClosureOperator.mk' {α : Type u_1} [PartialOrder α] (f : αα) (hf₁ : Monotone f) (hf₂ : ∀ (x : α), x f x) (hf₃ : ∀ (x : α), f (f x) f x) :

      Constructor for a closure operator using the weaker idempotency axiom: f (f x) ≤ f x.

      Equations
      • ClosureOperator.mk' f hf₁ hf₂ hf₃ = { toOrderHom := { toFun := f, monotone' := hf₁ }, le_closure' := hf₂, idempotent' := (_ : ∀ (x : α), f (f x) = f x) }
      Instances For
        @[simp]
        theorem ClosureOperator.mk₂_apply {α : Type u_1} [PartialOrder α] (f : αα) (hf : ∀ (x : α), x f x) (hmin : ∀ ⦃x y : α⦄, x f yf x f y) :
        ∀ (a : α), ↑(ClosureOperator.mk₂ f hf hmin) a = f a
        def ClosureOperator.mk₂ {α : Type u_1} [PartialOrder α] (f : αα) (hf : ∀ (x : α), x f x) (hmin : ∀ ⦃x y : α⦄, x f yf x f y) :

        Convenience constructor for a closure operator using the weaker minimality axiom: x ≤ f y → f x ≤ f y, which is sometimes easier to prove in practice.

        Equations
        • ClosureOperator.mk₂ f hf hmin = { toOrderHom := { toFun := f, monotone' := (_ : ∀ (x y : α), x yf x f y) }, le_closure' := hf, idempotent' := (_ : ∀ (x : α), f (f x) = f x) }
        Instances For
          @[simp]
          theorem ClosureOperator.mk₃_apply {α : Type u_1} [PartialOrder α] (f : αα) (p : αProp) (hf : ∀ (x : α), x f x) (hfp : (x : α) → p (f x)) (hmin : ∀ ⦃x y : α⦄, x yp yf x y) :
          ∀ (a : α), ↑(ClosureOperator.mk₃ f p hf hfp hmin) a = f a
          def ClosureOperator.mk₃ {α : Type u_1} [PartialOrder α] (f : αα) (p : αProp) (hf : ∀ (x : α), x f x) (hfp : (x : α) → p (f x)) (hmin : ∀ ⦃x y : α⦄, x yp yf x y) :

          Expanded out version of mk₂. p implies being closed. This constructor should be used when you already know a sufficient condition for being closed and using mem_mk₃_closed will avoid you the (slight) hassle of having to prove it both inside and outside the constructor.

          Equations
          Instances For
            theorem ClosureOperator.closure_mem_mk₃ {α : Type u_1} [PartialOrder α] {f : αα} {p : αProp} {hf : ∀ (x : α), x f x} {hfp : (x : α) → p (f x)} {hmin : ∀ ⦃x y : α⦄, x yp yf x y} (x : α) :
            p (↑(ClosureOperator.mk₃ f p hf hfp hmin) x)

            This lemma shows that the image of x of a closure operator built from the mk₃ constructor respects p, the property that was fed into it.

            theorem ClosureOperator.closure_le_mk₃_iff {α : Type u_1} [PartialOrder α] {f : αα} {p : αProp} {hf : ∀ (x : α), x f x} {hfp : (x : α) → p (f x)} {hmin : ∀ ⦃x y : α⦄, x yp yf x y} {x : α} {y : α} (hxy : x y) (hy : p y) :
            ↑(ClosureOperator.mk₃ f p hf hfp hmin) x y

            Analogue of closure_le_closed_iff_le but with the p that was fed into the mk₃ constructor.

            theorem ClosureOperator.le_closure {α : Type u_1} [PartialOrder α] (c : ClosureOperator α) (x : α) :
            x c x

            Every element is less than its closure. This property is sometimes referred to as extensivity or inflationarity.

            @[simp]
            theorem ClosureOperator.idempotent {α : Type u_1} [PartialOrder α] (c : ClosureOperator α) (x : α) :
            c (c x) = c x
            theorem ClosureOperator.le_closure_iff {α : Type u_1} [PartialOrder α] (c : ClosureOperator α) (x : α) (y : α) :
            x c y c x c y

            An element x is closed for the closure operator c if it is a fixed point for it.

            Equations
            Instances For
              @[simp]

              The set of closed elements for c is exactly its range.

              Send an x to an element of the set of closed elements (by taking the closure).

              Equations
              Instances For
                @[simp]
                theorem ClosureOperator.closure_le_closed_iff_le {α : Type u_1} [PartialOrder α] (c : ClosureOperator α) (x : α) {y : α} (hy : ClosureOperator.closed c y) :
                c x y x y
                theorem ClosureOperator.eq_mk₃_closed {α : Type u_1} [PartialOrder α] (c : ClosureOperator α) :
                c = ClosureOperator.mk₃ (c) (ClosureOperator.closed c) (_ : ∀ (x : α), x c x) (_ : ∀ (x : α), c x ClosureOperator.closed c) (_ : ∀ (x y : α), x yClosureOperator.closed c yc x y)

                A closure operator is equal to the closure operator obtained by feeding c.closed into the mk₃ constructor.

                @[simp]
                theorem ClosureOperator.mem_mk₃_closed {α : Type u_1} [PartialOrder α] {f : αα} {p : αProp} {hf : ∀ (x : α), x f x} {hfp : (x : α) → p (f x)} {hmin : ∀ ⦃x y : α⦄, x yp yf x y} {x : α} :

                The property p fed into the mk₃ constructor exactly corresponds to being closed.

                @[simp]
                theorem ClosureOperator.closure_top {α : Type u_1} [PartialOrder α] [OrderTop α] (c : ClosureOperator α) :
                c =
                theorem ClosureOperator.closure_inf_le {α : Type u_1} [SemilatticeInf α] (c : ClosureOperator α) (x : α) (y : α) :
                c (x y) c x c y
                theorem ClosureOperator.closure_sup_closure_le {α : Type u_1} [SemilatticeSup α] (c : ClosureOperator α) (x : α) (y : α) :
                c x c y c (x y)
                theorem ClosureOperator.closure_sup_closure_left {α : Type u_1} [SemilatticeSup α] (c : ClosureOperator α) (x : α) (y : α) :
                c (c x y) = c (x y)
                theorem ClosureOperator.closure_sup_closure_right {α : Type u_1} [SemilatticeSup α] (c : ClosureOperator α) (x : α) (y : α) :
                c (x c y) = c (x y)
                theorem ClosureOperator.closure_sup_closure {α : Type u_1} [SemilatticeSup α] (c : ClosureOperator α) (x : α) (y : α) :
                c (c x c y) = c (x y)
                @[simp]
                theorem ClosureOperator.closure_iSup_closure {α : Type u_1} {ι : Sort u_2} [CompleteLattice α] (c : ClosureOperator α) (f : ια) :
                c (⨆ (i : ι), c (f i)) = c (⨆ (i : ι), f i)
                @[simp]
                theorem ClosureOperator.closure_iSup₂_closure {α : Type u_1} {ι : Sort u_2} {κ : ιSort u_3} [CompleteLattice α] (c : ClosureOperator α) (f : (i : ι) → κ iα) :
                c (⨆ (i : ι) (j : κ i), c (f i j)) = c (⨆ (i : ι) (j : κ i), f i j)

                Lower adjoint #

                structure LowerAdjoint {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] (u : βα) :
                Type (max u_1 u_4)
                • toFun : αβ

                  The underlying function

                • gc' : GaloisConnection s.toFun u

                  The underlying function is a lower adjoint.

                A lower adjoint of u on the preorder α is a function l such that l and u form a Galois connection. It allows us to define closure operators whose output does not match the input. In practice, u is often (↑) : β → α.

                Instances For
                  @[simp]
                  theorem LowerAdjoint.id_toFun (α : Type u_1) [Preorder α] (x : α) :
                  def LowerAdjoint.id (α : Type u_1) [Preorder α] :

                  The identity function as a lower adjoint to itself.

                  Equations
                  Instances For
                    Equations
                    instance LowerAdjoint.instCoeFunLowerAdjointForAll {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} :
                    CoeFun (LowerAdjoint u) fun x => αβ
                    Equations
                    • LowerAdjoint.instCoeFunLowerAdjointForAll = { coe := LowerAdjoint.toFun }
                    theorem LowerAdjoint.gc {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l : LowerAdjoint u) :
                    theorem LowerAdjoint.ext {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l₁ : LowerAdjoint u) (l₂ : LowerAdjoint u) :
                    l₁.toFun = l₂.toFunl₁ = l₂
                    theorem LowerAdjoint.monotone {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l : LowerAdjoint u) :
                    Monotone (u l.toFun)
                    theorem LowerAdjoint.le_closure {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) :

                    Every element is less than its closure. This property is sometimes referred to as extensivity or inflationarity.

                    @[simp]
                    theorem LowerAdjoint.closureOperator_apply {α : Type u_1} {β : Type u_4} [PartialOrder α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) :
                    def LowerAdjoint.closureOperator {α : Type u_1} {β : Type u_4} [PartialOrder α] [Preorder β] {u : βα} (l : LowerAdjoint u) :

                    Every lower adjoint induces a closure operator given by the composition. This is the partial order version of the statement that every adjunction induces a monad.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      theorem LowerAdjoint.idempotent {α : Type u_1} {β : Type u_4} [PartialOrder α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) :
                      theorem LowerAdjoint.le_closure_iff {α : Type u_1} {β : Type u_4} [PartialOrder α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) (y : α) :
                      def LowerAdjoint.closed {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l : LowerAdjoint u) :
                      Set α

                      An element x is closed for l : LowerAdjoint u if it is a fixed point: u (l x) = x

                      Equations
                      Instances For
                        theorem LowerAdjoint.mem_closed_iff {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) :
                        theorem LowerAdjoint.closure_eq_self_of_mem_closed {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {u : βα} (l : LowerAdjoint u) {x : α} (h : x LowerAdjoint.closed l) :
                        theorem LowerAdjoint.mem_closed_iff_closure_le {α : Type u_1} {β : Type u_4} [PartialOrder α] [PartialOrder β] {u : βα} (l : LowerAdjoint u) (x : α) :
                        @[simp]
                        theorem LowerAdjoint.closure_is_closed {α : Type u_1} {β : Type u_4} [PartialOrder α] [PartialOrder β] {u : βα} (l : LowerAdjoint u) (x : α) :
                        theorem LowerAdjoint.closed_eq_range_close {α : Type u_1} {β : Type u_4} [PartialOrder α] [PartialOrder β] {u : βα} (l : LowerAdjoint u) :

                        The set of closed elements for l is the range of u ∘ l.

                        def LowerAdjoint.toClosed {α : Type u_1} {β : Type u_4} [PartialOrder α] [PartialOrder β] {u : βα} (l : LowerAdjoint u) (x : α) :

                        Send an x to an element of the set of closed elements (by taking the closure).

                        Equations
                        Instances For
                          @[simp]
                          theorem LowerAdjoint.closure_le_closed_iff_le {α : Type u_1} {β : Type u_4} [PartialOrder α] [PartialOrder β] {u : βα} (l : LowerAdjoint u) (x : α) {y : α} (hy : LowerAdjoint.closed l y) :
                          theorem LowerAdjoint.closure_top {α : Type u_1} {β : Type u_4} [PartialOrder α] [OrderTop α] [Preorder β] {u : βα} (l : LowerAdjoint u) :
                          theorem LowerAdjoint.closure_inf_le {α : Type u_1} {β : Type u_4} [SemilatticeInf α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) (y : α) :
                          theorem LowerAdjoint.closure_sup_closure_le {α : Type u_1} {β : Type u_4} [SemilatticeSup α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) (y : α) :
                          theorem LowerAdjoint.closure_sup_closure_left {α : Type u_1} {β : Type u_4} [SemilatticeSup α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) (y : α) :
                          theorem LowerAdjoint.closure_sup_closure_right {α : Type u_1} {β : Type u_4} [SemilatticeSup α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) (y : α) :
                          theorem LowerAdjoint.closure_sup_closure {α : Type u_1} {β : Type u_4} [SemilatticeSup α] [Preorder β] {u : βα} (l : LowerAdjoint u) (x : α) (y : α) :
                          theorem LowerAdjoint.closure_iSup_closure {α : Type u_1} {ι : Sort u_2} {β : Type u_4} [CompleteLattice α] [Preorder β] {u : βα} (l : LowerAdjoint u) (f : ια) :
                          u (LowerAdjoint.toFun l (⨆ (i : ι), u (LowerAdjoint.toFun l (f i)))) = u (LowerAdjoint.toFun l (⨆ (i : ι), f i))
                          theorem LowerAdjoint.closure_iSup₂_closure {α : Type u_1} {ι : Sort u_2} {κ : ιSort u_3} {β : Type u_4} [CompleteLattice α] [Preorder β] {u : βα} (l : LowerAdjoint u) (f : (i : ι) → κ iα) :
                          u (LowerAdjoint.toFun l (⨆ (i : ι) (j : κ i), u (LowerAdjoint.toFun l (f i j)))) = u (LowerAdjoint.toFun l (⨆ (i : ι) (j : κ i), f i j))
                          theorem LowerAdjoint.subset_closure {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (s : Set β) :
                          theorem LowerAdjoint.not_mem_of_not_mem_closure {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) {s : Set β} {P : β} (hP : ¬P LowerAdjoint.toFun l s) :
                          ¬P s
                          theorem LowerAdjoint.le_iff_subset {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (s : Set β) (S : α) :
                          theorem LowerAdjoint.mem_iff {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (s : Set β) (x : β) :
                          x LowerAdjoint.toFun l s ∀ (S : α), s Sx S
                          theorem LowerAdjoint.eq_of_le {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) {s : Set β} {S : α} (h₁ : s S) (h₂ : S LowerAdjoint.toFun l s) :
                          theorem LowerAdjoint.closure_union_closure_subset {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (x : α) (y : α) :
                          ↑(LowerAdjoint.toFun l x) ↑(LowerAdjoint.toFun l y) ↑(LowerAdjoint.toFun l (x y))
                          @[simp]
                          theorem LowerAdjoint.closure_union_closure_left {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (x : α) (y : α) :
                          @[simp]
                          theorem LowerAdjoint.closure_union_closure_right {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (x : α) (y : α) :
                          theorem LowerAdjoint.closure_union_closure {α : Type u_1} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (x : α) (y : α) :
                          @[simp]
                          theorem LowerAdjoint.closure_iUnion_closure {α : Type u_1} {ι : Sort u_2} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (f : ια) :
                          LowerAdjoint.toFun l (⋃ (i : ι), ↑(LowerAdjoint.toFun l ↑(f i))) = LowerAdjoint.toFun l (⋃ (i : ι), ↑(f i))
                          @[simp]
                          theorem LowerAdjoint.closure_iUnion₂_closure {α : Type u_1} {ι : Sort u_2} {κ : ιSort u_3} {β : Type u_4} [SetLike α β] (l : LowerAdjoint SetLike.coe) (f : (i : ι) → κ iα) :
                          LowerAdjoint.toFun l (⋃ (i : ι) (j : κ i), ↑(LowerAdjoint.toFun l ↑(f i j))) = LowerAdjoint.toFun l (⋃ (i : ι) (j : κ i), ↑(f i j))

                          Translations between GaloisConnection, LowerAdjoint, ClosureOperator #

                          @[simp]
                          theorem GaloisConnection.lowerAdjoint_toFun {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {l : αβ} {u : βα} (gc : GaloisConnection l u) :
                          def GaloisConnection.lowerAdjoint {α : Type u_1} {β : Type u_4} [Preorder α] [Preorder β] {l : αβ} {u : βα} (gc : GaloisConnection l u) :

                          Every Galois connection induces a lower adjoint.

                          Equations
                          Instances For
                            @[simp]
                            theorem GaloisConnection.closureOperator_apply {α : Type u_1} {β : Type u_4} [PartialOrder α] [Preorder β] {l : αβ} {u : βα} (gc : GaloisConnection l u) (x : α) :
                            def GaloisConnection.closureOperator {α : Type u_1} {β : Type u_4} [PartialOrder α] [Preorder β] {l : αβ} {u : βα} (gc : GaloisConnection l u) :

                            Every Galois connection induces a closure operator given by the composition. This is the partial order version of the statement that every adjunction induces a monad.

                            Equations
                            Instances For

                              The set of closed elements has a Galois insertion to the underlying type.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                @[simp]

                                The Galois insertion associated to a closure operator can be used to reconstruct the closure operator. Note that the inverse in the opposite direction does not hold in general.