Skip to content

Commit 67f52fe

Browse files
chore: backports for leanprover/lean4#4814 (part 6) (#15313)
See #15245 Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com>
1 parent 037c8b9 commit 67f52fe

File tree

16 files changed

+131
-87
lines changed

16 files changed

+131
-87
lines changed

Mathlib/Algebra/BigOperators/Group/Multiset.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -22,7 +22,7 @@ and sums indexed by finite sets.
2222

2323
assert_not_exists MonoidWithZero
2424

25-
variable {F ι α β γ : Type*}
25+
variable {F ι α β β' γ : Type*}
2626

2727
namespace Multiset
2828

@@ -124,27 +124,27 @@ theorem pow_count [DecidableEq α] (a : α) : a ^ s.count a = (s.filter (Eq a)).
124124
rw [filter_eq, prod_replicate]
125125

126126
@[to_additive]
127-
theorem prod_hom [CommMonoid β] (s : Multiset α) {F : Type*} [FunLike F α β]
127+
theorem prod_hom (s : Multiset α) {F : Type*} [FunLike F α β]
128128
[MonoidHomClass F α β] (f : F) :
129129
(s.map f).prod = f s.prod :=
130130
Quotient.inductionOn s fun l => by simp only [l.prod_hom f, quot_mk_to_coe, map_coe, prod_coe]
131131

132132
@[to_additive]
133-
theorem prod_hom' [CommMonoid β] (s : Multiset ι) {F : Type*} [FunLike F α β]
133+
theorem prod_hom' (s : Multiset ι) {F : Type*} [FunLike F α β]
134134
[MonoidHomClass F α β] (f : F)
135135
(g : ι → α) : (s.map fun i => f <| g i).prod = f (s.map g).prod := by
136136
convert (s.map g).prod_hom f
137137
exact (map_map _ _ _).symm
138138

139139
@[to_additive]
140-
theorem prod_hom₂ [CommMonoid β] [CommMonoid γ] (s : Multiset ι) (f : α → β → γ)
140+
theorem prod_hom₂ [CommMonoid γ] (s : Multiset ι) (f : α → β → γ)
141141
(hf : ∀ a b c d, f (a * b) (c * d) = f a c * f b d) (hf' : f 1 1 = 1) (f₁ : ι → α)
142142
(f₂ : ι → β) : (s.map fun i => f (f₁ i) (f₂ i)).prod = f (s.map f₁).prod (s.map f₂).prod :=
143143
Quotient.inductionOn s fun l => by
144144
simp only [l.prod_hom₂ f hf hf', quot_mk_to_coe, map_coe, prod_coe]
145145

146146
@[to_additive]
147-
theorem prod_hom_rel [CommMonoid β] (s : Multiset ι) {r : α → β → Prop} {f : ι → α} {g : ι → β}
147+
theorem prod_hom_rel (s : Multiset ι) {r : α → β → Prop} {f : ι → α} {g : ι → β}
148148
(h₁ : r 1 1) (h₂ : ∀ ⦃a b c⦄, r b c → r (f a * b) (g a * c)) :
149149
r (s.map f).prod (s.map g).prod :=
150150
Quotient.inductionOn s fun l => by
@@ -163,7 +163,7 @@ theorem prod_map_pow {n : ℕ} : (m.map fun i => f i ^ n).prod = (m.map f).prod
163163
m.prod_hom' (powMonoidHom n : α →* α) f
164164

165165
@[to_additive]
166-
theorem prod_map_prod_map (m : Multiset β) (n : Multiset γ) {f : β → γ → α} :
166+
theorem prod_map_prod_map (m : Multiset β') (n : Multiset γ) {f : β' → γ → α} :
167167
prod (m.map fun a => prod <| n.map fun b => f a b) =
168168
prod (n.map fun b => prod <| m.map fun a => f a b) :=
169169
Multiset.induction_on m (by simp) fun a m ih => by simp [ih]

Mathlib/Algebra/Module/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -148,7 +148,7 @@ end Function
148148

149149
namespace Set
150150
section SMulZeroClass
151-
variable [Zero R] [Zero M] [SMulZeroClass R M]
151+
variable [Zero M] [SMulZeroClass R M]
152152

153153
lemma indicator_smul_apply (s : Set α) (r : α → R) (f : α → M) (a : α) :
154154
indicator s (fun a ↦ r a • f a) a = r a • indicator s f a := by

Mathlib/Algebra/Module/Defs.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -546,11 +546,12 @@ end AddCommGroup
546546

547547
section Module
548548

549-
variable [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M]
549+
variable [Ring R] [AddCommGroup M] [Module R M]
550550

551551
section SMulInjective
552552

553553
variable (R)
554+
variable [NoZeroSMulDivisors R M]
554555

555556
theorem smul_left_injective {x : M} (hx : x ≠ 0) : Function.Injective fun c : R => c • x :=
556557
fun c d h =>

Mathlib/Algebra/Order/Group/Lattice.lean

Lines changed: 29 additions & 14 deletions
Original file line numberDiff line numberDiff line change
@@ -40,50 +40,65 @@ open Function
4040
variable {α β : Type*}
4141

4242
section Group
43-
variable [Lattice α] [Group α] [CovariantClass α α (· * ·) (· ≤ ·)]
44-
[CovariantClass α α (swap (· * ·)) (· ≤ ·)]
43+
variable [Lattice α] [Group α]
4544

4645
-- Special case of Bourbaki A.VI.9 (1)
4746
@[to_additive]
48-
lemma mul_sup (a b c : α) : c * (a ⊔ b) = c * a ⊔ c * b := (OrderIso.mulLeft _).map_sup _ _
47+
lemma mul_sup [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) :
48+
c * (a ⊔ b) = c * a ⊔ c * b :=
49+
(OrderIso.mulLeft _).map_sup _ _
4950

5051
@[to_additive]
51-
lemma sup_mul (a b c : α) : (a ⊔ b) * c = a * c ⊔ b * c := (OrderIso.mulRight _).map_sup _ _
52+
lemma sup_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) :
53+
(a ⊔ b) * c = a * c ⊔ b * c :=
54+
(OrderIso.mulRight _).map_sup _ _
5255

5356
@[to_additive]
54-
lemma mul_inf (a b c : α) : c * (a ⊓ b) = c * a ⊓ c * b := (OrderIso.mulLeft _).map_inf _ _
57+
lemma mul_inf [CovariantClass α α (· * ·) (· ≤ ·)] (a b c : α) :
58+
c * (a ⊓ b) = c * a ⊓ c * b :=
59+
(OrderIso.mulLeft _).map_inf _ _
5560

5661
@[to_additive]
57-
lemma inf_mul (a b c : α) : (a ⊓ b) * c = a * c ⊓ b * c := (OrderIso.mulRight _).map_inf _ _
62+
lemma inf_mul [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) :
63+
(a ⊓ b) * c = a * c ⊓ b * c :=
64+
(OrderIso.mulRight _).map_inf _ _
65+
66+
@[to_additive]
67+
lemma sup_div [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) :
68+
(a ⊔ b) / c = a / c ⊔ b / c :=
69+
(OrderIso.divRight _).map_sup _ _
70+
71+
@[to_additive]
72+
lemma inf_div [CovariantClass α α (swap (· * ·)) (· ≤ ·)] (a b c : α) :
73+
(a ⊓ b) / c = a / c ⊓ b / c :=
74+
(OrderIso.divRight _).map_inf _ _
75+
76+
section
77+
variable [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)]
5878

59-
-- Special case of Bourbaki A.VI.9 (2)
6079
@[to_additive] lemma inv_sup (a b : α) : (a ⊔ b)⁻¹ = a⁻¹ ⊓ b⁻¹ := (OrderIso.inv α).map_sup _ _
6180

6281
@[to_additive] lemma inv_inf (a b : α) : (a ⊓ b)⁻¹ = a⁻¹ ⊔ b⁻¹ := (OrderIso.inv α).map_inf _ _
6382

6483
@[to_additive]
6584
lemma div_sup (a b c : α) : c / (a ⊔ b) = c / a ⊓ c / b := (OrderIso.divLeft c).map_sup _ _
6685

67-
@[to_additive]
68-
lemma sup_div (a b c : α) : (a ⊔ b) / c = a / c ⊔ b / c := (OrderIso.divRight _).map_sup _ _
69-
7086
@[to_additive]
7187
lemma div_inf (a b c : α) : c / (a ⊓ b) = c / a ⊔ c / b := (OrderIso.divLeft c).map_inf _ _
7288

73-
@[to_additive]
74-
lemma inf_div (a b c : α) : (a ⊓ b) / c = a / c ⊓ b / c := (OrderIso.divRight _).map_inf _ _
75-
7689
-- In fact 0 ≤ n•a implies 0 ≤ a, see L. Fuchs, "Partially ordered algebraic systems"
7790
-- Chapter V, 1.E
7891
-- See also `one_le_pow_iff` for the existing version in linear orders
7992
@[to_additive]
8093
lemma pow_two_semiclosed
81-
[CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a : α} (ha : 1 ≤ a ^ 2) : 1 ≤ a := by
94+
{a : α} (ha : 1 ≤ a ^ 2) : 1 ≤ a := by
8295
suffices this : (a ⊓ 1) * (a ⊓ 1) = a ⊓ 1 by
8396
rwa [← inf_eq_right, ← mul_right_eq_self]
8497
rw [mul_inf, inf_mul, ← pow_two, mul_one, one_mul, inf_assoc, inf_left_idem, inf_comm,
8598
inf_assoc, inf_of_le_left ha]
8699

100+
end
101+
87102
end Group
88103

89104
variable [Lattice α] [CommGroup α]

Mathlib/Algebra/Order/Group/Unbundled/Abs.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -203,6 +203,8 @@ variable [Group α] [LinearOrder α] {a b : α}
203203
@[to_additive] lemma isSquare_mabs : IsSquare |a|ₘ ↔ IsSquare a :=
204204
mabs_by_cases (IsSquare · ↔ _) Iff.rfl isSquare_inv
205205

206+
@[to_additive] lemma lt_of_mabs_lt : |a|ₘ < b → a < b := (le_mabs_self _).trans_lt
207+
206208
variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α}
207209

