Skip to content

Commit

Permalink
feat(category_theory/category/Kleisli): Fix lint errors (#8244)
Browse files Browse the repository at this point in the history
Fixes some lint errors for this file: unused arguments, module doc, inhabited instances
  • Loading branch information
b-mehta committed Jul 9, 2021
1 parent a444e81 commit 70320f7
Showing 1 changed file with 23 additions and 12 deletions.
35 changes: 23 additions & 12 deletions src/category_theory/category/Kleisli.lean
Expand Up @@ -9,31 +9,42 @@ TODO: generalise this to work with category_theory.monad
-/
import category_theory.category

/-!
# The Kleisli construction on the Type category
Define the Kleisli category for (control) monads.
`category_theory/monad/kleisli` defines the general version for a monad on `C`, and demonstrates
the equivalence between the two.
-/

universes u v

namespace category_theory

def Kleisli (m) [monad.{u v} m] := Type u
/-- The Kleisli category on the (type-)monad `m`. Note that the monad is not assumed to be lawful
yet. -/
@[nolint unused_arguments]
def Kleisli (m : Type u → Type v) := Type u

def Kleisli.mk (m) [monad.{u v} m] (α : Type u) : Kleisli m := α
/-- Construct an object of the Kleisli category from a type. -/
def Kleisli.mk (m) (α : Type u) : Kleisli m := α

instance Kleisli.category_struct {m} [monad m] : category_struct (Kleisli m) :=
instance Kleisli.category_struct {m} [monad.{u v} m] : category_struct (Kleisli m) :=
{ hom := λ α β, α → m β,
id := λ α x, (pure x : m α),
id := λ α x, pure x,
comp := λ X Y Z f g, f >=> g }

instance Kleisli.category {m} [monad m] [is_lawful_monad m] : category (Kleisli m) :=
by refine { hom := λ α β, α → m β,
id := λ α x, (pure x : m α),
comp := λ X Y Z f g, f >=> g,
id_comp' := _, comp_id' := _, assoc' := _ };
intros; ext; simp only [(>=>)] with functor_norm
instance Kleisli.category {m} [monad.{u v} m] [is_lawful_monad m] : category (Kleisli m) :=
by refine { id_comp' := _, comp_id' := _, assoc' := _ };
intros; ext; unfold_projs; simp only [(>=>)] with functor_norm

@[simp] lemma Kleisli.id_def {m} [monad m] [is_lawful_monad m] (α : Kleisli m) :
@[simp] lemma Kleisli.id_def {m} [monad m] (α : Kleisli m) :
𝟙 α = @pure m _ α := rfl

lemma Kleisli.comp_def {m} [monad m] [is_lawful_monad m] (α β γ : Kleisli m)
lemma Kleisli.comp_def {m} [monad m] (α β γ : Kleisli m)
(xs : α ⟶ β) (ys : β ⟶ γ) (a : α) :
(xs ≫ ys) a = xs a >>= ys := rfl

instance : inhabited (Kleisli id) := ⟨punit⟩
instance {α : Type u} [inhabited α] : inhabited (Kleisli.mk id α) := ⟨(default α : _)⟩
end category_theory

0 comments on commit 70320f7

Please sign in to comment.