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

Commit d6a1fc0

Browse files
committed
feat(algebra/ordered_monoid): correct definition (#8877)
Our definition of `ordered_monoid` is not the usual one, and this PR corrects that. The standard definition just says ``` (mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) ``` while we currently have an extra axiom ``` (lt_of_add_lt_add_left : ∀ a b c : α, a + b < a + c → b < c) ``` (This was introduced in ancient times, 7d8e3f3, with no indication of where the definition came from. I couldn't find it in other sources.) As @urkud pointed out a while ago [on Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/ordered_comm_monoid), these really are different. The second axiom *does* automatically hold for cancellative ordered monoids, however. This PR simply drops the axiom. In practice this causes no harm, because all the interesting examples are cancellative. There's a little bit of cleanup to do. In `src/algebra/ordered_sub.lean` four lemmas break, so I just added the necessary `[contravariant_class _ _ (+) (<)]` hypothesis. These lemmas aren't currently used in mathlib, but presumably where they are needed this typeclass will be available. Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent ce31c1c commit d6a1fc0

File tree

9 files changed

+41
-105
lines changed

9 files changed

+41
-105
lines changed

counterexamples/canonically_ordered_comm_semiring_two_mul.lean

Lines changed: 1 addition & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -222,9 +222,6 @@ instance order_bot : order_bot L :=
222222
bot_le := bot_le,
223223
..(infer_instance : partial_order L) }
224224

225-
lemma lt_of_add_lt_add_left : ∀ (a b c : L), a + b < a + c → b < c :=
226-
λ (a : L) {b c : L}, (add_lt_add_iff_left a).mp
227-
228225
lemma le_iff_exists_add : ∀ (a b : L), a ≤ b ↔ ∃ (c : L), b = a + c :=
229226
begin
230227
rintros ⟨⟨an, a2⟩, ha⟩ ⟨⟨bn, b2⟩, hb⟩,
@@ -271,8 +268,7 @@ begin
271268
end
272269

273270
instance can : canonically_ordered_comm_semiring L :=
274-
{ lt_of_add_lt_add_left := lt_of_add_lt_add_left,
275-
le_iff_exists_add := le_iff_exists_add,
271+
{ le_iff_exists_add := le_iff_exists_add,
276272
eq_zero_or_eq_zero_of_mul_eq_zero := eq_zero_or_eq_zero_of_mul_eq_zero,
277273
..(infer_instance : order_bot L),
278274
..(infer_instance : ordered_comm_semiring L) }

counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,8 @@
11
/-
2-
Copyright (c) 2021 Johan Commelin (inspired by Kevin Buzzard, copied by Damiano Testa).
2+
Copyright (c) 2021 Johan Commelin.
33
All rights reserved.
44
Released under Apache 2.0 license as described in the file LICENSE.
5-
Authors: Johan Commelin (inspired by Kevin Buzzard, copied by Damiano Testa)
5+
Authors: Johan Commelin, Damiano Testa, Kevin Buzzard
66
-/
77
import algebra.ordered_monoid
88

@@ -68,7 +68,6 @@ instance : linear_ordered_comm_monoid_with_zero foo :=
6868
zero_mul := by boom,
6969
mul_zero := by boom,
7070
mul_le_mul_left := by { rintro ⟨⟩ ⟨⟩ h ⟨⟩; revert h; dec_trivial },
71-
lt_of_mul_lt_mul_left := by { rintro ⟨⟩ ⟨⟩ ⟨⟩; dec_trivial },
7271
zero_le_one := dec_trivial,
7372
.. foo.linear_order,
7473
.. foo.comm_monoid }

src/algebra/ordered_monoid.lean

Lines changed: 27 additions & 76 deletions
Original file line numberDiff line numberDiff line change
@@ -28,24 +28,18 @@ universe u
2828
variable {α : Type u}
2929

3030
/-- 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)
3432
-/
3533
@[protect_proj, ancestor comm_monoid partial_order]
3634
class ordered_comm_monoid (α : Type*) extends comm_monoid α, partial_order α :=
3735
(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)
3936

