Multilinear maps
We define multilinear maps as maps from Π(i : ι), M₁ i
to M₂
which are linear in each
coordinate. Here, M₁ i
and M₂
are modules over a ring R
, and ι
is an arbitrary type
(although some statements will require it to be a fintype). This space, denoted by
multilinear_map R M₁ M₂
, inherits a module structure by pointwise addition and multiplication.
Main definitions
multilinear_map R M₁ M₂
is the space of multilinear maps fromΠ(i : ι), M₁ i
toM₂
.f.map_smul
is the multiplicativity of the multilinear mapf
along each coordinate.f.map_add
is the additivity of the multilinear mapf
along each coordinate.f.map_smul_univ
expresses the multiplicativity off
over all coordinates at the same time, writingf (λi, c i • m i)
as(∏ i, c i) • f m
.f.map_add_univ
expresses the additivity off
over all coordinates at the same time, writingf (m + m')
as the sum over all subsetss
ofι
off (s.piecewise m m')
.f.map_sum
expressesf (Σ_{j₁} g₁ j₁, ..., Σ_{jₙ} gₙ jₙ)
as the sum off (g₁ (r 1), ..., gₙ (r n))
wherer
ranges over all possible functions.
We also register isomorphisms corresponding to currying or uncurrying variables, transforming a
multilinear function f
on n+1
variables into a linear function taking values in multilinear
functions in n
variables, and into a multilinear function in n
variables taking values in linear
functions. These operations are called f.curry_left
and f.curry_right
respectively
(with inverses f.uncurry_left
and f.uncurry_right
). These operations induce linear equivalences
between spaces of multilinear functions in n+1
variables and spaces of linear functions into
multilinear functions in n
variables (resp. multilinear functions in n
variables taking values
in linear functions), called respectively multilinear_curry_left_equiv
and
multilinear_curry_right_equiv
.
Implementation notes
Expressing that a map is linear along the i
-th coordinate when all other coordinates are fixed
can be done in two (equivalent) different ways:
- fixing a vector
m : Π(j : ι - i), M₁ j.val
, and then choosing separately thei
-th coordinate - fixing a vector
m : Πj, M₁ j
, and then modifying itsi
-th coordinate The second way is more artificial as the value ofm
ati
is not relevant, but it has the advantage of avoiding subtype inclusion issues. This is the definition we use, based onfunction.update
that allows to change the value ofm
ati
.
- to_fun : (Π (i : ι), M₁ i) → M₂
- map_add' : ∀ (m : Π (i : ι), M₁ i) (i : ι) (x y : M₁ i), c.to_fun (function.update m i (x + y)) = c.to_fun (function.update m i x) + c.to_fun (function.update m i y)
- map_smul' : ∀ (m : Π (i : ι), M₁ i) (i : ι) (c_1 : R) (x : M₁ i), c.to_fun (function.update m i (c_1 • x)) = c_1 • c.to_fun (function.update m i x)
Multilinear maps over the ring R
, from Πi, M₁ i
to M₂
where M₁ i
and M₂
are modules
over R
.
Equations
- multilinear_map.has_coe_to_fun = {F := λ (x : multilinear_map R M₁ M₂), (Π (i : ι), M₁ i) → M₂, coe := multilinear_map.to_fun _inst_10}
Equations
Equations
- multilinear_map.add_comm_monoid = {add := has_add.add multilinear_map.has_add, add_assoc := _, zero := 0, zero_add := _, add_zero := _, add_comm := _}
If f
is a multilinear map, then f.to_linear_map m i
is the linear map obtained by fixing all
coordinates but i
equal to those of m
, and varying the i
-th coordinate.
Equations
- f.to_linear_map m i = {to_fun := λ (x : M₁ i), ⇑f (function.update m i x), map_add' := _, map_smul' := _}
The cartesian product of two multilinear maps, as a multilinear map.
Given a multilinear map f
on n
variables (parameterized by fin n
) and a subset s
of k
of these variables, one gets a new multilinear map on fin k
by varying these variables, and fixing
the other ones equal to a given value z
. It is denoted by f.restr s hk z
, where hk
is a
proof that the cardinality of s
is k
. The implicit identification between fin k
and s
that
we use is the canonical (increasing) bijection.
In the specific case of multilinear maps on spaces indexed by fin (n+1)
, where one can build
an element of Π(i : fin (n+1)), M i
using cons
, one can express directly the additivity of a
multilinear map along the first variable.
In the specific case of multilinear maps on spaces indexed by fin (n+1)
, where one can build
an element of Π(i : fin (n+1)), M i
using cons
, one can express directly the multiplicativity
of a multilinear map along the first variable.
In the specific case of multilinear maps on spaces indexed by fin (n+1)
, where one can build
an element of Π(i : fin (n+1)), M i
using snoc
, one can express directly the additivity of a
multilinear map along the first variable.
In the specific case of multilinear maps on spaces indexed by fin (n+1)
, where one can build
an element of Π(i : fin (n+1)), M i
using cons
, one can express directly the multiplicativity
of a multilinear map along the first variable.
If g
is multilinear and f
is linear, then g (f m₁, ..., f mₙ)
is again a multilinear
function, that we call g.comp_linear_map f
.
If one adds to a vector m'
another vector m
, but only for coordinates in a finset t
, then
the image under a multilinear map f
is the sum of f (s.piecewise m m')
along all subsets s
of
t
. This is mainly an auxiliary statement to prove the result when t = univ
, given in
map_add_univ
, although it can be useful in its own right as it does not require the index set ι
to be finite.
Additivity of a multilinear map along all coordinates at the same time,
writing f (m + m')
as the sum of f (s.piecewise m m')
over all sets s
.
If f
is multilinear, then f (Σ_{j₁ ∈ A₁} g₁ j₁, ..., Σ_{jₙ ∈ Aₙ} gₙ jₙ)
is the sum of
f (g₁ (r 1), ..., gₙ (r n))
where r
ranges over all functions with r 1 ∈ A₁
, ...,
r n ∈ Aₙ
. This follows from multilinearity by expanding successively with respect to each
coordinate. Here, we give an auxiliary statement tailored for an inductive proof. Use instead
map_sum_finset
.
If f
is multilinear, then f (Σ_{j₁ ∈ A₁} g₁ j₁, ..., Σ_{jₙ ∈ Aₙ} gₙ jₙ)
is the sum of
f (g₁ (r 1), ..., gₙ (r n))
where r
ranges over all functions with r 1 ∈ A₁
, ...,
r n ∈ Aₙ
. This follows from multilinearity by expanding successively with respect to each
coordinate.
If f
is multilinear, then f (Σ_{j₁} g₁ j₁, ..., Σ_{jₙ} gₙ jₙ)
is the sum of
f (g₁ (r 1), ..., gₙ (r n))
where r
ranges over all functions r
. This follows from
multilinearity by expanding successively with respect to each coordinate.
If one multiplies by c i
the coordinates in a finset s
, then the image under a multilinear
map is multiplied by ∏ i in s, c i
. This is mainly an auxiliary statement to prove the result when
s = univ
, given in map_smul_univ
, although it can be useful in its own right as it does not
require the index set ι
to be finite.
Multiplicativity of a multilinear map along all coordinates at the same time,
writing f (λi, c i • m i)
as (∏ i, c i) • f m
.
The canonical multilinear map on R^ι
when ι
is finite, associating to m
the product of
all the m i
(multiplied by a fixed reference element z
in the target module)
Equations
- multilinear_map.mk_pi_ring R ι z = {to_fun := λ (m : ι → R), finset.univ.prod (λ (i : ι), m i) • z, map_add' := _, map_smul' := _}
Equations
- multilinear_map.add_comm_group = {add := has_add.add multilinear_map.has_add, add_assoc := _, zero := 0, zero_add := _, add_zero := _, neg := has_neg.neg multilinear_map.has_neg, add_left_neg := _, add_comm := _}
The space of multilinear maps is a module over R
, for the pointwise addition and scalar
multiplication.
Equations
- multilinear_map.semimodule R ι M₁ M₂ = semimodule.of_core {to_has_scalar := {smul := has_scalar.smul multilinear_map.has_scalar}, smul_add := _, add_smul := _, mul_smul := _, one_smul := _}
Equations
- multilinear_map.semimodule_ring R ι M₂ = multilinear_map.semimodule R ι (λ (i : ι), R) M₂
When ι
is finite, multilinear maps on R^ι
with values in M₂
are in bijection with M₂
,
as such a multilinear map is completely determined by its value on the constant vector made of ones.
We register this bijection as a linear equivalence in multilinear_map.pi_ring_equiv
.
Equations
- multilinear_map.pi_ring_equiv R ι M₂ = {to_fun := λ (z : M₂), multilinear_map.mk_pi_ring R ι z, map_add' := _, map_smul' := _, inv_fun := λ (f : multilinear_map R (λ (i : ι), R) M₂), ⇑f (λ (i : ι), 1), left_inv := _, right_inv := _}
Composing a multilinear map with a linear map gives again a multilinear map.
Currying
We associate to a multilinear map in n+1
variables (i.e., based on fin n.succ
) two
curried functions, named f.curry_left
(which is a linear map on E 0
taking values
in multilinear maps in n
variables) and f.curry_right
(wich is a multilinear map in n
variables taking values in linear maps on E 0
). In both constructions, the variable that is
singled out is 0
, to take advantage of the operations cons
and tail
on fin n
.
The inverse operations are called uncurry_left
and uncurry_right
.
We also register linear equiv versions of these correspondences, in
multilinear_curry_left_equiv
and multilinear_curry_right_equiv
.
Left currying
Given a linear map f
from M 0
to multilinear maps on n
variables,
construct the corresponding multilinear map on n+1
variables obtained by concatenating
the variables, given by m ↦ f (m 0) (tail m)
Given a multilinear map f
in n+1
variables, split the first variable to obtain
a linear map into multilinear maps in n
variables, given by x ↦ (m ↦ f (cons x m))
.
The space of multilinear maps on Π(i : fin (n+1)), M i
is canonically isomorphic to
the space of linear maps from M 0
to the space of multilinear maps on
Π(i : fin n), M i.succ
, by separating the first variable. We register this isomorphism as a
linear isomorphism in multilinear_curry_left_equiv R M M₂
.
The direct and inverse maps are given by f.uncurry_left
and f.curry_left
. Use these
unless you need the full framework of linear equivs.
Equations
- multilinear_curry_left_equiv R M M₂ = {to_fun := linear_map.uncurry_left _inst_8, map_add' := _, map_smul' := _, inv_fun := multilinear_map.curry_left _inst_8, left_inv := _, right_inv := _}
Right currying
Given a multilinear map f
in n
variables to the space of linear maps from M (last n)
to
M₂
, construct the corresponding multilinear map on n+1
variables obtained by concatenating
the variables, given by m ↦ f (init m) (m (last n))
Given a multilinear map f
in n+1
variables, split the last variable to obtain
a multilinear map in n
variables taking values in linear maps from M (last n)
to M₂
, given by
m ↦ (x ↦ f (snoc m x))
.
The space of multilinear maps on Π(i : fin (n+1)), M i
is canonically isomorphic to
the space of linear maps from the space of multilinear maps on Π(i : fin n), M i.cast_succ
to the
space of linear maps on M (last n)
, by separating the last variable. We register this isomorphism
as a linear isomorphism in multilinear_curry_right_equiv R M M₂
.
The direct and inverse maps are given by f.uncurry_right
and f.curry_right
. Use these
unless you need the full framework of linear equivs.
Equations
- multilinear_curry_right_equiv R M M₂ = {to_fun := multilinear_map.uncurry_right _inst_8, map_add' := _, map_smul' := _, inv_fun := multilinear_map.curry_right _inst_8, left_inv := _, right_inv := _}