208210
@[to_additive (attr := simp) abs_pos] lemma one_lt_mabs : 1 < |a|ₘ ↔ a ≠ 1 := by
@@ -245,8 +247,6 @@ variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)]
245247

246248
@[to_additive] lemma inv_lt_of_mabs_lt (h : |a|ₘ < b) : b⁻¹ < a := (mabs_lt.mp h).1
247249

248-
@[to_additive] lemma lt_of_mabs_lt : |a|ₘ < b → a < b := (le_mabs_self _).trans_lt
249-
250250
@[to_additive] lemma max_div_min_eq_mabs' (a b : α) : max a b / min a b = |a / b|ₘ := by
251251
rcases le_total a b with ab | ba
252252
· rw [max_eq_right ab, min_eq_left ab, mabs_of_le_one, inv_div]

Mathlib/Algebra/Order/Group/Unbundled/Basic.lean

Lines changed: 22 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -210,9 +210,21 @@ end TypeclassesRightLT
210210

211211
section TypeclassesLeftRightLE
212212

213-
variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)]
213+
variable [LE α] [CovariantClass α α (· * ·) (· ≤ ·)]
214214
{a b c d : α}
215215

216+
@[to_additive (attr := simp)]
217+
theorem div_le_self_iff (a : α) {b : α} : a / b ≤ a ↔ 1 ≤ b := by
218+
simp [div_eq_mul_inv]
219+
220+
@[to_additive (attr := simp)]
221+
theorem le_div_self_iff (a : α) {b : α} : a ≤ a / b ↔ b ≤ 1 := by
222+
simp [div_eq_mul_inv]
223+
224+
alias ⟨_, sub_le_self⟩ := sub_le_self_iff
225+
226+
variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)]
227+
216228
@[to_additive (attr := simp)]
217229
theorem inv_le_inv_iff : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := by
218230
rw [← mul_le_mul_iff_left a, ← mul_le_mul_iff_right b]
@@ -225,23 +237,21 @@ theorem mul_inv_le_inv_mul_iff : a * b⁻¹ ≤ d⁻¹ * c ↔ d * a ≤ c * b :
225237
rw [← mul_le_mul_iff_left d, ← mul_le_mul_iff_right b, mul_inv_cancel_left, mul_assoc,
226238
inv_mul_cancel_right]
227239