4037
/-- 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)
4439
-/
4540
@[protect_proj, ancestor add_comm_monoid partial_order]
4641
class ordered_add_comm_monoid (α : Type*) extends add_comm_monoid α, partial_order α :=
4742
(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)
4943

5044
attribute [to_additive] ordered_comm_monoid
5145

@@ -56,27 +50,13 @@ instance ordered_comm_monoid.to_covariant_class_left (M : Type*) [ordered_comm_m
5650
covariant_class M M (*) (≤) :=
5751
{ elim := λ a b c bc, ordered_comm_monoid.mul_le_mul_left _ _ bc a }
5852

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-
6453
/- This instance can be proven with `by apply_instance`. However, `with_bot ℕ` does not
6554
pick up a `covariant_class M M (function.swap (*)) (≤)` instance without it (see PR #7940). -/
6655
@[to_additive]
6756
instance ordered_comm_monoid.to_covariant_class_right (M : Type*) [ordered_comm_monoid M] :
6857
covariant_class M M (swap (*)) (≤) :=
6958
covariant_swap_mul_le_of_covariant_mul_le M
7059

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-
8060
end ordered_instances
8161

8262
/-- 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)
10080
/-- A linearly ordered additive commutative monoid. -/
10181
@[protect_proj, ancestor linear_order ordered_add_comm_monoid]
10282
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 α.
11084

11185
/-- A linearly ordered commutative monoid. -/
11286
@[protect_proj, ancestor linear_order ordered_comm_monoid, to_additive]
11387
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 α.
12189

12290
/-- A linearly ordered commutative monoid with a zero element. -/
12391
class linear_ordered_comm_monoid_with_zero (α : Type*)
@@ -158,8 +126,6 @@ def function.injective.ordered_comm_monoid [ordered_comm_monoid α] {β : Type*}
158126
ordered_comm_monoid β :=
159127
{ mul_le_mul_left := λ a b ab c, show f (c * a) ≤ f (c * b), by
160128
{ 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) * _ < _),
163129
..partial_order.lift f hf,
164130
..hf.comm_monoid f one mul }
165131

@@ -273,7 +239,6 @@ end
273239

274240
instance [ordered_comm_monoid α] : ordered_comm_monoid (with_zero α) :=
275241
{ 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,
277242
..with_zero.comm_monoid_with_zero,
278243
..with_zero.partial_order }
279244

@@ -295,17 +260,6 @@ begin
295260
add_le_add_left := this,
296261
..with_zero.partial_order,
297262
..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), } } },
309263
{ intros a b h c ca h₂,
310264
cases b with b,
311265
{ rw le_antisymm h bot_le at h₂,
@@ -420,15 +374,6 @@ instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (with_top α) :=
420374
simp only [some_eq_coe, ← coe_add, coe_le_coe] at h ⊢,
421375
exact add_le_add_left h c
422376
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,
432377
..with_top.partial_order, ..with_top.add_comm_monoid }
433378

434379
instance [linear_ordered_add_comm_monoid α] :
@@ -469,16 +414,6 @@ begin
469414
add_le_add_left := this,
470415
..with_bot.partial_order,
471416
..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)⟩ } },
482417
{ intros a b h c ca h₂,
483418
cases c with c, {cases h₂},
484419
cases a with a; cases h₂,
@@ -733,12 +668,31 @@ class ordered_cancel_comm_monoid (α : Type u)
733668
section ordered_cancel_comm_monoid
734669
variables [ordered_cancel_comm_monoid α] {a b c d : α}
735670

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+
736693
@[priority 100, to_additive] -- see Note [lower instance priority]
737694
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 α› }
742696

743697
/-- Pullback an `ordered_cancel_comm_monoid` under an injective map.
744698
See note [reducible non-instances]. -/
@@ -943,7 +897,6 @@ instance covariant_class_swap_mul_lt [has_lt α] [has_mul α]
943897
@[to_additive]
944898
instance [ordered_comm_monoid α] : ordered_comm_monoid (order_dual α) :=
945899
{ 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',
947900
.. order_dual.partial_order α,
948901
.. order_dual.comm_monoid }
949902

