Documentation

Init.Data.Order.FactoriesExtra

@[inline]
def LE.ofOrd (α : Type u) [Ord α] :
LE α

Creates an LE α instance from an Ord α instance.

OrientedOrd α must be satisfied so that the resulting LE α instance faithfully represents the Ord α instance.

Equations
Instances For
    @[inline]
    def DecidableLE.ofOrd (α : Type u) [LE α] [Ord α] [Std.LawfulOrderOrd α] :

    Creates an DecidableLE α instance using a well-behaved Ord α instance.

    Equations
    Instances For
      @[inline]
      def LT.ofOrd (α : Type u) [Ord α] :
      LT α

      Creates an LT α instance from an Ord α instance.

      OrientedOrd α must be satisfied so that the resulting LT α instance faithfully represents the Ord α instance.

      Equations
      Instances For
        theorem Std.isLE_compare {α : Type u} [Ord α] [LE α] [LawfulOrderOrd α] {a b : α} :
        (compare a b).isLE = true a b
        theorem Std.isGE_compare {α : Type u} [Ord α] [LE α] [LawfulOrderOrd α] {a b : α} :
        (compare a b).isGE = true b a
        theorem Std.compare_eq_lt {α : Type u} [Ord α] [LT α] [LE α] [LawfulOrderOrd α] [LawfulOrderLT α] {a b : α} :
        theorem Std.compare_eq_gt {α : Type u} [Ord α] [LT α] [LE α] [LawfulOrderOrd α] [LawfulOrderLT α] {a b : α} :
        theorem Std.compare_eq_eq {α : Type u} [Ord α] [BEq α] [LE α] [LawfulOrderOrd α] [LawfulOrderBEq α] {a b : α} :
        @[inline]

        Creates a DecidableLT α instance using a well-behaved Ord α instance.

        Equations
        Instances For
          @[inline]
          def BEq.ofOrd (α : Type u) [Ord α] :
          BEq α

          Creates a BEq α instance from an Ord α instance.

          Equations
          Instances For

            The LE α instance obtained from an Ord α instance is compatible with said Ord α instance if compare is oriented, i.e., compare a b = .lt ↔ compare b a = .gt.

            instance Std.instLawfulOrderLT_ofOrd {α : Type u} [Ord α] [LE α] [LawfulOrderOrd α] :

            The LT α instance obtained from an Ord α instance is compatible with the LE α instance instance if Ord α is compatible with it.

            The BEq α instance obtained from an Ord α instance is compatible with the LE α instance instance if Ord α is compatible with it.