Skip to content

Commit 40b64f7

Browse files
kim-emleanprover-community-mathlib4-boteric-wiesercollares
committed
chore: bump to v4.3.0-rc2 (#8366)
# PR contents This is the supremum of - #8284 - #8056 - #8023 - #8332 - #8226 (already approved) - #7834 (already approved) along with some minor fixes from failures on nightly-testing as Mathlib `master` is merged into it. Note that some PRs for changes that are already compatible with the current toolchain and will be necessary have already been split out: #8380. I am hopeful that in future we will be able to progressively merge adaptation PRs into a `bump/v4.X.0` branch, so we never end up with a "big merge" like this. However one of these adaptation PRs (#8056) predates my new scheme for combined CI, and it wasn't possible to keep that PR viable in the meantime. # Lean PRs involved in this bump In particular this includes adjustments for the Lean PRs * leanprover/lean4#2778 * leanprover/lean4#2790 * leanprover/lean4#2783 * leanprover/lean4#2825 * leanprover/lean4#2722 ## leanprover/lean4#2778 We can get rid of all the ``` local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 ``` macros across Mathlib (and in any projects that want to write natural number powers of reals). ## leanprover/lean4#2722 Changes the default behaviour of `simp` to `(config := {decide := false})`. This makes `simp` (and consequentially `norm_num`) less powerful, but also more consistent, and less likely to blow up in long failures. This requires a variety of changes: changing some previously by `simp` or `norm_num` to `decide` or `rfl`, or adding `(config := {decide := true})`. ## leanprover/lean4#2783 This changed the behaviour of `simp` so that `simp [f]` will only unfold "fully applied" occurrences of `f`. The old behaviour can be recovered with `simp (config := { unfoldPartialApp := true })`. We may in future add a syntax for this, e.g. `simp [!f]`; please provide feedback! In the meantime, we have made the following changes: * switching to using explicit lemmas that have the intended level of application * `(config := { unfoldPartialApp := true })` in some places, to recover the old behaviour * Using `@[eqns]` to manually adjust the equation lemmas for a particular definition, recovering the old behaviour just for that definition. See #8371, where we do this for `Function.comp` and `Function.flip`. This change in Lean may require further changes down the line (e.g. adding the `!f` syntax, and/or upstreaming the special treatment for `Function.comp` and `Function.flip`, and/or removing this special treatment). Please keep an open and skeptical mind about these changes! Co-authored-by: leanprover-community-mathlib4-bot <[email protected]> Co-authored-by: Scott Morrison <[email protected]> Co-authored-by: Eric Wieser <[email protected]> Co-authored-by: Mauricio Collares <[email protected]>
1 parent a2ca43e commit 40b64f7

File tree

446 files changed

+1136
-1253
lines changed

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

446 files changed

+1136
-1253
lines changed

.gitignore

+2-2
Original file line numberDiff line numberDiff line change
@@ -2,10 +2,10 @@
22
.DS_Store
33
# Created by `lake exe cache get` if no home directory is available
44
/.cache
5-
# Prior to https://github.com/leanprover/lean4/pull/2749 lake stored files in these locations.
5+
# Prior to v4.3.0-rc2 lake stored files in these locations.
66
# We'll leave them in the `.gitignore` for a while for users switching between toolchains.
77
/build/
88
/lake-packages/
99
/lakefile.olean
10-
# After https://github.com/leanprover/lean4/pull/2749 lake stores its files here:
10+
# After v4.3.0-rc2 lake stores its files here:
1111
/.lake/

Archive/Examples/IfNormalization/Result.lean

+1-8
Original file line numberDiff line numberDiff line change
@@ -29,20 +29,13 @@ attribute [local simp] normalized hasNestedIf hasConstantIf hasRedundantIf disjo
2929
/-!
3030
Adding these lemmas to the simp set allows Lean to handle the termination proof automatically.
3131
-/
32-
attribute [local simp] Nat.lt_add_one_iff le_add_of_le_right
32+
attribute [local simp] Nat.lt_add_one_iff le_add_of_le_right max_add_add_right max_mul_mul_left
3333

3434
/-!
3535
Some further simp lemmas for handling if-then-else statements.
3636
-/
3737
attribute [local simp] apply_ite ite_eq_iff'
3838

39-
-- A copy of Lean's `decide_eq_true_eq` which unifies the `Decidable` instance
40-
-- rather than finding it by typeclass search.
41-
-- See https://github.com/leanprover/lean4/pull/2816
42-
@[simp] theorem decide_eq_true_eq {i : Decidable p} : (@decide p i = true) = p :=
43-
_root_.decide_eq_true_eq
44-
45-
4639
/-!
4740
Simp lemmas for `eval`.
4841
We don't want a `simp` lemma for `(ite i t e).eval` in general, only once we know the shape of `i`.

Archive/Examples/IfNormalization/WithoutAesop.lean

+2-8
Original file line numberDiff line numberDiff line change
@@ -22,12 +22,6 @@ namespace IfExpr
2222
attribute [local simp] eval normalized hasNestedIf hasConstantIf hasRedundantIf disjoint vars
2323
List.disjoint max_add_add_right max_mul_mul_left Nat.lt_add_one_iff le_add_of_le_right
2424

25-
-- A copy of Lean's `decide_eq_true_eq` which unifies the `Decidable` instance
26-
-- rather than finding it by typeclass search.
27-
-- See https://github.com/leanprover/lean4/pull/2816
28-
@[simp] theorem decide_eq_true_eq' {i : Decidable p} : (@decide p i = true) = p :=
29-
_root_.decide_eq_true_eq
30-
3125
theorem eval_ite_ite' :
3226
(ite (ite a b c) d e).eval f = (ite a (ite b d e) (ite c d e)).eval f := by
3327
cases h : eval f a <;> simp_all
@@ -106,8 +100,8 @@ def normalize' (l : AList (fun _ : ℕ => Bool)) :
106100
simp_all
107101
· simp_all? says simp_all only [ne_eq, hasNestedIf, Bool.or_self, hasConstantIf,
108102
and_self, hasRedundantIf, Bool.or_false, beq_eq_false_iff_ne, not_false_eq_true,
109-
disjoint, List.disjoint, Bool.and_true, Bool.and_eq_true, decide_eq_true_eq',
110-
true_and]
103+
disjoint, List.disjoint, decide_not, Bool.and_true, Bool.and_eq_true,
104+
Bool.not_eq_true', decide_eq_false_iff_not, true_and]
111105
constructor <;> assumption
112106
· have := ht₃ w
113107
have := he₃ w

Archive/Imo/Imo1960Q1.lean

+6-1
Original file line numberDiff line numberDiff line change
@@ -95,7 +95,12 @@ theorem right_direction {n : ℕ} : ProblemPredicate n → SolutionPredicate n :
9595
have := searchUpTo_start
9696
iterate 82
9797
replace :=
98-
searchUpTo_step this (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num)
98+
searchUpTo_step this (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num1; rfl)
99+
(by -- This used to be just `norm_num`, but after leanprover/lean4#2790,
100+
-- that triggers a max recursion depth exception. As a workaround, we
101+
-- manually rewrite with `Nat.digits_of_two_le_of_pos` first.
102+
rw [Nat.digits_of_two_le_of_pos (by norm_num) (by norm_num)]
103+
norm_num <;> decide)
99104
exact searchUpTo_end this
100105
#align imo1960_q1.right_direction Imo1960Q1.right_direction
101106

Archive/Imo/Imo1962Q1.lean

+12-1
Original file line numberDiff line numberDiff line change
@@ -131,7 +131,18 @@ Now we combine these cases to show that 153846 is the smallest solution.
131131

132132

133133
theorem satisfied_by_153846 : ProblemPredicate 153846 := by
134-
norm_num [ProblemPredicate]
134+
-- This proof used to be the single line `norm_num [ProblemPredicate]`.
135+
-- After leanprover/lean4#2790, that triggers a max recursion depth exception.
136+
-- As a workaround, we manually apply `Nat.digits_of_two_le_of_pos` a few times
137+
-- before invoking `norm_num`.
138+
unfold ProblemPredicate
139+
have two_le_ten : 210 := by norm_num
140+
rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)]
141+
rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)]
142+
rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)]
143+
rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)]
144+
norm_num
145+
decide
135146
#align imo1962_q1.satisfied_by_153846 Imo1962Q1.satisfied_by_153846
136147

