|
| 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 Mathlib.LinearAlgebra.BilinearForm.TensorProduct |
| 7 | +import Mathlib.LinearAlgebra.QuadraticForm.Basic |
| 8 | + |
| 9 | +/-! |
| 10 | +# The quadratic form on a tensor product |
| 11 | +
|
| 12 | +## Main definitions |
| 13 | +
|
| 14 | +* `QuadraticForm.tensorDistrib (Q₁ ⊗ₜ Q₂)`: the quadratic form on `M₁ ⊗ M₂` constructed by applying |
| 15 | + `Q₁` on `M₁` and `Q₂` on `M₂`. This construction is not available in characteristic two. |
| 16 | +
|
| 17 | +-/ |
| 18 | + |
| 19 | +universe uR uA uM₁ uM₂ |
| 20 | + |
| 21 | +variable {R : Type uR} {A : Type uA} {M₁ : Type uM₁} {M₂ : Type uM₂} |
| 22 | + |
| 23 | +open TensorProduct |
| 24 | + |
| 25 | +namespace QuadraticForm |
| 26 | + |
| 27 | +section CommRing |
| 28 | +variable [CommRing R] [CommRing A] |
| 29 | +variable [AddCommGroup M₁] [AddCommGroup M₂] |
| 30 | +variable [Algebra R A] [Module R M₁] [Module A M₁] |
| 31 | +variable [SMulCommClass R A M₁] [SMulCommClass A R M₁] [IsScalarTower R A M₁] |
| 32 | +variable [Module R M₂] [Invertible (2 : R)] |
| 33 | + |
| 34 | + |
| 35 | +variable (R A) in |
| 36 | +/-- The tensor product of two quadratic forms injects into quadratic forms on tensor products. |
| 37 | +
|
| 38 | +Note this is heterobasic; the quadratic form on the left can take values in a larger ring than |
| 39 | +the one on the right. -/ |
| 40 | +def tensorDistrib : QuadraticForm A M₁ ⊗[R] QuadraticForm R M₂ →ₗ[A] QuadraticForm A (M₁ ⊗[R] M₂) := |
| 41 | + letI : Invertible (2 : A) := (Invertible.map (algebraMap R A) 2).copy 2 (map_ofNat _ _).symm |
| 42 | + -- while `letI`s would produce a better term than `let`, they would make this already-slow |
| 43 | + -- definition even slower. |
| 44 | + let toQ := BilinForm.toQuadraticFormLinearMap A A (M₁ ⊗[R] M₂) |
| 45 | + let tmulB := BilinForm.tensorDistrib R A (M₁ := M₁) (M₂ := M₂) |
| 46 | + let toB := AlgebraTensorModule.map |
| 47 | + (QuadraticForm.associated : QuadraticForm A M₁ →ₗ[A] BilinForm A M₁) |
| 48 | + (QuadraticForm.associated : QuadraticForm R M₂ →ₗ[R] BilinForm R M₂) |
| 49 | + toQ ∘ₗ tmulB ∘ₗ toB |
| 50 | + |
| 51 | +-- TODO: make the RHS `MulOpposite.op (Q₂ m₂) • Q₁ m₁` so that this has a nicer defeq for |
| 52 | +-- `R = A` of `Q₁ m₁ * Q₂ m₂`. |
| 53 | +@[simp] |
| 54 | +theorem tensorDistrib_tmul (Q₁ : QuadraticForm A M₁) (Q₂ : QuadraticForm R M₂) (m₁ : M₁) (m₂ : M₂) : |
| 55 | + tensorDistrib R A (Q₁ ⊗ₜ Q₂) (m₁ ⊗ₜ m₂) = Q₂ m₂ • Q₁ m₁ := |
| 56 | + letI : Invertible (2 : A) := (Invertible.map (algebraMap R A) 2).copy 2 (map_ofNat _ _).symm |
| 57 | + (BilinForm.tensorDistrib_tmul _ _ _ _ _ _).trans <| congr_arg₂ _ |
| 58 | + (associated_eq_self_apply _ _ _) (associated_eq_self_apply _ _ _) |
| 59 | + |
| 60 | +/-- The tensor product of two quadratic forms, a shorthand for dot notation. -/ |
| 61 | +protected abbrev tmul (Q₁ : QuadraticForm A M₁) (Q₂ : QuadraticForm R M₂) : |
| 62 | + QuadraticForm A (M₁ ⊗[R] M₂) := |
| 63 | + tensorDistrib R A (Q₁ ⊗ₜ[R] Q₂) |
| 64 | + |
| 65 | +theorem associated_tmul [Invertible (2 : A)] (Q₁ : QuadraticForm A M₁) (Q₂ : QuadraticForm R M₂) : |
| 66 | + associated (R₁ := A) (Q₁.tmul Q₂) |
| 67 | + = (associated (R₁ := A) Q₁).tmul (associated (R₁ := R) Q₂) := by |
| 68 | + rw [QuadraticForm.tmul, tensorDistrib, BilinForm.tmul] |
| 69 | + dsimp |
| 70 | + convert associated_left_inverse A ((associated_isSymm A Q₁).tmul (associated_isSymm R Q₂)) |
| 71 | + |
| 72 | +variable (A) in |
| 73 | +/-- The base change of a quadratic form. -/ |
| 74 | +protected def baseChange (Q : QuadraticForm R M₂) : QuadraticForm A (A ⊗[R] M₂) := |
| 75 | + QuadraticForm.tmul (R := R) (A := A) (M₁ := A) (M₂ := M₂) (QuadraticForm.sq (R := A)) Q |
| 76 | + |
| 77 | +@[simp] |
| 78 | +theorem baseChange_tmul (Q : QuadraticForm R M₂) (a : A) (m₂ : M₂) : |
| 79 | + Q.baseChange A (a ⊗ₜ m₂) = Q m₂ • (a * a) := |
| 80 | + tensorDistrib_tmul _ _ _ _ |
| 81 | + |
| 82 | + |
| 83 | +theorem associated_baseChange [Invertible (2 : A)] (Q : QuadraticForm R M₂) : |
| 84 | + associated (R₁ := A) (Q.baseChange A) = (associated (R₁ := R) Q).baseChange A := by |
| 85 | + dsimp only [QuadraticForm.baseChange, BilinForm.baseChange] |
| 86 | + rw [associated_tmul (QuadraticForm.sq (R := A)) Q, associated_sq] |
| 87 | + |
| 88 | +end CommRing |
| 89 | + |
| 90 | +end QuadraticForm |
0 commit comments