|
| 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