PhysLean/HepLean/Mathematics/LinearMaps.lean

300 lines
9.8 KiB
Text
Raw Normal View History

2024-04-16 15:34:00 -04:00
/-
Copyright (c) 2024 Joseph Tooby-Smith. All rights reserved.
2024-07-12 16:39:44 -04:00
Released under Apache 2.0 license as described in the file LICENSE.
2024-04-16 15:34:00 -04:00
Authors: Joseph Tooby-Smith
-/
import Mathlib.Tactic.Polyrith
2024-07-26 16:32:54 -04:00
import Mathlib.Algebra.Module.LinearMap.Defs
2024-06-25 07:06:32 -04:00
import Mathlib.Data.Fintype.BigOperators
2024-04-16 15:34:00 -04:00
/-!
# Linear maps
2024-06-08 01:50:33 +02:00
Some definitions and properties of linear, bilinear, and trilinear maps, along with homogeneous
2024-04-16 15:34:00 -04:00
quadratic and cubic equations.
-/
2024-07-09 19:22:16 -04:00
/-! TODO: Replace the definitions in this file with Mathlib definitions. -/
2024-04-16 15:34:00 -04:00
/-- The structure defining a homogeneous quadratic equation. -/
2024-04-22 08:41:50 -04:00
@[simp]
def HomogeneousQuadratic (V : Type) [AddCommMonoid V] [Module V] : Type :=
V →ₑ[((fun a => a ^ 2) : )]
2024-04-16 15:34:00 -04:00
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
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 symmetric bilinear function. -/
2024-04-22 08:41:50 -04:00
structure BiLinearSymm (V : Type) [AddCommMonoid V] [Module V] extends V →ₗ[] V →ₗ[] where
swap' : ∀ S T, toFun S T = toFun T S
2024-04-16 15:34:00 -04:00
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
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
2024-04-16 15:34:00 -04:00
coe f := f.toFun
coe_injective' f g h := by
cases f
cases g
simp_all
2024-06-08 01:50:33 +02:00
/-- The construction of a symmetric bilinear map from `smul` and `map_add` in the first factor,
2024-04-22 08:46:29 -04:00
and swap. -/
2024-04-22 08:41:50 -04:00
@[simps!]
def mk₂ (f : V × V → ) (map_smul : ∀ a S T, f (a • S, T) = a * f (S, T))
(map_add : ∀ S1 S2 T, f (S1 + S2, T) = f (S1, T) + f (S2, T))
(swap : ∀ S T, f (S, T) = f (T, S)) : BiLinearSymm V where
toFun := fun S => {
toFun := fun T => f (S, T)
map_add' := by
intro T1 T2
2024-04-22 08:46:29 -04:00
simp only
2024-04-22 08:41:50 -04:00
rw [swap, map_add]
2024-06-13 16:49:01 -04:00
exact Mathlib.Tactic.LinearCombination.add_pf (swap T1 S) (swap T2 S)
2024-07-12 16:22:06 -04:00
map_smul' := by
2024-04-22 08:41:50 -04:00
intro a T
2024-04-22 08:46:29 -04:00
simp only [eq_ratCast, Rat.cast_eq_id, id_eq, smul_eq_mul]
2024-04-22 08:41:50 -04:00
rw [swap, map_smul]
2024-06-13 16:49:01 -04:00
exact congrArg (HMul.hMul a) (swap T S)
2024-04-22 08:41:50 -04:00
}
2024-07-12 11:23:02 -04:00
map_smul' := fun a S => LinearMap.ext fun T => map_smul a S T
map_add' := fun S1 S2 => LinearMap.ext fun T => map_add S1 S2 T
2024-04-22 08:41:50 -04:00
swap' := swap
lemma map_smul₁ (f : BiLinearSymm V) (a : ) (S T : V) : f (a • S) T = a * f S T := by
erw [f.map_smul a S]
rfl
2024-04-16 15:34:00 -04:00
2024-08-19 17:23:55 +02:00
lemma swap (f : BiLinearSymm V) (S T : V) : f S T = f T S := f.swap' S T
2024-04-16 15:34:00 -04:00
2024-07-12 11:23:02 -04:00
lemma map_smul₂ (f : BiLinearSymm V) (a : ) (S : V) (T : V) : f S (a • T) = a * f S T := by
2024-04-16 15:34:00 -04:00
rw [f.swap, f.map_smul₁, f.swap]
2024-04-22 08:41:50 -04:00
lemma map_add₁ (f : BiLinearSymm V) (S1 S2 T : V) : f (S1 + S2) T = f S1 T + f S2 T := by
erw [f.map_add]
rfl
2024-04-16 15:34:00 -04:00
lemma map_add₂ (f : BiLinearSymm V) (S : V) (T1 T2 : V) :
2024-04-22 08:41:50 -04:00
f S (T1 + T2) = f S T1 + f S T2 := by
2024-04-16 15:34:00 -04:00
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. -/
def toLinear₁ (f : BiLinearSymm V) (T : V) : V →ₗ[] where
2024-04-22 08:41:50 -04:00
toFun S := f S T
2024-06-13 16:49:01 -04:00
map_add' S1 S2 := map_add₁ f S1 S2 T
map_smul' a S := by
simp only [f.map_smul₁]
rfl
2024-04-22 08:41:50 -04:00
lemma toLinear₁_apply (f : BiLinearSymm V) (S T : V) : f S T = f.toLinear₁ T S := rfl
2024-07-12 11:23:02 -04:00
lemma map_sum₁ {n : } (f : BiLinearSymm V) (S : Fin n → V) (T : V) :
2024-04-22 08:41:50 -04:00
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) :
2024-08-19 17:23:55 +02:00
f T (∑ i, S i) = ∑ i, f T (S i) := map_sum (f T) S Finset.univ
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
2024-04-22 08:41:50 -04:00
toFun S := τ S S
2024-04-16 15:34:00 -04:00
map_smul' a S := by
simp only
rw [τ.map_smul₁, τ.map_smul₂]
2024-04-22 08:41:50 -04:00
ring_nf
rfl
2024-04-16 15:34:00 -04:00
lemma toHomogeneousQuad_add {V : Type} [AddCommMonoid V] [Module V]
(τ : BiLinearSymm V) (S T : V) :
τ.toHomogeneousQuad (S + T) = τ.toHomogeneousQuad S +
2024-04-22 08:41:50 -04:00
τ.toHomogeneousQuad T + 2 * τ S T := by
simp only [HomogeneousQuadratic, toHomogeneousQuad_apply, map_add]
2024-04-22 08:41:50 -04:00
rw [τ.map_add₁, τ.map_add₁, τ.swap T S]
2024-04-16 15:34:00 -04:00
ring
end BiLinearSymm
2024-04-16 15:44:20 -04:00
/-- The structure of a homogeneous cubic equation. -/
2024-04-22 08:41:50 -04:00
@[simp]
def HomogeneousCubic (V : Type) [AddCommMonoid V] [Module V] : Type :=
V →ₑ[((fun a => a ^ 3) : )]
2024-04-16 15:34:00 -04:00
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
/-- 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
namespace TriLinearSymm
open BigOperators
2024-04-16 15:34:00 -04:00
variable {V : Type} [AddCommMonoid V] [Module V]
2024-07-12 16:22:06 -04:00
instance instFun : FunLike (TriLinearSymm V) V (V →ₗ[] V →ₗ[] ) where
2024-04-16 15:34:00 -04:00
coe f := f.toFun
coe_injective' f g h := by
cases f
cases g
simp_all
2024-06-08 01:50:33 +02:00
/-- The construction of a symmetric trilinear map from `smul` and `map_add` in the first factor,
and two swap. -/
@[simps!]
def mk₃ (f : V × V × V→ ) (map_smul : ∀ a S T L, f (a • S, T, L) = a * f (S, T, L))
(map_add : ∀ S1 S2 T L, f (S1 + S2, T, L) = f (S1, T, L) + f (S2, T, L))
(swap₁ : ∀ S T L, f (S, T, L) = f (T, S, L))
(swap₂ : ∀ S T L, f (S, T, L) = f (S, L, T)) : TriLinearSymm V where
toFun := fun S => (BiLinearSymm.mk₂ (fun T => f (S, T))
(by
intro a T L
simp only
rw [swap₁, map_smul, swap₁])
(by
intro S1 S2 T
simp only
rw [swap₁, map_add, swap₁, swap₁ S2 S T])
(by
intro L T
2024-08-30 10:11:55 -04:00
exact swap₂ S L T)).toLinearMap
map_add' S1 S2 := by
apply LinearMap.ext
intro T
apply LinearMap.ext
intro S
2024-08-30 10:11:55 -04:00
exact map_add S1 S2 T S
map_smul' a S :=
LinearMap.ext fun T => LinearMap.ext fun L => map_smul a S T L
swap₁' := swap₁
swap₂' := swap₂
lemma swap₁ (f : TriLinearSymm V) (S T L : V) : f S T L = f T S L :=
2024-04-16 15:34:00 -04:00
f.swap₁' S T L
lemma swap₂ (f : TriLinearSymm V) (S T L : V) : f S T L = f S L T :=
2024-04-16 15:34:00 -04:00
f.swap₂' S T L
lemma swap₃ (f : TriLinearSymm V) (S T L : V) : f S T L = f L T S := by
2024-04-16 15:34:00 -04:00
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 := by
erw [f.map_smul a S]
rfl
2024-04-16 15:34:00 -04:00
lemma map_smul₂ (f : TriLinearSymm V) (S : V) (a : ) (T L : V) :
f S (a • T) L = a * f S T L := by
2024-04-16 15:34:00 -04:00
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
2024-04-16 15:34:00 -04:00
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 := by
erw [f.map_add]
rfl
2024-04-16 15:34:00 -04:00
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
2024-04-16 15:34:00 -04:00
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
2024-04-16 15:34:00 -04:00
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-07-12 11:23:02 -04:00
def toLinear₁ (f : TriLinearSymm V) (T L : V) : V →ₗ[] where
toFun S := f S T L
2024-08-30 10:11:55 -04:00
map_add' S1 S2 := map_add₁ f S1 S2 T L
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) :
2024-08-19 17:23:55 +02:00
f T L (∑ i, S i) = ∑ i, f T L (S i) := map_sum ((f T) L) S Finset.univ
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
2024-04-16 15:34:00 -04:00
map_smul' a S := by
2024-04-22 08:46:29 -04:00
simp only [smul_eq_mul]
2024-04-16 15:34:00 -04:00
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
2024-04-22 08:41:50 -04:00
simp only [HomogeneousCubic, toCubic_apply]
2024-04-16 15:34:00 -04:00
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