Skip to content

Commit

Permalink
fix: patch for std4#203
Browse files Browse the repository at this point in the history
  • Loading branch information
fgdorais committed Nov 1, 2023
1 parent 41442e1 commit eb40450
Show file tree
Hide file tree
Showing 22 changed files with 36 additions and 35 deletions.
4 changes: 2 additions & 2 deletions Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean
Original file line number Diff line number Diff line change
Expand Up @@ -161,7 +161,7 @@ theorem card_le_two_pow {x k : ℕ} :
exact
⟨p.pred, (Nat.pred_lt (Nat.Prime.ne_zero hp.1)).trans_le ((hmp p) hp),
Nat.succ_pred_eq_of_pos (Nat.Prime.pos hp.1)⟩
· simp [Nat.prod_factors m.succ_ne_zero, m.succ_sub_one]
· simp [Nat.prod_factors m.succ_ne_zero, m.add_one_sub_one]
-- The number of elements of `M x k` with `e + 1` squarefree is bounded by the number of subsets
-- of `[1, k]`.
calc
Expand Down Expand Up @@ -191,7 +191,7 @@ theorem card_le_two_pow_mul_sqrt {x k : ℕ} : card (M x k) ≤ 2 ^ k * Nat.sqrt
obtain ⟨a, b, hab₁, hab₂⟩ := Nat.sq_mul_squarefree_of_pos' hm'
obtain ⟨ham, hbm⟩ := Dvd.intro_left _ hab₁, Dvd.intro _ hab₁
refine'
⟨a, b, ⟨⟨⟨_, fun p hp => _⟩, hab₂⟩, ⟨_, fun p hp => _⟩⟩, by simp_rw [hab₁, m.succ_sub_one]⟩
⟨a, b, ⟨⟨⟨_, fun p hp => _⟩, hab₂⟩, ⟨_, fun p hp => _⟩⟩, by simp_rw [hab₁, m.add_one_sub_one]⟩
· exact (Nat.succ_le_succ_iff.mp (Nat.le_of_dvd hm' ham)).trans_lt hm.1
· exact hm.2 p ⟨hp.1, hp.2.trans ham⟩
· calc
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/AlgebraicTopology/DoldKan/Homotopies.lean
Original file line number Diff line number Diff line change
Expand Up @@ -94,7 +94,7 @@ theorem cs_down_0_not_rel_left (j : ℕ) : ¬c.Rel 0 j := by
/-- The sequence of maps which gives the null homotopic maps `Hσ` that shall be in
the inductive construction of the projections `P q : K[X] ⟶ K[X]` -/
def (q : ℕ) (n : ℕ) : X _[n] ⟶ X _[n + 1] :=
if n < q then 0 else (-1 : ℤ) ^ (n - q) • X.σ ⟨n - q, Nat.sub_lt_succ n q
if n < q then 0 else (-1 : ℤ) ^ (n - q) • X.σ ⟨n - q, Nat.lt_succ_of_le (Nat.sub_le _ _)
#align algebraic_topology.dold_kan.hσ AlgebraicTopology.DoldKan.hσ

/-- We can turn `hσ` into a datum that can be passed to `nullHomotopicMap'`. -/
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Control/Random.lean
Original file line number Diff line number Diff line change
Expand Up @@ -104,7 +104,7 @@ instance : BoundedRandom Nat where
let z ← rand (Fin (hi - lo).succ)
pure ⟨
lo + z.val, Nat.le_add_right _ _,
Nat.add_le_of_le_sub_left h (Nat.le_of_succ_le_succ z.isLt)
Nat.add_le_of_le_sub' h (Nat.le_of_succ_le_succ z.isLt)

instance : BoundedRandom Int where
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Fin/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1222,7 +1222,7 @@ theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a
theorem coe_neg_one : ↑(-1 : Fin (n + 1)) = n := by
cases n
· simp
rw [Fin.coe_neg, Fin.val_one, Nat.succ_sub_one, Nat.mod_eq_of_lt]
rw [Fin.coe_neg, Fin.val_one, Nat.add_one_sub_one, Nat.mod_eq_of_lt]
constructor
#align fin.coe_neg_one Fin.coe_neg_one

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Fin/Interval.lean
Original file line number Diff line number Diff line change
Expand Up @@ -219,7 +219,7 @@ theorem card_Ici : (Ici a).card = n - a := by
cases n with
| zero => exact Fin.elim0 a
| succ =>
rw [← card_map, map_valEmbedding_Ici, Nat.card_Icc, Nat.succ_sub_one]
rw [← card_map, map_valEmbedding_Ici, Nat.card_Icc, Nat.add_one_sub_one]
#align fin.card_Ici Fin.card_Ici

@[simp]
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/List/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -646,8 +646,8 @@ theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <;
theorem length_dropLast : ∀ l : List α, length l.dropLast = length l - 1
| [] | [_] => rfl
| a::b::l => by
rw [dropLast, length_cons, length_cons, length_dropLast (b::l), succ_sub_one, length_cons,
succ_sub_one]
rw [dropLast, length_cons, length_cons, length_dropLast (b::l), Nat.add_one_sub_one,
length_cons, Nat.add_one_sub_one]
simp
#align list.length_init List.length_dropLast

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/List/Chain.lean
Original file line number Diff line number Diff line change
Expand Up @@ -355,7 +355,7 @@ theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔
simp only [length_cons, get_cons_succ, Fin.zero_eta, get_cons_zero, zero_add, Fin.mk_one,
get_cons_cons_one, succ_sub_succ_eq_sub, nonpos_iff_eq_zero, add_eq_zero_iff, and_false,
tsub_zero, add_pos_iff, or_true, forall_true_left, and_congr_right_iff]
dsimp [succ_sub_one]
dsimp [Nat.add_one_sub_one]
exact fun _ => ⟨fun h i hi => h i (Nat.lt_of_succ_lt_succ hi),
fun h i hi => h i (Nat.succ_lt_succ hi)⟩

Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/Nat/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -295,11 +295,11 @@ theorem exists_eq_add_of_lt (h : m < n) : ∃ k : ℕ, n = m + k + 1 :=
/-! ### `pred` -/

@[simp]
theorem add_succ_sub_one (n m : ℕ) : n + succ m - 1 = n + m := by rw [add_succ, succ_sub_one]
theorem add_succ_sub_one (n m : ℕ) : n + succ m - 1 = n + m := by rw [add_succ, Nat.add_one_sub_one]
#align nat.add_succ_sub_one Nat.add_succ_sub_one

@[simp]
theorem succ_add_sub_one (n m : ℕ) : succ n + m - 1 = n + m := by rw [succ_add, succ_sub_one]
theorem succ_add_sub_one (n m : ℕ) : succ n + m - 1 = n + m := by rw [succ_add, Nat.add_one_sub_one]
#align nat.succ_add_sub_one Nat.succ_add_sub_one

theorem pred_eq_sub_one (n : ℕ) : pred n = n - 1 :=
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/Nat/Factorial/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -281,8 +281,8 @@ theorem ascFactorial_of_sub {n k : ℕ} (h : k < n) :
(n - k) * (n - k).ascFactorial k = (n - (k + 1)).ascFactorial (k + 1) := by
let t := n - k.succ
let ht : t = n - k.succ := rfl
suffices h' : n - k = t.succ by rw [← ht, h', succ_ascFactorial, ascFactorial_succ]
rw [ht, succ_eq_add_one, ← tsub_tsub_assoc (succ_le_of_lt h) (succ_pos _), succ_sub_one]
suffices h' : n - k = t.succ; · rw [← ht, h', succ_ascFactorial, ascFactorial_succ]
rw [ht, succ_eq_add_one, ← tsub_tsub_assoc (succ_le_of_lt h) (succ_pos _), Nat.add_one_sub_one]
#align nat.asc_factorial_of_sub Nat.ascFactorial_of_sub

