Documentation

Init.Data.Function

@[inline]
def Function.curry {α : Type u_1} {β : Type u_2} {φ : Sort u_3} :
(α × βφ)αβφ

Transforms a function from pairs into an equivalent two-parameter function.

Examples:

Equations
Instances For
    @[inline]
    def Function.uncurry {α : Type u_1} {β : Type u_2} {φ : Sort u_3} :
    (αβφ)α × βφ

    Transforms a two-parameter function into an equivalent function from pairs.

    Examples:

    Equations
    Instances For
      @[simp]
      theorem Function.curry_uncurry {α : Type u_1} {β : Type u_2} {φ : Sort u_3} (f : αβφ) :
      @[simp]
      theorem Function.uncurry_curry {α : Type u_1} {β : Type u_2} {φ : Sort u_3} (f : α × βφ) :
      @[simp]
      theorem Function.uncurry_apply_pair {α : Type u_1} {β : Type u_2} {γ : Sort u_3} (f : αβγ) (x : α) (y : β) :
      uncurry f (x, y) = f x y
      @[simp]
      theorem Function.curry_apply {α : Type u_1} {β : Type u_2} {γ : Sort u_3} (f : α × βγ) (x : α) (y : β) :
      curry f x y = f (x, y)
      def Function.Injective {α : Sort u_1} {β : Sort u_2} (f : αβ) :

      A function f : α → β is called injective if f x = f y implies x = y.

      Equations
      Instances For
        theorem Function.Injective.comp {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {g : βγ} {f : αβ} (hg : Injective g) (hf : Injective f) :
        def Function.Surjective {α : Sort u_1} {β : Sort u_2} (f : αβ) :

        A function f : α → β is called surjective if every b : β is equal to f a for some a : α.

        Equations
        Instances For
          theorem Function.Surjective.comp {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} {g : βγ} {f : αβ} (hg : Surjective g) (hf : Surjective f) :
          def Function.LeftInverse {α : Sort u_1} {β : Sort u_2} (g : βα) (f : αβ) :

          LeftInverse g f means that g is a left inverse to f. That is, g ∘ f = id.

          Equations
          Instances For
            def Function.HasLeftInverse {α : Sort u_1} {β : Sort u_2} (f : αβ) :

            HasLeftInverse f means that f has an unspecified left inverse.

            Equations
            Instances For
              def Function.RightInverse {α : Sort u_1} {β : Sort u_2} (g : βα) (f : αβ) :

              RightInverse g f means that g is a right inverse to f. That is, f ∘ g = id.

              Equations
              Instances For
                def Function.HasRightInverse {α : Sort u_1} {β : Sort u_2} (f : αβ) :

                HasRightInverse f means that f has an unspecified right inverse.

                Equations
                Instances For
                  theorem Function.LeftInverse.injective {α : Sort u_1} {β : Sort u_2} {g : βα} {f : αβ} :
                  theorem Function.HasLeftInverse.injective {α : Sort u_1} {β : Sort u_2} {f : αβ} :
                  theorem Function.rightInverse_of_injective_of_leftInverse {α : Sort u_1} {β : Sort u_2} {f : αβ} {g : βα} (injf : Injective f) (lfg : LeftInverse f g) :
                  theorem Function.RightInverse.surjective {α : Sort u_1} {β : Sort u_2} {f : αβ} {g : βα} (h : RightInverse g f) :
                  theorem Function.HasRightInverse.surjective {α : Sort u_1} {β : Sort u_2} {f : αβ} :
                  theorem Function.leftInverse_of_surjective_of_rightInverse {α : Sort u_1} {β : Sort u_2} {f : αβ} {g : βα} (surjf : Surjective f) (rfg : RightInverse f g) :
                  theorem Function.Injective.eq_iff {α : Sort u_1} {β : Sort u_2} {f : αβ} (I : Injective f) {a b : α} :
                  f a = f b a = b
                  theorem Function.Injective.eq_iff' {α : Sort u_1} {β : Sort u_2} {f : αβ} (I : Injective f) {a b : α} {c : β} (h : f b = c) :
                  f a = c a = b
                  theorem Function.Injective.ne {α : Sort u_1} {β : Sort u_2} {f : αβ} (hf : Injective f) {a₁ a₂ : α} :
                  a₁ a₂f a₁ f a₂
                  theorem Function.Injective.ne_iff {α : Sort u_1} {β : Sort u_2} {f : αβ} (hf : Injective f) {x y : α} :
                  f x f y x y
                  theorem Function.Injective.ne_iff' {α : Sort u_1} {β : Sort u_2} {f : αβ} (hf : Injective f) {x y : α} {z : β} (h : f y = z) :
                  f x z x y
                  theorem Function.LeftInverse.id {α : Sort u_1} {β : Sort u_2} {g : βα} {f : αβ} (h : LeftInverse g f) :
                  g f = id
                  theorem Function.RightInverse.id {α : Sort u_1} {β : Sort u_2} {g : βα} {f : αβ} (h : RightInverse g f) :
                  f g = id