Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit d078950

Browse files
committed
feat(linear_algebra/bilinear_form): equivalence with matrices, given a basis (#5095)
This PR defines the equivalence between bilinear forms on an arbitrary module and matrices, given a basis of that module. It updates the existing equivalence between bilinear forms on `n → R` so that the overall structure of the file matches that of `linear_algebra.matrix`, i.e. there are two pairs of equivs, one for `n → R` and one for any `M` with a basis. Changes: - `bilin_form_equiv_matrix`, `bilin_form.to_matrix` and `matrix.to_bilin_form` have been consolidated as linear equivs `bilin_form.to_matrix'` and `matrix.to_bilin'`. Other declarations have been renamed accordingly. - `quadratic_form.to_matrix` and `matrix.to_quadratic_form` are renamed by analogy to `quadratic_form.to_matrix'` and `matrix.to_quadratic_form'` - replaced some `classical.decidable_eq` in lemma statements with instance parameters, because otherwise we have to use `congr` to apply these lemmas when a `decidable_eq` instance is available Additions: - `bilin_form.to_matrix` and `matrix.to_bilin`: given a basis, the equivalences between bilinear forms on any module and matrices - lemmas from `to_matrix'` and `to_bilin'` have been ported to `to_matrix` and `to_bilin` - `bilin_form.congr`, a dependency of `bilin_form.to_matrix`, states that `bilin_form R M` and `bilin_form R M'` are linearly equivalent if `M` and `M'` are - assorted lemmas involving `std_basis` Deletions: - `bilin_form.to_matrix_smul`: use `linear_equiv.map_smul` instead Co-authored-by: Anne Baanen <Vierkantor@users.noreply.github.com>
1 parent c06c616 commit d078950

File tree

4 files changed

+399
-117
lines changed

4 files changed

+399
-117
lines changed

src/linear_algebra/basis.lean

Lines changed: 21 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -515,7 +515,7 @@ section module
515515
variables {η : Type*} {ιs : η → Type*} {Ms : η → Type*}
516516
variables [ring R] [∀i, add_comm_group (Ms i)] [∀i, module R (Ms i)]
517517

518-
lemma linear_independent_std_basis
518+
lemma linear_independent_std_basis [decidable_eq η]
519519
(v : Πj, ιs j → (Ms j)) (hs : ∀i, linear_independent R (v i)) :
520520
linear_independent R (λ (ji : Σ j, ιs j), std_basis R Ms ji.1 (v ji.1 ji.2)) :=
521521
begin
@@ -546,7 +546,7 @@ end
546546

547547
variable [fintype η]
548548

549-
lemma is_basis_std_basis (s : Πj, ιs j → (Ms j)) (hs : ∀j, is_basis R (s j)) :
549+
lemma is_basis_std_basis [decidable_eq η] (s : Πj, ιs j → (Ms j)) (hs : ∀j, is_basis R (s j)) :
550550
is_basis R (λ (ji : Σ j, ιs j), std_basis R Ms ji.1 (s ji.1 ji.2)) :=
551551
begin
552552
split,
@@ -569,20 +569,36 @@ end
569569
section
570570
variables (R η)
571571

572-
lemma is_basis_fun₀ : is_basis R
572+
lemma is_basis_fun₀ [decidable_eq η] : is_basis R
573573
(λ (ji : Σ (j : η), unit),
574574
(std_basis R (λ (i : η), R) (ji.fst)) 1) :=
575-
@is_basis_std_basis R η (λi:η, unit) (λi:η, R) _ _ _ _ (λ _ _, (1 : R))
575+
@is_basis_std_basis R η (λi:η, unit) (λi:η, R) _ _ _ _ _ (λ _ _, (1 : R))
576576
(assume i, @is_basis_singleton_one _ _ _ _)
577577

578-
lemma is_basis_fun : is_basis R (λ i, std_basis R (λi:η, R) i 1) :=
578+
lemma is_basis_fun [decidable_eq η] : is_basis R (λ i, std_basis R (λi:η, R) i 1) :=
579579
begin
580580
apply (is_basis_fun₀ R η).comp (λ i, ⟨i, punit.star⟩),
581581
apply bijective_iff_has_inverse.2,
582582
use sigma.fst,
583583
simp [function.left_inverse, function.right_inverse]
584584
end
585585

586+
@[simp] lemma is_basis_fun_repr [decidable_eq η] (x : η → R) (i : η) :
587+
(pi.is_basis_fun R η).repr x i = x i :=
588+
begin
589+
conv_rhs { rw ← (pi.is_basis_fun R η).total_repr x },
590+
rw [finsupp.total_apply, finsupp.sum_fintype],
591+
show (pi.is_basis_fun R η).repr x i =
592+
(∑ j, λ i, (pi.is_basis_fun R η).repr x j • std_basis R (λ _, R) j 1 i) i,
593+
rw [finset.sum_apply, finset.sum_eq_single i],
594+
{ simp only [pi.smul_apply, smul_eq_mul, std_basis_same, mul_one] },
595+
{ rintros b - hb, simp only [std_basis_ne _ _ _ _ hb.symm, smul_zero] },
596+
{ intro,
597+
have := finset.mem_univ i,
598+
contradiction },
599+
{ intros, apply zero_smul },
600+
end
601+
586602
end
587603

588604
end module

0 commit comments

Comments
 (0)