2015-02-26 18:19:54 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2014 Jakob von Raumer. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2015-10-01 19:52:28 +00:00
|
|
|
Authors: Jakob von Raumer, Floris van Doorn
|
2015-02-26 18:19:54 +00:00
|
|
|
|
|
|
|
Ported from Coq HoTT
|
|
|
|
-/
|
|
|
|
|
2015-04-25 04:20:59 +00:00
|
|
|
import .iso ..group
|
2015-02-26 18:19:54 +00:00
|
|
|
|
2015-05-05 21:25:35 +00:00
|
|
|
open eq is_trunc iso category algebra nat unit
|
2015-02-26 18:19:54 +00:00
|
|
|
|
|
|
|
namespace category
|
|
|
|
|
|
|
|
structure groupoid [class] (ob : Type) extends parent : precategory ob :=
|
2015-04-25 03:04:24 +00:00
|
|
|
mk' :: (all_iso : Π ⦃a b : ob⦄ (f : hom a b), @is_iso ob parent a b f)
|
2015-02-26 18:19:54 +00:00
|
|
|
|
|
|
|
abbreviation all_iso := @groupoid.all_iso
|
2015-03-11 20:03:10 +00:00
|
|
|
attribute groupoid.all_iso [instance] [priority 100000]
|
2015-02-26 18:19:54 +00:00
|
|
|
|
2015-04-25 03:04:24 +00:00
|
|
|
definition groupoid.mk [reducible] {ob : Type} (C : precategory ob)
|
2015-02-26 18:19:54 +00:00
|
|
|
(H : Π (a b : ob) (f : a ⟶ b), is_iso f) : groupoid ob :=
|
2015-04-25 03:04:24 +00:00
|
|
|
precategory.rec_on C groupoid.mk' H
|
2015-02-26 18:19:54 +00:00
|
|
|
|
2015-09-24 02:44:36 +00:00
|
|
|
-- We can turn each group into a groupoid on the unit type
|
2015-06-24 21:59:17 +00:00
|
|
|
definition groupoid_of_group.{l} (A : Type.{l}) [G : group A] : groupoid.{0 l} unit :=
|
2015-02-26 18:19:54 +00:00
|
|
|
begin
|
2015-04-25 03:04:24 +00:00
|
|
|
fapply groupoid.mk, fapply precategory.mk,
|
2015-02-26 18:19:54 +00:00
|
|
|
intros, exact A,
|
2015-05-05 21:25:35 +00:00
|
|
|
intros, apply (@group.is_hset_carrier A G),
|
2015-03-28 00:26:06 +00:00
|
|
|
intros [a, b, c, g, h], exact (@group.mul A G g h),
|
2015-02-26 18:19:54 +00:00
|
|
|
intro a, exact (@group.one A G),
|
2015-02-28 00:02:18 +00:00
|
|
|
intros, exact (@group.mul_assoc A G h g f)⁻¹,
|
2015-02-26 18:19:54 +00:00
|
|
|
intros, exact (@group.one_mul A G f),
|
|
|
|
intros, exact (@group.mul_one A G f),
|
2015-04-25 03:04:24 +00:00
|
|
|
intros, esimp [precategory.mk], apply is_iso.mk,
|
2015-05-05 21:25:35 +00:00
|
|
|
apply mul.left_inv,
|
|
|
|
apply mul.right_inv,
|
2015-02-26 18:19:54 +00:00
|
|
|
end
|
|
|
|
|
2015-09-24 02:44:36 +00:00
|
|
|
definition hom_group {A : Type} [G : groupoid A] (a : A) :
|
2015-02-26 18:19:54 +00:00
|
|
|
group (hom a a) :=
|
|
|
|
begin
|
|
|
|
fapply group.mk,
|
2015-04-30 18:00:39 +00:00
|
|
|
intro f g, apply (comp f g),
|
2015-04-25 03:04:24 +00:00
|
|
|
apply is_hset_hom,
|
2015-04-30 18:00:39 +00:00
|
|
|
intros f g h, apply (assoc f g h)⁻¹,
|
2015-02-26 18:19:54 +00:00
|
|
|
apply (ID a),
|
|
|
|
intro f, apply id_left,
|
|
|
|
intro f, apply id_right,
|
2015-02-28 06:16:20 +00:00
|
|
|
intro f, exact (iso.inverse f),
|
|
|
|
intro f, exact (iso.left_inverse f),
|
2015-02-26 18:19:54 +00:00
|
|
|
end
|
|
|
|
|
2015-09-24 02:44:36 +00:00
|
|
|
definition group_of_is_contr_groupoid {ob : Type} [H : is_contr ob]
|
|
|
|
[G : groupoid ob] : group (hom (center ob) (center ob)) := !hom_group
|
|
|
|
definition group_of_groupoid_unit [G : groupoid unit] : group (hom ⋆ ⋆) := !hom_group
|
|
|
|
|
2015-02-26 18:19:54 +00:00
|
|
|
-- Bundled version of categories
|
|
|
|
-- we don't use Groupoid.carrier explicitly, but rather use Groupoid.carrier (to_Precategory C)
|
|
|
|
structure Groupoid : Type :=
|
|
|
|
(carrier : Type)
|
|
|
|
(struct : groupoid carrier)
|
|
|
|
|
|
|
|
attribute Groupoid.struct [instance] [coercion]
|
|
|
|
|
|
|
|
definition Groupoid.to_Precategory [coercion] [reducible] (C : Groupoid) : Precategory :=
|
|
|
|
Precategory.mk (Groupoid.carrier C) C
|
|
|
|
|
|
|
|
definition groupoid.Mk [reducible] := Groupoid.mk
|
2015-03-13 22:33:53 +00:00
|
|
|
definition groupoid.MK [reducible] (C : Precategory) (H : Π (a b : C) (f : a ⟶ b), is_iso f)
|
2015-02-26 18:19:54 +00:00
|
|
|
: Groupoid :=
|
2015-04-25 03:04:24 +00:00
|
|
|
Groupoid.mk C (groupoid.mk C H)
|
2015-02-26 18:19:54 +00:00
|
|
|
|
|
|
|
definition Groupoid.eta (C : Groupoid) : Groupoid.mk C C = C :=
|
|
|
|
Groupoid.rec (λob c, idp) C
|
|
|
|
|
|
|
|
end category
|