Skip to content

Commit 7a94bed

Browse files
feat: port LinearAlgebra.Multilinear.Basis (#3330)
1 parent 5a04e70 commit 7a94bed

File tree

2 files changed

+71
-0
lines changed

2 files changed

+71
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1236,6 +1236,7 @@ import Mathlib.LinearAlgebra.Matrix.Orthogonal
12361236
import Mathlib.LinearAlgebra.Matrix.Symmetric
12371237
import Mathlib.LinearAlgebra.Matrix.Trace
12381238
import Mathlib.LinearAlgebra.Multilinear.Basic
1239+
import Mathlib.LinearAlgebra.Multilinear.Basis
12391240
import Mathlib.LinearAlgebra.Multilinear.TensorProduct
12401241
import Mathlib.LinearAlgebra.Pi
12411242
import Mathlib.LinearAlgebra.Prod
Lines changed: 70 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,70 @@
1+
/-
2+
Copyright (c) 2021 Joseph Myers. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Joseph Myers
5+
6+
! This file was ported from Lean 3 source module linear_algebra.multilinear.basis
7+
! leanprover-community/mathlib commit ce11c3c2a285bbe6937e26d9792fda4e51f3fe1a
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.Basis
12+
import Mathlib.LinearAlgebra.Multilinear.Basic
13+
14+
/-!
15+
# Multilinear maps in relation to bases.
16+
17+
This file proves lemmas about the action of multilinear maps on basis vectors.
18+
19+
## TODO
20+
21+
* Refactor the proofs in terms of bases of tensor products, once there is an equivalent of
22+
`Basis.tensorProduct` for `pi_tensor_product`.
23+
24+
-/
25+
26+
27+
open MultilinearMap
28+
29+
variable {R : Type _} {ι : Type _} {n : ℕ} {M : Fin n → Type _} {M₂ : Type _} {M₃ : Type _}
30+
31+
variable [CommSemiring R] [AddCommMonoid M₂] [AddCommMonoid M₃] [∀ i, AddCommMonoid (M i)]
32+
33+
variable [∀ i, Module R (M i)] [Module R M₂] [Module R M₃]
34+
35+
/-- Two multilinear maps indexed by `Fin n` are equal if they are equal when all arguments are
36+
basis vectors. -/
37+
theorem Basis.ext_multilinear_fin {f g : MultilinearMap R M M₂} {ι₁ : Fin n → Type _}
38+
(e : ∀ i, Basis (ι₁ i) R (M i))
39+
(h : ∀ v : ∀ i, ι₁ i, (f fun i => e i (v i)) = g fun i => e i (v i)) : f = g := by
40+
induction' n with m hm
41+
· ext x
42+
convert h finZeroElim <;>
43+
-- Porting note: added below
44+
· rename_i x
45+
apply finZeroElim x
46+
· apply Function.LeftInverse.injective uncurry_curryLeft
47+
refine' Basis.ext (e 0) _
48+
intro i
49+
apply hm (Fin.tail e)
50+
intro j
51+
convert h (Fin.cons i j)
52+
iterate 2
53+
rw [curryLeft_apply]
54+
congr 1 with x
55+
refine' Fin.cases rfl (fun x => _) x
56+
dsimp [Fin.tail]
57+
rw [Fin.cons_succ, Fin.cons_succ]
58+
#align basis.ext_multilinear_fin Basis.ext_multilinear_fin
59+
60+
/-- Two multilinear maps indexed by a `Fintype` are equal if they are equal when all arguments
61+
are basis vectors. Unlike `Basis.ext_multilinear_fin`, this only uses a single basis; a
62+
dependently-typed version would still be true, but the proof would need a dependently-typed
63+
version of `dom_dom_congr`. -/
64+
theorem Basis.ext_multilinear [Finite ι] {f g : MultilinearMap R (fun _ : ι => M₂) M₃} {ι₁ : Type _}
65+
(e : Basis ι₁ R M₂) (h : ∀ v : ι → ι₁, (f fun i => e (v i)) = g fun i => e (v i)) : f = g := by
66+
cases nonempty_fintype ι
67+
exact
68+
(domDomCongr_eq_iff (Fintype.equivFin ι) f g).mp
69+
(Basis.ext_multilinear_fin (fun _ => e) fun i => h (i ∘ _))
70+
#align basis.ext_multilinear Basis.ext_multilinear

0 commit comments

Comments
 (0)