Currying alternating forms #
In this file we define AlternatingMap.curryLeft
which interprets an alternating map in n + 1
variables
as a linear map in the 0th variable taking values in the alternating maps in n
variables.
def
AlternatingMap.curryLeft
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[Module R M]
[Module R N]
{n : ℕ}
(f : M [⋀^Fin n.succ]→ₗ[R] N)
:
Given an alternating map f
in n+1
variables, split the first variable to obtain
a linear map into alternating maps in n
variables, given by x ↦ (m ↦ f (Matrix.vecCons x m))
.
It can be thought of as a map $Hom(\bigwedge^{n+1} M, N) \to Hom(M, Hom(\bigwedge^n M, N))$.
This is MultilinearMap.curryLeft
for AlternatingMap
. See also
AlternatingMap.curryLeftLinearMap
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
AlternatingMap.curryLeft_apply_apply
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[Module R M]
[Module R N]
{n : ℕ}
(f : M [⋀^Fin n.succ]→ₗ[R] N)
(m : M)
(v : Fin n → M)
:
@[simp]
theorem
AlternatingMap.curryLeft_zero
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[Module R M]
[Module R N]
{n : ℕ}
:
def
AlternatingMap.curryLeftLinearMap
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[Module R M]
[Module R N]
{n : ℕ}
:
AlternatingMap.curryLeft
as a LinearMap
. This is a separate definition as dot notation
does not work for this version.
Equations
Instances For
@[simp]
theorem
AlternatingMap.curryLeftLinearMap_apply
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[Module R M]
[Module R N]
{n : ℕ}
(f : M [⋀^Fin n.succ]→ₗ[R] N)
:
@[simp]
theorem
AlternatingMap.curryLeft_same
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[Module R M]
[Module R N]
{n : ℕ}
(f : M [⋀^Fin n.succ.succ]→ₗ[R] N)
(m : M)
:
Currying with the same element twice gives the zero map.
@[simp]
theorem
AlternatingMap.curryLeft_compAlternatingMap
{R : Type u_1}
{M : Type u_2}
{N : Type u_4}
{N₂ : Type u_5}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid N]
[AddCommMonoid N₂]
[Module R M]
[Module R N]
[Module R N₂]
{n : ℕ}
(g : N →ₗ[R] N₂)
(f : M [⋀^Fin n.succ]→ₗ[R] N)
(m : M)
:
@[simp]
theorem
AlternatingMap.curryLeft_compLinearMap
{R : Type u_1}
{M : Type u_2}
{M₂ : Type u_3}
{N : Type u_4}
[CommSemiring R]
[AddCommMonoid M]
[AddCommMonoid M₂]
[AddCommMonoid N]
[Module R M]
[Module R M₂]
[Module R N]
{n : ℕ}
(g : M₂ →ₗ[R] M)
(f : M [⋀^Fin n.succ]→ₗ[R] N)
(m : M₂)
: