Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(category_theory): make defining groupoids easier #2360

Closed
wants to merge 4 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
12 changes: 12 additions & 0 deletions src/category_theory/epi_mono.lean
Original file line number Diff line number Diff line change
Expand Up @@ -116,6 +116,18 @@ instance split_mono.mono {X Y : C} (f : X ⟶ Y) [split_mono f] : mono f :=
instance split_epi.epi {X Y : C} (f : X ⟶ Y) [split_epi f] : epi f :=
{ left_cancellation := λ Z g h w, begin replace w := section_ f ≫= w, simpa using w, end }

/-- Every split mono whose retraction is mono is an iso. -/
def is_iso.of_mono_retraction {X Y : C} {f : X ⟶ Y} [split_mono f] [mono $ retraction f]
: is_iso f :=
{ inv := retraction f,
inv_hom_id' := (cancel_mono_id $ retraction f).mp (by simp) }

/-- Every split epi whose section is epi is an iso. -/
def is_iso.of_epi_section {X Y : C} {f : X ⟶ Y} [split_epi f] [epi $ section_ f]
: is_iso f :=
{ inv := section_ f,
hom_inv_id' := (cancel_epi_id $ section_ f).mp (by simp) }

section
variables {D : Type u₂} [𝒟 : category.{v₂} D]
include 𝒟
Expand Down
26 changes: 25 additions & 1 deletion src/category_theory/groupoid.lean
Original file line number Diff line number Diff line change
@@ -1,11 +1,12 @@
/-
Copyright (c) 2018 Reid Barton All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Reid Barton
Authors: Reid Barton, Scott Morrison, David Wärn
-/

import category_theory.category
import category_theory.isomorphism
import category_theory.epi_mono
import data.equiv.basic

namespace category_theory
Expand Down Expand Up @@ -48,4 +49,27 @@ def groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y) :=

end

section

variables {C : Type u} [𝒞 : category.{v} C]
include 𝒞

/-- A category where every morphism `is_iso` is a groupoid. -/
def groupoid.of_is_iso (all_is_iso : ∀ {X Y : C} (f : X ⟶ Y), is_iso.{v} f) : groupoid.{v} C :=
{ inv := λ X Y f, (all_is_iso f).inv }

/-- A category where every morphism has a `trunc` retraction is computably a groupoid. -/
def groupoid.of_trunc_split_mono
(all_split_mono : ∀ {X Y : C} (f : X ⟶ Y), trunc (split_mono.{v} f)) :
groupoid.{v} C :=
begin
apply groupoid.of_is_iso,
intros X Y f,
trunc_cases all_split_mono f,
trunc_cases all_split_mono (retraction f),
apply is_iso.of_mono_retraction,
end

end

end category_theory
1 change: 1 addition & 0 deletions src/tactic/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -29,5 +29,6 @@ import
tactic.simps
tactic.split_ifs
tactic.squeeze
tactic.trunc_cases
tactic.where
tactic.hint