theorem pow_succ_le_ascFactorial (n : ℕ) : ∀ k : ℕ, (n + 1) ^ k ≤ n.ascFactorial k
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Nat/Factorial/Cast.lean
Original file line number Diff line number Diff line change
Expand Up @@ -41,7 +41,7 @@ theorem cast_descFactorial :
rw [← ascPochhammer_eval_cast, ascPochhammer_nat_eq_descFactorial]
induction' b with b
· simp
· simp_rw [add_succ, succ_sub_one]
· simp_rw [add_succ, Nat.add_one_sub_one]
obtain h | h := le_total a b
· rw [descFactorial_of_lt (lt_succ_of_le h), descFactorial_of_lt (lt_succ_of_le _)]
rw [tsub_eq_zero_iff_le.mpr h, zero_add]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Nat/Order/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -323,7 +323,7 @@ theorem lt_mul_self_iff : ∀ {n : ℕ}, n < n * n ↔ 1 < n
theorem add_sub_one_le_mul (hm : m ≠ 0) (hn : n ≠ 0) : m + n - 1 ≤ m * n := by
cases m
· cases hm rfl
· rw [succ_add, succ_sub_one, succ_mul]
· rw [succ_add, Nat.add_one_sub_one, succ_mul]
exact add_le_add_right (le_mul_of_one_le_right' $ succ_le_iff.2 $ pos_iff_ne_zero.2 hn) _
#align nat.add_sub_one_le_mul Nat.add_sub_one_le_mul

Expand Down
6 changes: 3 additions & 3 deletions Mathlib/Data/Polynomial/Derivative.lean
Original file line number Diff line number Diff line change
Expand Up @@ -64,7 +64,7 @@ theorem coeff_derivative (p : R[X]) (n : ℕ) :
· intros
rw [Nat.cast_zero, mul_zero, zero_mul]
· intro _ H
rw [Nat.succ_sub_one, if_neg (mt (congr_arg Nat.succ) H.symm), mul_zero]
rw [Nat.add_one_sub_one, if_neg (mt (congr_arg Nat.succ) H.symm), mul_zero]
· rw [if_pos (add_tsub_cancel_right n 1).symm, mul_one, Nat.cast_add, Nat.cast_one,
mem_support_iff]
intro h
Expand Down Expand Up @@ -314,7 +314,7 @@ theorem derivative_mul {f g : R[X]} : derivative (f * g) = derivative f * g + f
cases n <;> cases m <;>
simp_rw [add_smul, mul_smul_comm, smul_mul_assoc, X_pow_mul_assoc, ← mul_assoc, ←
C_mul, mul_assoc, ← pow_add] <;>
simp [Nat.add_succ, Nat.succ_add, Nat.succ_sub_one, zero_smul, add_comm])
simp [Nat.add_succ, Nat.succ_add, Nat.add_one_sub_one, zero_smul, add_comm])
_ = derivative f * g + f * derivative g := by
conv =>
rhs
Expand Down Expand Up @@ -468,7 +468,7 @@ theorem derivative_pow_succ (p : R[X]) (n : ℕ) :
theorem derivative_pow (p : R[X]) (n : ℕ) :
derivative (p ^ n) = C (n : R) * p ^ (n - 1) * derivative p :=
Nat.casesOn n (by rw [pow_zero, derivative_one, Nat.cast_zero, C_0, zero_mul, zero_mul]) fun n =>
by rw [p.derivative_pow_succ n, n.succ_sub_one, n.cast_succ]
by rw [p.derivative_pow_succ n, Nat.add_one_sub_one, n.cast_succ]
#align polynomial.derivative_pow Polynomial.derivative_pow

