module Algebra.Group.Ab.Tensor where
Bilinear maps🔗
A function , where all types involved are equipped with abelian group structures, is called bilinear when it satisfies and : it is a group homomorphism in each of its arguments.
record Bilinear (A : Abelian-group ℓ) (B : Abelian-group ℓ') (C : Abelian-group ℓ'') : Type (ℓ ⊔ ℓ' ⊔ ℓ'') where
private
module A = Abelian-group-on (A .snd)
module B = Abelian-group-on (B .snd)
module C = Abelian-group-on (C .snd)
field
: ⌞ A ⌟ → ⌞ B ⌟ → ⌞ C ⌟
map : ∀ x y z → map (x A.* y) z ≡ map x z C.* map y z
pres-*l : ∀ x y z → map x (y B.* z) ≡ map x y C.* map x z pres-*r
We have already noted that the set of group homomorphisms between a pair of abelian groups is an abelian group, under pointwise multiplication. The type of bilinear maps is equivalent to the type of group homomorphisms .
: Bilinear A B C → Ab.Hom A Ab[ B , C ]
curry-bilinear .hom a .hom = f .Bilinear.map a
curry-bilinear f .hom a .preserves = Bilinear.fixl-is-group-hom f a
curry-bilinear f .preserves .is-group-hom.pres-⋆ x y =
curry-bilinear f λ z → f .Bilinear.pres-*l _ _ _
Homomorphism-path
: is-equiv curry-bilinear
curry-bilinear-is-equiv = is-iso→is-equiv morp where
curry-bilinear-is-equiv : is-iso curry-bilinear
morp .is-iso.inv uc .Bilinear.map x y = uc # x # y
morp .is-iso.inv uc .Bilinear.pres-*l x y z = ap (_# _) (uc .preserves .is-group-hom.pres-⋆ _ _)
morp .is-iso.inv uc .Bilinear.pres-*r x y z = (uc # _) .preserves .is-group-hom.pres-⋆ _ _
morp .is-iso.rinv uc = trivial!
morp .is-iso.linv uc = Bilinear-path λ x y → refl morp
The tensor product🔗
Thinking about the currying isomorphism
,
we set out to search for an abelian group which lets us “associate”
Bilinear
in the other direction:
Bilinear maps
are equivalent to group homomorphisms
,
but is there a construction
“”,
playing the role of product type, such that
is also the type of bilinear maps
?
The answer is yes: rather than , we write this infix as , and refer to it as the tensor product of abelian groups. Since is determined by the maps out of it, we can construct it directly as a higher inductive type. We add a constructor for every operation we want, and for the equations these should satisfy: should be an Abelian group, and it should admit a bilinear map .
data Tensor : Type (ℓ ⊔ ℓ') where
: Tensor
:1 _:*_ : Tensor → Tensor → Tensor
: Tensor → Tensor
:inv : is-set Tensor
squash : ∀ {x} → :inv x :* x ≡ :1
t-invl : ∀ {x} → :1 :* x ≡ x
t-idl : ∀ {x y z} → x :* (y :* z) ≡ (x :* y) :* z
t-assoc : ∀ {x y} → x :* y ≡ y :* x
t-comm
_,_ : ⌞ A ⌟ → ⌞ B ⌟ → Tensor
: ∀ {x y z} → (x , y B.* z) ≡ (x , y) :* (x , z)
t-pres-*r : ∀ {x y z} → (x A.* y , z) ≡ (x , z) :* (y , z) t-pres-*l
The first 8 constructors are, by definition, exactly what we need to
make Tensor
into an abelian group.
The latter three are the bilinear map
.
Since this is an inductive type, it’s the initial object equipped with
these data.
open make-abelian-group
: make-abelian-group Tensor
make-abelian-tensor .ab-is-set = squash
make-abelian-tensor .mul = _:*_
make-abelian-tensor .inv = :inv
make-abelian-tensor .1g = :1
make-abelian-tensor .idl x = t-idl
make-abelian-tensor .assoc x y z = t-assoc
make-abelian-tensor .invl x = t-invl
make-abelian-tensor .comm x y = t-comm
make-abelian-tensor
_⊗_ : Abelian-group (ℓ ⊔ ℓ')
_⊗_ = to-ab make-abelian-tensor
: Bilinear A B _⊗_
to-tensor .Bilinear.map = _,_
to-tensor .Bilinear.pres-*l x y z = t-pres-*l
to-tensor .Bilinear.pres-*r x y z = t-pres-*r to-tensor
If we have any bilinear map , we can first extend it to a function of sets by recursion, then prove that this is a group homomorphism. It turns out that this extension is definitionally a group homomorphism.
: Bilinear A B C → Tensor A B → ⌞ C ⌟
bilinear-map→function = go module bilinear-map→function where
bilinear-map→function bilin : Tensor A B → ⌞ C ⌟
go = C.1g
go :1 (x :* y) = go x C.* go y
go (:inv x) = go x C.⁻¹
go (x , y) = Bilinear.map bilin x y
go
(squash x y p q i j) = C.has-is-set (go x) (go y) (λ i → go (p i)) (λ i → go (q i)) i j
go (t-invl {x} i) = C.inversel {x = go x} i
go (t-idl {x} i) = C.idl {x = go x} i
go (t-assoc {x} {y} {z} i) = C.associative {x = go x} {go y} {go z} i
go (t-comm {x} {y} i) = C.commutes {x = go x} {y = go y} i
go (t-pres-*r {a} {b} {c} i) = Bilinear.pres-*r bilin a b c i
go (t-pres-*l {a} {b} {c} i) = Bilinear.pres-*l bilin a b c i
go
{-# DISPLAY bilinear-map→function.go b = bilinear-map→function b #-}
: Bilinear A B C → Ab.Hom (A ⊗ B) C
from-bilinear-map .hom = bilinear-map→function A B C bilin
from-bilinear-map bilin .preserves .is-group-hom.pres-⋆ x y = refl from-bilinear-map bilin
It’s also easy to construct a function in the opposite direction,
turning group homomorphisms into bilinear maps, but proving that this is
an equivalence requires appealing to an induction principle of Tensor
which handles the equational
fields: Tensor-elim-prop
.
: Ab.Hom (A ⊗ B) C → Bilinear A B C
to-bilinear-map .Bilinear.map x y = gh # (x , y)
to-bilinear-map gh .Bilinear.pres-*l x y z =
to-bilinear-map gh (apply gh) t-pres-*l ∙ gh .preserves .is-group-hom.pres-⋆ _ _
ap .Bilinear.pres-*r x y z =
to-bilinear-map gh (apply gh) t-pres-*r ∙ gh .preserves .is-group-hom.pres-⋆ _ _
ap
: is-equiv from-bilinear-map
from-bilinear-map-is-equiv = is-iso→is-equiv morp where
from-bilinear-map-is-equiv : is-iso from-bilinear-map
morp .is-iso.inv = to-bilinear-map
morp .is-iso.rinv hom = Homomorphism-path $ Tensor-elim-prop A B (λ x → C.has-is-set _ _)
morp (λ x y → refl)
(λ x y → ap₂ C._*_ x y ∙ sym (hom .preserves .is-group-hom.pres-⋆ _ _))
(λ x → ap C._⁻¹ x ∙ sym (is-group-hom.pres-inv (hom .preserves)))
(sym (is-group-hom.pres-id (hom .preserves)))
.is-iso.linv x = Bilinear-path (λ x y → refl) morp
The tensor-hom adjunction🔗
Since we have a construction satisfying , we’re driven, being category theorists, to question its naturality: Is the tensor product a functor, and is this equivalence of Hom-sets an adjunction?
The answer is yes, and the proofs are essentially plugging together existing definitions, other than the construction of the functorial action of .
: Functor (Ab ℓ ×ᶜ Ab ℓ) (Ab ℓ)
Ab-tensor-functor .F₀ (A , B) = A ⊗ B
Ab-tensor-functor .F₁ (f , g) = from-bilinear-map _ _ _ bilin where
Ab-tensor-functor : Bilinear _ _ _
bilin .Bilinear.map x y = f # x , g # y
bilin .Bilinear.pres-*l x y z = ap (_, g # z) (f .preserves .is-group-hom.pres-⋆ _ _) ∙ t-pres-*l
bilin .Bilinear.pres-*r x y z = ap (f # x ,_) (g .preserves .is-group-hom.pres-⋆ _ _) ∙ t-pres-*r
bilin .F-id = Hom≃Bilinear.injective _ _ _ (Bilinear-path λ x y → refl)
Ab-tensor-functor .F-∘ f g = Hom≃Bilinear.injective _ _ _ (Bilinear-path λ x y → refl)
Ab-tensor-functor
: (A : Abelian-group ℓ) → Bifunctor.Left Ab-tensor-functor A ⊣ Bifunctor.Right Ab-hom-functor A
Tensor⊣Hom = hom-iso→adjoints to to-eqv nat where
Tensor⊣Hom A to : ∀ {x y} → Ab.Hom (x ⊗ A) y → Ab.Hom x Ab[ A , y ]
to f = curry-bilinear _ _ _ (to-bilinear-map _ _ _ f)
: ∀ {x y} → is-equiv (to {x} {y})
to-eqv = ∙-is-equiv
to-eqv (Hom≃Bilinear _ _ _ .snd)
(curry-bilinear-is-equiv _ _ _)
: hom-iso-natural {L = Bifunctor.Left Ab-tensor-functor A} {R = Bifunctor.Right Ab-hom-functor A} to
nat = trivial! nat f g h