228-
@[to_additive (attr := simp)]
229-
theorem div_le_self_iff (a : α) {b : α} : a / b ≤ a ↔ 1 ≤ b := by
230-
simp [div_eq_mul_inv]
231-
232-
@[to_additive (attr := simp)]
233-
theorem le_div_self_iff (a : α) {b : α} : a ≤ a / b ↔ b ≤ 1 := by
234-
simp [div_eq_mul_inv]
235-
236-
alias ⟨_, sub_le_self⟩ := sub_le_self_iff
237-
238240
end TypeclassesLeftRightLE
239241

240242
section TypeclassesLeftRightLT
241243

242-
variable [LT α] [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (swap (· * ·)) (· < ·)]
244+
variable [LT α] [CovariantClass α α (· * ·) (· < ·)]
243245
{a b c d : α}
244246

247+
@[to_additive (attr := simp)]
248+
theorem div_lt_self_iff (a : α) {b : α} : a / b < a ↔ 1 < b := by
249+
simp [div_eq_mul_inv]
250+
251+
alias ⟨_, sub_lt_self⟩ := sub_lt_self_iff
252+
253+
variable [CovariantClass α α (swap (· * ·)) (· < ·)]
254+
245255
@[to_additive (attr := simp)]
246256
theorem inv_lt_inv_iff : a⁻¹ < b⁻¹ ↔ b < a := by
247257
rw [← mul_lt_mul_iff_left a, ← mul_lt_mul_iff_right b]
@@ -266,12 +276,6 @@ theorem mul_inv_lt_inv_mul_iff : a * b⁻¹ < d⁻¹ * c ↔ d * a < c * b := by
266276
rw [← mul_lt_mul_iff_left d, ← mul_lt_mul_iff_right b, mul_inv_cancel_left, mul_assoc,
267277
inv_mul_cancel_right]
268278

