|
| 1 | +/- |
| 2 | +Copyright (c) 2019 Johannes Hölzl. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen, Antoine Labelle |
| 5 | +
|
| 6 | +! This file was ported from Lean 3 source module linear_algebra.trace |
| 7 | +! leanprover-community/mathlib commit 4cf7ca0e69e048b006674cf4499e5c7d296a89e0 |
| 8 | +! Please do not edit these lines, except to modify the commit id |
| 9 | +! if you have ported upstream changes. |
| 10 | +-/ |
| 11 | +import Mathlib.LinearAlgebra.Matrix.ToLin |
| 12 | +import Mathlib.LinearAlgebra.Matrix.Trace |
| 13 | +import Mathlib.LinearAlgebra.Contraction |
| 14 | +import Mathlib.LinearAlgebra.TensorProductBasis |
| 15 | +import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition |
| 16 | +import Mathlib.LinearAlgebra.FreeModule.Finite.Rank |
| 17 | +import Mathlib.LinearAlgebra.Projection |
| 18 | + |
| 19 | +/-! |
| 20 | +# Trace of a linear map |
| 21 | +
|
| 22 | +This file defines the trace of a linear map. |
| 23 | +
|
| 24 | +See also `LinearAlgebra/Matrix/Trace.lean` for the trace of a matrix. |
| 25 | +
|
| 26 | +## Tags |
| 27 | +
|
| 28 | +linear_map, trace, diagonal |
| 29 | +
|
| 30 | +-/ |
| 31 | + |
| 32 | + |
| 33 | +noncomputable section |
| 34 | + |
| 35 | +universe u v w |
| 36 | + |
| 37 | +namespace LinearMap |
| 38 | + |
| 39 | +open BigOperators |
| 40 | + |
| 41 | +open Matrix |
| 42 | + |
| 43 | +open FiniteDimensional |
| 44 | + |
| 45 | +open TensorProduct |
| 46 | + |
| 47 | +section |
| 48 | + |
| 49 | +variable (R : Type u) [CommSemiring R] {M : Type v} [AddCommMonoid M] [Module R M] |
| 50 | + |
| 51 | +variable {ι : Type w} [DecidableEq ι] [Fintype ι] |
| 52 | + |
| 53 | +variable {κ : Type _} [DecidableEq κ] [Fintype κ] |
| 54 | + |
| 55 | +variable (b : Basis ι R M) (c : Basis κ R M) |
| 56 | + |
| 57 | +/-- The trace of an endomorphism given a basis. -/ |
| 58 | +def traceAux : (M →ₗ[R] M) →ₗ[R] R := |
| 59 | + Matrix.traceLinearMap ι R R ∘ₗ ↑(LinearMap.toMatrix b b) |
| 60 | +#align linear_map.trace_aux LinearMap.traceAux |
| 61 | + |
| 62 | +-- Can't be `simp` because it would cause a loop. |
| 63 | +theorem traceAux_def (b : Basis ι R M) (f : M →ₗ[R] M) : |
| 64 | + traceAux R b f = Matrix.trace (LinearMap.toMatrix b b f) := |
| 65 | + rfl |
| 66 | +#align linear_map.trace_aux_def LinearMap.traceAux_def |
| 67 | + |
| 68 | +theorem traceAux_eq : traceAux R b = traceAux R c := |
| 69 | + LinearMap.ext fun f => |
| 70 | + calc |
| 71 | + Matrix.trace (LinearMap.toMatrix b b f) = |
| 72 | + Matrix.trace (LinearMap.toMatrix b b ((LinearMap.id.comp f).comp LinearMap.id)) := by |
| 73 | + rw [LinearMap.id_comp, LinearMap.comp_id] |
| 74 | + _ = Matrix.trace (LinearMap.toMatrix c b LinearMap.id ⬝ LinearMap.toMatrix c c f ⬝ |
| 75 | + LinearMap.toMatrix b c LinearMap.id) := by |
| 76 | + rw [LinearMap.toMatrix_comp _ c, LinearMap.toMatrix_comp _ c] |
| 77 | + _ = Matrix.trace (LinearMap.toMatrix c c f ⬝ LinearMap.toMatrix b c LinearMap.id ⬝ |
| 78 | + LinearMap.toMatrix c b LinearMap.id) := by |
| 79 | + rw [Matrix.mul_assoc, Matrix.trace_mul_comm] |
| 80 | + _ = Matrix.trace (LinearMap.toMatrix c c ((f.comp LinearMap.id).comp LinearMap.id)) := by |
| 81 | + rw [LinearMap.toMatrix_comp _ b, LinearMap.toMatrix_comp _ c] |
| 82 | + _ = Matrix.trace (LinearMap.toMatrix c c f) := by rw [LinearMap.comp_id, LinearMap.comp_id] |
| 83 | +#align linear_map.trace_aux_eq LinearMap.traceAux_eq |
| 84 | + |
| 85 | +open Classical |
| 86 | + |
| 87 | +variable (M) |
| 88 | + |
| 89 | +/-- Trace of an endomorphism independent of basis. -/ |
| 90 | +def trace : (M →ₗ[R] M) →ₗ[R] R := |
| 91 | + if H : ∃ s : Finset M, Nonempty (Basis s R M) then traceAux R H.choose_spec.some else 0 |
| 92 | +#align linear_map.trace LinearMap.trace |
| 93 | + |
| 94 | +variable {M} |
| 95 | + |
| 96 | +/-- Auxiliary lemma for `trace_eq_matrix_trace`. -/ |
| 97 | +theorem trace_eq_matrix_trace_of_finset {s : Finset M} (b : Basis s R M) (f : M →ₗ[R] M) : |
| 98 | + trace R M f = Matrix.trace (LinearMap.toMatrix b b f) := by |
| 99 | + have : ∃ s : Finset M, Nonempty (Basis s R M) := ⟨s, ⟨b⟩⟩ |
| 100 | + rw [trace, dif_pos this, ← traceAux_def] |
| 101 | + congr 1 |
| 102 | + apply traceAux_eq |
| 103 | +#align linear_map.trace_eq_matrix_trace_of_finset LinearMap.trace_eq_matrix_trace_of_finset |
| 104 | + |
| 105 | +theorem trace_eq_matrix_trace (f : M →ₗ[R] M) : |
| 106 | + trace R M f = Matrix.trace (LinearMap.toMatrix b b f) := by |
| 107 | + rw [trace_eq_matrix_trace_of_finset R b.reindexFinsetRange, ← traceAux_def, ← traceAux_def, |
| 108 | + traceAux_eq R b b.reindexFinsetRange] |
| 109 | +#align linear_map.trace_eq_matrix_trace LinearMap.trace_eq_matrix_trace |
| 110 | + |
| 111 | +theorem trace_mul_comm (f g : M →ₗ[R] M) : trace R M (f * g) = trace R M (g * f) := |
| 112 | + if H : ∃ s : Finset M, Nonempty (Basis s R M) then by |
| 113 | + let ⟨s, ⟨b⟩⟩ := H |
| 114 | + simp_rw [trace_eq_matrix_trace R b, LinearMap.toMatrix_mul] |
| 115 | + apply Matrix.trace_mul_comm |
| 116 | + else by rw [trace, dif_neg H, LinearMap.zero_apply, LinearMap.zero_apply] |
| 117 | +#align linear_map.trace_mul_comm LinearMap.trace_mul_comm |
| 118 | + |
| 119 | +/-- The trace of an endomorphism is invariant under conjugation -/ |
| 120 | +@[simp] |
| 121 | +theorem trace_conj (g : M →ₗ[R] M) (f : (M →ₗ[R] M)ˣ) : |
| 122 | + trace R M (↑f * g * ↑f⁻¹) = trace R M g := by |
| 123 | + rw [trace_mul_comm] |
| 124 | + simp |
| 125 | +#align linear_map.trace_conj LinearMap.trace_conj |
| 126 | + |
| 127 | +end |
| 128 | + |
| 129 | +section |
| 130 | + |
| 131 | +variable {R : Type _} [CommRing R] {M : Type _} [AddCommGroup M] [Module R M] |
| 132 | + |
| 133 | +variable (N : Type _) [AddCommGroup N] [Module R N] |
| 134 | + |
| 135 | +variable {ι : Type _} |
| 136 | + |
| 137 | +/-- The trace of a linear map correspond to the contraction pairing under the isomorphism |
| 138 | + `End(M) ≃ M* ⊗ M`-/ |
| 139 | +theorem trace_eq_contract_of_basis [Finite ι] (b : Basis ι R M) : |
| 140 | + LinearMap.trace R M ∘ₗ dualTensorHom R M M = contractLeft R M := by |
| 141 | + classical |
| 142 | + cases nonempty_fintype ι |
| 143 | + apply Basis.ext (Basis.tensorProduct (Basis.dualBasis b) b) |
| 144 | + rintro ⟨i, j⟩ |
| 145 | + simp only [Function.comp_apply, Basis.tensorProduct_apply, Basis.coe_dualBasis, coe_comp] |
| 146 | + rw [trace_eq_matrix_trace R b, toMatrix_dualTensorHom] |
| 147 | + by_cases hij : i = j |
| 148 | + · rw [hij] |
| 149 | + simp |
| 150 | + rw [Matrix.StdBasisMatrix.trace_zero j i (1 : R) hij] |
| 151 | + simp [Finsupp.single_eq_pi_single, hij] |
| 152 | +#align linear_map.trace_eq_contract_of_basis LinearMap.trace_eq_contract_of_basis |
| 153 | + |
| 154 | +/-- The trace of a linear map correspond to the contraction pairing under the isomorphism |
| 155 | + `End(M) ≃ M* ⊗ M`-/ |
| 156 | +theorem trace_eq_contract_of_basis' [Fintype ι] [DecidableEq ι] (b : Basis ι R M) : |
| 157 | + LinearMap.trace R M = contractLeft R M ∘ₗ (dualTensorHomEquivOfBasis b).symm.toLinearMap := by |
| 158 | + simp [LinearEquiv.eq_comp_toLinearMap_symm, trace_eq_contract_of_basis b] |
| 159 | +#align linear_map.trace_eq_contract_of_basis' LinearMap.trace_eq_contract_of_basis' |
| 160 | + |
| 161 | +variable (R M) |
| 162 | + |
| 163 | +variable [Module.Free R M] [Module.Finite R M] [Module.Free R N] [Module.Finite R N] [Nontrivial R] |
| 164 | + |
| 165 | +/-- When `M` is finite free, the trace of a linear map correspond to the contraction pairing under |
| 166 | +the isomorphism `End(M) ≃ M* ⊗ M`-/ |
| 167 | +@[simp] |
| 168 | +theorem trace_eq_contract : LinearMap.trace R M ∘ₗ dualTensorHom R M M = contractLeft R M := |
| 169 | + trace_eq_contract_of_basis (Module.Free.chooseBasis R M) |
| 170 | +#align linear_map.trace_eq_contract LinearMap.trace_eq_contract |
| 171 | + |
| 172 | +@[simp] |
| 173 | +theorem trace_eq_contract_apply (x : Module.Dual R M ⊗[R] M) : |
| 174 | + (LinearMap.trace R M) ((dualTensorHom R M M) x) = contractLeft R M x := by |
| 175 | + rw [← comp_apply, trace_eq_contract] |
| 176 | +#align linear_map.trace_eq_contract_apply LinearMap.trace_eq_contract_apply |
| 177 | + |
| 178 | +open Classical |
| 179 | + |
| 180 | +/-- When `M` is finite free, the trace of a linear map correspond to the contraction pairing under |
| 181 | +the isomorphism `End(M) ≃ M* ⊗ M`-/ |
| 182 | +theorem trace_eq_contract' : |
| 183 | + LinearMap.trace R M = contractLeft R M ∘ₗ (dualTensorHomEquiv R M M).symm.toLinearMap := |
| 184 | + trace_eq_contract_of_basis' (Module.Free.chooseBasis R M) |
| 185 | +#align linear_map.trace_eq_contract' LinearMap.trace_eq_contract' |
| 186 | + |
| 187 | +/-- The trace of the identity endomorphism is the dimension of the free module -/ |
| 188 | +@[simp] |
| 189 | +theorem trace_one : trace R M 1 = (finrank R M : R) := by |
| 190 | + have b := Module.Free.chooseBasis R M |
| 191 | + rw [trace_eq_matrix_trace R b, toMatrix_one, finrank_eq_card_chooseBasisIndex] |
| 192 | + simp |
| 193 | +#align linear_map.trace_one LinearMap.trace_one |
| 194 | + |
| 195 | +/-- The trace of the identity endomorphism is the dimension of the free module -/ |
| 196 | +@[simp] |
| 197 | +theorem trace_id : trace R M id = (finrank R M : R) := by rw [← one_eq_id, trace_one] |
| 198 | +#align linear_map.trace_id LinearMap.trace_id |
| 199 | + |
| 200 | +@[simp] |
| 201 | +theorem trace_transpose : trace R (Module.Dual R M) ∘ₗ Module.Dual.transpose = trace R M := by |
| 202 | + let e := dualTensorHomEquiv R M M |
| 203 | + have h : Function.Surjective e.toLinearMap := e.surjective |
| 204 | + refine' (cancel_right h).1 _ |
| 205 | + ext (f m); simp |
| 206 | +#align linear_map.trace_transpose LinearMap.trace_transpose |
| 207 | + |
| 208 | +theorem trace_prodMap : |
| 209 | + trace R (M × N) ∘ₗ prodMapLinear R M N M N R = |
| 210 | + (coprod id id : R × R →ₗ[R] R) ∘ₗ prodMap (trace R M) (trace R N) := by |
| 211 | + let e := (dualTensorHomEquiv R M M).prod (dualTensorHomEquiv R N N) |
| 212 | + have h : Function.Surjective e.toLinearMap := e.surjective |
| 213 | + refine' (cancel_right h).1 _ |
| 214 | + ext |
| 215 | + · simp only [dualTensorHomEquiv, AlgebraTensorModule.curry_apply_apply, coe_comp, |
| 216 | + LinearEquiv.coe_toLinearMap, coe_inl, Function.comp_apply, LinearEquiv.prod_apply, |
| 217 | + dualTensorHomEquivOfBasis_apply, map_zero, prodMapLinear_apply, dualTensorHom_prodMap_zero, |
| 218 | + trace_eq_contract_apply, contractLeft_apply, fst_apply, prodMap_apply, coprod_apply, id_coe, |
| 219 | + id.def, add_zero] |
| 220 | + · simp only [dualTensorHomEquiv, AlgebraTensorModule.curry_apply_apply, coe_comp, |
| 221 | + LinearEquiv.coe_toLinearMap, coe_inr, Function.comp_apply, LinearEquiv.prod_apply, map_zero, |
| 222 | + dualTensorHomEquivOfBasis_apply, prodMapLinear_apply, zero_prodMap_dualTensorHom, |
| 223 | + trace_eq_contract_apply, contractLeft_apply, snd_apply, prodMap_apply, coprod_apply, id_coe, |
| 224 | + id.def, zero_add] |
| 225 | +#align linear_map.trace_prod_map LinearMap.trace_prodMap |
| 226 | + |
| 227 | +variable {R M N} |
| 228 | + |
| 229 | +theorem trace_prodMap' (f : M →ₗ[R] M) (g : N →ₗ[R] N) : |
| 230 | + trace R (M × N) (prodMap f g) = trace R M f + trace R N g := by |
| 231 | + have h := ext_iff.1 (trace_prodMap R M N) (f, g) |
| 232 | + simp only [coe_comp, Function.comp_apply, prodMap_apply, coprod_apply, id_coe, id.def, |
| 233 | + prodMapLinear_apply] at h |
| 234 | + exact h |
| 235 | +#align linear_map.trace_prod_map' LinearMap.trace_prodMap' |
| 236 | + |
| 237 | +variable (R M N) |
| 238 | + |
| 239 | +open TensorProduct Function |
| 240 | + |
| 241 | +theorem trace_tensorProduct : compr₂ (mapBilinear R M N M N) (trace R (M ⊗ N)) = |
| 242 | + compl₁₂ (lsmul R R : R →ₗ[R] R →ₗ[R] R) (trace R M) (trace R N) := by |
| 243 | + apply |
| 244 | + (compl₁₂_inj (show Surjective (dualTensorHom R M M) from (dualTensorHomEquiv R M M).surjective) |
| 245 | + (show Surjective (dualTensorHom R N N) from (dualTensorHomEquiv R N N).surjective)).1 |
| 246 | + ext (f m g n) |
| 247 | + simp only [AlgebraTensorModule.curry_apply_apply, toFun_eq_coe, TensorProduct.curry_apply, |
| 248 | + coe_restrictScalars, compl₁₂_apply, compr₂_apply, mapBilinear_apply, |
| 249 | + trace_eq_contract_apply, contractLeft_apply, lsmul_apply, Algebra.id.smul_eq_mul, |
| 250 | + map_dualTensorHom, dualDistrib_apply] |
| 251 | +#align linear_map.trace_tensor_product LinearMap.trace_tensorProduct |
| 252 | + |
| 253 | +theorem trace_comp_comm : |
| 254 | + compr₂ (llcomp R M N M) (trace R M) = compr₂ (llcomp R N M N).flip (trace R N) := by |
| 255 | + apply |
| 256 | + (compl₁₂_inj (show Surjective (dualTensorHom R N M) from (dualTensorHomEquiv R N M).surjective) |
| 257 | + (show Surjective (dualTensorHom R M N) from (dualTensorHomEquiv R M N).surjective)).1 |
| 258 | + ext (g m f n) |
| 259 | + simp only [AlgebraTensorModule.curry_apply_apply, compl₁₂_apply, compr₂_apply, llcomp_apply', |
| 260 | + comp_dualTensorHom, map_smulₛₗ, RingHom.id_apply, trace_eq_contract_apply, contractLeft_apply, |
| 261 | + smul_eq_mul, mul_comm, flip_apply] |
| 262 | +#align linear_map.trace_comp_comm LinearMap.trace_comp_comm |
| 263 | + |
| 264 | +variable {R M N} |
| 265 | + |
| 266 | +@[simp] |
| 267 | +theorem trace_transpose' (f : M →ₗ[R] M) : |
| 268 | + trace R _ (Module.Dual.transpose (R := R) f) = trace R M f := by |
| 269 | + rw [← comp_apply, trace_transpose] |
| 270 | +#align linear_map.trace_transpose' LinearMap.trace_transpose' |
| 271 | + |
| 272 | +theorem trace_tensorProduct' (f : M →ₗ[R] M) (g : N →ₗ[R] N) : |
| 273 | + trace R (M ⊗ N) (map f g) = trace R M f * trace R N g := by |
| 274 | + have h := ext_iff.1 (ext_iff.1 (trace_tensorProduct R M N) f) g |
| 275 | + simp only [compr₂_apply, mapBilinear_apply, compl₁₂_apply, lsmul_apply, |
| 276 | + Algebra.id.smul_eq_mul] at h |
| 277 | + exact h |
| 278 | +#align linear_map.trace_tensor_product' LinearMap.trace_tensorProduct' |
| 279 | + |
| 280 | +theorem trace_comp_comm' (f : M →ₗ[R] N) (g : N →ₗ[R] M) : |
| 281 | + trace R M (g ∘ₗ f) = trace R N (f ∘ₗ g) := by |
| 282 | + have h := ext_iff.1 (ext_iff.1 (trace_comp_comm R M N) g) f |
| 283 | + simp only [llcomp_apply', compr₂_apply, flip_apply] at h |
| 284 | + exact h |
| 285 | +#align linear_map.trace_comp_comm' LinearMap.trace_comp_comm' |
| 286 | + |
| 287 | +@[simp] |
| 288 | +theorem trace_conj' (f : M →ₗ[R] M) (e : M ≃ₗ[R] N) : trace R N (e.conj f) = trace R M f := by |
| 289 | + rw [e.conj_apply, trace_comp_comm', ← comp_assoc, LinearEquiv.comp_coe, |
| 290 | + LinearEquiv.self_trans_symm, LinearEquiv.refl_toLinearMap, id_comp] |
| 291 | +#align linear_map.trace_conj' LinearMap.trace_conj' |
| 292 | + |
| 293 | +theorem IsProj.trace {p : Submodule R M} {f : M →ₗ[R] M} (h : IsProj p f) [Module.Free R p] |
| 294 | + [Module.Finite R p] [Module.Free R (ker f)] [Module.Finite R (ker f)] : |
| 295 | + trace R M f = (finrank R p : R) := by |
| 296 | + rw [h.eq_conj_prodMap, trace_conj', trace_prodMap', trace_id, map_zero, add_zero] |
| 297 | +#align linear_map.is_proj.trace LinearMap.IsProj.trace |
| 298 | + |
| 299 | +end |
| 300 | + |
| 301 | +end LinearMap |
0 commit comments