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

Commit bf22ab3

Browse files
committed
feat(linear_algebra/bilinear_form): Unique adjoints with respect to a nondegenerate bilinear form (#7071)
1 parent a4f59bd commit bf22ab3

File tree

2 files changed

+62
-1
lines changed

2 files changed

+62
-1
lines changed

docs/undergrad.yaml

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -205,7 +205,7 @@ Bilinear and Quadratic Forms Over a Vector Space:
205205
polar form of a quadratic: 'quadratic_form.polar'
206206
Orthogonality:
207207
orthogonal elements: 'bilin_form.is_ortho'
208-
adjoint endomorphism:
208+
adjoint endomorphism: 'bilin_form.left_adjoint_of_nondegenerate'
209209
Sylvester's law of inertia:
210210
real classification:
211211
complex classification:

src/linear_algebra/bilinear_form.lean

Lines changed: 61 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1453,4 +1453,65 @@ begin
14531453
erw [add_right, show B m.1 y = 0, by rw hB₂; exact m.2 y hy, hm, add_zero]
14541454
end
14551455

1456+
section linear_adjoints
1457+
1458+
lemma comp_left_injective (B : bilin_form R₁ M₁) (hB : B.nondegenerate) :
1459+
function.injective B.comp_left :=
1460+
λ φ ψ h, begin
1461+
ext w,
1462+
refine eq_of_sub_eq_zero (hB _ _),
1463+
intro v,
1464+
rw [sub_left, ← comp_left_apply, ← comp_left_apply, ← h, sub_self]
1465+
end
1466+
1467+
lemma is_adjoint_pair_unique_of_nondegenerate (B : bilin_form R₁ M₁) (hB : B.nondegenerate)
1468+
(φ ψ₁ ψ₂ : M₁ →ₗ[R₁] M₁) (hψ₁ : is_adjoint_pair B B ψ₁ φ) (hψ₂ : is_adjoint_pair B B ψ₂ φ) :
1469+
ψ₁ = ψ₂ :=
1470+
B.comp_left_injective hB $ ext $ λ v w, by rw [comp_left_apply, comp_left_apply, hψ₁, hψ₂]
1471+
1472+
variable [finite_dimensional K V]
1473+
1474+
/-- Given bilinear forms `B₁, B₂` where `B₂` is nondegenerate, `symm_comp_of_nondegenerate`
1475+
is the linear map `B₂.to_lin⁻¹ ∘ B₁.to_lin`. -/
1476+
noncomputable def symm_comp_of_nondegenerate
1477+
(B₁ B₂ : bilin_form K V) (hB₂ : B₂.nondegenerate) : V →ₗ[K] V :=
1478+
(B₂.to_dual hB₂).symm.to_linear_map.comp B₁.to_lin
1479+
1480+
lemma comp_symm_comp_of_nondegenerate_apply (B₁ : bilin_form K V)
1481+
{B₂ : bilin_form K V} (hB₂ : B₂.nondegenerate) (v : V) :
1482+
to_lin B₂ (B₁.symm_comp_of_nondegenerate B₂ hB₂ v) = to_lin B₁ v :=
1483+
by erw [symm_comp_of_nondegenerate, linear_equiv.apply_symm_apply (B₂.to_dual hB₂) _]
1484+
1485+
@[simp]
1486+
lemma symm_comp_of_nondegenerate_left_apply (B₁ : bilin_form K V)
1487+
{B₂ : bilin_form K V} (hB₂ : B₂.nondegenerate) (v w : V) :
1488+
B₂ (symm_comp_of_nondegenerate B₁ B₂ hB₂ w) v = B₁ w v :=
1489+
begin
1490+
conv_lhs { rw [← bilin_form.to_lin_apply, comp_symm_comp_of_nondegenerate_apply] },
1491+
refl,
1492+
end
1493+
1494+
/-- Given the nondegenerate bilinear form `B` and the linear map `φ`,
1495+
`left_adjoint_of_nondegenerate` provides the left adjoint of `φ` with respect to `B`.
1496+
The lemma proving this property is `bilin_form.is_adjoint_pair_left_adjoint_of_nondegenerate`. -/
1497+
noncomputable def left_adjoint_of_nondegenerate
1498+
(B : bilin_form K V) (hB : B.nondegenerate) (φ : V →ₗ[K] V) : V →ₗ[K] V :=
1499+
symm_comp_of_nondegenerate (B.comp_right φ) B hB
1500+
1501+
lemma is_adjoint_pair_left_adjoint_of_nondegenerate
1502+
(B : bilin_form K V) (hB : B.nondegenerate) (φ : V →ₗ[K] V) :
1503+
is_adjoint_pair B B (B.left_adjoint_of_nondegenerate hB φ) φ :=
1504+
λ x y, (B.comp_right φ).symm_comp_of_nondegenerate_left_apply hB y x
1505+
1506+
/-- Given the nondegenerate bilinear form `B`, the linear map `φ` has a unique left adjoint given by
1507+
`bilin_form.left_adjoint_of_nondegenerate`. -/
1508+
theorem is_adjoint_pair_iff_eq_of_nondegenerate
1509+
(B : bilin_form K V) (hB : B.nondegenerate) (ψ φ : V →ₗ[K] V) :
1510+
is_adjoint_pair B B ψ φ ↔ ψ = B.left_adjoint_of_nondegenerate hB φ :=
1511+
⟨λ h, B.is_adjoint_pair_unique_of_nondegenerate hB φ ψ _ h
1512+
(is_adjoint_pair_left_adjoint_of_nondegenerate _ _ _),
1513+
λ h, h.symm ▸ is_adjoint_pair_left_adjoint_of_nondegenerate _ _ _⟩
1514+
1515+
end linear_adjoints
1516+
14561517
end bilin_form

0 commit comments

Comments
 (0)