Documentation

Mathlib.Logic.Function.Coequalizer

Coequalizer of a pair of functions #

The coequalizer of two functions f g : α → β is the pair (μ, p : β → μ) that satisfies the following universal property: Every function u : β → γ with u ∘ f = u ∘ g factors uniquely via p.

In this file we define the coequalizer and provide the basic API.

inductive Function.Coequalizer.Rel {α : Type u_1} {β : Type u_2} (f g : αβ) :
ββProp

The relation generating the equivalence relation used for defining Function.coequalizer.

  • intro {α : Type u_1} {β : Type u_2} {f g : αβ} (x : α) : Rel f g (f x) (g x)
Instances For
    def Function.Coequalizer {α : Type u_1} {β : Type v} (f g : αβ) :

    The coequalizer of two functions f g : α → β is the pair (μ, p : β → μ) that satisfies the following universal property: Every function u : β → γ with u ∘ f = u ∘ g factors uniquely via p.

    Equations
    Instances For
      def Function.Coequalizer.mk {α : Type u_1} {β : Type u_2} (f g : αβ) (x : β) :

      The canonical projection to the coequalizer.

      Equations
      Instances For
        theorem Function.Coequalizer.condition {α : Type u_1} {β : Type u_2} (f g : αβ) (x : α) :
        mk f g (f x) = mk f g (g x)
        theorem Function.Coequalizer.mk_surjective {α : Type u_1} {β : Type u_2} (f g : αβ) :
        def Function.Coequalizer.desc {α : Type u_1} {β : Type u_2} (f g : αβ) {γ : Type u_3} (u : βγ) (hu : u f = u g) :
        Coequalizer f gγ

        Any map u : β → γ with u ∘ f = u ∘ g factors via Function.Coequalizer.mk.

        Equations
        Instances For
          @[simp]
          theorem Function.Coequalizer.desc_mk {α : Type u_1} {β : Type u_2} (f g : αβ) {γ : Type u_3} (u : βγ) (hu : u f = u g) (x : β) :
          desc f g u hu (mk f g x) = u x