Documentation

Mathlib.Order.Monotone.Basic

Monotonicity #

This file defines (strictly) monotone/antitone functions. Contrary to standard mathematical usage, "monotone"/"mono" here means "increasing", not "increasing or decreasing". We use "antitone"/"anti" to mean "decreasing".

Definitions #

Main theorems #

Implementation notes #

Some of these definitions used to only require LE α or LT α. The advantage of this is unclear and it led to slight elaboration issues. Now, everything requires Preorder α and seems to work fine. Related Zulip discussion: https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Order.20diamond/near/254353352.

TODO #

The above theorems are also true in ℕ+, Fin n... To make that work, we need SuccOrder α and IsSuccArchimedean α.

Tags #

monotone, strictly monotone, antitone, strictly antitone, increasing, strictly increasing, decreasing, strictly decreasing

def Monotone {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) :

A function f is monotone if a ≤ b implies f a ≤ f b.

Equations
Instances For
    def Antitone {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) :

    A function f is antitone if a ≤ b implies f b ≤ f a.

    Equations
    Instances For
      def MonotoneOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) (s : Set α) :

      A function f is monotone on s if, for all a, b ∈ s, a ≤ b implies f a ≤ f b.

      Equations
      Instances For
        def AntitoneOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) (s : Set α) :

        A function f is antitone on s if, for all a, b ∈ s, a ≤ b implies f b ≤ f a.

        Equations
        Instances For
          def StrictMono {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) :

          A function f is strictly monotone if a < b implies f a < f b.

          Equations
          Instances For
            def StrictAnti {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) :

            A function f is strictly antitone if a < b implies f b < f a.

            Equations
            Instances For
              def StrictMonoOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) (s : Set α) :

              A function f is strictly monotone on s if, for all a, b ∈ s, a < b implies f a < f b.

              Equations
              Instances For
                def StrictAntiOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] (f : αβ) (s : Set α) :

                A function f is strictly antitone on s if, for all a, b ∈ s, a < b implies f b < f a.

                Equations
                Instances For
                  instance instDecidableMonotone {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} [i : Decidable (∀ (a b : α), a bf a f b)] :
                  Equations
                  • instDecidableMonotone = i
                  instance instDecidableAntitone {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} [i : Decidable (∀ (a b : α), a bf b f a)] :
                  Equations
                  • instDecidableAntitone = i
                  instance instDecidableMonotoneOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} [i : Decidable (∀ (a : α), a s∀ (b : α), b sa bf a f b)] :
                  Equations
                  • instDecidableMonotoneOn = i
                  instance instDecidableAntitoneOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} [i : Decidable (∀ (a : α), a s∀ (b : α), b sa bf b f a)] :
                  Equations
                  • instDecidableAntitoneOn = i
                  instance instDecidableStrictMono {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} [i : Decidable (∀ (a b : α), a < bf a < f b)] :
                  Equations
                  • instDecidableStrictMono = i
                  instance instDecidableStrictAnti {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} [i : Decidable (∀ (a b : α), a < bf b < f a)] :
                  Equations
                  • instDecidableStrictAnti = i
                  instance instDecidableStrictMonoOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} [i : Decidable (∀ (a : α), a s∀ (b : α), b sa < bf a < f b)] :
                  Equations
                  • instDecidableStrictMonoOn = i
                  instance instDecidableStrictAntiOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} [i : Decidable (∀ (a : α), a s∀ (b : α), b sa < bf b < f a)] :
                  Equations
                  • instDecidableStrictAntiOn = i

                  Monotonicity on the dual order #

                  Strictly, many of the *On.dual lemmas in this section should use ofDual ⁻¹' s instead of s, but right now this is not possible as Set.preimage is not defined yet, and importing it creates an import cycle.

                  Often, you should not need the rewriting lemmas. Instead, you probably want to add .dual, .dual_left or .dual_right to your Monotone/Antitone hypothesis.

                  @[simp]
                  theorem monotone_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Monotone (f OrderDual.ofDual) Antitone f
                  @[simp]
                  theorem antitone_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Antitone (f OrderDual.ofDual) Monotone f
                  @[simp]
                  theorem monotone_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Monotone (OrderDual.toDual f) Antitone f
                  @[simp]
                  theorem antitone_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Antitone (OrderDual.toDual f) Monotone f
                  @[simp]
                  theorem monotoneOn_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn (f OrderDual.ofDual) s AntitoneOn f s
                  @[simp]
                  theorem antitoneOn_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn (f OrderDual.ofDual) s MonotoneOn f s
                  @[simp]
                  theorem monotoneOn_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn (OrderDual.toDual f) s AntitoneOn f s
                  @[simp]
                  theorem antitoneOn_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn (OrderDual.toDual f) s MonotoneOn f s
                  @[simp]
                  theorem strictMono_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictMono (f OrderDual.ofDual) StrictAnti f
                  @[simp]
                  theorem strictAnti_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictAnti (f OrderDual.ofDual) StrictMono f
                  @[simp]
                  theorem strictMono_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictMono (OrderDual.toDual f) StrictAnti f
                  @[simp]
                  theorem strictAnti_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictAnti (OrderDual.toDual f) StrictMono f
                  @[simp]
                  theorem strictMonoOn_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictMonoOn (f OrderDual.ofDual) s StrictAntiOn f s
                  @[simp]
                  theorem strictAntiOn_comp_ofDual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictAntiOn (f OrderDual.ofDual) s StrictMonoOn f s
                  @[simp]
                  theorem strictMonoOn_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictMonoOn (OrderDual.toDual f) s StrictAntiOn f s
                  @[simp]
                  theorem strictAntiOn_toDual_comp_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictAntiOn (OrderDual.toDual f) s StrictMonoOn f s
                  theorem monotone_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Monotone (OrderDual.toDual f OrderDual.ofDual) Monotone f
                  theorem antitone_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Antitone (OrderDual.toDual f OrderDual.ofDual) Antitone f
                  theorem monotoneOn_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn (OrderDual.toDual f OrderDual.ofDual) s MonotoneOn f s
                  theorem antitoneOn_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn (OrderDual.toDual f OrderDual.ofDual) s AntitoneOn f s
                  theorem strictMono_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictMono (OrderDual.toDual f OrderDual.ofDual) StrictMono f
                  theorem strictAnti_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictAnti (OrderDual.toDual f OrderDual.ofDual) StrictAnti f
                  theorem strictMonoOn_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictMonoOn (OrderDual.toDual f OrderDual.ofDual) s StrictMonoOn f s
                  theorem strictAntiOn_dual_iff {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictAntiOn (OrderDual.toDual f OrderDual.ofDual) s StrictAntiOn f s
                  theorem Monotone.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Monotone fAntitone (f OrderDual.ofDual)

                  Alias of the reverse direction of antitone_comp_ofDual_iff.

                  theorem Antitone.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Antitone fMonotone (f OrderDual.ofDual)

                  Alias of the reverse direction of monotone_comp_ofDual_iff.

                  theorem Monotone.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Monotone fAntitone (OrderDual.toDual f)

                  Alias of the reverse direction of antitone_toDual_comp_iff.

                  theorem Antitone.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Antitone fMonotone (OrderDual.toDual f)

                  Alias of the reverse direction of monotone_toDual_comp_iff.

                  theorem MonotoneOn.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn f sAntitoneOn (f OrderDual.ofDual) s

                  Alias of the reverse direction of antitoneOn_comp_ofDual_iff.

                  theorem AntitoneOn.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn f sMonotoneOn (f OrderDual.ofDual) s

                  Alias of the reverse direction of monotoneOn_comp_ofDual_iff.

                  theorem MonotoneOn.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn f sAntitoneOn (OrderDual.toDual f) s

                  Alias of the reverse direction of antitoneOn_toDual_comp_iff.

                  theorem AntitoneOn.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn f sMonotoneOn (OrderDual.toDual f) s

                  Alias of the reverse direction of monotoneOn_toDual_comp_iff.

                  theorem StrictMono.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictMono fStrictAnti (f OrderDual.ofDual)

                  Alias of the reverse direction of strictAnti_comp_ofDual_iff.

                  theorem StrictAnti.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictAnti fStrictMono (f OrderDual.ofDual)

                  Alias of the reverse direction of strictMono_comp_ofDual_iff.

                  theorem StrictMono.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictMono fStrictAnti (OrderDual.toDual f)

                  Alias of the reverse direction of strictAnti_toDual_comp_iff.

                  theorem StrictAnti.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictAnti fStrictMono (OrderDual.toDual f)

                  Alias of the reverse direction of strictMono_toDual_comp_iff.

                  theorem StrictMonoOn.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictMonoOn f sStrictAntiOn (f OrderDual.ofDual) s

                  Alias of the reverse direction of strictAntiOn_comp_ofDual_iff.

                  theorem StrictAntiOn.dual_left {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictAntiOn f sStrictMonoOn (f OrderDual.ofDual) s

                  Alias of the reverse direction of strictMonoOn_comp_ofDual_iff.

                  theorem StrictMonoOn.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictMonoOn f sStrictAntiOn (OrderDual.toDual f) s

                  Alias of the reverse direction of strictAntiOn_toDual_comp_iff.

                  theorem StrictAntiOn.dual_right {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictAntiOn f sStrictMonoOn (OrderDual.toDual f) s

                  Alias of the reverse direction of strictMonoOn_toDual_comp_iff.

                  theorem Monotone.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Monotone fMonotone (OrderDual.toDual f OrderDual.ofDual)

                  Alias of the reverse direction of monotone_dual_iff.

                  theorem Antitone.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  Antitone fAntitone (OrderDual.toDual f OrderDual.ofDual)

                  Alias of the reverse direction of antitone_dual_iff.

                  theorem MonotoneOn.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn f sMonotoneOn (OrderDual.toDual f OrderDual.ofDual) s

                  Alias of the reverse direction of monotoneOn_dual_iff.

                  theorem AntitoneOn.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn f sAntitoneOn (OrderDual.toDual f OrderDual.ofDual) s

                  Alias of the reverse direction of antitoneOn_dual_iff.

                  theorem StrictMono.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictMono fStrictMono (OrderDual.toDual f OrderDual.ofDual)

                  Alias of the reverse direction of strictMono_dual_iff.

                  theorem StrictAnti.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  StrictAnti fStrictAnti (OrderDual.toDual f OrderDual.ofDual)

                  Alias of the reverse direction of strictAnti_dual_iff.

                  theorem StrictMonoOn.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictMonoOn f sStrictMonoOn (OrderDual.toDual f OrderDual.ofDual) s

                  Alias of the reverse direction of strictMonoOn_dual_iff.

                  theorem StrictAntiOn.dual {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {s : Set α} :
                  StrictAntiOn f sStrictAntiOn (OrderDual.toDual f OrderDual.ofDual) s

                  Alias of the reverse direction of strictAntiOn_dual_iff.

                  Monotonicity in function spaces #

                  theorem Monotone.comp_le_comp_left {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] {f : βα} {g : γβ} {h : γβ} (hf : Monotone f) (le_gh : g h) :
                  f g f h
                  theorem monotone_lam {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder γ] {f : αβγ} (hf : ∀ (b : β), Monotone fun (a : α) => f a b) :
                  theorem monotone_app {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder γ] (f : βαγ) (b : β) (hf : Monotone fun (a : α) (b : β) => f b a) :
                  Monotone (f b)
                  theorem antitone_lam {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder γ] {f : αβγ} (hf : ∀ (b : β), Antitone fun (a : α) => f a b) :
                  theorem antitone_app {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder γ] (f : βαγ) (b : β) (hf : Antitone fun (a : α) (b : β) => f b a) :
                  Antitone (f b)
                  theorem Function.monotone_eval {ι : Type u} {α : ιType v} [(i : ι) → Preorder (α i)] (i : ι) :

                  Monotonicity hierarchy #

                  These four lemmas are there to strip off the semi-implicit arguments ⦃a b : α⦄. This is useful when you do not want to apply a Monotone assumption (i.e. your goal is a ≤ b → f a ≤ f b). However if you find yourself writing hf.imp h, then you should have written hf h instead.

                  theorem Monotone.imp {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} {b : α} (hf : Monotone f) (h : a b) :
                  f a f b
                  theorem Antitone.imp {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} {b : α} (hf : Antitone f) (h : a b) :
                  f b f a
                  theorem StrictMono.imp {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} {b : α} (hf : StrictMono f) (h : a < b) :
                  f a < f b
                  theorem StrictAnti.imp {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} {b : α} (hf : StrictAnti f) (h : a < b) :
                  f b < f a
                  theorem Monotone.monotoneOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Monotone f) (s : Set α) :
                  theorem Antitone.antitoneOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : Antitone f) (s : Set α) :
                  @[simp]
                  theorem monotoneOn_univ {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  MonotoneOn f Set.univ Monotone f
                  @[simp]
                  theorem antitoneOn_univ {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  AntitoneOn f Set.univ Antitone f
                  theorem StrictMono.strictMonoOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : StrictMono f) (s : Set α) :
                  theorem StrictAnti.strictAntiOn {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (hf : StrictAnti f) (s : Set α) :
                  @[simp]
                  theorem strictMonoOn_univ {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  @[simp]
                  theorem strictAntiOn_univ {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} :
                  theorem Monotone.strictMono_of_injective {α : Type u} {β : Type v} [Preorder α] [PartialOrder β] {f : αβ} (h₁ : Monotone f) (h₂ : Function.Injective f) :
                  theorem Antitone.strictAnti_of_injective {α : Type u} {β : Type v} [Preorder α] [PartialOrder β] {f : αβ} (h₁ : Antitone f) (h₂ : Function.Injective f) :
                  theorem monotone_iff_forall_lt {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} :
                  Monotone f ∀ ⦃a b : α⦄, a < bf a f b
                  theorem antitone_iff_forall_lt {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} :
                  Antitone f ∀ ⦃a b : α⦄, a < bf b f a
                  theorem monotoneOn_iff_forall_lt {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} {s : Set α} :
                  MonotoneOn f s ∀ ⦃a : α⦄, a s∀ ⦃b : α⦄, b sa < bf a f b
                  theorem antitoneOn_iff_forall_lt {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} {s : Set α} :
                  AntitoneOn f s ∀ ⦃a : α⦄, a s∀ ⦃b : α⦄, b sa < bf b f a
                  theorem StrictMonoOn.monotoneOn {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictMonoOn f s) :
                  theorem StrictAntiOn.antitoneOn {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictAntiOn f s) :
                  theorem StrictMono.monotone {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) :
                  theorem StrictAnti.antitone {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) :

                  Monotonicity from and to subsingletons #

                  theorem Subsingleton.monotone {α : Type u} {β : Type v} [Preorder α] [Preorder β] [Subsingleton α] (f : αβ) :
                  theorem Subsingleton.antitone {α : Type u} {β : Type v} [Preorder α] [Preorder β] [Subsingleton α] (f : αβ) :
                  theorem Subsingleton.monotone' {α : Type u} {β : Type v} [Preorder α] [Preorder β] [Subsingleton β] (f : αβ) :
                  theorem Subsingleton.antitone' {α : Type u} {β : Type v} [Preorder α] [Preorder β] [Subsingleton β] (f : αβ) :
                  theorem Subsingleton.strictMono {α : Type u} {β : Type v} [Preorder α] [Preorder β] [Subsingleton α] (f : αβ) :
                  theorem Subsingleton.strictAnti {α : Type u} {β : Type v} [Preorder α] [Preorder β] [Subsingleton α] (f : αβ) :

                  Miscellaneous monotonicity results #

                  theorem monotone_id {α : Type u} [Preorder α] :
                  theorem monotoneOn_id {α : Type u} [Preorder α] {s : Set α} :
                  theorem strictMono_id {α : Type u} [Preorder α] :
                  theorem strictMonoOn_id {α : Type u} [Preorder α] {s : Set α} :
                  theorem monotone_const {α : Type u} {β : Type v} [Preorder α] [Preorder β] {c : β} :
                  Monotone fun (x : α) => c
                  theorem monotoneOn_const {α : Type u} {β : Type v} [Preorder α] [Preorder β] {c : β} {s : Set α} :
                  MonotoneOn (fun (x : α) => c) s
                  theorem antitone_const {α : Type u} {β : Type v} [Preorder α] [Preorder β] {c : β} :
                  Antitone fun (x : α) => c
                  theorem antitoneOn_const {α : Type u} {β : Type v} [Preorder α] [Preorder β] {c : β} {s : Set α} :
                  AntitoneOn (fun (x : α) => c) s
                  theorem strictMono_of_le_iff_le {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (h : ∀ (x y : α), x y f x f y) :
                  theorem strictAnti_of_le_iff_le {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} (h : ∀ (x y : α), x y f y f x) :
                  theorem injective_of_lt_imp_ne {α : Type u} {β : Type v} [LinearOrder α] {f : αβ} (h : ∀ (x y : α), x < yf x f y) :
                  theorem injective_of_le_imp_le {α : Type u} {β : Type v} [PartialOrder α] [Preorder β] (f : αβ) (h : ∀ {x y : α}, f x f yx y) :
                  theorem StrictMono.isMax_of_apply {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} (hf : StrictMono f) (ha : IsMax (f a)) :
                  theorem StrictMono.isMin_of_apply {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} (hf : StrictMono f) (ha : IsMin (f a)) :
                  theorem StrictAnti.isMax_of_apply {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} (hf : StrictAnti f) (ha : IsMin (f a)) :
                  theorem StrictAnti.isMin_of_apply {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {a : α} (hf : StrictAnti f) (ha : IsMax (f a)) :
                  theorem StrictMono.ite' {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {g : αβ} (hf : StrictMono f) (hg : StrictMono g) {p : αProp} [DecidablePred p] (hp : ∀ ⦃x y : α⦄, x < yp yp x) (hfg : ∀ ⦃x y : α⦄, p x¬p yx < yf x < g y) :
                  StrictMono fun (x : α) => if p x then f x else g x
                  theorem StrictMono.ite {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {g : αβ} (hf : StrictMono f) (hg : StrictMono g) {p : αProp} [DecidablePred p] (hp : ∀ ⦃x y : α⦄, x < yp yp x) (hfg : ∀ (x : α), f x g x) :
                  StrictMono fun (x : α) => if p x then f x else g x
                  theorem StrictAnti.ite' {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {g : αβ} (hf : StrictAnti f) (hg : StrictAnti g) {p : αProp} [DecidablePred p] (hp : ∀ ⦃x y : α⦄, x < yp yp x) (hfg : ∀ ⦃x y : α⦄, p x¬p yx < yg y < f x) :
                  StrictAnti fun (x : α) => if p x then f x else g x
                  theorem StrictAnti.ite {α : Type u} {β : Type v} [Preorder α] [Preorder β] {f : αβ} {g : αβ} (hf : StrictAnti f) (hg : StrictAnti g) {p : αProp} [DecidablePred p] (hp : ∀ ⦃x y : α⦄, x < yp yp x) (hfg : ∀ (x : α), g x f x) :
                  StrictAnti fun (x : α) => if p x then f x else g x

                  Monotonicity under composition #

                  theorem Monotone.comp {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : Monotone g) (hf : Monotone f) :
                  theorem Monotone.comp_antitone {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : Monotone g) (hf : Antitone f) :
                  theorem Antitone.comp {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : Antitone g) (hf : Antitone f) :
                  theorem Antitone.comp_monotone {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : Antitone g) (hf : Monotone f) :
                  theorem Monotone.iterate {α : Type u} [Preorder α] {f : αα} (hf : Monotone f) (n : ) :
                  theorem Monotone.comp_monotoneOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : Monotone g) (hf : MonotoneOn f s) :
                  MonotoneOn (g f) s
                  theorem Monotone.comp_antitoneOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : Monotone g) (hf : AntitoneOn f s) :
                  AntitoneOn (g f) s
                  theorem Antitone.comp_antitoneOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : Antitone g) (hf : AntitoneOn f s) :
                  MonotoneOn (g f) s
                  theorem Antitone.comp_monotoneOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : Antitone g) (hf : MonotoneOn f s) :
                  AntitoneOn (g f) s
                  theorem StrictMono.comp {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : StrictMono g) (hf : StrictMono f) :
                  theorem StrictMono.comp_strictAnti {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : StrictMono g) (hf : StrictAnti f) :
                  theorem StrictAnti.comp {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : StrictAnti g) (hf : StrictAnti f) :
                  theorem StrictAnti.comp_strictMono {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} (hg : StrictAnti g) (hf : StrictMono f) :
                  theorem StrictMono.iterate {α : Type u} [Preorder α] {f : αα} (hf : StrictMono f) (n : ) :
                  theorem StrictMono.comp_strictMonoOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : StrictMono g) (hf : StrictMonoOn f s) :
                  theorem StrictMono.comp_strictAntiOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : StrictMono g) (hf : StrictAntiOn f s) :
                  theorem StrictAnti.comp_strictAntiOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : StrictAnti g) (hf : StrictAntiOn f s) :
                  theorem StrictAnti.comp_strictMonoOn {α : Type u} {β : Type v} {γ : Type w} [Preorder α] [Preorder β] [Preorder γ] {g : βγ} {f : αβ} {s : Set α} (hg : StrictAnti g) (hf : StrictMonoOn f s) :
                  theorem List.foldl_monotone {α : Type u} {β : Type v} [Preorder α] {f : αβα} (H : ∀ (b : β), Monotone fun (a : α) => f a b) (l : List β) :
                  Monotone fun (a : α) => List.foldl f a l
                  theorem List.foldr_monotone {α : Type u} {β : Type v} [Preorder β] {f : αββ} (H : ∀ (a : α), Monotone (f a)) (l : List α) :
                  Monotone fun (b : β) => List.foldr f b l
                  theorem List.foldl_strictMono {α : Type u} {β : Type v} [Preorder α] {f : αβα} (H : ∀ (b : β), StrictMono fun (a : α) => f a b) (l : List β) :
                  StrictMono fun (a : α) => List.foldl f a l
                  theorem List.foldr_strictMono {α : Type u} {β : Type v} [Preorder β] {f : αββ} (H : ∀ (a : α), StrictMono (f a)) (l : List α) :
                  StrictMono fun (b : β) => List.foldr f b l

                  Monotonicity in linear orders #

                  theorem Monotone.reflect_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : Monotone f) {a : α} {b : α} (h : f a < f b) :
                  a < b
                  theorem Antitone.reflect_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : Antitone f) {a : α} {b : α} (h : f a < f b) :
                  b < a
                  theorem MonotoneOn.reflect_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : MonotoneOn f s) {a : α} {b : α} (ha : a s) (hb : b s) (h : f a < f b) :
                  a < b
                  theorem AntitoneOn.reflect_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : AntitoneOn f s) {a : α} {b : α} (ha : a s) (hb : b s) (h : f a < f b) :
                  b < a
                  theorem StrictMonoOn.le_iff_le {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictMonoOn f s) {a : α} {b : α} (ha : a s) (hb : b s) :
                  f a f b a b
                  theorem StrictAntiOn.le_iff_le {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictAntiOn f s) {a : α} {b : α} (ha : a s) (hb : b s) :
                  f a f b b a
                  theorem StrictMonoOn.eq_iff_eq {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictMonoOn f s) {a : α} {b : α} (ha : a s) (hb : b s) :
                  f a = f b a = b
                  theorem StrictAntiOn.eq_iff_eq {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictAntiOn f s) {a : α} {b : α} (ha : a s) (hb : b s) :
                  f a = f b b = a
                  theorem StrictMonoOn.lt_iff_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictMonoOn f s) {a : α} {b : α} (ha : a s) (hb : b s) :
                  f a < f b a < b
                  theorem StrictAntiOn.lt_iff_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictAntiOn f s) {a : α} {b : α} (ha : a s) (hb : b s) :
                  f a < f b b < a
                  theorem StrictMono.le_iff_le {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) {a : α} {b : α} :
                  f a f b a b
                  theorem StrictAnti.le_iff_le {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) {a : α} {b : α} :
                  f a f b b a
                  theorem StrictMono.lt_iff_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) {a : α} {b : α} :
                  f a < f b a < b
                  theorem StrictAnti.lt_iff_lt {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) {a : α} {b : α} :
                  f a < f b b < a
                  theorem StrictMonoOn.compares {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictMonoOn f s) {a : α} {b : α} (ha : a s) (hb : b s) {o : Ordering} :
                  theorem StrictAntiOn.compares {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} {s : Set α} (hf : StrictAntiOn f s) {a : α} {b : α} (ha : a s) (hb : b s) {o : Ordering} :
                  theorem StrictMono.compares {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) {a : α} {b : α} {o : Ordering} :
                  theorem StrictAnti.compares {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) {a : α} {b : α} {o : Ordering} :
                  theorem StrictMono.injective {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) :
                  theorem StrictAnti.injective {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) :
                  theorem StrictMono.maximal_of_maximal_image {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) {a : α} (hmax : ∀ (p : β), p f a) (x : α) :
                  x a
                  theorem StrictMono.minimal_of_minimal_image {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictMono f) {a : α} (hmin : ∀ (p : β), f a p) (x : α) :
                  a x
                  theorem StrictAnti.minimal_of_maximal_image {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) {a : α} (hmax : ∀ (p : β), p f a) (x : α) :
                  a x
                  theorem StrictAnti.maximal_of_minimal_image {α : Type u} {β : Type v} [LinearOrder α] [Preorder β] {f : αβ} (hf : StrictAnti f) {a : α} (hmin : ∀ (p : β), f a p) (x : α) :
                  x a
                  theorem Monotone.strictMono_iff_injective {α : Type u} {β : Type v} [LinearOrder α] [PartialOrder β] {f : αβ} (hf : Monotone f) :
                  theorem Antitone.strictAnti_iff_injective {α : Type u} {β : Type v} [LinearOrder α] [PartialOrder β] {f : αβ} (hf : Antitone f) :
                  theorem not_monotone_not_antitone_iff_exists_le_le {α : Type u} {β : Type v} [LinearOrder α] [LinearOrder β] {f : αβ} :
                  ¬Monotone f ¬Antitone f ∃ (a : α), ∃ (b : α), ∃ (c : α), a b b c (f a < f b f c < f b f b < f a f b < f c)

                  A function between linear orders which is neither monotone nor antitone makes a dent upright or downright.

                  theorem not_monotone_not_antitone_iff_exists_lt_lt {α : Type u} {β : Type v} [LinearOrder α] [LinearOrder β] {f : αβ} :
                  ¬Monotone f ¬Antitone f ∃ (a : α), ∃ (b : α), ∃ (c : α), a < b b < c (f a < f b f c < f b f b < f a f b < f c)

                  A function between linear orders which is neither monotone nor antitone makes a dent upright or downright.

                  Strictly monotone functions and cmp #

                  theorem StrictMonoOn.cmp_map_eq {α : Type u} {β : Type v} [LinearOrder α] [LinearOrder β] {f : αβ} {s : Set α} {x : α} {y : α} (hf : StrictMonoOn f s) (hx : x s) (hy : y s) :
                  cmp (f x) (f y) = cmp x y
                  theorem StrictMono.cmp_map_eq {α : Type u} {β : Type v} [LinearOrder α] [LinearOrder β] {f : αβ} (hf : StrictMono f) (x : α) (y : α) :
                  cmp (f x) (f y) = cmp x y
                  theorem StrictAntiOn.cmp_map_eq {α : Type u} {β : Type v} [LinearOrder α] [LinearOrder β] {f : αβ} {s : Set α} {x : α} {y : α} (hf : StrictAntiOn f s) (hx : x s) (hy : y s) :
                  cmp (f x) (f y) = cmp y x
                  theorem StrictAnti.cmp_map_eq {α : Type u} {β : Type v} [LinearOrder α] [LinearOrder β] {f : αβ} (hf : StrictAnti f) (x : α) (y : α) :
                  cmp (f x) (f y) = cmp y x

                  Monotonicity in and #

                  theorem Nat.rel_of_forall_rel_succ_of_le_of_lt {β : Type v} (r : ββProp) [IsTrans β r] {f : β} {a : } (h : ∀ (n : ), a nr (f n) (f (n + 1))) ⦃b : ⦃c : (hab : a b) (hbc : b < c) :
                  r (f b) (f c)
                  theorem Nat.rel_of_forall_rel_succ_of_le_of_le {β : Type v} (r : ββProp) [IsRefl β r] [IsTrans β r] {f : β} {a : } (h : ∀ (n : ), a nr (f n) (f (n + 1))) ⦃b : ⦃c : (hab : a b) (hbc : b c) :
                  r (f b) (f c)
                  theorem Nat.rel_of_forall_rel_succ_of_lt {β : Type v} (r : ββProp) [IsTrans β r] {f : β} (h : ∀ (n : ), r (f n) (f (n + 1))) ⦃a : ⦃b : (hab : a < b) :
                  r (f a) (f b)
                  theorem Nat.rel_of_forall_rel_succ_of_le {β : Type v} (r : ββProp) [IsRefl β r] [IsTrans β r] {f : β} (h : ∀ (n : ), r (f n) (f (n + 1))) ⦃a : ⦃b : (hab : a b) :
                  r (f a) (f b)
                  theorem monotone_nat_of_le_succ {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f n f (n + 1)) :
                  theorem antitone_nat_of_succ_le {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f (n + 1) f n) :
                  theorem strictMono_nat_of_lt_succ {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f n < f (n + 1)) :
                  theorem strictAnti_nat_of_succ_lt {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f (n + 1) < f n) :
                  theorem Nat.exists_strictMono' {α : Type u} [Preorder α] [NoMaxOrder α] (a : α) :
                  ∃ (f : α), StrictMono f f 0 = a

                  If α is a preorder with no maximal elements, then there exists a strictly monotone function ℕ → α with any prescribed value of f 0.

                  theorem Nat.exists_strictAnti' {α : Type u} [Preorder α] [NoMinOrder α] (a : α) :
                  ∃ (f : α), StrictAnti f f 0 = a

                  If α is a preorder with no maximal elements, then there exists a strictly antitone function ℕ → α with any prescribed value of f 0.

                  theorem Nat.exists_strictMono (α : Type u) [Preorder α] [Nonempty α] [NoMaxOrder α] :
                  ∃ (f : α), StrictMono f

                  If α is a nonempty preorder with no maximal elements, then there exists a strictly monotone function ℕ → α.

                  theorem Nat.exists_strictAnti (α : Type u) [Preorder α] [Nonempty α] [NoMinOrder α] :
                  ∃ (f : α), StrictAnti f

                  If α is a nonempty preorder with no minimal elements, then there exists a strictly antitone function ℕ → α.

                  theorem Int.rel_of_forall_rel_succ_of_lt {β : Type v} (r : ββProp) [IsTrans β r] {f : β} (h : ∀ (n : ), r (f n) (f (n + 1))) ⦃a : ⦃b : (hab : a < b) :
                  r (f a) (f b)
                  theorem Int.rel_of_forall_rel_succ_of_le {β : Type v} (r : ββProp) [IsRefl β r] [IsTrans β r] {f : β} (h : ∀ (n : ), r (f n) (f (n + 1))) ⦃a : ⦃b : (hab : a b) :
                  r (f a) (f b)
                  theorem monotone_int_of_le_succ {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f n f (n + 1)) :
                  theorem antitone_int_of_succ_le {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f (n + 1) f n) :
                  theorem strictMono_int_of_lt_succ {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f n < f (n + 1)) :
                  theorem strictAnti_int_of_succ_lt {α : Type u} [Preorder α] {f : α} (hf : ∀ (n : ), f (n + 1) < f n) :
                  theorem Int.exists_strictMono (α : Type u) [Preorder α] [Nonempty α] [NoMinOrder α] [NoMaxOrder α] :
                  ∃ (f : α), StrictMono f

                  If α is a nonempty preorder with no minimal or maximal elements, then there exists a strictly monotone function f : ℤ → α.

                  theorem Int.exists_strictAnti (α : Type u) [Preorder α] [Nonempty α] [NoMinOrder α] [NoMaxOrder α] :
                  ∃ (f : α), StrictAnti f

                  If α is a nonempty preorder with no minimal or maximal elements, then there exists a strictly antitone function f : ℤ → α.

                  theorem Monotone.ne_of_lt_of_lt_nat {α : Type u} [Preorder α] {f : α} (hf : Monotone f) (n : ) {x : α} (h1 : f n < x) (h2 : x < f (n + 1)) (a : ) :
                  f a x

                  If f is a monotone function from to a preorder such that x lies between f n and f (n + 1), then x doesn't lie in the range of f.

                  theorem Antitone.ne_of_lt_of_lt_nat {α : Type u} [Preorder α] {f : α} (hf : Antitone f) (n : ) {x : α} (h1 : f (n + 1) < x) (h2 : x < f n) (a : ) :
                  f a x

                  If f is an antitone function from to a preorder such that x lies between f (n + 1) and f n, then x doesn't lie in the range of f.

                  theorem Monotone.ne_of_lt_of_lt_int {α : Type u} [Preorder α] {f : α} (hf : Monotone f) (n : ) {x : α} (h1 : f n < x) (h2 : x < f (n + 1)) (a : ) :
                  f a x

                  If f is a monotone function from to a preorder and x lies between f n and f (n + 1), then x doesn't lie in the range of f.

                  theorem Antitone.ne_of_lt_of_lt_int {α : Type u} [Preorder α] {f : α} (hf : Antitone f) (n : ) {x : α} (h1 : f (n + 1) < x) (h2 : x < f n) (a : ) :
                  f a x

                  If f is an antitone function from to a preorder and x lies between f (n + 1) and f n, then x doesn't lie in the range of f.

                  theorem StrictMono.id_le {φ : } (h : StrictMono φ) (n : ) :
                  n φ n
                  theorem Subtype.mono_coe {α : Type u} [Preorder α] (t : Set α) :
                  Monotone Subtype.val
                  theorem Subtype.strictMono_coe {α : Type u} [Preorder α] (t : Set α) :
                  StrictMono Subtype.val
                  theorem monotone_fst {α : Type u} {β : Type v} [Preorder α] [Preorder β] :
                  Monotone Prod.fst
                  theorem monotone_snd {α : Type u} {β : Type v} [Preorder α] [Preorder β] :
                  Monotone Prod.snd
                  theorem Monotone.prod_map {α : Type u} {β : Type v} {γ : Type w} {δ : Type u_2} [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] {f : αγ} {g : βδ} (hf : Monotone f) (hg : Monotone g) :
                  theorem Antitone.prod_map {α : Type u} {β : Type v} {γ : Type w} {δ : Type u_2} [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] {f : αγ} {g : βδ} (hf : Antitone f) (hg : Antitone g) :
                  theorem StrictMono.prod_map {α : Type u} {β : Type v} {γ : Type w} {δ : Type u_2} [PartialOrder α] [PartialOrder β] [Preorder γ] [Preorder δ] {f : αγ} {g : βδ} (hf : StrictMono f) (hg : StrictMono g) :
                  theorem StrictAnti.prod_map {α : Type u} {β : Type v} {γ : Type w} {δ : Type u_2} [PartialOrder α] [PartialOrder β] [Preorder γ] [Preorder δ] {f : αγ} {g : βδ} (hf : StrictAnti f) (hg : StrictAnti g) :

                  Pi types #

                  theorem Function.update_mono {ι : Type u_1} {π : ιType u_3} [DecidableEq ι] [(i : ι) → Preorder (π i)] {f : (i : ι) → π i} {i : ι} :
                  theorem Function.update_strictMono {ι : Type u_1} {π : ιType u_3} [DecidableEq ι] [(i : ι) → Preorder (π i)] {f : (i : ι) → π i} {i : ι} :
                  theorem Function.const_mono {α : Type u} {β : Type v} [Preorder α] :