Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit a1c0ad5

Browse files
urkudmergify[bot]
authored andcommitted
feat(category_theory): def is_isomorphic_setoid, groupoid.iso_equiv_hom (#1506)
* feat(category_theory): def `is_isomorphic_setoid`, `groupoid.iso_equiv_hom` * Move to a dedicated file, define `isomorphic_class_functor` * explicit/implicit arguments * Update src/category_theory/groupoid.lean * Update src/category_theory/groupoid.lean * Update src/category_theory/isomorphism_classes.lean * Update src/category_theory/isomorphism_classes.lean * Update src/category_theory/isomorphism_classes.lean
1 parent e5fc2a7 commit a1c0ad5

File tree

2 files changed

+68
-2
lines changed

2 files changed

+68
-2
lines changed

src/category_theory/groupoid.lean

Lines changed: 19 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -6,11 +6,13 @@ Authors: Reid Barton
66

77
import category_theory.category
88
import category_theory.isomorphism
9+
import data.equiv.basic
910

1011
namespace category_theory
1112

1213
universes v u -- declare the `v`'s first; see `category_theory.category` for an explanation
1314

15+
/-- A `groupoid` is a category such that all morphisms are isomorphisms. -/
1416
class groupoid (obj : Type u) extends category.{v} obj : Type (max u (v+1)) :=
1517
(inv : Π {X Y : obj}, (X ⟶ Y) → (Y ⟶ X))
1618
(inv_comp' : ∀ {X Y : obj} (f : X ⟶ Y), comp (inv f) f = id Y . obviously)
@@ -24,7 +26,22 @@ attribute [simp] groupoid.inv_comp groupoid.comp_inv
2426
abbreviation large_groupoid (C : Type (u+1)) : Type (u+1) := groupoid.{u} C
2527
abbreviation small_groupoid (C : Type u) : Type (u+1) := groupoid.{u} C
2628

27-
instance of_groupoid {C : Type u} [groupoid.{v} C] {X Y : C} (f : X ⟶ Y) : is_iso f :=
28-
{ inv := groupoid.inv f }
29+
section
30+
31+
variables {C : Type u} [𝒞 : groupoid.{v} C] {X Y : C}
32+
include 𝒞
33+
34+
instance is_iso.of_groupoid (f : X ⟶ Y) : is_iso f := { inv := groupoid.inv f }
35+
36+
variables (X Y)
37+
38+
/-- In a groupoid, isomorphisms are equivalent to morphisms. -/
39+
def groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y) :=
40+
{ to_fun := iso.hom,
41+
inv_fun := λ f, as_iso f,
42+
left_inv := λ i, iso.ext rfl,
43+
right_inv := λ f, rfl }
44+
45+
end
2946

3047
end category_theory
Lines changed: 49 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,49 @@
1+
/-
2+
Copyright (c) 2019 Yury Kudryashov. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yury Kudryashov
5+
-/
6+
import category_theory.category.Cat category_theory.groupoid data.quot
7+
8+
/-!
9+
# Objects of a category up to an isomorphism
10+
11+
`is_isomorphic X Y := nonempty (X ≅ Y)` is an equivalence relation on the objects of a category.
12+
The quotient with respect to this relation defines a functor from our category to `Type`.
13+
-/
14+
15+
universes v u
16+
17+
namespace category_theory
18+
19+
section category
20+
21+
variables {C : Type u} [𝒞 : category.{v} C]
22+
include 𝒞
23+
24+
/-- An object `X` is isomorphic to an object `Y`, if `X ≅ Y` is not empty. -/
25+
def is_isomorphic : C → C → Prop := λ X Y, nonempty (X ≅ Y)
26+
27+
variable (C)
28+
29+
/-- `is_isomorphic` defines a setoid. -/
30+
def is_isomorphic_setoid : setoid C :=
31+
{ r := is_isomorphic,
32+
iseqv := ⟨λ X, ⟨iso.refl X⟩, λ X Y ⟨α⟩, ⟨α.symm⟩, λ X Y Z ⟨α⟩ ⟨β⟩, ⟨α.trans β⟩⟩ }
33+
34+
end category
35+
36+
/--
37+
The functor that sends each category to the quotient space of its objects up to an isomorphism.
38+
-/
39+
def isomorphism_classes : Cat.{v u} ⥤ Type u :=
40+
{ obj := λ C, quotient (is_isomorphic_setoid C.α),
41+
map := λ C D F, quot.map F.obj $ λ X Y ⟨f⟩, ⟨F.map_iso f⟩ }
42+
43+
lemma groupoid.is_isomorphic_iff_nonempty_hom {C : Type u} [groupoid.{v} C] {X Y : C} :
44+
is_isomorphic X Y ↔ nonempty (X ⟶ Y) :=
45+
(groupoid.iso_equiv_hom X Y).nonempty_iff_nonempty
46+
47+
-- PROJECT: define `skeletal`, and show every category is equivalent to a skeletal category,
48+
-- using the axiom of choice to pick a representative of every isomorphism class.
49+
end category_theory

0 commit comments

Comments
 (0)