Skip to content

Commit e521a25

Browse files
committed
refactor(LinearAlgebra/BilinearForm/Basic): Derive the *_left and *_right results from the map_*₂ and map_* results (#12124)
Following #11278 a number of the results in `LinearAlgebra/BilinearForm/Basic` are just special cases of results in `LinearAlgebra/BilinearMap`.
1 parent 1c6ef0c commit e521a25

File tree

1 file changed

+10
-20
lines changed

1 file changed

+10
-20
lines changed

Mathlib/LinearAlgebra/BilinearForm/Basic.lean

Lines changed: 10 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -69,12 +69,10 @@ theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y'
6969
| _, _, _, _, rfl, rfl => rfl
7070
#align bilin_form.coe_fn_congr LinearMap.BilinForm.coeFn_congr
7171

72-
theorem add_left (x y z : M) : B (x + y) z = B x z + B y z := by
73-
simp only [map_add, LinearMap.add_apply]
72+
theorem add_left (x y z : M) : B (x + y) z = B x z + B y z := map_add₂ _ _ _ _
7473
#align bilin_form.add_left LinearMap.BilinForm.add_left
7574

76-
theorem smul_left (a : R) (x y : M) : B (a • x) y = a * B x y := by
77-
simp only [map_smul, LinearMap.smul_apply, smul_eq_mul]
75+
theorem smul_left (a : R) (x y : M) : B (a • x) y = a * B x y := map_smul₂ _ _ _ _
7876
#align bilin_form.smul_left LinearMap.BilinForm.smul_left
7977

8078
theorem add_right (x y z : M) : B x (y + z) = B x y + B x z := map_add _ _ _
@@ -83,28 +81,22 @@ theorem add_right (x y z : M) : B x (y + z) = B x y + B x z := map_add _ _ _
8381
theorem smul_right (a : R) (x y : M) : B x (a • y) = a * B x y := map_smul _ _ _
8482
#align bilin_form.smul_right LinearMap.BilinForm.smul_right
8583

86-
theorem zero_left (x : M) : B 0 x = 0 := by
87-
rw [← @zero_smul R _ _ _ _ (0 : M), smul_left, zero_mul]
84+
theorem zero_left (x : M) : B 0 x = 0 := map_zero₂ _ _
8885
#align bilin_form.zero_left LinearMap.BilinForm.zero_left
8986

90-
theorem zero_right (x : M) : B x 0 = 0 := by
91-
rw [← @zero_smul R _ _ _ _ (0 : M), smul_right, zero_mul]
87+
theorem zero_right (x : M) : B x 0 = 0 := map_zero _
9288
#align bilin_form.zero_right LinearMap.BilinForm.zero_right
9389

94-
theorem neg_left (x y : M₁) : B₁ (-x) y = -B₁ x y := by
95-
rw [← @neg_one_smul R₁ _ _, smul_left, neg_one_mul]
90+
theorem neg_left (x y : M₁) : B₁ (-x) y = -B₁ x y := map_neg₂ _ _ _
9691
#align bilin_form.neg_left LinearMap.BilinForm.neg_left
9792

98-
theorem neg_right (x y : M₁) : B₁ x (-y) = -B₁ x y := by
99-
rw [← @neg_one_smul R₁ _ _, smul_right, neg_one_mul]
93+
theorem neg_right (x y : M₁) : B₁ x (-y) = -B₁ x y := map_neg _ _
10094
#align bilin_form.neg_right LinearMap.BilinForm.neg_right
10195

102-
theorem sub_left (x y z : M₁) : B₁ (x - y) z = B₁ x z - B₁ y z := by
103-
rw [sub_eq_add_neg, sub_eq_add_neg, add_left, neg_left]
96+
theorem sub_left (x y z : M₁) : B₁ (x - y) z = B₁ x z - B₁ y z := map_sub₂ _ _ _ _
10497
#align bilin_form.sub_left LinearMap.BilinForm.sub_left
10598

106-
theorem sub_right (x y z : M₁) : B₁ x (y - z) = B₁ x y - B₁ x z := by
107-
rw [sub_eq_add_neg, sub_eq_add_neg, add_right, neg_right]
99+
theorem sub_right (x y z : M₁) : B₁ x (y - z) = B₁ x y - B₁ x z := map_sub _ _ _
108100
#align bilin_form.sub_right LinearMap.BilinForm.sub_right
109101

110102
lemma smul_left_of_tower (r : S) (x y : M) : B (r • x) y = r • B x y := by
@@ -126,12 +118,10 @@ theorem coe_injective : Function.Injective ((fun B x y => B x y) : BilinForm R M
126118
theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H
127119
#align bilin_form.ext LinearMap.BilinForm.ext
128120

129-
theorem congr_fun (h : B = D) (x y : M) : B x y = D x y :=
130-
h ▸ rfl
121+
theorem congr_fun (h : B = D) (x y : M) : B x y = D x y := congr_fun₂ h _ _
131122
#align bilin_form.congr_fun LinearMap.BilinForm.congr_fun
132123

133-
theorem ext_iff : B = D ↔ ∀ x y, B x y = D x y :=
134-
⟨congr_fun, ext₂⟩
124+
theorem ext_iff : B = D ↔ ∀ x y, B x y = D x y := ext_iff₂
135125
#align bilin_form.ext_iff LinearMap.BilinForm.ext_iff
136126

137127
@[deprecated]

0 commit comments

Comments
 (0)