137148
theorem no_smaller_solutions (n : ℕ) (h1 : ProblemPredicate n) : n ≥ 153846 := by

Archive/Imo/Imo1962Q4.lean

-2
Original file line numberDiff line numberDiff line change
@@ -17,8 +17,6 @@ Since Lean does not have a concept of "simplest form", we just express what is
1717
in fact the simplest form of the set of solutions, and then prove it equals the set of solutions.
1818
-/
1919

20-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
21-
2220
open Real
2321

2422
open scoped Real

Archive/Imo/Imo1964Q1.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,7 @@ theorem two_pow_mod_seven (n : ℕ) : 2 ^ n ≡ 2 ^ (n % 3) [MOD 7] :=
3434
let t := n % 3
3535
calc 2 ^ n = 2 ^ (3 * (n / 3) + t) := by rw [Nat.div_add_mod]
3636
_ = (2 ^ 3) ^ (n / 3) * 2 ^ t := by rw [pow_add, pow_mul]
37-
_ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; norm_num
37+
_ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; decide
3838
_ = 2 ^ t := by ring
3939

4040
/-!
@@ -68,5 +68,5 @@ theorem imo1964_q1b (n : ℕ) : ¬7 ∣ 2 ^ n + 1 := by
6868
have H : 2 ^ t + 10 [MOD 7]
6969
· calc 2 ^ t + 12 ^ n + 1 [MOD 7 ] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm
7070
_ ≡ 0 [MOD 7] := h.modEq_zero_nat
71-
interval_cases t <;> norm_num at H
71+
interval_cases t <;> contradiction
7272
#align imo1964_q1b imo1964_q1b

