Documentation

Mathlib.Order.Defs.PartialOrder

Orders #

Defines classes for preorders and partial orders and proves some basic lemmas about them.

We also define covering relations on a preorder. We say that b covers a if a < b and there is no element in between. We say that b weakly covers a if a ≤ b and there is no element between a and b. In a partial order this is equivalent to a ⋖ b ∨ a = b, in a preorder this is equivalent to a ⋖ b ∨ (a ≤ b ∧ b ≤ a)

Notation #

Definition of Preorder and lemmas about types with a Preorder #

class Preorder (α : Type u_2) extends LE α, LT α :
Type u_2

A preorder is a reflexive, transitive relation with a < b defined in the obvious way.

Instances
    def Preorder.mk' {α : Type u_1} [LE α] [LT α] (le_refl : ∀ (a : α), a a) (ge_trans : ∀ (a b c : α), b ac bc a) (lt_iff_le_not_ge : ∀ (a b : α), b < a b a ¬a b) :

    A variant of Preorder.mk which allows to_dual to dualize a Preorder instance.

    Equations
    • Preorder.mk' le_refl ge_trans lt_iff_le_not_ge = { toLE := inst✝¹, toLT := inst✝, le_refl := le_refl, le_trans := , lt_iff_le_not_ge := }
    Instances For
      @[simp]
      theorem le_refl {α : Type u_1} [Preorder α] (a : α) :
      a a

      The relation on a preorder is reflexive.

      theorem le_rfl {α : Type u_1} [Preorder α] {a : α} :
      a a

      A version of le_refl where the argument is implicit

      theorem le_trans {α : Type u_1} [Preorder α] {a b c : α} :
      a bb ca c

      The relation on a preorder is transitive.

      theorem ge_trans {α : Type u_1} [Preorder α] {a b c : α} :
      b ac bc a
      theorem lt_iff_le_not_ge {α : Type u_1} [Preorder α] {a b : α} :
      a < b a b ¬b a
      theorem lt_of_le_not_ge {α : Type u_1} [Preorder α] {a b : α} (hab : a b) (hba : ¬b a) :
      a < b
      theorem le_of_eq {α : Type u_1} [Preorder α] {a b : α} (hab : a = b) :
      a b
      theorem ge_of_eq {α : Type u_1} [Preorder α] {a b : α} (hab : a = b) :
      b a
      theorem le_of_lt {α : Type u_1} [Preorder α] {a b : α} (hab : a < b) :
      a b
      theorem not_le_of_gt {α : Type u_1} [Preorder α] {a b : α} (hab : a < b) :
      ¬b a
      theorem not_lt_of_ge {α : Type u_1} [Preorder α] {a b : α} (hab : a b) :
      ¬b < a
      theorem LT.lt.not_ge {α : Type u_1} [Preorder α] {a b : α} (hab : a < b) :
      ¬b a

      Alias of not_le_of_gt.

      theorem LE.le.not_gt {α : Type u_1} [Preorder α] {a b : α} (hab : a b) :
      ¬b < a

      Alias of not_lt_of_ge.

      @[deprecated LT.lt.not_ge (since := "2025-06-07")]
      theorem LT.lt.not_le {α : Type u_1} [Preorder α] {a b : α} (hab : a < b) :
      ¬b a

      Alias of not_le_of_gt.


      Alias of not_le_of_gt.

      @[deprecated LE.le.not_gt (since := "2025-06-07")]
      theorem LE.le.not_lt {α : Type u_1} [Preorder α] {a b : α} (hab : a b) :
      ¬b < a

      Alias of not_lt_of_ge.


      Alias of not_lt_of_ge.

      theorem lt_irrefl {α : Type u_1} [Preorder α] (a : α) :
      ¬a < a
      @[deprecated lt_irrefl (since := "2025-06-07")]
      theorem gt_irrefl {α : Type u_1} [Preorder α] (a : α) :
      ¬a < a

      Alias of lt_irrefl.

      theorem lt_of_lt_of_le {α : Type u_1} [Preorder α] {a b c : α} (hab : a < b) (hbc : b c) :
      a < c
      theorem lt_of_lt_of_le' {α : Type u_1} [Preorder α] {a b c : α} (hab : b < a) (hbc : c b) :
      c < a
      theorem lt_of_le_of_lt {α : Type u_1} [Preorder α] {a b c : α} (hab : a b) (hbc : b < c) :
      a < c
      theorem lt_of_le_of_lt' {α : Type u_1} [Preorder α] {a b c : α} (hab : b a) (hbc : c < b) :
      c < a
      @[deprecated lt_of_lt_of_le' (since := "2025-06-07")]
      theorem gt_of_gt_of_ge {α : Type u_1} [Preorder α] {a b c : α} (hab : b < a) (hbc : c b) :
      c < a

      Alias of lt_of_lt_of_le'.

      @[deprecated lt_of_le_of_lt' (since := "2025-06-07")]
      theorem gt_of_ge_of_gt {α : Type u_1} [Preorder α] {a b c : α} (hab : b a) (hbc : c < b) :
      c < a

      Alias of lt_of_le_of_lt'.

      theorem lt_trans {α : Type u_1} [Preorder α] {a b c : α} :
      a < bb < ca < c
      theorem gt_trans {α : Type u_1} [Preorder α] {a b c : α} :
      b < ac < bc < a
      theorem ne_of_lt {α : Type u_1} [Preorder α] {a b : α} (h : a < b) :
      a b
      theorem ne_of_gt {α : Type u_1} [Preorder α] {a b : α} (h : b < a) :
      a b
      theorem lt_asymm {α : Type u_1} [Preorder α] {a b : α} (h : a < b) :
      ¬b < a
      theorem not_lt_of_gt {α : Type u_1} [Preorder α] {a b : α} (h : a < b) :
      ¬b < a

      Alias of lt_asymm.

      theorem le_of_lt_or_eq {α : Type u_1} [Preorder α] {a b : α} (h : a < b a = b) :
      a b
      theorem le_of_lt_or_eq' {α : Type u_1} [Preorder α] {a b : α} (h : b < a a = b) :
      b a
      theorem le_of_eq_or_lt {α : Type u_1} [Preorder α] {a b : α} (h : a = b a < b) :
      a b
      theorem le_of_eq_or_lt' {α : Type u_1} [Preorder α] {a b : α} (h : a = b b < a) :
      b a
      instance instTransLE {α : Type u_1} [Preorder α] :
      Equations
      instance instTransLT {α : Type u_1} [Preorder α] :
      Equations
      instance instTransLTLE {α : Type u_1} [Preorder α] :
      Equations
      instance instTransLELT {α : Type u_1} [Preorder α] :
      Equations
      instance instTransGE {α : Type u_1} [Preorder α] :
      Equations
      instance instTransGT {α : Type u_1} [Preorder α] :
      Equations
      instance instTransGTGE {α : Type u_1} [Preorder α] :
      Equations
      instance instTransGEGT {α : Type u_1} [Preorder α] :
      Equations

      < is decidable if is.

      Equations
      Instances For

        < is decidable if is.

        Equations
        Instances For
          @[deprecated decidableLT'OfDecidableLE' (since := "2025-12-09")]

          Alias of decidableLT'OfDecidableLE'.


          < is decidable if is.

          Equations
          Instances For
            def WCovBy {α : Type u_1} [Preorder α] (a b : α) :

            WCovBy a b means that a = b or b covers a. This means that a ≤ b and there is no element in between. This is denoted a ⩿ b.

            Equations
            Instances For

              WCovBy a b means that a = b or b covers a. This means that a ≤ b and there is no element in between. This is denoted a ⩿ b.

              Equations
              Instances For
                def CovBy {α : Type u_2} [LT α] (a b : α) :

                CovBy a b means that b covers a. This means that a < b and there is no element in between. This is denoted a ⋖ b.

                Equations
                Instances For

                  CovBy a b means that b covers a. This means that a < b and there is no element in between. This is denoted a ⋖ b.

                  Equations
                  Instances For

                    Definition of PartialOrder and lemmas about types with a partial order #

                    class PartialOrder (α : Type u_2) extends Preorder α :
                    Type u_2

                    A partial order is a reflexive, transitive, antisymmetric relation .

                    Instances
                      def PartialOrder.mk' {α : Type u_1} [Preorder α] (le_antisymm : ∀ (a b : α), b aa ba = b) :

                      A variant of PartialOrder.mk which allows to_dual to dualize a PartialOrder instance.

                      Equations
                      Instances For
                        theorem le_antisymm {α : Type u_1} [PartialOrder α] {a b : α} :
                        a bb aa = b
                        theorem ge_antisymm {α : Type u_1} [PartialOrder α] {a b : α} :
                        b aa ba = b
                        theorem eq_of_le_of_ge {α : Type u_1} [PartialOrder α] {a b : α} :
                        a bb aa = b

                        Alias of le_antisymm.

                        theorem eq_of_ge_of_le {α : Type u_1} [PartialOrder α] {a b : α} :
                        b aa ba = b
                        @[deprecated eq_of_le_of_ge (since := "2025-06-07")]
                        theorem eq_of_le_of_le {α : Type u_1} [PartialOrder α] {a b : α} :
                        a bb aa = b

                        Alias of le_antisymm.


                        Alias of le_antisymm.

                        theorem le_antisymm_iff {α : Type u_1} [PartialOrder α] {a b : α} :
                        a = b a b b a
                        theorem ge_antisymm_iff {α : Type u_1} [PartialOrder α] {a b : α} :
                        a = b b a a b
                        theorem lt_of_le_of_ne {α : Type u_1} [PartialOrder α] {a b : α} :
                        a ba ba < b
                        theorem lt_of_le_of_ne' {α : Type u_1} [PartialOrder α] {a b : α} :
                        b aa bb < a

                        Equality is decidable if is.

                        Equations
                        Instances For

                          Equality is decidable if is.

                          Equations
                          Instances For
                            @[deprecated decidableEqOfDecidableLE' (since := "2025-12-09")]

                            Alias of decidableEqOfDecidableLE'.


                            Equality is decidable if is.

                            Equations
                            Instances For
                              @[deprecated decidableEqOfDecidableLE' (since := "2025-12-09")]

                              Alias of decidableEqOfDecidableLE'.


                              Equality is decidable if is.

                              Equations
                              Instances For
                                theorem Decidable.lt_or_eq_of_le {α : Type u_1} [PartialOrder α] {a b : α} [DecidableLE α] (hab : a b) :
                                a < b a = b
                                theorem Decidable.lt_or_eq_of_le' {α : Type u_1} [PartialOrder α] {a b : α} [DecidableLE' α] (hab : b a) :
                                b < a a = b
                                theorem Decidable.le_iff_lt_or_eq {α : Type u_1} [PartialOrder α] {a b : α} [DecidableLE α] :
                                a b a < b a = b
                                theorem Decidable.le_iff_lt_or_eq' {α : Type u_1} [PartialOrder α] {a b : α} [DecidableLE' α] :
                                b a b < a a = b
                                theorem lt_or_eq_of_le {α : Type u_1} [PartialOrder α] {a b : α} :
                                a ba < b a = b
                                theorem lt_or_eq_of_le' {α : Type u_1} [PartialOrder α] {a b : α} :
                                b ab < a a = b
                                theorem le_iff_lt_or_eq {α : Type u_1} [PartialOrder α] {a b : α} :
                                a b a < b a = b
                                theorem le_iff_lt_or_eq' {α : Type u_1} [PartialOrder α] {a b : α} :
                                b a b < a a = b