Skip to content

Commit a62da66

Browse files
chore: bump dependencies (#10954)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Riccardo Brasca <riccardo.brasca@gmail.com>
1 parent a3692d9 commit a62da66

File tree

14 files changed

+27
-68
lines changed

14 files changed

+27
-68
lines changed

Mathlib/Data/BitVec/Lemmas.lean

Lines changed: 1 addition & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -170,14 +170,9 @@ lemma toFin_nsmul (n : ℕ) (x : BitVec w) : toFin (n • x) = n • x.toFin :=
170170
lemma toFin_zsmul (z : ℤ) (x : BitVec w) : toFin (z • x) = z • x.toFin := rfl
171171
@[simp] lemma toFin_pow (n : ℕ) : toFin (x ^ n) = x.toFin ^ n := rfl
172172

173-
@[simp] lemma toFin_natCast (n : ℕ) : toFin (n : BitVec w) = n := by
173+
lemma toFin_natCast (n : ℕ) : toFin (n : BitVec w) = n := by
174174
apply toFin_inj.mpr; simp only [ofFin_natCast]
175175

176-
-- See Note [no_index around OfNat.ofNat]
177-
lemma toFin_ofNat (n : ℕ) :
178-
toFin (no_index (OfNat.ofNat n : BitVec w)) = OfNat.ofNat n := by
179-
simp only [OfNat.ofNat, BitVec.ofNat, and_pow_two_is_mod, Fin.ofNat']
180-
181176
end
182177

183178
/-!

Mathlib/Data/Int/Order/Basic.lean

Lines changed: 0 additions & 49 deletions
Original file line numberDiff line numberDiff line change
@@ -445,74 +445,25 @@ protected theorem sign_eq_ediv_abs (a : ℤ) : sign a = a / |a| :=
445445

446446
/-! ### `/` and ordering -/
447447

448-
449-
protected theorem ediv_mul_le (a : ℤ) {b : ℤ} (H : b ≠ 0) : a / b * b ≤ a :=
450-
le_of_sub_nonneg <| by rw [mul_comm, ← emod_def]; apply emod_nonneg _ H
451448
#align int.div_mul_le Int.ediv_mul_le
452-
453-
protected theorem ediv_le_of_le_mul {a b c : ℤ} (H : 0 < c) (H' : a ≤ b * c) : a / c ≤ b :=
454-
le_of_mul_le_mul_right (le_trans (Int.ediv_mul_le _ (ne_of_gt H)) H') H
455449
#align int.div_le_of_le_mul Int.ediv_le_of_le_mul
456-
457-
protected theorem mul_lt_of_lt_ediv {a b c : ℤ} (H : 0 < c) (H3 : a < b / c) : a * c < b :=
458-
lt_of_not_ge <| mt (Int.ediv_le_of_le_mul H) (not_le_of_gt H3)
459450
#align int.mul_lt_of_lt_div Int.mul_lt_of_lt_ediv
460-
461-
protected theorem mul_le_of_le_ediv {a b c : ℤ} (H1 : 0 < c) (H2 : a ≤ b / c) : a * c ≤ b :=
462-
le_trans (mul_le_mul_of_nonneg_right H2 (le_of_lt H1)) (Int.ediv_mul_le _ (ne_of_gt H1))
463451
#align int.mul_le_of_le_div Int.mul_le_of_le_ediv
464-
465-
protected theorem le_ediv_of_mul_le {a b c : ℤ} (H1 : 0 < c) (H2 : a * c ≤ b) : a ≤ b / c :=
466-
le_of_lt_add_one <|
467-
lt_of_mul_lt_mul_right (lt_of_le_of_lt H2 (lt_ediv_add_one_mul_self _ H1)) (le_of_lt H1)
468452
#align int.le_div_of_mul_le Int.le_ediv_of_mul_le
469-
470-
protected theorem le_ediv_iff_mul_le {a b c : ℤ} (H : 0 < c) : a ≤ b / c ↔ a * c ≤ b :=
471-
⟨Int.mul_le_of_le_ediv H, Int.le_ediv_of_mul_le H⟩
472453
#align int.le_div_iff_mul_le Int.le_ediv_iff_mul_le
473-
474-
protected theorem ediv_le_ediv {a b c : ℤ} (H : 0 < c) (H' : a ≤ b) : a / c ≤ b / c :=
475-
Int.le_ediv_of_mul_le H (le_trans (Int.ediv_mul_le _ (ne_of_gt H)) H')
476454
#align int.div_le_div Int.ediv_le_ediv
477-
478-
protected theorem ediv_lt_of_lt_mul {a b c : ℤ} (H : 0 < c) (H' : a < b * c) : a / c < b :=
479-
lt_of_not_ge <| mt (Int.mul_le_of_le_ediv H) (not_le_of_gt H')
480455
#align int.div_lt_of_lt_mul Int.ediv_lt_of_lt_mul
481-
482-
protected theorem lt_mul_of_ediv_lt {a b c : ℤ} (H1 : 0 < c) (H2 : a / c < b) : a < b * c :=
483-
lt_of_not_ge <| mt (Int.le_ediv_of_mul_le H1) (not_le_of_gt H2)
484456
#align int.lt_mul_of_div_lt Int.lt_mul_of_ediv_lt
485-
486-
protected theorem ediv_lt_iff_lt_mul {a b c : ℤ} (H : 0 < c) : a / c < b ↔ a < b * c :=
487-
⟨Int.lt_mul_of_ediv_lt H, Int.ediv_lt_of_lt_mul H⟩
488457
#align int.div_lt_iff_lt_mul Int.ediv_lt_iff_lt_mul
489-
490-
protected theorem le_mul_of_ediv_le {a b c : ℤ} (H1 : 0 ≤ b) (H2 : b ∣ a) (H3 : a / b ≤ c) :
491-
a ≤ c * b := by rw [← Int.ediv_mul_cancel H2]; exact mul_le_mul_of_nonneg_right H3 H1
492458
#align int.le_mul_of_div_le Int.le_mul_of_ediv_le
493-
494-
protected theorem lt_ediv_of_mul_lt {a b c : ℤ} (H1 : 0 ≤ b) (H2 : b ∣ c) (H3 : a * b < c) :
495-
a < c / b :=
496-
lt_of_not_ge <| mt (Int.le_mul_of_ediv_le H1 H2) (not_le_of_gt H3)
497459
#align int.lt_div_of_mul_lt Int.lt_ediv_of_mul_lt
498-
499-
protected theorem lt_ediv_iff_mul_lt {a b : ℤ} (c : ℤ) (H : 0 < c) (H' : c ∣ b) :
500-
a < b / c ↔ a * c < b :=
501-
⟨Int.mul_lt_of_lt_ediv H, Int.lt_ediv_of_mul_lt (le_of_lt H) H'⟩
502460
#align int.lt_div_iff_mul_lt Int.lt_ediv_iff_mul_lt
503-
504-
theorem ediv_pos_of_pos_of_dvd {a b : ℤ} (H1 : 0 < a) (H2 : 0 ≤ b) (H3 : b ∣ a) : 0 < a / b :=
505-
Int.lt_ediv_of_mul_lt H2 H3 (by rwa [zero_mul])
506461
#align int.div_pos_of_pos_of_dvd Int.ediv_pos_of_pos_of_dvd
507462

508463
theorem natAbs_eq_of_dvd_dvd {s t : ℤ} (hst : s ∣ t) (hts : t ∣ s) : natAbs s = natAbs t :=
509464
Nat.dvd_antisymm (natAbs_dvd_natAbs.mpr hst) (natAbs_dvd_natAbs.mpr hts)
510465
#align int.nat_abs_eq_of_dvd_dvd Int.natAbs_eq_of_dvd_dvd
511466

512-
theorem ediv_eq_ediv_of_mul_eq_mul {a b c d : ℤ} (H2 : d ∣ c) (H3 : b ≠ 0) (H4 : d ≠ 0)
513-
(H5 : a * d = b * c) : a / b = c / d :=
514-
Int.ediv_eq_of_eq_mul_right H3 <| by
515-
rw [← Int.mul_ediv_assoc _ H2]; exact (Int.ediv_eq_of_eq_mul_left H4 H5.symm).symm
516467
#align int.div_eq_div_of_mul_eq_mul Int.ediv_eq_ediv_of_mul_eq_mul
517468

518469
theorem ediv_dvd_of_dvd {s t : ℤ} (hst : s ∣ t) : t / s ∣ t := by

Mathlib/Data/List/Basic.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -3663,8 +3663,8 @@ variable [DecidableEq α]
36633663

36643664
theorem map_erase [DecidableEq β] {f : α → β} (finj : Injective f) {a : α} (l : List α) :
36653665
map f (l.erase a) = (map f l).erase (f a) := by
3666-
have this : Eq a = Eq (f a) ∘ f := by ext b; simp [finj.eq_iff]
3667-
simp [erase_eq_eraseP, erase_eq_eraseP, eraseP_map, this]; rfl
3666+
have this : (a == ·) = (f a == f ·) := by ext b; simp [beq_eq_decide, finj.eq_iff]
3667+
rw [erase_eq_eraseP, erase_eq_eraseP, eraseP_map, this]; rfl
36683668
#align list.map_erase List.map_erase
36693669

36703670
theorem map_foldl_erase [DecidableEq β] {f : α → β} (finj : Injective f) {l₁ l₂ : List α} :
@@ -3720,7 +3720,7 @@ theorem erase_diff_erase_sublist_of_sublist {a : α} :
37203720
| b :: l₁, l₂, h =>
37213721
if heq : b = a then by simp only [heq, erase_cons_head, diff_cons]; rfl
37223722
else by
3723-
simp only [erase_cons_head b l₁, erase_cons_tail l₁ heq,
3723+
simp only [erase_cons_head b l₁, erase_cons_tail l₁ (not_beq_of_ne heq),
37243724
diff_cons ((List.erase l₂ a)) (List.erase l₁ a) b, diff_cons l₂ l₁ b, erase_comm a b l₂]
37253725
have h' := h.erase b
37263726
rw [erase_cons_head] at h'

Mathlib/Data/List/BigOperators/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -581,7 +581,7 @@ theorem prod_map_erase [DecidableEq ι] [CommMonoid M] (f : ι → M) {a} :
581581
| b :: l, h => by
582582
obtain rfl | ⟨ne, h⟩ := Decidable.List.eq_or_ne_mem_of_mem h
583583
· simp only [map, erase_cons_head, prod_cons]
584-
· simp only [map, erase_cons_tail _ ne.symm, prod_cons, prod_map_erase _ h,
584+
· simp only [map, erase_cons_tail _ (not_beq_of_ne ne.symm), prod_cons, prod_map_erase _ h,
585585
mul_left_comm (f a) (f b)]
586586
#align list.prod_map_erase List.prod_map_erase
587587
#align list.sum_map_erase List.sum_map_erase

Mathlib/Data/List/Nodup.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -310,7 +310,7 @@ theorem Nodup.erase_eq_filter [DecidableEq α] {l} (d : Nodup l) (a : α) :
310310
symm
311311
rw [filter_eq_self]
312312
simpa [@eq_comm α] using m
313-
· rw [erase_cons_tail _ h, filter_cons_of_pos, IH]
313+
· rw [erase_cons_tail _ (not_beq_of_ne h), filter_cons_of_pos, IH]
314314
simp [h]
315315
#align list.nodup.erase_eq_filter List.Nodup.erase_eq_filter
316316

Mathlib/Data/List/Sigma.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -453,6 +453,10 @@ theorem mem_keys_kerase_of_ne {a₁ a₂} {l : List (Sigma β)} (h : a₁ ≠ a
453453

454454
theorem keys_kerase {a} {l : List (Sigma β)} : (kerase a l).keys = l.keys.erase a := by
455455
rw [keys, kerase, erase_eq_eraseP, eraseP_map, Function.comp]
456+
simp only [beq_eq_decide]
457+
congr
458+
funext
459+
simp
456460
#align list.keys_kerase List.keys_kerase
457461

458462
theorem kerase_kerase {a a'} {l : List (Sigma β)} :

Mathlib/Data/Multiset/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1052,7 +1052,7 @@ theorem erase_cons_head (a : α) (s : Multiset α) : (a ::ₘ s).erase a = s :=
10521052
@[simp]
10531053
theorem erase_cons_tail {a b : α} (s : Multiset α) (h : b ≠ a) :
10541054
(b ::ₘ s).erase a = b ::ₘ s.erase a :=
1055-
Quot.inductionOn s fun l => congr_arg _ <| List.erase_cons_tail l h
1055+
Quot.inductionOn s fun l => congr_arg _ <| List.erase_cons_tail l (not_beq_of_ne h)
10561056
#align multiset.erase_cons_tail Multiset.erase_cons_tail
10571057

10581058
@[simp]

Mathlib/Data/NNRat/Defs.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -328,7 +328,7 @@ theorem lt_toNNRat_iff_coe_lt {q : ℚ≥0} : q < toNNRat p ↔ ↑q < p :=
328328

329329
theorem toNNRat_mul (hp : 0 ≤ p) : toNNRat (p * q) = toNNRat p * toNNRat q := by
330330
rcases le_total 0 q with hq | hq
331-
· ext <;> simp [toNNRat, hp, hq, max_eq_left, mul_nonneg]
331+
· ext; simp [toNNRat, hp, hq, max_eq_left, mul_nonneg]
332332
· have hpq := mul_nonpos_of_nonneg_of_nonpos hp hq
333333
rw [toNNRat_eq_zero.2 hq, toNNRat_eq_zero.2 hpq, mul_zero]
334334
#align rat.to_nnrat_mul Rat.toNNRat_mul
@@ -378,7 +378,7 @@ theorem natAbs_num_coe : (q : ℚ).num.natAbs = q.num := rfl
378378
@[simp, norm_cast] lemma den_natCast (n : ℕ) : den n = 1 := rfl
379379

380380
theorem ext_num_den (hn : p.num = q.num) (hd : p.den = q.den) : p = q := by
381-
ext
381+
refine ext <| Rat.ext ?_ ?_
382382
· apply (Int.natAbs_inj_of_nonneg_of_nonneg _ _).1 hn
383383
exact Rat.num_nonneg.2 p.2
384384
exact Rat.num_nonneg.2 q.2

Mathlib/Init/Data/Rat/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
-/
66

7-
import Std.Data.Rat.Basic
7+
import Std.Data.Rat.Lemmas
88
import Mathlib.Mathport.Rename
99

1010
/-!
@@ -13,5 +13,5 @@ import Mathlib.Mathport.Rename
1313

1414
#align rat.ext Rat.ext
1515
#align rat Rat
16-
#align rat.ext_iff Rat.ext_iff
16+
#noalign rat.ext_iff
1717
#align rat.floor Rat.floor

Mathlib/Logic/Basic.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1285,3 +1285,10 @@ theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} [Decidable P] :
12851285
#align if_false_left_eq_and if_false_left
12861286

12871287
end ite
1288+
1289+
theorem not_beq_of_ne [BEq α] [LawfulBEq α] {a b : α} (ne : a ≠ b) : ¬(a == b) :=
1290+
fun h => ne (eq_of_beq h)
1291+
1292+
theorem beq_eq_decide [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by
1293+
rw [← beq_iff_eq a b]
1294+
cases a == b <;> simp

0 commit comments

Comments
 (0)