Archive/Imo/Imo1981Q3.lean

+4-2
Original file line numberDiff line numberDiff line change
@@ -129,7 +129,7 @@ theorem imp_fib {n : ℕ} : ∀ m : ℕ, NatPredicate N m n → ∃ k : ℕ, m =
129129
obtain (rfl : 1 = n) | (h4 : 1 < n) := (succ_le_iff.mpr h2.n_pos).eq_or_lt
130130
· use 1
131131
have h5 : 1 ≤ m := succ_le_iff.mpr h2.m_pos
132-
simpa [fib_one, fib_two] using (h3.antisymm h5 : m = 1)
132+
simpa [fib_one, fib_two, (by decide : 1 + 1 = 2)] using (h3.antisymm h5 : m = 1)
133133
· obtain (rfl : m = n) | (h6 : m < n) := h3.eq_or_lt
134134
· exact absurd h2.eq_imp_1 (Nat.ne_of_gt h4)
135135
· have h7 : NatPredicate N (n - m) m := h2.reduction h4
@@ -205,5 +205,7 @@ theorem imo1981_q3 : IsGreatest (specifiedSet 1981) 3524578 := by
205205
have := fun h => @solution_greatest 1981 16 h 3524578
206206
norm_num at this
207207
apply this
208-
norm_num [ProblemPredicate_iff]
208+
· decide
209+
· decide
210+
· norm_num [ProblemPredicate_iff]; decide
209211
#align imo1981_q3 imo1981_q3

Archive/Imo/Imo2001Q2.lean

-2
Original file line numberDiff line numberDiff line change
@@ -32,8 +32,6 @@ open Real
3232

3333
variable {a b c : ℝ}
3434

35-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
36-
3735
namespace Imo2001Q2
3836

3937
theorem bound (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) :