269-
@[to_additive (attr := simp)]
270-
theorem div_lt_self_iff (a : α) {b : α} : a / b < a ↔ 1 < b := by
271-
simp [div_eq_mul_inv]
272-
273-
alias ⟨_, sub_lt_self⟩ := sub_lt_self_iff
274-
275279
end TypeclassesLeftRightLT
276280

277281
section Preorder
@@ -796,5 +800,3 @@ theorem StrictAntiOn.inv (hf : StrictAntiOn f s) : StrictMonoOn (fun x => (f x)
796800
fun _ hx _ hy hxy => inv_lt_inv_iff.2 (hf hx hy hxy)
797801

798802
end
799-
800-

Mathlib/Algebra/Order/Hom/Monoid.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -205,7 +205,7 @@ theorem antitone_iff_map_nonneg : Antitone (f : α → β) ↔ ∀ a ≤ 0, 0
205205

206206
variable [CovariantClass β β (· + ·) (· < ·)]
207207

208-
theorem strictMono_iff_map_pos [iamhc : AddMonoidHomClass F α β] :
208+
theorem strictMono_iff_map_pos :
209209
StrictMono (f : α → β) ↔ ∀ a, 0 < a → 0 < f a := by
210210
refine ⟨fun h a => ?_, fun h a b hl => ?_⟩
211211
· rw [← map_zero f]

Mathlib/Algebra/Order/Monoid/Unbundled/MinMax.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ variable {α β : Type*}
1818
rules relating them. -/
1919

2020
section CommSemigroup
21-
variable [LinearOrder α] [CommSemigroup α] [CommSemigroup β]
21+
variable [LinearOrder α] [CommSemigroup β]
2222

2323
@[to_additive]
2424
lemma fn_min_mul_fn_max (f : α → β) (a b : α) : f (min a b) * f (max a b) = f a * f b := by
@@ -28,6 +28,8 @@ lemma fn_min_mul_fn_max (f : α → β) (a b : α) : f (min a b) * f (max a b)
2828
lemma fn_max_mul_fn_min (f : α → β) (a b : α) : f (max a b) * f (min a b) = f a * f b := by
2929
obtain h | h := le_total a b <;> simp [h, mul_comm]
3030

31+
variable [CommSemigroup α]
32+
3133
@[to_additive (attr := simp)]
3234
lemma min_mul_max (a b : α) : min a b * max a b = a * b := fn_min_mul_fn_max id _ _
3335

Mathlib/Algebra/Order/Sub/Canonical.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -247,7 +247,7 @@ theorem tsub_tsub_tsub_cancel_left (h : b ≤ a) : a - c - (a - b) = b - c :=
247247

248248
-- note: not generalized to `AddLECancellable` because `add_tsub_add_eq_tsub_left` isn't
249249
/-- The `tsub` version of `sub_sub_eq_add_sub`. -/
250-
theorem tsub_tsub_eq_add_tsub_of_le [ContravariantClass α α HAdd.hAdd LE.le]
250+
theorem tsub_tsub_eq_add_tsub_of_le
251251
(h : c ≤ b) : a - (b - c) = a + c - b := by
252252
obtain ⟨d, rfl⟩ := exists_add_of_le h
253253
rw [add_tsub_cancel_left c, add_comm a c, add_tsub_add_eq_tsub_left]
@@ -433,7 +433,7 @@ theorem tsub_add_min : a - b + min a b = a := by
433433
apply min_le_left
434434

435435
-- `Odd.tsub` requires `CanonicallyLinearOrderedSemiring`, which we don't have
436-
lemma Even.tsub [CanonicallyLinearOrderedAddCommMonoid α] [Sub α] [OrderedSub α]
436+
lemma Even.tsub
437437
[ContravariantClass α α (· + ·) (· ≤ ·)] {m n : α} (hm : Even m) (hn : Even n) :
438438
Even (m - n) := by
439439
obtain ⟨a, rfl⟩ := hm

Mathlib/Data/Finset/Basic.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2292,6 +2292,7 @@ theorem filter_cons {a : α} (s : Finset α) (ha : a ∉ s) :
22922292
· rw [filter_cons_of_pos _ _ _ ha h, singleton_disjUnion]
22932293
· rw [filter_cons_of_neg _ _ _ ha h, empty_disjUnion]
22942294

2295+
section
22952296
variable [DecidableEq α]
22962297

22972298
theorem filter_union (s₁ s₂ : Finset α) : (s₁ ∪ s₂).filter p = s₁.filter p ∪ s₂.filter p :=
@@ -2425,6 +2426,8 @@ theorem filter_union_filter_neg_eq [∀ x, Decidable (¬p x)] (s : Finset α) :
24252426
(s.filter p ∪ s.filter fun a => ¬p a) = s :=
24262427
filter_union_filter_of_codisjoint _ _ _ <| @codisjoint_hnot_right _ _ p
24272428

2429+
end
2430+
24282431
lemma filter_inj : s.filter p = t.filter p ↔ ∀ ⦃a⦄, p a → (a ∈ s ↔ a ∈ t) := by simp [ext_iff]
24292432

24302433
lemma filter_inj' : s.filter p = s.filter q ↔ ∀ ⦃a⦄, a ∈ s → (p a ↔ q a) := by simp [ext_iff]

0 commit comments

Comments
 (0)