2024-04-16 15:34:00 -04:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license.
|
|
|
|
|
Authors: Joseph Tooby-Smith
|
|
|
|
|
-/
|
|
|
|
|
import Mathlib.Tactic.Polyrith
|
|
|
|
|
import Mathlib.Algebra.Module.LinearMap.Basic
|
2024-04-19 09:57:30 -04:00
|
|
|
|
import Mathlib.Algebra.BigOperators.Fin
|
2024-04-16 15:34:00 -04:00
|
|
|
|
/-!
|
|
|
|
|
# Linear maps
|
|
|
|
|
|
|
|
|
|
Some definitions and properites of linear, bilinear, and trilinear maps, along with homogeneous
|
|
|
|
|
quadratic and cubic equations.
|
|
|
|
|
|
|
|
|
|
## TODO
|
|
|
|
|
|
|
|
|
|
Use definitions in `Mathlib4` for definitions where possible.
|
2024-04-17 06:23:48 -04:00
|
|
|
|
In particular a HomogeneousQuadratic should be a map `V →ₗ[ℚ] V →ₗ[ℚ] ℚ` etc.
|
2024-04-16 15:34:00 -04:00
|
|
|
|
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
/-- The structure defining a homogeneous quadratic equation. -/
|
|
|
|
|
structure HomogeneousQuadratic (V : Type) [AddCommMonoid V] [Module ℚ V] where
|
|
|
|
|
/-- The quadratic equation. -/
|
|
|
|
|
toFun : V → ℚ
|
|
|
|
|
/-- The equation is homogeneous. -/
|
|
|
|
|
map_smul' : ∀ a S, toFun (a • S) = a ^ 2 * toFun S
|
|
|
|
|
|
|
|
|
|
namespace HomogeneousQuadratic
|
|
|
|
|
|
|
|
|
|
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
|
|
|
|
|
instance instFun : FunLike (HomogeneousQuadratic V) V ℚ where
|
|
|
|
|
coe f := f.toFun
|
|
|
|
|
coe_injective' f g h := by
|
|
|
|
|
cases f
|
|
|
|
|
cases g
|
|
|
|
|
simp_all
|
|
|
|
|
|
2024-04-17 10:25:05 -04:00
|
|
|
|
/-- Given an equivariant function from `V` to `ℚ` we get a `HomogeneousQuadratic V`. -/
|
|
|
|
|
def fromEquivariant (f : V →ₑ[((fun a => a ^ 2) : ℚ → ℚ)] ℚ) : HomogeneousQuadratic V where
|
|
|
|
|
toFun := f
|
|
|
|
|
map_smul' a S := by
|
|
|
|
|
rw [map_smulₛₗ]
|
|
|
|
|
rfl
|
|
|
|
|
|
2024-04-16 15:34:00 -04:00
|
|
|
|
lemma map_smul (f : HomogeneousQuadratic V) (a : ℚ) (S : V) : f (a • S) = a ^ 2 * f S :=
|
|
|
|
|
f.map_smul' a S
|
|
|
|
|
|
|
|
|
|
end HomogeneousQuadratic
|
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The structure of a bilinear map. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
structure BiLinear (V : Type) [AddCommMonoid V] [Module ℚ V] where
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The underling function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
toFun : V × V → ℚ
|
|
|
|
|
map_smul₁' : ∀ a S T, toFun (a • S, T) = a * toFun (S, T)
|
|
|
|
|
map_smul₂' : ∀ a S T , toFun (S, a • T) = a * toFun (S, T)
|
|
|
|
|
map_add₁' : ∀ S1 S2 T, toFun (S1 + S2, T) = toFun (S1, T) + toFun (S2, T)
|
|
|
|
|
map_add₂' : ∀ S T1 T2, toFun (S, T1 + T2) = toFun (S, T1) + toFun (S, T2)
|
|
|
|
|
|
|
|
|
|
namespace BiLinear
|
|
|
|
|
|
|
|
|
|
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
|
|
|
|
|
instance instFun : FunLike (BiLinear V) (V × V) ℚ where
|
|
|
|
|
coe f := f.toFun
|
|
|
|
|
coe_injective' f g h := by
|
|
|
|
|
cases f
|
|
|
|
|
cases g
|
|
|
|
|
simp_all
|
|
|
|
|
|
2024-04-17 10:25:05 -04:00
|
|
|
|
/-- A bilinear map from a linear function ` V →ₗ[ℚ] V →ₗ[ℚ] ℚ`. -/
|
|
|
|
|
def fromLinearToHom (f : V →ₗ[ℚ] V →ₗ[ℚ] ℚ) : BiLinear V where
|
|
|
|
|
toFun S := f S.1 S.2
|
|
|
|
|
map_smul₁' a S T := by
|
|
|
|
|
simp only [LinearMapClass.map_smul]
|
|
|
|
|
rfl
|
|
|
|
|
map_smul₂' a S T := (f S).map_smul a T
|
|
|
|
|
map_add₁' S1 S2 T := by
|
|
|
|
|
simp only [f.map_add]
|
|
|
|
|
rfl
|
|
|
|
|
map_add₂' S T1 T2 := (f S).map_add T1 T2
|
|
|
|
|
|
2024-04-16 15:34:00 -04:00
|
|
|
|
lemma map_smul₁ (f : BiLinear V) (a : ℚ) (S T : V) : f (a • S, T) = a * f (S, T) :=
|
|
|
|
|
f.map_smul₁' a S T
|
|
|
|
|
|
|
|
|
|
lemma map_smul₂ (f : BiLinear V) (S : V) (a : ℚ) (T : V) : f (S, a • T) = a * f (S, T) :=
|
|
|
|
|
f.map_smul₂' a S T
|
|
|
|
|
|
|
|
|
|
lemma map_add₁ (f : BiLinear V) (S1 S2 T : V) : f (S1 + S2, T) = f (S1, T) + f (S2, T) :=
|
|
|
|
|
f.map_add₁' S1 S2 T
|
|
|
|
|
|
|
|
|
|
lemma map_add₂ (f : BiLinear V) (S : V) (T1 T2 : V) : f (S, T1 + T2) = f (S, T1) + f (S, T2) :=
|
|
|
|
|
f.map_add₂' S T1 T2
|
|
|
|
|
|
|
|
|
|
end BiLinear
|
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The structure of a symmetric bilinear function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
structure BiLinearSymm (V : Type) [AddCommMonoid V] [Module ℚ V] where
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The underlying function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
toFun : V × V → ℚ
|
|
|
|
|
map_smul₁' : ∀ a S T, toFun (a • S, T) = a * toFun (S, T)
|
|
|
|
|
map_add₁' : ∀ S1 S2 T, toFun (S1 + S2, T) = toFun (S1, T) + toFun (S2, T)
|
|
|
|
|
swap' : ∀ S T, toFun (S, T) = toFun (T, S)
|
|
|
|
|
|
2024-04-17 10:25:05 -04:00
|
|
|
|
/-- A symmetric bilinear function. -/
|
|
|
|
|
class IsSymmetric {V : Type} [AddCommMonoid V] [Module ℚ V] (f : V →ₗ[ℚ] V →ₗ[ℚ] ℚ) : Prop where
|
|
|
|
|
swap : ∀ S T, f S T = f T S
|
|
|
|
|
|
2024-04-16 15:34:00 -04:00
|
|
|
|
namespace BiLinearSymm
|
2024-04-19 09:57:30 -04:00
|
|
|
|
open BigOperators
|
2024-04-16 15:34:00 -04:00
|
|
|
|
|
|
|
|
|
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
|
|
|
|
|
instance instFun (V : Type) [AddCommMonoid V] [Module ℚ V] :
|
|
|
|
|
FunLike (BiLinearSymm V) (V × V) ℚ where
|
|
|
|
|
coe f := f.toFun
|
|
|
|
|
coe_injective' f g h := by
|
|
|
|
|
cases f
|
|
|
|
|
cases g
|
|
|
|
|
simp_all
|
|
|
|
|
|
2024-04-17 10:25:05 -04:00
|
|
|
|
/-- A bilinear symmetric function from a symemtric `V →ₗ[ℚ] V →ₗ[ℚ] ℚ`. -/
|
|
|
|
|
def fromSymmToHom (f : V →ₗ[ℚ] V →ₗ[ℚ] ℚ) [IsSymmetric f] : BiLinearSymm V where
|
|
|
|
|
toFun S := f S.1 S.2
|
|
|
|
|
map_smul₁' a S T := by
|
|
|
|
|
simp only [LinearMapClass.map_smul]
|
|
|
|
|
rfl
|
|
|
|
|
map_add₁' S1 S2 T := by
|
|
|
|
|
simp only [f.map_add]
|
|
|
|
|
rfl
|
|
|
|
|
swap' S T := IsSymmetric.swap S T
|
2024-04-16 15:34:00 -04:00
|
|
|
|
|
|
|
|
|
lemma toFun_eq_coe (f : BiLinearSymm V) : f.toFun = f := rfl
|
|
|
|
|
|
|
|
|
|
lemma map_smul₁ (f : BiLinearSymm V) (a : ℚ) (S T : V) : f (a • S, T) = a * f (S, T) :=
|
|
|
|
|
f.map_smul₁' a S T
|
|
|
|
|
|
|
|
|
|
lemma swap (f : BiLinearSymm V) (S T : V) : f (S, T) = f (T, S) :=
|
|
|
|
|
f.swap' S T
|
|
|
|
|
|
|
|
|
|
lemma map_smul₂ (f : BiLinearSymm V) (a : ℚ) (S : V) (T : V) : f (S, a • T) = a * f (S, T) := by
|
|
|
|
|
rw [f.swap, f.map_smul₁, f.swap]
|
|
|
|
|
|
|
|
|
|
lemma map_add₁ (f : BiLinearSymm V) (S1 S2 T : V) : f (S1 + S2, T) = f (S1, T) + f (S2, T) :=
|
|
|
|
|
f.map_add₁' S1 S2 T
|
|
|
|
|
|
|
|
|
|
lemma map_add₂ (f : BiLinearSymm V) (S : V) (T1 T2 : V) :
|
|
|
|
|
f (S, T1 + T2) = f (S, T1) + f (S, T2) := by
|
|
|
|
|
rw [f.swap, f.map_add₁, f.swap T1 S, f.swap T2 S]
|
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- Fixing the second input vectors, the resulting linear map. -/
|
2024-04-19 09:57:30 -04:00
|
|
|
|
def toLinear₁ (f : BiLinearSymm V) (T : V) : V →ₗ[ℚ] ℚ where
|
|
|
|
|
toFun S := f (S, T)
|
|
|
|
|
map_add' S1 S2 := by
|
|
|
|
|
simp only [f.map_add₁]
|
|
|
|
|
map_smul' a S := by
|
|
|
|
|
simp only [f.map_smul₁]
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma toLinear₁_apply (f : BiLinearSymm V) (S T : V) : f (S, T) = f.toLinear₁ T S := rfl
|
|
|
|
|
|
|
|
|
|
lemma map_sum₁ {n : ℕ} (f : BiLinearSymm V) (S : Fin n → V) (T : V) :
|
|
|
|
|
f (∑ i, S i, T) = ∑ i, f (S i, T) := by
|
|
|
|
|
rw [f.toLinear₁_apply]
|
|
|
|
|
rw [map_sum]
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma map_sum₂ {n : ℕ} (f : BiLinearSymm V) (S : Fin n → V) (T : V) :
|
|
|
|
|
f ( T, ∑ i, S i) = ∑ i, f (T, S i) := by
|
|
|
|
|
rw [swap, map_sum₁]
|
|
|
|
|
apply Fintype.sum_congr
|
|
|
|
|
intro i
|
|
|
|
|
rw [swap]
|
|
|
|
|
|
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The homogenous quadratic equation obtainable from a bilinear function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
@[simps!]
|
|
|
|
|
def toHomogeneousQuad {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
(τ : BiLinearSymm V) : HomogeneousQuadratic V where
|
|
|
|
|
toFun S := τ (S, S)
|
|
|
|
|
map_smul' a S := by
|
|
|
|
|
simp only
|
|
|
|
|
rw [τ.map_smul₁, τ.map_smul₂]
|
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
lemma toHomogeneousQuad_add {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
(τ : BiLinearSymm V) (S T : V) :
|
|
|
|
|
τ.toHomogeneousQuad (S + T) = τ.toHomogeneousQuad S +
|
|
|
|
|
τ.toHomogeneousQuad T + 2 * τ (S, T) := by
|
|
|
|
|
simp only [toHomogeneousQuad_toFun]
|
|
|
|
|
rw [τ.map_add₁, τ.map_add₂, τ.map_add₂, τ.swap T S]
|
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
end BiLinearSymm
|
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The structure of a homogeneous cubic equation. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
structure HomogeneousCubic (V : Type) [AddCommMonoid V] [Module ℚ V] where
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The underlying function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
toFun : V → ℚ
|
|
|
|
|
map_smul' : ∀ a S, toFun (a • S) = a ^ 3 * toFun S
|
|
|
|
|
|
|
|
|
|
namespace HomogeneousCubic
|
|
|
|
|
|
|
|
|
|
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
|
|
|
|
|
instance instFun : FunLike (HomogeneousCubic V) V ℚ where
|
|
|
|
|
coe f := f.toFun
|
|
|
|
|
coe_injective' f g h := by
|
|
|
|
|
cases f
|
|
|
|
|
cases g
|
|
|
|
|
simp_all
|
|
|
|
|
|
|
|
|
|
lemma map_smul (f : HomogeneousCubic V) (a : ℚ) (S : V) : f (a • S) = a ^ 3 * f S :=
|
|
|
|
|
f.map_smul' a S
|
|
|
|
|
|
|
|
|
|
end HomogeneousCubic
|
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The structure of a trilinear function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
structure TriLinear (V : Type) [AddCommMonoid V] [Module ℚ V] where
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The underlying function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
toFun : V × V × V → ℚ
|
|
|
|
|
map_smul₁' : ∀ a S T L, toFun (a • S, T, L) = a * toFun (S, T, L)
|
|
|
|
|
map_smul₂' : ∀ a S T L, toFun (S, a • T, L) = a * toFun (S, T, L)
|
|
|
|
|
map_smul₃' : ∀ a S T L, toFun (S, T, a • L) = a * toFun (S, T, L)
|
|
|
|
|
map_add₁' : ∀ S1 S2 T L, toFun (S1 + S2, T, L) = toFun (S1, T, L) + toFun (S2, T, L)
|
|
|
|
|
map_add₂' : ∀ S T1 T2 L, toFun (S, T1 + T2, L) = toFun (S, T1, L) + toFun (S, T2, L)
|
|
|
|
|
map_add₃' : ∀ S T L1 L2, toFun (S, T, L1 + L2) = toFun (S, T, L1) + toFun (S, T, L2)
|
|
|
|
|
|
|
|
|
|
namespace TriLinear
|
|
|
|
|
|
|
|
|
|
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
|
|
|
|
|
instance instFun : FunLike (TriLinear V) (V × V × V) ℚ where
|
|
|
|
|
coe f := f.toFun
|
|
|
|
|
coe_injective' f g h := by
|
|
|
|
|
cases f
|
|
|
|
|
cases g
|
|
|
|
|
simp_all
|
|
|
|
|
|
|
|
|
|
end TriLinear
|
2024-04-19 09:57:30 -04:00
|
|
|
|
/-- The structure of a symmetric trilinear function. -/
|
|
|
|
|
structure TriLinearSymm' (V : Type) [AddCommMonoid V] [Module ℚ V] extends
|
|
|
|
|
V →ₗ[ℚ] V →ₗ[ℚ] V →ₗ[ℚ] ℚ where
|
|
|
|
|
swap₁' : ∀ S T L, toFun S T L = toFun T S L
|
|
|
|
|
swap₂' : ∀ S T L, toFun S T L = toFun S L T
|
2024-04-16 15:34:00 -04:00
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The structure of a symmetric trilinear function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
structure TriLinearSymm (V : Type) [AddCommMonoid V] [Module ℚ V] where
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The underlying function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
toFun : V × V × V → ℚ
|
|
|
|
|
map_smul₁' : ∀ a S T L, toFun (a • S, T, L) = a * toFun (S, T, L)
|
|
|
|
|
map_add₁' : ∀ S1 S2 T L, toFun (S1 + S2, T, L) = toFun (S1, T, L) + toFun (S2, T, L)
|
|
|
|
|
swap₁' : ∀ S T L, toFun (S, T, L) = toFun (T, S, L)
|
|
|
|
|
swap₂' : ∀ S T L, toFun (S, T, L) = toFun (S, L, T)
|
|
|
|
|
|
|
|
|
|
namespace TriLinearSymm
|
2024-04-19 09:57:30 -04:00
|
|
|
|
open BigOperators
|
2024-04-16 15:34:00 -04:00
|
|
|
|
variable {V : Type} [AddCommMonoid V] [Module ℚ V]
|
|
|
|
|
|
|
|
|
|
instance instFun : FunLike (TriLinearSymm V) (V × V × V) ℚ where
|
|
|
|
|
coe f := f.toFun
|
|
|
|
|
coe_injective' f g h := by
|
|
|
|
|
cases f
|
|
|
|
|
cases g
|
|
|
|
|
simp_all
|
|
|
|
|
|
|
|
|
|
lemma toFun_eq_coe (f : TriLinearSymm V) : f.toFun = f := rfl
|
|
|
|
|
|
|
|
|
|
lemma swap₁ (f : TriLinearSymm V) (S T L : V) : f (S, T, L) = f (T, S, L) :=
|
|
|
|
|
f.swap₁' S T L
|
|
|
|
|
|
|
|
|
|
lemma swap₂ (f : TriLinearSymm V) (S T L : V) : f (S, T, L) = f (S, L, T) :=
|
|
|
|
|
f.swap₂' S T L
|
|
|
|
|
|
|
|
|
|
lemma swap₃ (f : TriLinearSymm V) (S T L : V) : f (S, T, L) = f (L, T, S) := by
|
|
|
|
|
rw [f.swap₁, f.swap₂, f.swap₁]
|
|
|
|
|
|
|
|
|
|
lemma map_smul₁ (f : TriLinearSymm V) (a : ℚ) (S T L : V) :
|
|
|
|
|
f (a • S, T, L) = a * f (S, T, L) :=
|
|
|
|
|
f.map_smul₁' a S T L
|
|
|
|
|
|
|
|
|
|
lemma map_smul₂ (f : TriLinearSymm V) (S : V) (a : ℚ) (T L : V) :
|
|
|
|
|
f (S, a • T, L) = a * f (S, T, L) := by
|
|
|
|
|
rw [f.swap₁, f.map_smul₁, f.swap₁]
|
|
|
|
|
|
|
|
|
|
lemma map_smul₃ (f : TriLinearSymm V) (S T : V) (a : ℚ) (L : V) :
|
|
|
|
|
f (S, T, a • L) = a * f (S, T, L) := by
|
|
|
|
|
rw [f.swap₃, f.map_smul₁, f.swap₃]
|
|
|
|
|
|
|
|
|
|
lemma map_add₁ (f : TriLinearSymm V) (S1 S2 T L : V) :
|
|
|
|
|
f (S1 + S2, T, L) = f (S1, T, L) + f (S2, T, L) :=
|
|
|
|
|
f.map_add₁' S1 S2 T L
|
|
|
|
|
|
|
|
|
|
lemma map_add₂ (f : TriLinearSymm V) (S T1 T2 L : V) :
|
|
|
|
|
f (S, T1 + T2, L) = f (S, T1, L) + f (S, T2, L) := by
|
|
|
|
|
rw [f.swap₁, f.map_add₁, f.swap₁ S T1, f.swap₁ S T2]
|
|
|
|
|
|
|
|
|
|
lemma map_add₃ (f : TriLinearSymm V) (S T L1 L2 : V) :
|
|
|
|
|
f (S, T, L1 + L2) = f (S, T, L1) + f (S, T, L2) := by
|
|
|
|
|
rw [f.swap₃, f.map_add₁, f.swap₃, f.swap₃ L2 T S]
|
|
|
|
|
|
2024-04-19 10:08:56 -04:00
|
|
|
|
/-- Fixing the second and third input vectors, the resulting linear map. -/
|
2024-04-19 09:57:30 -04:00
|
|
|
|
def toLinear₁ (f : TriLinearSymm V) (T L : V) : V →ₗ[ℚ] ℚ where
|
|
|
|
|
toFun S := f (S, T, L)
|
|
|
|
|
map_add' S1 S2 := by
|
|
|
|
|
simp only [f.map_add₁]
|
|
|
|
|
map_smul' a S := by
|
|
|
|
|
simp only [f.map_smul₁]
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma toLinear₁_apply (f : TriLinearSymm V) (S T L : V) : f (S, T, L) = f.toLinear₁ T L S := rfl
|
|
|
|
|
|
|
|
|
|
lemma map_sum₁ {n : ℕ} (f : TriLinearSymm V) (S : Fin n → V) (T : V) (L : V) :
|
|
|
|
|
f (∑ i, S i, T, L) = ∑ i, f (S i, T, L) := by
|
|
|
|
|
rw [f.toLinear₁_apply]
|
|
|
|
|
rw [map_sum]
|
|
|
|
|
rfl
|
|
|
|
|
|
|
|
|
|
lemma map_sum₂ {n : ℕ} (f : TriLinearSymm V) (S : Fin n → V) (T : V) (L : V) :
|
|
|
|
|
f ( T, ∑ i, S i, L) = ∑ i, f (T, S i, L) := by
|
|
|
|
|
rw [swap₁, map_sum₁]
|
|
|
|
|
apply Fintype.sum_congr
|
|
|
|
|
intro i
|
|
|
|
|
rw [swap₁]
|
|
|
|
|
|
|
|
|
|
lemma map_sum₃ {n : ℕ} (f : TriLinearSymm V) (S : Fin n → V) (T : V) (L : V) :
|
|
|
|
|
f ( T, L, ∑ i, S i) = ∑ i, f (T, L, S i) := by
|
|
|
|
|
rw [swap₃, map_sum₁]
|
|
|
|
|
apply Fintype.sum_congr
|
|
|
|
|
intro i
|
|
|
|
|
rw [swap₃]
|
|
|
|
|
|
|
|
|
|
lemma map_sum₁₂₃ {n1 n2 n3 : ℕ} (f : TriLinearSymm V) (S : Fin n1 → V)
|
|
|
|
|
(T : Fin n2 → V) (L : Fin n3 → V) :
|
|
|
|
|
f (∑ i, S i, ∑ i, T i, ∑ i, L i) = ∑ i, ∑ k, ∑ l, f (S i, T k, L l) := by
|
|
|
|
|
rw [map_sum₁]
|
|
|
|
|
apply Fintype.sum_congr
|
|
|
|
|
intro i
|
|
|
|
|
rw [map_sum₂]
|
|
|
|
|
apply Fintype.sum_congr
|
|
|
|
|
intro i
|
|
|
|
|
rw [map_sum₃]
|
|
|
|
|
|
|
|
|
|
|
2024-04-16 15:44:20 -04:00
|
|
|
|
/-- The homogenous cubic equation obtainable from a symmetric trilinear function. -/
|
2024-04-16 15:34:00 -04:00
|
|
|
|
@[simps!]
|
|
|
|
|
def toCubic {charges : Type} [AddCommMonoid charges] [Module ℚ charges]
|
|
|
|
|
(τ : TriLinearSymm charges) : HomogeneousCubic charges where
|
|
|
|
|
toFun S := τ (S, S, S)
|
|
|
|
|
map_smul' a S := by
|
|
|
|
|
simp only
|
|
|
|
|
rw [τ.map_smul₁, τ.map_smul₂, τ.map_smul₃]
|
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
lemma toCubic_add {charges : Type} [AddCommMonoid charges] [Module ℚ charges]
|
|
|
|
|
(τ : TriLinearSymm charges) (S T : charges) :
|
|
|
|
|
τ.toCubic (S + T) = τ.toCubic S +
|
|
|
|
|
τ.toCubic T + 3 * τ (S, S, T) + 3 * τ (T, T, S) := by
|
|
|
|
|
simp only [toCubic_toFun]
|
|
|
|
|
rw [τ.map_add₁, τ.map_add₂, τ.map_add₂, τ.map_add₃, τ.map_add₃, τ.map_add₃, τ.map_add₃]
|
|
|
|
|
rw [τ.swap₂ S T S, τ.swap₁ T S S, τ.swap₂ S T S, τ.swap₂ T S T, τ.swap₁ S T T, τ.swap₂ T S T]
|
|
|
|
|
ring
|
|
|
|
|
|
|
|
|
|
end TriLinearSymm
|