@@ -995,13 +948,11 @@ instance : Π [linear_order α], linear_order (additive α) := id
995948

996949
instance [ordered_add_comm_monoid α] : ordered_comm_monoid (multiplicative α) :=
997950
{ 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 α _,
999951
..multiplicative.partial_order,
1000952
..multiplicative.comm_monoid }
1001953

1002954
instance [ordered_comm_monoid α] : ordered_add_comm_monoid (additive α) :=
1003955
{ 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 α _,
1005956
..additive.partial_order,
1006957
..additive.add_comm_monoid }
1007958

src/algebra/ordered_sub.lean

Lines changed: 9 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -413,8 +413,8 @@ protected lemma lt_sub_iff_left_of_le (hc : add_le_cancellable c) (h : c ≤ b)
413413
a < b - c ↔ c + a < b :=
414414
by { rw [add_comm], exact hc.lt_sub_iff_right_of_le h }
415415

416-
protected lemma lt_of_sub_lt_sub_left_of_le (hb : add_le_cancellable b) (hca : c ≤ a)
417-
(h : a - b < a - c) : c < b :=
416+
protected lemma lt_of_sub_lt_sub_left_of_le [contravariant_class α α (+) (<)]
417+
(hb : add_le_cancellable b) (hca : c ≤ a) (h : a - b < a - c) : c < b :=
418418
begin
419419
conv_lhs at h { rw [← sub_add_cancel_of_le hca] },
420420
exact lt_of_add_lt_add_left (hb.lt_add_of_sub_lt_right h),
@@ -441,8 +441,9 @@ protected lemma sub_inj_right (hab : add_le_cancellable (a - b)) (h₁ : b ≤ a
441441
(h₃ : a - b = a - c) : b = c :=
442442
by { rw ← hab.inj, rw [sub_add_cancel_of_le h₁, h₃, sub_add_cancel_of_le h₂] }
443443

444-
protected lemma sub_lt_sub_iff_left_of_le_of_le (hb : add_le_cancellable b)
445-
(hab : add_le_cancellable (a - b)) (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < a - c ↔ c < b :=
444+
protected lemma sub_lt_sub_iff_left_of_le_of_le [contravariant_class α α (+) (<)]
445+
(hb : add_le_cancellable b) (hab : add_le_cancellable (a - b)) (h₁ : b ≤ a) (h₂ : c ≤ a) :
446+
a - b < a - c ↔ c < b :=
446447
begin
447448
refine ⟨hb.lt_of_sub_lt_sub_left_of_le h₂, _⟩,
448449
intro h, refine (sub_le_sub_left' h.le _).lt_of_ne _,
@@ -516,7 +517,8 @@ lemma lt_sub_iff_left_of_le (h : c ≤ b) : a < b - c ↔ c + a < b :=
516517
contravariant.add_le_cancellable.lt_sub_iff_left_of_le h
517518

518519
/-- See `lt_of_sub_lt_sub_left` for a stronger statement in a linear order. -/
519-
lemma lt_of_sub_lt_sub_left_of_le (hca : c ≤ a) (h : a - b < a - c) : c < b :=
520+
lemma lt_of_sub_lt_sub_left_of_le [contravariant_class α α (+) (<)]
521+
(hca : c ≤ a) (h : a - b < a - c) : c < b :=
520522
contravariant.add_le_cancellable.lt_of_sub_lt_sub_left_of_le hca h
521523

522524
lemma sub_le_sub_iff_left' (h : c ≤ a) : a - b ≤ a - c ↔ c ≤ b :=
@@ -533,7 +535,8 @@ lemma sub_inj_right (h₁ : b ≤ a) (h₂ : c ≤ a) (h₃ : a - b = a - c) : b
533535
contravariant.add_le_cancellable.sub_inj_right h₁ h₂ h₃
534536

535537
/-- See `sub_lt_sub_iff_left_of_le` for a stronger statement in a linear order. -/
536-
lemma sub_lt_sub_iff_left_of_le_of_le (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < a - c ↔ c < b :=
538+
lemma sub_lt_sub_iff_left_of_le_of_le [contravariant_class α α (+) (<)]
539+
(h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < a - c ↔ c < b :=
537540
contravariant.add_le_cancellable.sub_lt_sub_iff_left_of_le_of_le
538541
contravariant.add_le_cancellable h₁ h₂
539542

src/algebra/punit_instances.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -56,8 +56,7 @@ intros; trivial <|> simp only [eq_iff_true_of_subsingleton]
5656

5757
instance : canonically_ordered_add_monoid punit :=
5858
by refine
59-
{ lt_of_add_lt_add_left := λ _ _ _, id,
60-
le_iff_exists_add := λ _ _, iff_of_true _ ⟨star, subsingleton.elim _ _⟩,
59+
{ le_iff_exists_add := λ _ _, iff_of_true _ ⟨star, subsingleton.elim _ _⟩,
6160
.. punit.comm_ring, .. punit.complete_boolean_algebra, .. };
6261
intros; trivial
6362

src/data/multiset/basic.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -410,8 +410,7 @@ theorem le_iff_exists_add {s t : multiset α} : s ≤ t ↔ ∃ u, t = s + u :=
410410
λ ⟨u, e⟩, e.symm ▸ le_add_right _ _⟩
411411

412412
instance : canonically_ordered_add_monoid (multiset α) :=
413-
{ lt_of_add_lt_add_left := λ a b c, (add_lt_add_iff_left a).mp,
414-
le_iff_exists_add := @le_iff_exists_add _,
413+
{ le_iff_exists_add := @le_iff_exists_add _,
415414
bot := 0,
416415
bot_le := multiset.zero_le,
417416
..multiset.ordered_cancel_add_comm_monoid }

src/data/nat/enat.lean

Lines changed: 0 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -259,14 +259,6 @@ instance : ordered_add_comm_monoid enat :=
259259
enat.cases_on c (by simp)
260260
(λ c, ⟨λ h, and.intro trivial (h₁ h.2),
261261
λ _, add_le_add_left (h₂ _) c⟩),
262-
lt_of_add_lt_add_left := λ a b c, enat.cases_on a
263-
(λ h, by simpa [lt_irrefl] using h)
264-
(λ a, enat.cases_on b
265-
(λ h, absurd h (not_lt_of_ge (by rw add_top; exact le_top)))
266-
(λ b, enat.cases_on c
267-
(λ _, coe_lt_top _)
268-
(λ c h, coe_lt_coe.2 (by rw [← coe_add, ← coe_add, coe_lt_coe] at h;
269-
exact lt_of_add_lt_add_left h)))),
270262
..enat.linear_order,
271263
..enat.add_comm_monoid }
272264

src/data/real/nnreal.lean

Lines changed: 0 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -267,8 +267,6 @@ instance : order_bot ℝ≥0 :=
267267
instance : canonically_linear_ordered_add_monoid ℝ≥0 :=
268268
{ add_le_add_left := assume a b h c,
269269
nnreal.coe_le_coe.mp $ (add_le_add_left (nnreal.coe_le_coe.mpr h) c),
270-
lt_of_add_lt_add_left := assume a b c bc,
271-
nnreal.coe_lt_coe.mp $ lt_of_add_lt_add_left (nnreal.coe_lt_coe.mpr bc),
272270
le_iff_exists_add := assume ⟨a, ha⟩ ⟨b, hb⟩,
273271
iff.intro
274272
(assume h : a ≤ b,

src/set_theory/cardinal.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -306,7 +306,6 @@ instance : order_bot cardinal.{u} :=
306306

307307
instance : canonically_ordered_comm_semiring cardinal.{u} :=
308308
{ add_le_add_left := λ a b h c, cardinal.add_le_add_left _ h,
309-
lt_of_add_lt_add_left := λ a b c, lt_imp_lt_of_le_imp_le (cardinal.add_le_add_left _),
310309
le_iff_exists_add := @cardinal.le_iff_exists_add,
311310
eq_zero_or_eq_zero_of_mul_eq_zero := @cardinal.eq_zero_or_eq_zero_of_mul_eq_zero,
312311
..cardinal.order_bot,

0 commit comments

Comments
 (0)