Archive/Imo/Imo2005Q4.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ def a (n : ℕ) : ℤ :=
2929
theorem find_specified_factor {p : ℕ} (hp : Nat.Prime p) (hp2 : p ≠ 2) (hp3 : p ≠ 3) :
3030
↑p ∣ a (p - 2) := by
3131
-- Since `p` is neither `2` nor `3`, it is coprime with `2`, `3`, and `6`
32-
rw [Ne.def, ← Nat.prime_dvd_prime_iff_eq hp (by norm_num), ← Nat.Prime.coprime_iff_not_dvd hp]
32+
rw [Ne.def, ← Nat.prime_dvd_prime_iff_eq hp (by decide), ← Nat.Prime.coprime_iff_not_dvd hp]
3333
at hp2 hp3
3434
have : Int.gcd p 6 = 1 := Nat.coprime_mul_iff_right.2 ⟨hp2, hp3⟩
3535
-- Nat arithmetic needed to deal with powers
@@ -71,10 +71,10 @@ theorem imo2005_q4 {k : ℕ} (hk : 0 < k) : (∀ n : ℕ, 1 ≤ n → IsCoprime
7171
-- For `p = 2` and `p = 3`, take `n = 1` and `n = 2`, respectively
7272
by_cases hp2 : p = 2
7373
· rw [hp2] at h
74-
apply h 1 <;> norm_num
74+
apply h 1 <;> decide
7575
by_cases hp3 : p = 3
7676
· rw [hp3] at h
77-
apply h 2 <;> norm_num
77+
apply h 2 <;> decide
7878
-- Otherwise, take `n = p - 2`
7979
refine h (p - 2) ?_ (find_specified_factor hp hp2 hp3)
8080
calc

Archive/Imo/Imo2008Q3.lean

-2
Original file line numberDiff line numberDiff line change
@@ -31,8 +31,6 @@ Then `p = 2n + k ≥ 2n + √(p - 4) = 2n + √(2n + k - 4) > √(2n)` and we ar
3131

3232
open Real
3333

34-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
35-
3634
namespace Imo2008Q3
3735

3836
theorem p_lemma (p : ℕ) (hpp : Nat.Prime p) (hp_mod_4_eq_1 : p ≡ 1 [MOD 4]) (hp_gt_20 : p > 20) :

Archive/Imo/Imo2013Q5.lean

-2
Original file line numberDiff line numberDiff line change
@@ -31,8 +31,6 @@ https://www.imo-official.org/problems/IMO2013SL.pdf
3131

3232
open scoped BigOperators
3333

34-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
35-
3634
namespace Imo2013Q5
3735

3836
theorem le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y)

Archive/Imo/Imo2019Q1.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -27,7 +27,7 @@ theorem imo2019_q1 (f : ℤ → ℤ) :
2727
(∀ a b : ℤ, f (2 * a) + 2 * f b = f (f (a + b))) ↔ f = 0 ∨ ∃ c, f = fun x => 2 * x + c := by
2828
constructor; swap
2929
-- easy way: f(x)=0 and f(x)=2x+c work.
30-
· rintro (rfl | ⟨c, rfl⟩) <;> intros <;> simp only [Pi.zero_apply]; ring
30+
· rintro (rfl | ⟨c, rfl⟩) <;> intros <;> norm_num; ring
3131
-- hard way.
3232
intro hf
3333
-- functional equation

Archive/Imo/Imo2019Q2.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -442,7 +442,7 @@ theorem not_collinear_QPA₂ : ¬Collinear ℝ ({cfg.Q, cfg.P, cfg.A₂} : Set P
442442
have h : Cospherical ({cfg.B, cfg.A, cfg.A₂} : Set Pt) := by
443443
refine' cfg.triangleABC.circumsphere.cospherical.subset _
444444
simp only [Set.insert_subset_iff, cfg.A_mem_circumsphere, cfg.B_mem_circumsphere,
445-
cfg.A₂_mem_circumsphere, Sphere.mem_coe, Set.singleton_subset_iff]
445+
cfg.A₂_mem_circumsphere, Sphere.mem_coe, Set.singleton_subset_iff, and_true]
446446
exact h.affineIndependent_of_ne cfg.A_ne_B.symm cfg.A₂_ne_B.symm cfg.A₂_ne_A.symm
447447
#align imo2019_q2.imo2019q2_cfg.not_collinear_QPA₂ Imo2019Q2.Imo2019q2Cfg.not_collinear_QPA₂
448448

Archive/Imo/Imo2019Q4.lean

+5-8
Original file line numberDiff line numberDiff line change
@@ -29,8 +29,6 @@ individually.
2929

3030
open scoped Nat BigOperators
3131

32-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
33-
3432
open Nat hiding zero_le Prime
3533

