|
| 1 | +/- |
| 2 | +Copyright (c) 2023 Eric Wieser. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Eric Wieser |
| 5 | +-/ |
| 6 | +import linear_algebra.bilinear_form |
| 7 | +import linear_algebra.tensor_product |
| 8 | +import linear_algebra.contraction |
| 9 | + |
| 10 | +/-! |
| 11 | +# The bilinear form on a tensor product |
| 12 | +
|
| 13 | +## Main definitions |
| 14 | +
|
| 15 | +* `bilin_form.tensor_distrib (B₁ ⊗ₜ B₂)`: the bilinear form on `M₁ ⊗ M₂` constructed by applying |
| 16 | + `B₁` on `M₁` and `B₂` on `M₂`. |
| 17 | +* `bilin_form.tensor_distrib_equiv`: `bilin_form.tensor_distrib` as an equivalence on finite free |
| 18 | + modules. |
| 19 | +
|
| 20 | +-/ |
| 21 | + |
| 22 | +universes u v w |
| 23 | +variables {ι : Type*} {R : Type*} {M₁ M₂ : Type*} |
| 24 | + |
| 25 | +open_locale tensor_product |
| 26 | + |
| 27 | +namespace bilin_form |
| 28 | + |
| 29 | +section comm_semiring |
| 30 | +variables [comm_semiring R] |
| 31 | +variables [add_comm_monoid M₁] [add_comm_monoid M₂] |
| 32 | +variables [module R M₁] [module R M₂] |
| 33 | + |
| 34 | +/-- The tensor product of two bilinear forms injects into bilinear forms on tensor products. -/ |
| 35 | +def tensor_distrib : bilin_form R M₁ ⊗[R] bilin_form R M₂ →ₗ[R] bilin_form R (M₁ ⊗[R] M₂) := |
| 36 | +((tensor_product.tensor_tensor_tensor_comm R _ _ _ _).dual_map |
| 37 | + ≪≫ₗ (tensor_product.lift.equiv R _ _ _).symm |
| 38 | + ≪≫ₗ linear_map.to_bilin).to_linear_map |
| 39 | + ∘ₗ tensor_product.dual_distrib R _ _ |
| 40 | + ∘ₗ (tensor_product.congr |
| 41 | + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _) |
| 42 | + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _)).to_linear_map |
| 43 | + |
| 44 | +@[simp] lemma tensor_distrib_tmul (B₁ : bilin_form R M₁) (B₂ : bilin_form R M₂) |
| 45 | + (m₁ : M₁) (m₂ : M₂) (m₁' : M₁) (m₂' : M₂) : |
| 46 | + tensor_distrib (B₁ ⊗ₜ B₂) (m₁ ⊗ₜ m₂) (m₁' ⊗ₜ m₂') = B₁ m₁ m₁' * B₂ m₂ m₂' := |
| 47 | +rfl |
| 48 | + |
| 49 | +/-- The tensor product of two bilinear forms, a shorthand for dot notation. -/ |
| 50 | +@[reducible] |
| 51 | +protected def tmul (B₁ : bilin_form R M₁) (B₂ : bilin_form R M₂) : bilin_form R (M₁ ⊗[R] M₂) := |
| 52 | +tensor_distrib (B₁ ⊗ₜ[R] B₂) |
| 53 | + |
| 54 | +end comm_semiring |
| 55 | + |
| 56 | +section comm_ring |
| 57 | +variables [comm_ring R] |
| 58 | +variables [add_comm_group M₁] [add_comm_group M₂] |
| 59 | +variables [module R M₁] [module R M₂] |
| 60 | +variables [module.free R M₁] [module.finite R M₁] |
| 61 | +variables [module.free R M₂] [module.finite R M₂] |
| 62 | +variables [nontrivial R] |
| 63 | + |
| 64 | +/-- `tensor_distrib` as an equivalence. -/ |
| 65 | +noncomputable def tensor_distrib_equiv : |
| 66 | + bilin_form R M₁ ⊗[R] bilin_form R M₂ ≃ₗ[R] bilin_form R (M₁ ⊗[R] M₂) := |
| 67 | +-- the same `linear_equiv`s as from `tensor_distrib`, but with the inner linear map also as an |
| 68 | +-- equiv |
| 69 | +tensor_product.congr |
| 70 | + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _) |
| 71 | + (bilin_form.to_lin ≪≫ₗ tensor_product.lift.equiv R _ _ _) |
| 72 | + ≪≫ₗ tensor_product.dual_distrib_equiv R (M₁ ⊗ M₁) (M₂ ⊗ M₂) |
| 73 | + ≪≫ₗ (tensor_product.tensor_tensor_tensor_comm R _ _ _ _).dual_map |
| 74 | + ≪≫ₗ (tensor_product.lift.equiv R _ _ _).symm |
| 75 | + ≪≫ₗ linear_map.to_bilin |
| 76 | + |
| 77 | +@[simp] |
| 78 | +lemma tensor_distrib_equiv_apply (B : bilin_form R M₁ ⊗ bilin_form R M₂) : |
| 79 | + tensor_distrib_equiv B = tensor_distrib B := rfl |
| 80 | + |
| 81 | +end comm_ring |
| 82 | + |
| 83 | +end bilin_form |
0 commit comments