Documentation

Mathlib.Order.Defs.Unbundled

Orders #

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

Unbundled classes #

def EmptyRelation {α : Sort u_1} :
ααProp

An empty relation does not relate any elements.

Equations
Instances For
    class IsIrrefl (α : Sort u_1) (r : ααProp) :

    IsIrrefl X r means the binary relation r on X is irreflexive (that is, r x x never holds).

    • irrefl : ∀ (a : α), ¬r a a
    Instances
      class IsRefl (α : Sort u_1) (r : ααProp) :

      IsRefl X r means the binary relation r on X is reflexive.

      • refl : ∀ (a : α), r a a
      Instances
        class IsSymm (α : Sort u_1) (r : ααProp) :

        IsSymm X r means the binary relation r on X is symmetric.

        • symm : ∀ (a b : α), r a br b a
        Instances
          class IsAsymm (α : Sort u_1) (r : ααProp) :

          IsAsymm X r means that the binary relation r on X is asymmetric, that is, r a b → ¬ r b a.

          • asymm : ∀ (a b : α), r a b¬r b a
          Instances
            class IsAntisymm (α : Sort u_1) (r : ααProp) :

            IsAntisymm X r means the binary relation r on X is antisymmetric.

            • antisymm : ∀ (a b : α), r a br b aa = b
            Instances
              theorem IsAsymm.toIsAntisymm {α : Sort u_1} (r : ααProp) [IsAsymm α r] :
              class IsTrans (α : Sort u_1) (r : ααProp) :

              IsTrans X r means the binary relation r on X is transitive.

              • trans : ∀ (a b c : α), r a br b cr a c
              Instances
                instance instTransOfIsTrans {α : Sort u_1} {r : ααProp} [IsTrans α r] :
                Trans r r r
                Equations
                • instTransOfIsTrans = { trans := }
                theorem instIsTransOfTrans {α : Sort u_1} {r : ααProp} [Trans r r r] :
                IsTrans α r
                class IsTotal (α : Sort u_1) (r : ααProp) :

                IsTotal X r means that the binary relation r on X is total, that is, that for any x y : X we have r x y or r y x.

                • total : ∀ (a b : α), r a b r b a
                Instances
                  class IsPreorder (α : Sort u_1) (r : ααProp) extends IsRefl α r, IsTrans α r :

                  IsPreorder X r means that the binary relation r on X is a pre-order, that is, reflexive and transitive.

                  • refl : ∀ (a : α), r a a
                  • trans : ∀ (a b c : α), r a br b cr a c
                  Instances
                    class IsPartialOrder (α : Sort u_1) (r : ααProp) extends IsPreorder α r, IsAntisymm α r :

                    IsPartialOrder X r means that the binary relation r on X is a partial order, that is, IsPreorder X r and IsAntisymm X r.

                    • refl : ∀ (a : α), r a a
                    • trans : ∀ (a b c : α), r a br b cr a c
                    • antisymm : ∀ (a b : α), r a br b aa = b
                    Instances
                      class IsLinearOrder (α : Sort u_1) (r : ααProp) extends IsPartialOrder α r, IsTotal α r :

                      IsLinearOrder X r means that the binary relation r on X is a linear order, that is, IsPartialOrder X r and IsTotal X r.

                      • refl : ∀ (a : α), r a a
                      • trans : ∀ (a b c : α), r a br b cr a c
                      • antisymm : ∀ (a b : α), r a br b aa = b
                      • total : ∀ (a b : α), r a b r b a
                      Instances
                        class IsEquiv (α : Sort u_1) (r : ααProp) extends IsPreorder α r, IsSymm α r :

                        IsEquiv X r means that the binary relation r on X is an equivalence relation, that is, IsPreorder X r and IsSymm X r.

                        • refl : ∀ (a : α), r a a
                        • trans : ∀ (a b c : α), r a br b cr a c
                        • symm : ∀ (a b : α), r a br b a
                        Instances
                          class IsStrictOrder (α : Sort u_1) (r : ααProp) extends IsIrrefl α r, IsTrans α r :

                          IsStrictOrder X r means that the binary relation r on X is a strict order, that is, IsIrrefl X r and IsTrans X r.

                          • irrefl : ∀ (a : α), ¬r a a
                          • trans : ∀ (a b c : α), r a br b cr a c
                          Instances
                            class IsStrictWeakOrder (α : Sort u_1) (lt : ααProp) extends IsStrictOrder α lt :

                            IsStrictWeakOrder X lt means that the binary relation lt on X is a strict weak order, that is, IsStrictOrder X lt and ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a.

                            • irrefl : ∀ (a : α), ¬lt a a
                            • trans : ∀ (a b c : α), lt a blt b clt a c
                            • incomp_trans : ∀ (a b c : α), ¬lt a b ¬lt b a¬lt b c ¬lt c b¬lt a c ¬lt c a
                            Instances
                              class IsTrichotomous (α : Sort u_1) (lt : ααProp) :

                              IsTrichotomous X lt means that the binary relation lt on X is trichotomous, that is, either lt a b or a = b or lt b a for any a and b.

                              • trichotomous : ∀ (a b : α), lt a b a = b lt b a
                              Instances
                                class IsStrictTotalOrder (α : Sort u_1) (lt : ααProp) extends IsTrichotomous α lt, IsStrictOrder α lt :

                                IsStrictTotalOrder X lt means that the binary relation lt on X is a strict total order, that is, IsTrichotomous X lt and IsStrictOrder X lt.

                                Instances
                                  theorem eq_isEquiv (α : Sort u_1) :
                                  IsEquiv α fun (x1 x2 : α) => x1 = x2

                                  Equality is an equivalence relation.

                                  Iff is an equivalence relation.

                                  theorem irrefl {α : Sort u_1} {r : ααProp} [IsIrrefl α r] (a : α) :
                                  ¬r a a
                                  theorem refl {α : Sort u_1} {r : ααProp} [IsRefl α r] (a : α) :
                                  r a a
                                  theorem trans {α : Sort u_1} {r : ααProp} {a b c : α} [IsTrans α r] :
                                  r a br b cr a c
                                  theorem symm {α : Sort u_1} {r : ααProp} {a b : α} [IsSymm α r] :
                                  r a br b a
                                  theorem antisymm {α : Sort u_1} {r : ααProp} {a b : α} [IsAntisymm α r] :
                                  r a br b aa = b
                                  theorem asymm {α : Sort u_1} {r : ααProp} {a b : α} [IsAsymm α r] :
                                  r a b¬r b a
                                  theorem trichotomous {α : Sort u_1} {r : ααProp} [IsTrichotomous α r] (a b : α) :
                                  r a b a = b r b a
                                  theorem isAsymm_of_isTrans_of_isIrrefl {α : Sort u_1} {r : ααProp} [IsTrans α r] [IsIrrefl α r] :
                                  IsAsymm α r
                                  theorem IsIrrefl.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsIrrefl α r] :
                                  IsIrrefl α fun (a b : α) => decide (r a b) = true
                                  theorem IsRefl.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsRefl α r] :
                                  IsRefl α fun (a b : α) => decide (r a b) = true
                                  theorem IsTrans.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsTrans α r] :
                                  IsTrans α fun (a b : α) => decide (r a b) = true
                                  theorem IsSymm.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsSymm α r] :
                                  IsSymm α fun (a b : α) => decide (r a b) = true
                                  theorem IsAntisymm.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsAntisymm α r] :
                                  IsAntisymm α fun (a b : α) => decide (r a b) = true
                                  theorem IsAsymm.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsAsymm α r] :
                                  IsAsymm α fun (a b : α) => decide (r a b) = true
                                  theorem IsTotal.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsTotal α r] :
                                  IsTotal α fun (a b : α) => decide (r a b) = true
                                  theorem IsTrichotomous.decide {α : Sort u_1} {r : ααProp} [DecidableRel r] [IsTrichotomous α r] :
                                  IsTrichotomous α fun (a b : α) => decide (r a b) = true
                                  @[elab_without_expected_type]
                                  theorem irrefl_of {α : Sort u_1} (r : ααProp) [IsIrrefl α r] (a : α) :
                                  ¬r a a
                                  @[elab_without_expected_type]
                                  theorem refl_of {α : Sort u_1} (r : ααProp) [IsRefl α r] (a : α) :
                                  r a a
                                  @[elab_without_expected_type]
                                  theorem trans_of {α : Sort u_1} (r : ααProp) {a b c : α} [IsTrans α r] :
                                  r a br b cr a c
                                  @[elab_without_expected_type]
                                  theorem symm_of {α : Sort u_1} (r : ααProp) {a b : α} [IsSymm α r] :
                                  r a br b a
                                  @[elab_without_expected_type]
                                  theorem asymm_of {α : Sort u_1} (r : ααProp) {a b : α} [IsAsymm α r] :
                                  r a b¬r b a
                                  @[elab_without_expected_type]
                                  theorem total_of {α : Sort u_1} (r : ααProp) [IsTotal α r] (a b : α) :
                                  r a b r b a
                                  @[elab_without_expected_type]
                                  theorem trichotomous_of {α : Sort u_1} (r : ααProp) [IsTrichotomous α r] (a b : α) :
                                  r a b a = b r b a
                                  def Reflexive {α : Sort u_1} (r : ααProp) :

                                  IsRefl as a definition, suitable for use in proofs.

                                  Equations
                                  Instances For
                                    def Symmetric {α : Sort u_1} (r : ααProp) :

                                    IsSymm as a definition, suitable for use in proofs.

                                    Equations
                                    Instances For
                                      def Transitive {α : Sort u_1} (r : ααProp) :

                                      IsTrans as a definition, suitable for use in proofs.

                                      Equations
                                      • Transitive r = ∀ ⦃x y z : α⦄, r x yr y zr x z
                                      Instances For
                                        def Irreflexive {α : Sort u_1} (r : ααProp) :

                                        IsIrrefl as a definition, suitable for use in proofs.

                                        Equations
                                        Instances For
                                          def AntiSymmetric {α : Sort u_1} (r : ααProp) :

                                          IsAntisymm as a definition, suitable for use in proofs.

                                          Equations
                                          Instances For
                                            def Total {α : Sort u_1} (r : ααProp) :

                                            IsTotal as a definition, suitable for use in proofs.

                                            Equations
                                            Instances For
                                              @[deprecated Equivalence.refl]
                                              theorem Equivalence.reflexive {α : Sort u_1} (r : ααProp) (h : Equivalence r) :
                                              @[deprecated Equivalence.symm]
                                              theorem Equivalence.symmetric {α : Sort u_1} (r : ααProp) (h : Equivalence r) :
                                              @[deprecated Equivalence.trans]
                                              theorem Equivalence.transitive {α : Sort u_1} (r : ααProp) (h : Equivalence r) :
                                              @[deprecated]
                                              theorem InvImage.trans {α : Sort u_1} {β : Sort u_2} (r : ββProp) (f : αβ) (h : Transitive r) :
                                              @[deprecated]
                                              theorem InvImage.irreflexive {α : Sort u_1} {β : Sort u_2} (r : ββProp) (f : αβ) (h : Irreflexive r) :

                                              Minimal and maximal #

                                              def Minimal {α : Type u_1} [LE α] (P : αProp) (x : α) :

                                              Minimal P x means that x is a minimal element satisfying P.

                                              Equations
                                              Instances For
                                                def Maximal {α : Type u_1} [LE α] (P : αProp) (x : α) :

                                                Maximal P x means that x is a maximal element satisfying P.

                                                Equations
                                                Instances For
                                                  theorem Minimal.prop {α : Type u_1} [LE α] {P : αProp} {x : α} (h : Minimal P x) :
                                                  P x
                                                  theorem Maximal.prop {α : Type u_1} [LE α] {P : αProp} {x : α} (h : Maximal P x) :
                                                  P x
                                                  theorem Minimal.le_of_le {α : Type u_1} [LE α] {P : αProp} {x y : α} (h : Minimal P x) (hy : P y) (hle : y x) :
                                                  x y
                                                  theorem Maximal.le_of_ge {α : Type u_1} [LE α] {P : αProp} {x y : α} (h : Maximal P x) (hy : P y) (hge : x y) :
                                                  y x

                                                  Upper and lower sets #

                                                  def IsUpperSet {α : Type u_1} [LE α] (s : Set α) :

                                                  An upper set in an order α is a set such that any element greater than one of its members is also a member. Also called up-set, upward-closed set.

                                                  Equations
                                                  Instances For
                                                    def IsLowerSet {α : Type u_1} [LE α] (s : Set α) :

                                                    A lower set in an order α is a set such that any element less than one of its members is also a member. Also called down-set, downward-closed set.

                                                    Equations
                                                    Instances For
                                                      structure UpperSet (α : Type u_1) [LE α] :
                                                      Type u_1

                                                      The type of upper sets of an order.

                                                      An upper set in an order α is a set such that any element greater than one of its members is also a member. Also called up-set, upward-closed set.

                                                      Instances For
                                                        structure LowerSet (α : Type u_1) [LE α] :
                                                        Type u_1

                                                        The type of lower sets of an order.

                                                        A lower set in an order α is a set such that any element less than one of its members is also a member. Also called down-set, downward-closed set.

                                                        Instances For