@@ -232,12 +232,10 @@ lemma pos_mul_strict_mono_iff_pos_mul_mono_rev : pos_mul_strict_mono α ↔ pos_
232
232
lemma mul_pos_strict_mono_iff_mul_pos_mono_rev : mul_pos_strict_mono α ↔ mul_pos_mono_rev α :=
233
233
⟨@mul_pos_strict_mono.to_mul_pos_mono_rev _ _ _ _, @mul_pos_mono_rev.to_mul_pos_strict_mono _ _ _ _⟩
234
234
235
- lemma pos_mul_reflect_lt.to_pos_mul_mono {α : Type *} [mul_zero_class α] [linear_order α]
236
- [pos_mul_reflect_lt α] : pos_mul_mono α :=
235
+ lemma pos_mul_reflect_lt.to_pos_mul_mono [pos_mul_reflect_lt α] : pos_mul_mono α :=
237
236
⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ lt_of_mul_lt_mul_left h' x.prop⟩
238
237
239
- lemma mul_pos_reflect_lt.to_mul_pos_mono {α : Type *} [mul_zero_class α] [linear_order α]
240
- [mul_pos_reflect_lt α] : mul_pos_mono α :=
238
+ lemma mul_pos_reflect_lt.to_mul_pos_mono [mul_pos_reflect_lt α] : mul_pos_mono α :=
241
239
⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ lt_of_mul_lt_mul_right h' x.prop⟩
242
240
243
241
lemma pos_mul_mono.to_pos_mul_reflect_lt [pos_mul_mono α] : pos_mul_reflect_lt α :=
@@ -246,13 +244,11 @@ lemma pos_mul_mono.to_pos_mul_reflect_lt [pos_mul_mono α] : pos_mul_reflect_lt
246
244
lemma mul_pos_mono.to_mul_pos_reflect_lt [mul_pos_mono α] : mul_pos_reflect_lt α :=
247
245
⟨λ x a b h, lt_of_not_le $ λ h', h.not_le $ mul_le_mul_of_nonneg_right h' x.prop⟩
248
246
249
- lemma pos_mul_mono_iff_pos_mul_reflect_lt {α : Type *} [mul_zero_class α] [linear_order α] :
250
- pos_mul_mono α ↔ pos_mul_reflect_lt α :=
251
- ⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _⟩
247
+ lemma pos_mul_mono_iff_pos_mul_reflect_lt : pos_mul_mono α ↔ pos_mul_reflect_lt α :=
248
+ ⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _ _⟩
252
249
253
- lemma mul_pos_mono_iff_mul_pos_reflect_lt {α : Type *} [mul_zero_class α] [linear_order α] :
254
- mul_pos_mono α ↔ mul_pos_reflect_lt α :=
255
- ⟨@mul_pos_mono.to_mul_pos_reflect_lt _ _ _ _, @mul_pos_reflect_lt.to_mul_pos_mono _ _ _⟩
250
+ lemma mul_pos_mono_iff_mul_pos_reflect_lt : mul_pos_mono α ↔ mul_pos_reflect_lt α :=
251
+ ⟨@mul_pos_mono.to_mul_pos_reflect_lt _ _ _ _, @mul_pos_reflect_lt.to_mul_pos_mono _ _ _ _⟩
256
252
257
253
end linear_order
258
254
@@ -268,7 +264,7 @@ variables [preorder α]
268
264
lemma left.mul_pos [pos_mul_strict_mono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b :=
269
265
by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha
270
266
271
- alias left.mul_pos ← _root_. mul_pos
267
+ alias left.mul_pos ← mul_pos
272
268
273
269
lemma mul_neg_of_pos_of_neg [pos_mul_strict_mono α] (ha : 0 < a) (hb : b < 0 ) : a * b < 0 :=
274
270
by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha
0 commit comments