@@ -28,24 +28,18 @@ universe u
28
28
variable {α : Type u}
29
29
30
30
/-- An ordered commutative monoid is a commutative monoid
31
- with a partial order such that
32
- * `a ≤ b → c * a ≤ c * b` (multiplication is monotone)
33
- * `a * b < a * c → b < c`.
31
+ with a partial order such that `a ≤ b → c * a ≤ c * b` (multiplication is monotone)
34
32
-/
35
33
@[protect_proj, ancestor comm_monoid partial_order]
36
34
class ordered_comm_monoid (α : Type *) extends comm_monoid α, partial_order α :=
37
35
(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b)
38
- (lt_of_mul_lt_mul_left : ∀ a b c : α, a * b < a * c → b < c)
39
36
40
37
/-- An ordered (additive) commutative monoid is a commutative monoid
41
- with a partial order such that
42
- * `a ≤ b → c + a ≤ c + b` (addition is monotone)
43
- * `a + b < a + c → b < c`.
38
+ with a partial order such that `a ≤ b → c + a ≤ c + b` (addition is monotone)
44
39
-/
45
40
@[protect_proj, ancestor add_comm_monoid partial_order]
46
41
class ordered_add_comm_monoid (α : Type *) extends add_comm_monoid α, partial_order α :=
47
42
(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b)
48
- (lt_of_add_lt_add_left : ∀ a b c : α, a + b < a + c → b < c)
49
43
50
44
attribute [to_additive] ordered_comm_monoid
51
45
@@ -56,27 +50,13 @@ instance ordered_comm_monoid.to_covariant_class_left (M : Type*) [ordered_comm_m
56
50
covariant_class M M (*) (≤) :=
57
51
{ elim := λ a b c bc, ordered_comm_monoid.mul_le_mul_left _ _ bc a }
58
52
59
- @[to_additive]
60
- instance ordered_comm_monoid.to_contravariant_class_left (M : Type *) [ordered_comm_monoid M] :
61
- contravariant_class M M (*) (<) :=
62
- { elim := λ a b c, ordered_comm_monoid.lt_of_mul_lt_mul_left _ _ _ }
63
-
64
53
/- This instance can be proven with `by apply_instance`. However, `with_bot ℕ` does not
65
54
pick up a `covariant_class M M (function.swap (*)) (≤)` instance without it (see PR #7940). -/
66
55
@[to_additive]
67
56
instance ordered_comm_monoid.to_covariant_class_right (M : Type *) [ordered_comm_monoid M] :
68
57
covariant_class M M (swap (*)) (≤) :=
69
58
covariant_swap_mul_le_of_covariant_mul_le M
70
59
71
- /- This instance can be proven with `by apply_instance`. However, by analogy with the
72
- instance `ordered_comm_monoid.to_covariant_class_right` above, I imagine that without
73
- this instance, some Type would not have a `contravariant_class M M (function.swap (*)) (≤)`
74
- instance. -/
75
- @[to_additive]
76
- instance ordered_comm_monoid.to_contravariant_class_right (M : Type *) [ordered_comm_monoid M] :
77
- contravariant_class M M (swap (*)) (<) :=
78
- contravariant_swap_mul_lt_of_contravariant_mul_lt M
79
-
80
60
end ordered_instances
81
61
82
62
/-- An `ordered_comm_monoid` with one-sided 'division' in the sense that
@@ -100,24 +80,12 @@ export has_exists_add_of_le (exists_add_of_le)
100
80
/-- A linearly ordered additive commutative monoid. -/
101
81
@[protect_proj, ancestor linear_order ordered_add_comm_monoid]
102
82
class linear_ordered_add_comm_monoid (α : Type *)
103
- extends linear_order α, ordered_add_comm_monoid α :=
104
- (lt_of_add_lt_add_left := λ x y z, by {
105
- -- type-class inference uses `a : linear_order α` which it can't unfold, unless we provide this!
106
- -- `lt_iff_le_not_le` gets filled incorrectly with `autoparam` if we don't provide that field.
107
- letI : linear_order α := by refine { le := le, lt := lt, lt_iff_le_not_le := _, .. }; assumption,
108
- apply lt_imp_lt_of_le_imp_le,
109
- exact λ h, add_le_add_left _ _ h _ })
83
+ extends linear_order α, ordered_add_comm_monoid α.
110
84
111
85
/-- A linearly ordered commutative monoid. -/
112
86
@[protect_proj, ancestor linear_order ordered_comm_monoid, to_additive]
113
87
class linear_ordered_comm_monoid (α : Type *)
114
- extends linear_order α, ordered_comm_monoid α :=
115
- (lt_of_mul_lt_mul_left := λ x y z, by {
116
- -- type-class inference uses `a : linear_order α` which it can't unfold, unless we provide this!
117
- -- `lt_iff_le_not_le` gets filled incorrectly with `autoparam` if we don't provide that field.
118
- letI : linear_order α := by refine { le := le, lt := lt, lt_iff_le_not_le := _, .. }; assumption,
119
- apply lt_imp_lt_of_le_imp_le,
120
- exact λ h, mul_le_mul_left _ _ h _ })
88
+ extends linear_order α, ordered_comm_monoid α.
121
89
122
90
/-- A linearly ordered commutative monoid with a zero element. -/
123
91
class linear_ordered_comm_monoid_with_zero (α : Type *)
@@ -158,8 +126,6 @@ def function.injective.ordered_comm_monoid [ordered_comm_monoid α] {β : Type*}
158
126
ordered_comm_monoid β :=
159
127
{ mul_le_mul_left := λ a b ab c, show f (c * a) ≤ f (c * b), by
160
128
{ rw [mul, mul], apply mul_le_mul_left', exact ab },
161
- lt_of_mul_lt_mul_left :=
162
- λ a b c bc, show f b < f c, from lt_of_mul_lt_mul_left' (by rwa [← mul, ← mul] : (f a) * _ < _),
163
129
..partial_order.lift f hf,
164
130
..hf.comm_monoid f one mul }
165
131
273
239
274
240
instance [ordered_comm_monoid α] : ordered_comm_monoid (with_zero α) :=
275
241
{ mul_le_mul_left := with_zero.mul_le_mul_left,
276
- lt_of_mul_lt_mul_left := with_zero.lt_of_mul_lt_mul_left,
277
242
..with_zero.comm_monoid_with_zero,
278
243
..with_zero.partial_order }
279
244
@@ -295,17 +260,6 @@ begin
295
260
add_le_add_left := this ,
296
261
..with_zero.partial_order,
297
262
..with_zero.add_comm_monoid, .. },
298
- { intros a b c h,
299
- have h' := lt_iff_le_not_le.1 h,
300
- rw lt_iff_le_not_le at ⊢,
301
- refine ⟨λ b h₂, _, λ h₂, h'.2 $ this _ _ h₂ _⟩,
302
- cases h₂, cases c with c,
303
- { cases h'.2 (this _ _ bot_le a) },
304
- { refine ⟨_, rfl, _⟩,
305
- cases a with a,
306
- { exact with_bot.some_le_some.1 h'.1 },
307
- { exact le_of_lt (lt_of_add_lt_add_left $
308
- with_bot.some_lt_some.1 h), } } },
309
263
{ intros a b h c ca h₂,
310
264
cases b with b,
311
265
{ rw le_antisymm h bot_le at h₂,
@@ -420,15 +374,6 @@ instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (with_top α) :=
420
374
simp only [some_eq_coe, ← coe_add, coe_le_coe] at h ⊢,
421
375
exact add_le_add_left h c
422
376
end ,
423
- lt_of_add_lt_add_left :=
424
- begin
425
- intros a b c h,
426
- rcases lt_iff_exists_coe.1 h with ⟨ab, hab, hlt⟩,
427
- rcases add_eq_coe.1 hab with ⟨a, b, rfl, rfl, rfl⟩,
428
- rw coe_lt_iff,
429
- rintro c rfl,
430
- exact lt_of_add_lt_add_left (coe_lt_coe.1 hlt)
431
- end ,
432
377
..with_top.partial_order, ..with_top.add_comm_monoid }
433
378
434
379
instance [linear_ordered_add_comm_monoid α] :
@@ -469,16 +414,6 @@ begin
469
414
add_le_add_left := this ,
470
415
..with_bot.partial_order,
471
416
..with_bot.add_comm_monoid, ..},
472
- { intros a b c h,
473
- have h' := h,
474
- rw lt_iff_le_not_le at h' ⊢,
475
- refine ⟨λ b h₂, _, λ h₂, h'.2 $ this _ _ h₂ _⟩,
476
- cases h₂, cases a with a,
477
- { exact (not_le_of_lt h).elim bot_le },
478
- cases c with c,
479
- { exact (not_le_of_lt h).elim bot_le },
480
- { exact ⟨_, rfl, le_of_lt (lt_of_add_lt_add_left $
481
- with_bot.some_lt_some.1 h)⟩ } },
482
417
{ intros a b h c ca h₂,
483
418
cases c with c, {cases h₂},
484
419
cases a with a; cases h₂,
@@ -733,12 +668,31 @@ class ordered_cancel_comm_monoid (α : Type u)
733
668
section ordered_cancel_comm_monoid
734
669
variables [ordered_cancel_comm_monoid α] {a b c d : α}
735
670
671
+ @[to_additive]
672
+ lemma ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left : ∀ a b c : α, a * b < a * c → b < c :=
673
+ λ a b c h, lt_of_le_not_le
674
+ (ordered_cancel_comm_monoid.le_of_mul_le_mul_left a b c h.le) $
675
+ mt (λ h, ordered_cancel_comm_monoid.mul_le_mul_left _ _ h _) (not_le_of_gt h)
676
+
677
+ @[to_additive]
678
+ instance ordered_cancel_comm_monoid.to_contravariant_class_left
679
+ (M : Type *) [ordered_cancel_comm_monoid M] :
680
+ contravariant_class M M (*) (<) :=
681
+ { elim := λ a b c, ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left _ _ _ }
682
+
683
+ /- This instance can be proven with `by apply_instance`. However, by analogy with the
684
+ instance `ordered_cancel_comm_monoid.to_covariant_class_right` above, I imagine that without
685
+ this instance, some Type would not have a `contravariant_class M M (function.swap (*)) (<)`
686
+ instance. -/
687
+ @[to_additive]
688
+ instance ordered_cancel_comm_monoid.to_contravariant_class_right
689
+ (M : Type *) [ordered_cancel_comm_monoid M] :
690
+ contravariant_class M M (swap (*)) (<) :=
691
+ contravariant_swap_mul_lt_of_contravariant_mul_lt M
692
+
736
693
@[priority 100 , to_additive] -- see Note [lower instance priority]
737
694
instance ordered_cancel_comm_monoid.to_ordered_comm_monoid : ordered_comm_monoid α :=
738
- { lt_of_mul_lt_mul_left := λ a b c h, lt_of_le_not_le
739
- (ordered_cancel_comm_monoid.le_of_mul_le_mul_left a b c h.le) $
740
- mt (λ h, ordered_cancel_comm_monoid.mul_le_mul_left _ _ h _) (not_le_of_gt h),
741
- ..‹ordered_cancel_comm_monoid α› }
695
+ { ..‹ordered_cancel_comm_monoid α› }
742
696
743
697
/-- Pullback an `ordered_cancel_comm_monoid` under an injective map.
744
698
See note [reducible non-instances]. -/
@@ -943,7 +897,6 @@ instance covariant_class_swap_mul_lt [has_lt α] [has_mul α]
943
897
@[to_additive]
944
898
instance [ordered_comm_monoid α] : ordered_comm_monoid (order_dual α) :=
945
899
{ mul_le_mul_left := λ a b h c, mul_le_mul_left' h c,
946
- lt_of_mul_lt_mul_left := λ a b c, lt_of_mul_lt_mul_left',
947
900
.. order_dual.partial_order α,
948
901
.. order_dual.comm_monoid }
949
902
@@ -995,13 +948,11 @@ instance : Π [linear_order α], linear_order (additive α) := id
995
948
996
949
instance [ordered_add_comm_monoid α] : ordered_comm_monoid (multiplicative α) :=
997
950
{ mul_le_mul_left := @ordered_add_comm_monoid.add_le_add_left α _,
998
- lt_of_mul_lt_mul_left := @ordered_add_comm_monoid.lt_of_add_lt_add_left α _,
999
951
..multiplicative.partial_order,
1000
952
..multiplicative.comm_monoid }
1001
953
1002
954
instance [ordered_comm_monoid α] : ordered_add_comm_monoid (additive α) :=
1003
955
{ add_le_add_left := @ordered_comm_monoid.mul_le_mul_left α _,
1004
- lt_of_add_lt_add_left := @ordered_comm_monoid.lt_of_mul_lt_mul_left α _,
1005
956
..additive.partial_order,
1006
957
..additive.add_comm_monoid }
1007
958
0 commit comments