3634
open Finset multiplicity
@@ -65,7 +63,7 @@ theorem upper_bound {k n : ℕ} (hk : k > 0)
6563
_ ≤ k ! := by gcongr
6664
clear h h2
6765
induction' n, hn using Nat.le_induction with n' hn' IH
68-
· norm_num
66+
· decide
6967
let A := ∑ i in range n', i
7068
have le_sum : ∑ i in range 6, i ≤ A
7169
· apply sum_le_sum_of_subset
@@ -87,8 +85,7 @@ theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) :
8785
-- The implication `←` holds.
8886
constructor
8987
swap
90-
· rintro (h | h) <;> simp [Prod.ext_iff] at h <;> rcases h with ⟨rfl, rfl⟩ <;>
91-
norm_num [prod_range_succ, succ_mul]
88+
· rintro (h | h) <;> simp [Prod.ext_iff] at h <;> rcases h with ⟨rfl, rfl⟩ <;> decide
9289
intro h
9390
-- We know that n < 6.
9491
have := Imo2019Q4.upper_bound hk h
@@ -101,9 +98,9 @@ theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) :
10198
exact h; rw [h]; norm_num
10299
all_goals exfalso; norm_num [prod_range_succ] at h; norm_cast at h
103100
-- n = 3
104-
· refine' monotone_factorial.ne_of_lt_of_lt_nat 5 _ _ _ h <;> norm_num
101+
· refine' monotone_factorial.ne_of_lt_of_lt_nat 5 _ _ _ h <;> decide
105102
-- n = 4
106-
· refine' monotone_factorial.ne_of_lt_of_lt_nat 7 _ _ _ h <;> norm_num
103+
· refine' monotone_factorial.ne_of_lt_of_lt_nat 7 _ _ _ h <;> decide
107104
-- n = 5
108-
· refine' monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h <;> norm_num
105+
· refine' monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h <;> decide
109106
#align imo2019_q4 imo2019_q4

Archive/MiuLanguage/DecisionNec.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -61,7 +61,7 @@ theorem mod3_eq_1_or_mod3_eq_2 {a b : ℕ} (h1 : a % 3 = 1 ∨ a % 3 = 2)
6161
· rw [h2]; exact h1
6262
· cases' h1 with h1 h1
6363
· right; simp [h2, mul_mod, h1, Nat.succ_lt_succ]
64-
· left; simp [h2, mul_mod, h1]
64+
· left; simp only [h2, mul_mod, h1, mod_mod]; decide
6565
#align miu.mod3_eq_1_or_mod3_eq_2 Miu.mod3_eq_1_or_mod3_eq_2
6666

6767
/-- `count_equiv_one_or_two_mod3_of_derivable` shows any derivable string must have a `count I` that
@@ -80,7 +80,7 @@ theorem count_equiv_one_or_two_mod3_of_derivable (en : Miustr) :
8080
simp_rw [count_cons_self, count_nil, count_cons, ite_false, add_right_comm, add_mod_right]
8181
simp
8282
· left; rw [count_append, count_append, count_append]
83-
simp only [ne_eq, count_cons_of_ne, count_nil, add_zero]
83+
simp only [ne_eq, not_false_eq_true, count_cons_of_ne, count_nil, add_zero]
8484
#align miu.count_equiv_one_or_two_mod3_of_derivable Miu.count_equiv_one_or_two_mod3_of_derivable
8585