theorem derivative_sq (p : R[X]) : derivative (p ^ 2) = C 2 * p * derivative p := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Stream/Init.lean
Original file line number Diff line number Diff line change
Expand Up @@ -608,7 +608,7 @@ theorem get?_take_succ (n : Nat) (s : Stream' α) :
@[simp] theorem dropLast_take {xs : Stream' α} :
(Stream'.take n xs).dropLast = Stream'.take (n-1) xs := by
cases n; case zero => simp
case succ n => rw [take_succ', List.dropLast_concat, Nat.succ_sub_one]
case succ n => rw [take_succ', List.dropLast_concat, Nat.add_one_sub_one]

@[simp]
theorem append_take_drop : ∀ (n : Nat) (s : Stream' α),
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/Sym/Card.lean
Original file line number Diff line number Diff line change
Expand Up @@ -193,8 +193,8 @@ theorem card_subtype_not_diag [Fintype α] :
/-- Finset **stars and bars** for the case `n = 2`. -/
theorem _root_.Finset.card_sym2 (s : Finset α) : s.sym2.card = s.card * (s.card + 1) / 2 := by
rw [← image_diag_union_image_offDiag, card_union_eq, Sym2.card_image_diag,
Sym2.card_image_offDiag, Nat.choose_two_right, add_comm, ← Nat.triangle_succ, Nat.succ_sub_one,
mul_comm]
Sym2.card_image_offDiag, Nat.choose_two_right, add_comm, ← Nat.triangle_succ,
Nat.add_one_sub_one, mul_comm]
rw [disjoint_left]
rintro m ha hb
rw [mem_image] at ha hb
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Init/Data/Nat/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -416,7 +416,7 @@ Many lemmas are proven more generally in mathlib `algebra/order/sub` -/

#align nat.sub_one Nat.sub_one

#align nat.succ_sub_one Nat.succ_sub_one
#align nat.succ_sub_one Nat.add_one_sub_one

#align nat.succ_pred_eq_of_pos Nat.succ_pred_eq_of_pos

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/NumberTheory/Cyclotomic/Discriminant.lean
Original file line number Diff line number Diff line change
Expand Up @@ -72,7 +72,7 @@ theorem discr_prime_pow_ne_two [IsCyclotomicExtension {p ^ (k + 1)} K L] [hp : F
haveI se : IsSeparable K L := (isGalois (p ^ (k + 1)) K L).to_isSeparable
rw [discr_powerBasis_eq_norm, finrank L hirr, hζ.powerBasis_gen _, ←
hζ.minpoly_eq_cyclotomic_of_irreducible hirr, PNat.pow_coe,
totient_prime_pow hp.out (succ_pos k), succ_sub_one]
totient_prime_pow hp.out (succ_pos k), Nat.add_one_sub_one]
have coe_two : ((2 : ℕ+) : ℕ) = 2 := rfl
have hp2 : p = 2 → k ≠ 0 := by
rintro rfl rfl
Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -69,7 +69,7 @@ private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ}
calc
(a ^ (p / 2) * (p / 2)! : ZMod p) = ∏ x in Ico 1 (p / 2).succ, a * x := by
rw [prod_mul_distrib, ← prod_natCast, prod_Ico_id_eq_factorial, prod_const, card_Ico,
succ_sub_one]; simp
Nat.add_one_sub_one]; simp
_ = ∏ x in Ico 1 (p / 2).succ, ↑((a * x : ZMod p).val) := by simp
_ = ∏ x in Ico 1 (p / 2).succ, (if (a * x : ZMod p).val ≤ p / 2 then (1 : ZMod p) else -1) *
(a * x : ZMod p).valMinAbs.natAbs :=
Expand Down
6 changes: 3 additions & 3 deletions Mathlib/NumberTheory/Wilson.lean
Original file line number Diff line number Diff line change
Expand Up @@ -54,15 +54,15 @@ theorem wilsons_lemma : ((p - 1)! : ZMod p) = -1 := by
symm
refine' prod_bij (fun a _ => (a : ZMod p).val) _ _ _ _
· intro a ha
rw [mem_Ico, ← Nat.succ_sub hp, Nat.succ_sub_one]
rw [mem_Ico, ← Nat.succ_sub hp, Nat.add_one_sub_one]
constructor
· apply Nat.pos_of_ne_zero; rw [← @val_zero p]
intro h; apply Units.ne_zero a (val_injective p h)
· exact val_lt _
· rintro a -; simp only [cast_id, nat_cast_val]
· intro _ _ _ _ h; rw [Units.ext_iff]; exact val_injective p h
· intro b hb
rw [mem_Ico, Nat.succ_le_iff, ← succ_sub hp, succ_sub_one, pos_iff_ne_zero] at hb
rw [mem_Ico, Nat.succ_le_iff, ← succ_sub hp, Nat.add_one_sub_one, pos_iff_ne_zero] at hb
refine' ⟨Units.mk0 b _, Finset.mem_univ _, _⟩
· intro h; apply hb.1; apply_fun val at h
simpa only [val_cast_of_lt hb.right, val_zero] using h
Expand All @@ -75,7 +75,7 @@ theorem prod_Ico_one_prime : ∏ x in Ico 1 p, (x : ZMod p) = -1 := by
conv =>
congr
congr
rw [← succ_sub_one p, succ_sub (Fact.out (p := p.Prime)).pos]
rw [← Nat.add_one_sub_one p, succ_sub (Fact.out (p := p.Prime)).pos]
rw [← prod_natCast, Finset.prod_Ico_id_eq_factorial, wilsons_lemma]
#align zmod.prod_Ico_one_prime ZMod.prod_Ico_one_prime

Expand Down
9 changes: 5 additions & 4 deletions Mathlib/Order/JordanHolder.lean
Original file line number Diff line number Diff line change
Expand Up @@ -287,7 +287,7 @@ theorem length_ofList (l : List X) (hl : l ≠ []) (hc : List.Chain' IsMaximal l
theorem ofList_toList (s : CompositionSeries X) :
ofList s.toList s.toList_ne_nil s.chain'_toList = s := by
refine' ext_fun _ _
· rw [length_ofList, length_toList, Nat.succ_sub_one]
· rw [length_ofList, length_toList, Nat.add_one_sub_one]
· rintro ⟨i, hi⟩
-- Porting note: Was `dsimp [ofList, toList]; rw [List.nthLe_ofFn']`.
simp [ofList, toList, -List.ofFn_succ]
Expand Down Expand Up @@ -411,7 +411,8 @@ theorem mem_eraseTop_of_ne_of_mem {s : CompositionSeries X} {x : X} (hx : x ≠
x ∈ s.eraseTop := by
rcases hxs with ⟨i, rfl⟩
have hi : (i : ℕ) < (s.length - 1).succ := by
conv_rhs => rw [← Nat.succ_sub (length_pos_of_mem_ne ⟨i, rfl⟩ s.top_mem hx), Nat.succ_sub_one]
conv_rhs => rw [← Nat.succ_sub (length_pos_of_mem_ne ⟨i, rfl⟩ s.top_mem hx),
Nat.add_one_sub_one]
exact lt_of_le_of_ne (Nat.le_of_lt_succ i.2) (by simpa [top, s.inj, Fin.ext_iff] using hx)
refine' ⟨Fin.castSucc (n := s.length + 1) i, _⟩
simp [Fin.ext_iff, Nat.mod_eq_of_lt hi]
Expand All @@ -424,7 +425,7 @@ theorem mem_eraseTop {s : CompositionSeries X} {x : X} (h : 0 < s.length) :
constructor
· rintro ⟨i, rfl⟩
have hi : (i : ℕ) < s.length := by
conv_rhs => rw [← Nat.succ_sub_one s.length, Nat.succ_sub h]
conv_rhs => rw [← Nat.add_one_sub_one s.length, Nat.succ_sub h]
exact i.2
-- Porting note: Was `simp [top, Fin.ext_iff, ne_of_lt hi]`.
simp [top, Fin.ext_iff, ne_of_lt hi, -Set.mem_range, Set.mem_range_self]
Expand All @@ -440,7 +441,7 @@ theorem lt_top_of_mem_eraseTop {s : CompositionSeries X} {x : X} (h : 0 < s.leng
theorem isMaximal_eraseTop_top {s : CompositionSeries X} (h : 0 < s.length) :
IsMaximal s.eraseTop.top s.top := by
have : s.length - 1 + 1 = s.length := by
conv_rhs => rw [← Nat.succ_sub_one s.length]; rw [Nat.succ_sub h]
conv_rhs => rw [← Nat.add_one_sub_one s.length]; rw [Nat.succ_sub h]
rw [top_eraseTop, top]
convert s.step ⟨s.length - 1, Nat.sub_lt h zero_lt_one⟩; ext; simp [this]
#align composition_series.is_maximal_erase_top_top CompositionSeries.isMaximal_eraseTop_top
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean
Original file line number Diff line number Diff line change
Expand Up @@ -95,7 +95,7 @@ theorem cyclotomic_prime_pow_comp_X_add_one_isEisensteinAt [hp : Fact p.Prime] (
· simp only [ite_eq_right_iff]
rintro ⟨k, hk⟩
rw [natDegree_comp, show (X + 1 : ℤ[X]) = X + C 1 by simp, natDegree_X_add_C, mul_one,
natDegree_cyclotomic, Nat.totient_prime_pow hp.out (Nat.succ_pos _), Nat.succ_sub_one]
natDegree_cyclotomic, Nat.totient_prime_pow hp.out (Nat.succ_pos _), Nat.add_one_sub_one]
at hn hi
rw [hk, pow_succ, mul_assoc] at hi
rw [hk, mul_comm, Nat.mul_div_cancel _ hp.out.pos]
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/RingTheory/Polynomial/Hermite/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -156,7 +156,7 @@ theorem coeff_hermite_explicit :
| n + 1, 0 => by
convert coeff_hermite_succ_zero (2 * n + 1) using 1
-- porting note: ring_nf did not solve the goal on line 165
rw [coeff_hermite_explicit n 1, (by rw [Nat.left_distrib, mul_one, Nat.succ_sub_one] :
rw [coeff_hermite_explicit n 1, (by rw [Nat.left_distrib, mul_one, Nat.add_one_sub_one] :
2 * (n + 1) - 1 = 2 * n + 1), Nat.doubleFactorial_add_one, Nat.choose_zero_right,
Nat.choose_one_right, pow_succ]
push_cast
Expand All @@ -178,7 +178,7 @@ theorem coeff_hermite_explicit :
-- Factor out double factorials.
norm_cast
-- porting note: ring_nf did not solve the goal on line 186
rw [(by rw [Nat.left_distrib, mul_one, Nat.succ_sub_one] : 2 * (n + 1) - 1 = 2 * n + 1),
rw [(by rw [Nat.left_distrib, mul_one, Nat.add_one_sub_one] : 2 * (n + 1) - 1 = 2 * n + 1),
Nat.doubleFactorial_add_one, mul_comm (2 * n + 1)]
simp only [mul_assoc, ← mul_add]
congr 1
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/RingTheory/Polynomial/Vieta.lean
Original file line number Diff line number Diff line change
Expand Up @@ -68,7 +68,7 @@ theorem prod_X_add_C_coeff (s : Multiset R) {k : ℕ} (h : k ≤ Multiset.card s
rw [hn, Nat.sub_sub_self (Nat.lt_succ_iff.mp (Finset.mem_range.mp hj1))] at hj2
exact Ne.irrefl hj2
· rw [Finset.mem_range]
exact Nat.sub_lt_succ (Multiset.card s) k
exact Nat.lt_succ_of_le (Nat.sub_le (Multiset.card s) k)
set_option linter.uppercaseLean3 false in
#align multiset.prod_X_add_C_coeff Multiset.prod_X_add_C_coeff

Expand Down

0 comments on commit eb40450

Please sign in to comment.