Documentation

Mathlib.Data.Set.Intervals.Basic

Intervals #

In any preorder α, we define intervals (which on each side can be either infinite, open, or closed) using the following naming conventions:

Each interval has the name I + letter for left side + letter for right side. For instance, Ioc a b denotes the interval (a, b].

This file contains these definitions, and basic facts on inclusion, intersection, difference of intervals (where the precise statements may depend on the properties of the order, in particular for some statements it should be LinearOrder or DenselyOrdered).

TODO: This is just the beginning; a lot of rules are missing

def Set.Ioo {α : Type u_1} [Preorder α] (a : α) (b : α) :
Set α

Left-open right-open interval

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

    Left-closed right-open interval

    Equations
    Instances For
      def Set.Iio {α : Type u_1} [Preorder α] (a : α) :
      Set α

      Left-infinite right-open interval

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

        Left-closed right-closed interval

        Equations
        Instances For
          def Set.Iic {α : Type u_1} [Preorder α] (b : α) :
          Set α

          Left-infinite right-closed interval

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

            Left-open right-closed interval

            Equations
            Instances For
              def Set.Ici {α : Type u_1} [Preorder α] (a : α) :
              Set α

              Left-closed right-infinite interval

              Equations
              Instances For
                def Set.Ioi {α : Type u_1} [Preorder α] (a : α) :
                Set α

                Left-open right-infinite interval

                Equations
                Instances For
                  theorem Set.Ioo_def {α : Type u_1} [Preorder α] (a : α) (b : α) :
                  {x : α | a < x x < b} = Set.Ioo a b
                  theorem Set.Ico_def {α : Type u_1} [Preorder α] (a : α) (b : α) :
                  {x : α | a x x < b} = Set.Ico a b
                  theorem Set.Iio_def {α : Type u_1} [Preorder α] (a : α) :
                  {x : α | x < a} = Set.Iio a
                  theorem Set.Icc_def {α : Type u_1} [Preorder α] (a : α) (b : α) :
                  {x : α | a x x b} = Set.Icc a b
                  theorem Set.Iic_def {α : Type u_1} [Preorder α] (b : α) :
                  {x : α | x b} = Set.Iic b
                  theorem Set.Ioc_def {α : Type u_1} [Preorder α] (a : α) (b : α) :
                  {x : α | a < x x b} = Set.Ioc a b
                  theorem Set.Ici_def {α : Type u_1} [Preorder α] (a : α) :
                  {x : α | a x} = Set.Ici a
                  theorem Set.Ioi_def {α : Type u_1} [Preorder α] (a : α) :
                  {x : α | a < x} = Set.Ioi a
                  @[simp]
                  theorem Set.mem_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} :
                  x Set.Ioo a b a < x x < b
                  @[simp]
                  theorem Set.mem_Ico {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} :
                  x Set.Ico a b a x x < b
                  @[simp]
                  theorem Set.mem_Iio {α : Type u_1} [Preorder α] {b : α} {x : α} :
                  x Set.Iio b x < b
                  @[simp]
                  theorem Set.mem_Icc {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} :
                  x Set.Icc a b a x x b
                  @[simp]
                  theorem Set.mem_Iic {α : Type u_1} [Preorder α] {b : α} {x : α} :
                  x Set.Iic b x b
                  @[simp]
                  theorem Set.mem_Ioc {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} :
                  x Set.Ioc a b a < x x b
                  @[simp]
                  theorem Set.mem_Ici {α : Type u_1} [Preorder α] {a : α} {x : α} :
                  x Set.Ici a a x
                  @[simp]
                  theorem Set.mem_Ioi {α : Type u_1} [Preorder α] {a : α} {x : α} :
                  x Set.Ioi a a < x
                  instance Set.decidableMemIoo {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} [Decidable (a < x x < b)] :
                  Equations
                  • Set.decidableMemIoo = inst
                  instance Set.decidableMemIco {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} [Decidable (a x x < b)] :
                  Equations
                  • Set.decidableMemIco = inst
                  instance Set.decidableMemIio {α : Type u_1} [Preorder α] {b : α} {x : α} [Decidable (x < b)] :
                  Equations
                  • Set.decidableMemIio = inst
                  instance Set.decidableMemIcc {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} [Decidable (a x x b)] :
                  Equations
                  • Set.decidableMemIcc = inst
                  instance Set.decidableMemIic {α : Type u_1} [Preorder α] {b : α} {x : α} [Decidable (x b)] :
                  Equations
                  • Set.decidableMemIic = inst
                  instance Set.decidableMemIoc {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} [Decidable (a < x x b)] :
                  Equations
                  • Set.decidableMemIoc = inst
                  instance Set.decidableMemIci {α : Type u_1} [Preorder α] {a : α} {x : α} [Decidable (a x)] :
                  Equations
                  • Set.decidableMemIci = inst
                  instance Set.decidableMemIoi {α : Type u_1} [Preorder α] {a : α} {x : α} [Decidable (a < x)] :
                  Equations
                  • Set.decidableMemIoi = inst
                  theorem Set.left_mem_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.left_mem_Ico {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  a Set.Ico a b a < b
                  theorem Set.left_mem_Icc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  a Set.Icc a b a b
                  theorem Set.left_mem_Ioc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.left_mem_Ici {α : Type u_1} [Preorder α] {a : α} :
                  theorem Set.right_mem_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.right_mem_Ico {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.right_mem_Icc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  b Set.Icc a b a b
                  theorem Set.right_mem_Ioc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  b Set.Ioc a b a < b
                  theorem Set.right_mem_Iic {α : Type u_1} [Preorder α] {a : α} :
                  @[simp]
                  theorem Set.dual_Ici {α : Type u_1} [Preorder α] {a : α} :
                  Set.Ici (OrderDual.toDual a) = OrderDual.ofDual ⁻¹' Set.Iic a
                  @[simp]
                  theorem Set.dual_Iic {α : Type u_1} [Preorder α] {a : α} :
                  Set.Iic (OrderDual.toDual a) = OrderDual.ofDual ⁻¹' Set.Ici a
                  @[simp]
                  theorem Set.dual_Ioi {α : Type u_1} [Preorder α] {a : α} :
                  Set.Ioi (OrderDual.toDual a) = OrderDual.ofDual ⁻¹' Set.Iio a
                  @[simp]
                  theorem Set.dual_Iio {α : Type u_1} [Preorder α] {a : α} :
                  Set.Iio (OrderDual.toDual a) = OrderDual.ofDual ⁻¹' Set.Ioi a
                  @[simp]
                  theorem Set.dual_Icc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  Set.Icc (OrderDual.toDual a) (OrderDual.toDual b) = OrderDual.ofDual ⁻¹' Set.Icc b a
                  @[simp]
                  theorem Set.dual_Ioc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  Set.Ioc (OrderDual.toDual a) (OrderDual.toDual b) = OrderDual.ofDual ⁻¹' Set.Ico b a
                  @[simp]
                  theorem Set.dual_Ico {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  Set.Ico (OrderDual.toDual a) (OrderDual.toDual b) = OrderDual.ofDual ⁻¹' Set.Ioc b a
                  @[simp]
                  theorem Set.dual_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  Set.Ioo (OrderDual.toDual a) (OrderDual.toDual b) = OrderDual.ofDual ⁻¹' Set.Ioo b a
                  @[simp]
                  theorem Set.nonempty_Icc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.nonempty_Ico {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.nonempty_Ioc {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.nonempty_Ici {α : Type u_1} [Preorder α] {a : α} :
                  @[simp]
                  theorem Set.nonempty_Iic {α : Type u_1} [Preorder α] {a : α} :
                  @[simp]
                  theorem Set.nonempty_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} [DenselyOrdered α] :
                  @[simp]
                  theorem Set.nonempty_Ioi {α : Type u_1} [Preorder α] {a : α} [NoMaxOrder α] :
                  @[simp]
                  theorem Set.nonempty_Iio {α : Type u_1} [Preorder α] {a : α} [NoMinOrder α] :
                  theorem Set.nonempty_Icc_subtype {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a b) :
                  Nonempty (Set.Icc a b)
                  theorem Set.nonempty_Ico_subtype {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a < b) :
                  Nonempty (Set.Ico a b)
                  theorem Set.nonempty_Ioc_subtype {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a < b) :
                  Nonempty (Set.Ioc a b)
                  instance Set.nonempty_Ici_subtype {α : Type u_1} [Preorder α] {a : α} :

                  An interval Ici a is nonempty.

                  Equations
                  • =
                  instance Set.nonempty_Iic_subtype {α : Type u_1} [Preorder α] {a : α} :

                  An interval Iic a is nonempty.

                  Equations
                  • =
                  theorem Set.nonempty_Ioo_subtype {α : Type u_1} [Preorder α] {a : α} {b : α} [DenselyOrdered α] (h : a < b) :
                  Nonempty (Set.Ioo a b)
                  instance Set.nonempty_Ioi_subtype {α : Type u_1} [Preorder α] {a : α} [NoMaxOrder α] :

                  In an order without maximal elements, the intervals Ioi are nonempty.

                  Equations
                  • =
                  instance Set.nonempty_Iio_subtype {α : Type u_1} [Preorder α] {a : α} [NoMinOrder α] :

                  In an order without minimal elements, the intervals Iio are nonempty.

                  Equations
                  • =
                  Equations
                  • =
                  Equations
                  • =
                  Equations
                  • =
                  Equations
                  • =
                  @[simp]
                  theorem Set.Icc_eq_empty {α : Type u_1} [Preorder α] {a : α} {b : α} (h : ¬a b) :
                  @[simp]
                  theorem Set.Ico_eq_empty {α : Type u_1} [Preorder α] {a : α} {b : α} (h : ¬a < b) :
                  @[simp]
                  theorem Set.Ioc_eq_empty {α : Type u_1} [Preorder α] {a : α} {b : α} (h : ¬a < b) :
                  @[simp]
                  theorem Set.Ioo_eq_empty {α : Type u_1} [Preorder α] {a : α} {b : α} (h : ¬a < b) :
                  @[simp]
                  theorem Set.Icc_eq_empty_of_lt {α : Type u_1} [Preorder α] {a : α} {b : α} (h : b < a) :
                  @[simp]
                  theorem Set.Ico_eq_empty_of_le {α : Type u_1} [Preorder α] {a : α} {b : α} (h : b a) :
                  @[simp]
                  theorem Set.Ioc_eq_empty_of_le {α : Type u_1} [Preorder α] {a : α} {b : α} (h : b a) :
                  @[simp]
                  theorem Set.Ioo_eq_empty_of_le {α : Type u_1} [Preorder α] {a : α} {b : α} (h : b a) :
                  theorem Set.Ico_self {α : Type u_1} [Preorder α] (a : α) :
                  theorem Set.Ioc_self {α : Type u_1} [Preorder α] (a : α) :
                  theorem Set.Ioo_self {α : Type u_1} [Preorder α] (a : α) :
                  theorem Set.Ici_subset_Ici {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Iic_subset_Iic {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ici_subset_Ioi {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Iic_subset_Iio {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioo_subset_Ioo {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₂ a₁) (h₂ : b₁ b₂) :
                  Set.Ioo a₁ b₁ Set.Ioo a₂ b₂
                  theorem Set.Ioo_subset_Ioo_left {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b : α} (h : a₁ a₂) :
                  Set.Ioo a₂ b Set.Ioo a₁ b
                  theorem Set.Ioo_subset_Ioo_right {α : Type u_1} [Preorder α] {a : α} {b₁ : α} {b₂ : α} (h : b₁ b₂) :
                  Set.Ioo a b₁ Set.Ioo a b₂
                  theorem Set.Ico_subset_Ico {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₂ a₁) (h₂ : b₁ b₂) :
                  Set.Ico a₁ b₁ Set.Ico a₂ b₂
                  theorem Set.Ico_subset_Ico_left {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b : α} (h : a₁ a₂) :
                  Set.Ico a₂ b Set.Ico a₁ b
                  theorem Set.Ico_subset_Ico_right {α : Type u_1} [Preorder α] {a : α} {b₁ : α} {b₂ : α} (h : b₁ b₂) :
                  Set.Ico a b₁ Set.Ico a b₂
                  theorem Set.Icc_subset_Icc {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₂ a₁) (h₂ : b₁ b₂) :
                  Set.Icc a₁ b₁ Set.Icc a₂ b₂
                  theorem Set.Icc_subset_Icc_left {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b : α} (h : a₁ a₂) :
                  Set.Icc a₂ b Set.Icc a₁ b
                  theorem Set.Icc_subset_Icc_right {α : Type u_1} [Preorder α] {a : α} {b₁ : α} {b₂ : α} (h : b₁ b₂) :
                  Set.Icc a b₁ Set.Icc a b₂
                  theorem Set.Icc_subset_Ioo {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (ha : a₂ < a₁) (hb : b₁ < b₂) :
                  Set.Icc a₁ b₁ Set.Ioo a₂ b₂
                  theorem Set.Icc_subset_Ici_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Icc_subset_Iic_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioc_subset_Iic_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioc_subset_Ioc {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₂ a₁) (h₂ : b₁ b₂) :
                  Set.Ioc a₁ b₁ Set.Ioc a₂ b₂
                  theorem Set.Ioc_subset_Ioc_left {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b : α} (h : a₁ a₂) :
                  Set.Ioc a₂ b Set.Ioc a₁ b
                  theorem Set.Ioc_subset_Ioc_right {α : Type u_1} [Preorder α] {a : α} {b₁ : α} {b₂ : α} (h : b₁ b₂) :
                  Set.Ioc a b₁ Set.Ioc a b₂
                  theorem Set.Ico_subset_Ioo_left {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b : α} (h₁ : a₁ < a₂) :
                  Set.Ico a₂ b Set.Ioo a₁ b
                  theorem Set.Ioc_subset_Ioo_right {α : Type u_1} [Preorder α] {a : α} {b₁ : α} {b₂ : α} (h : b₁ < b₂) :
                  Set.Ioc a b₁ Set.Ioo a b₂
                  theorem Set.Icc_subset_Ico_right {α : Type u_1} [Preorder α] {a : α} {b₁ : α} {b₂ : α} (h₁ : b₁ < b₂) :
                  Set.Icc a b₁ Set.Ico a b₂
                  theorem Set.Ioo_subset_Ico_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioo_subset_Ioc_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ico_subset_Icc_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioc_subset_Icc_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioo_subset_Icc_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ico_subset_Iio_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioo_subset_Iio_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioc_subset_Ioi_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioo_subset_Ioi_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioi_subset_Ici_self {α : Type u_1} [Preorder α] {a : α} :
                  theorem Set.Iio_subset_Iic_self {α : Type u_1} [Preorder α] {a : α} :
                  theorem Set.Ico_subset_Ici_self {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioi_ssubset_Ici_self {α : Type u_1} [Preorder α] {a : α} :
                  theorem Set.Iio_ssubset_Iic_self {α : Type u_1} [Preorder α] {a : α} :
                  theorem Set.Icc_subset_Icc_iff {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Icc a₂ b₂ a₂ a₁ b₁ b₂
                  theorem Set.Icc_subset_Ioo_iff {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Ioo a₂ b₂ a₂ < a₁ b₁ < b₂
                  theorem Set.Icc_subset_Ico_iff {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Ico a₂ b₂ a₂ a₁ b₁ < b₂
                  theorem Set.Icc_subset_Ioc_iff {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Ioc a₂ b₂ a₂ < a₁ b₁ b₂
                  theorem Set.Icc_subset_Iio_iff {α : Type u_1} [Preorder α] {a₁ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Iio b₂ b₁ < b₂
                  theorem Set.Icc_subset_Ioi_iff {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Ioi a₂ a₂ < a₁
                  theorem Set.Icc_subset_Iic_iff {α : Type u_1} [Preorder α] {a₁ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Iic b₂ b₁ b₂
                  theorem Set.Icc_subset_Ici_iff {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} (h₁ : a₁ b₁) :
                  Set.Icc a₁ b₁ Set.Ici a₂ a₂ a₁
                  theorem Set.Icc_ssubset_Icc_left {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (hI : a₂ b₂) (ha : a₂ < a₁) (hb : b₁ b₂) :
                  Set.Icc a₁ b₁ Set.Icc a₂ b₂
                  theorem Set.Icc_ssubset_Icc_right {α : Type u_1} [Preorder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (hI : a₂ b₂) (ha : a₂ a₁) (hb : b₁ < b₂) :
                  Set.Icc a₁ b₁ Set.Icc a₂ b₂
                  theorem Set.Ioi_subset_Ioi {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a b) :

                  If a ≤ b, then (b, +∞) ⊆ (a, +∞). In preorders, this is just an implication. If you need the equivalence in linear orders, use Ioi_subset_Ioi_iff.

                  theorem Set.Ioi_subset_Ici {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a b) :

                  If a ≤ b, then (b, +∞) ⊆ [a, +∞). In preorders, this is just an implication. If you need the equivalence in dense linear orders, use Ioi_subset_Ici_iff.

                  theorem Set.Iio_subset_Iio {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a b) :

                  If a ≤ b, then (-∞, a) ⊆ (-∞, b). In preorders, this is just an implication. If you need the equivalence in linear orders, use Iio_subset_Iio_iff.

                  theorem Set.Iio_subset_Iic {α : Type u_1} [Preorder α] {a : α} {b : α} (h : a b) :

                  If a ≤ b, then (-∞, a) ⊆ (-∞, b]. In preorders, this is just an implication. If you need the equivalence in dense linear orders, use Iio_subset_Iic_iff.

                  theorem Set.Ici_inter_Iic {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ici_inter_Iio {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioi_inter_Iic {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ioi_inter_Iio {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Iic_inter_Ici {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Iio_inter_Ici {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Iic_inter_Ioi {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Iio_inter_Ioi {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.mem_Icc_of_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} (h : x Set.Ioo a b) :
                  x Set.Icc a b
                  theorem Set.mem_Ico_of_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} (h : x Set.Ioo a b) :
                  x Set.Ico a b
                  theorem Set.mem_Ioc_of_Ioo {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} (h : x Set.Ioo a b) :
                  x Set.Ioc a b
                  theorem Set.mem_Icc_of_Ico {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} (h : x Set.Ico a b) :
                  x Set.Icc a b
                  theorem Set.mem_Icc_of_Ioc {α : Type u_1} [Preorder α] {a : α} {b : α} {x : α} (h : x Set.Ioc a b) :
                  x Set.Icc a b
                  theorem Set.mem_Ici_of_Ioi {α : Type u_1} [Preorder α] {a : α} {x : α} (h : x Set.Ioi a) :
                  theorem Set.mem_Iic_of_Iio {α : Type u_1} [Preorder α] {a : α} {x : α} (h : x Set.Iio a) :
                  theorem Set.Icc_eq_empty_iff {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  theorem Set.Ico_eq_empty_iff {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  Set.Ico a b = ¬a < b
                  theorem Set.Ioc_eq_empty_iff {α : Type u_1} [Preorder α] {a : α} {b : α} :
                  Set.Ioc a b = ¬a < b
                  theorem Set.Ioo_eq_empty_iff {α : Type u_1} [Preorder α] {a : α} {b : α} [DenselyOrdered α] :
                  Set.Ioo a b = ¬a < b
                  theorem IsTop.Iic_eq {α : Type u_1} [Preorder α] {a : α} (h : IsTop a) :
                  Set.Iic a = Set.univ
                  theorem IsBot.Ici_eq {α : Type u_1} [Preorder α] {a : α} (h : IsBot a) :
                  Set.Ici a = Set.univ
                  theorem IsMax.Ioi_eq {α : Type u_1} [Preorder α] {a : α} (h : IsMax a) :
                  theorem IsMin.Iio_eq {α : Type u_1} [Preorder α] {a : α} (h : IsMin a) :
                  theorem Set.Iic_inter_Ioc_of_le {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (h : a c) :
                  theorem Set.not_mem_Icc_of_lt {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (ha : c < a) :
                  cSet.Icc a b
                  theorem Set.not_mem_Icc_of_gt {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (hb : b < c) :
                  cSet.Icc a b
                  theorem Set.not_mem_Ico_of_lt {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (ha : c < a) :
                  cSet.Ico a b
                  theorem Set.not_mem_Ioc_of_gt {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (hb : b < c) :
                  cSet.Ioc a b
                  theorem Set.not_mem_Ioi_self {α : Type u_1} [Preorder α] {a : α} :
                  aSet.Ioi a
                  theorem Set.not_mem_Iio_self {α : Type u_1} [Preorder α] {b : α} :
                  bSet.Iio b
                  theorem Set.not_mem_Ioc_of_le {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (ha : c a) :
                  cSet.Ioc a b
                  theorem Set.not_mem_Ico_of_ge {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (hb : b c) :
                  cSet.Ico a b
                  theorem Set.not_mem_Ioo_of_le {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (ha : c a) :
                  cSet.Ioo a b
                  theorem Set.not_mem_Ioo_of_ge {α : Type u_1} [Preorder α] {a : α} {b : α} {c : α} (hb : b c) :
                  cSet.Ioo a b
                  @[simp]
                  theorem Set.Icc_self {α : Type u_1} [PartialOrder α] (a : α) :
                  Set.Icc a a = {a}
                  instance Set.instIccUnique {α : Type u_1} [PartialOrder α] {a : α} :
                  Unique (Set.Icc a a)
                  Equations
                  • Set.instIccUnique = { toInhabited := { default := { val := a, property := } }, uniq := }
                  @[simp]
                  theorem Set.Icc_eq_singleton_iff {α : Type u_1} [PartialOrder α] {a : α} {b : α} {c : α} :
                  Set.Icc a b = {c} a = c b = c
                  theorem Set.subsingleton_Icc_of_ge {α : Type u_1} [PartialOrder α] {a : α} {b : α} (hba : b a) :
                  @[simp]
                  theorem Set.subsingleton_Icc_iff {α : Type u_3} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Icc_diff_left {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  Set.Icc a b \ {a} = Set.Ioc a b
                  @[simp]
                  theorem Set.Icc_diff_right {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  Set.Icc a b \ {b} = Set.Ico a b
                  @[simp]
                  theorem Set.Ico_diff_left {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  Set.Ico a b \ {a} = Set.Ioo a b
                  @[simp]
                  theorem Set.Ioc_diff_right {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  Set.Ioc a b \ {b} = Set.Ioo a b
                  @[simp]
                  theorem Set.Icc_diff_both {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  Set.Icc a b \ {a, b} = Set.Ioo a b
                  @[simp]
                  theorem Set.Ici_diff_left {α : Type u_1} [PartialOrder α] {a : α} :
                  Set.Ici a \ {a} = Set.Ioi a
                  @[simp]
                  theorem Set.Iic_diff_right {α : Type u_1} [PartialOrder α] {a : α} :
                  Set.Iic a \ {a} = Set.Iio a
                  @[simp]
                  theorem Set.Ico_diff_Ioo_same {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a < b) :
                  Set.Ico a b \ Set.Ioo a b = {a}
                  @[simp]
                  theorem Set.Ioc_diff_Ioo_same {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a < b) :
                  Set.Ioc a b \ Set.Ioo a b = {b}
                  @[simp]
                  theorem Set.Icc_diff_Ico_same {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a b) :
                  Set.Icc a b \ Set.Ico a b = {b}
                  @[simp]
                  theorem Set.Icc_diff_Ioc_same {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a b) :
                  Set.Icc a b \ Set.Ioc a b = {a}
                  @[simp]
                  theorem Set.Icc_diff_Ioo_same {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a b) :
                  Set.Icc a b \ Set.Ioo a b = {a, b}
                  @[simp]
                  theorem Set.Ici_diff_Ioi_same {α : Type u_1} [PartialOrder α] {a : α} :
                  Set.Ici a \ Set.Ioi a = {a}
                  @[simp]
                  theorem Set.Iic_diff_Iio_same {α : Type u_1} [PartialOrder α] {a : α} :
                  Set.Iic a \ Set.Iio a = {a}
                  theorem Set.Ioi_union_left {α : Type u_1} [PartialOrder α] {a : α} :
                  theorem Set.Iio_union_right {α : Type u_1} [PartialOrder α] {a : α} :
                  theorem Set.Ioo_union_left {α : Type u_1} [PartialOrder α] {a : α} {b : α} (hab : a < b) :
                  Set.Ioo a b {a} = Set.Ico a b
                  theorem Set.Ioo_union_right {α : Type u_1} [PartialOrder α] {a : α} {b : α} (hab : a < b) :
                  Set.Ioo a b {b} = Set.Ioc a b
                  theorem Set.Ioo_union_both {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a b) :
                  Set.Ioo a b {a, b} = Set.Icc a b
                  theorem Set.Ioc_union_left {α : Type u_1} [PartialOrder α] {a : α} {b : α} (hab : a b) :
                  Set.Ioc a b {a} = Set.Icc a b
                  theorem Set.Ico_union_right {α : Type u_1} [PartialOrder α] {a : α} {b : α} (hab : a b) :
                  Set.Ico a b {b} = Set.Icc a b
                  @[simp]
                  theorem Set.Ico_insert_right {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a b) :
                  insert b (Set.Ico a b) = Set.Icc a b
                  @[simp]
                  theorem Set.Ioc_insert_left {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a b) :
                  insert a (Set.Ioc a b) = Set.Icc a b
                  @[simp]
                  theorem Set.Ioo_insert_left {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a < b) :
                  insert a (Set.Ioo a b) = Set.Ico a b
                  @[simp]
                  theorem Set.Ioo_insert_right {α : Type u_1} [PartialOrder α] {a : α} {b : α} (h : a < b) :
                  insert b (Set.Ioo a b) = Set.Ioc a b
                  @[simp]
                  theorem Set.Iio_insert {α : Type u_1} [PartialOrder α] {a : α} :
                  @[simp]
                  theorem Set.Ioi_insert {α : Type u_1} [PartialOrder α] {a : α} :
                  theorem Set.mem_Ici_Ioi_of_subset_of_subset {α : Type u_1} [PartialOrder α] {a : α} {s : Set α} (ho : Set.Ioi a s) (hc : s Set.Ici a) :
                  theorem Set.mem_Iic_Iio_of_subset_of_subset {α : Type u_1} [PartialOrder α] {a : α} {s : Set α} (ho : Set.Iio a s) (hc : s Set.Iic a) :
                  theorem Set.mem_Icc_Ico_Ioc_Ioo_of_subset_of_subset {α : Type u_1} [PartialOrder α] {a : α} {b : α} {s : Set α} (ho : Set.Ioo a b s) (hc : s Set.Icc a b) :
                  s {Set.Icc a b, Set.Ico a b, Set.Ioc a b, Set.Ioo a b}
                  theorem Set.eq_left_or_mem_Ioo_of_mem_Ico {α : Type u_1} [PartialOrder α] {a : α} {b : α} {x : α} (hmem : x Set.Ico a b) :
                  x = a x Set.Ioo a b
                  theorem Set.eq_right_or_mem_Ioo_of_mem_Ioc {α : Type u_1} [PartialOrder α] {a : α} {b : α} {x : α} (hmem : x Set.Ioc a b) :
                  x = b x Set.Ioo a b
                  theorem Set.eq_endpoints_or_mem_Ioo_of_mem_Icc {α : Type u_1} [PartialOrder α] {a : α} {b : α} {x : α} (hmem : x Set.Icc a b) :
                  x = a x = b x Set.Ioo a b
                  theorem IsMax.Ici_eq {α : Type u_1} [PartialOrder α] {a : α} (h : IsMax a) :
                  Set.Ici a = {a}
                  theorem IsMin.Iic_eq {α : Type u_1} [PartialOrder α] {a : α} (h : IsMin a) :
                  Set.Iic a = {a}
                  theorem Set.Ici_inj {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  theorem Set.Iic_inj {α : Type u_1} [PartialOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ici_top {α : Type u_1} [PartialOrder α] [OrderTop α] :
                  @[simp]
                  theorem Set.Ioi_top {α : Type u_1} [Preorder α] [OrderTop α] :
                  @[simp]
                  theorem Set.Iic_top {α : Type u_1} [Preorder α] [OrderTop α] :
                  Set.Iic = Set.univ
                  @[simp]
                  theorem Set.Icc_top {α : Type u_1} [Preorder α] [OrderTop α] {a : α} :
                  @[simp]
                  theorem Set.Ioc_top {α : Type u_1} [Preorder α] [OrderTop α] {a : α} :
                  @[simp]
                  theorem Set.Iic_bot {α : Type u_1} [PartialOrder α] [OrderBot α] :
                  @[simp]
                  theorem Set.Iio_bot {α : Type u_1} [Preorder α] [OrderBot α] :
                  @[simp]
                  theorem Set.Ici_bot {α : Type u_1} [Preorder α] [OrderBot α] :
                  Set.Ici = Set.univ
                  @[simp]
                  theorem Set.Icc_bot {α : Type u_1} [Preorder α] [OrderBot α] {a : α} :
                  @[simp]
                  theorem Set.Ico_bot {α : Type u_1} [Preorder α] [OrderBot α] {a : α} :
                  theorem Set.Icc_bot_top {α : Type u_1} [PartialOrder α] [BoundedOrder α] :
                  Set.Icc = Set.univ
                  theorem Set.not_mem_Ici {α : Type u_1} [LinearOrder α] {a : α} {c : α} :
                  cSet.Ici a c < a
                  theorem Set.not_mem_Iic {α : Type u_1} [LinearOrder α] {b : α} {c : α} :
                  cSet.Iic b b < c
                  theorem Set.not_mem_Ioi {α : Type u_1} [LinearOrder α] {a : α} {c : α} :
                  cSet.Ioi a c a
                  theorem Set.not_mem_Iio {α : Type u_1} [LinearOrder α] {b : α} {c : α} :
                  cSet.Iio b b c
                  @[simp]
                  theorem Set.compl_Iic {α : Type u_1} [LinearOrder α] {a : α} :
                  @[simp]
                  theorem Set.compl_Ici {α : Type u_1} [LinearOrder α] {a : α} :
                  @[simp]
                  theorem Set.compl_Iio {α : Type u_1} [LinearOrder α] {a : α} :
                  @[simp]
                  theorem Set.compl_Ioi {α : Type u_1} [LinearOrder α] {a : α} :
                  @[simp]
                  theorem Set.Ici_diff_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ici_diff_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioi_diff_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioi_diff_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iic_diff_Iic {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iio_diff_Iic {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iic_diff_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iio_diff_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  theorem Set.Ioi_inj {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  theorem Set.Iio_inj {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  theorem Set.Ico_subset_Ico_iff {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ < b₁) :
                  Set.Ico a₁ b₁ Set.Ico a₂ b₂ a₂ a₁ b₁ b₂
                  theorem Set.Ioc_subset_Ioc_iff {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h₁ : a₁ < b₁) :
                  Set.Ioc a₁ b₁ Set.Ioc a₂ b₂ b₁ b₂ a₂ a₁
                  theorem Set.Ioo_subset_Ioo_iff {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} [DenselyOrdered α] (h₁ : a₁ < b₁) :
                  Set.Ioo a₁ b₁ Set.Ioo a₂ b₂ a₂ a₁ b₁ b₂
                  theorem Set.Ico_eq_Ico_iff {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h : a₁ < b₁ a₂ < b₂) :
                  Set.Ico a₁ b₁ = Set.Ico a₂ b₂ a₁ = a₂ b₁ = b₂
                  theorem Set.Ici_eq_singleton_iff_isTop {α : Type u_1} [LinearOrder α] {x : α} :
                  Set.Ici x = {x} IsTop x
                  @[simp]
                  theorem Set.Ioi_subset_Ioi_iff {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioi_subset_Ici_iff {α : Type u_1} [LinearOrder α] {a : α} {b : α} [DenselyOrdered α] :
                  @[simp]
                  theorem Set.Iio_subset_Iio_iff {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iio_subset_Iic_iff {α : Type u_1} [LinearOrder α] {a : α} {b : α} [DenselyOrdered α] :

                  Unions of adjacent intervals #

                  Two infinite intervals #

                  theorem Set.Iic_union_Ioi_of_le {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  Set.Iic b Set.Ioi a = Set.univ
                  theorem Set.Iio_union_Ici_of_le {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  Set.Iio b Set.Ici a = Set.univ
                  theorem Set.Iic_union_Ici_of_le {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  Set.Iic b Set.Ici a = Set.univ
                  theorem Set.Iio_union_Ioi_of_lt {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a < b) :
                  Set.Iio b Set.Ioi a = Set.univ
                  @[simp]
                  theorem Set.Iic_union_Ici {α : Type u_1} [LinearOrder α] {a : α} :
                  Set.Iic a Set.Ici a = Set.univ
                  @[simp]
                  theorem Set.Iio_union_Ici {α : Type u_1} [LinearOrder α] {a : α} :
                  Set.Iio a Set.Ici a = Set.univ
                  @[simp]
                  theorem Set.Iic_union_Ioi {α : Type u_1} [LinearOrder α] {a : α} :
                  Set.Iic a Set.Ioi a = Set.univ
                  @[simp]
                  theorem Set.Iio_union_Ioi {α : Type u_1} [LinearOrder α] {a : α} :

                  A finite and an infinite interval #

                  theorem Set.Ioo_union_Ioi' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : c < b) :
                  theorem Set.Ioo_union_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h : c < max a b) :
                  theorem Set.Ioi_subset_Ioo_union_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioo_union_Ici_eq_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a < b) :
                  theorem Set.Ici_subset_Ico_union_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ico_union_Ici_eq_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Ico_union_Ici' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : c b) :
                  theorem Set.Ico_union_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h : c max a b) :
                  theorem Set.Ioi_subset_Ioc_union_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioc_union_Ioi_eq_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Ioc_union_Ioi' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : c b) :
                  theorem Set.Ioc_union_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h : c max a b) :
                  theorem Set.Ici_subset_Icc_union_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Icc_union_Ioi_eq_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Ioi_subset_Ioc_union_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioc_union_Ici_eq_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a < b) :
                  theorem Set.Ici_subset_Icc_union_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Icc_union_Ici_eq_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Icc_union_Ici' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : c b) :
                  theorem Set.Icc_union_Ici {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h : c max a b) :

                  An infinite and a finite interval #

                  theorem Set.Iic_subset_Iio_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iio_union_Icc_eq_Iic {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Iio_subset_Iio_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iio_union_Ico_eq_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Iio_union_Ico' {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h₁ : c b) :
                  theorem Set.Iio_union_Ico {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h : min c d b) :
                  theorem Set.Iic_subset_Iic_union_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iic_union_Ioc_eq_Iic {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Iic_union_Ioc' {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h₁ : c < b) :
                  theorem Set.Iic_union_Ioc {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h : min c d < b) :
                  theorem Set.Iio_subset_Iic_union_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iic_union_Ioo_eq_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a < b) :
                  theorem Set.Iio_union_Ioo' {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h₁ : c < b) :
                  theorem Set.Iio_union_Ioo {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h : min c d < b) :
                  theorem Set.Iic_subset_Iic_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iic_union_Icc_eq_Iic {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a b) :
                  theorem Set.Iic_union_Icc' {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h₁ : c b) :
                  theorem Set.Iic_union_Icc {α : Type u_1} [LinearOrder α] {b : α} {c : α} {d : α} (h : min c d b) :
                  theorem Set.Iio_subset_Iic_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iic_union_Ico_eq_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} (h : a < b) :

                  Two finite intervals, I?o and Ic? #

                  theorem Set.Ioo_subset_Ioo_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ioo_union_Ico_eq_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a < b) (h₂ : b c) :
                  theorem Set.Ico_subset_Ico_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ico_union_Ico_eq_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b c) :
                  theorem Set.Ico_union_Ico' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : c b) (h₂ : a d) :
                  Set.Ico a b Set.Ico c d = Set.Ico (min a c) (max b d)
                  theorem Set.Ico_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : min a b max c d) (h₂ : min c d max a b) :
                  Set.Ico a b Set.Ico c d = Set.Ico (min a c) (max b d)
                  theorem Set.Icc_subset_Ico_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ico_union_Icc_eq_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b c) :
                  theorem Set.Ioc_subset_Ioo_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ioo_union_Icc_eq_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a < b) (h₂ : b c) :

                  Two finite intervals, I?c and Io? #

                  theorem Set.Ioo_subset_Ioc_union_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ioc_union_Ioo_eq_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b < c) :
                  theorem Set.Ico_subset_Icc_union_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Icc_union_Ioo_eq_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b < c) :
                  theorem Set.Icc_subset_Icc_union_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Icc_union_Ioc_eq_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b c) :
                  theorem Set.Ioc_subset_Ioc_union_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ioc_union_Ioc_eq_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b c) :
                  theorem Set.Ioc_union_Ioc' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : c b) (h₂ : a d) :
                  Set.Ioc a b Set.Ioc c d = Set.Ioc (min a c) (max b d)
                  theorem Set.Ioc_union_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : min a b max c d) (h₂ : min c d max a b) :
                  Set.Ioc a b Set.Ioc c d = Set.Ioc (min a c) (max b d)

                  Two finite intervals with a common point #

                  theorem Set.Ioo_subset_Ioc_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ioc_union_Ico_eq_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a < b) (h₂ : b < c) :
                  theorem Set.Ico_subset_Icc_union_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Icc_union_Ico_eq_Ico {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b < c) :
                  theorem Set.Icc_subset_Icc_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Icc_union_Icc_eq_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a b) (h₂ : b c) :
                  theorem Set.Icc_union_Icc' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : c b) (h₂ : a d) :
                  Set.Icc a b Set.Icc c d = Set.Icc (min a c) (max b d)
                  theorem Set.Icc_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : min a b < max c d) (h₂ : min c d < max a b) :
                  Set.Icc a b Set.Icc c d = Set.Icc (min a c) (max b d)

                  We cannot replace < by in the hypotheses. Otherwise for b < a = d < c the l.h.s. is and the r.h.s. is {a}.

                  theorem Set.Ioc_subset_Ioc_union_Icc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  @[simp]
                  theorem Set.Ioc_union_Icc_eq_Ioc {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} (h₁ : a < b) (h₂ : b c) :
                  theorem Set.Ioo_union_Ioo' {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : c < b) (h₂ : a < d) :
                  Set.Ioo a b Set.Ioo c d = Set.Ioo (min a c) (max b d)
                  theorem Set.Ioo_union_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} {d : α} (h₁ : min a b < max c d) (h₂ : min c d < max a b) :
                  Set.Ioo a b Set.Ioo c d = Set.Ioo (min a c) (max b d)
                  @[simp]
                  theorem Set.Iic_inter_Iic {α : Type u_1} [SemilatticeInf α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ioc_inter_Iic {α : Type u_1} [SemilatticeInf α] (a : α) (b : α) (c : α) :
                  Set.Ioc a b Set.Iic c = Set.Ioc a (b c)
                  @[simp]
                  theorem Set.Ici_inter_Ici {α : Type u_1} [SemilatticeSup α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Ico_inter_Ici {α : Type u_1} [SemilatticeSup α] (a : α) (b : α) (c : α) :
                  Set.Ico a b Set.Ici c = Set.Ico (a c) b
                  theorem Set.Icc_inter_Icc {α : Type u_1} [Lattice α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} :
                  Set.Icc a₁ b₁ Set.Icc a₂ b₂ = Set.Icc (a₁ a₂) (b₁ b₂)
                  @[simp]
                  theorem Set.Icc_inter_Icc_eq_singleton {α : Type u_1} [Lattice α] {a : α} {b : α} {c : α} (hab : a b) (hbc : b c) :
                  Set.Icc a b Set.Icc b c = {b}
                  @[simp]
                  theorem Set.Ioi_inter_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  @[simp]
                  theorem Set.Iio_inter_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  theorem Set.Ico_inter_Ico {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} :
                  Set.Ico a₁ b₁ Set.Ico a₂ b₂ = Set.Ico (a₁ a₂) (b₁ b₂)
                  theorem Set.Ioc_inter_Ioc {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} :
                  Set.Ioc a₁ b₁ Set.Ioc a₂ b₂ = Set.Ioc (a₁ a₂) (b₁ b₂)
                  theorem Set.Ioo_inter_Ioo {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} :
                  Set.Ioo a₁ b₁ Set.Ioo a₂ b₂ = Set.Ioo (a₁ a₂) (b₁ b₂)
                  theorem Set.Ioo_inter_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioo a b Set.Iio c = Set.Ioo a (min b c)
                  theorem Set.Iio_inter_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Iio a Set.Ioo b c = Set.Ioo b (min a c)
                  theorem Set.Ioo_inter_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioo a b Set.Ioi c = Set.Ioo (max a c) b
                  theorem Set.Ioi_inter_Ioo {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioi a Set.Ioo b c = Set.Ioo (max a b) c
                  theorem Set.Ioc_inter_Ioo_of_left_lt {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h : b₁ < b₂) :
                  Set.Ioc a₁ b₁ Set.Ioo a₂ b₂ = Set.Ioc (max a₁ a₂) b₁
                  theorem Set.Ioc_inter_Ioo_of_right_le {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h : b₂ b₁) :
                  Set.Ioc a₁ b₁ Set.Ioo a₂ b₂ = Set.Ioo (max a₁ a₂) b₂
                  theorem Set.Ioo_inter_Ioc_of_left_le {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h : b₁ b₂) :
                  Set.Ioo a₁ b₁ Set.Ioc a₂ b₂ = Set.Ioo (max a₁ a₂) b₁
                  theorem Set.Ioo_inter_Ioc_of_right_lt {α : Type u_1} [LinearOrder α] {a₁ : α} {a₂ : α} {b₁ : α} {b₂ : α} (h : b₂ < b₁) :
                  Set.Ioo a₁ b₁ Set.Ioc a₂ b₂ = Set.Ioc (max a₁ a₂) b₂
                  @[simp]
                  theorem Set.Ico_diff_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ico a b \ Set.Iio c = Set.Ico (max a c) b
                  @[simp]
                  theorem Set.Ioc_diff_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioc a b \ Set.Ioi c = Set.Ioc a (min b c)
                  @[simp]
                  theorem Set.Ioc_inter_Ioi {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioc a b Set.Ioi c = Set.Ioc (a c) b
                  @[simp]
                  theorem Set.Ico_inter_Iio {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ico a b Set.Iio c = Set.Ico a (min b c)
                  @[simp]
                  theorem Set.Ioc_diff_Iic {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioc a b \ Set.Iic c = Set.Ioc (max a c) b
                  @[simp]
                  theorem Set.Ioc_union_Ioc_right {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioc a b Set.Ioc a c = Set.Ioc a (max b c)
                  @[simp]
                  theorem Set.Ioc_union_Ioc_left {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioc a c Set.Ioc b c = Set.Ioc (min a b) c
                  @[simp]
                  theorem Set.Ioc_union_Ioc_symm {α : Type u_1} [LinearOrder α] {a : α} {b : α} :
                  Set.Ioc a b Set.Ioc b a = Set.Ioc (min a b) (max a b)
                  @[simp]
                  theorem Set.Ioc_union_Ioc_union_Ioc_cycle {α : Type u_1} [LinearOrder α] {a : α} {b : α} {c : α} :
                  Set.Ioc a b Set.Ioc b c Set.Ioc c a = Set.Ioc (min a (min b c)) (max a (max b c))

                  Closed intervals in α × β #

                  @[simp]
                  theorem Set.Iic_prod_Iic {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] (a : α) (b : β) :
                  @[simp]
                  theorem Set.Ici_prod_Ici {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] (a : α) (b : β) :
                  theorem Set.Ici_prod_eq {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] (a : α × β) :
                  theorem Set.Iic_prod_eq {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] (a : α × β) :
                  @[simp]
                  theorem Set.Icc_prod_Icc {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] (a₁ : α) (a₂ : α) (b₁ : β) (b₂ : β) :
                  Set.Icc a₁ a₂ ×ˢ Set.Icc b₁ b₂ = Set.Icc (a₁, b₁) (a₂, b₂)
                  theorem Set.Icc_prod_eq {α : Type u_1} {β : Type u_2} [Preorder α] [Preorder β] (a : α × β) (b : α × β) :
                  Set.Icc a b = Set.Icc a.1 b.1 ×ˢ Set.Icc a.2 b.2

                  Lemmas about intervals in dense orders #

                  instance instNoMinOrderElemIooLtToLTMemSetInstMembershipSet (α : Type u_1) [Preorder α] [DenselyOrdered α] {x : α} {y : α} :
                  Equations
                  • =
                  instance instNoMinOrderElemIocLtToLTMemSetInstMembershipSet (α : Type u_1) [Preorder α] [DenselyOrdered α] {x : α} {y : α} :
                  Equations
                  • =
                  Equations
                  • =
                  instance instNoMaxOrderElemIooLtToLTMemSetInstMembershipSet (α : Type u_1) [Preorder α] [DenselyOrdered α] {x : α} {y : α} :
                  Equations
                  • =
                  instance instNoMaxOrderElemIcoLtToLTMemSetInstMembershipSet (α : Type u_1) [Preorder α] [DenselyOrdered α] {x : α} {y : α} :
                  Equations
                  • =
                  Equations
                  • =