refactor: Linting
This commit is contained in:
parent
a65fb06605
commit
a438af453d
14 changed files with 87 additions and 76 deletions
|
@ -34,6 +34,7 @@ open TensorProduct
|
|||
|
||||
variable {R : Type} [CommSemiring R]
|
||||
|
||||
/-- An auxillary function to contract the vector space `V1` and `V2` in `V1 ⊗[R] V2 ⊗[R] V3`. -/
|
||||
def contrDualLeftAux {V1 V2 V3 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
|
||||
[Module R V1] [Module R V2] [Module R V3] (f : V1 ⊗[R] V2 →ₗ[R] R) :
|
||||
V1 ⊗[R] V2 ⊗[R] V3 →ₗ[R] V3 :=
|
||||
|
@ -41,9 +42,11 @@ def contrDualLeftAux {V1 V2 V3 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [Ad
|
|||
TensorProduct.map (f) (LinearEquiv.refl R V3).toLinearMap
|
||||
∘ₗ (TensorProduct.assoc R _ _ _).symm.toLinearMap
|
||||
|
||||
/-- An auxillary function to contract the vector space `V1` and `V2` in
|
||||
`V4 ⊗[R] V1 ⊗[R] V2 ⊗[R] V3`. -/
|
||||
def contrDualMidAux {V1 V2 V3 V4 : Type} [AddCommMonoid V1] [AddCommMonoid V2] [AddCommMonoid V3]
|
||||
[AddCommMonoid V4] [Module R V1] [Module R V2] [Module R V3] [Module R V4] (f : V1 ⊗[R] V2 →ₗ[R] R) :
|
||||
(V4 ⊗[R] V1) ⊗[R] (V2 ⊗[R] V3) →ₗ[R] V4 ⊗[R] V3 :=
|
||||
[AddCommMonoid V4] [Module R V1] [Module R V2] [Module R V3] [Module R V4]
|
||||
(f : V1 ⊗[R] V2 →ₗ[R] R) : (V4 ⊗[R] V1) ⊗[R] (V2 ⊗[R] V3) →ₗ[R] V4 ⊗[R] V3 :=
|
||||
(TensorProduct.map (LinearEquiv.refl R V4).toLinearMap (contrDualLeftAux f)) ∘ₗ
|
||||
(TensorProduct.assoc R _ _ _).toLinearMap
|
||||
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue