Skip to content

Commit

Permalink
refactor(LinearAlgebra/BilinForm): Remove structure BilinForm from …
Browse files Browse the repository at this point in the history
…Mathlib, migrate all of `_root_.BilinForm` to `LinearMap.BilinForm` (#11278)

Remove `structure BilinForm` from `LinearAlgebra/BilinearForm/Basic` and migrate all of `_root_.BilinForm` to `LinearMap.BilinForm`

Closes: #10553 

This isn't the end of the story, as there's still a lot of overlap between `LinearAlgebra/BilinearForm` and `LinearAlgebra/SesquilinearForm` but that can be sorted out in subsequent PRs.

Supersedes: 

- #11057
- #11032
- #10432
- #10422



Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
Co-authored-by: Christopher Hoskin <christopher.hoskin@overleaf.com>
Co-authored-by: Christopher Hoskin <mans0954@users.noreply.github.com>
Co-authored-by: Vierkantor <vierkantor@vierkantor.com>
  • Loading branch information
5 people authored and uniwuni committed Apr 19, 2024
1 parent 078d254 commit 1ce336b
Show file tree
Hide file tree
Showing 12 changed files with 404 additions and 507 deletions.
254 changes: 84 additions & 170 deletions Mathlib/LinearAlgebra/BilinearForm/Basic.lean

Large diffs are not rendered by default.

14 changes: 10 additions & 4 deletions Mathlib/LinearAlgebra/BilinearForm/DualLattice.lean
Original file line number Diff line number Diff line change
Expand Up @@ -17,9 +17,13 @@ import Mathlib.LinearAlgebra.BilinearForm.Properties
Properly develop the material in the context of lattices.
-/

open LinearMap (BilinForm)

variable {R S M} [CommRing R] [Field S] [AddCommGroup M]
variable [Algebra R S] [Module R M] [Module S M] [IsScalarTower R S M]

namespace LinearMap

namespace BilinForm

variable (B : BilinForm S M)
Expand All @@ -31,7 +35,8 @@ def dualSubmodule (N : Submodule R M) : Submodule R M where
zero_mem' y _ := by rw [B.zero_left]; exact zero_mem _
smul_mem' r a ha y hy := by
convert (1 : Submodule R S).smul_mem r (ha y hy)
rw [← IsScalarTower.algebraMap_smul S r a, bilin_smul_left, Algebra.smul_def]
rw [← IsScalarTower.algebraMap_smul S r a]
simp only [algebraMap_smul, map_smul_of_tower, LinearMap.smul_apply]

lemma mem_dualSubmodule {N : Submodule R M} {x} :
x ∈ B.dualSubmodule N ↔ ∀ y ∈ N, B x y ∈ (1 : Submodule R S) := Iff.rfl
Expand Down Expand Up @@ -96,11 +101,12 @@ lemma dualSubmodule_span_of_basis {ι} [Finite ι] [DecidableEq ι]
· rw [Submodule.span_le]
rintro _ ⟨i, rfl⟩ y hy
obtain ⟨f, rfl⟩ := (mem_span_range_iff_exists_fun _).mp hy
simp only [sum_right, bilin_smul_right]
simp only [map_sum, map_smul]
apply sum_mem
rintro j -
rw [← IsScalarTower.algebraMap_smul S (f j), B.bilin_smul_right, apply_dualBasis_left,
mul_ite, mul_one, mul_zero, ← (algebraMap R S).map_zero, ← apply_ite]
rw [← IsScalarTower.algebraMap_smul S (f j), map_smul]
simp_rw [apply_dualBasis_left]
rw [smul_eq_mul, mul_ite, mul_one, mul_zero, ← (algebraMap R S).map_zero, ← apply_ite]
exact ⟨_, rfl⟩

lemma dualSubmodule_dualSubmodule_flip_of_basis {ι : Type*} [Finite ι]
Expand Down

0 comments on commit 1ce336b

Please sign in to comment.