2016-10-13 19:04:57 +00:00
|
|
|
|
/-
|
2017-06-08 23:51:25 +00:00
|
|
|
|
Copyright (c) 2015 Floris van Doorn, Egbert Rijke, Favonia. All rights reserved.
|
2016-10-13 19:04:57 +00:00
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2017-06-08 23:51:25 +00:00
|
|
|
|
Authors: Floris van Doorn, Egbert Rijke, Favonia
|
2016-10-13 19:04:57 +00:00
|
|
|
|
|
|
|
|
|
Constructions with groups
|
|
|
|
|
-/
|
|
|
|
|
|
2017-07-04 15:11:21 +00:00
|
|
|
|
import .quotient_group .free_abelian_group .product_group
|
2016-10-13 20:01:17 +00:00
|
|
|
|
|
2017-06-09 16:08:21 +00:00
|
|
|
|
open eq is_equiv algebra is_trunc set_quotient relation sigma prod sum list trunc function equiv sigma.ops lift
|
2016-10-13 19:04:57 +00:00
|
|
|
|
|
|
|
|
|
namespace group
|
|
|
|
|
|
|
|
|
|
section
|
|
|
|
|
|
2017-06-07 15:30:09 +00:00
|
|
|
|
parameters {I : Type} [is_set I] (Y : I → AbGroup)
|
2017-06-07 05:01:19 +00:00
|
|
|
|
variables {A' : AbGroup} {Y' : I → AbGroup}
|
2016-10-13 19:04:57 +00:00
|
|
|
|
|
2017-06-07 15:30:09 +00:00
|
|
|
|
definition dirsum_carrier : AbGroup := free_ab_group (Σi, Y i)
|
2016-11-24 04:54:57 +00:00
|
|
|
|
local abbreviation ι [constructor] := @free_ab_group_inclusion
|
2016-10-13 19:04:57 +00:00
|
|
|
|
inductive dirsum_rel : dirsum_carrier → Type :=
|
2017-06-07 05:01:19 +00:00
|
|
|
|
| rmk : Πi y₁ y₂, dirsum_rel (ι ⟨i, y₁⟩ * ι ⟨i, y₂⟩ * (ι ⟨i, y₁ * y₂⟩)⁻¹)
|
2016-10-13 19:04:57 +00:00
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum : AbGroup := quotient_ab_group_gen dirsum_carrier (λg, ∥dirsum_rel g∥)
|
2016-11-14 19:44:29 +00:00
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
-- definition dirsum_carrier_incl [constructor] (i : I) : Y i →g dirsum_carrier :=
|
2016-11-14 19:44:29 +00:00
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_incl [constructor] (i : I) : Y i →g dirsum :=
|
|
|
|
|
homomorphism.mk (λy, class_of (ι ⟨i, y⟩))
|
2016-11-14 19:44:29 +00:00
|
|
|
|
begin intro g h, symmetry, apply gqg_eq_of_rel, apply tr, apply dirsum_rel.rmk end
|
|
|
|
|
|
2017-03-31 22:21:02 +00:00
|
|
|
|
parameter {Y}
|
|
|
|
|
definition dirsum.rec {P : dirsum → Type} [H : Πg, is_prop (P g)]
|
2017-06-07 05:01:19 +00:00
|
|
|
|
(h₁ : Πi (y : Y i), P (dirsum_incl i y)) (h₂ : P 1) (h₃ : Πg h, P g → P h → P (g * h)) :
|
2017-03-31 22:21:02 +00:00
|
|
|
|
Πg, P g :=
|
|
|
|
|
begin
|
|
|
|
|
refine @set_quotient.rec_prop _ _ _ H _,
|
|
|
|
|
refine @set_quotient.rec_prop _ _ _ (λx, !H) _,
|
|
|
|
|
esimp, intro l, induction l with s l ih,
|
|
|
|
|
exact h₂,
|
|
|
|
|
induction s with v v,
|
|
|
|
|
induction v with i y,
|
|
|
|
|
exact h₃ _ _ (h₁ i y) ih,
|
|
|
|
|
induction v with i y,
|
|
|
|
|
refine h₃ (gqg_map _ _ (class_of [inr ⟨i, y⟩])) _ _ ih,
|
2017-06-07 05:01:19 +00:00
|
|
|
|
refine transport P _ (h₁ i y⁻¹),
|
2017-03-31 22:21:02 +00:00
|
|
|
|
refine _ ⬝ !one_mul,
|
2017-04-07 19:56:37 +00:00
|
|
|
|
refine _ ⬝ ap (λx, mul x _) (to_respect_zero (dirsum_incl i)),
|
2017-03-31 22:21:02 +00:00
|
|
|
|
apply gqg_eq_of_rel',
|
|
|
|
|
apply tr, esimp,
|
2017-06-07 05:01:19 +00:00
|
|
|
|
refine transport dirsum_rel _ (dirsum_rel.rmk i y⁻¹ y),
|
|
|
|
|
rewrite [mul.left_inv, mul.assoc],
|
2017-03-31 22:21:02 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_homotopy {φ ψ : dirsum →g A'}
|
2017-03-31 22:21:02 +00:00
|
|
|
|
(h : Πi (y : Y i), φ (dirsum_incl i y) = ψ (dirsum_incl i y)) : φ ~ ψ :=
|
|
|
|
|
begin
|
|
|
|
|
refine dirsum.rec _ _ _,
|
|
|
|
|
exact h,
|
2017-04-07 19:56:37 +00:00
|
|
|
|
refine !to_respect_zero ⬝ !to_respect_zero⁻¹,
|
2017-06-07 05:01:19 +00:00
|
|
|
|
intro g₁ g₂ h₁ h₂, rewrite [* to_respect_mul, h₁, h₂]
|
2017-03-31 22:21:02 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_elim_resp_quotient (f : Πi, Y i →g A') (g : dirsum_carrier)
|
2016-11-24 04:54:57 +00:00
|
|
|
|
(r : ∥dirsum_rel g∥) : free_ab_group_elim (λv, f v.1 v.2) g = 1 :=
|
2016-11-14 19:44:29 +00:00
|
|
|
|
begin
|
2016-11-18 20:20:22 +00:00
|
|
|
|
induction r with r, induction r,
|
2017-06-07 05:01:19 +00:00
|
|
|
|
rewrite [to_respect_mul, to_respect_inv, to_respect_mul, ▸*, ↑foldl, *one_mul,
|
|
|
|
|
to_respect_mul], apply mul.right_inv
|
2016-11-14 19:44:29 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_elim [constructor] (f : Πi, Y i →g A') : dirsum →g A' :=
|
2016-11-24 04:54:57 +00:00
|
|
|
|
gqg_elim _ (free_ab_group_elim (λv, f v.1 v.2)) (dirsum_elim_resp_quotient f)
|
2016-11-18 20:20:22 +00:00
|
|
|
|
|
2017-06-08 20:04:58 +00:00
|
|
|
|
definition dirsum_elim_compute (f : Πi, Y i →g A') (i : I) (y : Y i) :
|
|
|
|
|
dirsum_elim f (dirsum_incl i y) = f i y :=
|
2016-11-14 19:44:29 +00:00
|
|
|
|
begin
|
2017-06-08 20:04:58 +00:00
|
|
|
|
apply one_mul
|
2016-11-14 19:44:29 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_elim_unique (f : Πi, Y i →g A') (k : dirsum →g A')
|
2016-11-14 19:44:29 +00:00
|
|
|
|
(H : Πi, k ∘g dirsum_incl i ~ f i) : k ~ dirsum_elim f :=
|
2016-11-18 20:20:22 +00:00
|
|
|
|
begin
|
|
|
|
|
apply gqg_elim_unique,
|
2016-11-24 04:54:57 +00:00
|
|
|
|
apply free_ab_group_elim_unique,
|
2016-11-18 20:20:22 +00:00
|
|
|
|
intro x, induction x with i y, exact H i y
|
|
|
|
|
end
|
2016-11-14 19:44:29 +00:00
|
|
|
|
|
2017-03-31 22:21:02 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-06-08 20:04:58 +00:00
|
|
|
|
definition binary_dirsum (G H : AbGroup) : dirsum (bool.rec G H) ≃g G ×ag H :=
|
|
|
|
|
let branch := bool.rec G H in
|
|
|
|
|
let to_hom := (dirsum_elim (bool.rec (product_inl G H) (product_inr G H))
|
|
|
|
|
: dirsum (bool.rec G H) →g G ×ag H) in
|
|
|
|
|
let from_hom := (Group_sum_elim (dirsum (bool.rec G H))
|
|
|
|
|
(dirsum_incl branch bool.ff) (dirsum_incl branch bool.tt)
|
|
|
|
|
: G ×g H →g dirsum branch) in
|
|
|
|
|
begin
|
|
|
|
|
fapply isomorphism.mk,
|
|
|
|
|
{ exact dirsum_elim (bool.rec (product_inl G H) (product_inr G H)) },
|
|
|
|
|
fapply adjointify,
|
|
|
|
|
{ exact from_hom },
|
|
|
|
|
{ intro gh, induction gh with g h,
|
|
|
|
|
exact prod_eq (mul_one (1 * g) ⬝ one_mul g) (ap (λ o, o * h) (mul_one 1) ⬝ one_mul h) },
|
|
|
|
|
{ refine dirsum.rec _ _ _,
|
|
|
|
|
{ intro b x,
|
|
|
|
|
refine ap from_hom (dirsum_elim_compute (bool.rec (product_inl G H) (product_inr G H)) b x) ⬝ _,
|
|
|
|
|
induction b,
|
|
|
|
|
{ exact ap (λ y, dirsum_incl branch bool.ff x * y) (to_respect_one (dirsum_incl branch bool.tt)) ⬝ mul_one _ },
|
|
|
|
|
{ exact ap (λ y, y * dirsum_incl branch bool.tt x) (to_respect_one (dirsum_incl branch bool.ff)) ⬝ one_mul _ }
|
|
|
|
|
},
|
|
|
|
|
{ refine ap from_hom (to_respect_one to_hom) ⬝ to_respect_one from_hom },
|
|
|
|
|
{ intro g h gβ hβ,
|
|
|
|
|
refine ap from_hom (to_respect_mul to_hom _ _) ⬝ to_respect_mul from_hom _ _ ⬝ _,
|
|
|
|
|
exact ap011 mul gβ hβ
|
|
|
|
|
}
|
|
|
|
|
}
|
|
|
|
|
end
|
|
|
|
|
|
2017-06-09 16:08:21 +00:00
|
|
|
|
variables {I J : Type} [is_set I] [is_set J] {Y Y' Y'' : I → AbGroup}
|
2016-10-13 19:04:57 +00:00
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_functor [constructor] (f : Πi, Y i →g Y' i) : dirsum Y →g dirsum Y' :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
dirsum_elim (λi, dirsum_incl Y' i ∘g f i)
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
theorem dirsum_functor_compose (f' : Πi, Y' i →g Y'' i) (f : Πi, Y i →g Y' i) :
|
2017-06-08 23:48:26 +00:00
|
|
|
|
dirsum_functor f' ∘g dirsum_functor f ~ dirsum_functor (λi, f' i ∘g f i) :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
begin
|
|
|
|
|
apply dirsum_homotopy,
|
|
|
|
|
intro i y, reflexivity,
|
|
|
|
|
end
|
|
|
|
|
|
|
|
|
|
variable (Y)
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_functor_gid : dirsum_functor (λi, gid (Y i)) ~ gid (dirsum Y) :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
begin
|
|
|
|
|
apply dirsum_homotopy,
|
|
|
|
|
intro i y, reflexivity,
|
|
|
|
|
end
|
|
|
|
|
variable {Y}
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_functor_mul (f f' : Πi, Y i →g Y' i) :
|
|
|
|
|
homomorphism_mul (dirsum_functor f) (dirsum_functor f') ~
|
|
|
|
|
dirsum_functor (λi, homomorphism_mul (f i) (f' i)) :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
begin
|
|
|
|
|
apply dirsum_homotopy,
|
2017-06-07 15:39:33 +00:00
|
|
|
|
intro i y, exact sorry
|
2016-10-13 19:04:57 +00:00
|
|
|
|
end
|
|
|
|
|
|
2017-06-09 16:08:21 +00:00
|
|
|
|
definition dirsum_functor_homotopy (f f' : Πi, Y i →g Y' i) (p : f ~2 f') :
|
2017-03-31 22:21:02 +00:00
|
|
|
|
dirsum_functor f ~ dirsum_functor f' :=
|
|
|
|
|
begin
|
|
|
|
|
apply dirsum_homotopy,
|
|
|
|
|
intro i y, exact sorry
|
|
|
|
|
end
|
|
|
|
|
|
2017-06-07 05:01:19 +00:00
|
|
|
|
definition dirsum_functor_left [constructor] (f : J → I) : dirsum (Y ∘ f) →g dirsum Y :=
|
2017-03-31 22:21:02 +00:00
|
|
|
|
dirsum_elim (λj, dirsum_incl Y (f j))
|
|
|
|
|
|
2017-06-09 00:06:59 +00:00
|
|
|
|
definition dirsum_isomorphism [constructor] (f : Πi, Y i ≃g Y' i) : dirsum Y ≃g dirsum Y' :=
|
2017-06-08 23:51:25 +00:00
|
|
|
|
let to_hom := dirsum_functor (λ i, f i) in
|
|
|
|
|
let from_hom := dirsum_functor (λ i, (f i)⁻¹ᵍ) in
|
|
|
|
|
begin
|
|
|
|
|
fapply isomorphism.mk,
|
|
|
|
|
exact dirsum_functor (λ i, f i),
|
|
|
|
|
fapply is_equiv.adjointify,
|
|
|
|
|
exact dirsum_functor (λ i, (f i)⁻¹ᵍ),
|
|
|
|
|
{ intro ds,
|
2018-09-11 15:06:46 +00:00
|
|
|
|
refine (homomorphism_compose_eq (dirsum_functor (λ i, f i)) (dirsum_functor (λ i, (f i)⁻¹ᵍ)) _)⁻¹ ⬝ _,
|
2017-06-08 23:51:25 +00:00
|
|
|
|
refine dirsum_functor_compose (λ i, f i) (λ i, (f i)⁻¹ᵍ) ds ⬝ _,
|
2017-06-09 16:08:21 +00:00
|
|
|
|
refine dirsum_functor_homotopy _ (λ i, !gid) (λ i, to_right_inv (equiv_of_isomorphism (f i))) ds ⬝ _,
|
2017-06-08 23:51:25 +00:00
|
|
|
|
exact !dirsum_functor_gid
|
|
|
|
|
},
|
|
|
|
|
{ intro ds,
|
2018-09-11 15:06:46 +00:00
|
|
|
|
refine (homomorphism_compose_eq (dirsum_functor (λ i, (f i)⁻¹ᵍ)) (dirsum_functor (λ i, f i)) _)⁻¹ ⬝ _,
|
2017-06-08 23:51:25 +00:00
|
|
|
|
refine dirsum_functor_compose (λ i, (f i)⁻¹ᵍ) (λ i, f i) ds ⬝ _,
|
2017-06-09 16:08:21 +00:00
|
|
|
|
refine dirsum_functor_homotopy _ (λ i, !gid) (λ i x,
|
2017-06-08 23:51:25 +00:00
|
|
|
|
proof
|
|
|
|
|
to_left_inv (equiv_of_isomorphism (f i)) x
|
|
|
|
|
qed
|
|
|
|
|
) ds ⬝ _,
|
|
|
|
|
exact !dirsum_functor_gid
|
|
|
|
|
}
|
|
|
|
|
end
|
|
|
|
|
|
2016-10-13 19:04:57 +00:00
|
|
|
|
end group
|
2017-06-09 16:08:21 +00:00
|
|
|
|
|
|
|
|
|
namespace group
|
|
|
|
|
|
2017-06-09 21:00:56 +00:00
|
|
|
|
definition dirsum_down_left.{u v w} {I : Type.{u}} [is_set I] (Y : I → AbGroup.{w})
|
2017-06-09 16:08:21 +00:00
|
|
|
|
: dirsum (Y ∘ down.{u v}) ≃g dirsum Y :=
|
2017-06-09 21:00:56 +00:00
|
|
|
|
proof
|
2017-06-09 16:08:21 +00:00
|
|
|
|
let to_hom := @dirsum_functor_left _ _ _ _ Y down.{u v} in
|
2017-06-09 21:00:56 +00:00
|
|
|
|
let from_hom := dirsum_elim (λi, dirsum_incl (Y ∘ down.{u v}) (up.{u v} i)) in
|
2017-06-09 16:08:21 +00:00
|
|
|
|
begin
|
|
|
|
|
fapply isomorphism.mk,
|
|
|
|
|
{ exact to_hom },
|
|
|
|
|
fapply adjointify,
|
|
|
|
|
{ exact from_hom },
|
|
|
|
|
{ intro ds,
|
2018-09-11 15:06:46 +00:00
|
|
|
|
refine (homomorphism_compose_eq to_hom from_hom ds)⁻¹ ⬝ _,
|
2017-06-09 16:08:21 +00:00
|
|
|
|
refine @dirsum_homotopy I _ Y (dirsum Y) (to_hom ∘g from_hom) !gid _ ds,
|
|
|
|
|
intro i y,
|
2018-09-11 15:06:46 +00:00
|
|
|
|
refine homomorphism_compose_eq to_hom from_hom _ ⬝ _,
|
2017-06-09 21:00:56 +00:00
|
|
|
|
refine ap to_hom (dirsum_elim_compute (λi, dirsum_incl (Y ∘ down.{u v}) (up.{u v} i)) i y) ⬝ _,
|
|
|
|
|
refine dirsum_elim_compute _ (up.{u v} i) y ⬝ _,
|
2017-06-09 16:08:21 +00:00
|
|
|
|
reflexivity
|
|
|
|
|
},
|
|
|
|
|
{ intro ds,
|
2018-09-11 15:06:46 +00:00
|
|
|
|
refine (homomorphism_compose_eq from_hom to_hom ds)⁻¹ ⬝ _,
|
2017-06-09 21:00:56 +00:00
|
|
|
|
refine @dirsum_homotopy _ _ (Y ∘ down.{u v}) (dirsum (Y ∘ down.{u v})) (from_hom ∘g to_hom) !gid _ ds,
|
2017-06-09 16:08:21 +00:00
|
|
|
|
intro i y, induction i with i,
|
2018-09-11 15:06:46 +00:00
|
|
|
|
refine homomorphism_compose_eq from_hom to_hom _ ⬝ _,
|
2017-06-09 21:00:56 +00:00
|
|
|
|
refine ap from_hom (dirsum_elim_compute (λi, dirsum_incl Y (down.{u v} i)) (up.{u v} i) y) ⬝ _,
|
2017-06-09 16:08:21 +00:00
|
|
|
|
refine dirsum_elim_compute _ i y ⬝ _,
|
|
|
|
|
reflexivity
|
|
|
|
|
}
|
|
|
|
|
end
|
2017-06-09 21:00:56 +00:00
|
|
|
|
qed
|
2017-06-09 16:08:21 +00:00
|
|
|
|
|
|
|
|
|
end group
|