@@ -69,12 +69,10 @@ theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y'
69
69
| _, _, _, _, rfl, rfl => rfl
70
70
#align bilin_form.coe_fn_congr LinearMap.BilinForm.coeFn_congr
71
71
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₂ _ _ _ _
74
73
#align bilin_form.add_left LinearMap.BilinForm.add_left
75
74
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₂ _ _ _ _
78
76
#align bilin_form.smul_left LinearMap.BilinForm.smul_left
79
77
80
78
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 _ _ _
83
81
theorem smul_right (a : R) (x y : M) : B x (a • y) = a * B x y := map_smul _ _ _
84
82
#align bilin_form.smul_right LinearMap.BilinForm.smul_right
85
83
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₂ _ _
88
85
#align bilin_form.zero_left LinearMap.BilinForm.zero_left
89
86
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 _
92
88
#align bilin_form.zero_right LinearMap.BilinForm.zero_right
93
89
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₂ _ _ _
96
91
#align bilin_form.neg_left LinearMap.BilinForm.neg_left
97
92
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 _ _
100
94
#align bilin_form.neg_right LinearMap.BilinForm.neg_right
101
95
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₂ _ _ _ _
104
97
#align bilin_form.sub_left LinearMap.BilinForm.sub_left
105
98
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 _ _ _
108
100
#align bilin_form.sub_right LinearMap.BilinForm.sub_right
109
101
110
102
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
126
118
theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H
127
119
#align bilin_form.ext LinearMap.BilinForm.ext
128
120
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 _ _
131
122
#align bilin_form.congr_fun LinearMap.BilinForm.congr_fun
132
123
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₂
135
125
#align bilin_form.ext_iff LinearMap.BilinForm.ext_iff
136
126
137
127
@[deprecated]
0 commit comments