8686
/-- Using the above theorem, we solve the MU puzzle, showing that `"MU"` is not derivable.
@@ -134,7 +134,7 @@ theorem goodm_of_rule1 (xs : Miustr) (h₁ : Derivable (xs ++ ↑[I])) (h₂ : G
134134
exact mhead
135135
· change ¬M ∈ tail (xs ++ ↑([I] ++ [U]))
136136
rw [← append_assoc, tail_append_singleton_of_ne_nil]
137-
· simp_rw [mem_append, not_or, and_true]; exact nmtail
137+
· simp_rw [mem_append, mem_singleton, or_false]; exact nmtail
138138
· exact append_ne_nil_of_ne_nil_left _ _ this
139139
#align miu.goodm_of_rule1 Miu.goodm_of_rule1
140140

Archive/MiuLanguage/DecisionSuf.lean

+4-3
Original file line numberDiff line numberDiff line change
@@ -267,7 +267,7 @@ theorem count_I_eq_length_of_count_U_zero_and_neg_mem {ys : Miustr} (hu : count
267267
· simpa only [count]
268268
· rw [mem_cons, not_or] at hm; exact hm.2
269269
· -- case `x = U` gives a contradiction.
270-
exfalso; simp only [count, countP_cons_of_pos] at hu
270+
exfalso; simp only [count, countP_cons_of_pos (· == U) _ (rfl : U == U)] at hu
271271
exact succ_ne_zero _ hu
272272
set_option linter.uppercaseLean3 false in
273273
#align miu.count_I_eq_length_of_count_U_zero_and_neg_mem Miu.count_I_eq_length_of_count_U_zero_and_neg_mem
@@ -277,7 +277,8 @@ set_option linter.uppercaseLean3 false in
277277
theorem base_case_suf (en : Miustr) (h : Decstr en) (hu : count U en = 0) : Derivable en := by
278278
rcases h with ⟨⟨mhead, nmtail⟩, hi⟩
279279
have : en ≠ nil := by
280-
intro k; simp only [k, count, countP, if_false, zero_mod, zero_ne_one, false_or_iff] at hi
280+
intro k
281+
simp only [k, count, countP, countP.go, if_false, zero_mod, zero_ne_one, false_or_iff] at hi
281282
rcases exists_cons_of_ne_nil this with ⟨y, ys, rfl⟩
282283
rcases mhead
283284
rsuffices ⟨c, rfl, hc⟩ : ∃ c, replicate c I = ys ∧ (c % 3 = 1 ∨ c % 3 = 2)
@@ -331,7 +332,7 @@ theorem ind_hyp_suf (k : ℕ) (ys : Miustr) (hu : count U ys = succ k) (hdec : D
331332
rw [cons_append, cons_append]
332333
dsimp [tail] at nmtail ⊢
333334
rw [mem_append] at nmtail
334-
simpa only [mem_append, mem_cons, false_or_iff, or_false_iff] using nmtail
335+
simpa only [append_assoc, cons_append, nil_append, mem_append, mem_cons, false_or] using nmtail
335336
· rw [count_append, count_append]; rw [← cons_append, count_append] at hic
336337
simp only [count_cons_self, count_nil, count_cons, if_false] at hic ⊢
337338
rw [add_right_comm, add_mod_right]; exact hic

Archive/Sensitivity.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -401,7 +401,7 @@ theorem exists_eigenvalue (H : Set (Q m.succ)) (hH : Card H ≥ 2 ^ m + 1) :
401401
let img := range (g m)
402402
suffices 0 < dim (W ⊓ img) by
403403
exact_mod_cast exists_mem_ne_zero_of_rank_pos this
404-
have dim_le : dim (W ⊔ img) ≤ 2 ^ (m + 1) := by
404+
have dim_le : dim (W ⊔ img) ≤ 2 ^ (m + 1 : Cardinal) := by
405405
convert ← rank_submodule_le (W ⊔ img)
406406
rw [← Nat.cast_succ]
407407
apply dim_V

Archive/Wiedijk100Theorems/AbelRuffini.lean

+7-8
Original file line numberDiff line numberDiff line change
@@ -31,8 +31,6 @@ namespace AbelRuffini
3131

3232
set_option linter.uppercaseLean3 false
3333

34-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
35-
3634
open Function Polynomial Polynomial.Gal Ideal
3735

3836
open scoped Polynomial
@@ -94,10 +92,11 @@ theorem irreducible_Phi (p : ℕ) (hp : p.Prime) (hpa : p ∣ a) (hpb : p ∣ b)
9492
rw [mem_span_singleton]
9593
rw [degree_Phi] at hn; norm_cast at hn
9694
interval_cases hn : n <;>
97-
simp only [Φ, coeff_X_pow, coeff_C, Int.coe_nat_dvd.mpr, hpb, if_true, coeff_C_mul, if_false,
98-
coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, coeff_sub, add_zero, zero_sub, dvd_neg,
99-
neg_zero, dvd_mul_of_dvd_left]
95+
simp (config := {decide := true}) only [Φ, coeff_X_pow, coeff_C, Int.coe_nat_dvd.mpr, hpb,
96+
if_true, coeff_C_mul, if_false, coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, coeff_sub,
97+
add_zero, zero_sub, dvd_neg, neg_zero, dvd_mul_of_dvd_left]
10098
· simp only [degree_Phi, ← WithBot.coe_zero, WithBot.coe_lt_coe, Nat.succ_pos']
99+
decide
101100
· rw [coeff_zero_Phi, span_singleton_pow, mem_span_singleton]
102101
exact mt Int.coe_nat_dvd.mp hp2b
103102
all_goals exact Monic.isPrimitive (monic_Phi a b)
@@ -136,7 +135,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) :
136135
have hfa :=
137136
calc
138137
f (-a) = (a : ℝ) ^ 2 - (a : ℝ) ^ 5 + b := by
139-
norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow]
138+
norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow (by decide : Odd 5)]
140139
_ ≤ (a : ℝ) ^ 2 - (a : ℝ) ^ 3 + (a - 1) := by
141140
refine' add_le_add (sub_le_sub_left (pow_le_pow ha _) _) _ <;> linarith
142141
_ = -((a : ℝ) - 1) ^ 2 * (a + 1) := by ring
@@ -163,7 +162,7 @@ theorem complex_roots_Phi (h : (Φ ℚ a b).Separable) : Fintype.card ((Φ ℚ a
163162
theorem gal_Phi (hab : b < a) (h_irred : Irreducible (Φ ℚ a b)) :
164163
Bijective (galActionHom (Φ ℚ a b) ℂ) := by
165164
apply galActionHom_bijective_of_prime_degree' h_irred
166-
· norm_num [natDegree_Phi]
165+
· simp only [natDegree_Phi]; decide
167166
· rw [complex_roots_Phi a b h_irred.separable, Nat.succ_le_succ_iff]
168167
exact (real_roots_Phi_le a b).trans (Nat.le_succ 3)
169168
· simp_rw [complex_roots_Phi a b h_irred.separable, Nat.succ_le_succ_iff]
@@ -181,7 +180,7 @@ theorem not_solvable_by_rad (p : ℕ) (x : ℂ) (hx : aeval x (Φ ℚ a b) = 0)
181180
#align abel_ruffini.not_solvable_by_rad AbelRuffini.not_solvable_by_rad
182181

183182
theorem not_solvable_by_rad' (x : ℂ) (hx : aeval x (Φ ℚ 4 2) = 0) : ¬IsSolvableByRad ℚ x := by
184-
apply not_solvable_by_rad 4 2 2 x hx <;> norm_num
183+
apply not_solvable_by_rad 4 2 2 x hx <;> decide
185184
#align abel_ruffini.not_solvable_by_rad' AbelRuffini.not_solvable_by_rad'
186185

187186
/-- **Abel-Ruffini Theorem** -/

Archive/Wiedijk100Theorems/AreaOfACircle.lean

-3
Original file line numberDiff line numberDiff line change
@@ -43,9 +43,6 @@ to the n-ball.
4343
-/
4444

4545

46-
47-
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220
48-
4946
open Set Real MeasureTheory intervalIntegral
5047

5148
open scoped Real NNReal

Archive/Wiedijk100Theorems/BallotProblem.lean

+3-1
Original file line numberDiff line numberDiff line change
@@ -221,7 +221,9 @@ theorem first_vote_pos :
221221
simp [ENNReal.div_self _ _]
222222
| 0, q + 1, _ => by
223223
rw [counted_left_zero, condCount_singleton]
224-
simp
224+
simp only [List.replicate, Nat.add_eq, add_zero, mem_setOf_eq, List.headI_cons, Nat.cast_zero,
225+
ENNReal.zero_div, ite_eq_right_iff]
226+
decide
225227
| p + 1, q + 1, _ => by
226228
simp_rw [counted_succ_succ]
227229
rw [← condCount_disjoint_union ((countedSequence_finite _ _).image _)

0 commit comments

Comments
 (0)