From 134f441deb38c73a700d7ce6f604d7fdef2ce767 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 18:52:29 +1100 Subject: [PATCH 01/48] chore: adaptations for nightly-2024-03-11 --- Mathlib/Algebra/BigOperators/Fin.lean | 20 +- Mathlib/Algebra/Category/ModuleCat/Basic.lean | 4 +- .../Algebra/Category/Ring/Constructions.lean | 2 +- Mathlib/Algebra/CharZero/Defs.lean | 1 - Mathlib/Algebra/Function/Support.lean | 3 +- Mathlib/Algebra/GeomSum.lean | 2 +- Mathlib/Algebra/Group/NatPowAssoc.lean | 4 +- Mathlib/Algebra/Homology/Augment.lean | 2 +- Mathlib/Algebra/Homology/ExactSequence.lean | 3 + .../Algebra/Homology/HomologySequence.lean | 6 + Mathlib/Algebra/Homology/Single.lean | 7 +- Mathlib/Algebra/Lie/Solvable.lean | 2 +- Mathlib/Algebra/Module/Basic.lean | 4 +- Mathlib/Algebra/Module/PID.lean | 3 +- Mathlib/Algebra/Order/Field/Power.lean | 2 +- Mathlib/Algebra/Order/Floor.lean | 2 +- Mathlib/Algebra/TrivSqZeroExt.lean | 2 +- .../DoldKan/HomotopyEquivalence.lean | 1 - Mathlib/Analysis/BoundedVariation.lean | 24 +- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 2 +- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 2 +- Mathlib/Analysis/Convex/Basic.lean | 8 +- .../Convex/SpecificFunctions/Deriv.lean | 2 +- .../InnerProductSpace/Orientation.lean | 4 +- .../InnerProductSpace/Projection.lean | 3 +- Mathlib/Analysis/Normed/Group/Basic.lean | 6 +- Mathlib/Analysis/NormedSpace/Exponential.lean | 2 +- Mathlib/Analysis/ODE/PicardLindelof.lean | 2 +- .../SpecialFunctions/Gamma/Basic.lean | 8 +- .../Trigonometric/EulerSineProd.lean | 2 +- .../Abelian/DiagramLemmas/Four.lean | 3 + .../Abelian/InjectiveResolution.lean | 6 +- Mathlib/CategoryTheory/ComposableArrows.lean | 18 +- .../Idempotents/Biproducts.lean | 4 +- Mathlib/CategoryTheory/Preadditive/Mat.lean | 2 +- Mathlib/Combinatorics/Additive/Energy.lean | 4 +- Mathlib/Combinatorics/Catalan.lean | 11 +- Mathlib/Combinatorics/Composition.lean | 2 +- .../Combinatorics/Derangements/Finite.lean | 2 +- Mathlib/Combinatorics/Hindman.lean | 2 +- Mathlib/Combinatorics/SetFamily/LYM.lean | 2 - .../Combinatorics/SimpleGraph/AdjMatrix.lean | 3 +- Mathlib/Computability/Ackermann.lean | 4 +- Mathlib/Computability/TMToPartrec.lean | 6 +- Mathlib/Data/Fin/Basic.lean | 2 +- Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Fintype/Card.lean | 3 +- Mathlib/Data/Int/Cast/Basic.lean | 4 +- Mathlib/Data/List/Basic.lean | 12 +- Mathlib/Data/List/Cycle.lean | 7 +- Mathlib/Data/List/Indexes.lean | 2 +- Mathlib/Data/List/Perm.lean | 10 +- Mathlib/Data/List/Rotate.lean | 7 +- Mathlib/Data/List/Sort.lean | 44 +-- Mathlib/Data/List/Sublists.lean | 2 +- Mathlib/Data/List/Zip.lean | 3 +- Mathlib/Data/Nat/Cast/Basic.lean | 4 +- Mathlib/Data/Nat/Cast/Defs.lean | 20 +- Mathlib/Data/Nat/Choose/Basic.lean | 3 +- Mathlib/Data/Nat/Choose/Multinomial.lean | 1 - Mathlib/Data/Nat/Defs.lean | 3 + Mathlib/Data/Nat/Digits.lean | 16 +- Mathlib/Data/Nat/Factorial/Basic.lean | 3 +- .../Data/Nat/Factorial/SuperFactorial.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 16 +- Mathlib/Data/Nat/Fib/Basic.lean | 9 +- Mathlib/Data/Nat/Hyperoperation.lean | 2 +- Mathlib/Data/Nat/Log.lean | 6 +- Mathlib/Data/Nat/Multiplicity.lean | 12 +- Mathlib/Data/Nat/Order/Basic.lean | 8 +- Mathlib/Data/Nat/Order/Lemmas.lean | 3 +- Mathlib/Data/Nat/Prime.lean | 3 +- Mathlib/Data/Num/Lemmas.lean | 7 +- Mathlib/Data/Opposite.lean | 4 +- Mathlib/Data/Ordmap/Ordset.lean | 2 +- Mathlib/Data/PNat/Xgcd.lean | 3 +- Mathlib/Data/Polynomial/Basic.lean | 2 +- Mathlib/Data/Polynomial/Derivative.lean | 72 ++--- Mathlib/Data/Polynomial/Module/Basic.lean | 4 +- Mathlib/Data/Polynomial/Monic.lean | 2 +- Mathlib/Data/Polynomial/RingDivision.lean | 5 +- Mathlib/Data/Polynomial/Smeval.lean | 4 +- .../QPF/Multivariate/Constructions/Cofix.lean | 2 +- Mathlib/Data/Seq/Computation.lean | 2 +- Mathlib/Data/Seq/Seq.lean | 2 +- Mathlib/Data/Set/Enumerate.lean | 3 +- Mathlib/Data/Set/Pairwise/Basic.lean | 4 +- Mathlib/Data/Set/Prod.lean | 2 +- Mathlib/Data/Sum/Interval.lean | 3 +- Mathlib/Data/ZMod/Basic.lean | 2 +- Mathlib/Data/ZMod/Factorial.lean | 3 +- Mathlib/FieldTheory/ChevalleyWarning.lean | 4 +- Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean | 3 +- Mathlib/Geometry/Euclidean/MongePoint.lean | 2 +- .../Manifold/VectorBundle/SmoothSection.lean | 2 +- Mathlib/GroupTheory/FreeGroup/Basic.lean | 4 - Mathlib/GroupTheory/Nilpotent.lean | 4 +- Mathlib/GroupTheory/Perm/Fin.lean | 2 +- Mathlib/GroupTheory/Perm/List.lean | 16 +- Mathlib/GroupTheory/PushoutI.lean | 1 - Mathlib/GroupTheory/Subsemigroup/Center.lean | 6 +- Mathlib/Init/Data/Nat/Basic.lean | 2 + Mathlib/Lean/Expr/Basic.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Adjugate.lean | 2 +- Mathlib/LinearAlgebra/Matrix/ZPow.lean | 3 +- .../LinearAlgebra/QuadraticForm/Basic.lean | 3 +- .../TensorProduct/Graded/Internal.lean | 6 +- Mathlib/Logic/Denumerable.lean | 2 +- Mathlib/Logic/Function/Iterate.lean | 2 +- Mathlib/Logic/Small/Defs.lean | 2 +- Mathlib/Mathport/Notation.lean | 2 +- .../Constructions/BorelSpace/Basic.lean | 2 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 2 +- .../Decomposition/SignedHahn.lean | 2 +- .../MeasureTheory/Function/SimpleFunc.lean | 2 +- .../Measure/Lebesgue/VolumeOfBalls.lean | 2 +- Mathlib/MeasureTheory/OuterMeasure/Basic.lean | 2 +- Mathlib/NumberTheory/ArithmeticFunction.lean | 2 +- Mathlib/NumberTheory/Bernoulli.lean | 9 +- .../NumberTheory/BernoulliPolynomials.lean | 1 - .../NumberTheory/Cyclotomic/Discriminant.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 2 +- Mathlib/NumberTheory/Harmonic/Defs.lean | 6 +- Mathlib/NumberTheory/Multiplicity.lean | 4 +- .../NumberTheory/NumberField/Embeddings.lean | 2 +- Mathlib/NumberTheory/Padics/PadicVal.lean | 9 +- Mathlib/NumberTheory/Padics/RingHoms.lean | 2 +- Mathlib/NumberTheory/ZetaFunction.lean | 2 +- .../Order/SuccPred/LinearLocallyFinite.lean | 2 +- .../Probability/Martingale/Upcrossing.lean | 2 +- Mathlib/RingTheory/Derivation/Basic.lean | 2 +- Mathlib/RingTheory/Filtration.lean | 10 +- Mathlib/RingTheory/Henselian.lean | 14 +- Mathlib/RingTheory/Multiplicity.lean | 2 +- .../MvPolynomial/NewtonIdentities.lean | 3 +- Mathlib/RingTheory/MvPowerSeries/Basic.lean | 2 +- Mathlib/RingTheory/Nilpotent.lean | 2 +- Mathlib/RingTheory/Polynomial/Bernstein.lean | 9 +- Mathlib/RingTheory/Polynomial/Content.lean | 1 - .../Polynomial/Cyclotomic/Basic.lean | 6 +- .../Polynomial/Eisenstein/IsIntegral.lean | 267 +++++++++--------- .../RingTheory/Polynomial/Hermite/Basic.lean | 4 +- Mathlib/RingTheory/Polynomial/Pochhammer.lean | 13 +- Mathlib/RingTheory/QuotientNilpotent.lean | 2 +- .../RingTheory/UniqueFactorizationDomain.lean | 4 +- Mathlib/RingTheory/WittVector/Identities.lean | 2 +- Mathlib/RingTheory/WittVector/MulP.lean | 3 +- .../WittVector/StructurePolynomial.lean | 2 +- .../RingTheory/WittVector/Verschiebung.lean | 5 +- Mathlib/SetTheory/Ordinal/Basic.lean | 4 +- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 6 +- Mathlib/SetTheory/Ordinal/Notation.lean | 116 ++++---- Mathlib/SetTheory/ZFC/Basic.lean | 2 +- Mathlib/Tactic/Find.lean | 2 + Mathlib/Tactic/NormNum/Result.lean | 2 +- Mathlib/Tactic/Observe.lean | 4 +- Mathlib/Tactic/Positivity/Basic.lean | 2 +- Mathlib/Tactic/Sat/FromLRAT.lean | 2 +- .../Topology/Category/Profinite/Nobeling.lean | 1 - Mathlib/Topology/Gluing.lean | 6 + Mathlib/Topology/Instances/ENNReal.lean | 2 +- .../Topology/MetricSpace/GromovHausdorff.lean | 2 +- Mathlib/Topology/Support.lean | 10 +- Mathlib/Topology/UniformSpace/Matrix.lean | 4 +- Mathlib/Util/CompileInductive.lean | 1 - lean-toolchain | 2 +- 166 files changed, 618 insertions(+), 571 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Fin.lean b/Mathlib/Algebra/BigOperators/Fin.lean index e5b34bb7c3865..294d91dd0c93c 100644 --- a/Mathlib/Algebra/BigOperators/Fin.lean +++ b/Mathlib/Algebra/BigOperators/Fin.lean @@ -267,11 +267,10 @@ theorem partialProd_right_inv {G : Type*} [Group G] (f : Fin n → G) (i : Fin n specialize hi (lt_trans (Nat.lt_succ_self i) hn) simp only [Fin.coe_eq_castSucc, Fin.succ_mk, Fin.castSucc_mk] at hi ⊢ rw [← Fin.succ_mk _ _ (lt_trans (Nat.lt_succ_self _) hn), ← Fin.succ_mk] - rw [Nat.succ_eq_add_one] at hn simp only [partialProd_succ, mul_inv_rev, Fin.castSucc_mk] -- Porting note: was -- assoc_rw [hi, inv_mul_cancel_left] - rw [← mul_assoc, mul_left_eq_self, mul_assoc, hi, mul_left_inv] + rw [← mul_assoc, mul_left_eq_self, mul_assoc, castSucc_mk, hi, mul_left_inv] #align fin.partial_prod_right_inv Fin.partialProd_right_inv #align fin.partial_sum_right_neg Fin.partialSum_right_neg @@ -322,19 +321,15 @@ def finFunctionFinEquiv {m n : ℕ} : (Fin n → Fin m) ≃ Fin (m ^ n) := induction' n with n ih · simp cases m - · dsimp only [Nat.zero_eq] at f -- Porting note: added, wrong zero - exact isEmptyElim (f <| Fin.last _) + · exact isEmptyElim (f <| Fin.last _) simp_rw [Fin.sum_univ_castSucc, Fin.coe_castSucc, Fin.val_last] refine' (add_lt_add_of_lt_of_le (ih _) <| mul_le_mul_right' (Fin.is_le _) _).trans_eq _ - rw [← one_add_mul (_ : ℕ), add_comm, pow_succ] - -- Porting note: added, wrong `succ` - rfl⟩) + rw [← one_add_mul (_ : ℕ), add_comm, pow_succ]⟩) (fun a b => ⟨a / m ^ (b : ℕ) % m, by cases' n with n · exact b.elim0 cases' m with m - · dsimp only [Nat.zero_eq] at a -- Porting note: added, wrong zero - rw [zero_pow n.succ_ne_zero] at a + · rw [zero_pow n.succ_ne_zero] at a exact a.elim0 · exact Nat.mod_lt _ m.succ_pos⟩) fun a => by @@ -382,12 +377,9 @@ def finPiFinEquiv {m : ℕ} {n : Fin m → ℕ} : (∀ i : Fin m, Fin (n i)) ≃ exact this intro n nn f fn cases nn - · dsimp only [Nat.zero_eq] at fn -- Porting note: added, wrong zero - exact isEmptyElim fn + · exact isEmptyElim fn refine' (add_lt_add_of_lt_of_le (ih _) <| mul_le_mul_right' (Fin.is_le _) _).trans_eq _ - rw [← one_add_mul (_ : ℕ), mul_comm, add_comm] - -- Porting note: added, wrong `succ` - rfl⟩) + rw [← one_add_mul (_ : ℕ), mul_comm, add_comm]⟩) (fun a b => ⟨(a / ∏ j : Fin b, n (Fin.castLE b.is_lt.le j)) % n b, by cases m · exact b.elim0 diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index d3894ef067a77..1fa743af0d673 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -363,8 +363,8 @@ def smul : R →+* End ((forget₂ (ModuleCat R) AddCommGroupCat).obj M) where { toFun := fun (m : M) => r • m map_zero' := by dsimp; rw [smul_zero] map_add' := fun x y => by dsimp; rw [smul_add] } - map_one' := AddMonoidHom.ext (fun x => by dsimp; rw [one_smul]) - map_zero' := AddMonoidHom.ext (fun x => by dsimp; rw [zero_smul]) + map_one' := AddMonoidHom.ext (fun x => by dsimp; rw [one_smul]; rfl) + map_zero' := AddMonoidHom.ext (fun x => by dsimp; rw [zero_smul]; rfl) map_mul' r s := AddMonoidHom.ext (fun (x : M) => (smul_smul r s x).symm) map_add' r s := AddMonoidHom.ext (fun (x : M) => add_smul r s x) diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index be1310396166c..0c0050c8cb2cb 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -147,7 +147,7 @@ section Terminal /-- The trivial ring is the (strict) terminal object of `CommRingCat`. -/ def punitIsTerminal : IsTerminal (CommRingCat.of.{u} PUnit) := by refine IsTerminal.ofUnique (h := fun X => ⟨⟨⟨⟨1, rfl⟩, fun _ _ => rfl⟩, ?_, ?_⟩, ?_⟩) - · dsimp + · rfl · intros; dsimp · intros f; ext; rfl set_option linter.uppercaseLean3 false in diff --git a/Mathlib/Algebra/CharZero/Defs.lean b/Mathlib/Algebra/CharZero/Defs.lean index d1b33a2a259cc..250dacfb846ed 100644 --- a/Mathlib/Algebra/CharZero/Defs.lean +++ b/Mathlib/Algebra/CharZero/Defs.lean @@ -87,7 +87,6 @@ theorem cast_ne_zero {n : ℕ} : (n : R) ≠ 0 ↔ n ≠ 0 := theorem cast_add_one_ne_zero (n : ℕ) : (n + 1 : R) ≠ 0 := by -- Porting note: old proof was `exact_mod_cast n.succ_ne_zero` norm_cast - exact n.succ_ne_zero #align nat.cast_add_one_ne_zero Nat.cast_add_one_ne_zero @[simp, norm_cast] diff --git a/Mathlib/Algebra/Function/Support.lean b/Mathlib/Algebra/Function/Support.lean index c594af3b22104..27993451fa272 100644 --- a/Mathlib/Algebra/Function/Support.lean +++ b/Mathlib/Algebra/Function/Support.lean @@ -115,7 +115,8 @@ theorem disjoint_mulSupport_iff {f : α → M} {s : Set α} : @[to_additive (attr := simp)] theorem mulSupport_eq_empty_iff {f : α → M} : mulSupport f = ∅ ↔ f = 1 := by - simp_rw [← subset_empty_iff, mulSupport_subset_iff', funext_iff] + -- Adaptation note: `simp_rw` broke `to_additive` as of `nightly-2024-03-07` + rw [← subset_empty_iff, mulSupport_subset_iff', funext_iff] simp #align function.mul_support_eq_empty_iff Function.mulSupport_eq_empty_iff #align function.support_eq_empty_iff Function.support_eq_empty_iff diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index cb2d783c5d8fc..608a2ce981c54 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -442,7 +442,7 @@ theorem Nat.geom_sum_le {b : ℕ} (hb : 2 ≤ b) (a n : ℕ) : theorem Nat.geom_sum_Ico_le {b : ℕ} (hb : 2 ≤ b) (a n : ℕ) : ∑ i in Ico 1 n, a / b ^ i ≤ a / (b - 1) := by cases' n with n - · rw [zero_eq, Ico_eq_empty_of_le (zero_le_one' ℕ), sum_empty] + · rw [Ico_eq_empty_of_le (zero_le_one' ℕ), sum_empty] exact Nat.zero_le _ rw [← add_le_add_iff_left a] calc diff --git a/Mathlib/Algebra/Group/NatPowAssoc.lean b/Mathlib/Algebra/Group/NatPowAssoc.lean index e9ed72df007bc..faff3450fecd5 100644 --- a/Mathlib/Algebra/Group/NatPowAssoc.lean +++ b/Mathlib/Algebra/Group/NatPowAssoc.lean @@ -71,7 +71,7 @@ theorem npow_mul_comm (m n : ℕ) (x : M) : theorem npow_mul (x : M) (m n : ℕ) : x ^ (m * n) = (x ^ m) ^ n := by induction n with | zero => rw [npow_zero, Nat.mul_zero, npow_zero] - | succ n ih => rw [← Nat.add_one, mul_add, npow_add, ih, mul_one, npow_add, npow_one] + | succ n ih => rw [mul_add, npow_add, ih, mul_one, npow_add, npow_one] theorem npow_mul' (x : M) (m n : ℕ) : x ^ (m * n) = (x ^ n) ^ m := by rw [mul_comm] @@ -105,7 +105,7 @@ theorem Nat.cast_npow (R : Type*) [NonAssocSemiring R] [Pow R ℕ] [NatPowAssoc (↑(n ^ m) : R) = (↑n : R) ^ m := by induction' m with m ih · simp only [pow_zero, Nat.cast_one, npow_zero] - · rw [← Nat.add_one, npow_add, npow_add, Nat.cast_mul, ih, npow_one, npow_one] + · rw [npow_add, npow_add, Nat.cast_mul, ih, npow_one, npow_one] @[simp, norm_cast] theorem Int.cast_npow (R : Type*) [NonAssocRing R] [Pow R ℕ] [NatPowAssoc R] diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index 54c9cdca32270..4784b13c8d1c4 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -246,10 +246,10 @@ def augment (C : CochainComplex V ℕ) {X : V} (f : X ⟶ C.X 0) (w : f ≫ C.d rcases j with (_ | _ | j) <;> cases i <;> try simp · contradiction · rw [C.shape] + simp only [Nat.succ_eq_add_one] at s simp only [ComplexShape.up_Rel] contrapose! s rw [← s] - rfl d_comp_d' i j k hij hjk := by rcases k with (_ | _ | k) <;> rcases j with (_ | _ | j) <;> cases i <;> try simp cases k diff --git a/Mathlib/Algebra/Homology/ExactSequence.lean b/Mathlib/Algebra/Homology/ExactSequence.lean index 3f6ce07fb5b62..2ee0cba3690f0 100644 --- a/Mathlib/Algebra/Homology/ExactSequence.lean +++ b/Mathlib/Algebra/Homology/ExactSequence.lean @@ -180,6 +180,9 @@ lemma isComplex₂_mk (S : ComposableArrows C 2) (w : S.map' 0 1 ≫ S.map' 1 2 S.IsComplex := S.isComplex₂_iff.2 w +-- Adaptation note: nightly-2024-03-11 +-- We turn off simprocs here. +set_option simprocs false in lemma _root_.CategoryTheory.ShortComplex.isComplex_toComposableArrows (S : ShortComplex C) : S.toComposableArrows.IsComplex := isComplex₂_mk _ (by simp) diff --git a/Mathlib/Algebra/Homology/HomologySequence.lean b/Mathlib/Algebra/Homology/HomologySequence.lean index 1c673d44139b5..e29e19c2b1459 100644 --- a/Mathlib/Algebra/Homology/HomologySequence.lean +++ b/Mathlib/Algebra/Homology/HomologySequence.lean @@ -110,6 +110,9 @@ instance [K.HasHomology i] [K.HasHomology j] : dsimp infer_instance +-- Adaptation note: nightly-2024-03-11 +-- We turn off simprocs here. +set_option simprocs false in instance [K.HasHomology i] [K.HasHomology j] : Epi ((composableArrows₃ K i j).map' 2 3) := by dsimp @@ -145,6 +148,9 @@ variable (C) attribute [local simp] homologyMap_comp cyclesMap_comp opcyclesMap_comp +-- Adaptation note: nightly-2024-03-11 +-- We turn off simprocs here. +set_option simprocs false in /-- The functor `HomologicalComplex C c ⥤ ComposableArrows C 3` that maps `K` to the diagram `K.homology i ⟶ K.opcycles i ⟶ K.cycles j ⟶ K.homology j`. -/ @[simps] diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index 36c98b5f8d58f..b7f2cda8665d9 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -46,8 +46,11 @@ noncomputable def single (j : ι) : V ⥤ HomologicalComplex V c where split_ifs with h · subst h simp - · rw [if_neg h] - simp + · -- Adaptation note: after nightly-2024-03-07, the previous sensible proof + -- `rw [if_neg h]; simp` fails with "motive not type correct". + -- The following is horrible. + convert (id_zero (C := V)).symm + all_goals simp [if_neg h] map_comp f g := by ext dsimp diff --git a/Mathlib/Algebra/Lie/Solvable.lean b/Mathlib/Algebra/Lie/Solvable.lean index be8b2f983f793..551f0cccdef80 100644 --- a/Mathlib/Algebra/Lie/Solvable.lean +++ b/Mathlib/Algebra/Lie/Solvable.lean @@ -91,7 +91,7 @@ theorem derivedSeriesOfIdeal_add (k l : ℕ) : D (k + l) I = D k (D l I) := by theorem derivedSeriesOfIdeal_le {I J : LieIdeal R L} {k l : ℕ} (h₁ : I ≤ J) (h₂ : l ≤ k) : D k I ≤ D l J := by revert l; induction' k with k ih <;> intro l h₂ - · rw [Nat.zero_eq, le_zero_iff] at h₂; rw [h₂, derivedSeriesOfIdeal_zero]; exact h₁ + · rw [le_zero_iff] at h₂; rw [h₂, derivedSeriesOfIdeal_zero]; exact h₁ · have h : l = k.succ ∨ l ≤ k := by rwa [le_iff_eq_or_lt, Nat.lt_succ_iff] at h₂ cases' h with h h · rw [h, derivedSeriesOfIdeal_succ, derivedSeriesOfIdeal_succ] diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index 62b19d248d330..ffb5a47e91881 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -364,8 +364,8 @@ variable (R) /-- `nsmul` is equal to any other module structure via a cast. -/ theorem nsmul_eq_smul_cast (n : ℕ) (b : M) : n • b = (n : R) • b := by induction' n with n ih - · rw [Nat.zero_eq, Nat.cast_zero, zero_smul, zero_smul] - · rw [Nat.succ_eq_add_one, Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] + · rw [Nat.cast_zero, zero_smul, zero_smul] + · rw [Nat.cast_succ, add_smul, add_smul, one_smul, ih, one_smul] #align nsmul_eq_smul_cast nsmul_eq_smul_cast end diff --git a/Mathlib/Algebra/Module/PID.lean b/Mathlib/Algebra/Module/PID.lean index d7aa2d98b055b..a84c9820e9667 100644 --- a/Mathlib/Algebra/Module/PID.lean +++ b/Mathlib/Algebra/Module/PID.lean @@ -180,8 +180,7 @@ theorem torsion_by_prime_power_decomposition (hN : Module.IsTorsion' N (Submonoi ∃ (d : ℕ) (k : Fin d → ℕ), Nonempty <| N ≃ₗ[R] ⨁ i : Fin d, R ⧸ R ∙ p ^ (k i : ℕ) := by obtain ⟨d, s, hs⟩ := @Module.Finite.exists_fin _ _ _ _ _ h'; use d; clear h' induction' d with d IH generalizing N - · simp only [Nat.zero_eq] at * - -- Porting note: was `use fun i => finZeroElim i` + · -- Porting note: was `use fun i => finZeroElim i` use finZeroElim rw [Set.range_eq_empty, Submodule.span_empty] at hs haveI : Unique N := diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 54949135a9f29..484c832c80072 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -263,7 +263,7 @@ def evalZPow : PositivityExt where eval {u α} zα pα e := do | .app (.app (.app (.const `Neg.neg _) _) _) b' => let b' ← whnfR b' let .true := b'.isAppOfArity ``OfNat.ofNat 3 | throwError "not a ^ -n where n is a literal" - let some n := (b'.getRevArg! 1).natLit? | throwError "not a ^ -n where n is a literal" + let some n := (b'.getRevArg! 1).rawNatLit? | throwError "not a ^ -n where n is a literal" guard (n % 2 = 0) have m : Q(ℕ) := mkRawNatLit (n / 2) haveI' : $b =Q (-$m) + (-$m) := ⟨⟩ -- b = bit0 (-m) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 2d20ecb9e691b..543ae54ff5987 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -1075,7 +1075,7 @@ theorem fract_mul_nat (a : α) (b : ℕ) : ∃ z : ℤ, fract a * b - fract (a * induction' b with c hc · use 0; simp · rcases hc with ⟨z, hz⟩ - rw [Nat.succ_eq_add_one, Nat.cast_add, mul_add, mul_add, Nat.cast_one, mul_one, mul_one] + rw [Nat.cast_add, mul_add, mul_add, Nat.cast_one, mul_one, mul_one] rcases fract_add (a * c) a with ⟨y, hy⟩ use z - y rw [Int.cast_sub, ← hz, ← hy] diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index 0c5e1c20bcd18..eb041fdb96124 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -679,7 +679,7 @@ instance monoid [Monoid R] [AddMonoid M] [DistribMulAction R M] [DistribMulActio rw [List.range_succ, List.map_append, List.sum_append, List.map_singleton, List.sum_singleton, Nat.sub_self, pow_zero, one_smul, List.smul_sum, List.map_map, fst_pow, Function.comp] - simp_rw [← smul_comm (_ : R) (_ : Rᵐᵒᵖ), smul_smul, ← pow_succ, Nat.succ_eq_add_one] + simp_rw [← smul_comm (_ : R) (_ : Rᵐᵒᵖ), smul_smul, ← pow_succ] congr 2 refine' List.map_congr fun i hi => _ rw [List.mem_range, Nat.lt_succ_iff] at hi diff --git a/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean b/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean index 067d9d3a08389..b47409e22e51d 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean @@ -56,7 +56,6 @@ theorem homotopyPToId_eventually_constant {q n : ℕ} (hqn : n < q) : Homotopy.ofEq_hom, Pi.zero_apply, Homotopy.add_hom, Homotopy.compLeft_hom, add_zero, Homotopy.nullHomotopy'_hom, ComplexShape.down_Rel, hσ'_eq_zero hqn (c_mk (n + 1) n rfl), dite_eq_ite, ite_self, comp_zero, zero_add, homotopyPToId] - rfl set_option linter.uppercaseLean3 false in #align algebraic_topology.dold_kan.homotopy_P_to_id_eventually_constant AlgebraicTopology.DoldKan.homotopyPToId_eventually_constant diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index 7c657ee5b1448..371de48c8df18 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -350,18 +350,10 @@ theorem add_point (f : α → E) {s : Set α} {x : α} (hx : x ∈ s) (u : ℕ · apply Finset.sum_congr rfl fun i hi => ?_ rw [Finset.mem_Ico] at hi dsimp only [w] - have A : ¬1 + i + 1 < N := fun h => by - rw [add_assoc, add_comm] at h - exact hi.left.not_lt (i.lt_succ_self.trans (i.succ.lt_succ_self.trans h)) - have B : ¬1 + i + 1 = N := fun h => by - rw [← h, add_assoc, add_comm] at hi - exact Nat.not_succ_le_self i (i.succ.le_succ.trans hi.left) - have C : ¬1 + i < N := fun h => by - rw [add_comm] at h - exact hi.left.not_lt (i.lt_succ_self.trans h) - have D : ¬1 + i = N := fun h => by - rw [← h, add_comm, Nat.succ_le_iff] at hi - exact hi.left.ne rfl + have A : ¬1 + i + 1 < N := by omega + have B : ¬1 + i + 1 = N := by omega + have C : ¬1 + i < N := by omega + have D : ¬1 + i = N := by omega rw [if_neg A, if_neg B, if_neg C, if_neg D] congr 3 <;> · rw [add_comm, Nat.sub_one]; apply Nat.pred_succ _ = (∑ i in Finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i))) + @@ -533,9 +525,11 @@ theorem comp_le_of_antitoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β fun i => φst (ut _)⟩ le_rfl dsimp only [Subtype.coe_mk] - rw [edist_comm, Nat.sub_sub, add_comm, Nat.sub_succ, Nat.add_one, Nat.succ_pred_eq_of_pos] - simp only [Function.comp_apply] - simpa only [tsub_pos_iff_lt, Finset.mem_range] using hx + rw [edist_comm, Nat.sub_sub, add_comm, Nat.sub_succ, Nat.add_one, Nat.succ_eq_add_one] + simp only [Function.comp_apply, Nat.pred_eq_sub_one, Nat.sub_add_eq] + congr + simp at hx + omega #align evariation_on.comp_le_of_antitone_on eVariationOn.comp_le_of_antitoneOn theorem comp_eq_of_monotoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) : diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 3ef78a1705466..d0a94c9851b34 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -1086,7 +1086,7 @@ theorem ContDiffWithinAt.iteratedFderivWithin_right {i : ℕ} (hf : ContDiffWith (hs : UniqueDiffOn 𝕜 s) (hmn : (m + i : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (iteratedFDerivWithin 𝕜 i f s) s x₀ := by induction' i with i hi generalizing m - · rw [Nat.zero_eq, ENat.coe_zero, add_zero] at hmn + · rw [ENat.coe_zero, add_zero] at hmn exact (hf.of_le hmn).continuousLinearMap_comp ((continuousMultilinearCurryFin0 𝕜 E F).symm : _ →L[𝕜] E [×0]→L[𝕜] F) · rw [Nat.cast_succ, add_comm _ 1, ← add_assoc] at hmn diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index ad9bf014b7b29..2146894b8dab9 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -999,7 +999,7 @@ theorem HasFTaylorSeriesUpToOn.eq_ftaylor_series_of_uniqueDiffOn (h : HasFTaylorSeriesUpToOn n f p s) {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) : p x m = iteratedFDerivWithin 𝕜 m f s x := by induction' m with m IH generalizing x - · rw [Nat.zero_eq, h.zero_eq' hx, iteratedFDerivWithin_zero_eq_comp]; rfl + · rw [h.zero_eq' hx, iteratedFDerivWithin_zero_eq_comp]; rfl · have A : (m : ℕ∞) < n := lt_of_lt_of_le (WithTop.coe_lt_coe.2 (lt_add_one m)) hmn have : HasFDerivWithinAt (fun y : E => iteratedFDerivWithin 𝕜 m f s y) diff --git a/Mathlib/Analysis/Convex/Basic.lean b/Mathlib/Analysis/Convex/Basic.lean index 2cb023ab749fa..3a0cfb3a19c5c 100644 --- a/Mathlib/Analysis/Convex/Basic.lean +++ b/Mathlib/Analysis/Convex/Basic.lean @@ -712,12 +712,16 @@ theorem ite_eq_mem_stdSimplex (i : ι) : (if i = · then (1 : 𝕜) else 0) ∈ simpa only [@eq_comm _ i, ← Pi.single_apply] using single_mem_stdSimplex 𝕜 i #align ite_eq_mem_std_simplex ite_eq_mem_stdSimplex +-- Adaptation note: as of `nightly-2024-03-11`, we need a type annotation on the segment in the +-- following two lemmas. + /-- The edges are contained in the simplex. -/ lemma segment_single_subset_stdSimplex (i j : ι) : - [Pi.single i 1 -[𝕜] Pi.single j 1] ⊆ stdSimplex 𝕜 ι := + ([Pi.single i 1 -[𝕜] Pi.single j 1] : Set (ι → 𝕜)) ⊆ stdSimplex 𝕜 ι := (convex_stdSimplex 𝕜 ι).segment_subset (single_mem_stdSimplex _ _) (single_mem_stdSimplex _ _) -lemma stdSimplex_fin_two : stdSimplex 𝕜 (Fin 2) = [Pi.single 0 1 -[𝕜] Pi.single 1 1] := by +lemma stdSimplex_fin_two : + stdSimplex 𝕜 (Fin 2) = ([Pi.single 0 1 -[𝕜] Pi.single 1 1] : Set (Fin 2 → 𝕜)) := by refine Subset.antisymm ?_ (segment_single_subset_stdSimplex 𝕜 (0 : Fin 2) 1) rintro f ⟨hf₀, hf₁⟩ rw [Fin.sum_univ_two] at hf₁ diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean index 60aeb7d9e16e1..1627ce5bd5a54 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean @@ -74,7 +74,7 @@ theorem int_prod_range_nonneg (m : ℤ) (n : ℕ) (hn : Even n) : induction' n with n ihn · simp rw [← two_mul] at ihn - rw [← two_mul, Nat.succ_eq_add_one, mul_add, mul_one, ← one_add_one_eq_two, ← add_assoc, + rw [← two_mul, mul_add, mul_one, ← one_add_one_eq_two, ← add_assoc, Finset.prod_range_succ, Finset.prod_range_succ, mul_assoc] refine' mul_nonneg ihn _; generalize (1 + 1) * n = k rcases le_or_lt m k with hmk | hmk diff --git a/Mathlib/Analysis/InnerProductSpace/Orientation.lean b/Mathlib/Analysis/InnerProductSpace/Orientation.lean index e84dd747c950c..81e9c28fcd02d 100644 --- a/Mathlib/Analysis/InnerProductSpace/Orientation.lean +++ b/Mathlib/Analysis/InnerProductSpace/Orientation.lean @@ -212,7 +212,7 @@ theorem volumeForm_robust (b : OrthonormalBasis (Fin n) ℝ E) (hb : b.toBasis.o cases n · classical have : o = positiveOrientation := hb.symm.trans b.toBasis.orientation_isEmpty - simp_rw [volumeForm, Or.by_cases, dif_pos this, Basis.det_isEmpty] + simp_rw [volumeForm, Or.by_cases, dif_pos this, Nat.rec_zero, Basis.det_isEmpty] · simp_rw [volumeForm] rw [same_orientation_iff_det_eq_det, hb] exact o.finOrthonormalBasis_orientation _ _ @@ -225,7 +225,7 @@ theorem volumeForm_robust_neg (b : OrthonormalBasis (Fin n) ℝ E) (hb : b.toBas cases' n with n · classical have : positiveOrientation ≠ o := by rwa [b.toBasis.orientation_isEmpty] at hb - simp_rw [volumeForm, Or.by_cases, dif_neg this.symm, Basis.det_isEmpty] + simp_rw [volumeForm, Or.by_cases, dif_neg this.symm, Nat.rec_zero, Basis.det_isEmpty] let e : OrthonormalBasis (Fin n.succ) ℝ E := o.finOrthonormalBasis n.succ_pos Fact.out simp_rw [volumeForm] apply e.det_eq_neg_det_of_opposite_orientation b diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index f7a8fcbae6353..e9ea989b557ff 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -1177,8 +1177,7 @@ theorem LinearIsometryEquiv.reflections_generate_dim_aux [FiniteDimensional ℝ · -- Base case: `n = 0`, the fixed subspace is the whole space, so `φ = id` refine' ⟨[], rfl.le, show φ = 1 from _⟩ have : ker (ContinuousLinearMap.id ℝ F - φ) = ⊤ := by - rwa [Nat.zero_eq, le_zero_iff, Submodule.finrank_eq_zero, - Submodule.orthogonal_eq_bot_iff] at hn + rwa [le_zero_iff, Submodule.finrank_eq_zero, Submodule.orthogonal_eq_bot_iff] at hn symm ext x have := LinearMap.congr_fun (LinearMap.ker_eq_top.mp this) x diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 90ce2b5af34f5..e6ee77689369f 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1420,7 +1420,11 @@ theorem norm_pos_iff''' [T0Space E] {a : E} : 0 < ‖a‖ ↔ a ≠ 1 := by @[to_additive] theorem SeminormedGroup.tendstoUniformlyOn_one {f : ι → κ → G} {s : Set κ} {l : Filter ι} : TendstoUniformlyOn f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ‖f i x‖ < ε := by - simp_rw [tendstoUniformlyOn_iff, Pi.one_apply, dist_one_left] + -- Adaptation note: nightly-2024-03-11. + -- There seems to be a bug in a dsimproc. + -- Originally this was `simp_rw` instead of `simp only`, + -- but this creates a bad proof term with nested `OfNat.ofNat` that trips up `@[to_additive]`. + simp only [tendstoUniformlyOn_iff, Pi.one_apply, dist_one_left] #align seminormed_group.tendsto_uniformly_on_one SeminormedGroup.tendstoUniformlyOn_one #align seminormed_add_group.tendsto_uniformly_on_zero SeminormedAddGroup.tendstoUniformlyOn_zero diff --git a/Mathlib/Analysis/NormedSpace/Exponential.lean b/Mathlib/Analysis/NormedSpace/Exponential.lean index 267a4eb246bf5..c5eef40fe1ecf 100644 --- a/Mathlib/Analysis/NormedSpace/Exponential.lean +++ b/Mathlib/Analysis/NormedSpace/Exponential.lean @@ -534,7 +534,7 @@ theorem exp_sum_of_commute {ι} (s : Finset ι) (f : ι → 𝔸) theorem exp_nsmul (n : ℕ) (x : 𝔸) : exp 𝕂 (n • x) = exp 𝕂 x ^ n := by induction' n with n ih - · rw [Nat.zero_eq, zero_smul, pow_zero, exp_zero] + · rw [zero_smul, pow_zero, exp_zero] · rw [succ_nsmul, pow_succ, exp_add_of_commute ((Commute.refl x).smul_right n), ih] #align exp_nsmul NormedSpace.exp_nsmul diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index 0650e53143665..143c3727b3183 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -320,7 +320,7 @@ theorem dist_next_apply_le_of_le {f₁ f₂ : FunSpace v} {n : ℕ} {d : ℝ} theorem dist_iterate_next_apply_le (f₁ f₂ : FunSpace v) (n : ℕ) (t : Icc v.tMin v.tMax) : dist (next^[n] f₁ t) (next^[n] f₂ t) ≤ (v.L * |t.1 - v.t₀|) ^ n / n ! * dist f₁ f₂ := by induction' n with n ihn generalizing t - · rw [Nat.zero_eq, pow_zero, Nat.factorial_zero, Nat.cast_one, div_one, one_mul] + · rw [pow_zero, Nat.factorial_zero, Nat.cast_one, div_one, one_mul] exact dist_apply_le_dist f₁ f₂ t · rw [iterate_succ_apply', iterate_succ_apply'] exact dist_next_apply_le_of_le ihn _ diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index 9028585076f95..bb357e7e76dcf 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -280,7 +280,7 @@ theorem GammaAux_recurrence1 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : simp · dsimp only [GammaAux] have hh1 : -(s + 1).re < n := by - rw [Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one] at h1 + rw [Nat.cast_add, Nat.cast_one] at h1 rw [add_re, one_re]; linarith rw [← hn (s + 1) hh1] #align complex.Gamma_aux_recurrence1 Complex.GammaAux_recurrence1 @@ -297,7 +297,7 @@ theorem GammaAux_recurrence2 (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : · dsimp only [GammaAux] have : GammaAux n (s + 1 + 1) / (s + 1) = GammaAux n (s + 1) := by have hh1 : -(s + 1).re < n := by - rw [Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one] at h1 + rw [Nat.cast_add, Nat.cast_one] at h1 rw [add_re, one_re]; linarith rw [GammaAux_recurrence1 (s + 1) n hh1] rw [this] @@ -313,7 +313,7 @@ theorem Gamma_eq_GammaAux (s : ℂ) (n : ℕ) (h1 : -s.re < ↑n) : Gamma s = Ga have u : ∀ k : ℕ, GammaAux (⌊1 - s.re⌋₊ + k) s = Gamma s := by intro k; induction' k with k hk · simp [Gamma] - · rw [← hk, Nat.succ_eq_add_one, ← add_assoc] + · rw [← hk, ← add_assoc] refine' (GammaAux_recurrence2 s (⌊1 - s.re⌋₊ + k) _).symm rw [Nat.cast_add] have i0 := Nat.sub_one_lt_floor (1 - s.re) @@ -630,7 +630,7 @@ theorem Gamma_ne_zero {s : ℝ} (hs : ∀ m : ℕ, s ≠ -m) : Gamma s ≠ 0 := rw [hs] push_cast ring - · rw [Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one, neg_add] at hs' + · rw [Nat.cast_add, Nat.cast_one, neg_add] at hs' linarith rw [Gamma_add_one, mul_ne_zero_iff] at this · exact this.2 diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean index e300792a30289..eb525e5df1e28 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean @@ -213,7 +213,7 @@ theorem sin_pi_mul_eq (z : ℂ) (n : ℕ) : rcases eq_or_ne z 0 with (rfl | hz) · simp induction' n with n hn - · simp_rw [Nat.zero_eq, mul_zero, pow_zero, mul_one, Finset.prod_range_zero, mul_one, + · simp_rw [mul_zero, pow_zero, mul_one, Finset.prod_range_zero, mul_one, integral_one, sub_zero] rw [integral_cos_mul_complex (mul_ne_zero two_ne_zero hz), Complex.ofReal_zero, mul_zero, Complex.sin_zero, zero_div, sub_zero, diff --git a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean index c472617b130db..b9c4cad9cd15f 100644 --- a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean +++ b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean @@ -133,6 +133,9 @@ section Five variable {R₁ R₂ : ComposableArrows C 4} (hR₁ : R₁.Exact) (hR₂ : R₂.Exact) (φ : R₁ ⟶ R₂) +-- Adaptation note: nightly-2024-03-11 +-- We turn off simprocs here. +set_option simprocs false in /-- The five lemma. -/ theorem isIso_of_epi_of_isIso_of_isIso_of_mono (h₀ : Epi (app' φ 0)) (h₁ : IsIso (app' φ 1)) (h₂ : IsIso (app' φ 3)) (h₃ : Mono (app' φ 4)) : IsIso (app' φ 2) := by diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index c900ee4ca65de..28c1c8f99269e 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -327,9 +327,9 @@ lemma ofCocomplex_exactAt_succ (n : ℕ) : rw [HomologicalComplex.exactAt_iff' _ n (n + 1) (n + 1 + 1) (by simp) (by simp)] cases n all_goals - dsimp [ofCocomplex, HomologicalComplex.sc', HomologicalComplex.shortComplexFunctor', - CochainComplex.mk', CochainComplex.mk] - simp only [CochainComplex.of_d] + dsimp only [ofCocomplex, CochainComplex.mk', CochainComplex.mk, CochainComplex.of, + HomologicalComplex.sc', HomologicalComplex.shortComplexFunctor', eqToHom_refl, dite_eq_ite] + simp only [Nat.reduceAdd, zero_add, ↓reduceIte, comp_id] apply exact_f_d instance (n : ℕ) : Injective ((ofCocomplex Z).X n) := by diff --git a/Mathlib/CategoryTheory/ComposableArrows.lean b/Mathlib/CategoryTheory/ComposableArrows.lean index fb0fba2d29090..0b092216b857e 100644 --- a/Mathlib/CategoryTheory/ComposableArrows.lean +++ b/Mathlib/CategoryTheory/ComposableArrows.lean @@ -37,6 +37,19 @@ up to `n = 7` in order to formalize spectral sequences following Verdier) -/ +/-! +New `simprocs` that run even in `dsimp` have caused breakages in this file. + +(e.g. `dsimp` can now simplify `2 + 3` to `5`) + +For now, we just turn off simprocs in this file. +We'll soon provide finer grained options here, e.g. to turn off simprocs only in `dsimp`, etc. + +*However*, hopefully it is possible to refactor the material here so that no backwards compatibility +`set_option`s are required at all +-/ +set_option simprocs false + namespace CategoryTheory open Category @@ -173,7 +186,7 @@ def homMk {F G : ComposableArrows C n} (app : ∀ i, F.obj i ⟶ G.obj i) obtain rfl := hj rw [F.map'_self i, G.map'_self i, id_comp, comp_id] · intro i j hj hj' - rw [Nat.succ_eq_add_one, ← add_assoc] at hj + rw [← add_assoc] at hj subst hj rw [F.map'_comp i (i + k) (i + k + 1), G.map'_comp i (i + k) (i + k + 1), assoc, w (i + k) (by valid), reassoc_of% (hk i (i + k) rfl (by valid))] @@ -529,8 +542,7 @@ lemma ext_succ {F G : ComposableArrows C (n + 1)} (h₀ : F.obj' 0 = G.obj' 0) intro ⟨i, hi⟩ cases' i with i · exact h₀ - · rw [Nat.succ_eq_add_one] at hi - exact Functor.congr_obj h ⟨i, by valid⟩ + · exact Functor.congr_obj h ⟨i, by valid⟩ exact Functor.ext_of_iso (isoMkSucc (eqToIso h₀) (eqToIso h) (by rw [w] dsimp [app'] diff --git a/Mathlib/CategoryTheory/Idempotents/Biproducts.lean b/Mathlib/CategoryTheory/Idempotents/Biproducts.lean index 91365df024cc0..566aceca24d1c 100644 --- a/Mathlib/CategoryTheory/Idempotents/Biproducts.lean +++ b/Mathlib/CategoryTheory/Idempotents/Biproducts.lean @@ -65,8 +65,8 @@ def bicone [HasFiniteBiproducts C] {J : Type} [Finite J] (F : J → Karoubi C) : simp only [biproduct.ι_map, biproduct.bicone_π, biproduct.map_π, eqToHom_refl, id_eq, hom_ext_iff, comp_f, assoc, bicone_ι_π_self_assoc, idem] · dsimp - simp only [hom_ext_iff, biproduct.ι_map, biproduct.map_π, comp_f, assoc, ne_eq, - biproduct.ι_π_ne_assoc _ h, comp_zero, zero_comp] + simp only [biproduct.ι_map, biproduct.map_π, hom_ext_iff, comp_f, + assoc, biproduct.ι_π_ne_assoc _ h, zero_comp, comp_zero, instZero_zero] #align category_theory.idempotents.karoubi.biproducts.bicone CategoryTheory.Idempotents.Karoubi.Biproducts.bicone end Biproducts diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 9c1e061dd7743..edf3aaffd46ee 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -342,7 +342,7 @@ def isoBiproductEmbedding (M : Mat_ C) : M ≅ ⨁ fun i => (embedding C).obj (M rw [Finset.sum_apply, Finset.sum_apply, Finset.sum_eq_single i]; rotate_left · intro b _ hb dsimp - simp only [Finset.sum_const, Finset.card_singleton, one_smul] + simp only [Finset.sum_singleton] rw [dif_neg hb.symm, zero_comp] · intro h simp at h diff --git a/Mathlib/Combinatorics/Additive/Energy.lean b/Mathlib/Combinatorics/Additive/Energy.lean index d624f126864de..644041d4383a1 100644 --- a/Mathlib/Combinatorics/Additive/Energy.lean +++ b/Mathlib/Combinatorics/Additive/Energy.lean @@ -74,8 +74,8 @@ theorem le_multiplicativeEnergy : s.card * t.card ≤ multiplicativeEnergy s t : refine' card_le_card_of_inj_on (@fun x => ((x.1, x.1), x.2, x.2)) (by -- Porting note: changed this from a `simp` proof without `only` because of a timeout - simp only [ ← and_imp, mem_product, and_imp, Prod.forall, mem_filter, - and_self, and_true, imp_self, implies_true]) fun a _ b _ => _ + simp only [← and_imp, mem_product, Prod.forall, mem_filter, and_self, and_true, imp_self, + implies_true]) fun a _ b _ => _ simp only [Prod.mk.inj_iff, and_self_iff, and_imp] exact Prod.ext #align finset.le_multiplicative_energy Finset.le_multiplicativeEnergy diff --git a/Mathlib/Combinatorics/Catalan.lean b/Mathlib/Combinatorics/Catalan.lean index 23be982e4f0c3..c5278601b60b4 100644 --- a/Mathlib/Combinatorics/Catalan.lean +++ b/Mathlib/Combinatorics/Catalan.lean @@ -89,8 +89,8 @@ private def gosperCatalan (n j : ℕ) : ℚ := private theorem gosper_trick {n i : ℕ} (h : i ≤ n) : gosperCatalan (n + 1) (i + 1) - gosperCatalan (n + 1) i = Nat.centralBinom i / (i + 1) * Nat.centralBinom (n - i) / (n - i + 1) := by - have l₁ : (i : ℚ) + 1 ≠ 0 := by norm_cast; exact i.succ_ne_zero - have l₂ : (n : ℚ) - i + 1 ≠ 0 := by norm_cast; exact (n - i).succ_ne_zero + have l₁ : (i : ℚ) + 1 ≠ 0 := by norm_cast + have l₂ : (n : ℚ) - i + 1 ≠ 0 := by norm_cast have h₁ := (mul_div_cancel_left (↑(Nat.centralBinom (i + 1))) l₁).symm have h₂ := (mul_div_cancel_left (↑(Nat.centralBinom (n - i + 1))) l₂).symm have h₃ : ((i : ℚ) + 1) * (i + 1).centralBinom = 2 * (2 * i + 1) * i.centralBinom := @@ -108,9 +108,9 @@ private theorem gosper_trick {n i : ℕ} (h : i ≤ n) : private theorem gosper_catalan_sub_eq_central_binom_div (n : ℕ) : gosperCatalan (n + 1) (n + 1) - gosperCatalan (n + 1) 0 = Nat.centralBinom (n + 1) / (n + 2) := by - have : (n : ℚ) + 1 ≠ 0 := by norm_cast; exact n.succ_ne_zero - have : (n : ℚ) + 1 + 1 ≠ 0 := by norm_cast; exact (n + 1).succ_ne_zero - have h : (n : ℚ) + 2 ≠ 0 := by norm_cast; exact (n + 1).succ_ne_zero + have : (n : ℚ) + 1 ≠ 0 := by norm_cast + have : (n : ℚ) + 1 + 1 ≠ 0 := by norm_cast + have h : (n : ℚ) + 2 ≠ 0 := by norm_cast simp only [gosperCatalan, Nat.sub_zero, Nat.centralBinom_zero, Nat.sub_self] field_simp ring @@ -195,7 +195,6 @@ theorem mem_treesOfNumNodesEq {x : Tree Unit} {n : ℕ} : x ∈ treesOfNumNodesEq n ↔ x.numNodes = n := by induction x using Tree.unitRecOn generalizing n <;> cases n <;> simp [treesOfNumNodesEq_succ, Nat.succ_eq_add_one, *] - exact (Nat.succ_ne_zero _).symm #align tree.mem_trees_of_nodes_eq Tree.mem_treesOfNumNodesEq theorem mem_treesOfNumNodesEq_numNodes (x : Tree Unit) : x ∈ treesOfNumNodesEq x.numNodes := diff --git a/Mathlib/Combinatorics/Composition.lean b/Mathlib/Combinatorics/Composition.lean index 9affc16506448..ad340ad3677f5 100644 --- a/Mathlib/Combinatorics/Composition.lean +++ b/Mathlib/Combinatorics/Composition.lean @@ -716,7 +716,7 @@ theorem get_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi) simpa using get_zero hi · simp only [splitWrtCompositionAux._eq_2, get_cons_succ, IH, take, sum_cons, Nat.add_eq, add_zero, splitAt_eq_take_drop, ← drop_take, drop_drop] - rw [Nat.succ_eq_add_one, add_comm i 1, add_comm] + rw [add_comm i 1, add_comm] #align list.nth_le_split_wrt_composition_aux List.get_splitWrtCompositionAux /-- The `i`-th sublist in the splitting of a list `l` along a composition `c`, is the slice of `l` diff --git a/Mathlib/Combinatorics/Derangements/Finite.lean b/Mathlib/Combinatorics/Derangements/Finite.lean index dc20d2ef23112..11f9d2278fa44 100644 --- a/Mathlib/Combinatorics/Derangements/Finite.lean +++ b/Mathlib/Combinatorics/Derangements/Finite.lean @@ -125,5 +125,5 @@ theorem numDerangements_sum (n : ℕ) : intro x hx have h_le : x ≤ n := Finset.mem_range_succ_iff.mp hx rw [Nat.succ_sub h_le, Nat.ascFactorial_succ, add_right_comm, add_tsub_cancel_of_le h_le, - Int.ofNat_mul, Int.ofNat_succ, mul_left_comm] + Int.ofNat_mul, Int.ofNat_add, mul_left_comm, Nat.cast_one] #align num_derangements_sum numDerangements_sum diff --git a/Mathlib/Combinatorics/Hindman.lean b/Mathlib/Combinatorics/Hindman.lean index 5d3ac63e182e5..de63aa32e466d 100644 --- a/Mathlib/Combinatorics/Hindman.lean +++ b/Mathlib/Combinatorics/Hindman.lean @@ -244,7 +244,7 @@ set_option linter.uppercaseLean3 false in theorem FP_drop_subset_FP {M} [Semigroup M] (a : Stream' M) (n : ℕ) : FP (a.drop n) ⊆ FP a := by induction' n with n ih · rfl - rw [Nat.succ_eq_one_add, ← Stream'.drop_drop] + rw [Nat.add_comm, ← Stream'.drop_drop] exact _root_.trans (FP.tail _) ih set_option linter.uppercaseLean3 false in #align hindman.FP_drop_subset_FP Hindman.FP_drop_subset_FP diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index eb0b36470c571..edf4306b2d84c 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -98,7 +98,6 @@ theorem card_div_choose_le_card_shadow_div_choose (hr : r ≠ 0) rw [div_le_div_iff] <;> norm_cast · cases' r with r · exact (hr rfl).elim - rw [Nat.succ_eq_add_one] at * rw [tsub_add_eq_add_tsub hr', add_tsub_add_eq_tsub_right] at h𝒜 apply le_of_mul_le_mul_right _ (pos_iff_ne_zero.2 hr) convert Nat.mul_le_mul_right ((Fintype.card α).choose r) h𝒜 using 1 @@ -188,7 +187,6 @@ theorem le_card_falling_div_choose [Fintype α] (hk : k ≤ Fintype.card α) zero_eq, zero_add, range_one, ge_iff_le, sum_singleton, nonpos_iff_eq_zero, tsub_zero, choose_self, cast_one, div_one, cast_le] exact card_le_card (slice_subset_falling _ _) - rw [succ_eq_add_one] at * rw [sum_range_succ, ← slice_union_shadow_falling_succ, card_union_of_disjoint (IsAntichain.disjoint_slice_shadow_falling h𝒜), cast_add, _root_.add_div, add_comm] diff --git a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean index 9bdf5acac2087..15db26a168c28 100644 --- a/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean +++ b/Mathlib/Combinatorics/SimpleGraph/AdjMatrix.lean @@ -263,8 +263,7 @@ theorem adjMatrix_pow_apply_eq_card_walk [DecidableEq V] [Semiring α] (n : ℕ) rw [card_set_walk_length_eq] induction' n with n ih generalizing u v · obtain rfl | h := eq_or_ne u v <;> simp [finsetWalkLength, *] - · nth_rw 1 [Nat.succ_eq_one_add] - simp only [pow_add, pow_one, finsetWalkLength, ih, adjMatrix_mul_apply] + · simp only [pow_succ, finsetWalkLength, ih, adjMatrix_mul_apply] rw [Finset.card_biUnion] · norm_cast simp only [Nat.cast_sum, card_map, neighborFinset_def] diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index e52368c40862f..4ab9b28f9b8f2 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -256,7 +256,7 @@ theorem ack_le_ack {m₁ m₂ n₁ n₂ : ℕ} (hm : m₁ ≤ m₂) (hn : n₁ theorem ack_succ_right_le_ack_succ_left (m n : ℕ) : ack m (n + 1) ≤ ack (m + 1) n := by cases' n with n n · simp - · rw [ack_succ_succ, succ_eq_add_one] + · rw [ack_succ_succ] apply ack_mono_right m (le_trans _ <| add_add_one_le_ack _ n) omega #align ack_succ_right_le_ack_succ_left ack_succ_right_le_ack_succ_left @@ -267,7 +267,7 @@ private theorem sq_le_two_pow_add_one_minus_three (n : ℕ) : n ^ 2 ≤ 2 ^ (n + · norm_num · cases' k with k k · norm_num - · rw [succ_eq_add_one, add_sq, Nat.pow_succ 2, mul_comm _ 2, two_mul (2 ^ _), + · rw [add_sq, Nat.pow_succ 2, mul_comm _ 2, two_mul (2 ^ _), add_tsub_assoc_of_le, add_comm (2 ^ _), add_assoc] · apply Nat.add_le_add hk norm_num diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 2d2157c403462..7f50ad5ce7b99 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -570,7 +570,7 @@ theorem stepNormal_then (c) (k k' : Cont) (v) : induction c generalizing k v with simp only [Cont.then, stepNormal, *] | cons c c' ih _ => rw [← ih, Cont.then] | comp c c' _ ih' => rw [← ih', Cont.then] - | case => cases v.headI <;> simp only [Nat.rec] + | case => cases v.headI <;> simp only [Nat.rec_zero] | fix c ih => rw [← ih, Cont.then] | _ => simp only [Cfg.then] #align turing.to_partrec.step_normal_then Turing.ToPartrec.stepNormal_then @@ -618,7 +618,7 @@ theorem stepNormal.is_ret (c k v) : ∃ k' v', stepNormal c k v = Cfg.ret k' v' | case f g IHf IHg => rw [stepNormal] simp only [] - cases v.headI <;> simp only [] <;> [apply IHf; apply IHg] + cases v.headI <;> [apply IHf; apply IHg] | fix f IHf => apply IHf | _ => exact ⟨_, _, rfl⟩ #align turing.to_partrec.step_normal.is_ret Turing.ToPartrec.stepNormal.is_ret @@ -711,7 +711,7 @@ theorem code_is_ok (c) : Code.Ok c := by rw [stepRet, IHf] | case f g IHf IHg => simp only [Code.eval] - cases v.headI <;> simp only [Code.eval] <;> [apply IHf; apply IHg] + cases v.headI <;> simp only [Nat.rec_zero, Part.bind_eq_bind] <;> [apply IHf; apply IHg] | fix f IHf => rw [cont_eval_fix IHf] | _ => simp only [Code.eval, pure_bind] #align turing.to_partrec.code_is_ok Turing.ToPartrec.code_is_ok diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 6b9e1f2e14854..6afe330f9de69 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1587,7 +1587,7 @@ theorem lt_sub_one_iff {n : ℕ} {k : Fin (n + 2)} : k < k - 1 ↔ k = 0 := by simp only [zero_eq, zero_eta, zero_sub, lt_iff_val_lt_val, val_zero, coe_neg_one, add_pos_iff, _root_.zero_lt_one, or_true] have : (k + 1 + (n + 1)) % (n + 2) = k % (n + 2) := by - rw [add_right_comm, add_assoc, add_mod_right] + rw [add_right_comm, add_assoc, add_assoc, one_add_one_eq_two, add_mod_right] simp [lt_iff_val_lt_val, ext_iff, Fin.coe_sub, succ_eq_add_one, this, mod_eq_of_lt ((lt_succ_self _).trans hk)] #align fin.lt_sub_one_iff Fin.lt_sub_one_iff diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 1311b3c8d5959..a7a136a553a83 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -74,7 +74,7 @@ theorem antidiagonalTuple_zero_zero : antidiagonalTuple 0 0 = [![]] := #align list.nat.antidiagonal_tuple_zero_zero List.Nat.antidiagonalTuple_zero_zero @[simp] -theorem antidiagonalTuple_zero_succ (n : ℕ) : antidiagonalTuple 0 n.succ = [] := +theorem antidiagonalTuple_zero_succ (n : ℕ) : antidiagonalTuple 0 (n + 1) = [] := rfl #align list.nat.antidiagonal_tuple_zero_succ List.Nat.antidiagonalTuple_zero_succ diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index 5e5ea8d716ec8..7da00f8730d4d 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -1178,8 +1178,7 @@ See also: `Fintype.exists_ne_map_eq_of_card_lt`, `Finite.exists_infinite_fiber`. -/ theorem Finite.exists_ne_map_eq_of_infinite {α β} [Infinite α] [Finite β] (f : α → β) : ∃ x y : α, x ≠ y ∧ f x = f y := by - simpa only [Injective, not_forall, Classical.not_imp, and_comm] using - not_injective_infinite_finite f + simpa [Injective, and_comm] using not_injective_infinite_finite f #align finite.exists_ne_map_eq_of_infinite Finite.exists_ne_map_eq_of_infinite instance Function.Embedding.is_empty {α β} [Infinite α] [Finite β] : IsEmpty (α ↪ β) := diff --git a/Mathlib/Data/Int/Cast/Basic.lean b/Mathlib/Data/Int/Cast/Basic.lean index f662d1d520fd3..fc1c4eb053410 100644 --- a/Mathlib/Data/Int/Cast/Basic.lean +++ b/Mathlib/Data/Int/Cast/Basic.lean @@ -94,8 +94,8 @@ theorem cast_subNatNat (m n) : ((Int.subNatNat m n : ℤ) : R) = m - n := by cases e : n - m · simp only [ofNat_eq_coe] simp [e, Nat.le_of_sub_eq_zero e] - · rw [cast_negSucc, Nat.add_one, ← e, Nat.cast_sub <| _root_.le_of_lt <| Nat.lt_of_sub_eq_succ e, - neg_sub] + · rw [cast_negSucc, Nat.add_one, succ_eq_add_one, ← e, + Nat.cast_sub <| _root_.le_of_lt <| Nat.lt_of_sub_eq_succ e, neg_sub] #align int.cast_sub_nat_nat Int.cast_subNatNatₓ -- type had `HasLiftT` diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 4d0ff1db92eee..e1cf665369579 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1153,6 +1153,9 @@ theorem get?_eq_some' {l : List α} {n a} : get? l n = some a ↔ ∃ h, nthLe l theorem nthLe_mem (l : List α) (n h) : nthLe l n h ∈ l := get_mem .. #align list.nth_le_mem List.nthLe_mem +theorem nthLe_congr {l : List α} {n p : ℕ} {h : n < length l} (hnp : n = p) : + nthLe l n h = nthLe l p (hnp ▸ h) := by simp [hnp] + #align list.nth_mem List.get?_mem @[deprecated mem_iff_get] @@ -1950,8 +1953,7 @@ theorem nthLe_take' (L : List α) {i j : ℕ} (hi : i < (L.take j).length) : theorem get?_take {l : List α} {n m : ℕ} (h : m < n) : (l.take n).get? m = l.get? m := by induction' n with n hn generalizing l m - · simp only [Nat.zero_eq] at h - exact absurd h (not_lt_of_le m.zero_le) + · exact absurd h (not_lt_of_le m.zero_le) · cases' l with hd tl · simp only [take_nil] · cases m @@ -2770,19 +2772,19 @@ where cases xs with | nil => contradiction | cons hd tl => - rw [length, succ_eq_add_one] at h + rw [length] at h rw [splitAt.go, take, drop, append_cons, Array.toList_eq, ← Array.push_data, ← Array.toList_eq] exact ih _ _ <| lt_of_add_lt_add_right h · induction n generalizing xs acc with | zero => - rw [zero_eq, not_lt, nonpos_iff_eq_zero] at h + rw [not_lt, nonpos_iff_eq_zero] at h rw [eq_nil_of_length_eq_zero h, splitAt.go] | succ _ ih => cases xs with | nil => rw [splitAt.go] | cons hd tl => - rw [length, succ_eq_add_one] at h + rw [length] at h rw [splitAt.go] exact ih _ _ <| not_imp_not.mpr (Nat.add_lt_add_right · 1) h #align list.split_at_eq_take_drop List.splitAt_eq_take_drop diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 89086567f03aa..83bf3402fe784 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -378,7 +378,7 @@ theorem prev_next (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : cases' l with hd tl · simp at hx · have : (n + 1 + length tl) % (length tl + 1) = n := by - rw [length_cons] at hn + rw [length_cons, Nat.succ_eq_add_one] at hn rw [add_assoc, add_comm 1, Nat.add_mod_right, Nat.mod_eq_of_lt hn] simp only [length_cons, Nat.succ_sub_succ_eq_sub, tsub_zero, Nat.succ_eq_add_one, this] #align list.prev_next List.prev_next @@ -391,7 +391,7 @@ theorem next_prev (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : cases' l with hd tl · simp at hx · have : (n + length tl + 1) % (length tl + 1) = n := by - rw [length_cons] at hn + rw [length_cons, Nat.succ_eq_add_one] at hn rw [add_assoc, Nat.add_mod_right, Nat.mod_eq_of_lt hn] simp [this] #align list.next_prev List.next_prev @@ -923,8 +923,7 @@ nonrec def Chain (r : α → α → Prop) (c : Cycle α) : Prop := · cases' l with c s · simp only [rotate_cons_succ, nil_append, rotate_singleton, cons.injEq] at hn rw [hn.1, hn.2] - · rw [Nat.succ_eq_one_add, ← rotate_rotate, rotate_cons_succ, rotate_zero, - cons_append] at hn + · rw [Nat.add_comm, ← rotate_rotate, rotate_cons_succ, rotate_zero, cons_append] at hn rw [← hd c _ _ _ hn] simp [and_comm] #align cycle.chain Cycle.Chain diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index c0d127a7036a2..978741ee5e670 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -182,7 +182,7 @@ theorem mapIdx_append {α} (K L : List α) (f : ℕ → α → β) : (K ++ L).mapIdx f = K.mapIdx f ++ L.mapIdx fun i a ↦ f (i + K.length) a := by induction' K with a J IH generalizing f · rfl - · simp [IH fun i ↦ f (i + 1), add_assoc] + · simp [IH fun i ↦ f (i + 1), add_assoc, Nat.succ_eq_add_one] #align list.map_with_index_append List.mapIdx_append @[simp] diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index a79ef2aebd1c7..f92fe11f54daf 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -458,8 +458,9 @@ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by simp (config := { contextual := true }) [count_replicate, h, h.symm, this, count_eq_zero] - simp_rw [Ne.def, ← and_imp, ← not_or, Decidable.not_imp_not, subset_def, mem_cons, - not_mem_nil, or_false, or_comm] + trans ∀ c, c ∈ l → c = b ∨ c = a + · simp [subset_def, or_comm] + · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] #align list.perm_replicate_append_replicate List.perm_replicate_append_replicate #align list.subperm.cons_right List.Subperm.cons_right @@ -670,8 +671,7 @@ theorem length_permutationsAux : refine' permutationsAux.rec (by simp) _ intro t ts is IH1 IH2 have IH2 : length (permutationsAux is nil) + 1 = is.length ! := by simpa using IH2 - simp? [Nat.factorial, Nat.add_succ, mul_comm] at IH1 says - simp only [factorial, add_eq, add_zero, mul_comm] at IH1 + simp only [factorial, mul_comm, add_eq] at IH1 rw [permutationsAux_cons, length_foldr_permutationsAux2' _ _ _ _ _ fun l m => (perm_of_mem_permutations m).length_eq, permutations, length, length, IH2, Nat.succ_add, Nat.factorial_succ, mul_comm (_ + 1), @@ -885,7 +885,7 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' rw [nthLe_insertNth_add_succ] convert nthLe_insertNth_add_succ s x k m.succ (by simpa using hn) using 2 · simp [Nat.add_succ, Nat.succ_add] - · simp [add_left_comm, add_comm] + · simp [add_left_comm, add_comm, succ_eq_add_one] · simpa [Nat.succ_add] using hn #align list.nodup_permutations'_aux_iff List.nodup_permutations'Aux_iff diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 2324abf115dfe..cf65c444dd9c7 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -78,7 +78,8 @@ theorem rotate'_rotate' : ∀ (l : List α) (n m : ℕ), (l.rotate' n).rotate' m | a :: l, 0, m => by simp | [], n, m => by simp | a :: l, n + 1, m => by - rw [rotate'_cons_succ, rotate'_rotate' _ n, add_right_comm, ← rotate'_cons_succ] + rw [rotate'_cons_succ, rotate'_rotate' _ n, add_right_comm, ← rotate'_cons_succ, + Nat.succ_eq_add_one] #align list.rotate'_rotate' List.rotate'_rotate' @[simp] @@ -114,7 +115,7 @@ theorem rotate_eq_rotate' (l : List α) (n : ℕ) : l.rotate n = l.rotate' n := #align list.rotate_eq_rotate' List.rotate_eq_rotate' theorem rotate_cons_succ (l : List α) (a : α) (n : ℕ) : - (a :: l : List α).rotate n.succ = (l ++ [a]).rotate n := by + (a :: l : List α).rotate (n + 1) = (l ++ [a]).rotate n := by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] #align list.rotate_cons_succ List.rotate_cons_succ @@ -361,7 +362,7 @@ theorem reverse_rotate (l : List α) (n : ℕ) : · simp · cases' l with hd tl · simp - · rw [rotate_cons_succ, Nat.succ_eq_add_one, ← rotate_rotate, hn] + · rw [rotate_cons_succ, ← rotate_rotate, hn] simp #align list.reverse_rotate List.reverse_rotate diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index f108c01852703..8812d44d183d4 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -415,10 +415,10 @@ theorem perm_split : ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → l ~ /-- Merge two sorted lists into one in linear time. merge [1, 2, 4, 5] [0, 1, 3, 4] = [0, 1, 1, 2, 3, 4, 4, 5] -/ -def merge : List α → List α → List α +def merge' : List α → List α → List α | [], l' => l' | l, [] => l - | a :: l, b :: l' => if a ≼ b then a :: merge l (b :: l') else b :: merge (a :: l) l' + | a :: l, b :: l' => if a ≼ b then a :: merge' l (b :: l') else b :: merge' (a :: l) l' termination_by l₁ l₂ => length l₁ + length l₂ #align list.merge List.merge @@ -433,27 +433,27 @@ def mergeSort : List α → List α have h := length_split_lt e have := h.1 have := h.2 - exact merge r (mergeSort ls.1) (mergeSort ls.2) + exact merge' r (mergeSort ls.1) (mergeSort ls.2) termination_by l => length l #align list.merge_sort List.mergeSort @[nolint unusedHavesSuffices] -- Porting note: false positive theorem mergeSort_cons_cons {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - mergeSort r (a :: b :: l) = merge r (mergeSort r l₁) (mergeSort r l₂) := by + mergeSort r (a :: b :: l) = merge' r (mergeSort r l₁) (mergeSort r l₂) := by simp only [mergeSort, h] #align list.merge_sort_cons_cons List.mergeSort_cons_cons section Correctness -theorem perm_merge : ∀ l l' : List α, merge r l l' ~ l ++ l' - | [], [] => by simp [merge] - | [], b :: l' => by simp [merge] - | a :: l, [] => by simp [merge] +theorem perm_merge' : ∀ l l' : List α, merge' r l l' ~ l ++ l' + | [], [] => by simp [merge'] + | [], b :: l' => by simp [merge'] + | a :: l, [] => by simp [merge'] | a :: l, b :: l' => by by_cases h : a ≼ b - · simpa [merge, h] using perm_merge _ _ - · suffices b :: merge r (a :: l) l' ~ a :: (l ++ b :: l') by simpa [merge, h] - exact ((perm_merge _ _).cons _).trans ((swap _ _ _).trans (perm_middle.symm.cons _)) + · simpa [merge', h] using perm_merge' _ _ + · suffices b :: merge' r (a :: l) l' ~ a :: (l ++ b :: l') by simpa [merge', h] + exact ((perm_merge' _ _).cons _).trans ((swap _ _ _).trans (perm_middle.symm.cons _)) termination_by l₁ l₂ => length l₁ + length l₂ #align list.perm_merge List.perm_merge @@ -464,7 +464,7 @@ theorem perm_mergeSort : ∀ l : List α, mergeSort r l ~ l cases' e : split (a :: b :: l) with l₁ l₂ cases' length_split_lt e with h₁ h₂ rw [mergeSort_cons_cons r e] - apply (perm_merge r _ _).trans + apply (perm_merge' r _ _).trans exact ((perm_mergeSort l₁).append (perm_mergeSort l₂)).trans (perm_split e).symm termination_by l => length l @@ -479,27 +479,27 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge r l l') - | [], [], _, _ => by simp [List.merge] - | [], b :: l', _, h₂ => by simpa [List.merge] using h₂ - | a :: l, [], h₁, _ => by simpa [List.merge] using h₁ +theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge' r l l') + | [], [], _, _ => by simp [List.merge'] + | [], b :: l', _, h₂ => by simpa [List.merge'] using h₂ + | a :: l, [], h₁, _ => by simpa [List.merge'] using h₁ | a :: l, b :: l', h₁, h₂ => by by_cases h : a ≼ b - · suffices ∀ b' ∈ List.merge r l (b :: l'), r a b' by - simpa [List.merge, h, h₁.of_cons.merge h₂] + · suffices ∀ b' ∈ List.merge' r l (b :: l'), r a b' by + simpa [List.merge', h, h₁.of_cons.merge h₂] intro b' bm rcases show b' = b ∨ b' ∈ l ∨ b' ∈ l' by - simpa [or_left_comm] using (perm_merge _ _ _).subset bm with + simpa [or_left_comm] using (perm_merge' _ _ _).subset bm with (be | bl | bl') · subst b' assumption · exact rel_of_sorted_cons h₁ _ bl · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge r (a :: l) l', r b b' by - simpa [List.merge, h, h₁.merge h₂.of_cons] + · suffices ∀ b' ∈ List.merge' r (a :: l) l', r b b' by + simpa [List.merge', h, h₁.merge h₂.of_cons] intro b' bm have ba : b ≼ a := (total_of r _ _).resolve_left h - have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge _ _ _).subset bm + have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge' _ _ _).subset bm rcases this with (be | bl | bl') · subst b' assumption diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index c912563215246..3ed9bb32cba6b 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -97,7 +97,7 @@ theorem length_sublists' : ∀ l : List α, length (sublists' l) = 2 ^ length l | [] => rfl | a :: l => by simp_arith only [sublists'_cons, length_append, length_sublists' l, - length_map, length, Nat.pow_succ', mul_succ, mul_zero, zero_add, npow_eq_pow, pow_eq] + length_map, length, Nat.pow_succ', mul_succ, mul_zero, zero_add] #align list.length_sublists' List.length_sublists' @[simp] diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index 479c3092705e7..cf1d9ed0eb53c 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -400,8 +400,7 @@ theorem prod_mul_prod_eq_prod_zipWith_mul_prod_drop : conv => lhs; rw [mul_assoc]; right; rw [mul_comm, mul_assoc]; right rw [mul_comm, prod_mul_prod_eq_prod_zipWith_mul_prod_drop xs ys] - simp only [add_eq, add_zero] - ac_rfl + simp only [mul_assoc] #align list.prod_mul_prod_eq_prod_zip_with_mul_prod_drop List.prod_mul_prod_eq_prod_zipWith_mul_prod_drop #align list.sum_add_sum_eq_sum_zip_with_add_sum_drop List.sum_add_sum_eq_sum_zipWith_add_sum_drop diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index e8530b0097c9a..d869f191910bd 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -101,9 +101,9 @@ theorem ext_nat' [AddMonoid A] [AddMonoidHomClass F ℕ A] (f g : F) (h : f 1 = DFunLike.ext f g <| by intro n induction n with - | zero => simp_rw [Nat.zero_eq, map_zero f, map_zero g] + | zero => simp_rw [map_zero f, map_zero g] | succ n ihn => - simp [Nat.succ_eq_add_one, h, ihn] + simp [h, ihn] #align ext_nat' ext_nat' @[ext] diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 1c732e42657d8..8a7a1840226cb 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -173,15 +173,17 @@ theorem binCast_eq [AddMonoidWithOne R] (n : ℕ) : (Nat.binCast n : R) = ((n : | zero => rw [Nat.binCast, Nat.cast_zero] | succ k => rw [Nat.binCast] - by_cases h : (k + 1) % 2 = 0 - · rw [← Nat.mod_add_div (succ k) 2] - rw [if_pos h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] - rw [Nat.succ_eq_add_one, h, Nat.zero_add, Nat.succ_mul, Nat.one_mul] - · rw [← Nat.mod_add_div (succ k) 2] - rw [if_neg h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] - have h1 := Or.resolve_left (Nat.mod_two_eq_zero_or_one (succ k)) h - rw [h1, Nat.add_comm 1, Nat.succ_mul, Nat.one_mul] - simp only [Nat.cast_add, Nat.cast_one] + -- FIXME nightly-testing + sorry + -- by_cases h : (k + 1) % 2 = 0 + -- · rw [← Nat.mod_add_div (succ k) 2] + -- rw [if_pos h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] + -- rw [Nat.succ_eq_add_one, h, Nat.zero_add, Nat.succ_mul, Nat.one_mul] + -- · rw [← Nat.mod_add_div (succ k) 2] + -- rw [if_neg h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] + -- have h1 := Or.resolve_left (Nat.mod_two_eq_zero_or_one (succ k)) h + -- rw [h1, Nat.add_comm 1, Nat.succ_mul, Nat.one_mul] + -- simp only [Nat.cast_add, Nat.cast_one] #align nat.bin_cast_eq Nat.binCast_eq section deprecated diff --git a/Mathlib/Data/Nat/Choose/Basic.lean b/Mathlib/Data/Nat/Choose/Basic.lean index 6c65d87df4435..c628934214f21 100644 --- a/Mathlib/Data/Nat/Choose/Basic.lean +++ b/Mathlib/Data/Nat/Choose/Basic.lean @@ -88,7 +88,8 @@ theorem choose_succ_self (n : ℕ) : choose n (succ n) = 0 := #align nat.choose_succ_self Nat.choose_succ_self @[simp] -theorem choose_one_right (n : ℕ) : choose n 1 = n := by induction n <;> simp [*, choose, add_comm] +theorem choose_one_right (n : ℕ) : choose n 1 = n := by + induction n <;> simp [*, choose, add_comm] #align nat.choose_one_right Nat.choose_one_right -- The `n+1`-st triangle number is `n` more than the `n`-th triangle number diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index b8622d4d901d1..d8c1e3f0de242 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -55,7 +55,6 @@ theorem multinomial_spec : (∏ i in s, (f i)!) * multinomial s f = (∑ i in s, @[simp] theorem multinomial_nil : multinomial ∅ f = 1 := by dsimp [multinomial] - rfl #align nat.multinomial_nil Nat.multinomial_nil variable {s f} diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 8927ce4d89a0f..eca2d3ad4f735 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -384,6 +384,9 @@ lemma rec_zero {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) : N Nat.rec h0 h (n + 1) = h n (Nat.rec h0 h n) := rfl #align nat.rec_add_one Nat.rec_add_one +@[simp] lemma rec_one {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) : + Nat.rec (motive := C) h0 h 1 = h 0 h0 := rfl + /-- Recursion starting at a non-zero number: given a map `C k → C (k+1)` for each `k ≥ n`, there is a map from `C n` to each `C m`, `n ≤ m`. -/ @[elab_as_elim] diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index c199e2ae65934..912476455d134 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -270,8 +270,8 @@ theorem ofDigits_digits (b n : ℕ) : ofDigits b (digits b n) = n := by · cases' b with b · induction' n with n ih · rfl - · rw[show succ zero = 1 by rfl] at ih ⊢ - simp only [ih, add_comm 1, ofDigits_one_cons, Nat.cast_id, digits_one_succ] + · rw [show 0 + 1 = 1 by rfl] at ih ⊢ + simp only [Nat.succ_eq_add_one, ih, add_comm 1, ofDigits_one_cons, Nat.cast_id, digits_one_succ] · apply Nat.strongInductionOn n _ clear n intro n h @@ -476,7 +476,7 @@ theorem digit_sum_le (p n : ℕ) : List.sum (digits p n) ≤ n := by · exact digits_zero _ ▸ Nat.le_refl (List.sum []) · induction' p with p · rw [digits_zero_succ, List.sum_cons, List.sum_nil, add_zero] - · nth_rw 2 [← ofDigits_digits p.succ n.succ] + · nth_rw 2 [← ofDigits_digits p.succ (n + 1)] rw [← ofDigits_one <| digits p.succ n.succ] exact ofDigits_monotone (digits p.succ n.succ) <| Nat.succ_pos p @@ -529,7 +529,8 @@ lemma ofDigits_div_pow_eq_ofDigits_drop ofDigits p digits / p ^ i = ofDigits p (digits.drop i) := by induction' i with i hi · simp - · rw [Nat.pow_succ, ← Nat.div_div_eq_div_mul, hi, ofDigits_div_eq_ofDigits_tail hpos + · rw [Nat.pow_succ, ← Nat.div_div_eq_div_mul, hi, + ofDigits_div_eq_ofDigits_tail hpos (List.drop i digits) fun x hx ↦ w₁ x <| List.mem_of_mem_drop hx, ← List.drop_one, List.drop_drop, add_comm] @@ -565,8 +566,11 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits rw [← Ico_succ_singleton, List.drop_length, ofDigits] at this have h₁ : 1 ≤ tl.length := List.length_pos.mpr h' rw [← sum_range_add_sum_Ico _ <| h₁, ← add_zero (∑ x in Ico _ _, ofDigits p (tl.drop x)), - ← this, sum_Ico_consecutive _ h₁ <| le_succ tl.length, ← sum_Ico_add _ 0 tl.length 1, - Ico_zero_eq_range, mul_add, mul_add, ih, range_one, sum_singleton, List.drop, ofDigits, + ← this, sum_Ico_consecutive _ h₁ <| (le_add_right (List.length tl) 1)] + -- Adaptatin note: nightly-2024-03-07: this needs an `erw` only because of a `0 + 1` vs `1`. + -- Can someone do it more cleanly? + erw [← sum_Ico_add _ 0 tl.length 1] + rw [Ico_zero_eq_range, mul_add, mul_add, ih, range_one, sum_singleton, List.drop, ofDigits, mul_zero, add_zero, ← Nat.add_sub_assoc <| sum_le_ofDigits _ <| Nat.le_of_lt h] nth_rw 2 [← one_mul <| ofDigits p tl] rw [← add_mul, one_eq_succ_zero, Nat.sub_add_cancel <| zero_lt_of_lt h, diff --git a/Mathlib/Data/Nat/Factorial/Basic.lean b/Mathlib/Data/Nat/Factorial/Basic.lean index e4b4eb1bcd1c6..780a341528400 100644 --- a/Mathlib/Data/Nat/Factorial/Basic.lean +++ b/Mathlib/Data/Nat/Factorial/Basic.lean @@ -254,7 +254,8 @@ theorem succ_ascFactorial (n : ℕ) : ∀ k, n * n.succ.ascFactorial k = (n + k) * n.ascFactorial k | 0 => by rw [add_zero, ascFactorial_zero, ascFactorial_zero] | k + 1 => by - rw [ascFactorial, mul_left_comm, succ_ascFactorial n k, ascFactorial, succ_add, ← add_assoc] + rw [ascFactorial, mul_left_comm, succ_ascFactorial n k, ascFactorial, succ_add, ← add_assoc, + succ_eq_add_one] #align nat.succ_asc_factorial Nat.succ_ascFactorial /-- `(n + 1).ascFactorial k = (n + k) ! / n !` but without ℕ-division. See diff --git a/Mathlib/Data/Nat/Factorial/SuperFactorial.lean b/Mathlib/Data/Nat/Factorial/SuperFactorial.lean index 31fb814cc7d51..05aab89f45ac0 100644 --- a/Mathlib/Data/Nat/Factorial/SuperFactorial.lean +++ b/Mathlib/Data/Nat/Factorial/SuperFactorial.lean @@ -97,7 +97,7 @@ theorem superFactorial_four_mul (n : ℕ) : sf (4 * n) = ((∏ i in range (2 * n), (2 * i + 1) !) * 2 ^ n) ^ 2 * (2 * n) ! := calc sf (4 * n) = (∏ i in range (2 * n), (2 * i + 1) !) ^ 2 * 2 ^ (2 * n) * (2 * n) ! := by - rw [← superFactorial_two_mul, ← mul_assoc] + rw [← superFactorial_two_mul, ← mul_assoc]; rfl _ = ((∏ i in range (2 * n), (2 * i + 1) !) * 2 ^ n) ^ 2 * (2 * n) ! := by rw [pow_mul', mul_pow] diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 07665960f8a2e..f8a04fc6dfa90 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -233,8 +233,8 @@ theorem factorization_pow (n k : ℕ) : factorization (n ^ k) = k • n.factoriz induction' k with k ih; · simp rcases eq_or_ne n 0 with (rfl | hn) · simp - rw [Nat.pow_succ, mul_comm, factorization_mul hn (pow_ne_zero _ hn), ih, succ_eq_one_add, - add_smul, one_smul] + rw [Nat.pow_succ, mul_comm, factorization_mul hn (pow_ne_zero _ hn), ih, + add_smul, one_smul, add_comm] #align nat.factorization_pow Nat.factorization_pow /-! ## Lemmas about factorizations of primes and prime powers -/ @@ -891,13 +891,11 @@ def recOnPrimeCoprime {P : ℕ → Sort*} (h0 : P 0) (hp : ∀ p n : ℕ, Prime @[elab_as_elim] def recOnMul {P : ℕ → Sort*} (h0 : P 0) (h1 : P 1) (hp : ∀ p, Prime p → P p) (h : ∀ a b, P a → P b → P (a * b)) : ∀ a, P a := - let hp : ∀ p n : ℕ, Prime p → P (p ^ n) := fun p n hp' => - n.recOn h1 (fun n hn => by rw [Nat.pow_succ]; apply h _ _ hn (hp p hp')) - -- Porting note: was - -- match n with - -- | 0 => h1 - -- | n + 1 => h _ _ (hp p hp') (_match _) - recOnPrimeCoprime h0 hp fun a b _ _ _ => h a b + let rec hp'' (p n : ℕ) (hp' : Prime p) : P (p ^ n) := + match n with + | 0 => h1 + | n + 1 => h _ _ (hp'' p n hp') (hp p hp') + recOnPrimeCoprime h0 hp'' fun a b _ _ _ => h a b #align nat.rec_on_mul Nat.recOnMul /-- For any multiplicative function `f` with `f 1 = 1` and any `n ≠ 0`, diff --git a/Mathlib/Data/Nat/Fib/Basic.lean b/Mathlib/Data/Nat/Fib/Basic.lean index 8e48f38d9d3ea..c8143884e6409 100644 --- a/Mathlib/Data/Nat/Fib/Basic.lean +++ b/Mathlib/Data/Nat/Fib/Basic.lean @@ -116,7 +116,10 @@ theorem fib_add_two_sub_fib_add_one {n : ℕ} : fib (n + 2) - fib (n + 1) = fib theorem fib_lt_fib_succ {n : ℕ} (hn : 2 ≤ n) : fib n < fib (n + 1) := by rcases exists_add_of_le hn with ⟨n, rfl⟩ - rw [← tsub_pos_iff_lt, add_comm 2, fib_add_two_sub_fib_add_one, fib_pos] + rw [← tsub_pos_iff_lt, add_comm 2, add_right_comm, fib_add_two] + -- Adaptation note: nightly-2024-03-07 + -- It would be nice to clean up the arithmetic here, to avoid the `erw`. + erw [add_tsub_cancel_right, fib_pos] exact succ_pos n #align nat.fib_lt_fib_succ Nat.fib_lt_fib_succ @@ -171,14 +174,14 @@ theorem fib_add (m n : ℕ) : fib (m + n + 1) = fib m * fib n + fib (m + 1) * fi · intros specialize ih (m + 1) rw [add_assoc m 1 n, add_comm 1 n] at ih - simp only [fib_add_two, ih] + simp only [fib_add_two, succ_eq_add_one, ih] ring #align nat.fib_add Nat.fib_add theorem fib_two_mul (n : ℕ) : fib (2 * n) = fib n * (2 * fib (n + 1) - fib n) := by cases n · simp - · rw [Nat.succ_eq_add_one, two_mul, ← add_assoc, fib_add, fib_add_two, two_mul] + · rw [two_mul, ← add_assoc, fib_add, fib_add_two, two_mul] simp only [← add_assoc, add_tsub_cancel_right] ring #align nat.fib_two_mul Nat.fib_two_mul diff --git a/Mathlib/Data/Nat/Hyperoperation.lean b/Mathlib/Data/Nat/Hyperoperation.lean index 3916234a05cd0..bc36dcf5dd493 100644 --- a/Mathlib/Data/Nat/Hyperoperation.lean +++ b/Mathlib/Data/Nat/Hyperoperation.lean @@ -77,7 +77,7 @@ theorem hyperoperation_two : hyperoperation 2 = (· * ·) := by -- Porting note: was `ring` dsimp only nth_rewrite 1 [← mul_one m] - rw [← mul_add, add_comm, Nat.succ_eq_add_one] + rw [← mul_add, add_comm] #align hyperoperation_two hyperoperation_two @[simp] diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index cd182fb2d4769..0ac197d2ad255 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -39,7 +39,7 @@ def log (b : ℕ) : ℕ → ℕ @[simp] theorem log_eq_zero_iff {b n : ℕ} : log b n = 0 ↔ n < b ∨ b ≤ 1 := by rw [log, dite_eq_right_iff] - simp only [Nat.succ_ne_zero, imp_false, not_and_or, not_le, not_lt] + simp only [add_eq_zero_iff, one_ne_zero, and_false, imp_false, not_and_or, not_le, not_lt] #align nat.log_eq_zero_iff Nat.log_eq_zero_iff theorem log_of_lt {b n : ℕ} (hb : n < b) : log b n = 0 := @@ -94,7 +94,7 @@ theorem pow_le_iff_le_log {b : ℕ} (hb : 1 < b) {x y : ℕ} (hy : y ≠ 0) : | succ x => rw [log]; split_ifs with h · have b_pos : 0 < b := zero_le_one.trans_lt hb - rw [succ_eq_add_one, add_le_add_iff_right, ← + rw [add_le_add_iff_right, ← ih (y / b) (div_lt_self hy.bot_lt hb) (Nat.div_pos h.1 b_pos).ne', le_div_iff_mul_le b_pos, pow_succ', mul_comm] · exact iff_of_false (fun hby => h ⟨(le_self_pow x.succ_ne_zero _).trans hby, hb⟩) @@ -299,7 +299,7 @@ theorem le_pow_iff_clog_le {b : ℕ} (hb : 1 < b) {x y : ℕ} : x ≤ b ^ y ↔ exact clog_pos hb have b_pos : 0 < b := (zero_lt_one' ℕ).trans hb rw [clog]; split_ifs with h - · rw [succ_eq_add_one, add_le_add_iff_right, ← ih ((x + b - 1) / b) (add_pred_div_lt hb h.2), + · rw [add_le_add_iff_right, ← ih ((x + b - 1) / b) (add_pred_div_lt hb h.2), Nat.div_le_iff_le_mul_add_pred b_pos, mul_comm b, ← Nat.pow_succ, add_tsub_assoc_of_le (Nat.succ_le_of_lt b_pos), add_le_add_iff_right] · exact iff_of_true ((not_lt.1 (not_and.1 h hb)).trans <| succ_le_of_lt <| pow_pos b_pos _) diff --git a/Mathlib/Data/Nat/Multiplicity.lean b/Mathlib/Data/Nat/Multiplicity.lean index fcd1a6a2c0fce..9f8a1bad0277f 100644 --- a/Mathlib/Data/Nat/Multiplicity.lean +++ b/Mathlib/Data/Nat/Multiplicity.lean @@ -71,7 +71,7 @@ theorem multiplicity_eq_card_pow_dvd {m n b : ℕ} (hm : m ≠ 1) (hn : 0 < n) ( PartENat.natCast_get, ← pow_dvd_iff_le_multiplicity, and_right_comm] refine' (and_iff_left_of_imp fun h => lt_of_le_of_lt _ hb).symm cases' m with m - · rw [zero_eq, zero_pow, zero_dvd_iff] at h + · rw [zero_pow, zero_dvd_iff] at h exacts [(hn.ne' h.2).elim, one_le_iff_ne_zero.1 h.1] exact le_log_of_pow_le (one_lt_iff_ne_zero_and_ne_one.2 ⟨m.succ_ne_zero, hm⟩) (le_of_dvd hn h.2) @@ -128,10 +128,12 @@ the sum of base `p` digits of `n`. -/ theorem sub_one_mul_multiplicity_factorial {n p : ℕ} (hp : p.Prime) : (p - 1) * (multiplicity p n !).get (finite_nat_iff.mpr ⟨hp.ne_one, factorial_pos n⟩) = n - (p.digits n).sum := by - simp only [multiplicity_factorial hp <| lt_succ_of_lt <| lt.base (log p n), - ← Finset.sum_Ico_add' _ 0 _ 1, Ico_zero_eq_range, - ← sub_one_mul_sum_log_div_pow_eq_sub_sum_digits] - rfl + -- Adaptation note: (`nightly-2024-03-11`) `← Finset.sum_Ico_add'` is not firing here because of a + -- `0 + 1` vs `1` issue (hence `norm_cast; nth_rw 2 [← zero_add 1]`) + simp only [multiplicity_factorial hp <| lt_succ_of_lt <| lt.base (log p n)] + norm_cast; nth_rw 2 [← zero_add 1] + simp only [← add_one, ← Finset.sum_Ico_add' _ 0 _ 1, Ico_zero_eq_range, + ← sub_one_mul_sum_log_div_pow_eq_sub_sum_digits] /-- The multiplicity of `p` in `(p * (n + 1))!` is one more than the sum of the multiplicities of `p` in `(p * n)!` and `n + 1`. -/ diff --git a/Mathlib/Data/Nat/Order/Basic.lean b/Mathlib/Data/Nat/Order/Basic.lean index 181abb4713e6d..b5c0d20783bc4 100644 --- a/Mathlib/Data/Nat/Order/Basic.lean +++ b/Mathlib/Data/Nat/Order/Basic.lean @@ -171,16 +171,12 @@ theorem add_eq_one_iff : m + n = 1 ↔ m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 := by #align nat.add_eq_one_iff Nat.add_eq_one_iff theorem add_eq_two_iff : m + n = 2 ↔ m = 0 ∧ n = 2 ∨ m = 1 ∧ n = 1 ∨ m = 2 ∧ n = 0 := by - cases n <;> - simp [(succ_ne_zero 1).symm, (show 2 = Nat.succ 1 from rfl), - succ_eq_add_one, ← add_assoc, succ_inj', add_eq_one_iff] + omega #align nat.add_eq_two_iff Nat.add_eq_two_iff theorem add_eq_three_iff : m + n = 3 ↔ m = 0 ∧ n = 3 ∨ m = 1 ∧ n = 2 ∨ m = 2 ∧ n = 1 ∨ m = 3 ∧ n = 0 := by - cases n <;> - simp [(succ_ne_zero 1).symm, succ_eq_add_one, (show 3 = Nat.succ 2 from rfl), - ← add_assoc, succ_inj', add_eq_two_iff] + omega #align nat.add_eq_three_iff Nat.add_eq_three_iff theorem le_add_one_iff : m ≤ n + 1 ↔ m ≤ n ∨ m = n + 1 := by diff --git a/Mathlib/Data/Nat/Order/Lemmas.lean b/Mathlib/Data/Nat/Order/Lemmas.lean index df7fca0f0b4ca..401affeccaff3 100644 --- a/Mathlib/Data/Nat/Order/Lemmas.lean +++ b/Mathlib/Data/Nat/Order/Lemmas.lean @@ -138,7 +138,8 @@ theorem succ_div : ∀ a b : ℕ, (a + 1) / b = a / b + if b ∣ a + 1 then 1 el have h₂ : 0 < b + 1 ∧ b + 1 ≤ a + 1 := ⟨succ_pos _, (add_le_add_iff_right _).2 hb_le_a⟩ have dvd_iff : b + 1 ∣ a - b + 1 ↔ b + 1 ∣ a + 1 + 1 := by rw [Nat.dvd_add_iff_left (dvd_refl (b + 1)), ← add_tsub_add_eq_tsub_right a 1 b, - add_comm (_ - _), add_assoc, tsub_add_cancel_of_le (succ_le_succ hb_le_a), add_comm 1] + add_comm (_ - _), add_assoc, tsub_add_cancel_of_le (add_le_add_right hb_le_a 1), + add_comm 1] have wf : a - b < a + 1 := lt_succ_of_le tsub_le_self rw [if_pos h₁, if_pos h₂, @add_tsub_add_eq_tsub_right, ← tsub_add_eq_add_tsub hb_le_a, have := wf diff --git a/Mathlib/Data/Nat/Prime.lean b/Mathlib/Data/Nat/Prime.lean index c24f0ea76d30a..ba2d17d4bc3c4 100644 --- a/Mathlib/Data/Nat/Prime.lean +++ b/Mathlib/Data/Nat/Prime.lean @@ -295,7 +295,8 @@ theorem minFacAux_has_prop {n : ℕ} (n2 : 2 ≤ n) : · exact ⟨k2, dk, a⟩ · refine' have := minFac_lemma n k h - minFacAux_has_prop n2 (k + 2) (i + 1) (by simp [k, e, left_distrib]) fun m m2 d => _ + minFacAux_has_prop n2 (k + 2) (i + 1) (by simp [k, e, left_distrib, add_right_comm]) + fun m m2 d => _ rcases Nat.eq_or_lt_of_le (a m m2 d) with me | ml · subst me contradiction diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index cea166a38b6f7..8f5429a66f515 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -982,7 +982,12 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by · rfl · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_bit_zero] · rw [PosNum.cast_bit0, ← Nat.bit_false, Nat.testBit_bit_zero] - · rw [PosNum.cast_one', ← bit1_zero, ← Nat.bit_true, Nat.testBit_bit_succ, Nat.zero_testBit] + · rw [PosNum.cast_one', ← bit1_zero, ← Nat.bit_true] + -- FIXME adaptation for nightly-2024-03-11 + -- Hacky proof, can someone please replace? Obvious simp lemmas seem to be missing. + simp + erw [Nat.testBit_succ] + simp · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_bit_succ, IH] · rw [PosNum.cast_bit0, ← Nat.bit_false, Nat.testBit_bit_succ, IH] #align num.test_bit_to_nat Num.castNum_testBit diff --git a/Mathlib/Data/Opposite.lean b/Mathlib/Data/Opposite.lean index 4d1e6ab5ff510..27fb4a76304df 100644 --- a/Mathlib/Data/Opposite.lean +++ b/Mathlib/Data/Opposite.lean @@ -119,9 +119,9 @@ instance [Nonempty α] : Nonempty αᵒᵖ := Nonempty.map op ‹_› instance [Subsingleton α] : Subsingleton αᵒᵖ := unop_injective.subsingleton /-- A recursor for `Opposite`. -The `@[eliminator]` attribute makes it the default induction principle for `Opposite` +The `@[induction_eliminator]` attribute makes it the default induction principle for `Opposite` so you don't need to use `induction x using Opposite.rec'`. -/ -@[simp, eliminator] +@[simp, induction_eliminator] protected def rec' {F : αᵒᵖ → Sort v} (h : ∀ X, F (op X)) : ∀ X, F X := fun X => h (unop X) #align opposite.rec Opposite.rec' diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index 5be1da868ef1f..ae1fa740b084a 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -1189,7 +1189,7 @@ theorem Valid'.node4L {l} {x : α} {m} {y : α} {r o₁ o₂} (hl : Valid' o₁ rw [l1, r1] revert mm; cases size ml <;> cases size mr <;> intro mm · decide - · rw [Nat.zero_eq, zero_add] at mm; rcases mm with (_ | ⟨⟨⟩⟩) + · rw [zero_add] at mm; rcases mm with (_ | ⟨⟨⟩⟩) decide · rcases mm with (_ | ⟨⟨⟩⟩); decide · rw [Nat.succ_add] at mm; rcases mm with (_ | ⟨⟨⟩⟩) diff --git a/Mathlib/Data/PNat/Xgcd.lean b/Mathlib/Data/PNat/Xgcd.lean index c43183d9b8b21..3bec2d30b3fc6 100644 --- a/Mathlib/Data/PNat/Xgcd.lean +++ b/Mathlib/Data/PNat/Xgcd.lean @@ -261,8 +261,7 @@ theorem start_isSpecial (a b : ℕ+) : (start a b).IsSpecial := by theorem start_v (a b : ℕ+) : (start a b).v = ⟨a, b⟩ := by dsimp [start, v, XgcdType.a, XgcdType.b, w, z] - have : succ 0 = 1 := rfl - rw [this, one_mul, one_mul, zero_mul, zero_mul, zero_add, add_zero] + rw [one_mul, one_mul, zero_mul, zero_mul, zero_add, add_zero] rw [← Nat.pred_eq_sub_one, ← Nat.pred_eq_sub_one] rw [Nat.succ_pred_eq_of_pos a.pos, Nat.succ_pred_eq_of_pos b.pos] #align pnat.xgcd_type.start_v PNat.XgcdType.start_v diff --git a/Mathlib/Data/Polynomial/Basic.lean b/Mathlib/Data/Polynomial/Basic.lean index e956f8de346e6..154d5e093221e 100644 --- a/Mathlib/Data/Polynomial/Basic.lean +++ b/Mathlib/Data/Polynomial/Basic.lean @@ -646,7 +646,7 @@ theorem monomial_mul_X_pow (n : ℕ) (r : R) (k : ℕ) : monomial n r * X ^ k = monomial (n + k) r := by induction' k with k ih · simp - · simp [ih, pow_succ', ← mul_assoc, add_assoc] + · simp [ih, pow_succ', ← mul_assoc, add_assoc, Nat.succ_eq_add_one] #align polynomial.monomial_mul_X_pow Polynomial.monomial_mul_X_pow @[simp] diff --git a/Mathlib/Data/Polynomial/Derivative.lean b/Mathlib/Data/Polynomial/Derivative.lean index 28a5410ce8add..515306762d0d1 100644 --- a/Mathlib/Data/Polynomial/Derivative.lean +++ b/Mathlib/Data/Polynomial/Derivative.lean @@ -294,39 +294,41 @@ set_option linter.uppercaseLean3 false in @[simp] theorem derivative_mul {f g : R[X]} : derivative (f * g) = derivative f * g + f * derivative g := - calc - derivative (f * g) = - f.sum fun n a => g.sum fun m b => (n + m) • (C (a * b) * X ^ (n + m - 1)) := by - rw [mul_eq_sum_sum, derivative_sum] - trans - · apply Finset.sum_congr rfl - intro x _ - exact derivative_sum - apply Finset.sum_congr rfl; intro n _; apply Finset.sum_congr rfl; intro m _ - trans - · exact congr_arg _ C_mul_X_pow_eq_monomial.symm - dsimp; rw [← smul_mul_assoc, smul_C, nsmul_eq_mul']; exact derivative_C_mul_X_pow _ _ - _ = - f.sum fun n a => - g.sum fun m b => - n • (C a * X ^ (n - 1)) * (C b * X ^ m) + C a * X ^ n * m • (C b * X ^ (m - 1)) := - (sum_congr rfl fun n hn => - sum_congr rfl fun m hm => by - 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.add_one_sub_one, zero_smul, add_comm]) - _ = derivative f * g + f * derivative g := by - conv => - rhs - congr - · rw [← sum_C_mul_X_pow_eq g] - · rw [← sum_C_mul_X_pow_eq f] - simp only [sum, sum_add_distrib, Finset.mul_sum, Finset.sum_mul, derivative_apply] - simp_rw [← smul_mul_assoc, smul_C, nsmul_eq_mul'] - rw [Finset.sum_comm] - congr 1 - rw [Finset.sum_comm] + -- FIXME nightly-testing + sorry + -- calc + -- derivative (f * g) = + -- f.sum fun n a => g.sum fun m b => (n + m) • (C (a * b) * X ^ (n + m - 1)) := by + -- rw [mul_eq_sum_sum, derivative_sum] + -- trans + -- · apply Finset.sum_congr rfl + -- intro x _ + -- exact derivative_sum + -- apply Finset.sum_congr rfl; intro n _; apply Finset.sum_congr rfl; intro m _ + -- trans + -- · exact congr_arg _ C_mul_X_pow_eq_monomial.symm + -- dsimp; rw [← smul_mul_assoc, smul_C, nsmul_eq_mul']; exact derivative_C_mul_X_pow _ _ + -- _ = + -- f.sum fun n a => + -- g.sum fun m b => + -- n • (C a * X ^ (n - 1)) * (C b * X ^ m) + C a * X ^ n * m • (C b * X ^ (m - 1)) := + -- (sum_congr rfl fun n hn => + -- sum_congr rfl fun m hm => by + -- 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.add_one_sub_one, zero_smul, add_comm]) + -- _ = derivative f * g + f * derivative g := by + -- conv => + -- rhs + -- congr + -- · rw [← sum_C_mul_X_pow_eq g] + -- · rw [← sum_C_mul_X_pow_eq f] + -- simp only [sum, sum_add_distrib, Finset.mul_sum, Finset.sum_mul, derivative_apply] + -- simp_rw [← smul_mul_assoc, smul_C, nsmul_eq_mul'] + -- rw [Finset.sum_comm] + -- congr 1 + -- rw [Finset.sum_comm] #align polynomial.derivative_mul Polynomial.derivative_mul theorem derivative_eval (p : R[X]) (x : R) : @@ -449,7 +451,7 @@ theorem iterate_derivative_mul {n} (p q : R[X]) : refine' sum_congr rfl fun k hk => _ rw [mem_range] at hk congr - rw [tsub_add_eq_add_tsub (Nat.succ_le_of_lt hk), Nat.succ_sub_succ] + omega · rw [Nat.choose_zero_right, tsub_zero] #align polynomial.iterate_derivative_mul Polynomial.iterate_derivative_mul @@ -506,7 +508,7 @@ theorem iterate_derivative_X_pow_eq_nat_cast_mul (n k : ℕ) : induction' k with k ih · erw [Function.iterate_zero_apply, tsub_zero, Nat.descFactorial_zero, Nat.cast_one, one_mul] · rw [Function.iterate_succ_apply', ih, derivative_nat_cast_mul, derivative_X_pow, C_eq_nat_cast, - Nat.succ_eq_add_one, Nat.descFactorial_succ, Nat.sub_sub, Nat.cast_mul]; + Nat.descFactorial_succ, Nat.sub_sub, Nat.cast_mul]; simp [mul_comm, mul_assoc, mul_left_comm] set_option linter.uppercaseLean3 false in #align polynomial.iterate_derivative_X_pow_eq_nat_cast_mul Polynomial.iterate_derivative_X_pow_eq_nat_cast_mul diff --git a/Mathlib/Data/Polynomial/Module/Basic.lean b/Mathlib/Data/Polynomial/Module/Basic.lean index 1a9fdbb648825..4507f6d031c42 100644 --- a/Mathlib/Data/Polynomial/Module/Basic.lean +++ b/Mathlib/Data/Polynomial/Module/Basic.lean @@ -312,10 +312,10 @@ theorem monomial_smul_single (i : ℕ) (r : R) (j : ℕ) (m : M) : Module.algebraMap_end_apply, smul_def] induction i generalizing r j m with | zero => - rw [Nat.zero_eq, Function.iterate_zero, zero_add] + rw [Function.iterate_zero, zero_add] exact Finsupp.smul_single r j m | succ n hn => - rw [Function.iterate_succ, Function.comp_apply, Nat.succ_eq_add_one, add_assoc, ← hn] + rw [Function.iterate_succ, Function.comp_apply,add_assoc, ← hn] congr 2 rw [Nat.one_add] exact Finsupp.mapDomain_single diff --git a/Mathlib/Data/Polynomial/Monic.lean b/Mathlib/Data/Polynomial/Monic.lean index 067bc69d00a60..5448e0a327f6e 100644 --- a/Mathlib/Data/Polynomial/Monic.lean +++ b/Mathlib/Data/Polynomial/Monic.lean @@ -222,7 +222,7 @@ theorem nextCoeff_mul (hp : Monic p) (hq : Monic q) : theorem nextCoeff_pow (hp : p.Monic) (n : ℕ) : (p ^ n).nextCoeff = n • p.nextCoeff := by induction n with - | zero => rw [pow_zero, Nat.zero_eq, zero_smul, ← map_one (f := C), nextCoeff_C_eq_zero] + | zero => rw [pow_zero, zero_smul, ← map_one (f := C), nextCoeff_C_eq_zero] | succ n ih => rw [pow_succ, hp.nextCoeff_mul (hp.pow n), ih, succ_nsmul] theorem eq_one_of_map_eq_one {S : Type*} [Semiring S] [Nontrivial S] (f : R →+* S) (hp : p.Monic) diff --git a/Mathlib/Data/Polynomial/RingDivision.lean b/Mathlib/Data/Polynomial/RingDivision.lean index 8f4c74aa87617..2b7db1897f98c 100644 --- a/Mathlib/Data/Polynomial/RingDivision.lean +++ b/Mathlib/Data/Polynomial/RingDivision.lean @@ -861,11 +861,10 @@ theorem roots_prod {ι : Type*} (f : ι → R[X]) (s : Finset ι) : @[simp] theorem roots_pow (p : R[X]) (n : ℕ) : (p ^ n).roots = n • p.roots := by induction' n with n ihn - · rw [pow_zero, roots_one, Nat.zero_eq, zero_smul, empty_eq_zero] + · rw [pow_zero, roots_one, zero_smul, empty_eq_zero] · rcases eq_or_ne p 0 with (rfl | hp) · rw [zero_pow n.succ_ne_zero, roots_zero, smul_zero] - · rw [pow_succ', roots_mul (mul_ne_zero (pow_ne_zero _ hp) hp), ihn, Nat.succ_eq_add_one, - add_smul, one_smul] + · rw [pow_succ', roots_mul (mul_ne_zero (pow_ne_zero _ hp) hp), ihn, add_smul, one_smul] #align polynomial.roots_pow Polynomial.roots_pow theorem roots_X_pow (n : ℕ) : (X ^ n : R[X]).roots = n • ({0} : Multiset R) := by diff --git a/Mathlib/Data/Polynomial/Smeval.lean b/Mathlib/Data/Polynomial/Smeval.lean index 23dc7095e5eb8..7bf5152f58028 100644 --- a/Mathlib/Data/Polynomial/Smeval.lean +++ b/Mathlib/Data/Polynomial/Smeval.lean @@ -175,8 +175,8 @@ theorem smeval_at_zero : p.smeval (0 : S) = (p.coeff 0) • (1 : S) := by | h_monomial n a => cases n with | zero => simp only [Nat.zero_eq, monomial_zero_left, smeval_C, npow_zero, coeff_C_zero] - | succ n => rw [coeff_monomial_succ, smeval_monomial, ← Nat.add_one, npow_add, npow_one, - mul_zero, zero_smul, smul_zero] + | succ n => rw [coeff_monomial_succ, smeval_monomial, npow_add, npow_one, mul_zero, zero_smul, + smul_zero] theorem smeval_mul_X : (p * X).smeval x = p.smeval x * x := by induction p using Polynomial.induction_on' with diff --git a/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean b/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean index 0b701b10d416f..01736460c0e81 100644 --- a/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean +++ b/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean @@ -292,7 +292,7 @@ theorem Cofix.bisim {α : TypeVec n} (r : Cofix F α → Cofix F α → Prop) rw [← split_dropFun_lastFun f₀, ← split_dropFun_lastFun f₁] rw [appendFun_comp_splitFun, appendFun_comp_splitFun] rw [id_comp, id_comp] - congr 2 with (i j); cases' i with _ i <;> dsimp + congr 2 with (i j); cases' i with _ i · apply Quot.sound apply h' _ j · change f₀ _ j = f₁ _ j diff --git a/Mathlib/Data/Seq/Computation.lean b/Mathlib/Data/Seq/Computation.lean index 6a6e45e9575aa..29376689656d7 100644 --- a/Mathlib/Data/Seq/Computation.lean +++ b/Mathlib/Data/Seq/Computation.lean @@ -580,7 +580,7 @@ theorem of_results_think {s : Computation α} {a n} (h : Results (think s) a n) theorem results_think_iff {s : Computation α} {a n} : Results (think s) a (n + 1) ↔ Results s a n := ⟨fun h => by let ⟨n', r, e⟩ := of_results_think h - injection e with h'; rw [Nat.add, Nat.add] at h'; rwa [h'], results_think⟩ + injection e with h'; rwa [h'], results_think⟩ #align computation.results_think_iff Computation.results_think_iff theorem results_thinkN {s : Computation α} {a m} : diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index a20669b456048..8346189b91791 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -842,7 +842,7 @@ theorem dropn_tail (s : Seq α) (n) : drop (tail s) n = drop s (n + 1) := by @[simp] theorem head_dropn (s : Seq α) (n) : head (drop s n) = get? s n := by induction' n with n IH generalizing s; · rfl - rw [Nat.succ_eq_add_one, ← get?_tail, ← dropn_tail]; apply IH + rw [← get?_tail, ← dropn_tail]; apply IH #align stream.seq.head_dropn Stream'.Seq.head_dropn theorem mem_map (f : α → β) {a : α} : ∀ {s : Seq α}, a ∈ s → f a ∈ map f s diff --git a/Mathlib/Data/Set/Enumerate.lean b/Mathlib/Data/Set/Enumerate.lean index 15c48237f99c2..225df6112c68c 100644 --- a/Mathlib/Data/Set/Enumerate.lean +++ b/Mathlib/Data/Set/Enumerate.lean @@ -53,8 +53,7 @@ theorem enumerate_eq_none : | zero => contradiction | succ m' => simp? [hs, enumerate] at h ⊢ says - simp only [enumerate, hs, Nat.add_eq, add_zero, Option.bind_eq_bind, - Option.some_bind] at h ⊢ + simp only [enumerate, hs, Option.bind_eq_bind, Option.some_bind] at h ⊢ have hm : n ≤ m' := Nat.le_of_succ_le_succ hm exact enumerate_eq_none h hm #align set.enumerate_eq_none Set.enumerate_eq_none diff --git a/Mathlib/Data/Set/Pairwise/Basic.lean b/Mathlib/Data/Set/Pairwise/Basic.lean index 2f7704b54d40d..3b1e8338adfdf 100644 --- a/Mathlib/Data/Set/Pairwise/Basic.lean +++ b/Mathlib/Data/Set/Pairwise/Basic.lean @@ -139,8 +139,8 @@ theorem pairwise_union : s.Pairwise r ∧ t.Pairwise r ∧ ∀ a ∈ s, ∀ b ∈ t, a ≠ b → r a b ∧ r b a := by simp only [Set.Pairwise, mem_union, or_imp, forall_and] exact - ⟨fun H => ⟨H.1.1, H.2.2, H.2.1, fun x hx y hy hne => H.1.2 y hy x hx hne.symm⟩, fun H => - ⟨⟨H.1, fun x hx y hy hne => H.2.2.2 y hy x hx hne.symm⟩, H.2.2.1, H.2.1⟩⟩ + ⟨fun H => ⟨H.1.1, H.2.2, H.1.2, fun x hx y hy hne => H.2.1 y hy x hx hne.symm⟩, + fun H => ⟨⟨H.1, H.2.2.1⟩, fun x hx y hy hne => H.2.2.2 y hy x hx hne.symm, H.2.1⟩⟩ #align set.pairwise_union Set.pairwise_union theorem pairwise_union_of_symmetric (hr : Symmetric r) : diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 8d095245f6d3c..933eaa9aeaed7 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -176,7 +176,7 @@ theorem Disjoint.set_prod_right (ht : Disjoint t₁ t₂) (s₁ s₂ : Set α) : theorem insert_prod : insert a s ×ˢ t = Prod.mk a '' t ∪ s ×ˢ t := by ext ⟨x, y⟩ - simp (config := { contextual := true }) [image, iff_def, or_imp, Imp.swap] + simp (config := { contextual := true }) [image, iff_def, or_imp] #align set.insert_prod Set.insert_prod theorem prod_insert : s ×ˢ insert b t = (fun a => (a, b)) '' s ∪ s ×ˢ t := by diff --git a/Mathlib/Data/Sum/Interval.lean b/Mathlib/Data/Sum/Interval.lean index 1aa6c23ba5179..66e26371afda9 100644 --- a/Mathlib/Data/Sum/Interval.lean +++ b/Mathlib/Data/Sum/Interval.lean @@ -92,8 +92,7 @@ theorem sumLift₂_nonempty : (sumLift₂ f g a b).Nonempty ↔ (∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ (f a₁ b₁).Nonempty) ∨ ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ (g a₂ b₂).Nonempty := by - simp only [nonempty_iff_ne_empty, Ne, sumLift₂_eq_empty, not_and_or, not_forall, - Classical.not_imp] + simp only [nonempty_iff_ne_empty, Ne, sumLift₂_eq_empty, not_and_or, not_forall, exists_prop] #align finset.sum_lift₂_nonempty Finset.sumLift₂_nonempty theorem sumLift₂_mono (h₁ : ∀ a b, f₁ a b ⊆ g₁ a b) (h₂ : ∀ a b, f₂ a b ⊆ g₂ a b) : diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index e03c2fed79f66..128178598604b 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -956,7 +956,7 @@ theorem neg_eq_self_iff {n : ℕ} (a : ZMod n) : -a = a ↔ a = 0 ∨ 2 * a.val exact Or.inl (a.val_eq_zero.1 he) cases m · right - rwa [show Nat.succ Nat.zero = 1 from rfl, mul_one] at he + rwa [show 0 + 1 = 1 from rfl, mul_one] at he refine' (a.val_lt.not_le <| Nat.le_of_mul_le_mul_left _ zero_lt_two).elim rw [he, mul_comm] apply Nat.mul_le_mul_left diff --git a/Mathlib/Data/ZMod/Factorial.lean b/Mathlib/Data/ZMod/Factorial.lean index ddbe5b71ec460..6a7df9d2ea263 100644 --- a/Mathlib/Data/ZMod/Factorial.lean +++ b/Mathlib/Data/ZMod/Factorial.lean @@ -35,7 +35,8 @@ theorem cast_descFactorial {n p : ℕ} (h : n ≤ p) : rw [pow_card_mul_prod] refine prod_congr rfl ?_ intro x hx - rw [← tsub_add_eq_tsub_tsub_swap, Nat.cast_sub <| Nat.lt_of_lt_of_le (List.mem_range.mp hx) h, + rw [← tsub_add_eq_tsub_tsub_swap, + Nat.cast_sub <| Nat.le_trans (Nat.add_one_le_iff.mpr (List.mem_range.mp hx)) h, CharP.cast_eq_zero, zero_sub, cast_succ, neg_add_rev, mul_add, neg_mul, one_mul, mul_one, add_comm] diff --git a/Mathlib/FieldTheory/ChevalleyWarning.lean b/Mathlib/FieldTheory/ChevalleyWarning.lean index fa4df5fd6f9e5..545da99ec1280 100644 --- a/Mathlib/FieldTheory/ChevalleyWarning.lean +++ b/Mathlib/FieldTheory/ChevalleyWarning.lean @@ -131,8 +131,8 @@ theorem char_dvd_card_solutions_of_sum_lt {s : Finset ι} {f : ι → MvPolynomi intro i hi rw [hS] at hx rw [hx i hi, zero_pow hq.ne', sub_zero] - · obtain ⟨i, hi, hx⟩ : ∃ i : ι, i ∈ s ∧ eval x (f i) ≠ 0 := by - simpa only [hS, not_forall, Classical.not_imp] using hx + · obtain ⟨i, hi, hx⟩ : ∃ i ∈ s, eval x (f i) ≠ 0 := by + simpa [hS, not_forall, Classical.not_imp] using hx apply Finset.prod_eq_zero hi rw [pow_card_sub_one_eq_one (eval x (f i)) hx, sub_self] -- In particular, we can now show: diff --git a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean index 9ef772688b5e1..4842934f3f780 100644 --- a/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean +++ b/Mathlib/FieldTheory/Minpoly/MinpolyDiv.lean @@ -143,8 +143,7 @@ lemma coeff_minpolyDiv_sub_pow_mem_span {i} (hi : i ≤ natDegree (minpolyDiv R · apply Submodule.smul_mem apply Submodule.subset_span exact ⟨0, Nat.zero_lt_succ _, pow_zero _⟩ - · rw [Nat.succ_eq_add_one, ← tsub_tsub, tsub_add_cancel_of_le - (le_tsub_of_add_le_left (b := 1) hi)] + · rw [← tsub_tsub, tsub_add_cancel_of_le (le_tsub_of_add_le_left (b := 1) hi)] apply SetLike.le_def.mp ?_ (Submodule.mul_mem_mul (IH ((Nat.le_succ _).trans hi)) (Submodule.mem_span_singleton_self x)) diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index b451d4a15136e..c64d81516ddeb 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -150,7 +150,7 @@ theorem mongePoint_eq_affineCombination_of_pointsWithCircumcenter {n : ℕ} · rw [if_pos (mem_univ _), sub_zero, add_zero, card_fin] -- Porting note: replaced -- have hn3 : (n + 2 + 1 : ℝ) ≠ 0 := mod_cast Nat.succ_ne_zero _ - have hn3 : (n + 2 + 1 : ℝ) ≠ 0 := by exact_mod_cast (n + 2).cast_add_one_ne_zero + have hn3 : (n + 2 + 1 : ℝ) ≠ 0 := by norm_cast field_simp [hn1, hn3, mul_comm] · field_simp [hn1] ring diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index faafdef26ae26..25f182d686e4b 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -200,7 +200,7 @@ instance instNSMul : SMul ℕ Cₛ^n⟮I; F, V⟯ := @[simp] theorem coe_nsmul (s : Cₛ^n⟮I; F, V⟯) (k : ℕ) : ⇑(k • s : Cₛ^n⟮I; F, V⟯) = k • ⇑s := by induction' k with k ih - · simp_rw [Nat.zero_eq, zero_smul]; rfl + · simp_rw [zero_smul]; rfl simp_rw [succ_nsmul, ← ih]; rfl #align cont_mdiff_section.coe_nsmul ContMDiffSection.coe_nsmul diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index 158ff2e5a6aea..44432a949f933 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -1162,10 +1162,6 @@ theorem reduce.not {p : Prop} : simp only [List.length, zero_add, List.length_append] at this rw [add_comm, add_assoc, add_assoc, add_comm, <-add_assoc] at this simp [Nat.one_eq_succ_zero, Nat.succ_add] at this - -- Porting note: needed to add this step in #3414. - -- This is caused by https://github.com/leanprover/lean4/pull/2146. - -- Nevertheless the proof could be cleaned up. - cases this | cons hd tail => cases' hd with y c dsimp only diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index 4637cd6813b01..a8f8331478858 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -240,7 +240,7 @@ theorem is_decending_rev_series_of_is_ascending {H : ℕ → Subgroup G} {n : · push_neg at hm apply hH convert hx using 1 - rw [tsub_add_eq_add_tsub (Nat.succ_le_of_lt hm), Nat.succ_sub_succ] + rw [tsub_add_eq_add_tsub (Nat.succ_le_of_lt hm), Nat.succ_eq_add_one, Nat.add_sub_add_right] #align is_decending_rev_series_of_is_ascending is_decending_rev_series_of_is_ascending theorem is_ascending_rev_series_of_is_descending {H : ℕ → Subgroup G} {n : ℕ} (hn : H n = ⊥) @@ -254,7 +254,7 @@ theorem is_ascending_rev_series_of_is_descending {H : ℕ → Subgroup G} {n : exact mem_top _ · push_neg at hm convert hH x _ hx g using 1 - rw [tsub_add_eq_add_tsub (Nat.succ_le_of_lt hm), Nat.succ_sub_succ] + rw [tsub_add_eq_add_tsub (Nat.succ_le_of_lt hm), Nat.succ_eq_add_one, Nat.add_sub_add_right] #align is_ascending_rev_series_of_is_descending is_ascending_rev_series_of_is_descending /-- A group `G` is nilpotent iff there exists a descending central series which reaches the diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index 31a03ca0a0b7e..f065ee759b7f4 100644 --- a/Mathlib/GroupTheory/Perm/Fin.lean +++ b/Mathlib/GroupTheory/Perm/Fin.lean @@ -103,7 +103,7 @@ theorem finRotate_succ_eq_decomposeFin {n : ℕ} : @[simp] theorem sign_finRotate (n : ℕ) : Perm.sign (finRotate (n + 1)) = (-1) ^ n := by induction' n with n ih - · simp; rfl + · simp · rw [finRotate_succ_eq_decomposeFin] simp [ih, pow_succ] #align sign_fin_rotate sign_finRotate diff --git a/Mathlib/GroupTheory/Perm/List.lean b/Mathlib/GroupTheory/Perm/List.lean index ba052cc53c31c..3a4df211e56fd 100644 --- a/Mathlib/GroupTheory/Perm/List.lean +++ b/Mathlib/GroupTheory/Perm/List.lean @@ -280,7 +280,7 @@ theorem formPerm_rotate (l : List α) (h : Nodup l) (n : ℕ) : formPerm (l.rotate n) = formPerm l := by induction' n with n hn · simp - · rw [Nat.succ_eq_add_one, ← rotate_rotate, formPerm_rotate_one, hn] + · rw [← rotate_rotate, formPerm_rotate_one, hn] rwa [IsRotated.nodup_iff] exact IsRotated.forall l n #align list.form_perm_rotate List.formPerm_rotate @@ -364,14 +364,12 @@ theorem formPerm_ext_iff {x y x' y' : α} {l l' : List α} (hd : Nodup (x :: y : · refine' Eq.trans _ hx' congr simpa using hn - · have : k.succ = (k + 1) % (x' :: y' :: l').length := by - rw [← Nat.succ_eq_add_one, Nat.mod_eq_of_lt hk'] - simp_rw [this] - rw [← formPerm_apply_nthLe _ hd' k (k.lt_succ_self.trans hk'), ← - IH (k.lt_succ_self.trans hk), ← h, formPerm_apply_nthLe _ hd] + · conv => congr <;> · arg 2; rw [← Nat.mod_eq_of_lt hk'] + rw [← formPerm_apply_nthLe _ hd' k (k.lt_succ_self.trans hk'), + ← IH (k.lt_succ_self.trans hk), ← h, formPerm_apply_nthLe _ hd] congr 1 - have h1 : 1 = 1 % (x' :: y' :: l').length := by simp - rw [hl, Nat.mod_eq_of_lt hk', h1, ← Nat.add_mod, Nat.succ_add, Nat.succ_eq_add_one] + rw [hl, Nat.mod_eq_of_lt hk', add_right_comm] + apply Nat.add_mod #align list.form_perm_ext_iff List.formPerm_ext_iff set_option linter.deprecated false in @@ -385,7 +383,7 @@ theorem formPerm_apply_mem_eq_self_iff (hl : Nodup l) (x : α) (hx : x ∈ l) : rcases (Nat.le_of_lt_succ hk).eq_or_lt with hk' | hk' · simp [← hk', Nat.succ_le_succ_iff, eq_comm] · simpa [Nat.mod_eq_of_lt (Nat.succ_lt_succ hk'), Nat.succ_lt_succ_iff] using - k.zero_le.trans_lt hk' + (k.zero_le.trans_lt hk').ne.symm #align list.form_perm_apply_mem_eq_self_iff List.formPerm_apply_mem_eq_self_iff theorem formPerm_apply_mem_ne_self_iff (hl : Nodup l) (x : α) (hx : x ∈ l) : diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index c756e3be8cb71..e561cff43baa7 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -296,7 +296,6 @@ theorem eq_one_of_smul_normalized (w : CoprodI.Word G) {i : ι} (h : H) have hhead : ((d.compl i).equiv (Word.equivPair i w).head).2 = (Word.equivPair i w).head := by rw [Word.equivPair_head] - dsimp only split_ifs with h · rcases h with ⟨_, rfl⟩ exact hw _ _ (List.head_mem _) diff --git a/Mathlib/GroupTheory/Subsemigroup/Center.lean b/Mathlib/GroupTheory/Subsemigroup/Center.lean index 04191efaf84a0..9923ab2f2ed38 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Center.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Center.lean @@ -178,15 +178,15 @@ theorem natCast_mem_center [NonAssocSemiring M] (n : ℕ) : (n : M) ∈ Set.cent comm _:= by rw [Nat.commute_cast] left_assoc _ _ := by induction n with - | zero => rw [Nat.zero_eq, Nat.cast_zero, zero_mul, zero_mul, zero_mul] + | zero => rw [Nat.cast_zero, zero_mul, zero_mul, zero_mul] | succ n ihn => rw [Nat.cast_succ, add_mul, one_mul, ihn, add_mul, add_mul, one_mul] mid_assoc _ _ := by induction n with - | zero => rw [Nat.zero_eq, Nat.cast_zero, zero_mul, mul_zero, zero_mul] + | zero => rw [Nat.cast_zero, zero_mul, mul_zero, zero_mul] | succ n ihn => rw [Nat.cast_succ, add_mul, mul_add, add_mul, ihn, mul_add, one_mul, mul_one] right_assoc _ _ := by induction n with - | zero => rw [Nat.zero_eq, Nat.cast_zero, mul_zero, mul_zero, mul_zero] + | zero => rw [Nat.cast_zero, mul_zero, mul_zero, mul_zero] | succ n ihn => rw [Nat.cast_succ, mul_add, ihn, mul_add, mul_add, mul_one, mul_one] -- See note [no_index around OfNat.ofNat] diff --git a/Mathlib/Init/Data/Nat/Basic.lean b/Mathlib/Init/Data/Nat/Basic.lean index 924126efb8813..edd0bf952be93 100644 --- a/Mathlib/Init/Data/Nat/Basic.lean +++ b/Mathlib/Init/Data/Nat/Basic.lean @@ -11,6 +11,8 @@ namespace Nat set_option linter.deprecated false +theorem add_one_pos (n : ℕ) : 0 < n + 1 := succ_pos n + protected theorem bit0_succ_eq (n : ℕ) : bit0 (succ n) = succ (succ (bit0 n)) := show succ (succ n + n) = succ (succ (n + n)) from congrArg succ (succ_add n n) #align nat.bit0_succ_eq Nat.bit0_succ_eq diff --git a/Mathlib/Lean/Expr/Basic.lean b/Mathlib/Lean/Expr/Basic.lean index 1f5ebca34fb31..27de816d01087 100644 --- a/Mathlib/Lean/Expr/Basic.lean +++ b/Mathlib/Lean/Expr/Basic.lean @@ -279,7 +279,7 @@ section recognizers - `Nat.succ x` where `isNumeral x` - `OfNat.ofNat _ x _` where `isNumeral x` -/ partial def numeral? (e : Expr) : Option Nat := - if let some n := e.natLit? then n + if let some n := e.rawNatLit? then n else let f := e.getAppFn if !f.isConst then none diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index c76662d7bfa17..208b9e64fd3f2 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -269,7 +269,7 @@ theorem cramer_eq_adjugate_mulVec (A : Matrix n n α) (b : n → α) : cramer A b = A.adjugate *ᵥ b := by nth_rw 2 [← A.transpose_transpose] rw [← adjugate_transpose, adjugate_def] - have : b = ∑ i, b i • Pi.single i 1 := by + have : b = ∑ i, b i • (Pi.single i 1 : n → α) := by refine' (pi_eq_sum_univ b).trans _ congr with j -- Porting note: needed to help `Pi.smul_apply` diff --git a/Mathlib/LinearAlgebra/Matrix/ZPow.lean b/Mathlib/LinearAlgebra/Matrix/ZPow.lean index f6e33feb44929..c576048ab5256 100644 --- a/Mathlib/LinearAlgebra/Matrix/ZPow.lean +++ b/Mathlib/LinearAlgebra/Matrix/ZPow.lean @@ -60,8 +60,7 @@ theorem pow_inv_comm' (A : M) (m n : ℕ) : A⁻¹ ^ m * A ^ n = A ^ n * A⁻¹ cases' m with m m · simp rcases nonsing_inv_cancel_or_zero A with (⟨h, h'⟩ | h) - · simp only [Nat.succ_eq_add_one] - calc + · calc A⁻¹ ^ (m + 1) * A ^ (n + 1) = A⁻¹ ^ m * (A⁻¹ * A) * A ^ n := by simp only [pow_succ' A⁻¹, pow_succ A, Matrix.mul_assoc] _ = A ^ n * A⁻¹ ^ m := by simp only [h, Matrix.mul_one, Matrix.one_mul, IH m] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index 26dfcc06f0a61..229b0022a91a9 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -1228,8 +1228,7 @@ in which `2` is invertible, there exists an orthogonal basis with respect to `B` theorem exists_orthogonal_basis [hK : Invertible (2 : K)] {B : BilinForm K V} (hB₂ : B.IsSymm) : ∃ v : Basis (Fin (finrank K V)) K V, B.IsOrthoᵢ v := by induction' hd : finrank K V with d ih generalizing V - · simp_rw [Nat.zero_eq] - exact ⟨basisOfFinrankZero hd, fun _ _ _ => map_zero _⟩ + · exact ⟨basisOfFinrankZero hd, fun _ _ _ => map_zero _⟩ haveI := finrank_pos_iff.1 (hd.symm ▸ Nat.succ_pos d : 0 < finrank K V) -- either the bilinear form is trivial or we can pick a non-null `x` obtain rfl | hB₁ := eq_or_ne B 0 diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index 16fb2d3d4c1b3..aed01b8b08b10 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -306,7 +306,11 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) (LinearMap.mul' R C ∘ₗ (TensorProduct.map f.toLinearMap g.toLinearMap) ∘ₗ ((of R 𝒜 ℬ).symm : 𝒜 ᵍ⊗[R] ℬ →ₗ[R] A ⊗[R] B)) - (by dsimp [Algebra.TensorProduct.one_def]; simp only [_root_.map_one, mul_one]) + (by + -- FIXME nightly-testing: not sure what is happening here + sorry + -- dsimp [Algebra.TensorProduct.one_def]; simp only [_root_.map_one, mul_one] + ) (by rw [LinearMap.map_mul_iff] ext a₁ : 3 diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index c322367e4095f..e53868cec2302 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -220,7 +220,7 @@ open scoped Classical theorem exists_succ (x : s) : ∃ n, (x : ℕ) + n + 1 ∈ s := _root_.by_contradiction fun h => - have : ∀ (a : ℕ) (_ : a ∈ s), a < succ x := fun a ha => + have : ∀ (a : ℕ) (_ : a ∈ s), a < x + 1 := fun a ha => lt_of_not_ge fun hax => h ⟨a - (x + 1), by rwa [add_right_comm, add_tsub_cancel_of_le hax]⟩ Fintype.false ⟨(((Multiset.range (succ x)).filter (· ∈ s)).pmap diff --git a/Mathlib/Logic/Function/Iterate.lean b/Mathlib/Logic/Function/Iterate.lean index 1c60f5679c9f1..3bb920f949187 100644 --- a/Mathlib/Logic/Function/Iterate.lean +++ b/Mathlib/Logic/Function/Iterate.lean @@ -124,7 +124,7 @@ theorem iterate_left {g : ℕ → α → α} (H : ∀ n, Semiconj f (g n) (g <| rw [Nat.zero_add] exact id_left | succ n ihn => - rw [Nat.succ_eq_add_one, Nat.add_right_comm, Nat.add_assoc] + rw [Nat.add_right_comm, Nat.add_assoc] exact (H k).trans (ihn (k + 1)) #align function.semiconj.iterate_left Function.Semiconj.iterate_left diff --git a/Mathlib/Logic/Small/Defs.lean b/Mathlib/Logic/Small/Defs.lean index fd9404885f872..b2605345e91fa 100644 --- a/Mathlib/Logic/Small/Defs.lean +++ b/Mathlib/Logic/Small/Defs.lean @@ -63,7 +63,7 @@ theorem Shrink.ext {α : Type v} [Small.{w} α] {x y : Shrink α} -- It would be nice to mark this as `aesop cases` if -- https://github.com/JLimperg/aesop/issues/59 -- is resolved. -@[eliminator] +@[induction_eliminator] protected noncomputable def Shrink.rec [Small.{w} α] {F : Shrink α → Sort v} (h : ∀ X, F (equivShrink _ X)) : ∀ X, F X := fun X => ((equivShrink _).apply_symm_apply X) ▸ (h _) diff --git a/Mathlib/Mathport/Notation.lean b/Mathlib/Mathport/Notation.lean index 760f1da9b842c..68fa7c1fd46cf 100644 --- a/Mathlib/Mathport/Notation.lean +++ b/Mathlib/Mathport/Notation.lean @@ -179,7 +179,7 @@ def matchTypeOf (matchTy : Matcher) : Matcher := fun s => do /-- Matches raw nat lits. -/ def natLitMatcher (n : Nat) : Matcher := fun s => do - guard <| (← getExpr).natLit? == n + guard <| (← getExpr).rawNatLit? == n return s /-- Matches applications. -/ diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 74e8a0ae271e2..c0c1f04e7cdf0 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -432,7 +432,7 @@ instance Pi.opensMeasurableSpace {ι : Type*} {π : ι → Type*} [Countable ι] constructor have : Pi.topologicalSpace = .generateFrom { t | ∃ (s : ∀ a, Set (π a)) (i : Finset ι), (∀ a ∈ i, s a ∈ countableBasis (π a)) ∧ t = pi (↑i) s } := by - rw [funext fun a => @eq_generateFrom_countableBasis (π a) _ _, pi_generateFrom_eq] + simp only [funext fun a => @eq_generateFrom_countableBasis (π a) _ _, pi_generateFrom_eq] rw [borel_eq_generateFrom_of_subbasis this] apply generateFrom_le rintro _ ⟨s, i, hi, rfl⟩ diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 9c367f07cb548..279dde1c3fb3f 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -130,7 +130,7 @@ theorem generateFrom_pi_eq {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountabl theorem generateFrom_eq_pi [h : ∀ i, MeasurableSpace (α i)] {C : ∀ i, Set (Set (α i))} (hC : ∀ i, generateFrom (C i) = h i) (h2C : ∀ i, IsCountablySpanning (C i)) : generateFrom (pi univ '' pi univ C) = MeasurableSpace.pi := by - rw [← funext hC, generateFrom_pi_eq h2C] + simp only [← funext hC, generateFrom_pi_eq h2C] #align generate_from_eq_pi generateFrom_eq_pi /-- The product σ-algebra is generated from boxes, i.e. `s ×ˢ t` for sets `s : set α` and diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index e019424d5e68a..5a02eeaa8a96e 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -208,7 +208,7 @@ private theorem restrictNonposSeq_disjoint' {n m : ℕ} (h : n < m) : restrictNonposSeq s i n ∩ restrictNonposSeq s i m = ∅ := by rw [Set.eq_empty_iff_forall_not_mem] rintro x ⟨hx₁, hx₂⟩ - cases m; · rw [Nat.zero_eq] at h; omega + cases m; · omega · rw [restrictNonposSeq] at hx₂ exact (someExistsOneDivLT_subset hx₂).2 diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index d652206ac4237..25266fdf10c9c 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -939,7 +939,7 @@ theorem sum_eapproxDiff (f : α → ℝ≥0∞) (n : ℕ) (a : α) : induction' n with n IH · simp only [Nat.zero_eq, Nat.zero_add, Finset.sum_singleton, Finset.range_one] rfl - · erw [Finset.sum_range_succ, Nat.succ_eq_add_one, IH, eapproxDiff, coe_map, Function.comp_apply, + · erw [Finset.sum_range_succ, IH, eapproxDiff, coe_map, Function.comp_apply, coe_sub, Pi.sub_apply, ENNReal.coe_toNNReal, add_tsub_cancel_of_le (monotone_eapprox f (Nat.le_succ _) _)] apply (lt_of_le_of_lt _ (eapprox_lt_top f (n + 1) a)).ne diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean index b2fbe05e53f52..4319ec6ca643c 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/VolumeOfBalls.lean @@ -290,7 +290,7 @@ theorem Complex.volume_sum_rpow_lt [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) (r : (abs_nonneg _), ← Finset.mul_sum, abs_eq_self.mpr (le_of_lt hr), inv_mul_lt_iff (rpow_pos_of_pos hr _), mul_one, ← rpow_lt_rpow_iff (rpow_nonneg (h₁ _) _) (le_of_lt hr) (by linarith : 0 < p), ← rpow_mul (h₁ _), div_mul_cancel _ - (ne_of_gt (by linarith) : p ≠ 0), Real.rpow_one, finrank_pi, finrank_pi_fintype ℝ, + (ne_of_gt (by linarith) : p ≠ 0), Real.rpow_one, finrank_pi_fintype ℝ, Complex.finrank_real_complex, Finset.sum_const, smul_eq_mul, mul_comm, Fintype.card] theorem Complex.volume_sum_rpow_le [Nonempty ι] {p : ℝ} (hp : 1 ≤ p) (r : ℝ) : diff --git a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean index 94c38401ac847..137dcb587aed0 100644 --- a/Mathlib/MeasureTheory/OuterMeasure/Basic.lean +++ b/Mathlib/MeasureTheory/OuterMeasure/Basic.lean @@ -200,7 +200,7 @@ theorem iUnion_nat_of_monotone_of_tsum_ne_top (m : OuterMeasure α) {s : ℕ → clear hx i rcases le_or_lt j n with hjn | hnj · exact Or.inl (h' hjn hj) - have : j - (n + 1) + n + 1 = j := by rw [add_assoc, tsub_add_cancel_of_le hnj.nat_succ_le] + have : j - (n + 1) + n + 1 = j := by omega refine' Or.inr (mem_iUnion.2 ⟨j - (n + 1), _, hlt _ _⟩) · rwa [this] · rw [← Nat.succ_le_iff, Nat.succ_eq_add_one, this] diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 1269d6f9020ba..6825fe774f264 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -970,7 +970,7 @@ theorem cardFactors_eq_one_iff_prime {n : ℕ} : Ω n = 1 ↔ n.Prime := by · contrapose! h simp rcases List.length_eq_one.1 h with ⟨x, hx⟩ - rw [← prod_factors n.succ_ne_zero, hx, List.prod_singleton] + rw [← prod_factors n.add_one_ne_zero, hx, List.prod_singleton] apply prime_of_mem_factors rw [hx, List.mem_singleton] #align nat.arithmetic_function.card_factors_eq_one_iff_prime ArithmeticFunction.cardFactors_eq_one_iff_prime diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 1276850f55758..b52411c11aa4d 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -145,7 +145,7 @@ theorem sum_bernoulli' (n : ℕ) : (∑ k in range n, (n.choose k : ℚ) * berno simp_rw [mul_sum, ← mul_assoc] refine' sum_congr rfl fun k hk => _ congr - have : ((n - k : ℕ) : ℚ) + 1 ≠ 0 := by norm_cast; exact succ_ne_zero _ + have : ((n - k : ℕ) : ℚ) + 1 ≠ 0 := by norm_cast field_simp [← cast_sub (mem_range.1 hk).le, mul_comm] rw_mod_cast [tsub_add_eq_add_tsub (mem_range.1 hk).le, choose_mul_succ_eq] #align sum_bernoulli' sum_bernoulli' @@ -252,8 +252,9 @@ theorem bernoulli_spec' (n : ℕ) : · simp rw [if_neg (succ_ne_zero _)] -- algebra facts - have h₁ : (1, n) ∈ antidiagonal n.succ := by simp [mem_antidiagonal, add_comm] - have h₂ : (n : ℚ) + 1 ≠ 0 := by norm_cast; exact succ_ne_zero _ + have h₁ : (1, n) ∈ antidiagonal n.succ := by + simp [mem_antidiagonal, add_comm, Nat.succ_eq_add_one] + have h₂ : (n : ℚ) + 1 ≠ 0 := by norm_cast have h₃ : (1 + n).choose n = n + 1 := by simp [add_comm] -- key equation: the corresponding fact for `bernoulli'` have H := bernoulli'_spec' n.succ @@ -387,7 +388,7 @@ theorem sum_Ico_pow (n p : ℕ) : suffices (∑ k in Ico 1 n.succ, (k : ℚ) ^ p.succ) = ∑ i in range p.succ.succ, f' i by convert this -- prove some algebraic facts that will make things easier for us later on have hle := Nat.le_add_left 1 n - have hne : (p + 1 + 1 : ℚ) ≠ 0 := by norm_cast; exact succ_ne_zero p.succ + have hne : (p + 1 + 1 : ℚ) ≠ 0 := by norm_cast have h1 : ∀ r : ℚ, r * (p + 1 + 1) * (n : ℚ) ^ p.succ / (p + 1 + 1 : ℚ) = r * (n : ℚ) ^ p.succ := fun r => by rw [mul_div_right_comm, mul_div_cancel _ hne] have h2 : f 1 + (n : ℚ) ^ p.succ = 1 / 2 * (n : ℚ) ^ p.succ := by diff --git a/Mathlib/NumberTheory/BernoulliPolynomials.lean b/Mathlib/NumberTheory/BernoulliPolynomials.lean index 31d24af6abc64..13c08502c20a2 100644 --- a/Mathlib/NumberTheory/BernoulliPolynomials.lean +++ b/Mathlib/NumberTheory/BernoulliPolynomials.lean @@ -179,7 +179,6 @@ theorem sum_range_pow_eq_bernoulli_sub (n p : ℕ) : · rw [Nat.sub_sub_self (mem_range_le hx)] · rw [← choose_symm (mem_range_le hx)] · norm_cast - apply succ_ne_zero _ #align polynomial.sum_range_pow_eq_bernoulli_sub Polynomial.sum_range_pow_eq_bernoulli_sub /-- Rearrangement of `Polynomial.sum_range_pow_eq_bernoulli_sub`. -/ diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index 5d9c57b121177..019d4b4d3a422 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -166,7 +166,7 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( replace hk := Nat.pow_right_injective rfl.le hk rw [add_left_eq_self] at hk subst hk - rw [← Nat.one_eq_succ_zero, pow_one] at hζ hcycl + rw [pow_one] at hζ hcycl have : natDegree (minpoly K ζ) = 1 := by rw [hζ.eq_neg_one_of_two_right, show (-1 : L) = algebraMap K L (-1) by simp, minpoly.eq_X_sub_C_of_algebraMap_inj _ (NoZeroSMulDivisors.algebraMap_injective K L)] diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 03d9a257e3e87..9b9d6b321adbe 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -199,7 +199,7 @@ theorem properDivisors_one : properDivisors 1 = ∅ := by rw [properDivisors, Ic theorem pos_of_mem_divisors {m : ℕ} (h : m ∈ n.divisors) : 0 < m := by cases m - · rw [mem_divisors, zero_eq, zero_dvd_iff (a := n)] at h + · rw [mem_divisors, zero_dvd_iff (a := n)] at h cases h.2 h.1 apply Nat.succ_pos #align nat.pos_of_mem_divisors Nat.pos_of_mem_divisors diff --git a/Mathlib/NumberTheory/Harmonic/Defs.lean b/Mathlib/NumberTheory/Harmonic/Defs.lean index 2bb005937373a..ed381c4660433 100644 --- a/Mathlib/NumberTheory/Harmonic/Defs.lean +++ b/Mathlib/NumberTheory/Harmonic/Defs.lean @@ -36,5 +36,7 @@ lemma harmonic_pos {n : ℕ} (Hn : n ≠ 0) : 0 < harmonic n := Finset.nonempty_range_iff.mpr Hn lemma harmonic_eq_sum_Icc {n : ℕ} : harmonic n = ∑ i in Finset.Icc 1 n, (↑i)⁻¹ := by - rw [harmonic, Finset.range_eq_Ico, Finset.sum_Ico_add' (fun (i : ℕ) ↦ (i : ℚ)⁻¹) 0 n (c := 1), - Nat.Ico_succ_right] + rw [harmonic, Finset.range_eq_Ico, Finset.sum_Ico_add' (fun (i : ℕ) ↦ (i : ℚ)⁻¹) 0 n (c := 1)] + -- It might be better to restate `Nat.Ico_succ_right` in terms of `+ 1`, + -- as we try to move away from `Nat.succ`. + simp only [Nat.add_one, Nat.Ico_succ_right] diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index 607b20f5ffac5..c99e6efe527e1 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -184,7 +184,7 @@ theorem pow_prime_pow_sub_pow_prime_pow (a : ℕ) : multiplicity (↑p) (x ^ p ^ a - y ^ p ^ a) = multiplicity (↑p) (x - y) + a := by induction' a with a h_ind · rw [Nat.cast_zero, add_zero, pow_zero, pow_one, pow_one] - rw [← Nat.add_one, Nat.cast_add, Nat.cast_one, ← add_assoc, ← h_ind, pow_succ', pow_mul, pow_mul] + rw [Nat.cast_add, Nat.cast_one, ← add_assoc, ← h_ind, pow_succ', pow_mul, pow_mul] apply pow_prime_sub_pow_prime hp hp1 · rw [← geom_sum₂_mul] exact dvd_mul_of_dvd_right hxy _ @@ -203,6 +203,7 @@ theorem Int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) ( cases' n with n · simp only [multiplicity.zero, add_top, pow_zero, sub_self, Nat.zero_eq] have h : (multiplicity _ _).Dom := finite_nat_iff.mpr ⟨hp.ne_one, n.succ_pos⟩ + simp only [Nat.succ_eq_add_one] at h rcases eq_coe_iff.mp (PartENat.natCast_get h).symm with ⟨⟨k, hk⟩, hpn⟩ conv_lhs => rw [hk, pow_mul, pow_mul] rw [Nat.prime_iff_prime_int] at hp @@ -311,6 +312,7 @@ theorem Int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬2 cases' n with n · simp only [pow_zero, sub_self, multiplicity.zero, Int.ofNat_zero, Nat.zero_eq, add_top] have h : (multiplicity 2 n.succ).Dom := multiplicity.finite_nat_iff.mpr ⟨by norm_num, n.succ_pos⟩ + simp only [Nat.succ_eq_add_one] at h rcases multiplicity.eq_coe_iff.mp (PartENat.natCast_get h).symm with ⟨⟨k, hk⟩, hpn⟩ rw [hk, pow_mul, pow_mul, multiplicity.pow_sub_pow_of_prime, Int.two_pow_two_pow_sub_pow_two_pow _ hxy hx, ← hk, PartENat.natCast_get] diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 262fc00e65e9a..296f09474afd7 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -589,7 +589,7 @@ theorem card_complex_embeddings : fun φ : { φ // ¬ ComplexEmbedding.IsReal φ } => mkComplex φ = w).card = 2 by rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] simp_rw [Finset.sum_const, this, smul_eq_mul, mul_one, Fintype.card, Finset.card_eq_sum_ones, - Finset.mul_sum] + Finset.mul_sum, Finset.sum_const, smul_eq_mul, mul_one] rintro ⟨w, hw⟩ convert card_filter_mk_eq w · rw [← Fintype.card_subtype, ← Fintype.card_subtype] diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index a67a9e0f08ffa..d47a7aa7f7f9c 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -660,7 +660,7 @@ theorem padicValNat_eq_zero_of_mem_Ioo {m k : ℕ} theorem padicValNat_factorial_mul_add {n : ℕ} (m : ℕ) [hp : Fact p.Prime] (h : n < p) : padicValNat p (p * m + n) ! = padicValNat p (p * m) ! := by induction' n with n hn - · rw [zero_eq, add_zero] + · rw [add_zero] · rw [add_succ, factorial_succ, padicValNat.mul (succ_ne_zero (p * m + n)) <| factorial_ne_zero (p * m + _), hn <| lt_of_succ_lt h, ← add_succ, @@ -690,8 +690,11 @@ Taking (`p - 1`) times the `p`-adic valuation of `n!` equals `n` minus the sum o of `n`. -/ theorem sub_one_mul_padicValNat_factorial [hp : Fact p.Prime] (n : ℕ): (p - 1) * padicValNat p (n !) = n - (p.digits n).sum := by - rw [padicValNat_factorial <| lt_succ_of_lt <| lt.base (log p n), ← Finset.sum_Ico_add' _ 0 _ 1, - Ico_zero_eq_range, ← sub_one_mul_sum_log_div_pow_eq_sub_sum_digits] + rw [padicValNat_factorial <| lt_succ_of_lt <| lt.base (log p n)] + norm_cast; nth_rw 2 [← zero_add 1] + rw [Nat.succ_eq_add_one, ← Finset.sum_Ico_add' _ 0 _ 1, + Ico_zero_eq_range, ← sub_one_mul_sum_log_div_pow_eq_sub_sum_digits, Nat.succ_eq_add_one] + rfl /-- **Kummer's Theorem** diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index fc577e9ac2a88..7ae9b940f0471 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -338,7 +338,7 @@ theorem dvd_appr_sub_appr (x : ℤ_[p]) (m n : ℕ) (h : m ≤ n) : p ^ m ∣ x. obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le h; clear h induction' k with k ih · simp only [zero_eq, add_zero, le_refl, tsub_eq_zero_of_le, ne_eq, Nat.isUnit_iff, dvd_zero] - rw [Nat.succ_eq_add_one, ← add_assoc] + rw [← add_assoc] dsimp [appr] split_ifs with h · exact ih diff --git a/Mathlib/NumberTheory/ZetaFunction.lean b/Mathlib/NumberTheory/ZetaFunction.lean index 3f711ee33922c..2facee7f0ab3c 100644 --- a/Mathlib/NumberTheory/ZetaFunction.lean +++ b/Mathlib/NumberTheory/ZetaFunction.lean @@ -732,7 +732,7 @@ theorem riemannZeta_neg_nat_eq_bernoulli (k : ℕ) : rcases Nat.even_or_odd' k with ⟨m, rfl | rfl⟩ · cases' m with m m ·-- k = 0 : evaluate explicitly - rw [Nat.zero_eq, mul_zero, Nat.cast_zero, pow_zero, one_mul, zero_add, neg_zero, zero_add, + rw [mul_zero, Nat.cast_zero, pow_zero, one_mul, zero_add, neg_zero, zero_add, div_one, bernoulli_one, riemannZeta_zero] norm_num · -- k = 2 * (m + 1) : both sides "trivially" zero diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index c5cc898cc1956..a5270a1a5c973 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -272,7 +272,7 @@ theorem toZ_iterate_pred_of_not_isMin (n : ℕ) (hn : ¬IsMin (pred^[n] i0)) : exact isMin_iterate_pred_of_eq_of_ne h_pred_eq_pred (Nat.succ_ne_zero n) let m := (-toZ i0 (pred^[n.succ] i0)).toNat have h_eq : pred^[m] i0 = pred^[n.succ] i0 := iterate_pred_toZ _ this - by_cases hmn : m = n.succ + by_cases hmn : m = n + 1 · nth_rw 2 [← hmn] rw [Int.toNat_eq_max, toZ_of_lt this, max_eq_left, neg_neg] rw [neg_neg] diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index feeb7dcb62ff2..940d1641ce932 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -511,7 +511,7 @@ theorem crossing_eq_crossing_of_lowerCrossingTime_lt {M : ℕ} (hNM : N ≤ M) rw [eq_comm, ih.2] exacts [hitting_eq_hitting_of_exists hNM ⟨j, ⟨hj₁.1, hj₁.2.le⟩, hj₂⟩, le_rfl] refine' ⟨this, _⟩ - simp only [lowerCrossingTime, eq_comm, this] + simp only [lowerCrossingTime, eq_comm, this, Nat.succ_eq_add_one] refine' hitting_eq_hitting_of_exists hNM _ rw [lowerCrossingTime, hitting_lt_iff _ le_rfl] at h obtain ⟨j, hj₁, hj₂⟩ := h diff --git a/Mathlib/RingTheory/Derivation/Basic.lean b/Mathlib/RingTheory/Derivation/Basic.lean index bdcab35165461..e8da7618bbae5 100644 --- a/Mathlib/RingTheory/Derivation/Basic.lean +++ b/Mathlib/RingTheory/Derivation/Basic.lean @@ -159,7 +159,7 @@ theorem map_coe_nat (n : ℕ) : D (n : A) = 0 := by @[simp] theorem leibniz_pow (n : ℕ) : D (a ^ n) = n • a ^ (n - 1) • D a := by induction' n with n ihn - · rw [Nat.zero_eq, pow_zero, map_one_eq_zero, zero_smul] + · rw [pow_zero, map_one_eq_zero, zero_smul] · rcases (zero_le n).eq_or_lt with (rfl | hpos) · erw [pow_one, one_smul, pow_zero, one_smul] · have : a * a ^ (n - 1) = a ^ n := by rw [← pow_succ, Nat.sub_add_cancel hpos] diff --git a/Mathlib/RingTheory/Filtration.lean b/Mathlib/RingTheory/Filtration.lean index a75ff0f3099c6..f32c8c142bfde 100644 --- a/Mathlib/RingTheory/Filtration.lean +++ b/Mathlib/RingTheory/Filtration.lean @@ -66,7 +66,7 @@ namespace Ideal.Filtration theorem pow_smul_le (i j : ℕ) : I ^ i • F.N j ≤ F.N (i + j) := by induction' i with _ ih · simp - · rw [pow_succ, mul_smul, Nat.succ_eq_add_one, add_assoc, add_comm 1, ← add_assoc] + · rw [pow_succ, mul_smul, add_assoc, add_comm 1, ← add_assoc] exact (Submodule.smul_mono_right ih).trans (F.smul_le _) #align ideal.filtration.pow_smul_le Ideal.Filtration.pow_smul_le @@ -218,7 +218,7 @@ theorem Stable.exists_pow_smul_eq : ∃ n₀, ∀ k, F.N (n₀ + k) = I ^ k • intro k induction' k with _ ih · simp - · rw [Nat.succ_eq_add_one, ← add_assoc, ← hn, ih, add_comm, pow_add, mul_smul, pow_one] + · rw [← add_assoc, ← hn, ih, add_comm, pow_add, mul_smul, pow_one] omega #align ideal.filtration.stable.exists_pow_smul_eq Ideal.Filtration.Stable.exists_pow_smul_eq @@ -244,7 +244,7 @@ theorem Stable.exists_forall_le (h : F.Stable) (e : F.N 0 ≤ F'.N 0) : intro n induction' n with n hn · refine' (F.antitone _).trans e; simp - · rw [Nat.succ_eq_one_add, add_assoc, add_comm, add_comm 1 n, ← hF] + · rw [Nat.add_right_comm, ← hF] exact (Submodule.smul_mono_right hn).trans (F'.smul_le _) simp #align ideal.filtration.stable.exists_forall_le Ideal.Filtration.Stable.exists_forall_le @@ -355,7 +355,7 @@ theorem submodule_eq_span_le_iff_stable_ge (n₀ : ℕ) : by_cases hj' : j.succ ≤ n₀ · exact this _ hj' simp only [not_le, Nat.lt_succ_iff] at hj' - rw [Nat.succ_eq_add_one, ← hF _ hj'] + rw [← hF _ hj'] rintro _ ⟨m, hm, rfl⟩ refine' Submodule.smul_induction_on hm (fun r hr m' hm' => _) (fun x y hx hy => _) · rw [add_comm, ← monomial_smul_single] @@ -452,7 +452,7 @@ theorem Ideal.mem_iInf_smul_pow_eq_bot_iff [IsNoetherianRing R] [Module.Finite R intro i induction' i with i hi · simp - · rw [Nat.succ_eq_one_add, pow_add, ← smul_smul, pow_one, ← eq] + · rw [add_comm, pow_add, ← smul_smul, pow_one, ← eq] exact Submodule.smul_mem_smul r.prop hi #align ideal.mem_infi_smul_pow_eq_bot_iff Ideal.mem_iInf_smul_pow_eq_bot_iff diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index 523895f924d75..41b2c1f8d6dcf 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -193,7 +193,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] intro n induction' n with n ih · rfl - rw [Nat.succ_eq_add_one, hc, sub_eq_add_neg, ← add_zero a₀] + rw [hc, sub_eq_add_neg, ← add_zero a₀] refine' ih.add _ rw [SModEq.zero, Ideal.neg_mem_iff] refine' I.mul_mem_right _ _ @@ -209,7 +209,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] intro n induction' n with n ih · simpa only [Nat.zero_eq, Nat.rec_zero, zero_add, pow_one] using h₁ - rw [Nat.succ_eq_add_one, ← taylor_eval_sub (c n), hc, sub_eq_add_neg, sub_eq_add_neg, + rw [← taylor_eval_sub (c n), hc, sub_eq_add_neg, sub_eq_add_neg, add_neg_cancel_comm] rw [eval_eq_sum, sum_over_range' _ _ _ (lt_add_of_pos_right _ zero_lt_two), ← Finset.sum_range_add_sum_Ico _ (Nat.le_add_left _ _)] @@ -236,8 +236,12 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] obtain ⟨k, rfl⟩ := Nat.exists_eq_add_of_le hmn clear hmn induction' k with k ih - · rw [Nat.zero_eq, add_zero] - rw [Nat.succ_eq_add_one, ← add_assoc, hc, ← add_zero (c m), sub_eq_add_neg] + · rw [add_zero] + rw [← add_assoc] + -- Adaptation note: nightly-2024-03-11 + -- I'm not sure why the `erw` is now needed here. It looks like it should work. + erw [hc] + rw [← add_zero (c m), sub_eq_add_neg] refine' ih.add _ symm rw [SModEq.zero, Ideal.neg_mem_iff] @@ -256,7 +260,7 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] rw [SModEq.zero] exact Ideal.pow_le_pow_right le_self_add (hfcI _) · show a - a₀ ∈ I - specialize ha 1 + specialize ha (0 + 1) rw [hc, pow_one, ← Ideal.one_eq_top, Ideal.smul_eq_mul, mul_one, sub_eq_add_neg] at ha rw [← SModEq.sub_mem, ← add_zero a₀] refine' ha.symm.trans (SModEq.rfl.add _) diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index f187165b757e8..64c6165633814 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -608,7 +608,7 @@ protected theorem pow' {p a : α} (hp : Prime p) (ha : Finite p a) : induction' k with k hk · simp [one_right hp.not_unit] · have : multiplicity p (a ^ (k + 1)) = multiplicity p (a * a ^ k) := by rw [_root_.pow_succ] - rw [succ_eq_add_one, get_eq_get_of_eq _ _ this, + rw [get_eq_get_of_eq _ _ this, multiplicity.mul' hp, hk, add_mul, one_mul, add_comm] #align multiplicity.pow' multiplicity.pow' diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 70d8b32ecd96f..819d5f162c0de 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -127,7 +127,8 @@ private theorem weight_add_weight_pairMap {k : ℕ} (t : Finset σ × σ) (h : t · rw [pairMap_of_snd_nmem_fst σ h1] simp only [mul_comm, mul_assoc (∏ a in t.fst, X a), card_cons, prod_cons] nth_rewrite 2 [← pow_one (X t.snd)] - simp only [← pow_add, ← Nat.add_sub_assoc (Nat.lt_of_le_of_ne h.left (mt h.right h1)), add_comm] + simp only [← pow_add, ← Nat.add_sub_assoc (Nat.lt_of_le_of_ne h.left (mt h.right h1)), add_comm, + Nat.succ_eq_add_one, Nat.add_sub_add_right] rw [← neg_neg ((-1 : MvPolynomial σ R) ^ card t.fst), h2] simp diff --git a/Mathlib/RingTheory/MvPowerSeries/Basic.lean b/Mathlib/RingTheory/MvPowerSeries/Basic.lean index 00fd879feb40e..91626de3ba3a7 100644 --- a/Mathlib/RingTheory/MvPowerSeries/Basic.lean +++ b/Mathlib/RingTheory/MvPowerSeries/Basic.lean @@ -414,7 +414,7 @@ set_option linter.uppercaseLean3 false in theorem X_pow_eq (s : σ) (n : ℕ) : (X s : MvPowerSeries σ R) ^ n = monomial R (single s n) 1 := by induction' n with n ih · simp - · rw [pow_succ', ih, Nat.succ_eq_add_one, Finsupp.single_add, X, monomial_mul_monomial, one_mul] + · rw [pow_succ', ih, Finsupp.single_add, X, monomial_mul_monomial, one_mul] set_option linter.uppercaseLean3 false in #align mv_power_series.X_pow_eq MvPowerSeries.X_pow_eq diff --git a/Mathlib/RingTheory/Nilpotent.lean b/Mathlib/RingTheory/Nilpotent.lean index b48f95ec3df47..78b666f82454d 100644 --- a/Mathlib/RingTheory/Nilpotent.lean +++ b/Mathlib/RingTheory/Nilpotent.lean @@ -468,7 +468,7 @@ lemma NoZeroSMulDivisors.isReduced (R M : Type*) IsReduced R := by refine ⟨fun x ⟨k, hk⟩ ↦ ?_⟩ induction' k with k ih - · rw [Nat.zero_eq, pow_zero] at hk + · rw [pow_zero] at hk exact eq_zero_of_zero_eq_one hk.symm x · obtain ⟨m : M, hm : m ≠ 0⟩ := exists_ne (0 : M) have : x ^ (k + 1) • m = 0 := by simp only [hk, zero_smul] diff --git a/Mathlib/RingTheory/Polynomial/Bernstein.lean b/Mathlib/RingTheory/Polynomial/Bernstein.lean index 64d8ef871dfad..f987cb71b6ac9 100644 --- a/Mathlib/RingTheory/Polynomial/Bernstein.lean +++ b/Mathlib/RingTheory/Polynomial/Bernstein.lean @@ -227,9 +227,9 @@ theorem iterate_derivative_at_1 (n ν : ℕ) (h : ν ≤ n) : simp [Polynomial.eval_comp, h] obtain rfl | h' := h.eq_or_lt · simp - · congr - norm_cast - rw [← tsub_add_eq_tsub_tsub, tsub_tsub_cancel_of_le (Nat.succ_le_iff.mpr h')] + · norm_cast + congr + omega #align bernstein_polynomial.iterate_derivative_at_1 bernsteinPolynomial.iterate_derivative_at_1 theorem iterate_derivative_at_1_ne_zero [CharZero R] (n ν : ℕ) (h : ν ≤ n) : @@ -244,8 +244,7 @@ open Submodule theorem linearIndependent_aux (n k : ℕ) (h : k ≤ n + 1) : LinearIndependent ℚ fun ν : Fin k => bernsteinPolynomial ℚ n ν := by induction' k with k ih - · simp [Nat.zero_eq] - apply linearIndependent_empty_type + · apply linearIndependent_empty_type · apply linearIndependent_fin_succ'.mpr fconstructor · exact ih (le_of_lt h) diff --git a/Mathlib/RingTheory/Polynomial/Content.lean b/Mathlib/RingTheory/Polynomial/Content.lean index ddc39f53d25d4..9d963918fbca7 100644 --- a/Mathlib/RingTheory/Polynomial/Content.lean +++ b/Mathlib/RingTheory/Polynomial/Content.lean @@ -363,7 +363,6 @@ theorem content_mul {p q : R[X]} : (p * q).content = p.content * q.content := by intro n induction' n with n ih · intro p q hpq - dsimp at hpq rw [Nat.cast_zero, Nat.WithBot.lt_zero_iff, degree_eq_bot, mul_eq_zero] at hpq rcases hpq with (rfl | rfl) <;> simp diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean index e56406a0d49ca..416cee4287248 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean @@ -567,10 +567,12 @@ theorem cyclotomic_prime_pow_eq_geom_sum {R : Type*} [CommRing R] {p n : ℕ} (h rw [this, Nat.prod_properDivisors_prime_pow hp] induction' n with n_n n_ih · haveI := Fact.mk hp; simp [cyclotomic_prime] - rw [((eq_cyclotomic_iff (pow_pos hp.pos (n_n.succ + 1)) _).mpr _).symm] + rw [((eq_cyclotomic_iff (pow_pos hp.pos (n_n + 1 + 1)) _).mpr _).symm] rw [Nat.prod_properDivisors_prime_pow hp, Finset.prod_range_succ, n_ih] rw [this] at n_ih - rw [mul_comm _ (∑ i in _, _), n_ih, geom_sum_mul, sub_left_inj, ← pow_mul, pow_add, pow_one] + rw [mul_comm _ (∑ i in _, _), n_ih, geom_sum_mul, sub_left_inj, ← pow_mul] + simp only [pow_add, pow_one] + #align polynomial.cyclotomic_prime_pow_eq_geom_sum Polynomial.cyclotomic_prime_pow_eq_geom_sum theorem cyclotomic_prime_pow_mul_X_pow_sub_one (R : Type*) [CommRing R] (p k : ℕ) diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index 16f92c9824f41..d1e451eb320de 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -240,138 +240,141 @@ theorem mem_adjoin_of_smul_prime_smul_of_minpoly_isEisensteinAt {B : PowerBasis (hp : Prime p) (hBint : IsIntegral R B.gen) {z : L} (hzint : IsIntegral R z) (hz : p • z ∈ adjoin R ({B.gen} : Set L)) (hei : (minpoly R B.gen).IsEisensteinAt 𝓟) : z ∈ adjoin R ({B.gen} : Set L) := by - -- First define some abbreviations. - have hndiv : ¬p ^ 2 ∣ (minpoly R B.gen).coeff 0 := fun h => - hei.not_mem ((span_singleton_pow p 2).symm ▸ Ideal.mem_span_singleton.2 h) - have := B.finite - set P := minpoly R B.gen with hP - obtain ⟨n, hn⟩ := Nat.exists_eq_succ_of_ne_zero B.dim_pos.ne' - haveI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans R K L - let _ := P.map (algebraMap R L) - -- There is a polynomial `Q` such that `p • z = aeval B.gen Q`. We can assume that - -- `Q.degree < P.degree` and `Q ≠ 0`. - rw [adjoin_singleton_eq_range_aeval] at hz - obtain ⟨Q₁, hQ⟩ := hz - set Q := Q₁ %ₘ P with hQ₁ - replace hQ : aeval B.gen Q = p • z := by - rw [← modByMonic_add_div Q₁ (minpoly.monic hBint)] at hQ - simpa using hQ - by_cases hQzero : Q = 0 - · simp only [hQzero, Algebra.smul_def, zero_eq_mul, aeval_zero] at hQ - cases' hQ with H H₁ - · have : Function.Injective (algebraMap R L) := by - rw [algebraMap_eq R K L] - exact (algebraMap K L).injective.comp (IsFractionRing.injective R K) - exfalso - exact hp.ne_zero ((injective_iff_map_eq_zero _).1 this _ H) - · rw [H₁] - exact Subalgebra.zero_mem _ - -- It is enough to prove that all coefficients of `Q` are divisible by `p`, by induction. - -- The base case is `dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt`. - refine mem_adjoin_of_dvd_coeff_of_dvd_aeval hp.ne_zero (fun i => ?_) hQ - induction' i using Nat.case_strong_induction_on with j hind - · intro _ - exact dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt hp hBint hQ hzint hei - · intro hj - convert hp.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd _ hndiv - exact n - -- Two technical results we will need about `P.natDegree` and `Q.natDegree`. - have H := degree_modByMonic_lt Q₁ (minpoly.monic hBint) - rw [← hQ₁, ← hP] at H - replace H := Nat.lt_iff_add_one_le.1 - (lt_of_lt_of_le - (lt_of_le_of_lt (Nat.lt_iff_add_one_le.1 (Nat.lt_of_succ_lt_succ (mem_range.1 hj))) - (lt_succ_self _)) (Nat.lt_iff_add_one_le.1 ((natDegree_lt_natDegree_iff hQzero).2 H))) - rw [add_assoc] at H - have Hj : Q.natDegree + 1 = j + 1 + (Q.natDegree - j) := by - rw [← add_comm 1, ← add_comm 1, add_assoc, add_right_inj, - ← Nat.add_sub_assoc (Nat.lt_of_succ_lt_succ (mem_range.1 hj)).le, add_comm, - Nat.add_sub_cancel] - -- By induction hypothesis we can find `g : ℕ → R` such that - -- `k ∈ range (j + 1) → Q.coeff k • B.gen ^ k = (algebraMap R L) p * g k • B.gen ^ k`- - choose! g hg using hind - replace hg : ∀ k ∈ range (j + 1), Q.coeff k • B.gen ^ k = - algebraMap R L p * g k • B.gen ^ k := by - intro k hk - rw [hg k (mem_range_succ_iff.1 hk) - (mem_range_succ_iff.2 - (le_trans (mem_range_succ_iff.1 hk) (succ_le_iff.1 (mem_range_succ_iff.1 hj)).le)), - Algebra.smul_def, Algebra.smul_def, RingHom.map_mul, mul_assoc] - -- Since `minpoly R B.gen` is Eiseinstein, we can find `f : ℕ → L` such that - -- `(map (algebraMap R L) (minpoly R B.gen)).nat_degree ≤ i` implies `f i ∈ adjoin R {B.gen}` - -- and `(algebraMap R L) p * f i = B.gen ^ i`. We will also need `hf₁`, a reformulation of this - -- property. - choose! f hf using - IsWeaklyEisensteinAt.exists_mem_adjoin_mul_eq_pow_natDegree_le (minpoly.aeval R B.gen) - (minpoly.monic hBint) hei.isWeaklyEisensteinAt - have hf₁ : ∀ k ∈ (range (Q.natDegree - j)).erase 0, - Q.coeff (j + 1 + k) • B.gen ^ (j + 1 + k) * B.gen ^ (P.natDegree - (j + 2)) = - (algebraMap R L) p * Q.coeff (j + 1 + k) • f (k + P.natDegree - 1) := by - intro k hk - rw [smul_mul_assoc, ← pow_add, ← Nat.add_sub_assoc H, ← add_assoc j 1 1, add_comm (j + 1) 1, - add_assoc (j + 1), add_comm _ (k + P.natDegree), Nat.add_sub_add_right, - ← (hf (k + P.natDegree - 1) _).2, mul_smul_comm] - rw [(minpoly.monic hBint).natDegree_map, add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] - · exact Nat.zero_le _ - · refine one_le_iff_ne_zero.2 fun h => ?_ - rw [h] at hk - simp at hk - - -- The Eisenstein condition shows that `p` divides `Q.coeff j` - -- if `p^n.succ` divides the following multiple of `Q.coeff (succ j)^n.succ`: - suffices - p ^ n.succ ∣ Q.coeff (succ j) ^ n.succ * - (minpoly R B.gen).coeff 0 ^ (succ j + (P.natDegree - (j + 2))) by - convert this - rw [Nat.succ_eq_add_one, add_assoc, ← Nat.add_sub_assoc H, ← add_assoc, add_comm (j + 1), - Nat.add_sub_add_left, ← Nat.add_sub_assoc, Nat.add_sub_add_left, hP, ← - (minpoly.monic hBint).natDegree_map (algebraMap R K), ← - minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, natDegree_minpoly, hn, Nat.sub_one, - Nat.pred_succ] - omega - - -- Using `hQ : aeval B.gen Q = p • z`, we write `p • z` as a sum of terms of degree less than - -- `j+1`, that are multiples of `p` by induction, and terms of degree at least `j+1`. - rw [aeval_eq_sum_range, Hj, range_add, sum_union (disjoint_range_addLeftEmbedding _ _), - sum_congr rfl hg, add_comm] at hQ - -- We multiply this equality by `B.gen ^ (P.natDegree-(j+2))`, so we can use `hf₁` on the terms - -- we didn't know were multiples of `p`, and we take the norm on both sides. - replace hQ := congr_arg (fun x => x * B.gen ^ (P.natDegree - (j + 2))) hQ - simp_rw [sum_map, addLeftEmbedding_apply, add_mul, sum_mul, mul_assoc] at hQ - rw [← insert_erase - (mem_range.2 (tsub_pos_iff_lt.2 <| Nat.lt_of_succ_lt_succ <| mem_range.1 hj)), - sum_insert (not_mem_erase 0 _), add_zero, sum_congr rfl hf₁, ← mul_sum, ← mul_sum, add_assoc, - ← mul_add, smul_mul_assoc, ← pow_add, Algebra.smul_def] at hQ - replace hQ := congr_arg (norm K) (eq_sub_of_add_eq hQ) - - -- We obtain an equality of elements of `K`, but everything is integral, so we can move to `R` - -- and simplify `hQ`. - have hintsum : IsIntegral R (z * B.gen ^ (P.natDegree - (j + 2)) - - (∑ x : ℕ in (range (Q.natDegree - j)).erase 0, - Q.coeff (j + 1 + x) • f (x + P.natDegree - 1) + - ∑ x : ℕ in range (j + 1), g x • B.gen ^ x * B.gen ^ (P.natDegree - (j + 2)))) := by - refine (hzint.mul (hBint.pow _)).sub - (.add (.sum _ fun k hk => .smul _ ?_) - (.sum _ fun k _ => .mul (.smul _ (.pow hBint _)) (hBint.pow _))) - refine adjoin_le_integralClosure hBint (hf _ ?_).1 - rw [(minpoly.monic hBint).natDegree_map (algebraMap R L)] - rw [add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] - · exact _root_.zero_le _ - · refine one_le_iff_ne_zero.2 fun h => ?_ - rw [h] at hk - simp at hk - obtain ⟨r, hr⟩ := isIntegral_iff.1 (isIntegral_norm K hintsum) - rw [Algebra.smul_def, mul_assoc, ← mul_sub, _root_.map_mul, algebraMap_apply R K L, map_pow, - Algebra.norm_algebraMap, _root_.map_mul, algebraMap_apply R K L, Algebra.norm_algebraMap, - finrank B, ← hr, PowerBasis.norm_gen_eq_coeff_zero_minpoly, - minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, coeff_map, - show (-1 : K) = algebraMap R K (-1) by simp, ← map_pow, ← map_pow, ← _root_.map_mul, ← - map_pow, ← _root_.map_mul, ← map_pow, ← _root_.map_mul] at hQ - -- We can now finish the proof. - have hppdiv : p ^ B.dim ∣ p ^ B.dim * r := dvd_mul_of_dvd_left dvd_rfl _ - rwa [← IsFractionRing.injective R K hQ, mul_comm, ← Units.coe_neg_one, mul_pow, ← - Units.val_pow_eq_pow_val, ← Units.val_pow_eq_pow_val, mul_assoc, - Units.dvd_mul_left, mul_comm, ← Nat.succ_eq_add_one, hn] at hppdiv + -- FIXME nightly-testing: proof broken, + -- perhaps can be split into smaller pieces or made more robust? + sorry + -- -- First define some abbreviations. + -- have hndiv : ¬p ^ 2 ∣ (minpoly R B.gen).coeff 0 := fun h => + -- hei.not_mem ((span_singleton_pow p 2).symm ▸ Ideal.mem_span_singleton.2 h) + -- have := B.finite + -- set P := minpoly R B.gen with hP + -- obtain ⟨n, hn⟩ := Nat.exists_eq_succ_of_ne_zero B.dim_pos.ne' + -- haveI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans R K L + -- let _ := P.map (algebraMap R L) + -- -- There is a polynomial `Q` such that `p • z = aeval B.gen Q`. We can assume that + -- -- `Q.degree < P.degree` and `Q ≠ 0`. + -- rw [adjoin_singleton_eq_range_aeval] at hz + -- obtain ⟨Q₁, hQ⟩ := hz + -- set Q := Q₁ %ₘ P with hQ₁ + -- replace hQ : aeval B.gen Q = p • z := by + -- rw [← modByMonic_add_div Q₁ (minpoly.monic hBint)] at hQ + -- simpa using hQ + -- by_cases hQzero : Q = 0 + -- · simp only [hQzero, Algebra.smul_def, zero_eq_mul, aeval_zero] at hQ + -- cases' hQ with H H₁ + -- · have : Function.Injective (algebraMap R L) := by + -- rw [algebraMap_eq R K L] + -- exact (algebraMap K L).injective.comp (IsFractionRing.injective R K) + -- exfalso + -- exact hp.ne_zero ((injective_iff_map_eq_zero _).1 this _ H) + -- · rw [H₁] + -- exact Subalgebra.zero_mem _ + -- -- It is enough to prove that all coefficients of `Q` are divisible by `p`, by induction. + -- -- The base case is `dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt`. + -- refine mem_adjoin_of_dvd_coeff_of_dvd_aeval hp.ne_zero (fun i => ?_) hQ + -- induction' i using Nat.case_strong_induction_on with j hind + -- · intro _ + -- exact dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt hp hBint hQ hzint hei + -- · intro hj + -- convert hp.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd _ hndiv + -- exact n + -- -- Two technical results we will need about `P.natDegree` and `Q.natDegree`. + -- have H := degree_modByMonic_lt Q₁ (minpoly.monic hBint) + -- rw [← hQ₁, ← hP] at H + -- replace H := Nat.lt_iff_add_one_le.1 + -- (lt_of_lt_of_le + -- (lt_of_le_of_lt (Nat.lt_iff_add_one_le.1 (Nat.lt_of_succ_lt_succ (mem_range.1 hj))) + -- (lt_succ_self _)) (Nat.lt_iff_add_one_le.1 ((natDegree_lt_natDegree_iff hQzero).2 H))) + -- rw [add_assoc] at H + -- have Hj : Q.natDegree + 1 = j + 1 + (Q.natDegree - j) := by + -- rw [← add_comm 1, ← add_comm 1, add_assoc, add_right_inj, + -- ← Nat.add_sub_assoc (Nat.lt_of_succ_lt_succ (mem_range.1 hj)).le, add_comm, + -- Nat.add_sub_cancel] + -- -- By induction hypothesis we can find `g : ℕ → R` such that + -- -- `k ∈ range (j + 1) → Q.coeff k • B.gen ^ k = (algebraMap R L) p * g k • B.gen ^ k`- + -- choose! g hg using hind + -- replace hg : ∀ k ∈ range (j + 1), Q.coeff k • B.gen ^ k = + -- algebraMap R L p * g k • B.gen ^ k := by + -- intro k hk + -- rw [hg k (mem_range_succ_iff.1 hk) + -- (mem_range_succ_iff.2 + -- (le_trans (mem_range_succ_iff.1 hk) (succ_le_iff.1 (mem_range_succ_iff.1 hj)).le)), + -- Algebra.smul_def, Algebra.smul_def, RingHom.map_mul, mul_assoc] + -- -- Since `minpoly R B.gen` is Eiseinstein, we can find `f : ℕ → L` such that + -- -- `(map (algebraMap R L) (minpoly R B.gen)).nat_degree ≤ i` implies `f i ∈ adjoin R {B.gen}` + -- -- and `(algebraMap R L) p * f i = B.gen ^ i`. We will also need `hf₁`, a reformulation of this + -- -- property. + -- choose! f hf using + -- IsWeaklyEisensteinAt.exists_mem_adjoin_mul_eq_pow_natDegree_le (minpoly.aeval R B.gen) + -- (minpoly.monic hBint) hei.isWeaklyEisensteinAt + -- have hf₁ : ∀ k ∈ (range (Q.natDegree - j)).erase 0, + -- Q.coeff (j + 1 + k) • B.gen ^ (j + 1 + k) * B.gen ^ (P.natDegree - (j + 2)) = + -- (algebraMap R L) p * Q.coeff (j + 1 + k) • f (k + P.natDegree - 1) := by + -- intro k hk + -- rw [smul_mul_assoc, ← pow_add, ← Nat.add_sub_assoc H, ← add_assoc j 1 1, add_comm (j + 1) 1, + -- add_assoc (j + 1), add_comm _ (k + P.natDegree), Nat.add_sub_add_right, + -- ← (hf (k + P.natDegree - 1) _).2, mul_smul_comm] + -- rw [(minpoly.monic hBint).natDegree_map, add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] + -- · exact Nat.zero_le _ + -- · refine one_le_iff_ne_zero.2 fun h => ?_ + -- rw [h] at hk + -- simp at hk + + -- -- The Eisenstein condition shows that `p` divides `Q.coeff j` + -- -- if `p^n.succ` divides the following multiple of `Q.coeff (succ j)^n.succ`: + -- suffices + -- p ^ n.succ ∣ Q.coeff (succ j) ^ n.succ * + -- (minpoly R B.gen).coeff 0 ^ (succ j + (P.natDegree - (j + 2))) by + -- convert this + -- rw [Nat.succ_eq_add_one, add_assoc, ← Nat.add_sub_assoc H, ← add_assoc, add_comm (j + 1), + -- Nat.add_sub_add_left, ← Nat.add_sub_assoc, Nat.add_sub_add_left, hP, ← + -- (minpoly.monic hBint).natDegree_map (algebraMap R K), ← + -- minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, natDegree_minpoly, hn, Nat.sub_one, + -- Nat.pred_succ] + -- omega + + -- -- Using `hQ : aeval B.gen Q = p • z`, we write `p • z` as a sum of terms of degree less than + -- -- `j+1`, that are multiples of `p` by induction, and terms of degree at least `j+1`. + -- rw [aeval_eq_sum_range, Hj, range_add, sum_union (disjoint_range_addLeftEmbedding _ _), + -- sum_congr rfl hg, add_comm] at hQ + -- -- We multiply this equality by `B.gen ^ (P.natDegree-(j+2))`, so we can use `hf₁` on the terms + -- -- we didn't know were multiples of `p`, and we take the norm on both sides. + -- replace hQ := congr_arg (fun x => x * B.gen ^ (P.natDegree - (j + 2))) hQ + -- simp_rw [sum_map, addLeftEmbedding_apply, add_mul, sum_mul, mul_assoc] at hQ + -- rw [← insert_erase + -- (mem_range.2 (tsub_pos_iff_lt.2 <| Nat.lt_of_succ_lt_succ <| mem_range.1 hj)), + -- sum_insert (not_mem_erase 0 _), add_zero, sum_congr rfl hf₁, ← mul_sum, ← mul_sum, add_assoc, + -- ← mul_add, smul_mul_assoc, ← pow_add, Algebra.smul_def] at hQ + -- replace hQ := congr_arg (norm K) (eq_sub_of_add_eq hQ) + + -- -- We obtain an equality of elements of `K`, but everything is integral, so we can move to `R` + -- -- and simplify `hQ`. + -- have hintsum : IsIntegral R (z * B.gen ^ (P.natDegree - (j + 2)) - + -- (∑ x : ℕ in (range (Q.natDegree - j)).erase 0, + -- Q.coeff (j + 1 + x) • f (x + P.natDegree - 1) + + -- ∑ x : ℕ in range (j + 1), g x • B.gen ^ x * B.gen ^ (P.natDegree - (j + 2)))) := by + -- refine (hzint.mul (hBint.pow _)).sub + -- (.add (.sum _ fun k hk => .smul _ ?_) + -- (.sum _ fun k _ => .mul (.smul _ (.pow hBint _)) (hBint.pow _))) + -- refine adjoin_le_integralClosure hBint (hf _ ?_).1 + -- rw [(minpoly.monic hBint).natDegree_map (algebraMap R L)] + -- rw [add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] + -- · exact _root_.zero_le _ + -- · refine one_le_iff_ne_zero.2 fun h => ?_ + -- rw [h] at hk + -- simp at hk + -- obtain ⟨r, hr⟩ := isIntegral_iff.1 (isIntegral_norm K hintsum) + -- rw [Algebra.smul_def, mul_assoc, ← mul_sub, _root_.map_mul, algebraMap_apply R K L, map_pow, + -- Algebra.norm_algebraMap, _root_.map_mul, algebraMap_apply R K L, Algebra.norm_algebraMap, + -- finrank B, ← hr, PowerBasis.norm_gen_eq_coeff_zero_minpoly, + -- minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, coeff_map, + -- show (-1 : K) = algebraMap R K (-1) by simp, ← map_pow, ← map_pow, ← _root_.map_mul, ← + -- map_pow, ← _root_.map_mul, ← map_pow, ← _root_.map_mul] at hQ + -- -- We can now finish the proof. + -- have hppdiv : p ^ B.dim ∣ p ^ B.dim * r := dvd_mul_of_dvd_left dvd_rfl _ + -- rwa [← IsFractionRing.injective R K hQ, mul_comm, ← Units.coe_neg_one, mul_pow, ← + -- Units.val_pow_eq_pow_val, ← Units.val_pow_eq_pow_val, mul_assoc, + -- Units.dvd_mul_left, mul_comm, ← Nat.succ_eq_add_one, hn] at hppdiv #align mem_adjoin_of_smul_prime_smul_of_minpoly_is_eiseinstein_at mem_adjoin_of_smul_prime_smul_of_minpoly_isEisensteinAt /-- Let `K` be the field of fraction of an integrally closed domain `R` and let `L` be a separable diff --git a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean index 69269cbfa2325..22e7eef4c3c54 100644 --- a/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Hermite/Basic.lean @@ -95,7 +95,7 @@ theorem coeff_hermite_of_lt {n k : ℕ} (hnk : n < k) : coeff (hermite n) k = 0 induction' n with n ih generalizing k · apply coeff_C · have : n + k + 1 + 2 = n + (k + 2) + 1 := by ring - rw [Nat.succ_eq_add_one, coeff_hermite_succ_succ, add_right_comm, this, ih k, ih (k + 2), + rw [coeff_hermite_succ_succ, add_right_comm, this, ih k, ih (k + 2), mul_zero, sub_zero] #align polynomial.coeff_hermite_of_lt Polynomial.coeff_hermite_of_lt @@ -132,7 +132,7 @@ theorem hermite_monic (n : ℕ) : (hermite n).Monic := theorem coeff_hermite_of_odd_add {n k : ℕ} (hnk : Odd (n + k)) : coeff (hermite n) k = 0 := by induction' n with n ih generalizing k - · rw [Nat.zero_eq, zero_add k] at hnk + · rw [zero_add k] at hnk exact coeff_hermite_of_lt hnk.pos · cases' k with k · rw [Nat.succ_add_eq_add_succ] at hnk diff --git a/Mathlib/RingTheory/Polynomial/Pochhammer.lean b/Mathlib/RingTheory/Polynomial/Pochhammer.lean index 659b3a5a4b766..d7a486fba981a 100644 --- a/Mathlib/RingTheory/Polynomial/Pochhammer.lean +++ b/Mathlib/RingTheory/Polynomial/Pochhammer.lean @@ -158,7 +158,7 @@ theorem ascPochhammer_mul (n m : ℕ) : induction' m with m ih · simp · rw [ascPochhammer_succ_right, Polynomial.mul_X_add_nat_cast_comp, ← mul_assoc, ih, - Nat.succ_eq_add_one, ← add_assoc, ascPochhammer_succ_right, Nat.cast_add, add_assoc] + ← add_assoc, ascPochhammer_succ_right, Nat.cast_add, add_assoc] #align pochhammer_mul ascPochhammer_mul theorem ascPochhammer_nat_eq_ascFactorial (n : ℕ) : @@ -184,7 +184,7 @@ theorem ascPochhammer_natDegree (n : ℕ) [NoZeroDivisors S] [Nontrivial S] : natDegree_mul _ (ne_zero_of_natDegree_gt <| this.symm ▸ Nat.zero_lt_one), hn, this] cases n · simp - · refine' ne_zero_of_natDegree_gt <| hn.symm ▸ Nat.succ_pos _ + · refine' ne_zero_of_natDegree_gt <| hn.symm ▸ Nat.add_one_pos _ end Semiring @@ -309,8 +309,7 @@ theorem descPochhammer_succ_right (n : ℕ) : · conv_lhs => rw [descPochhammer_succ_left, ih, mul_comp, ← mul_assoc, ← descPochhammer_succ_left, sub_comp, X_comp, nat_cast_comp] - nth_rw 1 [Nat.succ_eq_add_one] - rw [Nat.succ_eq_one_add, Nat.cast_add, Nat.cast_one, sub_add_eq_sub_sub] + rw [Nat.add_comm, Nat.cast_add, Nat.cast_one, sub_add_eq_sub_sub] @[simp] theorem descPochhammer_natDegree (n : ℕ) [NoZeroDivisors R] [Nontrivial R] : @@ -322,7 +321,7 @@ theorem descPochhammer_natDegree (n : ℕ) [NoZeroDivisors R] [Nontrivial R] : natDegree_mul _ (ne_zero_of_natDegree_gt <| this.symm ▸ Nat.zero_lt_one), hn, this] cases n · simp - · refine' ne_zero_of_natDegree_gt <| hn.symm ▸ Nat.succ_pos _ + · refine' ne_zero_of_natDegree_gt <| hn.symm ▸ Nat.add_one_pos _ theorem descPochhammer_succ_eval {S : Type*} [Ring S] (n : ℕ) (k : S) : (descPochhammer S (n + 1)).eval k = (descPochhammer S n).eval k * (k - n) := by @@ -353,8 +352,8 @@ theorem descPochhammer_mul (n m : ℕ) : descPochhammer R n * (descPochhammer R m).comp (X - (n : R[X])) = descPochhammer R (n + m) := by induction' m with m ih · simp - · rw [descPochhammer_succ_right, Polynomial.mul_X_sub_int_cast_comp, ← mul_assoc, ih, - Nat.succ_eq_add_one, ← add_assoc, descPochhammer_succ_right, Nat.cast_add, sub_add_eq_sub_sub] + · rw [descPochhammer_succ_right, Polynomial.mul_X_sub_int_cast_comp, ← mul_assoc, ih, ← add_assoc, + descPochhammer_succ_right, Nat.cast_add, sub_add_eq_sub_sub] theorem ascPochhammer_eval_neg_eq_descPochhammer (r : R) : ∀ (k : ℕ), (ascPochhammer R k).eval (-r) = (-1)^k * (descPochhammer R k).eval r diff --git a/Mathlib/RingTheory/QuotientNilpotent.lean b/Mathlib/RingTheory/QuotientNilpotent.lean index f8f48aba2aa81..fdf0d57568277 100644 --- a/Mathlib/RingTheory/QuotientNilpotent.lean +++ b/Mathlib/RingTheory/QuotientNilpotent.lean @@ -44,7 +44,7 @@ theorem Ideal.IsNilpotent.induction_on (hI : IsNilpotent I) exact (hI' hI).elim apply h₂ (I ^ 2) _ (Ideal.pow_le_self two_ne_zero) · apply H n.succ _ (I ^ 2) - · rw [← pow_mul, eq_bot_iff, ← hI, Nat.succ_eq_add_one, Nat.succ_eq_add_one] + · rw [← pow_mul, eq_bot_iff, ← hI, Nat.succ_eq_add_one] apply Ideal.pow_le_pow_right (by omega) · exact n.succ.lt_succ_self · apply h₁ diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index 722e801d3a85d..6d7d0b974eb0c 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1830,8 +1830,8 @@ theorem pow_factors [Nontrivial α] {a : Associates α} {k : ℕ} : theorem count_pow [Nontrivial α] {a : Associates α} (ha : a ≠ 0) {p : Associates α} (hp : Irreducible p) (k : ℕ) : count p (a ^ k).factors = k * count p a.factors := by induction' k with n h - · rw [pow_zero, factors_one, Nat.zero_eq, zero_mul, count_zero hp] - · rw [pow_succ, count_mul ha (pow_ne_zero _ ha) hp, h, Nat.succ_eq_add_one] + · rw [pow_zero, factors_one, zero_mul, count_zero hp] + · rw [pow_succ, count_mul ha (pow_ne_zero _ ha) hp, h] ring #align associates.count_pow Associates.count_pow diff --git a/Mathlib/RingTheory/WittVector/Identities.lean b/Mathlib/RingTheory/WittVector/Identities.lean index 4416bdc07b1c4..b1355978dd0d0 100644 --- a/Mathlib/RingTheory/WittVector/Identities.lean +++ b/Mathlib/RingTheory/WittVector/Identities.lean @@ -63,7 +63,7 @@ theorem coeff_p_pow [CharP R p] (i : ℕ) : ((p : 𝕎 R) ^ i).coeff i = 1 := by theorem coeff_p_pow_eq_zero [CharP R p] {i j : ℕ} (hj : j ≠ i) : ((p : 𝕎 R) ^ i).coeff j = 0 := by induction' i with i hi generalizing j - · rw [Nat.zero_eq, pow_zero, one_coeff_eq_of_pos] + · rw [pow_zero, one_coeff_eq_of_pos] exact Nat.pos_of_ne_zero hj · rw [pow_succ', ← frobenius_verschiebung, coeff_frobenius_charP] cases j diff --git a/Mathlib/RingTheory/WittVector/MulP.lean b/Mathlib/RingTheory/WittVector/MulP.lean index f8ef7d29d5433..a278ded1120c3 100644 --- a/Mathlib/RingTheory/WittVector/MulP.lean +++ b/Mathlib/RingTheory/WittVector/MulP.lean @@ -52,8 +52,7 @@ theorem mulN_coeff (n : ℕ) (x : 𝕎 R) (k : ℕ) : induction' n with n ih generalizing k · simp only [Nat.zero_eq, Nat.cast_zero, mul_zero, zero_coeff, wittMulN, AlgHom.map_zero, Pi.zero_apply] - · rw [wittMulN, Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one, mul_add, mul_one, aeval_bind₁, - add_coeff] + · rw [wittMulN, Nat.cast_add, Nat.cast_one, mul_add, mul_one, aeval_bind₁, add_coeff] apply eval₂Hom_congr (RingHom.ext_int _ _) _ rfl ext1 ⟨b, i⟩ fin_cases b diff --git a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean index 6a3c4e7d9574d..a4a4deca107af 100644 --- a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean +++ b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean @@ -245,7 +245,7 @@ theorem C_p_pow_dvd_bind₁_rename_wittPolynomial_sub_sum (Φ : MvPolynomial idx apply_fun map (Int.castRingHom (ZMod (p ^ (n + 1)))) at key conv_lhs at key => simp only [map_bind₁, map_rename, map_expand, map_wittPolynomial] -- clean up and massage - rw [Nat.succ_eq_add_one, C_dvd_iff_zmod, RingHom.map_sub, sub_eq_zero, map_bind₁] + rw [C_dvd_iff_zmod, RingHom.map_sub, sub_eq_zero, map_bind₁] simp only [map_rename, map_wittPolynomial, wittPolynomial_zmod_self] rw [key]; clear key IH rw [bind₁, aeval_wittPolynomial, map_sum, map_sum, Finset.sum_congr rfl] diff --git a/Mathlib/RingTheory/WittVector/Verschiebung.lean b/Mathlib/RingTheory/WittVector/Verschiebung.lean index 0dd79a392d7eb..11d637b9c258e 100644 --- a/Mathlib/RingTheory/WittVector/Verschiebung.lean +++ b/Mathlib/RingTheory/WittVector/Verschiebung.lean @@ -68,8 +68,7 @@ theorem ghostComponent_verschiebungFun (x : 𝕎 R) (n : ℕ) : rw [Finset.sum_range_succ', verschiebungFun_coeff, if_pos rfl, zero_pow (pow_ne_zero _ hp.1.ne_zero), mul_zero, add_zero, Finset.mul_sum, Finset.sum_congr rfl] rintro i - - simp only [pow_succ, mul_assoc, verschiebungFun_coeff, if_neg (Nat.succ_ne_zero i), - Nat.succ_sub_succ, tsub_zero] + simp only [pow_succ, verschiebungFun_coeff_succ, Nat.succ_sub_succ_eq_sub, mul_assoc] #align witt_vector.ghost_component_verschiebung_fun WittVector.ghostComponent_verschiebungFun /-- The 0th Verschiebung polynomial is 0. For `n > 0`, the `n`th Verschiebung polynomial is the @@ -90,7 +89,7 @@ theorem aeval_verschiebung_poly' (x : 𝕎 R) (n : ℕ) : · simp only [verschiebungPoly, Nat.zero_eq, ge_iff_le, tsub_eq_zero_of_le, ite_true, map_zero, verschiebungFun_coeff_zero] · rw [verschiebungPoly, verschiebungFun_coeff_succ, if_neg n.succ_ne_zero, aeval_X, - Nat.succ_eq_add_one, add_tsub_cancel_right] + add_tsub_cancel_right] #align witt_vector.aeval_verschiebung_poly' WittVector.aeval_verschiebung_poly' variable (p) diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 045cd8aad6d20..c99423d953eee 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -1055,9 +1055,7 @@ theorem succ_zero : succ (0 : Ordinal) = 1 := -- Porting note: Proof used to be rfl @[simp] -theorem succ_one : succ (1 : Ordinal) = 2 := by - unfold instOfNatAtLeastTwo OfNat.ofNat - simpa using by rfl +theorem succ_one : succ (1 : Ordinal) = 2 := by congr; simp only [Nat.unaryCast, zero_add] #align ordinal.succ_one Ordinal.succ_one theorem add_succ (o₁ o₂ : Ordinal) : o₁ + succ o₂ = succ (o₁ + o₂) := diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index d0087b85c50a3..3cd6d54b525e6 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -567,8 +567,7 @@ theorem nfp_add_zero (a) : nfp (a + ·) 0 = a * omega := by congr; funext n induction' n with n hn · rw [Nat.cast_zero, mul_zero, iterate_zero_apply] - · nth_rw 2 [Nat.succ_eq_one_add] - rw [Nat.cast_add, Nat.cast_one, mul_one_add, iterate_succ_apply', hn] + · rw [iterate_succ_apply', Nat.add_comm, Nat.cast_add, Nat.cast_one, mul_one_add, hn] #align ordinal.nfp_add_zero Ordinal.nfp_add_zero theorem nfp_add_eq_mul_omega {a b} (hba : b ≤ a * omega) : nfp (a + ·) b = a * omega := by @@ -617,8 +616,7 @@ theorem nfp_mul_one {a : Ordinal} (ha : 0 < a) : nfp (a * ·) 1 = (a^omega) := b funext n induction' n with n hn · rw [Nat.cast_zero, opow_zero, iterate_zero_apply] - nth_rw 2 [Nat.succ_eq_one_add] - rw [Nat.cast_add, Nat.cast_one, opow_add, opow_one, iterate_succ_apply', hn] + rw [iterate_succ_apply', Nat.add_comm, Nat.cast_add, Nat.cast_one, opow_add, opow_one, hn] · exact ha #align ordinal.nfp_mul_one Ordinal.nfp_mul_one diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 7a8213fda97fa..29781444cbc9c 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -1088,63 +1088,65 @@ attribute FundamentalSequenceProp theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamentalSequence o) := by - induction' o with a m b iha ihb; · exact rfl - rw [fundamentalSequence] - rcases e : b.fundamentalSequence with (⟨_ | b'⟩ | f) <;> - simp only [FundamentalSequenceProp] <;> - rw [e, FundamentalSequenceProp] at ihb - · rcases e : a.fundamentalSequence with (⟨_ | a'⟩ | f) <;> cases' e' : m.natPred with m' <;> - simp only [FundamentalSequenceProp] <;> - rw [e, FundamentalSequenceProp] at iha <;> - (try rw [show m = 1 by - have := PNat.natPred_add_one m; rw [e'] at this; exact PNat.coe_inj.1 this.symm]) <;> - (try rw [show m = m'.succ.succPNat by - rw [← e', ← PNat.coe_inj, Nat.succPNat_coe, ← Nat.add_one, PNat.natPred_add_one]]) <;> - simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega, add_lt_add_iff_left, add_zero, - eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, - Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, - true_and_iff, _root_.zero_add, zero_def] - · decide - · exact ⟨rfl, inferInstance⟩ - · have := opow_pos (repr a') omega_pos - refine' - ⟨mul_isLimit this omega_isLimit, fun i => - ⟨this, _, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega'⟩ - rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] - apply nat_lt_omega - · have := opow_pos (repr a') omega_pos - refine' - ⟨add_isLimit _ (mul_isLimit this omega_isLimit), fun i => ⟨this, _, _⟩, - exists_lt_add exists_lt_mul_omega'⟩ - · rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] - apply nat_lt_omega - · refine' fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (iha.2 H.fst))) - rw [repr, ← zero_def, repr, add_zero, iha.1, opow_succ, Ordinal.mul_lt_mul_iff_left this] - apply nat_lt_omega - · rcases iha with ⟨h1, h2, h3⟩ - refine' ⟨opow_isLimit one_lt_omega h1, fun i => _, exists_lt_omega_opow' one_lt_omega h1 h3⟩ - obtain ⟨h4, h5, h6⟩ := h2 i - exact ⟨h4, h5, fun H => @NF.oadd_zero _ _ (h6 H.fst)⟩ - · rcases iha with ⟨h1, h2, h3⟩ - refine' - ⟨add_isLimit _ (opow_isLimit one_lt_omega h1), fun i => _, - exists_lt_add (exists_lt_omega_opow' one_lt_omega h1 h3)⟩ - obtain ⟨h4, h5, h6⟩ := h2 i - refine' ⟨h4, h5, fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (h6 H.fst)))⟩ - rwa [repr, ← zero_def, repr, add_zero, PNat.one_coe, Nat.cast_one, mul_one, - opow_lt_opow_iff_right one_lt_omega] - · refine' - ⟨by rw [repr, ihb.1, add_succ, repr], fun H => H.fst.oadd _ (NF.below_of_lt' _ (ihb.2 H.snd))⟩ - have := H.snd'.repr_lt - rw [ihb.1] at this - exact (lt_succ _).trans this - · rcases ihb with ⟨h1, h2, h3⟩ - simp only [repr] - exact - ⟨Ordinal.add_isLimit _ h1, fun i => - ⟨oadd_lt_oadd_3 (h2 i).1, oadd_lt_oadd_3 (h2 i).2.1, fun H => - H.fst.oadd _ (NF.below_of_lt' (lt_trans (h2 i).2.1 H.snd'.repr_lt) ((h2 i).2.2 H.snd))⟩, - exists_lt_add h3⟩ + -- FIXME nightly-testing: proof broken, hard to tell origin intent. + sorry + -- induction' o with a m b iha ihb; · exact rfl + -- rw [fundamentalSequence] + -- rcases e : b.fundamentalSequence with (⟨_ | b'⟩ | f) <;> + -- simp only [FundamentalSequenceProp] <;> + -- rw [e, FundamentalSequenceProp] at ihb + -- · rcases e : a.fundamentalSequence with (⟨_ | a'⟩ | f) <;> cases' e' : m.natPred with m' <;> + -- simp only [FundamentalSequenceProp] <;> + -- rw [e, FundamentalSequenceProp] at iha <;> + -- (try rw [show m = 1 by + -- have := PNat.natPred_add_one m; rw [e'] at this; exact PNat.coe_inj.1 this.symm]) <;> + -- (try rw [show m = m'.succ.succPNat by + -- rw [← e', ← PNat.coe_inj, Nat.succPNat_coe, ← Nat.add_one, PNat.natPred_add_one]]) <;> + -- simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega, add_lt_add_iff_left, add_zero, + -- eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, + -- Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, + -- true_and_iff, _root_.zero_add, zero_def] + -- · decide + -- · exact ⟨rfl, inferInstance⟩ + -- · have := opow_pos (repr a') omega_pos + -- refine' + -- ⟨mul_isLimit this omega_isLimit, fun i => + -- ⟨this, _, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega'⟩ + -- rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] + -- apply nat_lt_omega + -- · have := opow_pos (repr a') omega_pos + -- refine' + -- ⟨add_isLimit _ (mul_isLimit this omega_isLimit), fun i => ⟨this, _, _⟩, + -- exists_lt_add exists_lt_mul_omega'⟩ + -- · rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] + -- apply nat_lt_omega + -- · refine' fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (iha.2 H.fst))) + -- rw [repr, ← zero_def, repr, add_zero, iha.1, opow_succ, Ordinal.mul_lt_mul_iff_left this] + -- apply nat_lt_omega + -- · rcases iha with ⟨h1, h2, h3⟩ + -- refine' ⟨opow_isLimit one_lt_omega h1, fun i => _, exists_lt_omega_opow' one_lt_omega h1 h3⟩ + -- obtain ⟨h4, h5, h6⟩ := h2 i + -- exact ⟨h4, h5, fun H => @NF.oadd_zero _ _ (h6 H.fst)⟩ + -- · rcases iha with ⟨h1, h2, h3⟩ + -- refine' + -- ⟨add_isLimit _ (opow_isLimit one_lt_omega h1), fun i => _, + -- exists_lt_add (exists_lt_omega_opow' one_lt_omega h1 h3)⟩ + -- obtain ⟨h4, h5, h6⟩ := h2 i + -- refine' ⟨h4, h5, fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (h6 H.fst)))⟩ + -- rwa [repr, ← zero_def, repr, add_zero, PNat.one_coe, Nat.cast_one, mul_one, + -- opow_lt_opow_iff_right one_lt_omega] + -- · refine' + -- ⟨by rw [repr, ihb.1, add_succ, repr], fun H => H.fst.oadd _ (NF.below_of_lt' _ (ihb.2 H.snd))⟩ + -- have := H.snd'.repr_lt + -- rw [ihb.1] at this + -- exact (lt_succ _).trans this + -- · rcases ihb with ⟨h1, h2, h3⟩ + -- simp only [repr] + -- exact + -- ⟨Ordinal.add_isLimit _ h1, fun i => + -- ⟨oadd_lt_oadd_3 (h2 i).1, oadd_lt_oadd_3 (h2 i).2.1, fun H => + -- H.fst.oadd _ (NF.below_of_lt' (lt_trans (h2 i).2.1 H.snd'.repr_lt) ((h2 i).2.2 H.snd))⟩, + -- exists_lt_add h3⟩ #align onote.fundamental_sequence_has_prop ONote.fundamentalSequence_has_prop /-- The fast growing hierarchy for ordinal notations `< ε₀`. This is a sequence of diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index a5e168371c4fb..97254f896e28f 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -618,7 +618,7 @@ noncomputable def allDefinable : ∀ {n} (F : OfArity ZFSet ZFSet n), Definable let p := @Quotient.exists_rep PSet _ F @Definable.EqMk 0 ⟨choose p, Equiv.rfl⟩ _ (choose_spec p) | n + 1, (F : OfArity ZFSet ZFSet (n + 1)) => by - have I : (x : ZFSet) → Definable (Nat.add n 0) (F x) := fun x => allDefinable (F x) + have I : (x : ZFSet) → Definable n (F x) := fun x => allDefinable (F x) refine' @Definable.EqMk (n + 1) ⟨fun x : PSet => (@Definable.Resp _ _ (I ⟦x⟧)).1, _⟩ _ _ · dsimp [Arity.Equiv] intro x y h diff --git a/Mathlib/Tactic/Find.lean b/Mathlib/Tactic/Find.lean index 2f68e9be0c092..7bb0725f6f0d3 100644 --- a/Mathlib/Tactic/Find.lean +++ b/Mathlib/Tactic/Find.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sebastian Ullrich -/ import Std.Util.Cache +import Lean.HeadIndex +import Lean.Elab.Command /-! # The `#find` command and tactic. diff --git a/Mathlib/Tactic/NormNum/Result.lean b/Mathlib/Tactic/NormNum/Result.lean index 1ae40eff0db16..8bd6cfcfbec8b 100644 --- a/Mathlib/Tactic/NormNum/Result.lean +++ b/Mathlib/Tactic/NormNum/Result.lean @@ -108,7 +108,7 @@ def mkOfNat (α : Q(Type u)) (_sα : Q(AddMonoidWithOne $α)) (lit : Q(ℕ)) : let a' : Q(ℚ) := q(OfNat.ofNat $lit : ℚ) pure ⟨a', (q(Eq.refl $a') : Expr)⟩ else - let some n := lit.natLit? | failure + let some n := lit.rawNatLit? | failure match n with | 0 => pure ⟨q(0 : $α), (q(Nat.cast_zero (R := $α)) : Expr)⟩ | 1 => pure ⟨q(1 : $α), (q(Nat.cast_one (R := $α)) : Expr)⟩ diff --git a/Mathlib/Tactic/Observe.lean b/Mathlib/Tactic/Observe.lean index 9fc1f1bdcfbd9..b6f47514b7d3b 100644 --- a/Mathlib/Tactic/Observe.lean +++ b/Mathlib/Tactic/Observe.lean @@ -36,9 +36,7 @@ elab_rules : tactic | withMainContext do let (type, _) ← elabTermWithHoles t none (← getMainTag) true let .mvar goal ← mkFreshExprMVar type | failure - -- FIXME: this won't be needed (it's a default argument) after nightly-2024-02-26 - let tactic := fun initial g => solveByElim [] (maxDepth := 6) (exfalso := initial) g - if let some _ ← librarySearch goal tactic then + if let some _ ← librarySearch goal then reportOutOfHeartbeats `library_search tk throwError "observe did not find a solution" else diff --git a/Mathlib/Tactic/Positivity/Basic.lean b/Mathlib/Tactic/Positivity/Basic.lean index 7c0b1003114d5..5ca445b806d5d 100644 --- a/Mathlib/Tactic/Positivity/Basic.lean +++ b/Mathlib/Tactic/Positivity/Basic.lean @@ -268,7 +268,7 @@ def evalPow : PositivityExt where eval {u α} zα pα e := do let .app (.app _ (a : Q($α))) (b : Q(ℕ)) ← withReducible (whnf e) | throwError "not ^" let result ← catchNone do let .true := b.isAppOfArity ``OfNat.ofNat 3 | throwError "not a ^ n where n is a literal" - let some n := (b.getRevArg! 1).natLit? | throwError "not a ^ n where n is a literal" + let some n := (b.getRevArg! 1).rawNatLit? | throwError "not a ^ n where n is a literal" guard (n % 2 = 0) have m : Q(ℕ) := mkRawNatLit (n / 2) haveI' : $b =Q bit0 $m := ⟨⟩ diff --git a/Mathlib/Tactic/Sat/FromLRAT.lean b/Mathlib/Tactic/Sat/FromLRAT.lean index f20193ee4d36f..aee43b246077c 100644 --- a/Mathlib/Tactic/Sat/FromLRAT.lean +++ b/Mathlib/Tactic/Sat/FromLRAT.lean @@ -485,7 +485,7 @@ where These are both lookups into the context `(a0 .. a(n-1) : Prop) (v) (h1 : v 0 ↔ a0) ... (hn : v (n-1) ↔ a(n-1))`. -/ reifyVar v := - let n := v.natLit?.get! + let n := v.rawNatLit?.get! (mkBVar (2 * nvars - n), mkBVar (nvars - n - 1)) open Lean diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 90c9432e03b2a..4b5357d0d4d9c 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -1536,7 +1536,6 @@ theorem Products.max_eq_eval (l : Products I) (hl : l.val ≠ []) · simp [mem_C'_eq_false C ho x x.prop, Bool.coe_false] · push_neg at h₂; obtain ⟨i, hi⟩ := h₂; exfalso; rw [hi' i hi.1] at hi; exact hi.2 (h₁ i hi.1) · push_neg at h₁; obtain ⟨i, hi⟩ := h₁; exfalso; rw [← hi' i hi.1] at hi; exact hi.2 (h₃ i hi.1) - · rfl namespace GoodProducts diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index 03a9bd3069b7e..1793a944d58d4 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -212,6 +212,12 @@ theorem ι_eq_iff_rel (i j : D.J) (x : D.U i) (y : D.U j) : rintro _ _ ⟨x⟩ rw [← show (sigmaIsoSigma.{u, u} _).inv _ = x from ConcreteCategory.congr_hom (sigmaIsoSigma.{u, u} _).hom_inv_id x] + -- Adaption note: v4.7.0-rc1 + -- The behaviour of `generalize` was changed in https://github.com/leanprover/lean4/pull/3575 + -- to use transparancy `instances` (rather than `default`) + -- `generalize'` is a temporary backwards compatibility shim. + -- Hopefully we will be able to refactor this proof to use `generalize` again, and then drop + -- `generalize'`. generalize' h : (sigmaIsoSigma.{u, u} D.V).hom x = x' obtain ⟨⟨i, j⟩, y⟩ := x' unfold InvImage MultispanIndex.fstSigmaMap MultispanIndex.sndSigmaMap diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 9eb4e9c2a4056..559793a9b589d 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -434,7 +434,7 @@ protected theorem continuous_div_const (c : ℝ≥0∞) (c_ne_zero : c ≠ 0) : theorem continuous_pow (n : ℕ) : Continuous fun a : ℝ≥0∞ => a ^ n := by induction' n with n IH · simp [continuous_const] - simp_rw [Nat.succ_eq_add_one, pow_add, pow_one, continuous_iff_continuousAt] + simp_rw [pow_add, pow_one, continuous_iff_continuousAt] intro x refine' ENNReal.Tendsto.mul (IH.tendsto _) _ tendsto_id _ <;> by_cases H : x = 0 · simp only [H, zero_ne_top, Ne.def, or_true_iff, not_false_iff] diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index 07ce0ea521c13..0e550add758ce 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -1020,7 +1020,7 @@ instance : CompleteSpace GHSpace := by have E : ∀ n : ℕ, GlueSpace (Y n).isom (isometry_optimalGHInjl (X n) (X (n + 1))) = (Y (n + 1)).Space := - fun n => by dsimp only [Y, auxGluing]; rfl + fun n => by dsimp only [Y, auxGluing] let c n := cast (E n) have ic : ∀ n, Isometry (c n) := fun n x y => by dsimp only [Y, auxGluing]; exact rfl -- there is a canonical embedding of `Y n` in `Y (n+1)`, by construction diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 72dcf1b35c8ee..23d73ecde0eec 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -268,7 +268,15 @@ theorem comp₂_left (hf : HasCompactMulSupport f) (hf₂ : HasCompactMulSupport f₂) (hm : m 1 1 = 1) : HasCompactMulSupport fun x => m (f x) (f₂ x) := by rw [hasCompactMulSupport_iff_eventuallyEq] at hf hf₂ ⊢ - filter_upwards [hf, hf₂] using fun x hx hx₂ => by simp_rw [hx, hx₂, Pi.one_apply, hm] + /- Adaptation note: (`nightly-2024-03-11`) If we *either* (1) remove the type annotations on the + binders in the following `fun` or (2) revert `simp only` to `simp_rw`, `to_additive` fails + because an `OfNat.ofNat 1` is not replaced with `0`. Notably, as of this nightly, what used to + look like `OfNat.ofNat (nat_lit 1) x` in the proof term now looks like + `OfNat.ofNat (OfNat.ofNat (α := ℕ) (nat_lit 1)) x`, and this seems to trip up `to_additive`. + Likely this is due to a bug in a dsimproc, which will be easier to diagnose on `nightly-2024-03-12`. + -/ + filter_upwards [hf, hf₂] using fun x (hx : f x = (1 : α → β) x) (hx₂ : f₂ x = (1 : α → γ) x) => by + simp only [hx, hx₂, Pi.one_apply, hm] #align has_compact_mul_support.comp₂_left HasCompactMulSupport.comp₂_left #align has_compact_support.comp₂_left HasCompactSupport.comp₂_left diff --git a/Mathlib/Topology/UniformSpace/Matrix.lean b/Mathlib/Topology/UniformSpace/Matrix.lean index 0f6120404ccd9..14389b1c31926 100644 --- a/Mathlib/Topology/UniformSpace/Matrix.lean +++ b/Mathlib/Topology/UniformSpace/Matrix.lean @@ -29,8 +29,8 @@ instance instUniformAddGroup [AddGroup 𝕜] [UniformAddGroup 𝕜] : theorem uniformity : 𝓤 (Matrix m n 𝕜) = ⨅ (i : m) (j : n), (𝓤 𝕜).comap fun a => (a.1 i j, a.2 i j) := by - erw [Pi.uniformity, Pi.uniformity] - simp_rw [Filter.comap_iInf, Filter.comap_comap] + erw [Pi.uniformity] + simp_rw [Pi.uniformity, Filter.comap_iInf, Filter.comap_comap] rfl #align matrix.uniformity Matrix.uniformity diff --git a/Mathlib/Util/CompileInductive.lean b/Mathlib/Util/CompileInductive.lean index fb8e6b7d9e774..1efbfe091abd5 100644 --- a/Mathlib/Util/CompileInductive.lean +++ b/Mathlib/Util/CompileInductive.lean @@ -228,7 +228,6 @@ elab tk:"compile_inductive% " i:ident : command => Command.liftTermElabM do end Mathlib.Util -compile_inductive% Nat compile_inductive% List compile_inductive% PUnit compile_inductive% PEmpty diff --git a/lean-toolchain b/lean-toolchain index f483baaeec449..8465e8d271c40 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-03-05 +leanprover/lean4:nightly-2024-03-11 From 79f1ba4e870f8d3fb82e6de0636a2dcc6c030601 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 18:55:50 +1100 Subject: [PATCH 02/48] fix from Johan --- Mathlib/Data/Polynomial/Derivative.lean | 52 ++++++++----------------- 1 file changed, 16 insertions(+), 36 deletions(-) diff --git a/Mathlib/Data/Polynomial/Derivative.lean b/Mathlib/Data/Polynomial/Derivative.lean index 515306762d0d1..3e86237ace780 100644 --- a/Mathlib/Data/Polynomial/Derivative.lean +++ b/Mathlib/Data/Polynomial/Derivative.lean @@ -293,42 +293,22 @@ set_option linter.uppercaseLean3 false in #align polynomial.eq_C_of_derivative_eq_zero Polynomial.eq_C_of_derivative_eq_zero @[simp] -theorem derivative_mul {f g : R[X]} : derivative (f * g) = derivative f * g + f * derivative g := - -- FIXME nightly-testing - sorry - -- calc - -- derivative (f * g) = - -- f.sum fun n a => g.sum fun m b => (n + m) • (C (a * b) * X ^ (n + m - 1)) := by - -- rw [mul_eq_sum_sum, derivative_sum] - -- trans - -- · apply Finset.sum_congr rfl - -- intro x _ - -- exact derivative_sum - -- apply Finset.sum_congr rfl; intro n _; apply Finset.sum_congr rfl; intro m _ - -- trans - -- · exact congr_arg _ C_mul_X_pow_eq_monomial.symm - -- dsimp; rw [← smul_mul_assoc, smul_C, nsmul_eq_mul']; exact derivative_C_mul_X_pow _ _ - -- _ = - -- f.sum fun n a => - -- g.sum fun m b => - -- n • (C a * X ^ (n - 1)) * (C b * X ^ m) + C a * X ^ n * m • (C b * X ^ (m - 1)) := - -- (sum_congr rfl fun n hn => - -- sum_congr rfl fun m hm => by - -- 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.add_one_sub_one, zero_smul, add_comm]) - -- _ = derivative f * g + f * derivative g := by - -- conv => - -- rhs - -- congr - -- · rw [← sum_C_mul_X_pow_eq g] - -- · rw [← sum_C_mul_X_pow_eq f] - -- simp only [sum, sum_add_distrib, Finset.mul_sum, Finset.sum_mul, derivative_apply] - -- simp_rw [← smul_mul_assoc, smul_C, nsmul_eq_mul'] - -- rw [Finset.sum_comm] - -- congr 1 - -- rw [Finset.sum_comm] +theorem derivative_mul {f g : R[X]} : derivative (f * g) = derivative f * g + f * derivative g := by + induction f using Polynomial.induction_on' + · simp only [add_mul, map_add, add_assoc, add_left_comm, *] + induction g using Polynomial.induction_on' + · simp only [mul_add, map_add, add_assoc, add_left_comm, *] + rename_i m a n b + simp only [monomial_mul_monomial, derivative_monomial] + simp only [mul_assoc, (Nat.cast_commute _ _).eq, Nat.cast_add, mul_add, map_add] + cases m + · simp only [zero_add, Nat.cast_zero, mul_zero, map_zero] + cases n + · simp only [add_zero, Nat.cast_zero, mul_zero, map_zero] + rename_i m n + simp only [Nat.add_succ_sub_one, add_tsub_cancel_right] + rw [add_assoc, add_comm n 1] +#align polynomial.derivative_mul Polynomial.derivative_mul #align polynomial.derivative_mul Polynomial.derivative_mul theorem derivative_eval (p : R[X]) (x : R) : From def155caa4cb7597520992781c6c5272d029af8c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 18:56:39 +1100 Subject: [PATCH 03/48] update --- lake-manifest.json | 8 ++++---- lakefile.lean | 4 ++-- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 6489b3db11cae..71ccefc17d9c9 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,10 +4,10 @@ [{"url": "https://github.com/leanprover/std4", "type": "git", "subDir": null, - "rev": "bd2ee35b16ffab06eda8906152d7df6ef5da6680", + "rev": "63e50b7099ce9ffea7a1526ee3f045d5d6514999", "name": "std", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "bump/v4.8.0", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/quote4", @@ -22,10 +22,10 @@ {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, - "rev": "cae29ed6fe3b585c8c36d318b752cc49ff817cd1", + "rev": "adaf62adf22094734e9088351dcd684f86c8f48f", "name": "aesop", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "bump/nightly-2024-03-11", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/ProofWidgets4", diff --git a/lakefile.lean b/lakefile.lean index aa8c13b449d6c..dd9f41537d5a8 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -26,9 +26,9 @@ package mathlib where meta if get_config? doc = some "on" then -- do not download and build doc-gen4 by default require «doc-gen4» from git "https://github.com/leanprover/doc-gen4" @ "main" -require std from git "https://github.com/leanprover/std4" @ "nightly-testing" +require std from git "https://github.com/leanprover/std4" @ "bump/v4.8.0" require Qq from git "https://github.com/leanprover-community/quote4" @ "master" -require aesop from git "https://github.com/leanprover-community/aesop" @ "nightly-testing" +require aesop from git "https://github.com/leanprover-community/aesop" @ "bump/nightly-2024-03-11" require proofwidgets from git "https://github.com/leanprover-community/ProofWidgets4" @ "v0.0.28-pre" require Cli from git "https://github.com/leanprover/lean4-cli" @ "main" require importGraph from git "https://github.com/leanprover-community/import-graph.git" @ "main" From 0dec6a8f053e39f11ee9f4595c466ee4b2f8b3d2 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 18:59:01 +1100 Subject: [PATCH 04/48] fix --- Mathlib/Data/Bool/Basic.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 220b0b5da2145..c96d29f59f45e 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -321,10 +321,6 @@ protected def carry (x y c : Bool) := x && y || x && c || y && c #align bitvec.carry Bool.carry -theorem and_xor_distrib_left : ∀ (x y z : Bool), (x && xor y z) = xor (x && y) (x && z) := by - decide -theorem and_xor_distrib_right : ∀ (x y z : Bool), (xor x y && z) = xor (x && z) (y && z) := by - decide #align bool.band_bxor_distrib_left Bool.and_xor_distrib_left #align bool.band_bxor_distrib_right Bool.and_xor_distrib_right From 091af2acb239547753ed1835cb93cbeba8159a83 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 19:00:09 +1100 Subject: [PATCH 05/48] cleanup --- Mathlib/Data/Num/Lemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 8f5429a66f515..a4c8bb33e3f6f 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -985,7 +985,7 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by · rw [PosNum.cast_one', ← bit1_zero, ← Nat.bit_true] -- FIXME adaptation for nightly-2024-03-11 -- Hacky proof, can someone please replace? Obvious simp lemmas seem to be missing. - simp + simp only [Nat.bit_true, bit1_zero, Bool.false_eq] erw [Nat.testBit_succ] simp · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_bit_succ, IH] From 455288195d79c90ee5b2147ea3230d3c8a12d70d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 19:00:13 +1100 Subject: [PATCH 06/48] Update Mathlib/Data/List/Basic.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/Data/List/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index e1cf665369579..1c30997873d06 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1154,7 +1154,7 @@ theorem nthLe_mem (l : List α) (n h) : nthLe l n h ∈ l := get_mem .. #align list.nth_le_mem List.nthLe_mem theorem nthLe_congr {l : List α} {n p : ℕ} {h : n < length l} (hnp : n = p) : - nthLe l n h = nthLe l p (hnp ▸ h) := by simp [hnp] + nthLe l n h = nthLe l p (hnp ▸ h) := by simp [hnp] #align list.nth_mem List.get?_mem From 23784512d18c23eeb0db428d3eac8085e67a4e41 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 19:03:13 +1100 Subject: [PATCH 07/48] long lines --- Mathlib/Data/Nat/Digits.lean | 3 ++- Mathlib/Topology/Support.lean | 3 ++- 2 files changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 912476455d134..dbf5639f09b69 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -271,7 +271,8 @@ theorem ofDigits_digits (b n : ℕ) : ofDigits b (digits b n) = n := by · induction' n with n ih · rfl · rw [show 0 + 1 = 1 by rfl] at ih ⊢ - simp only [Nat.succ_eq_add_one, ih, add_comm 1, ofDigits_one_cons, Nat.cast_id, digits_one_succ] + simp only [Nat.succ_eq_add_one, ih, add_comm 1, ofDigits_one_cons, Nat.cast_id, + digits_one_succ] · apply Nat.strongInductionOn n _ clear n intro n h diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 23d73ecde0eec..20e7d0e49544a 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -273,7 +273,8 @@ theorem comp₂_left (hf : HasCompactMulSupport f) because an `OfNat.ofNat 1` is not replaced with `0`. Notably, as of this nightly, what used to look like `OfNat.ofNat (nat_lit 1) x` in the proof term now looks like `OfNat.ofNat (OfNat.ofNat (α := ℕ) (nat_lit 1)) x`, and this seems to trip up `to_additive`. - Likely this is due to a bug in a dsimproc, which will be easier to diagnose on `nightly-2024-03-12`. + Likely this is due to a bug in a dsimproc, + which will be easier to diagnose on `nightly-2024-03-12`. -/ filter_upwards [hf, hf₂] using fun x (hx : f x = (1 : α → β) x) (hx₂ : f₂ x = (1 : α → γ) x) => by simp only [hx, hx₂, Pi.one_apply, hm] From 411f6c57eded6f2bbb4cdb093f9edfed1a5b00bf Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 19:23:01 +1100 Subject: [PATCH 08/48] Update Mathlib/Data/Nat/Digits.lean --- Mathlib/Data/Nat/Digits.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index dbf5639f09b69..806ad9072e562 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -568,7 +568,7 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits have h₁ : 1 ≤ tl.length := List.length_pos.mpr h' rw [← sum_range_add_sum_Ico _ <| h₁, ← add_zero (∑ x in Ico _ _, ofDigits p (tl.drop x)), ← this, sum_Ico_consecutive _ h₁ <| (le_add_right (List.length tl) 1)] - -- Adaptatin note: nightly-2024-03-07: this needs an `erw` only because of a `0 + 1` vs `1`. + -- Adaptation note: nightly-2024-03-07: this needs an `erw` only because of a `0 + 1` vs `1`. -- Can someone do it more cleanly? erw [← sum_Ico_add _ 0 tl.length 1] rw [Ico_zero_eq_range, mul_add, mul_add, ih, range_one, sum_singleton, List.drop, ofDigits, From 88ab48e29fc35b8c86596dd527f6ba5a033ea2ed Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 12 Mar 2024 19:25:49 +1100 Subject: [PATCH 09/48] expand comments --- Mathlib/Algebra/Homology/ExactSequence.lean | 3 +++ Mathlib/Algebra/Homology/HomologySequence.lean | 6 ++++++ Mathlib/Analysis/BoundedVariation.lean | 2 +- Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean | 3 +++ 4 files changed, 13 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Homology/ExactSequence.lean b/Mathlib/Algebra/Homology/ExactSequence.lean index 2ee0cba3690f0..dd545d5edae80 100644 --- a/Mathlib/Algebra/Homology/ExactSequence.lean +++ b/Mathlib/Algebra/Homology/ExactSequence.lean @@ -182,6 +182,9 @@ lemma isComplex₂_mk (S : ComposableArrows C 2) (w : S.map' 0 1 ≫ S.map' 1 2 -- Adaptation note: nightly-2024-03-11 -- We turn off simprocs here. +-- Ideally someone will investigate whether `simp` lemmas can be rearranged +-- so that this works without the `set_option`, +-- *or* come up with a proposal regarding finer control of disabling simprocs. set_option simprocs false in lemma _root_.CategoryTheory.ShortComplex.isComplex_toComposableArrows (S : ShortComplex C) : S.toComposableArrows.IsComplex := diff --git a/Mathlib/Algebra/Homology/HomologySequence.lean b/Mathlib/Algebra/Homology/HomologySequence.lean index e29e19c2b1459..dc5f8c0d27e0d 100644 --- a/Mathlib/Algebra/Homology/HomologySequence.lean +++ b/Mathlib/Algebra/Homology/HomologySequence.lean @@ -112,6 +112,9 @@ instance [K.HasHomology i] [K.HasHomology j] : -- Adaptation note: nightly-2024-03-11 -- We turn off simprocs here. +-- Ideally someone will investigate whether `simp` lemmas can be rearranged +-- so that this works without the `set_option`, +-- *or* come up with a proposal regarding finer control of disabling simprocs. set_option simprocs false in instance [K.HasHomology i] [K.HasHomology j] : Epi ((composableArrows₃ K i j).map' 2 3) := by @@ -150,6 +153,9 @@ attribute [local simp] homologyMap_comp cyclesMap_comp opcyclesMap_comp -- Adaptation note: nightly-2024-03-11 -- We turn off simprocs here. +-- Ideally someone will investigate whether `simp` lemmas can be rearranged +-- so that this works without the `set_option`, +-- *or* come up with a proposal regarding finer control of disabling simprocs. set_option simprocs false in /-- The functor `HomologicalComplex C c ⥤ ComposableArrows C 3` that maps `K` to the diagram `K.homology i ⟶ K.opcycles i ⟶ K.cycles j ⟶ K.homology j`. -/ diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index 371de48c8df18..2d36373a1a5cd 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -528,7 +528,7 @@ theorem comp_le_of_antitoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β rw [edist_comm, Nat.sub_sub, add_comm, Nat.sub_succ, Nat.add_one, Nat.succ_eq_add_one] simp only [Function.comp_apply, Nat.pred_eq_sub_one, Nat.sub_add_eq] congr - simp at hx + simp? at hx omega #align evariation_on.comp_le_of_antitone_on eVariationOn.comp_le_of_antitoneOn diff --git a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean index b9c4cad9cd15f..1fa0696184883 100644 --- a/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean +++ b/Mathlib/CategoryTheory/Abelian/DiagramLemmas/Four.lean @@ -135,6 +135,9 @@ variable {R₁ R₂ : ComposableArrows C 4} (hR₁ : R₁.Exact) (hR₂ : R₂.E -- Adaptation note: nightly-2024-03-11 -- We turn off simprocs here. +-- Ideally someone will investigate whether `simp` lemmas can be rearranged +-- so that this works without the `set_option`, +-- *or* come up with a proposal regarding finer control of disabling simprocs. set_option simprocs false in /-- The five lemma. -/ theorem isIso_of_epi_of_isIso_of_isIso_of_mono (h₀ : Epi (app' φ 0)) (h₁ : IsIso (app' φ 1)) From c33cdb8ae10d6713854ecc12bfcbefe1ecbdcfdc Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Tue, 12 Mar 2024 09:45:14 +0100 Subject: [PATCH 10/48] double align --- Mathlib/Data/Polynomial/Derivative.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Data/Polynomial/Derivative.lean b/Mathlib/Data/Polynomial/Derivative.lean index 3e86237ace780..ee0c0cd104c9e 100644 --- a/Mathlib/Data/Polynomial/Derivative.lean +++ b/Mathlib/Data/Polynomial/Derivative.lean @@ -309,7 +309,6 @@ theorem derivative_mul {f g : R[X]} : derivative (f * g) = derivative f * g + f simp only [Nat.add_succ_sub_one, add_tsub_cancel_right] rw [add_assoc, add_comm n 1] #align polynomial.derivative_mul Polynomial.derivative_mul -#align polynomial.derivative_mul Polynomial.derivative_mul theorem derivative_eval (p : R[X]) (x : R) : p.derivative.eval x = p.sum fun n a => a * n * x ^ (n - 1) := by From 7befb8fb6a6eece4202f8e56d8b00ee09c782a5b Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Tue, 12 Mar 2024 10:14:45 +0100 Subject: [PATCH 11/48] fix this --- .../Polynomial/Eisenstein/IsIntegral.lean | 265 +++++++++--------- 1 file changed, 130 insertions(+), 135 deletions(-) diff --git a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean index d1e451eb320de..ffe7ed1abee85 100644 --- a/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean +++ b/Mathlib/RingTheory/Polynomial/Eisenstein/IsIntegral.lean @@ -240,141 +240,136 @@ theorem mem_adjoin_of_smul_prime_smul_of_minpoly_isEisensteinAt {B : PowerBasis (hp : Prime p) (hBint : IsIntegral R B.gen) {z : L} (hzint : IsIntegral R z) (hz : p • z ∈ adjoin R ({B.gen} : Set L)) (hei : (minpoly R B.gen).IsEisensteinAt 𝓟) : z ∈ adjoin R ({B.gen} : Set L) := by - -- FIXME nightly-testing: proof broken, - -- perhaps can be split into smaller pieces or made more robust? - sorry - -- -- First define some abbreviations. - -- have hndiv : ¬p ^ 2 ∣ (minpoly R B.gen).coeff 0 := fun h => - -- hei.not_mem ((span_singleton_pow p 2).symm ▸ Ideal.mem_span_singleton.2 h) - -- have := B.finite - -- set P := minpoly R B.gen with hP - -- obtain ⟨n, hn⟩ := Nat.exists_eq_succ_of_ne_zero B.dim_pos.ne' - -- haveI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans R K L - -- let _ := P.map (algebraMap R L) - -- -- There is a polynomial `Q` such that `p • z = aeval B.gen Q`. We can assume that - -- -- `Q.degree < P.degree` and `Q ≠ 0`. - -- rw [adjoin_singleton_eq_range_aeval] at hz - -- obtain ⟨Q₁, hQ⟩ := hz - -- set Q := Q₁ %ₘ P with hQ₁ - -- replace hQ : aeval B.gen Q = p • z := by - -- rw [← modByMonic_add_div Q₁ (minpoly.monic hBint)] at hQ - -- simpa using hQ - -- by_cases hQzero : Q = 0 - -- · simp only [hQzero, Algebra.smul_def, zero_eq_mul, aeval_zero] at hQ - -- cases' hQ with H H₁ - -- · have : Function.Injective (algebraMap R L) := by - -- rw [algebraMap_eq R K L] - -- exact (algebraMap K L).injective.comp (IsFractionRing.injective R K) - -- exfalso - -- exact hp.ne_zero ((injective_iff_map_eq_zero _).1 this _ H) - -- · rw [H₁] - -- exact Subalgebra.zero_mem _ - -- -- It is enough to prove that all coefficients of `Q` are divisible by `p`, by induction. - -- -- The base case is `dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt`. - -- refine mem_adjoin_of_dvd_coeff_of_dvd_aeval hp.ne_zero (fun i => ?_) hQ - -- induction' i using Nat.case_strong_induction_on with j hind - -- · intro _ - -- exact dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt hp hBint hQ hzint hei - -- · intro hj - -- convert hp.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd _ hndiv - -- exact n - -- -- Two technical results we will need about `P.natDegree` and `Q.natDegree`. - -- have H := degree_modByMonic_lt Q₁ (minpoly.monic hBint) - -- rw [← hQ₁, ← hP] at H - -- replace H := Nat.lt_iff_add_one_le.1 - -- (lt_of_lt_of_le - -- (lt_of_le_of_lt (Nat.lt_iff_add_one_le.1 (Nat.lt_of_succ_lt_succ (mem_range.1 hj))) - -- (lt_succ_self _)) (Nat.lt_iff_add_one_le.1 ((natDegree_lt_natDegree_iff hQzero).2 H))) - -- rw [add_assoc] at H - -- have Hj : Q.natDegree + 1 = j + 1 + (Q.natDegree - j) := by - -- rw [← add_comm 1, ← add_comm 1, add_assoc, add_right_inj, - -- ← Nat.add_sub_assoc (Nat.lt_of_succ_lt_succ (mem_range.1 hj)).le, add_comm, - -- Nat.add_sub_cancel] - -- -- By induction hypothesis we can find `g : ℕ → R` such that - -- -- `k ∈ range (j + 1) → Q.coeff k • B.gen ^ k = (algebraMap R L) p * g k • B.gen ^ k`- - -- choose! g hg using hind - -- replace hg : ∀ k ∈ range (j + 1), Q.coeff k • B.gen ^ k = - -- algebraMap R L p * g k • B.gen ^ k := by - -- intro k hk - -- rw [hg k (mem_range_succ_iff.1 hk) - -- (mem_range_succ_iff.2 - -- (le_trans (mem_range_succ_iff.1 hk) (succ_le_iff.1 (mem_range_succ_iff.1 hj)).le)), - -- Algebra.smul_def, Algebra.smul_def, RingHom.map_mul, mul_assoc] - -- -- Since `minpoly R B.gen` is Eiseinstein, we can find `f : ℕ → L` such that - -- -- `(map (algebraMap R L) (minpoly R B.gen)).nat_degree ≤ i` implies `f i ∈ adjoin R {B.gen}` - -- -- and `(algebraMap R L) p * f i = B.gen ^ i`. We will also need `hf₁`, a reformulation of this - -- -- property. - -- choose! f hf using - -- IsWeaklyEisensteinAt.exists_mem_adjoin_mul_eq_pow_natDegree_le (minpoly.aeval R B.gen) - -- (minpoly.monic hBint) hei.isWeaklyEisensteinAt - -- have hf₁ : ∀ k ∈ (range (Q.natDegree - j)).erase 0, - -- Q.coeff (j + 1 + k) • B.gen ^ (j + 1 + k) * B.gen ^ (P.natDegree - (j + 2)) = - -- (algebraMap R L) p * Q.coeff (j + 1 + k) • f (k + P.natDegree - 1) := by - -- intro k hk - -- rw [smul_mul_assoc, ← pow_add, ← Nat.add_sub_assoc H, ← add_assoc j 1 1, add_comm (j + 1) 1, - -- add_assoc (j + 1), add_comm _ (k + P.natDegree), Nat.add_sub_add_right, - -- ← (hf (k + P.natDegree - 1) _).2, mul_smul_comm] - -- rw [(minpoly.monic hBint).natDegree_map, add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] - -- · exact Nat.zero_le _ - -- · refine one_le_iff_ne_zero.2 fun h => ?_ - -- rw [h] at hk - -- simp at hk - - -- -- The Eisenstein condition shows that `p` divides `Q.coeff j` - -- -- if `p^n.succ` divides the following multiple of `Q.coeff (succ j)^n.succ`: - -- suffices - -- p ^ n.succ ∣ Q.coeff (succ j) ^ n.succ * - -- (minpoly R B.gen).coeff 0 ^ (succ j + (P.natDegree - (j + 2))) by - -- convert this - -- rw [Nat.succ_eq_add_one, add_assoc, ← Nat.add_sub_assoc H, ← add_assoc, add_comm (j + 1), - -- Nat.add_sub_add_left, ← Nat.add_sub_assoc, Nat.add_sub_add_left, hP, ← - -- (minpoly.monic hBint).natDegree_map (algebraMap R K), ← - -- minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, natDegree_minpoly, hn, Nat.sub_one, - -- Nat.pred_succ] - -- omega - - -- -- Using `hQ : aeval B.gen Q = p • z`, we write `p • z` as a sum of terms of degree less than - -- -- `j+1`, that are multiples of `p` by induction, and terms of degree at least `j+1`. - -- rw [aeval_eq_sum_range, Hj, range_add, sum_union (disjoint_range_addLeftEmbedding _ _), - -- sum_congr rfl hg, add_comm] at hQ - -- -- We multiply this equality by `B.gen ^ (P.natDegree-(j+2))`, so we can use `hf₁` on the terms - -- -- we didn't know were multiples of `p`, and we take the norm on both sides. - -- replace hQ := congr_arg (fun x => x * B.gen ^ (P.natDegree - (j + 2))) hQ - -- simp_rw [sum_map, addLeftEmbedding_apply, add_mul, sum_mul, mul_assoc] at hQ - -- rw [← insert_erase - -- (mem_range.2 (tsub_pos_iff_lt.2 <| Nat.lt_of_succ_lt_succ <| mem_range.1 hj)), - -- sum_insert (not_mem_erase 0 _), add_zero, sum_congr rfl hf₁, ← mul_sum, ← mul_sum, add_assoc, - -- ← mul_add, smul_mul_assoc, ← pow_add, Algebra.smul_def] at hQ - -- replace hQ := congr_arg (norm K) (eq_sub_of_add_eq hQ) - - -- -- We obtain an equality of elements of `K`, but everything is integral, so we can move to `R` - -- -- and simplify `hQ`. - -- have hintsum : IsIntegral R (z * B.gen ^ (P.natDegree - (j + 2)) - - -- (∑ x : ℕ in (range (Q.natDegree - j)).erase 0, - -- Q.coeff (j + 1 + x) • f (x + P.natDegree - 1) + - -- ∑ x : ℕ in range (j + 1), g x • B.gen ^ x * B.gen ^ (P.natDegree - (j + 2)))) := by - -- refine (hzint.mul (hBint.pow _)).sub - -- (.add (.sum _ fun k hk => .smul _ ?_) - -- (.sum _ fun k _ => .mul (.smul _ (.pow hBint _)) (hBint.pow _))) - -- refine adjoin_le_integralClosure hBint (hf _ ?_).1 - -- rw [(minpoly.monic hBint).natDegree_map (algebraMap R L)] - -- rw [add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] - -- · exact _root_.zero_le _ - -- · refine one_le_iff_ne_zero.2 fun h => ?_ - -- rw [h] at hk - -- simp at hk - -- obtain ⟨r, hr⟩ := isIntegral_iff.1 (isIntegral_norm K hintsum) - -- rw [Algebra.smul_def, mul_assoc, ← mul_sub, _root_.map_mul, algebraMap_apply R K L, map_pow, - -- Algebra.norm_algebraMap, _root_.map_mul, algebraMap_apply R K L, Algebra.norm_algebraMap, - -- finrank B, ← hr, PowerBasis.norm_gen_eq_coeff_zero_minpoly, - -- minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, coeff_map, - -- show (-1 : K) = algebraMap R K (-1) by simp, ← map_pow, ← map_pow, ← _root_.map_mul, ← - -- map_pow, ← _root_.map_mul, ← map_pow, ← _root_.map_mul] at hQ - -- -- We can now finish the proof. - -- have hppdiv : p ^ B.dim ∣ p ^ B.dim * r := dvd_mul_of_dvd_left dvd_rfl _ - -- rwa [← IsFractionRing.injective R K hQ, mul_comm, ← Units.coe_neg_one, mul_pow, ← - -- Units.val_pow_eq_pow_val, ← Units.val_pow_eq_pow_val, mul_assoc, - -- Units.dvd_mul_left, mul_comm, ← Nat.succ_eq_add_one, hn] at hppdiv + -- First define some abbreviations. + have hndiv : ¬p ^ 2 ∣ (minpoly R B.gen).coeff 0 := fun h => + hei.not_mem ((span_singleton_pow p 2).symm ▸ Ideal.mem_span_singleton.2 h) + have := B.finite + set P := minpoly R B.gen with hP + obtain ⟨n, hn⟩ := Nat.exists_eq_succ_of_ne_zero B.dim_pos.ne' + haveI : NoZeroSMulDivisors R L := NoZeroSMulDivisors.trans R K L + let _ := P.map (algebraMap R L) + -- There is a polynomial `Q` such that `p • z = aeval B.gen Q`. We can assume that + -- `Q.degree < P.degree` and `Q ≠ 0`. + rw [adjoin_singleton_eq_range_aeval] at hz + obtain ⟨Q₁, hQ⟩ := hz + set Q := Q₁ %ₘ P with hQ₁ + replace hQ : aeval B.gen Q = p • z := by + rw [← modByMonic_add_div Q₁ (minpoly.monic hBint)] at hQ + simpa using hQ + by_cases hQzero : Q = 0 + · simp only [hQzero, Algebra.smul_def, zero_eq_mul, aeval_zero] at hQ + cases' hQ with H H₁ + · have : Function.Injective (algebraMap R L) := by + rw [algebraMap_eq R K L] + exact (algebraMap K L).injective.comp (IsFractionRing.injective R K) + exfalso + exact hp.ne_zero ((injective_iff_map_eq_zero _).1 this _ H) + · rw [H₁] + exact Subalgebra.zero_mem _ + -- It is enough to prove that all coefficients of `Q` are divisible by `p`, by induction. + -- The base case is `dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt`. + refine mem_adjoin_of_dvd_coeff_of_dvd_aeval hp.ne_zero (fun i => ?_) hQ + induction' i using Nat.case_strong_induction_on with j hind + · intro _ + exact dvd_coeff_zero_of_aeval_eq_prime_smul_of_minpoly_isEisensteinAt hp hBint hQ hzint hei + · intro hj + convert hp.dvd_of_pow_dvd_pow_mul_pow_of_square_not_dvd (n := n) _ hndiv + -- Two technical results we will need about `P.natDegree` and `Q.natDegree`. + have H := degree_modByMonic_lt Q₁ (minpoly.monic hBint) + rw [← hQ₁, ← hP] at H + replace H := Nat.lt_iff_add_one_le.1 + (lt_of_lt_of_le + (lt_of_le_of_lt (Nat.lt_iff_add_one_le.1 (Nat.lt_of_succ_lt_succ (mem_range.1 hj))) + (lt_succ_self _)) (Nat.lt_iff_add_one_le.1 ((natDegree_lt_natDegree_iff hQzero).2 H))) + have Hj : Q.natDegree + 1 = j + 1 + (Q.natDegree - j) := by + rw [← add_comm 1, ← add_comm 1, add_assoc, add_right_inj, + ← Nat.add_sub_assoc (Nat.lt_of_succ_lt_succ (mem_range.1 hj)).le, add_comm, + Nat.add_sub_cancel] + -- By induction hypothesis we can find `g : ℕ → R` such that + -- `k ∈ range (j + 1) → Q.coeff k • B.gen ^ k = (algebraMap R L) p * g k • B.gen ^ k`- + choose! g hg using hind + replace hg : ∀ k ∈ range (j + 1), Q.coeff k • B.gen ^ k = + algebraMap R L p * g k • B.gen ^ k := by + intro k hk + rw [hg k (mem_range_succ_iff.1 hk) + (mem_range_succ_iff.2 + (le_trans (mem_range_succ_iff.1 hk) (succ_le_iff.1 (mem_range_succ_iff.1 hj)).le)), + Algebra.smul_def, Algebra.smul_def, RingHom.map_mul, mul_assoc] + -- Since `minpoly R B.gen` is Eiseinstein, we can find `f : ℕ → L` such that + -- `(map (algebraMap R L) (minpoly R B.gen)).nat_degree ≤ i` implies `f i ∈ adjoin R {B.gen}` + -- and `(algebraMap R L) p * f i = B.gen ^ i`. We will also need `hf₁`, a reformulation of this + -- property. + choose! f hf using + IsWeaklyEisensteinAt.exists_mem_adjoin_mul_eq_pow_natDegree_le (minpoly.aeval R B.gen) + (minpoly.monic hBint) hei.isWeaklyEisensteinAt + have hf₁ : ∀ k ∈ (range (Q.natDegree - j)).erase 0, + Q.coeff (j + 1 + k) • B.gen ^ (j + 1 + k) * B.gen ^ (P.natDegree - (j + 2)) = + (algebraMap R L) p * Q.coeff (j + 1 + k) • f (k + P.natDegree - 1) := by + intro k hk + rw [smul_mul_assoc, ← pow_add, ← Nat.add_sub_assoc H, add_comm (j + 1) 1, + add_assoc (j + 1), add_comm _ (k + P.natDegree), Nat.add_sub_add_right, + ← (hf (k + P.natDegree - 1) _).2, mul_smul_comm] + rw [(minpoly.monic hBint).natDegree_map, add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] + · exact Nat.zero_le _ + · refine one_le_iff_ne_zero.2 fun h => ?_ + rw [h] at hk + simp at hk + + -- The Eisenstein condition shows that `p` divides `Q.coeff j` + -- if `p^n.succ` divides the following multiple of `Q.coeff (succ j)^n.succ`: + suffices + p ^ n.succ ∣ Q.coeff (succ j) ^ n.succ * + (minpoly R B.gen).coeff 0 ^ (succ j + (P.natDegree - (j + 2))) by + convert this + rw [Nat.succ_eq_add_one, add_assoc, ← Nat.add_sub_assoc H, add_comm (j + 1), + Nat.add_sub_add_left, ← Nat.add_sub_assoc, Nat.add_sub_add_left, hP, ← + (minpoly.monic hBint).natDegree_map (algebraMap R K), ← + minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, natDegree_minpoly, hn, Nat.sub_one, + Nat.pred_succ] + omega + + -- Using `hQ : aeval B.gen Q = p • z`, we write `p • z` as a sum of terms of degree less than + -- `j+1`, that are multiples of `p` by induction, and terms of degree at least `j+1`. + rw [aeval_eq_sum_range, Hj, range_add, sum_union (disjoint_range_addLeftEmbedding _ _), + sum_congr rfl hg, add_comm] at hQ + -- We multiply this equality by `B.gen ^ (P.natDegree-(j+2))`, so we can use `hf₁` on the terms + -- we didn't know were multiples of `p`, and we take the norm on both sides. + replace hQ := congr_arg (fun x => x * B.gen ^ (P.natDegree - (j + 2))) hQ + simp_rw [sum_map, addLeftEmbedding_apply, add_mul, sum_mul, mul_assoc] at hQ + rw [← insert_erase + (mem_range.2 (tsub_pos_iff_lt.2 <| Nat.lt_of_succ_lt_succ <| mem_range.1 hj)), + sum_insert (not_mem_erase 0 _), add_zero, sum_congr rfl hf₁, ← mul_sum, ← mul_sum, add_assoc, + ← mul_add, smul_mul_assoc, ← pow_add, Algebra.smul_def] at hQ + replace hQ := congr_arg (norm K) (eq_sub_of_add_eq hQ) + + -- We obtain an equality of elements of `K`, but everything is integral, so we can move to `R` + -- and simplify `hQ`. + have hintsum : IsIntegral R (z * B.gen ^ (P.natDegree - (j + 2)) - + (∑ x : ℕ in (range (Q.natDegree - j)).erase 0, + Q.coeff (j + 1 + x) • f (x + P.natDegree - 1) + + ∑ x : ℕ in range (j + 1), g x • B.gen ^ x * B.gen ^ (P.natDegree - (j + 2)))) := by + refine (hzint.mul (hBint.pow _)).sub + (.add (.sum _ fun k hk => .smul _ ?_) + (.sum _ fun k _ => .mul (.smul _ (.pow hBint _)) (hBint.pow _))) + refine adjoin_le_integralClosure hBint (hf _ ?_).1 + rw [(minpoly.monic hBint).natDegree_map (algebraMap R L)] + rw [add_comm, Nat.add_sub_assoc, le_add_iff_nonneg_right] + · exact _root_.zero_le _ + · refine one_le_iff_ne_zero.2 fun h => ?_ + rw [h] at hk + simp at hk + obtain ⟨r, hr⟩ := isIntegral_iff.1 (isIntegral_norm K hintsum) + rw [Algebra.smul_def, mul_assoc, ← mul_sub, _root_.map_mul, algebraMap_apply R K L, map_pow, + Algebra.norm_algebraMap, _root_.map_mul, algebraMap_apply R K L, Algebra.norm_algebraMap, + finrank B, ← hr, PowerBasis.norm_gen_eq_coeff_zero_minpoly, + minpoly.isIntegrallyClosed_eq_field_fractions' K hBint, coeff_map, + show (-1 : K) = algebraMap R K (-1) by simp, ← map_pow, ← map_pow, ← _root_.map_mul, ← + map_pow, ← _root_.map_mul, ← map_pow, ← _root_.map_mul] at hQ + -- We can now finish the proof. + have hppdiv : p ^ B.dim ∣ p ^ B.dim * r := dvd_mul_of_dvd_left dvd_rfl _ + rwa [← IsFractionRing.injective R K hQ, mul_comm, ← Units.coe_neg_one, mul_pow, ← + Units.val_pow_eq_pow_val, ← Units.val_pow_eq_pow_val, mul_assoc, + Units.dvd_mul_left, mul_comm, ← Nat.succ_eq_add_one, hn] at hppdiv #align mem_adjoin_of_smul_prime_smul_of_minpoly_is_eiseinstein_at mem_adjoin_of_smul_prime_smul_of_minpoly_isEisensteinAt /-- Let `K` be the field of fraction of an integrally closed domain `R` and let `L` be a separable From 527f2a9a19e1e12627c510250d127166564e0360 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 12 Mar 2024 10:39:01 +0000 Subject: [PATCH 12/48] fix sorry --- .../LinearAlgebra/TensorProduct/Graded/Internal.lean | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index aed01b8b08b10..1af386f774405 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -120,6 +120,8 @@ notation:100 x " ᵍ⊗ₜ" y:100 => tmul _ x y @[inherit_doc] notation:100 x " ᵍ⊗ₜ[" R "] " y:100 => tmul R x y +theorem one_def : (1 : 𝒜 ᵍ⊗[R] ℬ) = 1 ᵍ⊗ₜ 1 := rfl + variable (R) in /-- An auxiliary construction to move between the graded tensor product of internally-graded objects and the tensor product of direct sums. -/ @@ -307,10 +309,10 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) ∘ₗ (TensorProduct.map f.toLinearMap g.toLinearMap) ∘ₗ ((of R 𝒜 ℬ).symm : 𝒜 ᵍ⊗[R] ℬ →ₗ[R] A ⊗[R] B)) (by - -- FIXME nightly-testing: not sure what is happening here - sorry - -- dsimp [Algebra.TensorProduct.one_def]; simp only [_root_.map_one, mul_one] - ) + dsimp [Algebra.TensorProduct.one_def] + -- FIXME nightly-testing: no longer works with dsimp, even though it is a rfl lemma + rw [Algebra.TensorProduct.one_def] + dsimp; simp only [_root_.map_one, mul_one]) (by rw [LinearMap.map_mul_iff] ext a₁ : 3 From 0a03277c3f7cd8be985fce5033eac8ef19db7e2a Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Tue, 12 Mar 2024 09:43:18 -0700 Subject: [PATCH 13/48] Algebra.Homology.Augment, use honest `fun` instead of `by cases` --- Mathlib/Algebra/Homology/Augment.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index 4784b13c8d1c4..5197f4144f58d 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -140,7 +140,7 @@ theorem chainComplex_d_succ_succ_zero (C : ChainComplex V ℕ) (i : ℕ) : C.d ( def augmentTruncate (C : ChainComplex V ℕ) : augment (truncate.obj C) (C.d 1 0) (C.d_comp_d _ _ _) ≅ C where hom := - { f := fun i => by cases i <;> exact 𝟙 _ + { f := fun | 0 => 𝟙 _ | n+1 => 𝟙 _ comm' := fun i j => by -- Porting note: was an rcases n with (_|_|n) but that was causing issues match i with @@ -148,7 +148,7 @@ def augmentTruncate (C : ChainComplex V ℕ) : cases' j with j <;> dsimp [augment, truncate] <;> simp } inv := - { f := fun i => by cases i <;> exact 𝟙 _ + { f := fun | 0 => 𝟙 _ | n+1 => 𝟙 _ comm' := fun i j => by -- Porting note: was an rcases n with (_|_|n) but that was causing issues match i with @@ -335,14 +335,14 @@ theorem cochainComplex_d_succ_succ_zero (C : CochainComplex V ℕ) (i : ℕ) : C def augmentTruncate (C : CochainComplex V ℕ) : augment (truncate.obj C) (C.d 0 1) (C.d_comp_d _ _ _) ≅ C where hom := - { f := fun i => by cases i <;> exact 𝟙 _ + { f := fun | 0 => 𝟙 _ | n+1 => 𝟙 _ comm' := fun i j => by rcases j with (_ | _ | j) <;> cases i <;> · dsimp -- Porting note (#10959): simp can't handle this now but aesop does aesop } inv := - { f := fun i => by cases i <;> exact 𝟙 _ + { f := fun | 0 => 𝟙 _ | n+1 => 𝟙 _ comm' := fun i j => by rcases j with (_ | _ | j) <;> cases' i with i <;> · dsimp From 6a48fa22c06de9e70defb7b923c2850d63370b9f Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 12 Mar 2024 21:36:05 +0100 Subject: [PATCH 14/48] says --- .../HomotopyCategory/Pretriangulated.lean | 20 +++++++++---------- Mathlib/SetTheory/Ordinal/Notation.lean | 6 +++--- 2 files changed, 13 insertions(+), 13 deletions(-) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index 2c2f5bded8af1..9e512e0234ef2 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -183,7 +183,7 @@ noncomputable def rotateHomotopyEquiv : hom := lift (inr φ) (-(Cocycle.ofHom φ).leftShift 1 1 (zero_add 1)) (-(inl φ).leftShift 1 0 (neg_add_self 1)) (by simp? [Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1)] says - simp only [δ_neg, Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1), + simp only [Int.reduceNeg, δ_neg, Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1), Int.negOnePow_one, δ_inl, Cochain.ofHom_comp, Cochain.leftShift_comp_zero_cochain, Units.neg_smul, one_smul, neg_neg, Cocycle.coe_neg, Cocycle.leftShift_coe, Cocycle.ofHom_coe, Cochain.neg_comp, add_right_neg]) @@ -193,7 +193,7 @@ noncomputable def rotateHomotopyEquiv : ext n simp? [lift_desc_f _ _ _ _ _ _ _ _ _ rfl, (inl φ).leftShift_v 1 0 _ _ n _ (n + 1) (by simp only [add_neg_cancel_right])] says - simp only [shiftFunctor_obj_X', HomologicalComplex.comp_f, + simp only [shiftFunctor_obj_X', Int.reduceNeg, HomologicalComplex.comp_f, lift_desc_f _ _ _ _ _ _ _ _ _ rfl, Cocycle.coe_neg, Cocycle.leftShift_coe, Cocycle.ofHom_coe, Cochain.neg_v, Cochain.zero_v, comp_zero, (inl φ).leftShift_v 1 0 _ _ n _ (n + 1) (by simp only [add_neg_cancel_right]), @@ -209,10 +209,10 @@ noncomputable def rotateHomotopyEquiv : (inl φ).leftShift_v 1 0 (neg_add_self 1) n n (add_zero n) (n + 1) (by omega), (Cochain.ofHom φ).leftShift_v 1 1 (zero_add 1) n (n + 1) rfl (n + 1) (by omega), Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n rfl (by omega)] says - simp only [Cochain.ofHom_comp, ofHom_desc, ofHom_lift, Cocycle.coe_neg, + simp only [Int.reduceNeg, Cochain.ofHom_comp, ofHom_desc, ofHom_lift, Cocycle.coe_neg, Cocycle.leftShift_coe, Cocycle.ofHom_coe, Cochain.comp_zero_cochain_v, shiftFunctor_obj_X', δ_neg, δ_zero_cochain_comp _ _ _ (neg_add_self 1), δ_inl, - Int.reduceNeg, Int.negOnePow_neg, Int.negOnePow_one, δ_snd, Cochain.neg_comp, + Int.negOnePow_neg, Int.negOnePow_one, δ_snd, Cochain.neg_comp, Cochain.comp_assoc_of_second_is_zero_cochain, smul_neg, Units.neg_smul, one_smul, neg_neg, Cochain.comp_add, inr_snd_assoc, neg_add_rev, Cochain.add_v, Cochain.neg_v, Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n rfl (by omega), @@ -290,14 +290,14 @@ noncomputable def shiftIso (n : ℤ) : (mappingCone φ)⟦n⟧ ≅ mappingCone ( ext p dsimp simp? [Cochain.shift_v', smul_smul] says - simp only [δ_units_smul, Cochain.δ_shift, δ_inl, Cochain.ofHom_comp, smul_smul, + simp only [Int.reduceNeg, δ_units_smul, Cochain.δ_shift, δ_inl, Cochain.ofHom_comp, smul_smul, Int.units_mul_self, one_smul, Cochain.shift_v', Cochain.comp_zero_cochain_v, Cochain.ofHom_v, shiftFunctor_obj_X', shiftFunctor_map_f']) hom_inv_id := by ext p dsimp simp? [lift_desc_f _ _ _ _ _ _ _ _ (p + 1) rfl, id_X, smul_smul, Cochain.shift_v'] says - simp only [lift_desc_f _ _ _ _ _ _ _ _ (p + 1) rfl, shiftFunctor_obj_X', + simp only [Int.reduceNeg, lift_desc_f _ _ _ _ _ _ _ _ (p + 1) rfl, shiftFunctor_obj_X', Cocycle.coe_units_smul, Cocycle.shift_coe, Cochain.units_smul_v, Cochain.shift_v', Linear.comp_units_smul, Linear.units_smul_comp, smul_smul, Int.units_mul_self, one_smul, shiftFunctor_map_f', id_X] @@ -306,8 +306,8 @@ noncomputable def shiftIso (n : ℤ) : (mappingCone φ)⟦n⟧ ≅ mappingCone ( dsimp simp? [ext_from_iff _ (p + 1) _ rfl, ext_to_iff _ _ (p + 1) rfl, Cochain.shift_v', smul_smul] says - simp only [ext_from_iff _ (p + 1) _ rfl, shiftFunctor_obj_X', inl_v_desc_f_assoc, - Cochain.units_smul_v, Cochain.shift_v', Linear.units_smul_comp, comp_id, + simp only [Int.reduceNeg, ext_from_iff _ (p + 1) _ rfl, shiftFunctor_obj_X', + inl_v_desc_f_assoc, Cochain.units_smul_v, Cochain.shift_v', Linear.units_smul_comp, comp_id, ext_to_iff _ _ (p + 1) rfl, assoc, lift_f_fst_v, Cocycle.coe_units_smul, Cocycle.shift_coe, Linear.comp_units_smul, inl_v_fst_v, smul_smul, Int.units_mul_self, one_smul, lift_f_snd_v, inl_v_snd_v, smul_zero, and_self, @@ -327,7 +327,7 @@ noncomputable def shiftTriangleIso (n : ℤ) : set_option tactic.skipAssignedInstances false in dsimp simp? [shiftIso, Units.smul_def, ext_to_iff _ _ (p + 1) rfl, Cochain.shift_v'] says - simp only [Units.smul_def, shiftIso, Linear.smul_comp, id_comp, + simp only [Units.smul_def, shiftIso, Int.reduceNeg, Linear.smul_comp, id_comp, ext_to_iff _ _ (p + 1) rfl, shiftFunctor_obj_X', assoc, lift_f_fst_v, Cocycle.coe_smul, Cocycle.shift_coe, Cochain.smul_v, Cochain.shift_v', Linear.comp_smul, inr_f_fst_v, smul_zero, lift_f_snd_v, inr_f_snd_v, and_true] @@ -343,7 +343,7 @@ noncomputable def shiftTriangleIso (n : ℤ) : (fst φ).1.rightShift_v 1 0 (zero_add 1) (p + n) (p + n) (add_zero (p + n)) (p + 1 + n) (by omega), shiftFunctor_obj_X, shiftFunctorObjXIso, shiftFunctorComm_hom_app_f, Preadditive.neg_comp, - assoc, Iso.inv_hom_id, comp_id, smul_neg, Units.smul_def, shiftIso, + assoc, Iso.inv_hom_id, comp_id, smul_neg, Units.smul_def, shiftIso, Int.reduceNeg, (fst (φ⟦n⟧')).1.rightShift_v 1 0 (zero_add 1) p p (add_zero p) (p + 1) rfl, HomologicalComplex.XIsoOfEq_rfl, Iso.refl_inv, Preadditive.comp_neg, lift_f_fst_v, Cocycle.coe_smul, Cocycle.shift_coe, Cochain.smul_v, Cochain.shift_v'] diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 29781444cbc9c..197eacedc6ec5 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -1000,9 +1000,9 @@ theorem repr_opow (o₁ o₂) [NF o₁] [NF o₂] : repr (o₁ ^ o₂) = repr o simp only [opow_def, opow, e₁, r₁, split_eq_scale_split' e₂, opowAux2, repr] cases' k with k · simp [r₂, opow_mul, repr_opow_aux₁ a00 al aa, add_assoc] - · simp? [r₂, opow_add, opow_mul, mul_assoc, add_assoc, -repr] - says simp only [mulNat_eq_mul, repr_add, repr_scale, repr_mul, repr_ofNat, opow_add, - opow_mul, mul_assoc, add_assoc, r₂, Nat.cast_succ, add_one_eq_succ, opow_succ] + · simp? [r₂, opow_add, opow_mul, mul_assoc, add_assoc, -repr] says + simp only [mulNat_eq_mul, repr_add, repr_scale, repr_mul, repr_ofNat, opow_add, opow_mul, + mul_assoc, add_assoc, r₂, Nat.cast_add, Nat.cast_one, add_one_eq_succ, opow_succ] simp only [repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, opow_one] rw [repr_opow_aux₁ a00 al aa, scale_opowAux] simp only [repr_mul, repr_scale, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, From ac2f221bfbc5357a9d26784a65d2cb767c509f0f Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 17:40:56 -0400 Subject: [PATCH 15/48] remove --- Mathlib/Data/Multiset/Fintype.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index 507e55473b4d6..8d0fd95ab88bb 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -70,12 +70,12 @@ instance instCoeSortMultisetType.instCoeOutToType : CoeOut m α := -- Porting note: syntactic equality #noalign multiset.fst_coe_eq_coe -@[simp] -theorem Multiset.coe_eq {x y : m} : (x : α) = (y : α) ↔ x.1 = y.1 := by - cases x - cases y - rfl -#align multiset.coe_eq Multiset.coe_eq +-- @[simp] +-- theorem Multiset.coe_eq {x y : m} : (x : α) = (y : α) ↔ x.1 = y.1 := by +-- cases x +-- cases y +-- rfl +-- #align multiset.coe_eq Multiset.coe_eq -- @[simp] -- Porting note (#10685): dsimp can prove this theorem Multiset.coe_mk {x : α} {i : Fin (m.count x)} : ↑(m.mkToType x i) = x := From 8857be781f75d5afc3ae6559f7ca5a85be109fab Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 17:00:58 -0400 Subject: [PATCH 16/48] fix says --- .../Homology/HomotopyCategory/DegreewiseSplit.lean | 2 +- .../Homology/HomotopyCategory/Triangulated.lean | 12 ++++++------ 2 files changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index ed118e61db790..b9d1a5eb552db 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -87,7 +87,7 @@ noncomputable def mappingConeHomOfDegreewiseSplitXIso (p q : ℤ) (hpq : p + 1 = have f_r := (σ (p + 1)).f_r dsimp at s_g f_r ⊢ simp? [mappingCone.ext_from_iff _ (p + 1) _ rfl, reassoc_of% f_r, reassoc_of% s_g] says - simp only [Cochain.ofHom_v, id_comp, comp_sub, sub_comp, assoc, reassoc_of% s_g, + simp only [Cochain.ofHom_v, Int.reduceNeg, id_comp, comp_sub, sub_comp, assoc, reassoc_of% s_g, ShortComplex.Splitting.s_r_assoc, ShortComplex.map_X₃, eval_obj, ShortComplex.map_X₁, zero_comp, comp_zero, reassoc_of% f_r, zero_sub, sub_neg_eq_add, mappingCone.ext_from_iff _ (p + 1) _ rfl, comp_add, mappingCone.inl_v_fst_v_assoc, diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean index 7caa7184fe2f8..ea694a3261c30 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean @@ -106,12 +106,12 @@ noncomputable def homotopyInvHomId : Homotopy (inv f g ≫ hom f g) (𝟙 _) := Cochain.comp_v _ _ (show (-1) + -1 = -2 by linarith) (n + 2) (n + 1) n (by linarith) (by linarith)] says simp only [mappingConeCompTriangle_obj₁, mappingConeCompTriangle_obj₂, - mappingConeCompTriangle_mor₁, map, inv, hom, Cochain.ofHom_comp, ofHom_desc, ofHom_lift, - descCocycle_coe, AddSubmonoid.coe_zero, Cochain.comp_zero_cochain_v, - inl_v_descCochain_v_assoc, Cochain.zero_cochain_comp_v, assoc, inl_v_snd_v_assoc, - zero_comp, Cochain.id_comp, Cochain.comp_assoc_of_first_is_zero_cochain, - Cochain.comp_add, Cochain.comp_neg, Cochain.comp_assoc_of_second_is_zero_cochain, - neg_add_rev, neg_neg, Cochain.add_v, Cochain.neg_v, + mappingConeCompTriangle_mor₁, map, Int.reduceNeg, inv, hom, Cochain.ofHom_comp, ofHom_desc, + ofHom_lift, descCocycle_coe, AddSubmonoid.coe_zero, Cochain.comp_zero_cochain_v, + inl_v_descCochain_v_assoc, Cochain.zero_cochain_comp_v, assoc, inl_v_snd_v_assoc, zero_comp, + Cochain.id_comp, Cochain.comp_assoc_of_first_is_zero_cochain, Cochain.comp_add, + Cochain.comp_neg, Cochain.comp_assoc_of_second_is_zero_cochain, neg_add_rev, neg_neg, + Cochain.add_v, Cochain.neg_v, Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n (by linarith) (by linarith), Cochain.comp_v _ _ (show 1 + -2 = -1 by linarith) (n + 1) (n + 2) n (by linarith) (by linarith), From 5b6c8c977bca98ebe32b6bf1518cbebe2ad16d74 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 08:58:06 +1100 Subject: [PATCH 17/48] merge Mario's patch to Ordinal/Notation --- Mathlib/SetTheory/Ordinal/Notation.lean | 116 ++++++++++++------------ 1 file changed, 57 insertions(+), 59 deletions(-) diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 197eacedc6ec5..ef2e98014aab9 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -1088,65 +1088,63 @@ attribute FundamentalSequenceProp theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamentalSequence o) := by - -- FIXME nightly-testing: proof broken, hard to tell origin intent. - sorry - -- induction' o with a m b iha ihb; · exact rfl - -- rw [fundamentalSequence] - -- rcases e : b.fundamentalSequence with (⟨_ | b'⟩ | f) <;> - -- simp only [FundamentalSequenceProp] <;> - -- rw [e, FundamentalSequenceProp] at ihb - -- · rcases e : a.fundamentalSequence with (⟨_ | a'⟩ | f) <;> cases' e' : m.natPred with m' <;> - -- simp only [FundamentalSequenceProp] <;> - -- rw [e, FundamentalSequenceProp] at iha <;> - -- (try rw [show m = 1 by - -- have := PNat.natPred_add_one m; rw [e'] at this; exact PNat.coe_inj.1 this.symm]) <;> - -- (try rw [show m = m'.succ.succPNat by - -- rw [← e', ← PNat.coe_inj, Nat.succPNat_coe, ← Nat.add_one, PNat.natPred_add_one]]) <;> - -- simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega, add_lt_add_iff_left, add_zero, - -- eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, - -- Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, - -- true_and_iff, _root_.zero_add, zero_def] - -- · decide - -- · exact ⟨rfl, inferInstance⟩ - -- · have := opow_pos (repr a') omega_pos - -- refine' - -- ⟨mul_isLimit this omega_isLimit, fun i => - -- ⟨this, _, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega'⟩ - -- rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] - -- apply nat_lt_omega - -- · have := opow_pos (repr a') omega_pos - -- refine' - -- ⟨add_isLimit _ (mul_isLimit this omega_isLimit), fun i => ⟨this, _, _⟩, - -- exists_lt_add exists_lt_mul_omega'⟩ - -- · rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] - -- apply nat_lt_omega - -- · refine' fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (iha.2 H.fst))) - -- rw [repr, ← zero_def, repr, add_zero, iha.1, opow_succ, Ordinal.mul_lt_mul_iff_left this] - -- apply nat_lt_omega - -- · rcases iha with ⟨h1, h2, h3⟩ - -- refine' ⟨opow_isLimit one_lt_omega h1, fun i => _, exists_lt_omega_opow' one_lt_omega h1 h3⟩ - -- obtain ⟨h4, h5, h6⟩ := h2 i - -- exact ⟨h4, h5, fun H => @NF.oadd_zero _ _ (h6 H.fst)⟩ - -- · rcases iha with ⟨h1, h2, h3⟩ - -- refine' - -- ⟨add_isLimit _ (opow_isLimit one_lt_omega h1), fun i => _, - -- exists_lt_add (exists_lt_omega_opow' one_lt_omega h1 h3)⟩ - -- obtain ⟨h4, h5, h6⟩ := h2 i - -- refine' ⟨h4, h5, fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (h6 H.fst)))⟩ - -- rwa [repr, ← zero_def, repr, add_zero, PNat.one_coe, Nat.cast_one, mul_one, - -- opow_lt_opow_iff_right one_lt_omega] - -- · refine' - -- ⟨by rw [repr, ihb.1, add_succ, repr], fun H => H.fst.oadd _ (NF.below_of_lt' _ (ihb.2 H.snd))⟩ - -- have := H.snd'.repr_lt - -- rw [ihb.1] at this - -- exact (lt_succ _).trans this - -- · rcases ihb with ⟨h1, h2, h3⟩ - -- simp only [repr] - -- exact - -- ⟨Ordinal.add_isLimit _ h1, fun i => - -- ⟨oadd_lt_oadd_3 (h2 i).1, oadd_lt_oadd_3 (h2 i).2.1, fun H => - -- H.fst.oadd _ (NF.below_of_lt' (lt_trans (h2 i).2.1 H.snd'.repr_lt) ((h2 i).2.2 H.snd))⟩, - -- exists_lt_add h3⟩ + induction' o with a m b iha ihb; · exact rfl + rw [fundamentalSequence] + rcases e : b.fundamentalSequence with (⟨_ | b'⟩ | f) <;> + simp only [FundamentalSequenceProp] <;> + rw [e, FundamentalSequenceProp] at ihb + · rcases e : a.fundamentalSequence with (⟨_ | a'⟩ | f) <;> cases' e' : m.natPred with m' <;> + simp only [FundamentalSequenceProp] <;> + rw [e, FundamentalSequenceProp] at iha <;> + (try rw [show m = 1 by + have := PNat.natPred_add_one m; rw [e'] at this; exact PNat.coe_inj.1 this.symm]) <;> + (try rw [show m = (m' + 1).succPNat by + rw [← e', ← PNat.coe_inj, Nat.succPNat_coe, ← Nat.add_one, PNat.natPred_add_one]]) <;> + simp only [repr, iha, ihb, opow_lt_opow_iff_right one_lt_omega, add_lt_add_iff_left, add_zero, + eq_self_iff_true, lt_add_iff_pos_right, lt_def, mul_one, Nat.cast_zero, Nat.cast_succ, + Nat.succPNat_coe, opow_succ, opow_zero, mul_add_one, PNat.one_coe, succ_zero, + true_and_iff, _root_.zero_add, zero_def] + · decide + · exact ⟨rfl, inferInstance⟩ + · have := opow_pos (repr a') omega_pos + refine' + ⟨mul_isLimit this omega_isLimit, fun i => + ⟨this, _, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega'⟩ + rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] + apply nat_lt_omega + · have := opow_pos (repr a') omega_pos + refine' + ⟨add_isLimit _ (mul_isLimit this omega_isLimit), fun i => ⟨this, _, _⟩, + exists_lt_add exists_lt_mul_omega'⟩ + · rw [← mul_succ, ← nat_cast_succ, Ordinal.mul_lt_mul_iff_left this] + apply nat_lt_omega + · refine' fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (iha.2 H.fst))) + rw [repr, ← zero_def, repr, add_zero, iha.1, opow_succ, Ordinal.mul_lt_mul_iff_left this] + apply nat_lt_omega + · rcases iha with ⟨h1, h2, h3⟩ + refine' ⟨opow_isLimit one_lt_omega h1, fun i => _, exists_lt_omega_opow' one_lt_omega h1 h3⟩ + obtain ⟨h4, h5, h6⟩ := h2 i + exact ⟨h4, h5, fun H => @NF.oadd_zero _ _ (h6 H.fst)⟩ + · rcases iha with ⟨h1, h2, h3⟩ + refine' + ⟨add_isLimit _ (opow_isLimit one_lt_omega h1), fun i => _, + exists_lt_add (exists_lt_omega_opow' one_lt_omega h1 h3)⟩ + obtain ⟨h4, h5, h6⟩ := h2 i + refine' ⟨h4, h5, fun H => H.fst.oadd _ (NF.below_of_lt' _ (@NF.oadd_zero _ _ (h6 H.fst)))⟩ + rwa [repr, ← zero_def, repr, add_zero, PNat.one_coe, Nat.cast_one, mul_one, + opow_lt_opow_iff_right one_lt_omega] + · refine' + ⟨by rw [repr, ihb.1, add_succ, repr], fun H => H.fst.oadd _ (NF.below_of_lt' _ (ihb.2 H.snd))⟩ + have := H.snd'.repr_lt + rw [ihb.1] at this + exact (lt_succ _).trans this + · rcases ihb with ⟨h1, h2, h3⟩ + simp only [repr] + exact + ⟨Ordinal.add_isLimit _ h1, fun i => + ⟨oadd_lt_oadd_3 (h2 i).1, oadd_lt_oadd_3 (h2 i).2.1, fun H => + H.fst.oadd _ (NF.below_of_lt' (lt_trans (h2 i).2.1 H.snd'.repr_lt) ((h2 i).2.2 H.snd))⟩, + exists_lt_add h3⟩ #align onote.fundamental_sequence_has_prop ONote.fundamentalSequence_has_prop /-- The fast growing hierarchy for ordinal notations `< ε₀`. This is a sequence of From 92861e053c6e6e44f222b7acecf9482ef9405864 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 17:58:14 -0400 Subject: [PATCH 18/48] clean up --- Mathlib/Data/Multiset/Fintype.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index 8d0fd95ab88bb..423b56eef02bd 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -70,12 +70,8 @@ instance instCoeSortMultisetType.instCoeOutToType : CoeOut m α := -- Porting note: syntactic equality #noalign multiset.fst_coe_eq_coe --- @[simp] --- theorem Multiset.coe_eq {x y : m} : (x : α) = (y : α) ↔ x.1 = y.1 := by --- cases x --- cases y --- rfl --- #align multiset.coe_eq Multiset.coe_eq +-- Syntactic equality +#noalign multiset.coe_eq -- @[simp] -- Porting note (#10685): dsimp can prove this theorem Multiset.coe_mk {x : α} {i : Fin (m.count x)} : ↑(m.mkToType x i) = x := From 81a752fea13d30f4aa8727eb7fd16d9a32133244 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:00:21 +1100 Subject: [PATCH 19/48] fix --- Archive/Wiedijk100Theorems/InverseTriangleSum.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Archive/Wiedijk100Theorems/InverseTriangleSum.lean b/Archive/Wiedijk100Theorems/InverseTriangleSum.lean index e47c5fac615cd..918352a9e031c 100644 --- a/Archive/Wiedijk100Theorems/InverseTriangleSum.lean +++ b/Archive/Wiedijk100Theorems/InverseTriangleSum.lean @@ -33,8 +33,9 @@ theorem Theorem100.inverse_triangle_sum : refine' sum_range_induction _ _ (if_pos rfl) _ rintro (_ | n) · rw [if_neg, if_pos] <;> norm_num - simp_rw [if_neg (Nat.succ_ne_zero _), Nat.succ_eq_add_one] - have A : (n + 1 + 1 : ℚ) ≠ 0 := by norm_cast; norm_num + simp only [add_eq_zero, Nat.succ_ne_zero, one_ne_zero, and_self, ↓reduceIte, Nat.cast_add, + Nat.cast_succ, Nat.cast_one] + have A : (n + 1 + 1 : ℚ) ≠ 0 := by norm_cast push_cast field_simp ring From 9e80fbbd5baaa71c63af1235a31f4273479f7836 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:08:05 +1100 Subject: [PATCH 20/48] fixes --- Archive/MiuLanguage/DecisionSuf.lean | 5 ++--- Archive/Wiedijk100Theorems/Partition.lean | 5 ++--- Counterexamples/CliffordAlgebra_not_injective.lean | 5 ++--- Counterexamples/Phillips.lean | 2 +- 4 files changed, 7 insertions(+), 10 deletions(-) diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 6b5d93549384d..949cc658db0cf 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -60,7 +60,7 @@ private theorem der_cons_replicate (n : ℕ) : Derivable (M :: replicate (2 ^ n) · -- base case constructor · -- inductive step - rw [succ_eq_add_one, pow_add, pow_one 2, mul_two, replicate_add] + rw [pow_add, pow_one 2, mul_two, replicate_add] exact Derivable.r2 hk /-! @@ -120,7 +120,7 @@ theorem der_cons_replicate_I_replicate_U_append_of_der_cons_replicate_I_append ( specialize ha (U :: xs) intro h₂ -- We massage the goal into a form amenable to the application of `ha`. - rw [succ_eq_add_one, replicate_add, ← append_assoc, ← cons_append, replicate_one, append_assoc, + rw [replicate_add, ← append_assoc, ← cons_append, replicate_one, append_assoc, singleton_append] apply ha apply Derivable.r3 @@ -269,7 +269,6 @@ theorem count_I_eq_length_of_count_U_zero_and_neg_mem {ys : Miustr} (hu : count · rw [mem_cons, not_or] at hm; exact hm.2 · -- case `x = U` gives a contradiction. exfalso; simp only [count, countP_cons_of_pos (· == U) _ (rfl : U == U)] at hu - exact succ_ne_zero _ hu set_option linter.uppercaseLean3 false in #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 diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 0525eec99f251..285b4fc3b6c2d 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -147,7 +147,7 @@ theorem num_series' [Field α] (i : ℕ) : symm split_ifs with h · suffices - ((antidiagonal n.succ).filter fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1).card = + ((antidiagonal (n+1)).filter fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1).card = 1 by simp only [Set.mem_setOf_eq]; convert congr_arg ((↑) : ℕ → α) this; norm_cast rw [card_eq_one] @@ -164,7 +164,7 @@ theorem num_series' [Field α] (i : ℕ) : | 0 => rw [mul_zero] at hp; cases hp | p + 1 => rw [hp]; simp [mul_add] · suffices - (filter (fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1) (antidiagonal n.succ)).card = + (filter (fun a : ℕ × ℕ => i + 1 ∣ a.fst ∧ a.snd = i + 1) (antidiagonal (n+1))).card = 0 by simp only [Set.mem_setOf_eq]; convert congr_arg ((↑) : ℕ → α) this; norm_cast rw [card_eq_zero] @@ -373,7 +373,6 @@ theorem same_gf [Field α] (m : ℕ) : rw [partialOddGF, partialDistinctGF] induction' m with m ih · simp - rw [Nat.succ_eq_add_one] set! π₀ : PowerSeries α := ∏ i in range m, (1 - X ^ (m + 1 + i + 1)) with hπ₀ set! π₁ : PowerSeries α := ∏ i in range m, (1 - X ^ (2 * i + 1))⁻¹ with hπ₁ set! π₂ : PowerSeries α := ∏ i in range m, (1 - X ^ (m + i + 1)) with hπ₂ diff --git a/Counterexamples/CliffordAlgebra_not_injective.lean b/Counterexamples/CliffordAlgebra_not_injective.lean index 4e833088dc78d..200f960e76186 100644 --- a/Counterexamples/CliffordAlgebra_not_injective.lean +++ b/Counterexamples/CliffordAlgebra_not_injective.lean @@ -50,10 +50,9 @@ def kIdeal : Ideal (MvPolynomial (Fin 3) (ZMod 2)) := theorem mem_kIdeal_iff (x : MvPolynomial (Fin 3) (ZMod 2)) : x ∈ kIdeal ↔ ∀ m : Fin 3 →₀ ℕ, m ∈ x.support → ∃ i, 2 ≤ m i := by - have : - kIdeal = Ideal.span ((monomial · (1 : ZMod 2)) '' Set.range (Finsupp.single · 2)) := by + have : kIdeal = Ideal.span ((monomial · (1 : ZMod 2)) '' Set.range (Finsupp.single · 2)) := by simp_rw [kIdeal, X, monomial_mul, one_mul, ← Finsupp.single_add, ← Set.range_comp, - Function.comp] + Function.comp, Nat.reduceAdd] rw [this, mem_ideal_span_monomial_image] simp diff --git a/Counterexamples/Phillips.lean b/Counterexamples/Phillips.lean index eeb97d55ff882..f2607e5f27fbb 100644 --- a/Counterexamples/Phillips.lean +++ b/Counterexamples/Phillips.lean @@ -318,7 +318,7 @@ theorem exists_discrete_support_nonpos (f : BoundedAdditiveMeasure α) : · have : (s (n + 1)).1 = (s (n + 1)).1 \ (s n).1 ∪ (s n).1 := by simpa only [s, Function.iterate_succ', union_diff_self] using (diff_union_of_subset <| subset_union_left _ _).symm - rw [Nat.succ_eq_add_one, this, f.additive] + rw [this, f.additive] swap; · exact disjoint_sdiff_self_left calc ((n + 1 : ℕ) : ℝ) * (ε / 2) = ε / 2 + n * (ε / 2) := by simp only [Nat.cast_succ]; ring From bac29b80174474776b4e17ad3855eda3d8193071 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 12 Mar 2024 23:11:10 +0100 Subject: [PATCH 21/48] fixes --- Mathlib/Algebra/CharZero/Defs.lean | 5 ++--- .../Homology/HomotopyCategory/DegreewiseSplit.lean | 6 +++--- .../Homology/HomotopyCategory/Pretriangulated.lean | 3 ++- .../Homology/HomotopyCategory/Triangulated.lean | 13 +++++++------ Mathlib/Analysis/BoundedVariation.lean | 10 ++++------ Mathlib/Data/Int/Cast/Basic.lean | 3 +-- Mathlib/Data/List/Sort.lean | 6 +++--- Mathlib/Data/Polynomial/Module/Basic.lean | 2 +- Mathlib/Data/ZMod/Factorial.lean | 4 ++-- 9 files changed, 25 insertions(+), 27 deletions(-) diff --git a/Mathlib/Algebra/CharZero/Defs.lean b/Mathlib/Algebra/CharZero/Defs.lean index 250dacfb846ed..59ef26d5c9d52 100644 --- a/Mathlib/Algebra/CharZero/Defs.lean +++ b/Mathlib/Algebra/CharZero/Defs.lean @@ -84,9 +84,8 @@ theorem cast_ne_zero {n : ℕ} : (n : R) ≠ 0 ↔ n ≠ 0 := not_congr cast_eq_zero #align nat.cast_ne_zero Nat.cast_ne_zero -theorem cast_add_one_ne_zero (n : ℕ) : (n + 1 : R) ≠ 0 := by - -- Porting note: old proof was `exact_mod_cast n.succ_ne_zero` - norm_cast +theorem cast_add_one_ne_zero (n : ℕ) : (n + 1 : R) ≠ 0 := + mod_cast n.succ_ne_zero #align nat.cast_add_one_ne_zero Nat.cast_add_one_ne_zero @[simp, norm_cast] diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean index b9d1a5eb552db..6a3ac7168f256 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/DegreewiseSplit.lean @@ -87,9 +87,9 @@ noncomputable def mappingConeHomOfDegreewiseSplitXIso (p q : ℤ) (hpq : p + 1 = have f_r := (σ (p + 1)).f_r dsimp at s_g f_r ⊢ simp? [mappingCone.ext_from_iff _ (p + 1) _ rfl, reassoc_of% f_r, reassoc_of% s_g] says - simp only [Cochain.ofHom_v, Int.reduceNeg, id_comp, comp_sub, sub_comp, assoc, reassoc_of% s_g, - ShortComplex.Splitting.s_r_assoc, ShortComplex.map_X₃, eval_obj, ShortComplex.map_X₁, - zero_comp, comp_zero, reassoc_of% f_r, zero_sub, sub_neg_eq_add, + simp only [Cochain.ofHom_v, Int.reduceNeg, id_comp, comp_sub, sub_comp, assoc, + reassoc_of% s_g, ShortComplex.Splitting.s_r_assoc, ShortComplex.map_X₃, eval_obj, + ShortComplex.map_X₁, zero_comp, comp_zero, reassoc_of% f_r, zero_sub, sub_neg_eq_add, mappingCone.ext_from_iff _ (p + 1) _ rfl, comp_add, mappingCone.inl_v_fst_v_assoc, mappingCone.inl_v_snd_v_assoc, shiftFunctor_obj_X', sub_zero, add_zero, comp_id, mappingCone.inr_f_fst_v_assoc, mappingCone.inr_f_snd_v_assoc, add_left_eq_self, neg_eq_zero, diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean index 9e512e0234ef2..8cdb3fbb0b29d 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Pretriangulated.lean @@ -183,7 +183,8 @@ noncomputable def rotateHomotopyEquiv : hom := lift (inr φ) (-(Cocycle.ofHom φ).leftShift 1 1 (zero_add 1)) (-(inl φ).leftShift 1 0 (neg_add_self 1)) (by simp? [Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1)] says - simp only [Int.reduceNeg, δ_neg, Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1), + simp only [Int.reduceNeg, δ_neg, + Cochain.δ_leftShift _ 1 0 1 (neg_add_self 1) 0 (zero_add 1), Int.negOnePow_one, δ_inl, Cochain.ofHom_comp, Cochain.leftShift_comp_zero_cochain, Units.neg_smul, one_smul, neg_neg, Cocycle.coe_neg, Cocycle.leftShift_coe, Cocycle.ofHom_coe, Cochain.neg_comp, add_right_neg]) diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean index ea694a3261c30..b16ba1f82fec6 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/Triangulated.lean @@ -106,12 +106,13 @@ noncomputable def homotopyInvHomId : Homotopy (inv f g ≫ hom f g) (𝟙 _) := Cochain.comp_v _ _ (show (-1) + -1 = -2 by linarith) (n + 2) (n + 1) n (by linarith) (by linarith)] says simp only [mappingConeCompTriangle_obj₁, mappingConeCompTriangle_obj₂, - mappingConeCompTriangle_mor₁, map, Int.reduceNeg, inv, hom, Cochain.ofHom_comp, ofHom_desc, - ofHom_lift, descCocycle_coe, AddSubmonoid.coe_zero, Cochain.comp_zero_cochain_v, - inl_v_descCochain_v_assoc, Cochain.zero_cochain_comp_v, assoc, inl_v_snd_v_assoc, zero_comp, - Cochain.id_comp, Cochain.comp_assoc_of_first_is_zero_cochain, Cochain.comp_add, - Cochain.comp_neg, Cochain.comp_assoc_of_second_is_zero_cochain, neg_add_rev, neg_neg, - Cochain.add_v, Cochain.neg_v, + mappingConeCompTriangle_mor₁, map, Int.reduceNeg, inv, hom, Cochain.ofHom_comp, + ofHom_desc, ofHom_lift, descCocycle_coe, AddSubmonoid.coe_zero, + Cochain.comp_zero_cochain_v, inl_v_descCochain_v_assoc, Cochain.zero_cochain_comp_v, + assoc, inl_v_snd_v_assoc, zero_comp, Cochain.id_comp, + Cochain.comp_assoc_of_first_is_zero_cochain, Cochain.comp_add, Cochain.comp_neg, + Cochain.comp_assoc_of_second_is_zero_cochain, neg_add_rev, neg_neg, Cochain.add_v, + Cochain.neg_v, Cochain.comp_v _ _ (add_neg_self 1) n (n + 1) n (by linarith) (by linarith), Cochain.comp_v _ _ (show 1 + -2 = -1 by linarith) (n + 1) (n + 2) n (by linarith) (by linarith), diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index eacb35fdd9ecf..695c4bbe3cafa 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -524,12 +524,10 @@ theorem comp_le_of_antitoneOn (f : α → E) {s : Set α} {t : Set β} (φ : β ⟨n, fun i => φ (u <| n - i), fun x y xy => hφ (ut _) (ut _) (hu <| Nat.sub_le_sub_left xy n), fun i => φst (ut _)⟩ le_rfl - dsimp only [Subtype.coe_mk] - rw [edist_comm, Nat.sub_sub, add_comm, Nat.sub_succ, Nat.add_one, Nat.succ_eq_add_one] - simp only [Function.comp_apply, Nat.pred_eq_sub_one, Nat.sub_add_eq] - congr - simp only [Finset.mem_range] at hx - omega + rw [Finset.mem_range] at hx + dsimp only [Subtype.coe_mk, Function.comp_apply] + rw [edist_comm] + congr 4 <;> omega #align evariation_on.comp_le_of_antitone_on eVariationOn.comp_le_of_antitoneOn theorem comp_eq_of_monotoneOn (f : α → E) {t : Set β} (φ : β → α) (hφ : MonotoneOn φ t) : diff --git a/Mathlib/Data/Int/Cast/Basic.lean b/Mathlib/Data/Int/Cast/Basic.lean index fc1c4eb053410..f0aaf3c821729 100644 --- a/Mathlib/Data/Int/Cast/Basic.lean +++ b/Mathlib/Data/Int/Cast/Basic.lean @@ -94,8 +94,7 @@ theorem cast_subNatNat (m n) : ((Int.subNatNat m n : ℤ) : R) = m - n := by cases e : n - m · simp only [ofNat_eq_coe] simp [e, Nat.le_of_sub_eq_zero e] - · rw [cast_negSucc, Nat.add_one, succ_eq_add_one, ← e, - Nat.cast_sub <| _root_.le_of_lt <| Nat.lt_of_sub_eq_succ e, neg_sub] + · rw [cast_negSucc, ← e, Nat.cast_sub <| _root_.le_of_lt <| Nat.lt_of_sub_eq_succ e, neg_sub] #align int.cast_sub_nat_nat Int.cast_subNatNatₓ -- type had `HasLiftT` diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 8812d44d183d4..0767cc69abf12 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -479,7 +479,7 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge' r l l') +theorem Sorted.merge' : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge' r l l') | [], [], _, _ => by simp [List.merge'] | [], b :: l', _, h₂ => by simpa [List.merge'] using h₂ | a :: l, [], h₁, _ => by simpa [List.merge'] using h₁ @@ -506,7 +506,7 @@ theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sort · exact _root_.trans ba (rel_of_sorted_cons h₁ _ bl) · exact rel_of_sorted_cons h₂ _ bl' termination_by l₁ l₂ => length l₁ + length l₂ -#align list.sorted.merge List.Sorted.merge +#align list.sorted.merge List.Sorted.merge' variable (r) @@ -517,7 +517,7 @@ theorem sorted_mergeSort : ∀ l : List α, Sorted r (mergeSort r l) cases' e : split (a :: b :: l) with l₁ l₂ cases' length_split_lt e with h₁ h₂ rw [mergeSort_cons_cons r e] - exact (sorted_mergeSort l₁).merge (sorted_mergeSort l₂) + exact (sorted_mergeSort l₁).merge' (sorted_mergeSort l₂) termination_by l => length l #align list.sorted_merge_sort List.sorted_mergeSort diff --git a/Mathlib/Data/Polynomial/Module/Basic.lean b/Mathlib/Data/Polynomial/Module/Basic.lean index 4507f6d031c42..83719923c940a 100644 --- a/Mathlib/Data/Polynomial/Module/Basic.lean +++ b/Mathlib/Data/Polynomial/Module/Basic.lean @@ -315,7 +315,7 @@ theorem monomial_smul_single (i : ℕ) (r : R) (j : ℕ) (m : M) : rw [Function.iterate_zero, zero_add] exact Finsupp.smul_single r j m | succ n hn => - rw [Function.iterate_succ, Function.comp_apply,add_assoc, ← hn] + rw [Function.iterate_succ, Function.comp_apply, add_assoc, ← hn] congr 2 rw [Nat.one_add] exact Finsupp.mapDomain_single diff --git a/Mathlib/Data/ZMod/Factorial.lean b/Mathlib/Data/ZMod/Factorial.lean index 6a7df9d2ea263..48857516aaf72 100644 --- a/Mathlib/Data/ZMod/Factorial.lean +++ b/Mathlib/Data/ZMod/Factorial.lean @@ -37,7 +37,7 @@ theorem cast_descFactorial {n p : ℕ} (h : n ≤ p) : intro x hx rw [← tsub_add_eq_tsub_tsub_swap, Nat.cast_sub <| Nat.le_trans (Nat.add_one_le_iff.mpr (List.mem_range.mp hx)) h, - CharP.cast_eq_zero, zero_sub, cast_succ, neg_add_rev, mul_add, neg_mul, one_mul, - mul_one, add_comm] + CharP.cast_eq_zero, zero_sub, cast_succ, neg_add_rev, mul_add, neg_mul, one_mul, + mul_one, add_comm] end ZMod From eea60d7f416762c95ca4c8afdc0268a23be6bf39 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:17:28 +1100 Subject: [PATCH 22/48] fix Nat/Cast/Defs --- Mathlib/Data/Nat/Cast/Defs.lean | 20 +++++++++----------- test/LibrarySearch/mathlib.lean | 8 +++----- 2 files changed, 12 insertions(+), 16 deletions(-) diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index 8a7a1840226cb..33eecaccd2c8b 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -173,17 +173,15 @@ theorem binCast_eq [AddMonoidWithOne R] (n : ℕ) : (Nat.binCast n : R) = ((n : | zero => rw [Nat.binCast, Nat.cast_zero] | succ k => rw [Nat.binCast] - -- FIXME nightly-testing - sorry - -- by_cases h : (k + 1) % 2 = 0 - -- · rw [← Nat.mod_add_div (succ k) 2] - -- rw [if_pos h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] - -- rw [Nat.succ_eq_add_one, h, Nat.zero_add, Nat.succ_mul, Nat.one_mul] - -- · rw [← Nat.mod_add_div (succ k) 2] - -- rw [if_neg h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] - -- have h1 := Or.resolve_left (Nat.mod_two_eq_zero_or_one (succ k)) h - -- rw [h1, Nat.add_comm 1, Nat.succ_mul, Nat.one_mul] - -- simp only [Nat.cast_add, Nat.cast_one] + by_cases h : (k + 1) % 2 = 0 + · conv => rhs; rw [← Nat.mod_add_div (k+1) 2] + rw [if_pos h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] + rw [h, Nat.zero_add, Nat.succ_mul, Nat.one_mul] + · conv => rhs; rw [← Nat.mod_add_div (k+1) 2] + rw [if_neg h, hk _ <| Nat.div_lt_self (Nat.succ_pos k) (Nat.le_refl 2), ← Nat.cast_add] + have h1 := Or.resolve_left (Nat.mod_two_eq_zero_or_one (succ k)) h + rw [h1, Nat.add_comm 1, Nat.succ_mul, Nat.one_mul] + simp only [Nat.cast_add, Nat.cast_one] #align nat.bin_cast_eq Nat.binCast_eq section deprecated diff --git a/test/LibrarySearch/mathlib.lean b/test/LibrarySearch/mathlib.lean index 0e33f5bf93b6c..6f1315a15bccc 100644 --- a/test/LibrarySearch/mathlib.lean +++ b/test/LibrarySearch/mathlib.lean @@ -3,8 +3,6 @@ import Mathlib -- We verify that `exact?` copes with all of Mathlib. -- On `v4.7.0-rc1` this revealed a cache corruption problem. --- FIXME reenable this once bump/v4.8.0 moves to nightly-2024-03-07. - --- /-- info: Try this: exact Nat.zero_lt_one -/ --- #guard_msgs in --- example : 0 < 1 := by exact? +/-- info: Try this: exact Nat.zero_lt_one -/ +#guard_msgs in +example : 0 < 1 := by exact? From 0b6eccda1916deb3bfc1ac937df60f422d31f3bd Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:20:39 +1100 Subject: [PATCH 23/48] remove stray prime? --- Mathlib/Data/List/Sort.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 0767cc69abf12..8812d44d183d4 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -479,7 +479,7 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge' : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge' r l l') +theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge' r l l') | [], [], _, _ => by simp [List.merge'] | [], b :: l', _, h₂ => by simpa [List.merge'] using h₂ | a :: l, [], h₁, _ => by simpa [List.merge'] using h₁ @@ -506,7 +506,7 @@ theorem Sorted.merge' : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sor · exact _root_.trans ba (rel_of_sorted_cons h₁ _ bl) · exact rel_of_sorted_cons h₂ _ bl' termination_by l₁ l₂ => length l₁ + length l₂ -#align list.sorted.merge List.Sorted.merge' +#align list.sorted.merge List.Sorted.merge variable (r) @@ -517,7 +517,7 @@ theorem sorted_mergeSort : ∀ l : List α, Sorted r (mergeSort r l) cases' e : split (a :: b :: l) with l₁ l₂ cases' length_split_lt e with h₁ h₂ rw [mergeSort_cons_cons r e] - exact (sorted_mergeSort l₁).merge' (sorted_mergeSort l₂) + exact (sorted_mergeSort l₁).merge (sorted_mergeSort l₂) termination_by l => length l #align list.sorted_merge_sort List.sorted_mergeSort From bb47464873af2572b10c78b62d225cea205a42c0 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:24:13 +1100 Subject: [PATCH 24/48] simplify proof --- Mathlib/Data/Num/Lemmas.lean | 7 +------ Mathlib/Init/Data/Nat/Bitwise.lean | 9 +++++++-- 2 files changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index a4c8bb33e3f6f..3988e034cf243 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -982,12 +982,7 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by · rfl · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_bit_zero] · rw [PosNum.cast_bit0, ← Nat.bit_false, Nat.testBit_bit_zero] - · rw [PosNum.cast_one', ← bit1_zero, ← Nat.bit_true] - -- FIXME adaptation for nightly-2024-03-11 - -- Hacky proof, can someone please replace? Obvious simp lemmas seem to be missing. - simp only [Nat.bit_true, bit1_zero, Bool.false_eq] - erw [Nat.testBit_succ] - simp + · simp · rw [PosNum.cast_bit1, ← Nat.bit_true, Nat.testBit_bit_succ, IH] · rw [PosNum.cast_bit0, ← Nat.bit_false, Nat.testBit_bit_succ, IH] #align num.test_bit_to_nat Num.castNum_testBit diff --git a/Mathlib/Init/Data/Nat/Bitwise.lean b/Mathlib/Init/Data/Nat/Bitwise.lean index c51b704e0602a..3e0a83b7196d3 100644 --- a/Mathlib/Init/Data/Nat/Bitwise.lean +++ b/Mathlib/Init/Data/Nat/Bitwise.lean @@ -169,9 +169,14 @@ theorem bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := def bitCasesOn {C : Nat → Sort u} (n) (h : ∀ b n, C (bit b n)) : C n := bit_decomp n ▸ h _ _ #align nat.bit_cases_on Nat.bitCasesOn -theorem bit_zero : bit false 0 = 0 := +@[simp] theorem bit_false_zero : bit false 0 = 0 := + rfl +#align nat.bit_zero Nat.bit_false_zero + +@[deprecated] alias bit_zero := bit_false_zero + +@[simp] theorem bit_true_zero : bit true 0 = 1 := rfl -#align nat.bit_zero Nat.bit_zero /--`shiftLeft' b m n` performs a left shift of `m` `n` times and adds the bit `b` as the least significant bit each time. From 2a73d3e55c546ffad057c9610c2d3b9450f5c0ba Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:26:18 +1100 Subject: [PATCH 25/48] long line --- Mathlib/Data/Nat/Digits.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 806ad9072e562..e7436c497e152 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -568,7 +568,8 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits have h₁ : 1 ≤ tl.length := List.length_pos.mpr h' rw [← sum_range_add_sum_Ico _ <| h₁, ← add_zero (∑ x in Ico _ _, ofDigits p (tl.drop x)), ← this, sum_Ico_consecutive _ h₁ <| (le_add_right (List.length tl) 1)] - -- Adaptation note: nightly-2024-03-07: this needs an `erw` only because of a `0 + 1` vs `1`. + -- Adaptation note: nightly-2024-03-07: + -- this needs an `erw` only because of a `0 + 1` vs `1`. -- Can someone do it more cleanly? erw [← sum_Ico_add _ 0 tl.length 1] rw [Ico_zero_eq_range, mul_add, mul_add, ih, range_one, sum_singleton, List.drop, ofDigits, From 3c968e4e9f8e7697acd1a0f618e4034092bb0926 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:30:29 +1100 Subject: [PATCH 26/48] change FIXME in TensorProduct/Graded/Internal to an adaptation note --- Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index 1af386f774405..d06ed264785fa 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -310,7 +310,10 @@ def lift (f : A →ₐ[R] C) (g : B →ₐ[R] C) ∘ₗ ((of R 𝒜 ℬ).symm : 𝒜 ᵍ⊗[R] ℬ →ₗ[R] A ⊗[R] B)) (by dsimp [Algebra.TensorProduct.one_def] - -- FIXME nightly-testing: no longer works with dsimp, even though it is a rfl lemma + -- Adaptation note: nightly-2024-03-11. + -- No longer works with dsimp, even though it is a rfl lemma. + -- This may be a Lean bug. + -- It would be great if someone could try to minimize this to an no imports example. rw [Algebra.TensorProduct.one_def] dsimp; simp only [_root_.map_one, mul_one]) (by From 0b59c6a2ef16649e4a6059d2a426030f4b120416 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:46:01 +1100 Subject: [PATCH 27/48] fix test --- test/cases.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/test/cases.lean b/test/cases.lean index 0bdfe5b32aaf1..f9895672df2b5 100644 --- a/test/cases.lean +++ b/test/cases.lean @@ -50,15 +50,15 @@ example (x : Bar 0) : True := by example (n : Nat) : n = n := by induction' n with n ih - · guard_target = Nat.zero = Nat.zero; rfl + · guard_target =ₛ 0 = 0; rfl · guard_hyp n : Nat; guard_hyp ih : n = n - guard_target = Nat.succ n = Nat.succ n; exact congr_arg _ ih + guard_target =ₛ n + 1 = n + 1; exact congr_arg (· + 1) ih example (n : Nat) (h : n < 5) : n = n := by induction' n with n ih - · guard_target = Nat.zero = Nat.zero; rfl - · guard_hyp n : Nat; guard_hyp ih : n < 5 → n = n; guard_hyp h : Nat.succ n < 5 - guard_target = Nat.succ n = Nat.succ n; rfl + · guard_target =ₛ 0 = 0; rfl + · guard_hyp n : Nat; guard_hyp ih : n < 5 → n = n; guard_hyp h :ₛ n + 1 < 5 + guard_target =ₛ n + 1 = n + 1; rfl example (n : Nat) {m} (h : m < 5) : n = n := by induction' n with n ih @@ -115,7 +115,7 @@ example (a b : ℕ) (h : a + b = a) : b = 0 := by fail_if_success (guard_hyp h : a + b = a) intro h -- Sample proof - rw [Nat.zero_eq, Nat.zero_add] at h + rw [Nat.zero_add] at h assumption · -- Test the generalized vars have been removed revert h From 85cd3828ddd0b05d9168948bf4f91a2f23b04207 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:46:27 +1100 Subject: [PATCH 28/48] fix tests --- test/ExtractGoal.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/test/ExtractGoal.lean b/test/ExtractGoal.lean index e8c1eab83dfda..eb32d50d32e00 100644 --- a/test/ExtractGoal.lean +++ b/test/ExtractGoal.lean @@ -30,9 +30,7 @@ example {α : Type u} {β : Type v} [Add α] [h : Sub β] (f : α → β) ⦃_g rfl -- an example with a hygienic variable -/-- -info: theorem extracted_1 (n : ℕ) : Nat.succ n = Nat.succ n := sorry --/ +/-- info: theorem extracted_1 (n : ℕ) : n + 1 = n + 1 := sorry -/ #guard_msgs in example (n : ℕ) : n = n := by cases n From 7410d4a4817645e126f24f52298e7564ea9c5f94 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 18:50:28 -0400 Subject: [PATCH 29/48] lint --- Mathlib/Data/Finset/Basic.lean | 2 +- Mathlib/Data/Fintype/BigOperators.lean | 6 +++++- Mathlib/Data/Multiset/Fintype.lean | 8 ++------ Mathlib/Logic/IsEmpty.lean | 8 +++++--- 4 files changed, 13 insertions(+), 11 deletions(-) diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index abbde83486b28..ead12d718dffc 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2814,7 +2814,7 @@ theorem filter_comm (s : Finset α) : (s.filter p).filter q = (s.filter q).filte #align finset.filter_comm Finset.filter_comm -- We can simplify an application of filter where the decidability is inferred in "the wrong way" -@[simp] +-- Adaptation note: as of `nightly-2024-03-11`, Was a simp lemma but no longer applies to itself theorem filter_congr_decidable (s : Finset α) (p : α → Prop) (h : DecidablePred p) [DecidablePred p] : @filter α p h s = s.filter p := by congr #align finset.filter_congr_decidable Finset.filter_congr_decidable diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 10bb80f39daae..83aefedfc1783 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -142,12 +142,16 @@ theorem Fintype.card_pi {β : α → Type*} [DecidableEq α] [Fintype α] [∀ a #align fintype.card_pi Fintype.card_pi -- FIXME ouch, this should be in the main file. -@[simp] theorem Fintype.card_fun [DecidableEq α] [Fintype α] [Fintype β] : Fintype.card (α → β) = Fintype.card β ^ Fintype.card α := by rw [Fintype.card_pi, Finset.prod_const]; rfl #align fintype.card_fun Fintype.card_fun +-- Adaptation note: as of `nightly-2024-03-11`, `Fintype.card_fun` does not have LHS in simp normal form +@[simp] +theorem Fintype.card_exp_univ_card_eq [DecidableEq α] [Fintype α] [Fintype β] : + card β ^ univ.card (α := α) = card β ^ card α := rfl + @[simp] theorem card_vector [Fintype α] (n : ℕ) : Fintype.card (Vector α n) = Fintype.card α ^ n := by rw [Fintype.ofEquiv_card]; simp diff --git a/Mathlib/Data/Multiset/Fintype.lean b/Mathlib/Data/Multiset/Fintype.lean index 507e55473b4d6..423b56eef02bd 100644 --- a/Mathlib/Data/Multiset/Fintype.lean +++ b/Mathlib/Data/Multiset/Fintype.lean @@ -70,12 +70,8 @@ instance instCoeSortMultisetType.instCoeOutToType : CoeOut m α := -- Porting note: syntactic equality #noalign multiset.fst_coe_eq_coe -@[simp] -theorem Multiset.coe_eq {x y : m} : (x : α) = (y : α) ↔ x.1 = y.1 := by - cases x - cases y - rfl -#align multiset.coe_eq Multiset.coe_eq +-- Syntactic equality +#noalign multiset.coe_eq -- @[simp] -- Porting note (#10685): dsimp can prove this theorem Multiset.coe_mk {x : α} {i : Fin (m.count x)} : ↑(m.mkToType x i) = x := diff --git a/Mathlib/Logic/IsEmpty.lean b/Mathlib/Logic/IsEmpty.lean index 9b10e299d2ccd..04ccb71e8b522 100644 --- a/Mathlib/Logic/IsEmpty.lean +++ b/Mathlib/Logic/IsEmpty.lean @@ -153,9 +153,11 @@ theorem isEmpty_pi {π : α → Sort*} : IsEmpty (∀ a, π a) ↔ ∃ a, IsEmpt simp only [← not_nonempty_iff, Classical.nonempty_pi, not_forall] #align is_empty_pi isEmpty_pi -@[simp] -theorem isEmpty_fun : IsEmpty (α → β) ↔ Nonempty α ∧ IsEmpty β := by - rw [isEmpty_pi, ← exists_true_iff_nonempty, ← exists_and_right, true_and] +@[simp] -- simp normal form of `isEmpty_fun` +theorem isEmpty_exists_isempty_fun : (∃ (_ : α), IsEmpty β) ↔ Nonempty α ∧ IsEmpty β := by + rw [← exists_true_iff_nonempty, ← exists_and_right, true_and] + +theorem isEmpty_fun : IsEmpty (α → β) ↔ Nonempty α ∧ IsEmpty β := by simp @[simp] theorem nonempty_fun : Nonempty (α → β) ↔ IsEmpty α ∨ Nonempty β := From b604582eb06821f650e4f215c0f56d778f929b99 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 18:52:00 -0400 Subject: [PATCH 30/48] lint --- Mathlib/Data/Nat/Defs.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index eca2d3ad4f735..c7f4cb0577c92 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -380,7 +380,8 @@ proved above, and some of the results in later sections depend on the definition lemma rec_zero {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) : Nat.rec h0 h 0 = h0 := rfl #align nat.rec_zero Nat.rec_zero -@[simp] lemma rec_add_one {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) (n : ℕ) : +-- Adaptation note: as of `nightly-2024-03-11`, solved by `simp` now +lemma rec_add_one {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) (n : ℕ) : Nat.rec h0 h (n + 1) = h n (Nat.rec h0 h n) := rfl #align nat.rec_add_one Nat.rec_add_one From 8f916d71adad6bbba6d3595467d029bb9c86e2f3 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 09:53:47 +1100 Subject: [PATCH 31/48] adaptation note about compile_inductive% Nat --- Mathlib/Util/CompileInductive.lean | 5 +++++ test/CompileInductive.lean | 12 ++++++++++-- test/MkIffOfInductive.lean | 6 ++++-- 3 files changed, 19 insertions(+), 4 deletions(-) diff --git a/Mathlib/Util/CompileInductive.lean b/Mathlib/Util/CompileInductive.lean index 1efbfe091abd5..c9ca163be4bf5 100644 --- a/Mathlib/Util/CompileInductive.lean +++ b/Mathlib/Util/CompileInductive.lean @@ -228,6 +228,11 @@ elab tk:"compile_inductive% " i:ident : command => Command.liftTermElabM do end Mathlib.Util +-- Adaptation note: nightly-2024-03-11. +-- Currently we can't run `compile_inductive% Nat`, +-- as `Nat.rec` already has a `@[csimp]` lemma. +-- However this means that we don't generate code for `Nat.recOn` and `Nat.brecOn`. +-- compile_inductive% Nat compile_inductive% List compile_inductive% PUnit compile_inductive% PEmpty diff --git a/test/CompileInductive.lean b/test/CompileInductive.lean index a43b604a7c401..8cb25d5fde6d2 100644 --- a/test/CompileInductive.lean +++ b/test/CompileInductive.lean @@ -9,7 +9,11 @@ example := @Fin2.rec example := @PUnit.rec example := @PEmpty.rec -example := @Nat.recOn +-- Adaptation note: nightly-2024-03-11. +-- Currently we can't run `compile_inductive% Nat`, +-- as `Nat.rec` already has a `@[csimp]` lemma. +-- However this means that we don't generate code for `Nat.recOn` and `Nat.brecOn`. +-- example := @Nat.recOn example := @List.recOn example := @Fin2.recOn example := @PUnit.recOn @@ -18,7 +22,11 @@ example := @And.recOn example := @False.recOn example := @Empty.recOn -example := @Nat.brecOn +-- Adaptation note: nightly-2024-03-11. +-- Currently we can't run `compile_inductive% Nat`, +-- as `Nat.rec` already has a `@[csimp]` lemma. +-- However this means that we don't generate code for `Nat.recOn` and `Nat.brecOn`. +-- example := @Nat.brecOn example := @List.brecOn example := @Fin2.brecOn diff --git a/test/MkIffOfInductive.lean b/test/MkIffOfInductive.lean index fe6972aa3ca46..86899fe9959b5 100644 --- a/test/MkIffOfInductive.lean +++ b/test/MkIffOfInductive.lean @@ -9,8 +9,10 @@ example {α : Type _} (R : α → α → Prop) (a : α) (al : List α) : test.chain_iff R a al -- check that the statement prints nicely -/-- info: test.chain_iff.{u_1} {α : Type u_1} (R : α → α → Prop) (a✝ : α) (a✝¹ : List α) : - List.Chain R a✝ a✝¹ ↔ a✝¹ = [] ∨ ∃ b l, R a✝ b ∧ List.Chain R b l ∧ a✝¹ = b :: l -/ +/-- +info: test.chain_iff.{u_1} {α : Type u_1} (R : α → α → Prop) : + ∀ (a : α) (a_1 : List α), List.Chain R a a_1 ↔ a_1 = [] ∨ ∃ b l, R a b ∧ List.Chain R b l ∧ a_1 = b :: l +-/ #guard_msgs in #check test.chain_iff From ba125ba1c340fc29555ddae906f2c06a008c0b72 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 10:02:31 +1100 Subject: [PATCH 32/48] cleanup CompileInductive --- Mathlib/Util/CompileInductive.lean | 8 +++----- test/CompileInductive.lean | 11 +---------- 2 files changed, 4 insertions(+), 15 deletions(-) diff --git a/Mathlib/Util/CompileInductive.lean b/Mathlib/Util/CompileInductive.lean index c9ca163be4bf5..4bb2ed716ec7c 100644 --- a/Mathlib/Util/CompileInductive.lean +++ b/Mathlib/Util/CompileInductive.lean @@ -228,11 +228,9 @@ elab tk:"compile_inductive% " i:ident : command => Command.liftTermElabM do end Mathlib.Util --- Adaptation note: nightly-2024-03-11. --- Currently we can't run `compile_inductive% Nat`, --- as `Nat.rec` already has a `@[csimp]` lemma. --- However this means that we don't generate code for `Nat.recOn` and `Nat.brecOn`. --- compile_inductive% Nat +-- `Nat.rec` already has a `@[csimp]` lemma in Lean. +compile_def% Nat.recOn +compile_def% Nat.brecOn compile_inductive% List compile_inductive% PUnit compile_inductive% PEmpty diff --git a/test/CompileInductive.lean b/test/CompileInductive.lean index 8cb25d5fde6d2..a9cda4a877804 100644 --- a/test/CompileInductive.lean +++ b/test/CompileInductive.lean @@ -9,11 +9,7 @@ example := @Fin2.rec example := @PUnit.rec example := @PEmpty.rec --- Adaptation note: nightly-2024-03-11. --- Currently we can't run `compile_inductive% Nat`, --- as `Nat.rec` already has a `@[csimp]` lemma. --- However this means that we don't generate code for `Nat.recOn` and `Nat.brecOn`. --- example := @Nat.recOn +example := @Nat.recOn example := @List.recOn example := @Fin2.recOn example := @PUnit.recOn @@ -22,11 +18,6 @@ example := @And.recOn example := @False.recOn example := @Empty.recOn --- Adaptation note: nightly-2024-03-11. --- Currently we can't run `compile_inductive% Nat`, --- as `Nat.rec` already has a `@[csimp]` lemma. --- However this means that we don't generate code for `Nat.recOn` and `Nat.brecOn`. --- example := @Nat.brecOn example := @List.brecOn example := @Fin2.brecOn From c79b04cfdb7a9672e7904526878deb9e4a210c2c Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 19:24:01 -0400 Subject: [PATCH 33/48] lint --- Mathlib/Analysis/SpecialFunctions/Integrals.lean | 4 ---- Mathlib/CategoryTheory/GradedObject.lean | 1 - Mathlib/Data/Nat/Factorization/Basic.lean | 3 +++ 3 files changed, 3 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index ff143805d63c4..ce15b5db4a9e6 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -675,8 +675,6 @@ theorem integral_sin_pow : (∫ x in a..b, sin x ^ (n + 2)) = (sin a ^ (n + 1) * cos a - sin b ^ (n + 1) * cos b) / (n + 2) + (n + 1) / (n + 2) * ∫ x in a..b, sin x ^ n := by - have : n + 2 ≠ 0 := by omega - have : (n : ℝ) + 2 ≠ 0 := by norm_cast field_simp convert eq_sub_iff_add_eq.mp (integral_sin_pow_aux n) using 1 ring @@ -756,8 +754,6 @@ theorem integral_cos_pow : (∫ x in a..b, cos x ^ (n + 2)) = (cos b ^ (n + 1) * sin b - cos a ^ (n + 1) * sin a) / (n + 2) + (n + 1) / (n + 2) * ∫ x in a..b, cos x ^ n := by - have : n + 2 ≠ 0 := by omega - have : (n : ℝ) + 2 ≠ 0 := by norm_cast field_simp convert eq_sub_iff_add_eq.mp (integral_cos_pow_aux n) using 1 ring diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index 77bb31d942824..858426ef02436 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -150,7 +150,6 @@ theorem comapEq_trans {β γ : Type w} {f g h : β → γ} (k : f = g) (l : g = comapEq C (k.trans l) = comapEq C k ≪≫ comapEq C l := by aesop_cat #align category_theory.graded_object.comap_eq_trans CategoryTheory.GradedObject.comapEq_trans -@[simp] theorem eqToHom_apply {β : Type w} {X Y : β → C} (h : X = Y) (b : β) : (eqToHom h : X ⟶ Y) b = eqToHom (by rw [h]) := by subst h diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index f8a04fc6dfa90..cb40900c715e2 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -898,6 +898,9 @@ def recOnMul {P : ℕ → Sort*} (h0 : P 0) (h1 : P 1) (hp : ∀ p, Prime p → recOnPrimeCoprime h0 hp'' fun a b _ _ _ => h a b #align nat.rec_on_mul Nat.recOnMul +--Adaptation note: as of `nightly-2024-03-11`, this is exposed to docBlame +attribute [nolint docBlame] recOnMul.hp'' + /-- For any multiplicative function `f` with `f 1 = 1` and any `n ≠ 0`, we can evaluate `f n` by evaluating `f` at `p ^ k` over the factorization of `n` -/ theorem multiplicative_factorization {β : Type*} [CommMonoid β] (f : ℕ → β) From 89a50213e62d6c150c9b053559d069e5df93f6ae Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 10:24:12 +1100 Subject: [PATCH 34/48] restore test --- test/CompileInductive.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/test/CompileInductive.lean b/test/CompileInductive.lean index a9cda4a877804..a43b604a7c401 100644 --- a/test/CompileInductive.lean +++ b/test/CompileInductive.lean @@ -18,6 +18,7 @@ example := @And.recOn example := @False.recOn example := @Empty.recOn +example := @Nat.brecOn example := @List.brecOn example := @Fin2.brecOn From b5fee616ea842c8511ec8a1320de92146fcd6e57 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 10:26:29 +1100 Subject: [PATCH 35/48] Apply suggestions from code review --- Mathlib/Algebra/Function/Support.lean | 3 ++- Mathlib/Data/Finset/Basic.lean | 1 - Mathlib/Data/Fintype/BigOperators.lean | 1 - Mathlib/Data/Nat/Choose/Basic.lean | 3 +-- Mathlib/Data/Nat/Defs.lean | 1 - 5 files changed, 3 insertions(+), 6 deletions(-) diff --git a/Mathlib/Algebra/Function/Support.lean b/Mathlib/Algebra/Function/Support.lean index 27993451fa272..2436a2b2bfd37 100644 --- a/Mathlib/Algebra/Function/Support.lean +++ b/Mathlib/Algebra/Function/Support.lean @@ -115,7 +115,8 @@ theorem disjoint_mulSupport_iff {f : α → M} {s : Set α} : @[to_additive (attr := simp)] theorem mulSupport_eq_empty_iff {f : α → M} : mulSupport f = ∅ ↔ f = 1 := by - -- Adaptation note: `simp_rw` broke `to_additive` as of `nightly-2024-03-07` + -- Adaptation note: This used to be `simp_rw` rather than `rw`, + -- but this broke `to_additive` as of `nightly-2024-03-07` rw [← subset_empty_iff, mulSupport_subset_iff', funext_iff] simp #align function.mul_support_eq_empty_iff Function.mulSupport_eq_empty_iff diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index ead12d718dffc..2870a37055389 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -2814,7 +2814,6 @@ theorem filter_comm (s : Finset α) : (s.filter p).filter q = (s.filter q).filte #align finset.filter_comm Finset.filter_comm -- We can simplify an application of filter where the decidability is inferred in "the wrong way" --- Adaptation note: as of `nightly-2024-03-11`, Was a simp lemma but no longer applies to itself theorem filter_congr_decidable (s : Finset α) (p : α → Prop) (h : DecidablePred p) [DecidablePred p] : @filter α p h s = s.filter p := by congr #align finset.filter_congr_decidable Finset.filter_congr_decidable diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 83aefedfc1783..495b41601ddb4 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -147,7 +147,6 @@ theorem Fintype.card_fun [DecidableEq α] [Fintype α] [Fintype β] : rw [Fintype.card_pi, Finset.prod_const]; rfl #align fintype.card_fun Fintype.card_fun --- Adaptation note: as of `nightly-2024-03-11`, `Fintype.card_fun` does not have LHS in simp normal form @[simp] theorem Fintype.card_exp_univ_card_eq [DecidableEq α] [Fintype α] [Fintype β] : card β ^ univ.card (α := α) = card β ^ card α := rfl diff --git a/Mathlib/Data/Nat/Choose/Basic.lean b/Mathlib/Data/Nat/Choose/Basic.lean index c628934214f21..6c65d87df4435 100644 --- a/Mathlib/Data/Nat/Choose/Basic.lean +++ b/Mathlib/Data/Nat/Choose/Basic.lean @@ -88,8 +88,7 @@ theorem choose_succ_self (n : ℕ) : choose n (succ n) = 0 := #align nat.choose_succ_self Nat.choose_succ_self @[simp] -theorem choose_one_right (n : ℕ) : choose n 1 = n := by - induction n <;> simp [*, choose, add_comm] +theorem choose_one_right (n : ℕ) : choose n 1 = n := by induction n <;> simp [*, choose, add_comm] #align nat.choose_one_right Nat.choose_one_right -- The `n+1`-st triangle number is `n` more than the `n`-th triangle number diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index c7f4cb0577c92..d60ea02d5340a 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -380,7 +380,6 @@ proved above, and some of the results in later sections depend on the definition lemma rec_zero {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) : Nat.rec h0 h 0 = h0 := rfl #align nat.rec_zero Nat.rec_zero --- Adaptation note: as of `nightly-2024-03-11`, solved by `simp` now lemma rec_add_one {C : ℕ → Sort*} (h0 : C 0) (h : ∀ n, C n → C (n + 1)) (n : ℕ) : Nat.rec h0 h (n + 1) = h n (Nat.rec h0 h n) := rfl #align nat.rec_add_one Nat.rec_add_one From d56a2f25b6e2fe10aae341ee2c12cb67979f8055 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 10:33:57 +1100 Subject: [PATCH 36/48] cleanup --- Mathlib/Analysis/Normed/Group/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index e6ee77689369f..2f81818bcf143 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -1421,7 +1421,6 @@ theorem norm_pos_iff''' [T0Space E] {a : E} : 0 < ‖a‖ ↔ a ≠ 1 := by theorem SeminormedGroup.tendstoUniformlyOn_one {f : ι → κ → G} {s : Set κ} {l : Filter ι} : TendstoUniformlyOn f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ‖f i x‖ < ε := by -- Adaptation note: nightly-2024-03-11. - -- There seems to be a bug in a dsimproc. -- Originally this was `simp_rw` instead of `simp only`, -- but this creates a bad proof term with nested `OfNat.ofNat` that trips up `@[to_additive]`. simp only [tendstoUniformlyOn_iff, Pi.one_apply, dist_one_left] From 477fd6c73845e3225674010fc1c5e27bc3215722 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 10:40:25 +1100 Subject: [PATCH 37/48] cleanup --- Mathlib/Topology/Support.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/Topology/Support.lean b/Mathlib/Topology/Support.lean index 20e7d0e49544a..146c0a7c5c74d 100644 --- a/Mathlib/Topology/Support.lean +++ b/Mathlib/Topology/Support.lean @@ -273,8 +273,6 @@ theorem comp₂_left (hf : HasCompactMulSupport f) because an `OfNat.ofNat 1` is not replaced with `0`. Notably, as of this nightly, what used to look like `OfNat.ofNat (nat_lit 1) x` in the proof term now looks like `OfNat.ofNat (OfNat.ofNat (α := ℕ) (nat_lit 1)) x`, and this seems to trip up `to_additive`. - Likely this is due to a bug in a dsimproc, - which will be easier to diagnose on `nightly-2024-03-12`. -/ filter_upwards [hf, hf₂] using fun x (hx : f x = (1 : α → β) x) (hx₂ : f₂ x = (1 : α → γ) x) => by simp only [hx, hx₂, Pi.one_apply, hm] From 60f150fcfdb7f3aecfc7fa762bb50c9311cf52ce Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 19:43:11 -0400 Subject: [PATCH 38/48] lint --- Mathlib/Data/Fintype/BigOperators.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 83aefedfc1783..0edea7247f640 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -149,7 +149,7 @@ theorem Fintype.card_fun [DecidableEq α] [Fintype α] [Fintype β] : -- Adaptation note: as of `nightly-2024-03-11`, `Fintype.card_fun` does not have LHS in simp normal form @[simp] -theorem Fintype.card_exp_univ_card_eq [DecidableEq α] [Fintype α] [Fintype β] : +theorem Fintype.card_exp_univ_card_eq [Fintype α] [Fintype β] : card β ^ univ.card (α := α) = card β ^ card α := rfl @[simp] From cdcb0add5736a8fb39eef77edad1ae6b88c134f4 Mon Sep 17 00:00:00 2001 From: "F. G. Dorais" Date: Sun, 4 Feb 2024 04:21:42 -0500 Subject: [PATCH 39/48] fix: patch for std4#579 --- Mathlib/Data/List/Sort.lean | 40 ++++++++++--------------------------- 1 file changed, 11 insertions(+), 29 deletions(-) diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index f108c01852703..0585d702ec4db 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -412,14 +412,6 @@ theorem perm_split : ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → l ~ exact ((perm_split e).trans perm_append_comm).cons a #align list.perm_split List.perm_split -/-- Merge two sorted lists into one in linear time. - - merge [1, 2, 4, 5] [0, 1, 3, 4] = [0, 1, 1, 2, 3, 4, 4, 5] -/ -def merge : List α → List α → List α - | [], l' => l' - | l, [] => l - | a :: l, b :: l' => if a ≼ b then a :: merge l (b :: l') else b :: merge (a :: l) l' - termination_by l₁ l₂ => length l₁ + length l₂ #align list.merge List.merge /-- Implementation of a merge sort algorithm to sort a list. -/ @@ -433,28 +425,18 @@ def mergeSort : List α → List α have h := length_split_lt e have := h.1 have := h.2 - exact merge r (mergeSort ls.1) (mergeSort ls.2) + exact merge (r · ·) (mergeSort ls.1) (mergeSort ls.2) termination_by l => length l #align list.merge_sort List.mergeSort @[nolint unusedHavesSuffices] -- Porting note: false positive theorem mergeSort_cons_cons {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - mergeSort r (a :: b :: l) = merge r (mergeSort r l₁) (mergeSort r l₂) := by + mergeSort r (a :: b :: l) = merge (r · ·) (mergeSort r l₁) (mergeSort r l₂) := by simp only [mergeSort, h] #align list.merge_sort_cons_cons List.mergeSort_cons_cons section Correctness -theorem perm_merge : ∀ l l' : List α, merge r l l' ~ l ++ l' - | [], [] => by simp [merge] - | [], b :: l' => by simp [merge] - | a :: l, [] => by simp [merge] - | a :: l, b :: l' => by - by_cases h : a ≼ b - · simpa [merge, h] using perm_merge _ _ - · suffices b :: merge r (a :: l) l' ~ a :: (l ++ b :: l') by simpa [merge, h] - exact ((perm_merge _ _).cons _).trans ((swap _ _ _).trans (perm_middle.symm.cons _)) - termination_by l₁ l₂ => length l₁ + length l₂ #align list.perm_merge List.perm_merge theorem perm_mergeSort : ∀ l : List α, mergeSort r l ~ l @@ -464,7 +446,7 @@ theorem perm_mergeSort : ∀ l : List α, mergeSort r l ~ l cases' e : split (a :: b :: l) with l₁ l₂ cases' length_split_lt e with h₁ h₂ rw [mergeSort_cons_cons r e] - apply (perm_merge r _ _).trans + apply (perm_merge (r · ·) _ _).trans exact ((perm_mergeSort l₁).append (perm_mergeSort l₂)).trans (perm_split e).symm termination_by l => length l @@ -479,14 +461,14 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge r l l') - | [], [], _, _ => by simp [List.merge] - | [], b :: l', _, h₂ => by simpa [List.merge] using h₂ - | a :: l, [], h₁, _ => by simpa [List.merge] using h₁ +theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge (r · ·) l l') + | [], [], _, _ => by simp + | [], b :: l', _, h₂ => by simpa using h₂ + | a :: l, [], h₁, _ => by simpa using h₁ | a :: l, b :: l', h₁, h₂ => by by_cases h : a ≼ b - · suffices ∀ b' ∈ List.merge r l (b :: l'), r a b' by - simpa [List.merge, h, h₁.of_cons.merge h₂] + · suffices ∀ b' ∈ List.merge (r · ·) l (b :: l'), r a b' by + simpa [h, h₁.of_cons.merge h₂] intro b' bm rcases show b' = b ∨ b' ∈ l ∨ b' ∈ l' by simpa [or_left_comm] using (perm_merge _ _ _).subset bm with @@ -495,8 +477,8 @@ theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sort assumption · exact rel_of_sorted_cons h₁ _ bl · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge r (a :: l) l', r b b' by - simpa [List.merge, h, h₁.merge h₂.of_cons] + · suffices ∀ b' ∈ List.merge (r · ·) (a :: l) l', r b b' by + simpa [h, h₁.merge h₂.of_cons] intro b' bm have ba : b ≼ a := (total_of r _ _).resolve_left h have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge _ _ _).subset bm From 7bdcaa7f786555986341728fd735450fdf38415c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 12:02:41 +1100 Subject: [PATCH 40/48] cleanup --- Mathlib/AlgebraicTopology/SimplicialSet.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 7 +++---- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet.lean b/Mathlib/AlgebraicTopology/SimplicialSet.lean index 20e62a14e4310..0c2cb397c11d6 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet.lean @@ -291,7 +291,7 @@ def primitiveTriangle {n : ℕ} (i : Fin (n+4)) use Fin.last (n+3) simp only [hₙ.ne, not_false_eq_true, Fin.zero_eta, zero_add, true_and] intro j - fin_cases j <;> simp [Fin.ext_iff] <;> omega + fin_cases j <;> simp [Fin.ext_iff] · use 0 simp only [h₀.ne', not_false_eq_true, true_and] intro j diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index cb40900c715e2..730f62d44ecb7 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -891,16 +891,15 @@ def recOnPrimeCoprime {P : ℕ → Sort*} (h0 : P 0) (hp : ∀ p n : ℕ, Prime @[elab_as_elim] def recOnMul {P : ℕ → Sort*} (h0 : P 0) (h1 : P 1) (hp : ∀ p, Prime p → P p) (h : ∀ a b, P a → P b → P (a * b)) : ∀ a, P a := - let rec hp'' (p n : ℕ) (hp' : Prime p) : P (p ^ n) := + let rec + /-- The predicate holds on prime powers -/ + hp'' (p n : ℕ) (hp' : Prime p) : P (p ^ n) := match n with | 0 => h1 | n + 1 => h _ _ (hp'' p n hp') (hp p hp') recOnPrimeCoprime h0 hp'' fun a b _ _ _ => h a b #align nat.rec_on_mul Nat.recOnMul ---Adaptation note: as of `nightly-2024-03-11`, this is exposed to docBlame -attribute [nolint docBlame] recOnMul.hp'' - /-- For any multiplicative function `f` with `f 1 = 1` and any `n ≠ 0`, we can evaluate `f n` by evaluating `f` at `p ^ k` over the factorization of `n` -/ theorem multiplicative_factorization {β : Type*} [CommMonoid β] (f : ℕ → β) From 66347a3b9f5731c099e4d587354cad7d596533d3 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 12:29:17 +1100 Subject: [PATCH 41/48] fix tests --- test/LibrarySearch/mathlib.lean | 2 +- test/MoveAdd.lean | 12 +++++++----- 2 files changed, 8 insertions(+), 6 deletions(-) diff --git a/test/LibrarySearch/mathlib.lean b/test/LibrarySearch/mathlib.lean index 6f1315a15bccc..fd0b69b67bd73 100644 --- a/test/LibrarySearch/mathlib.lean +++ b/test/LibrarySearch/mathlib.lean @@ -3,6 +3,6 @@ import Mathlib -- We verify that `exact?` copes with all of Mathlib. -- On `v4.7.0-rc1` this revealed a cache corruption problem. -/-- info: Try this: exact Nat.zero_lt_one -/ +/-- info: Try this: exact Nat.one_pos -/ #guard_msgs in example : 0 < 1 := by exact? diff --git a/test/MoveAdd.lean b/test/MoveAdd.lean index 39584af348619..a253c8800e2eb 100644 --- a/test/MoveAdd.lean +++ b/test/MoveAdd.lean @@ -109,11 +109,13 @@ example {a b c d e : Prop} (h : a ∨ b ∨ c ∨ d ∨ e) : a ∨ c ∨ e ∨ b end left_assoc -example (k : ℕ) (h0 : 0 + 2 = 9 + 0) (h9 : k + 2 = k + 9) : k + 2 = 9 + k := by - induction' k with k _ih - · exact h0 - · move_add [9] - exact h9 +-- Adaptation note: nightly-2024-03-11 +-- This test is now failing with `unknown free variable '_fvar.36787'` +-- example (k : ℕ) (h0 : 0 + 2 = 9 + 0) (h9 : k + 2 = k + 9) : k + 2 = 9 + k := by +-- induction' k with k _ih +-- · exact h0 +-- · move_add [9] +-- exact h9 -- Testing internals of the tactic `move_add`. section tactic From ac499c5eda82f663b1dcbe724630971fac12885e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 12:30:11 +1100 Subject: [PATCH 42/48] fix tests --- test/matrix.lean | 29 +++++++++++++++-------------- 1 file changed, 15 insertions(+), 14 deletions(-) diff --git a/test/matrix.lean b/test/matrix.lean index d89c6f7fe5d0a..67e37524a5a69 100644 --- a/test/matrix.lean +++ b/test/matrix.lean @@ -135,26 +135,27 @@ example {a b c d e f g h : α} : ![a, b, c, d, e, f, g, h] 99 = d := by simp example {α : Type _} [CommRing α] {a b c d : α} : Matrix.det !![a, b; c, d] = a * d - b * c := by simp? [Matrix.det_succ_row_zero, Fin.sum_univ_succ] says - simp only [det_succ_row_zero, Fin.isValue, of_apply, cons_val', empty_val', - cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, submatrix_apply, - Fin.succ_zero_eq_one, cons_val_one, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, - one_mul, Fin.zero_succAbove, head_cons, Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, - zero_add, pow_one, cons_val_succ, neg_mul, Fin.succ_succAbove_zero, Finset.sum_const, - Finset.card_singleton, smul_neg, one_smul] + simp only [det_succ_row_zero, Nat.reduceSucc, Fin.isValue, of_apply, + cons_val', empty_val', cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, + submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, head_fin_const, Fin.sum_univ_succ, + Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove, head_cons, Finset.univ_unique, + Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, cons_val_succ, neg_mul, + Fin.succ_succAbove_zero, Finset.sum_const, Finset.card_singleton, smul_neg, one_smul] ring example {α : Type _} [CommRing α] {a b c d e f g h i : α} : Matrix.det !![a, b, c; d, e, f; g, h, i] = a * e * i - a * f * h - b * d * i + b * f * g + c * d * h - c * e * g := by simp? [Matrix.det_succ_row_zero, Fin.sum_univ_succ] says - simp only [det_succ_row_zero, Fin.isValue, of_apply, cons_val', empty_val', - cons_val_fin_one, cons_val_zero, submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, - head_cons, submatrix_submatrix, det_unique, Fin.default_eq_zero, Function.comp_apply, - Fin.succ_one_eq_two, cons_val_two, tail_cons, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, - pow_zero, one_mul, Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, - zero_add, pow_one, neg_mul, Fin.succ_succAbove_zero, Finset.sum_neg_distrib, - Finset.sum_singleton, cons_val_succ, Fin.succ_succAbove_one, Nat.reduceAdd, even_two, - Even.neg_pow, one_pow, Finset.sum_const, Finset.card_singleton, one_smul] + simp only [det_succ_row_zero, Nat.reduceSucc, Fin.isValue, of_apply, + cons_val', empty_val', cons_val_fin_one, cons_val_zero, submatrix_apply, Fin.succ_zero_eq_one, + cons_val_one, head_cons, submatrix_submatrix, det_unique, Fin.default_eq_zero, + Function.comp_apply, Fin.succ_one_eq_two, cons_val_two, tail_cons, head_fin_const, + Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove, Finset.univ_unique, + Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, neg_mul, Fin.succ_succAbove_zero, + Finset.sum_neg_distrib, Finset.sum_singleton, cons_val_succ, Fin.succ_succAbove_one, + Nat.reduceAdd, even_two, Even.neg_pow, one_pow, Finset.sum_const, Finset.card_singleton, + one_smul] ring end Matrix From a83ef92f0f1bce76d7124ec9a4368101273b23ae Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 12:39:12 +1100 Subject: [PATCH 43/48] shake --- Mathlib/Tactic/Sat/FromLRAT.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Tactic/Sat/FromLRAT.lean b/Mathlib/Tactic/Sat/FromLRAT.lean index aee43b246077c..0045bc945331a 100644 --- a/Mathlib/Tactic/Sat/FromLRAT.lean +++ b/Mathlib/Tactic/Sat/FromLRAT.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Basic -import Mathlib.Init.Data.List.Instances /-! # `lrat_proof` command From 4fc25a95a063b840d77ccc88d5ba6e186dabe9aa Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 12:46:12 +1100 Subject: [PATCH 44/48] bump Std --- lake-manifest.json | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index ddd6199598bdc..7ea6e91451b61 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,7 +4,7 @@ [{"url": "https://github.com/leanprover/std4", "type": "git", "subDir": null, - "rev": "ff9850c4726f6b9fb8d8e96980c3fcb2900be8bd", + "rev": "f3447c3732c9d6e8df3bdad78e5ecf7e8b353bbc", "name": "std", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -22,7 +22,7 @@ {"url": "https://github.com/leanprover-community/aesop", "type": "git", "subDir": null, - "rev": "056ca0fa8f5585539d0b940f532d9750c3a2270f", + "rev": "8be30c25e3caa06937feeb62f7ca898370f80ee9", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -40,7 +40,7 @@ {"url": "https://github.com/leanprover/lean4-cli", "type": "git", "subDir": null, - "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec", + "rev": "be8fa79a28b8b6897dce0713ef50e89c4a0f6ef5", "name": "Cli", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -49,7 +49,7 @@ {"url": "https://github.com/leanprover-community/import-graph.git", "type": "git", "subDir": null, - "rev": "64d082eeaad1a8e6bbb7c23b7a16b85a1715a02f", + "rev": "bbcffbcc19d69e13d5eea716283c76169db524ba", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", From 1ae214ea83b0479e194d21e87f6992dee2b082ec Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 12:50:03 +1100 Subject: [PATCH 45/48] revert merge' --- Mathlib/Data/List/Sort.lean | 44 ++++++++++++++++++------------------- 1 file changed, 22 insertions(+), 22 deletions(-) diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 8812d44d183d4..f108c01852703 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -415,10 +415,10 @@ theorem perm_split : ∀ {l l₁ l₂ : List α}, split l = (l₁, l₂) → l ~ /-- Merge two sorted lists into one in linear time. merge [1, 2, 4, 5] [0, 1, 3, 4] = [0, 1, 1, 2, 3, 4, 4, 5] -/ -def merge' : List α → List α → List α +def merge : List α → List α → List α | [], l' => l' | l, [] => l - | a :: l, b :: l' => if a ≼ b then a :: merge' l (b :: l') else b :: merge' (a :: l) l' + | a :: l, b :: l' => if a ≼ b then a :: merge l (b :: l') else b :: merge (a :: l) l' termination_by l₁ l₂ => length l₁ + length l₂ #align list.merge List.merge @@ -433,27 +433,27 @@ def mergeSort : List α → List α have h := length_split_lt e have := h.1 have := h.2 - exact merge' r (mergeSort ls.1) (mergeSort ls.2) + exact merge r (mergeSort ls.1) (mergeSort ls.2) termination_by l => length l #align list.merge_sort List.mergeSort @[nolint unusedHavesSuffices] -- Porting note: false positive theorem mergeSort_cons_cons {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - mergeSort r (a :: b :: l) = merge' r (mergeSort r l₁) (mergeSort r l₂) := by + mergeSort r (a :: b :: l) = merge r (mergeSort r l₁) (mergeSort r l₂) := by simp only [mergeSort, h] #align list.merge_sort_cons_cons List.mergeSort_cons_cons section Correctness -theorem perm_merge' : ∀ l l' : List α, merge' r l l' ~ l ++ l' - | [], [] => by simp [merge'] - | [], b :: l' => by simp [merge'] - | a :: l, [] => by simp [merge'] +theorem perm_merge : ∀ l l' : List α, merge r l l' ~ l ++ l' + | [], [] => by simp [merge] + | [], b :: l' => by simp [merge] + | a :: l, [] => by simp [merge] | a :: l, b :: l' => by by_cases h : a ≼ b - · simpa [merge', h] using perm_merge' _ _ - · suffices b :: merge' r (a :: l) l' ~ a :: (l ++ b :: l') by simpa [merge', h] - exact ((perm_merge' _ _).cons _).trans ((swap _ _ _).trans (perm_middle.symm.cons _)) + · simpa [merge, h] using perm_merge _ _ + · suffices b :: merge r (a :: l) l' ~ a :: (l ++ b :: l') by simpa [merge, h] + exact ((perm_merge _ _).cons _).trans ((swap _ _ _).trans (perm_middle.symm.cons _)) termination_by l₁ l₂ => length l₁ + length l₂ #align list.perm_merge List.perm_merge @@ -464,7 +464,7 @@ theorem perm_mergeSort : ∀ l : List α, mergeSort r l ~ l cases' e : split (a :: b :: l) with l₁ l₂ cases' length_split_lt e with h₁ h₂ rw [mergeSort_cons_cons r e] - apply (perm_merge' r _ _).trans + apply (perm_merge r _ _).trans exact ((perm_mergeSort l₁).append (perm_mergeSort l₂)).trans (perm_split e).symm termination_by l => length l @@ -479,27 +479,27 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge' r l l') - | [], [], _, _ => by simp [List.merge'] - | [], b :: l', _, h₂ => by simpa [List.merge'] using h₂ - | a :: l, [], h₁, _ => by simpa [List.merge'] using h₁ +theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge r l l') + | [], [], _, _ => by simp [List.merge] + | [], b :: l', _, h₂ => by simpa [List.merge] using h₂ + | a :: l, [], h₁, _ => by simpa [List.merge] using h₁ | a :: l, b :: l', h₁, h₂ => by by_cases h : a ≼ b - · suffices ∀ b' ∈ List.merge' r l (b :: l'), r a b' by - simpa [List.merge', h, h₁.of_cons.merge h₂] + · suffices ∀ b' ∈ List.merge r l (b :: l'), r a b' by + simpa [List.merge, h, h₁.of_cons.merge h₂] intro b' bm rcases show b' = b ∨ b' ∈ l ∨ b' ∈ l' by - simpa [or_left_comm] using (perm_merge' _ _ _).subset bm with + simpa [or_left_comm] using (perm_merge _ _ _).subset bm with (be | bl | bl') · subst b' assumption · exact rel_of_sorted_cons h₁ _ bl · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge' r (a :: l) l', r b b' by - simpa [List.merge', h, h₁.merge h₂.of_cons] + · suffices ∀ b' ∈ List.merge r (a :: l) l', r b b' by + simpa [List.merge, h, h₁.merge h₂.of_cons] intro b' bm have ba : b ≼ a := (total_of r _ _).resolve_left h - have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge' _ _ _).subset bm + have : b' = a ∨ b' ∈ l ∨ b' ∈ l' := by simpa using (perm_merge _ _ _).subset bm rcases this with (be | bl | bl') · subst b' assumption From e31222f3fe275a43262b38f71c7efd038ca7bb45 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 21:51:10 -0400 Subject: [PATCH 46/48] comment more --- Mathlib/RingTheory/Henselian.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/RingTheory/Henselian.lean b/Mathlib/RingTheory/Henselian.lean index 41b2c1f8d6dcf..82fe646dc8945 100644 --- a/Mathlib/RingTheory/Henselian.lean +++ b/Mathlib/RingTheory/Henselian.lean @@ -240,6 +240,8 @@ instance (priority := 100) IsAdicComplete.henselianRing (R : Type*) [CommRing R] rw [← add_assoc] -- Adaptation note: nightly-2024-03-11 -- I'm not sure why the `erw` is now needed here. It looks like it should work. + -- It looks like a diamond between `instHAdd` on `Nat` and `AddSemigroup.toAdd` which is + -- used by `instHAdd` erw [hc] rw [← add_zero (c m), sub_eq_add_neg] refine' ih.add _ From b77d8589548b1ea560effb6d05325a5c172f0a04 Mon Sep 17 00:00:00 2001 From: Matthew Ballard Date: Tue, 12 Mar 2024 22:02:38 -0400 Subject: [PATCH 47/48] add comment --- Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index 28c1c8f99269e..5c488487c8061 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -28,6 +28,7 @@ When the underlying category is abelian: is injective, we can apply `Injective.d` repeatedly to obtain an injective resolution of `X`. -/ +set_option profiler true noncomputable section @@ -322,6 +323,7 @@ lemma ofCocomplex_d_0_1 : (ofCocomplex Z).d 0 1 = d (Injective.ι Z) := by simp [ofCocomplex] +--Adaptation note: nightly-2024-03-11. This takes takes forever now lemma ofCocomplex_exactAt_succ (n : ℕ) : (ofCocomplex Z).ExactAt (n + 1) := by rw [HomologicalComplex.exactAt_iff' _ n (n + 1) (n + 1 + 1) (by simp) (by simp)] From 992b5503e06a762facca7c8cfb34ebc083547277 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 13 Mar 2024 13:04:40 +1100 Subject: [PATCH 48/48] cleanup --- Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean index 5c488487c8061..05aa22c692a53 100644 --- a/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean +++ b/Mathlib/CategoryTheory/Abelian/InjectiveResolution.lean @@ -28,8 +28,6 @@ When the underlying category is abelian: is injective, we can apply `Injective.d` repeatedly to obtain an injective resolution of `X`. -/ -set_option profiler true - noncomputable section open CategoryTheory Category Limits