From 40b64f797572a7615ba4553a28d3d6ee5049aa3b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 17 Nov 2023 05:53:46 +0000 Subject: [PATCH] chore: bump to v4.3.0-rc2 (#8366) # PR contents This is the supremum of - #8284 - #8056 - #8023 - #8332 - #8226 (already approved) - #7834 (already approved) along with some minor fixes from failures on nightly-testing as Mathlib `master` is merged into it. Note that some PRs for changes that are already compatible with the current toolchain and will be necessary have already been split out: #8380. I am hopeful that in future we will be able to progressively merge adaptation PRs into a `bump/v4.X.0` branch, so we never end up with a "big merge" like this. However one of these adaptation PRs (#8056) predates my new scheme for combined CI, and it wasn't possible to keep that PR viable in the meantime. # Lean PRs involved in this bump In particular this includes adjustments for the Lean PRs * leanprover/lean4#2778 * leanprover/lean4#2790 * leanprover/lean4#2783 * leanprover/lean4#2825 * leanprover/lean4#2722 ## leanprover/lean4#2778 We can get rid of all the ``` local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 ``` macros across Mathlib (and in any projects that want to write natural number powers of reals). ## leanprover/lean4#2722 Changes the default behaviour of `simp` to `(config := {decide := false})`. This makes `simp` (and consequentially `norm_num`) less powerful, but also more consistent, and less likely to blow up in long failures. This requires a variety of changes: changing some previously by `simp` or `norm_num` to `decide` or `rfl`, or adding `(config := {decide := true})`. ## leanprover/lean4#2783 This changed the behaviour of `simp` so that `simp [f]` will only unfold "fully applied" occurrences of `f`. The old behaviour can be recovered with `simp (config := { unfoldPartialApp := true })`. We may in future add a syntax for this, e.g. `simp [!f]`; please provide feedback! In the meantime, we have made the following changes: * switching to using explicit lemmas that have the intended level of application * `(config := { unfoldPartialApp := true })` in some places, to recover the old behaviour * Using `@[eqns]` to manually adjust the equation lemmas for a particular definition, recovering the old behaviour just for that definition. See #8371, where we do this for `Function.comp` and `Function.flip`. This change in Lean may require further changes down the line (e.g. adding the `!f` syntax, and/or upstreaming the special treatment for `Function.comp` and `Function.flip`, and/or removing this special treatment). Please keep an open and skeptical mind about these changes! Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Scott Morrison Co-authored-by: Eric Wieser Co-authored-by: Mauricio Collares --- .gitignore | 4 +- Archive/Examples/IfNormalization/Result.lean | 9 +- .../IfNormalization/WithoutAesop.lean | 10 +- Archive/Imo/Imo1960Q1.lean | 7 +- Archive/Imo/Imo1962Q1.lean | 13 +- Archive/Imo/Imo1962Q4.lean | 2 - Archive/Imo/Imo1964Q1.lean | 4 +- Archive/Imo/Imo1981Q3.lean | 6 +- Archive/Imo/Imo2001Q2.lean | 2 - Archive/Imo/Imo2005Q4.lean | 6 +- Archive/Imo/Imo2008Q3.lean | 2 - Archive/Imo/Imo2013Q5.lean | 2 - Archive/Imo/Imo2019Q1.lean | 2 +- Archive/Imo/Imo2019Q2.lean | 2 +- Archive/Imo/Imo2019Q4.lean | 13 +- Archive/MiuLanguage/DecisionNec.lean | 6 +- Archive/MiuLanguage/DecisionSuf.lean | 7 +- Archive/Sensitivity.lean | 2 +- Archive/Wiedijk100Theorems/AbelRuffini.lean | 15 +- Archive/Wiedijk100Theorems/AreaOfACircle.lean | 3 - Archive/Wiedijk100Theorems/BallotProblem.lean | 4 +- .../Wiedijk100Theorems/BirthdayProblem.lean | 20 +- Archive/Wiedijk100Theorems/HeronsFormula.lean | 2 - Archive/Wiedijk100Theorems/Konigsberg.lean | 2 +- .../Wiedijk100Theorems/PerfectNumbers.lean | 4 +- .../SumOfPrimeReciprocalsDiverges.lean | 2 +- Archive/ZagierTwoSquares.lean | 9 +- Cache/IO.lean | 6 +- .../CliffordAlgebra_not_injective.lean | 3 +- Counterexamples/Cyclotomic105.lean | 2 +- Counterexamples/HomogeneousPrimeNotPrime.lean | 2 +- Mathlib.lean | 2 - .../Algebra/Algebra/NonUnitalSubalgebra.lean | 2 +- Mathlib/Algebra/BigOperators/Basic.lean | 6 +- .../Category/ModuleCat/ChangeOfRings.lean | 2 - Mathlib/Algebra/CharP/Basic.lean | 4 +- Mathlib/Algebra/CharP/LocalRing.lean | 2 +- Mathlib/Algebra/CharZero/Quotient.lean | 2 +- Mathlib/Algebra/GeomSum.lean | 11 +- Mathlib/Algebra/GradedMonoid.lean | 2 +- Mathlib/Algebra/Group/Defs.lean | 45 +++- Mathlib/Algebra/GroupPower/NegOnePow.lean | 3 +- Mathlib/Algebra/Homology/Augment.lean | 2 +- Mathlib/Algebra/Homology/Exact.lean | 2 +- Mathlib/Algebra/Homology/ExactSequence.lean | 7 +- Mathlib/Algebra/Homology/Flip.lean | 2 +- Mathlib/Algebra/Homology/Homotopy.lean | 4 +- Mathlib/Algebra/Homology/Single.lean | 1 + Mathlib/Algebra/IsPrimePow.lean | 2 +- Mathlib/Algebra/Lie/DirectSum.lean | 2 +- Mathlib/Algebra/Module/DedekindDomain.lean | 6 +- Mathlib/Algebra/Module/Injective.lean | 2 +- Mathlib/Algebra/Order/Interval.lean | 2 +- Mathlib/Algebra/Order/Positive/Ring.lean | 2 +- Mathlib/Algebra/Order/Ring/WithTop.lean | 2 +- Mathlib/Algebra/Order/Sub/WithTop.lean | 4 +- Mathlib/Algebra/Quaternion.lean | 2 +- Mathlib/Algebra/Regular/SMul.lean | 6 +- Mathlib/Algebra/RingQuot.lean | 2 +- Mathlib/Algebra/Star/NonUnitalSubalgebra.lean | 2 +- .../EllipticCurve/Weierstrass.lean | 6 +- .../DoldKan/NReflectsIso.lean | 2 +- Mathlib/Analysis/Analytic/Basic.lean | 6 +- Mathlib/Analysis/Analytic/Composition.lean | 8 +- Mathlib/Analysis/Analytic/RadiusLiminf.lean | 2 - .../Asymptotics/AsymptoticEquivalent.lean | 2 +- Mathlib/Analysis/BoxIntegral/Basic.lean | 2 - .../BoxIntegral/DivergenceTheorem.lean | 2 - .../Analysis/BoxIntegral/Partition/Split.lean | 6 +- .../Calculus/BumpFunction/Convolution.lean | 2 - .../Calculus/BumpFunction/Normed.lean | 2 - Mathlib/Analysis/Calculus/Monotone.lean | 2 - Mathlib/Analysis/Complex/AbsMax.lean | 1 - Mathlib/Analysis/Complex/Liouville.lean | 3 - .../Analysis/Complex/LocallyUniformLimit.lean | 2 - .../Analysis/Complex/PhragmenLindelof.lean | 1 - .../Complex/RemovableSingularity.lean | 2 - Mathlib/Analysis/ConstantSpeed.lean | 2 +- Mathlib/Analysis/Convex/Between.lean | 3 +- .../Convex/SpecificFunctions/Basic.lean | 3 - .../Convex/SpecificFunctions/Deriv.lean | 4 +- .../Convex/SpecificFunctions/Pow.lean | 2 - Mathlib/Analysis/Convex/Strong.lean | 2 - Mathlib/Analysis/Convolution.lean | 3 +- .../Analysis/Distribution/SchwartzSpace.lean | 2 - Mathlib/Analysis/Fourier/AddCircle.lean | 12 +- .../Fourier/RiemannLebesgueLemma.lean | 6 +- .../Analysis/InnerProductSpace/Adjoint.lean | 2 - .../Analysis/InnerProductSpace/Calculus.lean | 6 +- Mathlib/Analysis/InnerProductSpace/PiL2.lean | 2 - .../Analysis/InnerProductSpace/Rayleigh.lean | 6 +- .../Analysis/InnerProductSpace/Spectrum.lean | 2 - .../Analysis/InnerProductSpace/TwoDim.lean | 4 +- .../Analysis/InnerProductSpace/l2Space.lean | 2 - Mathlib/Analysis/Matrix.lean | 3 - Mathlib/Analysis/MeanInequalities.lean | 2 - Mathlib/Analysis/MeanInequalitiesPow.lean | 2 - Mathlib/Analysis/MellinTransform.lean | 4 +- Mathlib/Analysis/NormedSpace/DualNumber.lean | 2 +- Mathlib/Analysis/NormedSpace/Exponential.lean | 2 +- Mathlib/Analysis/NormedSpace/Multilinear.lean | 2 +- Mathlib/Analysis/NormedSpace/PiLp.lean | 4 +- Mathlib/Analysis/NormedSpace/ProdLp.lean | 9 +- .../NormedSpace/QuaternionExponential.lean | 6 +- Mathlib/Analysis/NormedSpace/Spectrum.lean | 5 +- Mathlib/Analysis/NormedSpace/Star/Matrix.lean | 4 +- Mathlib/Analysis/NormedSpace/lpSpace.lean | 3 - Mathlib/Analysis/ODE/PicardLindelof.lean | 2 - Mathlib/Analysis/PSeries.lean | 12 +- .../Analysis/SpecialFunctions/CompareExp.lean | 2 - .../SpecialFunctions/Gamma/Basic.lean | 2 +- .../Analysis/SpecialFunctions/Gamma/Beta.lean | 10 +- .../Analysis/SpecialFunctions/Gaussian.lean | 2 - .../Analysis/SpecialFunctions/Integrals.lean | 20 +- .../SpecialFunctions/JapaneseBracket.lean | 1 - .../Analysis/SpecialFunctions/Log/Base.lean | 4 +- .../Analysis/SpecialFunctions/Log/Basic.lean | 3 +- .../Analysis/SpecialFunctions/PolarCoord.lean | 2 - .../SpecialFunctions/Pow/Complex.lean | 2 - .../SpecialFunctions/Pow/Continuity.lean | 2 - .../Analysis/SpecialFunctions/Pow/Deriv.lean | 6 +- .../Analysis/SpecialFunctions/Pow/NNReal.lean | 4 +- .../Analysis/SpecialFunctions/Pow/Real.lean | 6 +- .../SpecialFunctions/SmoothTransition.lean | 1 - .../Analysis/SpecialFunctions/Stirling.lean | 4 +- .../SpecialFunctions/Trigonometric/Angle.lean | 4 +- .../Trigonometric/Arctan.lean | 2 - .../Trigonometric/ArctanDeriv.lean | 2 - .../Trigonometric/Bounds.lean | 4 +- .../Trigonometric/Complex.lean | 2 - .../Trigonometric/ComplexDeriv.lean | 2 - .../Trigonometric/EulerSineProd.lean | 7 +- .../Trigonometric/Series.lean | 4 +- Mathlib/Analysis/SpecificLimits/FloorPow.lean | 2 - Mathlib/CategoryTheory/Action.lean | 2 +- Mathlib/CategoryTheory/Category/Basic.lean | 1 + .../CategoryTheory/Category/KleisliCat.lean | 2 +- Mathlib/CategoryTheory/CofilteredSystem.lean | 8 +- .../CategoryTheory/DifferentialObject.lean | 4 +- Mathlib/CategoryTheory/Elements.lean | 2 +- .../CategoryTheory/Groupoid/Subgroupoid.lean | 2 +- .../Limits/Preserves/Shapes/Biproducts.lean | 4 +- Mathlib/CategoryTheory/Preadditive/Mat.lean | 4 +- Mathlib/CategoryTheory/Sites/Types.lean | 2 +- Mathlib/Combinatorics/Additive/Behrend.lean | 22 +- Mathlib/Combinatorics/Quiver/Covering.lean | 4 +- .../SetFamily/FourFunctions.lean | 3 +- .../SetFamily/HarrisKleitman.lean | 7 +- Mathlib/Combinatorics/SimpleGraph/Clique.lean | 2 +- .../SimpleGraph/Regularity/Bound.lean | 6 +- .../SimpleGraph/Regularity/Chunk.lean | 2 - .../SimpleGraph/Regularity/Energy.lean | 2 - .../SimpleGraph/Regularity/Increment.lean | 6 +- .../SimpleGraph/Regularity/Lemma.lean | 2 - .../SimpleGraph/StronglyRegular.lean | 3 +- Mathlib/Combinatorics/Young/YoungDiagram.lean | 2 +- Mathlib/Computability/Ackermann.lean | 2 +- Mathlib/Computability/Halting.lean | 2 +- Mathlib/Computability/Primrec.lean | 2 +- Mathlib/Computability/TMComputable.lean | 2 +- Mathlib/Computability/TMToPartrec.lean | 2 +- Mathlib/Control/Applicative.lean | 3 +- Mathlib/Control/Basic.lean | 14 +- Mathlib/Control/Functor/Multivariate.lean | 9 +- Mathlib/Control/Traversable/Equiv.lean | 5 +- Mathlib/Data/Bitvec/Lemmas.lean | 6 +- Mathlib/Data/Bool/Basic.lean | 4 +- Mathlib/Data/Complex/Exponential.lean | 2 +- Mathlib/Data/DFinsupp/WellFounded.lean | 2 +- Mathlib/Data/ENat/Basic.lean | 6 +- Mathlib/Data/Fin/Basic.lean | 6 +- Mathlib/Data/Fin/Tuple/Basic.lean | 8 +- Mathlib/Data/Fin/Tuple/Curry.lean | 4 +- Mathlib/Data/Fin/Tuple/Monotone.lean | 7 +- Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Finset/NAry.lean | 4 +- Mathlib/Data/Finsupp/BigOperators.lean | 2 +- Mathlib/Data/Finsupp/Defs.lean | 2 +- Mathlib/Data/Holor.lean | 3 +- Mathlib/Data/Int/Bitwise.lean | 18 +- Mathlib/Data/Int/Lemmas.lean | 2 +- Mathlib/Data/Int/ModEq.lean | 2 +- Mathlib/Data/Int/Parity.lean | 8 +- Mathlib/Data/Int/Range.lean | 2 +- Mathlib/Data/Int/Units.lean | 2 +- Mathlib/Data/KVMap.lean | 27 --- Mathlib/Data/LazyList/Basic.lean | 2 +- Mathlib/Data/List/Basic.lean | 10 +- Mathlib/Data/List/BigOperators/Basic.lean | 2 +- Mathlib/Data/List/Chain.lean | 2 +- Mathlib/Data/List/DropRight.lean | 213 +++++++++--------- Mathlib/Data/List/EditDistance/Defs.lean | 2 +- Mathlib/Data/List/Indexes.lean | 11 +- Mathlib/Data/List/Infix.lean | 2 +- Mathlib/Data/List/Intervals.lean | 6 +- Mathlib/Data/List/Lattice.lean | 4 +- Mathlib/Data/List/Perm.lean | 6 +- Mathlib/Data/List/Range.lean | 5 +- Mathlib/Data/List/Sigma.lean | 2 +- Mathlib/Data/Matrix/Basis.lean | 4 +- Mathlib/Data/Multiset/Basic.lean | 4 +- Mathlib/Data/MvPolynomial/Division.lean | 6 +- Mathlib/Data/MvPolynomial/Variables.lean | 2 +- Mathlib/Data/Nat/Bitwise.lean | 21 +- Mathlib/Data/Nat/Choose/Multinomial.lean | 2 +- Mathlib/Data/Nat/Digits.lean | 7 +- Mathlib/Data/Nat/Factorial/BigOperators.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 5 +- Mathlib/Data/Nat/Hyperoperation.lean | 2 +- Mathlib/Data/Nat/Log.lean | 2 +- Mathlib/Data/Nat/ModEq.lean | 2 +- Mathlib/Data/Nat/Parity.lean | 5 +- Mathlib/Data/Nat/PartENat.lean | 4 +- Mathlib/Data/Nat/Sqrt.lean | 4 +- Mathlib/Data/Nat/Squarefree.lean | 2 +- Mathlib/Data/Nat/WithBot.lean | 9 +- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/Option/Basic.lean | 2 +- Mathlib/Data/PFunctor/Multivariate/W.lean | 2 +- Mathlib/Data/PNat/Basic.lean | 4 +- Mathlib/Data/PNat/Prime.lean | 2 +- Mathlib/Data/Polynomial/Coeff.lean | 4 +- Mathlib/Data/Polynomial/Div.lean | 2 +- Mathlib/Data/Polynomial/FieldDivision.lean | 6 +- Mathlib/Data/Polynomial/RingDivision.lean | 4 +- Mathlib/Data/Polynomial/UnitTrinomial.lean | 1 + .../QPF/Multivariate/Constructions/Cofix.lean | 4 +- .../QPF/Multivariate/Constructions/Comp.lean | 14 +- Mathlib/Data/QPF/Univariate/Basic.lean | 4 +- Mathlib/Data/Quot.lean | 2 +- Mathlib/Data/Rat/Cast/CharZero.lean | 2 +- Mathlib/Data/Rat/Defs.lean | 6 +- Mathlib/Data/Rat/Lemmas.lean | 1 + Mathlib/Data/Real/ENNReal.lean | 16 +- Mathlib/Data/Real/Pi/Bounds.lean | 2 - Mathlib/Data/Real/Pi/Leibniz.lean | 2 - Mathlib/Data/Real/Pi/Wallis.lean | 6 +- Mathlib/Data/Rel.lean | 8 +- Mathlib/Data/Seq/Parallel.lean | 2 +- Mathlib/Data/Seq/WSeq.lean | 7 +- Mathlib/Data/Set/Card.lean | 2 +- Mathlib/Data/Set/Function.lean | 2 +- Mathlib/Data/Set/Intervals/Group.lean | 6 +- .../Set/Intervals/OrdConnectedComponent.lean | 3 +- Mathlib/Data/Set/Pointwise/Interval.lean | 18 +- Mathlib/Data/Sign.lean | 4 +- Mathlib/Data/String/Basic.lean | 8 +- Mathlib/Data/Sym/Card.lean | 4 +- Mathlib/Data/ZMod/Basic.lean | 2 +- Mathlib/Data/ZMod/IntUnitsPower.lean | 2 - Mathlib/FieldTheory/AbelRuffini.lean | 2 - Mathlib/FieldTheory/IsAlgClosed/Spectrum.lean | 2 - Mathlib/FieldTheory/PrimitiveElement.lean | 6 +- Mathlib/FieldTheory/RatFunc.lean | 3 +- Mathlib/FieldTheory/Subfield.lean | 2 +- .../Euclidean/Angle/Oriented/Basic.lean | 2 +- .../Angle/Unoriented/RightAngle.lean | 2 - .../Euclidean/Inversion/Calculus.lean | 5 +- Mathlib/Geometry/Euclidean/MongePoint.lean | 19 +- Mathlib/Geometry/Euclidean/Sphere/Basic.lean | 2 +- Mathlib/Geometry/Euclidean/Triangle.lean | 2 - Mathlib/Geometry/Manifold/ChartedSpace.lean | 1 - Mathlib/Geometry/Manifold/ContMDiff.lean | 2 +- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 8 +- .../Geometry/Manifold/Instances/Sphere.lean | 4 - .../Manifold/VectorBundle/Tangent.lean | 3 +- Mathlib/GroupTheory/CommutingProbability.lean | 4 +- Mathlib/GroupTheory/FiniteAbelian.lean | 2 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 2 +- .../GroupAction/SubMulAction/Pointwise.lean | 4 +- Mathlib/GroupTheory/HNNExtension.lean | 2 +- Mathlib/GroupTheory/OrderOfElement.lean | 2 +- Mathlib/GroupTheory/Perm/Fin.lean | 2 +- Mathlib/GroupTheory/PushoutI.lean | 2 +- .../SpecificGroups/Alternating.lean | 5 +- .../GroupTheory/SpecificGroups/Dihedral.lean | 4 +- .../GroupTheory/SpecificGroups/KleinFour.lean | 4 +- Mathlib/GroupTheory/Submonoid/Operations.lean | 2 +- Mathlib/Init/Data/Bool/Lemmas.lean | 2 +- Mathlib/Init/Data/Nat/Lemmas.lean | 2 +- Mathlib/Init/Function.lean | 2 - Mathlib/Init/Order/Defs.lean | 18 +- Mathlib/Init/Quot.lean | 2 + Mathlib/Lean/Linter.lean | 24 -- .../AffineSpace/FiniteDimensional.lean | 3 +- .../AffineSpace/Independent.lean | 2 +- Mathlib/LinearAlgebra/Alternating/Basic.lean | 3 +- Mathlib/LinearAlgebra/Basic.lean | 2 - Mathlib/LinearAlgebra/Determinant.lean | 2 +- .../ExteriorAlgebra/Grading.lean | 2 +- Mathlib/LinearAlgebra/Finrank.lean | 2 +- Mathlib/LinearAlgebra/LinearPMap.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Polynomial.lean | 11 +- .../LinearAlgebra/Matrix/Transvection.lean | 3 +- Mathlib/Logic/Basic.lean | 3 +- Mathlib/Logic/Embedding/Basic.lean | 2 +- .../Constructions/BorelSpace/Basic.lean | 15 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 3 +- .../MeasureTheory/Covering/Besicovitch.lean | 6 +- .../Covering/BesicovitchVectorSpace.lean | 2 - .../Covering/Differentiation.lean | 22 +- .../MeasureTheory/Decomposition/Lebesgue.lean | 3 +- .../Function/ConvergenceInMeasure.lean | 4 +- Mathlib/MeasureTheory/Function/Jacobian.lean | 3 - Mathlib/MeasureTheory/Function/L1Space.lean | 2 +- Mathlib/MeasureTheory/Function/L2Space.lean | 6 +- .../MeasureTheory/Function/LpSeminorm.lean | 6 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 2 - .../MeasureTheory/Function/SimpleFunc.lean | 2 +- .../Function/SimpleFuncDense.lean | 2 +- .../Function/SimpleFuncDenseLp.lean | 4 +- .../Function/StronglyMeasurable/Basic.lean | 2 +- .../Function/UniformIntegrable.lean | 6 +- Mathlib/MeasureTheory/Group/Measure.lean | 16 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 10 +- .../Integral/CircleIntegral.lean | 10 +- .../Integral/CircleTransform.lean | 2 - .../Integral/MeanInequalities.lean | 2 - .../MeasureTheory/Integral/PeakFunction.lean | 3 - .../MeasureTheory/Integral/SetIntegral.lean | 4 +- .../MeasureTheory/Integral/TorusIntegral.lean | 5 +- .../MeasureTheory/MeasurableSpace/Basic.lean | 2 +- Mathlib/MeasureTheory/Measure/Doubling.lean | 3 - Mathlib/MeasureTheory/Measure/Haar/Basic.lean | 2 +- .../Measure/Haar/InnerProductSpace.lean | 7 +- .../Measure/Haar/NormedSpace.lean | 4 +- .../MeasureTheory/Measure/Lebesgue/Basic.lean | 2 - .../Measure/Lebesgue/Complex.lean | 2 - .../Measure/Lebesgue/EqHaar.lean | 2 - .../Measure/ProbabilityMeasure.lean | 2 +- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- Mathlib/NumberTheory/Basic.lean | 2 +- Mathlib/NumberTheory/Bernoulli.lean | 9 +- Mathlib/NumberTheory/Bertrand.lean | 20 +- Mathlib/NumberTheory/ClassNumber/Finite.lean | 3 +- .../NumberTheory/Cyclotomic/Discriminant.lean | 8 +- .../Cyclotomic/PrimitiveRoots.lean | 9 +- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 6 +- Mathlib/NumberTheory/Dioph.lean | 4 +- Mathlib/NumberTheory/Divisors.lean | 3 +- Mathlib/NumberTheory/FLT/Four.lean | 4 +- .../LegendreSymbol/GaussEisensteinLemmas.lean | 4 +- .../LegendreSymbol/JacobiSymbol.lean | 6 +- .../LegendreSymbol/QuadraticChar/Basic.lean | 2 +- .../QuadraticChar/GaussSum.lean | 6 +- .../LegendreSymbol/QuadraticReciprocity.lean | 4 +- .../NumberTheory/LegendreSymbol/ZModChar.lean | 10 +- .../Liouville/LiouvilleNumber.lean | 4 +- .../NumberTheory/Liouville/LiouvilleWith.lean | 2 - Mathlib/NumberTheory/Liouville/Measure.lean | 2 - Mathlib/NumberTheory/Modular.lean | 2 +- .../ModularForms/JacobiTheta/Basic.lean | 16 +- .../ModularForms/SlashActions.lean | 2 - Mathlib/NumberTheory/Multiplicity.lean | 10 +- .../NumberField/CanonicalEmbedding.lean | 4 +- .../NumberTheory/NumberField/Embeddings.lean | 4 +- Mathlib/NumberTheory/NumberField/Units.lean | 10 +- .../NumberTheory/Padics/PadicIntegers.lean | 3 +- Mathlib/NumberTheory/Padics/PadicVal.lean | 3 +- Mathlib/NumberTheory/Padics/RingHoms.lean | 8 +- Mathlib/NumberTheory/Pell.lean | 2 +- Mathlib/NumberTheory/PellMatiyasevic.lean | 4 +- Mathlib/NumberTheory/PythagoreanTriples.lean | 19 +- Mathlib/NumberTheory/Rayleigh.lean | 4 +- Mathlib/NumberTheory/SmoothNumbers.lean | 3 +- Mathlib/NumberTheory/SumFourSquares.lean | 4 +- Mathlib/NumberTheory/SumTwoSquares.lean | 2 +- Mathlib/NumberTheory/ZetaFunction.lean | 4 +- Mathlib/NumberTheory/ZetaValues.lean | 2 - .../Zsqrtd/QuadraticReciprocity.lean | 4 +- Mathlib/Order/Filter/AtTopBot.lean | 3 +- Mathlib/Order/Filter/ENNReal.lean | 2 +- Mathlib/Order/Filter/IndicatorFunction.lean | 2 +- Mathlib/Order/LocallyFinite.lean | 4 +- .../Probability/Distributions/Gaussian.lean | 4 +- Mathlib/Probability/Integration.lean | 4 +- Mathlib/Probability/Kernel/Composition.lean | 5 +- .../Probability/Martingale/BorelCantelli.lean | 4 +- .../Probability/Martingale/Convergence.lean | 2 +- .../Probability/Martingale/Upcrossing.lean | 7 +- Mathlib/Probability/Notation.lean | 3 +- Mathlib/Probability/Process/Stopping.lean | 2 +- Mathlib/Probability/StrongLaw.lean | 4 +- Mathlib/Probability/Variance.lean | 10 +- .../GroupCohomology/Resolution.lean | 3 +- .../HomogeneousLocalization.lean | 2 +- Mathlib/RingTheory/LocalProperties.lean | 2 - .../MvPolynomial/NewtonIdentities.lean | 6 +- .../Polynomial/Cyclotomic/Basic.lean | 2 +- Mathlib/RingTheory/Polynomial/Selmer.lean | 2 - Mathlib/RingTheory/RootsOfUnity/Basic.lean | 4 +- Mathlib/RingTheory/WittVector/Basic.lean | 6 +- .../WittVector/DiscreteValuationRing.lean | 2 - Mathlib/RingTheory/WittVector/Frobenius.lean | 2 +- Mathlib/RingTheory/WittVector/IsPoly.lean | 11 +- Mathlib/RingTheory/WittVector/MulCoeff.lean | 5 +- .../WittVector/StructurePolynomial.lean | 3 +- Mathlib/SetTheory/Cardinal/Basic.lean | 30 +-- Mathlib/SetTheory/Cardinal/Divisibility.lean | 3 +- Mathlib/SetTheory/Cardinal/Finite.lean | 2 +- Mathlib/SetTheory/Lists.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 6 +- Mathlib/SetTheory/Ordinal/Basic.lean | 3 +- .../SetTheory/Ordinal/CantorNormalForm.lean | 6 +- Mathlib/SetTheory/Ordinal/Exponential.lean | 28 +-- Mathlib/SetTheory/Ordinal/Notation.lean | 33 +-- Mathlib/SetTheory/Ordinal/Principal.lean | 4 +- Mathlib/Tactic/ComputeDegree.lean | 6 +- Mathlib/Tactic/DeriveTraversable.lean | 4 +- Mathlib/Tactic/LibrarySearch.lean | 2 +- Mathlib/Tactic/NormNum/LegendreSymbol.lean | 1 + Mathlib/Tactic/Rewrites.lean | 2 +- Mathlib/Tactic/Simps/Basic.lean | 2 - Mathlib/Tactic/ToAdditive.lean | 1 - Mathlib/Tactic/UnsetOption.lean | 1 - .../Topology/Algebra/InfiniteSum/Basic.lean | 5 +- Mathlib/Topology/Algebra/Module/Basic.lean | 4 +- .../Topology/Algebra/Module/Cardinality.lean | 4 +- Mathlib/Topology/Algebra/Monoid.lean | 11 +- .../Algebra/Nonarchimedean/AdicTopology.lean | 4 +- Mathlib/Topology/Algebra/ValuedField.lean | 3 +- Mathlib/Topology/Basic.lean | 2 +- .../Topology/Category/Profinite/Nobeling.lean | 16 +- .../Category/TopCat/Limits/Pullbacks.lean | 2 +- Mathlib/Topology/Constructions.lean | 5 +- .../Topology/ContinuousFunction/Basic.lean | 3 +- .../Topology/EMetricSpace/Paracompact.lean | 2 +- Mathlib/Topology/Instances/ENNReal.lean | 8 +- .../Topology/MetricSpace/GromovHausdorff.lean | 2 - .../MetricSpace/HausdorffDimension.lean | 2 +- .../MetricSpace/ThickenedIndicator.lean | 2 +- Mathlib/Topology/Sequences.lean | 3 +- Mathlib/Topology/Sheaves/Presheaf.lean | 2 +- .../Sheaves/SheafCondition/UniqueGluing.lean | 2 +- Mathlib/Topology/UniformSpace/Pi.lean | 3 +- Mathlib/Topology/VectorBundle/Basic.lean | 2 +- lake-manifest.json | 92 ++++---- lakefile.lean | 4 +- lean-toolchain | 2 +- test/Clean.lean | 12 +- test/RewriteSearch/Basic.lean | 2 +- test/hint.lean | 2 +- test/linear_combination.lean | 2 +- test/matrix.lean | 15 +- test/norm_num_ext.lean | 12 +- test/positivity.lean | 2 +- 446 files changed, 1136 insertions(+), 1253 deletions(-) delete mode 100644 Mathlib/Data/KVMap.lean delete mode 100644 Mathlib/Lean/Linter.lean diff --git a/.gitignore b/.gitignore index 6c4a2e683c64f..9f41205d619e2 100644 --- a/.gitignore +++ b/.gitignore @@ -2,10 +2,10 @@ .DS_Store # Created by `lake exe cache get` if no home directory is available /.cache -# Prior to https://github.com/leanprover/lean4/pull/2749 lake stored files in these locations. +# Prior to v4.3.0-rc2 lake stored files in these locations. # We'll leave them in the `.gitignore` for a while for users switching between toolchains. /build/ /lake-packages/ /lakefile.olean -# After https://github.com/leanprover/lean4/pull/2749 lake stores its files here: +# After v4.3.0-rc2 lake stores its files here: /.lake/ diff --git a/Archive/Examples/IfNormalization/Result.lean b/Archive/Examples/IfNormalization/Result.lean index 79a203404983f..73f3849f2967b 100644 --- a/Archive/Examples/IfNormalization/Result.lean +++ b/Archive/Examples/IfNormalization/Result.lean @@ -29,20 +29,13 @@ attribute [local simp] normalized hasNestedIf hasConstantIf hasRedundantIf disjo /-! Adding these lemmas to the simp set allows Lean to handle the termination proof automatically. -/ -attribute [local simp] Nat.lt_add_one_iff le_add_of_le_right +attribute [local simp] Nat.lt_add_one_iff le_add_of_le_right max_add_add_right max_mul_mul_left /-! Some further simp lemmas for handling if-then-else statements. -/ attribute [local simp] apply_ite ite_eq_iff' --- A copy of Lean's `decide_eq_true_eq` which unifies the `Decidable` instance --- rather than finding it by typeclass search. --- See https://github.com/leanprover/lean4/pull/2816 -@[simp] theorem decide_eq_true_eq {i : Decidable p} : (@decide p i = true) = p := - _root_.decide_eq_true_eq - - /-! Simp lemmas for `eval`. We don't want a `simp` lemma for `(ite i t e).eval` in general, only once we know the shape of `i`. diff --git a/Archive/Examples/IfNormalization/WithoutAesop.lean b/Archive/Examples/IfNormalization/WithoutAesop.lean index 254ac3b982785..57f9cf2905f71 100644 --- a/Archive/Examples/IfNormalization/WithoutAesop.lean +++ b/Archive/Examples/IfNormalization/WithoutAesop.lean @@ -22,12 +22,6 @@ namespace IfExpr attribute [local simp] eval normalized hasNestedIf hasConstantIf hasRedundantIf disjoint vars List.disjoint max_add_add_right max_mul_mul_left Nat.lt_add_one_iff le_add_of_le_right --- A copy of Lean's `decide_eq_true_eq` which unifies the `Decidable` instance --- rather than finding it by typeclass search. --- See https://github.com/leanprover/lean4/pull/2816 -@[simp] theorem decide_eq_true_eq' {i : Decidable p} : (@decide p i = true) = p := - _root_.decide_eq_true_eq - theorem eval_ite_ite' : (ite (ite a b c) d e).eval f = (ite a (ite b d e) (ite c d e)).eval f := by cases h : eval f a <;> simp_all @@ -106,8 +100,8 @@ def normalize' (l : AList (fun _ : ℕ => Bool)) : simp_all · simp_all? says simp_all only [ne_eq, hasNestedIf, Bool.or_self, hasConstantIf, and_self, hasRedundantIf, Bool.or_false, beq_eq_false_iff_ne, not_false_eq_true, - disjoint, List.disjoint, Bool.and_true, Bool.and_eq_true, decide_eq_true_eq', - true_and] + disjoint, List.disjoint, decide_not, Bool.and_true, Bool.and_eq_true, + Bool.not_eq_true', decide_eq_false_iff_not, true_and] constructor <;> assumption · have := ht₃ w have := he₃ w diff --git a/Archive/Imo/Imo1960Q1.lean b/Archive/Imo/Imo1960Q1.lean index 5b6a4a08aea0e..d4e98c61b6659 100644 --- a/Archive/Imo/Imo1960Q1.lean +++ b/Archive/Imo/Imo1960Q1.lean @@ -95,7 +95,12 @@ theorem right_direction {n : ℕ} : ProblemPredicate n → SolutionPredicate n : have := searchUpTo_start iterate 82 replace := - searchUpTo_step this (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num) + searchUpTo_step this (by norm_num1; rfl) (by norm_num1; rfl) (by norm_num1; rfl) + (by -- This used to be just `norm_num`, but after leanprover/lean4#2790, + -- that triggers a max recursion depth exception. As a workaround, we + -- manually rewrite with `Nat.digits_of_two_le_of_pos` first. + rw [Nat.digits_of_two_le_of_pos (by norm_num) (by norm_num)] + norm_num <;> decide) exact searchUpTo_end this #align imo1960_q1.right_direction Imo1960Q1.right_direction diff --git a/Archive/Imo/Imo1962Q1.lean b/Archive/Imo/Imo1962Q1.lean index 951ab2d688130..97245b3333242 100644 --- a/Archive/Imo/Imo1962Q1.lean +++ b/Archive/Imo/Imo1962Q1.lean @@ -131,7 +131,18 @@ Now we combine these cases to show that 153846 is the smallest solution. theorem satisfied_by_153846 : ProblemPredicate 153846 := by - norm_num [ProblemPredicate] + -- This proof used to be the single line `norm_num [ProblemPredicate]`. + -- After leanprover/lean4#2790, that triggers a max recursion depth exception. + -- As a workaround, we manually apply `Nat.digits_of_two_le_of_pos` a few times + -- before invoking `norm_num`. + unfold ProblemPredicate + have two_le_ten : 2 ≤ 10 := by norm_num + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + rw [Nat.digits_of_two_le_of_pos two_le_ten (by norm_num)] + norm_num + decide #align imo1962_q1.satisfied_by_153846 Imo1962Q1.satisfied_by_153846 theorem no_smaller_solutions (n : ℕ) (h1 : ProblemPredicate n) : n ≥ 153846 := by diff --git a/Archive/Imo/Imo1962Q4.lean b/Archive/Imo/Imo1962Q4.lean index 4789a28129eca..575c66f1a28a7 100644 --- a/Archive/Imo/Imo1962Q4.lean +++ b/Archive/Imo/Imo1962Q4.lean @@ -17,8 +17,6 @@ Since Lean does not have a concept of "simplest form", we just express what is in fact the simplest form of the set of solutions, and then prove it equals the set of solutions. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Real open scoped Real diff --git a/Archive/Imo/Imo1964Q1.lean b/Archive/Imo/Imo1964Q1.lean index 0811661468c43..d0376117a076c 100644 --- a/Archive/Imo/Imo1964Q1.lean +++ b/Archive/Imo/Imo1964Q1.lean @@ -34,7 +34,7 @@ theorem two_pow_mod_seven (n : ℕ) : 2 ^ n ≡ 2 ^ (n % 3) [MOD 7] := let t := n % 3 calc 2 ^ n = 2 ^ (3 * (n / 3) + t) := by rw [Nat.div_add_mod] _ = (2 ^ 3) ^ (n / 3) * 2 ^ t := by rw [pow_add, pow_mul] - _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; norm_num + _ ≡ 1 ^ (n / 3) * 2 ^ t [MOD 7] := by gcongr; decide _ = 2 ^ t := by ring /-! @@ -68,5 +68,5 @@ theorem imo1964_q1b (n : ℕ) : ¬7 ∣ 2 ^ n + 1 := by have H : 2 ^ t + 1 ≡ 0 [MOD 7] · calc 2 ^ t + 1 ≡ 2 ^ n + 1 [MOD 7 ] := by gcongr ?_ + 1; exact (two_pow_mod_seven n).symm _ ≡ 0 [MOD 7] := h.modEq_zero_nat - interval_cases t <;> norm_num at H + interval_cases t <;> contradiction #align imo1964_q1b imo1964_q1b diff --git a/Archive/Imo/Imo1981Q3.lean b/Archive/Imo/Imo1981Q3.lean index 5b0765b4a9cf9..1a136a674c3ed 100644 --- a/Archive/Imo/Imo1981Q3.lean +++ b/Archive/Imo/Imo1981Q3.lean @@ -129,7 +129,7 @@ theorem imp_fib {n : ℕ} : ∀ m : ℕ, NatPredicate N m n → ∃ k : ℕ, m = obtain (rfl : 1 = n) | (h4 : 1 < n) := (succ_le_iff.mpr h2.n_pos).eq_or_lt · use 1 have h5 : 1 ≤ m := succ_le_iff.mpr h2.m_pos - simpa [fib_one, fib_two] using (h3.antisymm h5 : m = 1) + simpa [fib_one, fib_two, (by decide : 1 + 1 = 2)] using (h3.antisymm h5 : m = 1) · obtain (rfl : m = n) | (h6 : m < n) := h3.eq_or_lt · exact absurd h2.eq_imp_1 (Nat.ne_of_gt h4) · have h7 : NatPredicate N (n - m) m := h2.reduction h4 @@ -205,5 +205,7 @@ theorem imo1981_q3 : IsGreatest (specifiedSet 1981) 3524578 := by have := fun h => @solution_greatest 1981 16 h 3524578 norm_num at this apply this - norm_num [ProblemPredicate_iff] + · decide + · decide + · norm_num [ProblemPredicate_iff]; decide #align imo1981_q3 imo1981_q3 diff --git a/Archive/Imo/Imo2001Q2.lean b/Archive/Imo/Imo2001Q2.lean index 1b8045c2ba72a..3e8528f7ef2ef 100644 --- a/Archive/Imo/Imo2001Q2.lean +++ b/Archive/Imo/Imo2001Q2.lean @@ -32,8 +32,6 @@ open Real variable {a b c : ℝ} -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Imo2001Q2 theorem bound (ha : 0 < a) (hb : 0 < b) (hc : 0 < c) : diff --git a/Archive/Imo/Imo2005Q4.lean b/Archive/Imo/Imo2005Q4.lean index 33ff8b28afaed..808daebbb4698 100644 --- a/Archive/Imo/Imo2005Q4.lean +++ b/Archive/Imo/Imo2005Q4.lean @@ -29,7 +29,7 @@ def a (n : ℕ) : ℤ := theorem find_specified_factor {p : ℕ} (hp : Nat.Prime p) (hp2 : p ≠ 2) (hp3 : p ≠ 3) : ↑p ∣ a (p - 2) := by -- Since `p` is neither `2` nor `3`, it is coprime with `2`, `3`, and `6` - rw [Ne.def, ← Nat.prime_dvd_prime_iff_eq hp (by norm_num), ← Nat.Prime.coprime_iff_not_dvd hp] + rw [Ne.def, ← Nat.prime_dvd_prime_iff_eq hp (by decide), ← Nat.Prime.coprime_iff_not_dvd hp] at hp2 hp3 have : Int.gcd p 6 = 1 := Nat.coprime_mul_iff_right.2 ⟨hp2, hp3⟩ -- Nat arithmetic needed to deal with powers @@ -71,10 +71,10 @@ theorem imo2005_q4 {k : ℕ} (hk : 0 < k) : (∀ n : ℕ, 1 ≤ n → IsCoprime -- For `p = 2` and `p = 3`, take `n = 1` and `n = 2`, respectively by_cases hp2 : p = 2 · rw [hp2] at h - apply h 1 <;> norm_num + apply h 1 <;> decide by_cases hp3 : p = 3 · rw [hp3] at h - apply h 2 <;> norm_num + apply h 2 <;> decide -- Otherwise, take `n = p - 2` refine h (p - 2) ?_ (find_specified_factor hp hp2 hp3) calc diff --git a/Archive/Imo/Imo2008Q3.lean b/Archive/Imo/Imo2008Q3.lean index 607f73a355f72..32d6e8437a59b 100644 --- a/Archive/Imo/Imo2008Q3.lean +++ b/Archive/Imo/Imo2008Q3.lean @@ -31,8 +31,6 @@ Then `p = 2n + k ≥ 2n + √(p - 4) = 2n + √(2n + k - 4) > √(2n)` and we ar open Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Imo2008Q3 theorem p_lemma (p : ℕ) (hpp : Nat.Prime p) (hp_mod_4_eq_1 : p ≡ 1 [MOD 4]) (hp_gt_20 : p > 20) : diff --git a/Archive/Imo/Imo2013Q5.lean b/Archive/Imo/Imo2013Q5.lean index fa07e8ed305a3..635e12eeec70e 100644 --- a/Archive/Imo/Imo2013Q5.lean +++ b/Archive/Imo/Imo2013Q5.lean @@ -31,8 +31,6 @@ https://www.imo-official.org/problems/IMO2013SL.pdf open scoped BigOperators -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Imo2013Q5 theorem le_of_all_pow_lt_succ {x y : ℝ} (hx : 1 < x) (hy : 1 < y) diff --git a/Archive/Imo/Imo2019Q1.lean b/Archive/Imo/Imo2019Q1.lean index 6e486889066d0..2bcd5a9f405b9 100644 --- a/Archive/Imo/Imo2019Q1.lean +++ b/Archive/Imo/Imo2019Q1.lean @@ -27,7 +27,7 @@ theorem imo2019_q1 (f : ℤ → ℤ) : (∀ a b : ℤ, f (2 * a) + 2 * f b = f (f (a + b))) ↔ f = 0 ∨ ∃ c, f = fun x => 2 * x + c := by constructor; swap -- easy way: f(x)=0 and f(x)=2x+c work. - · rintro (rfl | ⟨c, rfl⟩) <;> intros <;> simp only [Pi.zero_apply]; ring + · rintro (rfl | ⟨c, rfl⟩) <;> intros <;> norm_num; ring -- hard way. intro hf -- functional equation diff --git a/Archive/Imo/Imo2019Q2.lean b/Archive/Imo/Imo2019Q2.lean index c301c615ed3ad..3c2131796191b 100644 --- a/Archive/Imo/Imo2019Q2.lean +++ b/Archive/Imo/Imo2019Q2.lean @@ -442,7 +442,7 @@ theorem not_collinear_QPA₂ : ¬Collinear ℝ ({cfg.Q, cfg.P, cfg.A₂} : Set P have h : Cospherical ({cfg.B, cfg.A, cfg.A₂} : Set Pt) := by refine' cfg.triangleABC.circumsphere.cospherical.subset _ simp only [Set.insert_subset_iff, cfg.A_mem_circumsphere, cfg.B_mem_circumsphere, - cfg.A₂_mem_circumsphere, Sphere.mem_coe, Set.singleton_subset_iff] + cfg.A₂_mem_circumsphere, Sphere.mem_coe, Set.singleton_subset_iff, and_true] exact h.affineIndependent_of_ne cfg.A_ne_B.symm cfg.A₂_ne_B.symm cfg.A₂_ne_A.symm #align imo2019_q2.imo2019q2_cfg.not_collinear_QPA₂ Imo2019Q2.Imo2019q2Cfg.not_collinear_QPA₂ diff --git a/Archive/Imo/Imo2019Q4.lean b/Archive/Imo/Imo2019Q4.lean index b52f6085c03ce..fd3367ae96956 100644 --- a/Archive/Imo/Imo2019Q4.lean +++ b/Archive/Imo/Imo2019Q4.lean @@ -29,8 +29,6 @@ individually. open scoped Nat BigOperators -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Nat hiding zero_le Prime open Finset multiplicity @@ -65,7 +63,7 @@ theorem upper_bound {k n : ℕ} (hk : k > 0) _ ≤ k ! := by gcongr clear h h2 induction' n, hn using Nat.le_induction with n' hn' IH - · norm_num + · decide let A := ∑ i in range n', i have le_sum : ∑ i in range 6, i ≤ A · apply sum_le_sum_of_subset @@ -87,8 +85,7 @@ theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) : -- The implication `←` holds. constructor swap - · rintro (h | h) <;> simp [Prod.ext_iff] at h <;> rcases h with ⟨rfl, rfl⟩ <;> - norm_num [prod_range_succ, succ_mul] + · rintro (h | h) <;> simp [Prod.ext_iff] at h <;> rcases h with ⟨rfl, rfl⟩ <;> decide intro h -- We know that n < 6. have := Imo2019Q4.upper_bound hk h @@ -101,9 +98,9 @@ theorem imo2019_q4 {k n : ℕ} (hk : k > 0) (hn : n > 0) : exact h; rw [h]; norm_num all_goals exfalso; norm_num [prod_range_succ] at h; norm_cast at h -- n = 3 - · refine' monotone_factorial.ne_of_lt_of_lt_nat 5 _ _ _ h <;> norm_num + · refine' monotone_factorial.ne_of_lt_of_lt_nat 5 _ _ _ h <;> decide -- n = 4 - · refine' monotone_factorial.ne_of_lt_of_lt_nat 7 _ _ _ h <;> norm_num + · refine' monotone_factorial.ne_of_lt_of_lt_nat 7 _ _ _ h <;> decide -- n = 5 - · refine' monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h <;> norm_num + · refine' monotone_factorial.ne_of_lt_of_lt_nat 10 _ _ _ h <;> decide #align imo2019_q4 imo2019_q4 diff --git a/Archive/MiuLanguage/DecisionNec.lean b/Archive/MiuLanguage/DecisionNec.lean index 89f01eb056d6e..ac22effdd4b50 100644 --- a/Archive/MiuLanguage/DecisionNec.lean +++ b/Archive/MiuLanguage/DecisionNec.lean @@ -61,7 +61,7 @@ theorem mod3_eq_1_or_mod3_eq_2 {a b : ℕ} (h1 : a % 3 = 1 ∨ a % 3 = 2) · rw [h2]; exact h1 · cases' h1 with h1 h1 · right; simp [h2, mul_mod, h1, Nat.succ_lt_succ] - · left; simp [h2, mul_mod, h1] + · left; simp only [h2, mul_mod, h1, mod_mod]; decide #align miu.mod3_eq_1_or_mod3_eq_2 Miu.mod3_eq_1_or_mod3_eq_2 /-- `count_equiv_one_or_two_mod3_of_derivable` shows any derivable string must have a `count I` that @@ -80,7 +80,7 @@ theorem count_equiv_one_or_two_mod3_of_derivable (en : Miustr) : simp_rw [count_cons_self, count_nil, count_cons, ite_false, add_right_comm, add_mod_right] simp · left; rw [count_append, count_append, count_append] - simp only [ne_eq, count_cons_of_ne, count_nil, add_zero] + simp only [ne_eq, not_false_eq_true, count_cons_of_ne, count_nil, add_zero] #align miu.count_equiv_one_or_two_mod3_of_derivable Miu.count_equiv_one_or_two_mod3_of_derivable /-- Using the above theorem, we solve the MU puzzle, showing that `"MU"` is not derivable. @@ -134,7 +134,7 @@ theorem goodm_of_rule1 (xs : Miustr) (h₁ : Derivable (xs ++ ↑[I])) (h₂ : G exact mhead · change ¬M ∈ tail (xs ++ ↑([I] ++ [U])) rw [← append_assoc, tail_append_singleton_of_ne_nil] - · simp_rw [mem_append, not_or, and_true]; exact nmtail + · simp_rw [mem_append, mem_singleton, or_false]; exact nmtail · exact append_ne_nil_of_ne_nil_left _ _ this #align miu.goodm_of_rule1 Miu.goodm_of_rule1 diff --git a/Archive/MiuLanguage/DecisionSuf.lean b/Archive/MiuLanguage/DecisionSuf.lean index 7f9fdaba0ac4a..6c7ff233e8c95 100644 --- a/Archive/MiuLanguage/DecisionSuf.lean +++ b/Archive/MiuLanguage/DecisionSuf.lean @@ -267,7 +267,7 @@ theorem count_I_eq_length_of_count_U_zero_and_neg_mem {ys : Miustr} (hu : count · simpa only [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] at hu + 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 @@ -277,7 +277,8 @@ set_option linter.uppercaseLean3 false in theorem base_case_suf (en : Miustr) (h : Decstr en) (hu : count U en = 0) : Derivable en := by rcases h with ⟨⟨mhead, nmtail⟩, hi⟩ have : en ≠ nil := by - intro k; simp only [k, count, countP, if_false, zero_mod, zero_ne_one, false_or_iff] at hi + intro k + simp only [k, count, countP, countP.go, if_false, zero_mod, zero_ne_one, false_or_iff] at hi rcases exists_cons_of_ne_nil this with ⟨y, ys, rfl⟩ rcases mhead rsuffices ⟨c, rfl, hc⟩ : ∃ c, replicate c I = ys ∧ (c % 3 = 1 ∨ c % 3 = 2) @@ -331,7 +332,7 @@ theorem ind_hyp_suf (k : ℕ) (ys : Miustr) (hu : count U ys = succ k) (hdec : D rw [cons_append, cons_append] dsimp [tail] at nmtail ⊢ rw [mem_append] at nmtail - simpa only [mem_append, mem_cons, false_or_iff, or_false_iff] using nmtail + simpa only [append_assoc, cons_append, nil_append, mem_append, mem_cons, false_or] using nmtail · rw [count_append, count_append]; rw [← cons_append, count_append] at hic simp only [count_cons_self, count_nil, count_cons, if_false] at hic ⊢ rw [add_right_comm, add_mod_right]; exact hic diff --git a/Archive/Sensitivity.lean b/Archive/Sensitivity.lean index 92320d5089a0e..cf65eedccc94f 100644 --- a/Archive/Sensitivity.lean +++ b/Archive/Sensitivity.lean @@ -401,7 +401,7 @@ theorem exists_eigenvalue (H : Set (Q m.succ)) (hH : Card H ≥ 2 ^ m + 1) : let img := range (g m) suffices 0 < dim (W ⊓ img) by exact_mod_cast exists_mem_ne_zero_of_rank_pos this - have dim_le : dim (W ⊔ img) ≤ 2 ^ (m + 1) := by + have dim_le : dim (W ⊔ img) ≤ 2 ^ (m + 1 : Cardinal) := by convert ← rank_submodule_le (W ⊔ img) rw [← Nat.cast_succ] apply dim_V diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index 72eb193a353d3..942f926b4464f 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -31,8 +31,6 @@ namespace AbelRuffini set_option linter.uppercaseLean3 false -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Function Polynomial Polynomial.Gal Ideal open scoped Polynomial @@ -94,10 +92,11 @@ theorem irreducible_Phi (p : ℕ) (hp : p.Prime) (hpa : p ∣ a) (hpb : p ∣ b) rw [mem_span_singleton] rw [degree_Phi] at hn; norm_cast at hn interval_cases hn : n <;> - simp only [Φ, coeff_X_pow, coeff_C, Int.coe_nat_dvd.mpr, hpb, if_true, coeff_C_mul, if_false, - coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, coeff_sub, add_zero, zero_sub, dvd_neg, - neg_zero, dvd_mul_of_dvd_left] + simp (config := {decide := true}) only [Φ, coeff_X_pow, coeff_C, Int.coe_nat_dvd.mpr, hpb, + if_true, coeff_C_mul, if_false, coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, coeff_sub, + add_zero, zero_sub, dvd_neg, neg_zero, dvd_mul_of_dvd_left] · simp only [degree_Phi, ← WithBot.coe_zero, WithBot.coe_lt_coe, Nat.succ_pos'] + decide · rw [coeff_zero_Phi, span_singleton_pow, mem_span_singleton] exact mt Int.coe_nat_dvd.mp hp2b all_goals exact Monic.isPrimitive (monic_Phi a b) @@ -136,7 +135,7 @@ theorem real_roots_Phi_ge_aux (hab : b < a) : have hfa := calc f (-a) = (a : ℝ) ^ 2 - (a : ℝ) ^ 5 + b := by - norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow] + norm_num [hf, ← sq, sub_eq_add_neg, add_comm, Odd.neg_pow (by decide : Odd 5)] _ ≤ (a : ℝ) ^ 2 - (a : ℝ) ^ 3 + (a - 1) := by refine' add_le_add (sub_le_sub_left (pow_le_pow ha _) _) _ <;> linarith _ = -((a : ℝ) - 1) ^ 2 * (a + 1) := by ring @@ -163,7 +162,7 @@ theorem complex_roots_Phi (h : (Φ ℚ a b).Separable) : Fintype.card ((Φ ℚ a theorem gal_Phi (hab : b < a) (h_irred : Irreducible (Φ ℚ a b)) : Bijective (galActionHom (Φ ℚ a b) ℂ) := by apply galActionHom_bijective_of_prime_degree' h_irred - · norm_num [natDegree_Phi] + · simp only [natDegree_Phi]; decide · rw [complex_roots_Phi a b h_irred.separable, Nat.succ_le_succ_iff] exact (real_roots_Phi_le a b).trans (Nat.le_succ 3) · simp_rw [complex_roots_Phi a b h_irred.separable, Nat.succ_le_succ_iff] @@ -181,7 +180,7 @@ theorem not_solvable_by_rad (p : ℕ) (x : ℂ) (hx : aeval x (Φ ℚ a b) = 0) #align abel_ruffini.not_solvable_by_rad AbelRuffini.not_solvable_by_rad theorem not_solvable_by_rad' (x : ℂ) (hx : aeval x (Φ ℚ 4 2) = 0) : ¬IsSolvableByRad ℚ x := by - apply not_solvable_by_rad 4 2 2 x hx <;> norm_num + apply not_solvable_by_rad 4 2 2 x hx <;> decide #align abel_ruffini.not_solvable_by_rad' AbelRuffini.not_solvable_by_rad' /-- **Abel-Ruffini Theorem** -/ diff --git a/Archive/Wiedijk100Theorems/AreaOfACircle.lean b/Archive/Wiedijk100Theorems/AreaOfACircle.lean index c98ccfc1e0236..e56ce51c6e24e 100644 --- a/Archive/Wiedijk100Theorems/AreaOfACircle.lean +++ b/Archive/Wiedijk100Theorems/AreaOfACircle.lean @@ -43,9 +43,6 @@ to the n-ball. -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Set Real MeasureTheory intervalIntegral open scoped Real NNReal diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 84660e6e21d11..f7a5a95024821 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -221,7 +221,9 @@ theorem first_vote_pos : simp [ENNReal.div_self _ _] | 0, q + 1, _ => by rw [counted_left_zero, condCount_singleton] - simp + simp only [List.replicate, Nat.add_eq, add_zero, mem_setOf_eq, List.headI_cons, Nat.cast_zero, + ENNReal.zero_div, ite_eq_right_iff] + decide | p + 1, q + 1, _ => by simp_rw [counted_succ_succ] rw [← condCount_disjoint_union ((countedSequence_finite _ _).image _) diff --git a/Archive/Wiedijk100Theorems/BirthdayProblem.lean b/Archive/Wiedijk100Theorems/BirthdayProblem.lean index 4c3224823ebb2..aadb154d8066e 100644 --- a/Archive/Wiedijk100Theorems/BirthdayProblem.lean +++ b/Archive/Wiedijk100Theorems/BirthdayProblem.lean @@ -29,7 +29,13 @@ local notation "‖" x "‖" => Fintype.card x /-- **Birthday Problem**: set cardinality interpretation. -/ theorem birthday : 2 * ‖Fin 23 ↪ Fin 365‖ < ‖Fin 23 → Fin 365‖ ∧ 2 * ‖Fin 22 ↪ Fin 365‖ > ‖Fin 22 → Fin 365‖ := by - simp only [Nat.descFactorial, Fintype.card_fin, Fintype.card_embedding_eq, Fintype.card_fun] + -- This used to be + -- `simp only [Nat.descFactorial, Fintype.card_fin, Fintype.card_embedding_eq, Fintype.card_fun]` + -- but after leanprover/lean4#2790 that triggers a max recursion depth exception. + -- As a workaround, we make some of the reduction steps more explicit. + rw [Fintype.card_embedding_eq, Fintype.card_fun, Fintype.card_fin, Fintype.card_fin] + rw [Fintype.card_embedding_eq, Fintype.card_fun, Fintype.card_fin, Fintype.card_fin] + decide #align theorems_100.birthday Theorems100.birthday section MeasureTheory @@ -72,12 +78,14 @@ theorem birthday_measure : generalize_proofs hfin have : |hfin.toFinset| = 42200819302092359872395663074908957253749760700776448000000 := by trans ‖Fin 23 ↪ Fin 365‖ - · simp_rw [← Fintype.card_coe, Set.Finite.coeSort_toFinset, Set.coe_setOf] - exact Fintype.card_congr (Equiv.subtypeInjectiveEquivEmbedding _ _) - · simp only [Fintype.card_embedding_eq, Fintype.card_fin, Nat.descFactorial] + · rw [← Fintype.card_coe] + apply Fintype.card_congr + rw [Set.Finite.coeSort_toFinset, Set.coe_setOf] + exact Equiv.subtypeInjectiveEquivEmbedding _ _ + · rw [Fintype.card_embedding_eq, Fintype.card_fin, Fintype.card_fin] + rfl rw [this, ENNReal.lt_div_iff_mul_lt, mul_comm, mul_div, ENNReal.div_lt_iff] - rotate_left; (iterate 2 right; norm_num); (iterate 2 left; norm_num) - norm_cast + rotate_left; (iterate 2 right; norm_num); decide; (iterate 2 left; norm_num) simp only [Fintype.card_pi] norm_num #align theorems_100.birthday_measure Theorems100.birthday_measure diff --git a/Archive/Wiedijk100Theorems/HeronsFormula.lean b/Archive/Wiedijk100Theorems/HeronsFormula.lean index 1bf4ced3eb470..81760996b9c26 100644 --- a/Archive/Wiedijk100Theorems/HeronsFormula.lean +++ b/Archive/Wiedijk100Theorems/HeronsFormula.lean @@ -25,8 +25,6 @@ open Real EuclideanGeometry open scoped Real EuclideanGeometry -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Theorems100 local notation "√" => Real.sqrt diff --git a/Archive/Wiedijk100Theorems/Konigsberg.lean b/Archive/Wiedijk100Theorems/Konigsberg.lean index 816ba1994e37b..a277213885050 100644 --- a/Archive/Wiedijk100Theorems/Konigsberg.lean +++ b/Archive/Wiedijk100Theorems/Konigsberg.lean @@ -75,7 +75,7 @@ lemma degree_eq_degree (v : Verts) : graph.degree v = degree v := by cases v <;> #align konigsberg.degree_eq_degree Konigsberg.degree_eq_degree lemma not_even_degree_iff (w : Verts) : ¬Even (degree w) ↔ w = V1 ∨ w = V2 ∨ w = V3 ∨ w = V4 := by - cases w <;> simp + cases w <;> decide lemma setOf_odd_degree_eq : {v | Odd (graph.degree v)} = {Verts.V1, Verts.V2, Verts.V3, Verts.V4} := by diff --git a/Archive/Wiedijk100Theorems/PerfectNumbers.lean b/Archive/Wiedijk100Theorems/PerfectNumbers.lean index ff1f483234934..c610c03973978 100644 --- a/Archive/Wiedijk100Theorems/PerfectNumbers.lean +++ b/Archive/Wiedijk100Theorems/PerfectNumbers.lean @@ -39,8 +39,8 @@ namespace Nat open Nat.ArithmeticFunction Finset theorem sigma_two_pow_eq_mersenne_succ (k : ℕ) : σ 1 (2 ^ k) = mersenne (k + 1) := by - simp [sigma_one_apply, mersenne, Nat.prime_two, show 2 = 1 + 1 from rfl, - ← geom_sum_mul_add 1 (k + 1)] + simp_rw [sigma_one_apply, mersenne, show 2 = 1 + 1 from rfl, ← geom_sum_mul_add 1 (k + 1)] + norm_num #align theorems_100.nat.sigma_two_pow_eq_mersenne_succ Theorems100.Nat.sigma_two_pow_eq_mersenne_succ /-- Euclid's theorem that Mersenne primes induce perfect numbers -/ diff --git a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean index 05b6a23992908..e91db756e5796 100644 --- a/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean +++ b/Archive/Wiedijk100Theorems/SumOfPrimeReciprocalsDiverges.lean @@ -240,7 +240,7 @@ theorem Real.tendsto_sum_one_div_prime_atTop : have h4 := calc (card M' : ℝ) ≤ 2 ^ k * x.sqrt := by exact_mod_cast card_le_two_pow_mul_sqrt - _ = 2 ^ k * ↑(2 ^ (k + 1)) := by rw [Nat.sqrt_eq] + _ = 2 ^ k * (2 ^ (k + 1) : ℕ) := by rw [Nat.sqrt_eq] _ = x / 2 := by field_simp [mul_right_comm, ← pow_succ'] refine' lt_irrefl (x : ℝ) _ calc diff --git a/Archive/ZagierTwoSquares.lean b/Archive/ZagierTwoSquares.lean index 686263c38918f..2b91f80fcb9e8 100644 --- a/Archive/ZagierTwoSquares.lean +++ b/Archive/ZagierTwoSquares.lean @@ -48,9 +48,10 @@ lemma zagierSet_lower_bound {x y z : ℕ} (h : (x, y, z) ∈ zagierSet k) : 0 < all_goals cases' (Nat.dvd_prime hk.out).1 (dvd_of_mul_left_eq _ h) with e e all_goals - simp only [e, self_eq_add_left, ne_eq, add_eq_zero, and_false, mul_eq_left₀] at h + simp only [e, self_eq_add_left, ne_eq, add_eq_zero, and_false, not_false_eq_true, + mul_eq_left₀] at h simp only [h, zero_add] at hk - exact hk.out + exact Nat.not_prime_one hk.out lemma zagierSet_upper_bound {x y z : ℕ} (h : (x, y, z) ∈ zagierSet k) : x ≤ k + 1 ∧ y ≤ k ∧ z ≤ k := by @@ -143,7 +144,8 @@ theorem eq_of_mem_fixedPoints {t : zagierSet k} (mem : t ∈ fixedPoints (comple rw [mem_fixedPoints_iff, complexInvo, Subtype.mk.injEq] at mem split_ifs at mem with less more <;> -- less (completely handled by the pre-applied `simp_all only`) - simp_all only [not_lt, Prod.mk.injEq, add_right_eq_self, mul_eq_zero, false_or] + simp_all only [not_lt, Prod.mk.injEq, add_right_eq_self, mul_eq_zero, false_or, + lt_self_iff_false] · -- more obtain ⟨_, _, _⟩ := mem; simp_all · -- middle (the one fixed point falls under this case) @@ -194,3 +196,4 @@ theorem Nat.Prime.sq_add_sq' {p : ℕ} [h : Fact p.Prime] (hp : p % 4 = 1) : contrapose key rw [Set.not_nonempty_iff_eq_empty] at key simp_rw [key, Fintype.card_of_isEmpty, card_fixedPoints_eq_one] + decide diff --git a/Cache/IO.lean b/Cache/IO.lean index 9687c497759b0..37127ba2903cf 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -35,11 +35,11 @@ open System (FilePath) /-- Target directory for build files -/ def LIBDIR : FilePath := - "build" / "lib" + ".lake" / "build" / "lib" /-- Target directory for IR files -/ def IRDIR : FilePath := - "build" / "ir" + ".lake" / "build" / "ir" /-- Target directory for caching -/ initialize CACHEDIR : FilePath ← do @@ -78,7 +78,7 @@ def LEANTARBIN := IO.CACHEDIR / s!"leantar-{LEANTARVERSION}{EXE}" def LAKEPACKAGESDIR : FilePath := - ⟨"lake-packages"⟩ + ".lake" / "packages" def getCurl : IO String := do return if (← CURLBIN.pathExists) then CURLBIN.toString else "curl" diff --git a/Counterexamples/CliffordAlgebra_not_injective.lean b/Counterexamples/CliffordAlgebra_not_injective.lean index cfa4ea3630e30..b012a3a673e77 100644 --- a/Counterexamples/CliffordAlgebra_not_injective.lean +++ b/Counterexamples/CliffordAlgebra_not_injective.lean @@ -56,6 +56,7 @@ theorem X0_X1_X2_not_mem_kIdeal : (X 0 * X 1 * X 2 : MvPolynomial (Fin 3) (ZMod simp_rw [mem_kIdeal_iff, support_mul_X, support_X, Finset.map_singleton, addRightEmbedding_apply, Finset.mem_singleton, forall_eq, ← Fin.sum_univ_three fun i => Finsupp.single i 1, ← Finsupp.equivFunOnFinite_symm_eq_sum] at h + contradiction theorem mul_self_mem_kIdeal_of_X0_X1_X2_mul_mem {x : MvPolynomial (Fin 3) (ZMod 2)} (h : X 0 * X 1 * X 2 * x ∈ kIdeal) : x * x ∈ kIdeal := by @@ -239,7 +240,7 @@ theorem quot_obv : α • x' - β • y' - γ • z' = 0 := by ← Submodule.Quotient.mk_sub] convert LinearMap.map_zero _ using 2 rw [Submodule.Quotient.mk_eq_zero] - norm_num [sub_zero, Ideal.span, Pi.single_apply] + simp (config := {decide := true}) [sub_zero, Ideal.span, Pi.single_apply] /-- The core of the proof - scaling `1` by `α * β * γ` gives zero -/ theorem αβγ_smul_eq_zero : (α * β * γ) • (1 : CliffordAlgebra Q) = 0 := by diff --git a/Counterexamples/Cyclotomic105.lean b/Counterexamples/Cyclotomic105.lean index f05b5d44f1eba..07a1cd6df40d5 100644 --- a/Counterexamples/Cyclotomic105.lean +++ b/Counterexamples/Cyclotomic105.lean @@ -109,7 +109,7 @@ theorem cyclotomic_105 : #align counterexample.cyclotomic_105 Counterexample.cyclotomic_105 theorem coeff_cyclotomic_105 : coeff (cyclotomic 105 ℤ) 7 = -2 := by - simp [cyclotomic_105, coeff_X_pow, coeff_one, coeff_X_of_ne_one, coeff_bit0_mul, two_mul] + simp [cyclotomic_105, coeff_one, coeff_X_of_ne_one] #align counterexample.coeff_cyclotomic_105 Counterexample.coeff_cyclotomic_105 theorem not_forall_coeff_cyclotomic_neg_one_zero_one : diff --git a/Counterexamples/HomogeneousPrimeNotPrime.lean b/Counterexamples/HomogeneousPrimeNotPrime.lean index dd776d4703975..d3d2dcdb4aec1 100644 --- a/Counterexamples/HomogeneousPrimeNotPrime.lean +++ b/Counterexamples/HomogeneousPrimeNotPrime.lean @@ -175,7 +175,7 @@ theorem homogeneous_mem_or_mem {x y : R × R} (hx : SetLike.Homogeneous (grading -- Porting note: added `h2` for later use; the proof is hideous have h2 : Prime (2:R) := by unfold Prime - simp only [true_and] + refine ⟨by decide, by decide, ?_⟩ intro a b have aux2 : (Fin.mk 2 _ : R) = 2 := rfl have aux3 : (Fin.mk 3 _ : R) = -1 := rfl diff --git a/Mathlib.lean b/Mathlib.lean index 4e9345bfa92f2..59d9062aaa727 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1572,7 +1572,6 @@ import Mathlib.Data.Int.SuccPred import Mathlib.Data.Int.Units import Mathlib.Data.IsROrC.Basic import Mathlib.Data.IsROrC.Lemmas -import Mathlib.Data.KVMap import Mathlib.Data.LazyList import Mathlib.Data.LazyList.Basic import Mathlib.Data.List.AList @@ -2255,7 +2254,6 @@ import Mathlib.Lean.Expr.ReplaceRec import Mathlib.Lean.Expr.Traverse import Mathlib.Lean.IO.Process import Mathlib.Lean.Json -import Mathlib.Lean.Linter import Mathlib.Lean.LocalContext import Mathlib.Lean.Message import Mathlib.Lean.Meta diff --git a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean index c1d7943722fc5..ae5dbed5d47f4 100644 --- a/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Algebra/NonUnitalSubalgebra.lean @@ -287,7 +287,7 @@ protected theorem coe_sub {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algeb @[simp, norm_cast] theorem coe_smul [Semiring R'] [SMul R' R] [Module R' A] [IsScalarTower R' R A] (r : R') (x : S) : - (r • x : A) = r • (x : A) := + ↑(r • x) = r • (x : A) := rfl protected theorem coe_eq_zero {x : S} : (x : A) = 0 ↔ x = 0 := diff --git a/Mathlib/Algebra/BigOperators/Basic.lean b/Mathlib/Algebra/BigOperators/Basic.lean index 47dfb43221cb7..bd32cb17f490b 100644 --- a/Mathlib/Algebra/BigOperators/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Basic.lean @@ -1572,7 +1572,7 @@ theorem prod_piecewise [DecidableEq α] (s t : Finset α) (f g : α → β) : theorem prod_inter_mul_prod_diff [DecidableEq α] (s t : Finset α) (f : α → β) : (∏ x in s ∩ t, f x) * ∏ x in s \ t, f x = ∏ x in s, f x := by convert (s.prod_piecewise t f f).symm - simp [Finset.piecewise] + simp (config := { unfoldPartialApp := true }) [Finset.piecewise] #align finset.prod_inter_mul_prod_diff Finset.prod_inter_mul_prod_diff #align finset.sum_inter_add_sum_diff Finset.sum_inter_add_sum_diff @@ -2151,7 +2151,7 @@ theorem finset_sum_eq_sup_iff_disjoint {β : Type*} {i : Finset β} {f : β → · simp only [Finset.not_mem_empty, IsEmpty.forall_iff, imp_true_iff, Finset.sum_empty, Finset.sup_empty, bot_eq_zero, eq_self_iff_true] · simp_rw [Finset.sum_cons hz, Finset.sup_cons, Finset.mem_cons, Multiset.sup_eq_union, - forall_eq_or_imp, Ne.def, IsEmpty.forall_iff, true_and_iff, + forall_eq_or_imp, Ne.def, not_true_eq_false, IsEmpty.forall_iff, true_and_iff, imp_and, forall_and, ← hr, @eq_comm _ z] have := fun x (H : x ∈ i) => ne_of_mem_of_not_mem H hz simp (config := { contextual := true }) only [this, not_false_iff, true_imp_iff] @@ -2310,7 +2310,7 @@ theorem nat_abs_sum_le {ι : Type*} (s : Finset ι) (f : ι → ℤ) : (∑ i in s, f i).natAbs ≤ ∑ i in s, (f i).natAbs := by classical induction' s using Finset.induction_on with i s his IH - · simp only [Finset.sum_empty, Int.natAbs_zero] + · simp only [Finset.sum_empty, Int.natAbs_zero, le_refl] · simp only [his, Finset.sum_insert, not_false_iff] exact (Int.natAbs_add_le _ _).trans (add_le_add le_rfl IH) #align nat_abs_sum_le nat_abs_sum_le diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index ef9cff13dd432..903df9af6bf13 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -124,8 +124,6 @@ instance (priority := 100) sMulCommClass_mk {R : Type u₁} {S : Type u₂} [Rin (f : R →+* S) (M : Type v) [I : AddCommGroup M] [Module S M] : haveI : SMul R M := (RestrictScalars.obj' f (ModuleCat.mk M)).isModule.toSMul SMulCommClass R S M := - -- Porting note: cannot synth SMul R M - have : SMul R M := (RestrictScalars.obj' f (ModuleCat.mk M)).isModule.toSMul @SMulCommClass.mk R S M (_) _ <| fun r s m => (by simp [← mul_smul, mul_comm] : f r • s • m = s • f r • m) #align category_theory.Module.smul_comm_class_mk ModuleCat.sMulCommClass_mk diff --git a/Mathlib/Algebra/CharP/Basic.lean b/Mathlib/Algebra/CharP/Basic.lean index f6dda71b19889..2dce592d4520c 100644 --- a/Mathlib/Algebra/CharP/Basic.lean +++ b/Mathlib/Algebra/CharP/Basic.lean @@ -716,8 +716,8 @@ theorem Int.cast_injOn_of_ringChar_ne_two {R : Type*} [NonAssocRing R] [Nontrivi rintro _ (rfl | rfl | rfl) _ (rfl | rfl | rfl) h <;> simp only [cast_neg, cast_one, cast_zero, neg_eq_zero, one_ne_zero, zero_ne_one, zero_eq_neg] at h ⊢ - · exact (Ring.neg_one_ne_one_of_char_ne_two hR).symm h - · exact (Ring.neg_one_ne_one_of_char_ne_two hR) h + · exact ((Ring.neg_one_ne_one_of_char_ne_two hR).symm h).elim + · exact ((Ring.neg_one_ne_one_of_char_ne_two hR) h).elim #align int.cast_inj_on_of_ring_char_ne_two Int.cast_injOn_of_ringChar_ne_two end diff --git a/Mathlib/Algebra/CharP/LocalRing.lean b/Mathlib/Algebra/CharP/LocalRing.lean index 2b4137755a246..a932451a0ca01 100644 --- a/Mathlib/Algebra/CharP/LocalRing.lean +++ b/Mathlib/Algebra/CharP/LocalRing.lean @@ -51,7 +51,7 @@ theorem charP_zero_or_prime_power (R : Type*) [CommRing R] [LocalRing R] (q : -- Let `b` be the inverse of `a`. cases' a_unit.exists_left_inv with a_inv h_inv_mul_a have rn_cast_zero : ↑(r ^ n) = (0 : R) := by - rw [← @mul_one R _ (r ^ n), mul_comm, ←Classical.choose_spec a_unit.exists_left_inv, + rw [← @mul_one R _ ↑(r ^ n), mul_comm, ←Classical.choose_spec a_unit.exists_left_inv, mul_assoc, ← Nat.cast_mul, ←q_eq_a_mul_rn, CharP.cast_eq_zero R q] simp have q_eq_rn := Nat.dvd_antisymm ((CharP.cast_eq_zero_iff R q (r ^ n)).mp rn_cast_zero) rn_dvd_q diff --git a/Mathlib/Algebra/CharZero/Quotient.lean b/Mathlib/Algebra/CharZero/Quotient.lean index d46c5853c577f..bf2cb7038ce32 100644 --- a/Mathlib/Algebra/CharZero/Quotient.lean +++ b/Mathlib/Algebra/CharZero/Quotient.lean @@ -53,7 +53,7 @@ end AddSubgroup namespace QuotientAddGroup theorem zmultiples_zsmul_eq_zsmul_iff {ψ θ : R ⧸ AddSubgroup.zmultiples p} {z : ℤ} (hz : z ≠ 0) : - z • ψ = z • θ ↔ ∃ k : Fin z.natAbs, ψ = θ + (k : ℕ) • (p / z : R) := by + z • ψ = z • θ ↔ ∃ k : Fin z.natAbs, ψ = θ + ((k : ℕ) • (p / z) : R) := by induction ψ using Quotient.inductionOn' induction θ using Quotient.inductionOn' -- Porting note: Introduced Zp notation to shorten lines diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 9980bd36baa50..f26c57e466078 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -143,8 +143,9 @@ theorem geom_sum₂_self {α : Type*} [CommRing α] (x : α) (n : ℕ) : ∑ i in range n, x ^ i * x ^ (n - 1 - i) = n * x ^ (n - 1) := calc ∑ i in Finset.range n, x ^ i * x ^ (n - 1 - i) = - ∑ i in Finset.range n, x ^ (i + (n - 1 - i)) := by simp_rw [← pow_add] - _ = ∑ i in Finset.range n, x ^ (n - 1) := + ∑ i in Finset.range n, x ^ (i + (n - 1 - i)) := + by simp_rw [← pow_add] + _ = ∑ _i in Finset.range n, x ^ (n - 1) := Finset.sum_congr rfl fun i hi => congr_arg _ <| add_tsub_cancel_of_le <| Nat.le_sub_one_of_lt <| Finset.mem_range.1 hi _ = (Finset.range n).card • x ^ (n - 1) := Finset.sum_const _ @@ -482,7 +483,7 @@ theorem geom_sum_alternating_of_le_neg_one [StrictOrderedRing α] (hx : x + 1 if Even n then (∑ i in range n, x ^ i) ≤ 0 else 1 ≤ ∑ i in range n, x ^ i := by have hx0 : x ≤ 0 := (le_add_of_nonneg_right zero_le_one).trans hx induction' n with n ih - · simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true] + · simp only [Nat.zero_eq, range_zero, sum_empty, le_refl, ite_true, even_zero] simp only [Nat.even_add_one, geom_sum_succ] split_ifs at ih with h · rw [if_neg (not_not_intro h), le_add_iff_nonneg_left] @@ -496,7 +497,7 @@ theorem geom_sum_alternating_of_lt_neg_one [StrictOrderedRing α] (hx : x + 1 < if Even n then (∑ i in range n, x ^ i) < 0 else 1 < ∑ i in range n, x ^ i := by have hx0 : x < 0 := ((le_add_iff_nonneg_right _).2 zero_le_one).trans_lt hx refine' Nat.le_induction _ _ n (show 2 ≤ n from hn) - · simp only [geom_sum_two, lt_add_iff_pos_left, ite_true, gt_iff_lt, hx] + · simp only [geom_sum_two, lt_add_iff_pos_left, ite_true, gt_iff_lt, hx, even_two] clear hn intro n _ ihn simp only [Nat.even_add_one, geom_sum_succ] @@ -570,7 +571,7 @@ theorem geom_sum_eq_zero_iff_neg_one [LinearOrderedRing α] (hn : n ≠ 0) : have hx := eq_or_ne x (-1) cases' hx with hx hx · rw [hx, neg_one_geom_sum] - simp only [h hx, ne_eq, ite_eq_left_iff, one_ne_zero, not_forall, exists_prop, and_true] + simp only [h hx, ite_false, ne_eq, one_ne_zero, not_false_eq_true] · exact geom_sum_ne_zero hx hn #align geom_sum_eq_zero_iff_neg_one geom_sum_eq_zero_iff_neg_one diff --git a/Mathlib/Algebra/GradedMonoid.lean b/Mathlib/Algebra/GradedMonoid.lean index 96c5d96362e20..aadfc5ba3fe89 100644 --- a/Mathlib/Algebra/GradedMonoid.lean +++ b/Mathlib/Algebra/GradedMonoid.lean @@ -320,7 +320,7 @@ section Monoid variable [AddMonoid ι] [GMonoid A] -instance : Pow (A 0) ℕ where +instance : NatPow (A 0) where pow x n := @Eq.rec ι (n • (0:ι)) (fun a _ => A a) (GMonoid.gnpow n x) 0 (nsmul_zero n) variable {A} diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 09b1504af0356..c2d441f39a6e8 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -56,10 +56,18 @@ class HVAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) where /-- The notation typeclass for heterogeneous scalar multiplication. This enables the notation `a • b : γ` where `a : α`, `b : β`. + +It is assumed to represent a left action in some sense. +The notation `a • b` is augmented with a macro (below) to have it elaborate as a left action. +Only the `b` argument participates in the elaboration algorithm: the algorithm uses the type of `b` +when calculating the type of the surrounding arithmetic expression +and it tries to insert coercions into `b` to get some `b'` +such that `a • b'` has the same type as `b'`. +See the module documentation near the macro for more details. -/ class HSMul (α : Type u) (β : Type v) (γ : outParam (Type w)) where /-- `a • b` computes the product of `a` and `b`. - The meaning of this notation is type-dependent. -/ + The meaning of this notation is type-dependent, but it is intended to be used for left actions. -/ hSMul : α → β → γ attribute [notation_class smul Simps.copySecond] HSMul @@ -86,6 +94,41 @@ infixl:65 " +ᵥ " => HVAdd.hVAdd infixl:65 " -ᵥ " => VSub.vsub infixr:73 " • " => HSMul.hSMul +/-! +We have a macro to make `x • y` notation participate in the expression tree elaborator, +like other arithmetic expressions such as `+`, `*`, `/`, `^`, `=`, inequalities, etc. +The macro is using the `leftact%` elaborator introduced in +[this RFC](https://github.com/leanprover/lean4/issues/2854). + +As a concrete example of the effect of this macro, consider +```lean +variable [Ring R] [AddCommMonoid M] [Module R M] (r : R) (N : Submodule R M) (m : M) (n : N) +#check m + r • n +``` +Without the macro, the expression would elaborate as `m + ↑(r • n : ↑N) : M`. +With the macro, the expression elaborates as `m + r • (↑n : M) : M`. +To get the first intepretation, one can write `m + (r • n :)`. + +Here is a quick review of the expression tree elaborator: +1. It builds up an expression tree of all the immediately accessible operations + that are marked with `binop%`, `unop%`, `leftact%`, `rightact%`, `binrel%`, etc. +2. It elaborates every leaf term of this tree + (without an expected type, so as if it were temporarily wrapped in `(... :)`). +3. Using the types of each elaborated leaf, it computes a supremum type they can all be + coerced to, if such a supremum exists. +4. It inserts coercions around leaf terms wherever needed. + +The hypothesis is that individual expression trees tend to be calculations with respect +to a single algebraic structure. + +Note(kmill): If we were to remove `HSMul` and switch to using `SMul` directly, +then the expression tree elaborator would not be able to insert coercions within the right operand; +they would likely appear as `↑(x • y)` rather than `x • ↑y`, unlike other arithmetic operations. +-/ + +@[inherit_doc HSMul.hSMul] +macro_rules | `($x • $y) => `(leftact% HSMul.hSMul $x $y) + attribute [to_additive existing] Mul Div HMul instHMul HDiv instHDiv HSMul attribute [to_additive (reorder := 1 2) SMul] Pow attribute [to_additive (reorder := 1 2)] HPow diff --git a/Mathlib/Algebra/GroupPower/NegOnePow.lean b/Mathlib/Algebra/GroupPower/NegOnePow.lean index f38c97f16b351..e3de679ff6df6 100644 --- a/Mathlib/Algebra/GroupPower/NegOnePow.lean +++ b/Mathlib/Algebra/GroupPower/NegOnePow.lean @@ -55,6 +55,7 @@ lemma negOnePow_eq_one_iff (n : ℤ) : n.negOnePow = 1 ↔ Even n := by rw [Int.even_iff_not_odd] intro h' simp only [negOnePow_odd _ h'] at h + contradiction · exact negOnePow_even n lemma negOnePow_eq_neg_one_iff (n : ℤ) : n.negOnePow = -1 ↔ Odd n := by @@ -63,7 +64,7 @@ lemma negOnePow_eq_neg_one_iff (n : ℤ) : n.negOnePow = -1 ↔ Odd n := by rw [Int.odd_iff_not_even] intro h' rw [negOnePow_even _ h'] at h - simp only at h + contradiction · exact negOnePow_odd n @[simp] diff --git a/Mathlib/Algebra/Homology/Augment.lean b/Mathlib/Algebra/Homology/Augment.lean index fe2d9689017c8..b19030893c688 100644 --- a/Mathlib/Algebra/Homology/Augment.lean +++ b/Mathlib/Algebra/Homology/Augment.lean @@ -239,7 +239,7 @@ def augment (C : CochainComplex V ℕ) {X : V} (f : X ⟶ C.X 0) (w : f ≫ C.d shape i j s := by simp at s rcases j with (_ | _ | j) <;> cases i <;> try simp - · simp at s + · contradiction · rw [C.shape] simp only [ComplexShape.up_Rel] contrapose! s diff --git a/Mathlib/Algebra/Homology/Exact.lean b/Mathlib/Algebra/Homology/Exact.lean index e2b84689772d6..3167d21fea54b 100644 --- a/Mathlib/Algebra/Homology/Exact.lean +++ b/Mathlib/Algebra/Homology/Exact.lean @@ -150,7 +150,7 @@ theorem imageToKernel_isIso_of_image_eq_kernel {A B C : V} (f : A ⟶ B) (g : B IsIso (imageToKernel f g (comp_eq_zero_of_image_eq_kernel f g p)) := by refine' ⟨⟨Subobject.ofLE _ _ p.ge, _⟩⟩ dsimp [imageToKernel] - simp only [Subobject.ofLE_comp_ofLE, Subobject.ofLE_refl] + simp only [Subobject.ofLE_comp_ofLE, Subobject.ofLE_refl, and_self] #align category_theory.image_to_kernel_is_iso_of_image_eq_kernel CategoryTheory.imageToKernel_isIso_of_image_eq_kernel -- We'll prove the converse later, when `V` is abelian. diff --git a/Mathlib/Algebra/Homology/ExactSequence.lean b/Mathlib/Algebra/Homology/ExactSequence.lean index 3022633ad1397..da5fec996eebb 100644 --- a/Mathlib/Algebra/Homology/ExactSequence.lean +++ b/Mathlib/Algebra/Homology/ExactSequence.lean @@ -70,7 +70,9 @@ lemma isComplex_iff_of_iso {S₁ S₂ : ComposableArrows C n} (e : S₁ ≅ S₂ ⟨isComplex_of_iso e, isComplex_of_iso e.symm⟩ lemma isComplex₀ (S : ComposableArrows C 0) : S.IsComplex where - zero i hi := by simp at hi + -- See https://github.com/leanprover/lean4/issues/2862 + -- Without `decide := true`, simp gets stuck at `hi : autoParam False _auto✝` + zero i hi := by simp (config := {decide := true}) at hi lemma isComplex₁ (S : ComposableArrows C 1) : S.IsComplex where zero i hi := by exfalso; linarith @@ -157,7 +159,8 @@ lemma exact_iff_of_iso {S₁ S₂ : ComposableArrows C n} (e : S₁ ≅ S₂) : lemma exact₀ (S : ComposableArrows C 0) : S.Exact where toIsComplex := S.isComplex₀ - exact i hi := by simp at hi + -- See https://github.com/leanprover/lean4/issues/2862 + exact i hi := by simp [autoParam] at hi lemma exact₁ (S : ComposableArrows C 1) : S.Exact where toIsComplex := S.isComplex₁ diff --git a/Mathlib/Algebra/Homology/Flip.lean b/Mathlib/Algebra/Homology/Flip.lean index 678caa3a87780..8e4ad84660353 100644 --- a/Mathlib/Algebra/Homology/Flip.lean +++ b/Mathlib/Algebra/Homology/Flip.lean @@ -39,7 +39,7 @@ def flipObj (C : HomologicalComplex (HomologicalComplex V c) c') : { X := fun j => (C.X j).X i d := fun j j' => (C.d j j').f i shape := fun j j' w => by - simp_all only [shape, zero_f] + simp_all only [not_false_eq_true, shape, zero_f] d_comp_d' := fun j₁ j₂ j₃ _ _ => congr_hom (C.d_comp_d j₁ j₂ j₃) i } d i i' := { f := fun j => (C.X j).d i i' diff --git a/Mathlib/Algebra/Homology/Homotopy.lean b/Mathlib/Algebra/Homology/Homotopy.lean index 55b7e688b2398..a0a7d0faf0a25 100644 --- a/Mathlib/Algebra/Homology/Homotopy.lean +++ b/Mathlib/Algebra/Homology/Homotopy.lean @@ -565,7 +565,7 @@ def mkInductive : Homotopy e 0 where · cases i · dsimp [fromNext, mkInductiveAux₂] rw [dif_neg] - simp only + decide · dsimp [fromNext] simp only [ChainComplex.next_nat_succ, dite_true] rw [mkInductiveAux₃ e zero comm_zero one comm_one succ] @@ -704,7 +704,7 @@ def mkCoinductive : Homotopy e 0 where · cases i · dsimp [toPrev, mkCoinductiveAux₂] rw [dif_neg] - simp only + decide · dsimp [toPrev] simp only [CochainComplex.prev_nat_succ, dite_true] rw [mkCoinductiveAux₃ e zero comm_zero one comm_one succ] diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index 67169d09ab035..b1be088bf7c86 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -229,6 +229,7 @@ noncomputable def fromSingle₀Equiv (C : ChainComplex V ℕ) (X : V) : invFun f := HomologicalComplex.mkHomFromSingle f (fun i hi => by simp at hi) left_inv := by aesop_cat right_inv := by aesop_cat +#align chain_complex.from_single₀_equiv ChainComplex.fromSingle₀Equiv @[simp] lemma fromSingle₀Equiv_symm_apply_f_zero diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index 8050955c70d8c..72ac069276671 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -101,7 +101,7 @@ theorem IsPrimePow.dvd {n m : ℕ} (hn : IsPrimePow n) (hm : m ∣ n) (hm₁ : m refine' ⟨p, i, hp, _, rfl⟩ apply Nat.pos_of_ne_zero rintro rfl - simp only [pow_zero, ne_eq] at hm₁ + simp only [pow_zero, ne_eq, not_true_eq_false] at hm₁ #align is_prime_pow.dvd IsPrimePow.dvd theorem Nat.disjoint_divisors_filter_isPrimePow {a b : ℕ} (hab : a.Coprime b) : diff --git a/Mathlib/Algebra/Lie/DirectSum.lean b/Mathlib/Algebra/Lie/DirectSum.lean index 41d401ae4a8d1..97eda2d1f63a6 100644 --- a/Mathlib/Algebra/Lie/DirectSum.lean +++ b/Mathlib/Algebra/Lie/DirectSum.lean @@ -145,7 +145,7 @@ theorem lie_of [DecidableEq ι] {i j : ι} (x : L i) (y : L j) : ⁅of L i x, of L j y⁆ = if hij : i = j then of L i ⁅x, hij.symm.recOn y⁆ else 0 := by obtain rfl | hij := Decidable.eq_or_ne i j · simp only [lie_of_same L x y, dif_pos] - · simp only [lie_of_of_ne L hij x y, hij, dif_neg] + · simp only [lie_of_of_ne L hij x y, hij, dif_neg, dite_false] #align direct_sum.lie_of DirectSum.lie_of instance lieAlgebra : LieAlgebra R (⨁ i, L i) := diff --git a/Mathlib/Algebra/Module/DedekindDomain.lean b/Mathlib/Algebra/Module/DedekindDomain.lean index 9fc146976d910..999d73a16f4f2 100644 --- a/Mathlib/Algebra/Module/DedekindDomain.lean +++ b/Mathlib/Algebra/Module/DedekindDomain.lean @@ -39,7 +39,7 @@ torsion submodules, where `I = ∏ i, p i ^ e i` is its unique decomposition in theorem isInternal_prime_power_torsion_of_is_torsion_by_ideal {I : Ideal R} (hI : I ≠ ⊥) (hM : Module.IsTorsionBySet R M I) : DirectSum.IsInternal fun p : (factors I).toFinset => - torsionBySet R M ((p : Ideal R) ^ (factors I).count ↑p) := by + torsionBySet R M (p ^ (factors I).count ↑p : Ideal R) := by let P := factors I have prime_of_mem := fun p (hp : p ∈ P.toFinset) => prime_of_factor p (Multiset.mem_toFinset.mp hp) @@ -66,7 +66,7 @@ theorem isInternal_prime_power_torsion_of_is_torsion_by_ideal {I : Ideal R} (hI `e i` are their multiplicities. -/ theorem isInternal_prime_power_torsion [Module.Finite R M] (hM : Module.IsTorsion R M) : DirectSum.IsInternal fun p : (factors (⊤ : Submodule R M).annihilator).toFinset => - torsionBySet R M ((p : Ideal R) ^ (factors (⊤ : Submodule R M).annihilator).count ↑p) := by + torsionBySet R M (p ^ (factors (⊤ : Submodule R M).annihilator).count ↑p : Ideal R) := by have hM' := Module.isTorsionBySet_annihilator_top R M have hI := Submodule.annihilator_top_inter_nonZeroDivisors hM refine' isInternal_prime_power_torsion_of_is_torsion_by_ideal _ hM' @@ -78,7 +78,7 @@ theorem isInternal_prime_power_torsion [Module.Finite R M] (hM : Module.IsTorsio `p i ^ e i`-torsion submodules for some prime ideals `p i` and numbers `e i`.-/ theorem exists_isInternal_prime_power_torsion [Module.Finite R M] (hM : Module.IsTorsion R M) : ∃ (P : Finset <| Ideal R) (_ : DecidableEq P) (_ : ∀ p ∈ P, Prime p) (e : P → ℕ), - DirectSum.IsInternal fun p : P => torsionBySet R M ((p : Ideal R) ^ e p) := + DirectSum.IsInternal fun p : P => torsionBySet R M (p ^ e p : Ideal R) := ⟨_, _, fun p hp => prime_of_factor p (Multiset.mem_toFinset.mp hp), _, isInternal_prime_power_torsion hM⟩ #align submodule.exists_is_internal_prime_power_torsion Submodule.exists_isInternal_prime_power_torsion diff --git a/Mathlib/Algebra/Module/Injective.lean b/Mathlib/Algebra/Module/Injective.lean index 1682aa34aca48..f9ea188004147 100644 --- a/Mathlib/Algebra/Module/Injective.lean +++ b/Mathlib/Algebra/Module/Injective.lean @@ -415,7 +415,7 @@ def extensionOfMaxAdjoin (h : Module.Baer R Q) (y : N) : ExtensionOf i f where ↑(r • ExtensionOfMaxAdjoin.fst i a) + (r • ExtensionOfMaxAdjoin.snd i a) • y := by rw [ExtensionOfMaxAdjoin.eqn, smul_add, smul_eq_mul, mul_smul] rfl - rw [ExtensionOfMaxAdjoin.extensionToFun_wd i f h (r • a) _ _ eq1, LinearMap.map_smul, + rw [ExtensionOfMaxAdjoin.extensionToFun_wd i f h (r • a :) _ _ eq1, LinearMap.map_smul, LinearPMap.map_smul, ← smul_add] congr } is_extension m := by diff --git a/Mathlib/Algebra/Order/Interval.lean b/Mathlib/Algebra/Order/Interval.lean index 2d26ed32e73b9..207ef40c4cd73 100644 --- a/Mathlib/Algebra/Order/Interval.lean +++ b/Mathlib/Algebra/Order/Interval.lean @@ -304,7 +304,7 @@ namespace NonemptyInterval @[to_additive] theorem coe_pow_interval [OrderedCommMonoid α] (s : NonemptyInterval α) (n : ℕ) : - (s ^ n : Interval α) = (s : Interval α) ^ n := + ↑(s ^ n) = (s : Interval α) ^ n := map_pow (⟨⟨(↑), coe_one_interval⟩, coe_mul_interval⟩ : NonemptyInterval α →* Interval α) _ _ #align nonempty_interval.coe_pow_interval NonemptyInterval.coe_pow_interval #align nonempty_interval.coe_nsmul_interval NonemptyInterval.coe_nsmul_interval diff --git a/Mathlib/Algebra/Order/Positive/Ring.lean b/Mathlib/Algebra/Order/Positive/Ring.lean index a9f802372208f..8d32a13e2effa 100644 --- a/Mathlib/Algebra/Order/Positive/Ring.lean +++ b/Mathlib/Algebra/Order/Positive/Ring.lean @@ -109,7 +109,7 @@ instance : Pow { x : R // 0 < x } ℕ := @[simp] theorem val_pow (x : { x : R // 0 < x }) (n : ℕ) : - (x ^ n : R) = (x : R) ^ n := + ↑(x ^ n) = (x : R) ^ n := rfl #align positive.coe_pow Positive.val_pow diff --git a/Mathlib/Algebra/Order/Ring/WithTop.lean b/Mathlib/Algebra/Order/Ring/WithTop.lean index 6a3992e306aab..a14188fab2fcd 100644 --- a/Mathlib/Algebra/Order/Ring/WithTop.lean +++ b/Mathlib/Algebra/Order/Ring/WithTop.lean @@ -67,7 +67,7 @@ theorem mul_eq_top_iff {a b : WithTop α} : a * b = ⊤ ↔ a ≠ 0 ∧ b = ⊤ theorem mul_lt_top' [LT α] {a b : WithTop α} (ha : a < ⊤) (hb : b < ⊤) : a * b < ⊤ := by rw [WithTop.lt_top_iff_ne_top] at * - simp only [Ne.def, mul_eq_top_iff, *, and_false, false_and, false_or] + simp only [Ne.def, mul_eq_top_iff, *, and_false, false_and, or_self, not_false_eq_true] #align with_top.mul_lt_top' WithTop.mul_lt_top' theorem mul_lt_top [LT α] {a b : WithTop α} (ha : a ≠ ⊤) (hb : b ≠ ⊤) : a * b < ⊤ := diff --git a/Mathlib/Algebra/Order/Sub/WithTop.lean b/Mathlib/Algebra/Order/Sub/WithTop.lean index 4baf09510cae5..dedf9ca024409 100644 --- a/Mathlib/Algebra/Order/Sub/WithTop.lean +++ b/Mathlib/Algebra/Order/Sub/WithTop.lean @@ -47,8 +47,8 @@ theorem sub_top {a : WithTop α} : a - ⊤ = 0 := by cases a <;> rfl @[simp] theorem sub_eq_top_iff {a b : WithTop α} : a - b = ⊤ ↔ a = ⊤ ∧ b ≠ ⊤ := by induction a using recTopCoe <;> induction b using recTopCoe <;> - simp only [← coe_sub, coe_ne_top, sub_top, zero_ne_top, coe_ne_top, top_sub_coe, false_and, - Ne.def] + simp only [← coe_sub, coe_ne_top, sub_top, zero_ne_top, top_sub_coe, false_and, + Ne.def, not_true_eq_false, not_false_eq_true, and_false, and_self] #align with_top.sub_eq_top_iff WithTop.sub_eq_top_iff theorem map_sub [Sub β] [Zero β] {f : α → β} (h : ∀ x y, f (x - y) = f x - f y) (h₀ : f 0 = 0) : diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 12ac392fbbbfa..9a1029e700d18 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -1454,7 +1454,7 @@ section QuaternionAlgebra variable {R : Type*} (c₁ c₂ : R) private theorem pow_four [Infinite R] : #R ^ℕ 4 = #R := - power_nat_eq (aleph0_le_mk R) <| by simp + power_nat_eq (aleph0_le_mk R) <| by decide /-- The cardinality of a quaternion algebra, as a type. -/ theorem mk_quaternionAlgebra : #(ℍ[R,c₁,c₂]) = #R ^ℕ 4 := by diff --git a/Mathlib/Algebra/Regular/SMul.lean b/Mathlib/Algebra/Regular/SMul.lean index 9fbd9eada1976..624f4175b549b 100644 --- a/Mathlib/Algebra/Regular/SMul.lean +++ b/Mathlib/Algebra/Regular/SMul.lean @@ -73,7 +73,7 @@ theorem smul (ra : IsSMulRegular M a) (rs : IsSMulRegular M s) : IsSMulRegular M element, then `b` is `M`-regular. -/ theorem of_smul (a : R) (ab : IsSMulRegular M (a • s)) : IsSMulRegular M s := @Function.Injective.of_comp _ _ _ (fun m : M => a • m) _ fun c d cd => by - dsimp only [Function.comp] at cd + dsimp only [Function.comp_def] at cd rw [←smul_assoc, ←smul_assoc] at cd exact ab cd #align is_smul_regular.of_smul IsSMulRegular.of_smul @@ -133,7 +133,7 @@ variable (M) /-- One is always `M`-regular. -/ @[simp] theorem one : IsSMulRegular M (1 : R) := fun a b ab => by - dsimp only [Function.comp] at ab + dsimp only [Function.comp_def] at ab rw [one_smul, one_smul] at ab assumption #align is_smul_regular.one IsSMulRegular.one @@ -186,7 +186,7 @@ variable [MonoidWithZero R] [MonoidWithZero S] [Zero M] [MulActionWithZero R M] /-- The element `0` is `M`-regular if and only if `M` is trivial. -/ protected theorem subsingleton (h : IsSMulRegular M (0 : R)) : Subsingleton M := - ⟨fun a b => h (by dsimp only [Function.comp]; repeat' rw [MulActionWithZero.zero_smul])⟩ + ⟨fun a b => h (by dsimp only [Function.comp_def]; repeat' rw [MulActionWithZero.zero_smul])⟩ #align is_smul_regular.subsingleton IsSMulRegular.subsingleton /-- The element `0` is `M`-regular if and only if `M` is trivial. -/ diff --git a/Mathlib/Algebra/RingQuot.lean b/Mathlib/Algebra/RingQuot.lean index 1142702d1ac84..a242981d9bb17 100644 --- a/Mathlib/Algebra/RingQuot.lean +++ b/Mathlib/Algebra/RingQuot.lean @@ -218,7 +218,7 @@ instance : Add (RingQuot r) := instance : Mul (RingQuot r) := ⟨mul r⟩ -instance : Pow (RingQuot r) ℕ := +instance : NatPow (RingQuot r) := ⟨fun x n ↦ npow r n x⟩ instance {R : Type uR} [Ring R] (r : R → R → Prop) : Neg (RingQuot r) := diff --git a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean index a46869adec3cd..f53ccb54c57a6 100644 --- a/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean +++ b/Mathlib/Algebra/Star/NonUnitalSubalgebra.lean @@ -282,7 +282,7 @@ protected theorem coe_sub {R : Type u} {A : Type v} [CommRing R] [NonUnitalRing @[simp, norm_cast] theorem coe_smul [Semiring R'] [SMul R' R] [Module R' A] [IsScalarTower R' R A] (r : R') (x : S) : - (r • x : A) = r • (x : A) := + ↑(r • x) = r • (x : A) := rfl protected theorem coe_eq_zero {x : S} : (x : A) = 0 ↔ x = 0 := diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean index 90e2e1ece6371..552e73d07efe7 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Weierstrass.lean @@ -638,8 +638,8 @@ lemma irreducible_polynomial [IsDomain R] : Irreducible W.polynomial := by apply (h1.symm.le.trans Cubic.degree_of_b_eq_zero').not_lt rcases Nat.WithBot.add_eq_three_iff.mp h0.symm with h | h | h | h -- porting note: replaced two `any_goals` proofs with two `iterate 2` proofs - iterate 2 rw [degree_add_eq_right_of_degree_lt] <;> simp only [h] - iterate 2 rw [degree_add_eq_left_of_degree_lt] <;> simp only [h] + iterate 2 rw [degree_add_eq_right_of_degree_lt] <;> simp only [h] <;> decide + iterate 2 rw [degree_add_eq_left_of_degree_lt] <;> simp only [h] <;> decide #align weierstrass_curve.irreducible_polynomial WeierstrassCurve.irreducible_polynomial -- porting note: removed `@[simp]` to avoid a `simpNF` linter error @@ -1029,7 +1029,7 @@ lemma norm_smul_basis (p q : R[X]) : simp_rw [Algebra.norm_eq_matrix_det <| CoordinateRing.basis W, Matrix.det_fin_two, Algebra.leftMulMatrix_eq_repr_mul, basis_zero, mul_one, basis_one, smul_basis_mul_Y, map_add, Finsupp.add_apply, map_smul, Finsupp.smul_apply, ← basis_zero, ← basis_one, - Basis.repr_self_apply, if_pos, if_neg, smul_eq_mul] + Basis.repr_self_apply, if_pos, one_ne_zero, if_false, smul_eq_mul] ring1 #align weierstrass_curve.coordinate_ring.norm_smul_basis WeierstrassCurve.CoordinateRing.norm_smul_basis diff --git a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean index 25f486dc08c1f..9425f8d75b786 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean @@ -63,7 +63,7 @@ instance : ReflectsIsomorphisms (N₁ : SimplicialObject C ⥤ Karoubi (ChainCom b := fun i => inv (f.app (op [n])) ≫ X.σ i } simp only [MorphComponents.id, ← id_φ, ← preComp_φ, preComp, ← postComp_φ, postComp, PInfty_f_naturality_assoc, IsIso.hom_inv_id_assoc, assoc, IsIso.inv_hom_id_assoc, - SimplicialObject.σ_naturality, h₁, h₂, h₃]⟩ + SimplicialObject.σ_naturality, h₁, h₂, h₃, and_self]⟩ theorem compatibility_N₂_N₁_karoubi : N₂ ⋙ (karoubiChainComplexEquivalence C ℕ).functor = diff --git a/Mathlib/Analysis/Analytic/Basic.lean b/Mathlib/Analysis/Analytic/Basic.lean index de6f459f47a87..8cc1543cc7e50 100644 --- a/Mathlib/Analysis/Analytic/Basic.lean +++ b/Mathlib/Analysis/Analytic/Basic.lean @@ -323,7 +323,7 @@ theorem radius_eq_top_iff_summable_norm (p : FormalMultilinearSeries 𝕜 E F) : /-- If the radius of `p` is positive, then `‖pₙ‖` grows at most geometrically. -/ theorem le_mul_pow_of_radius_pos (p : FormalMultilinearSeries 𝕜 E F) (h : 0 < p.radius) : - ∃ (C r : _) (hC : 0 < C) (hr : 0 < r), ∀ n, ‖p n‖ ≤ C * r ^ n := by + ∃ (C r : _) (hC : 0 < C) (_ : 0 < r), ∀ n, ‖p n‖ ≤ C * r ^ n := by rcases ENNReal.lt_iff_exists_nnreal_btwn.1 h with ⟨r, r0, rlt⟩ have rpos : 0 < (r : ℝ) := by simp [ENNReal.coe_pos.1 r0] rcases norm_le_div_pow_of_pos_of_lt_radius p rpos rlt with ⟨C, Cpos, hCp⟩ @@ -1220,7 +1220,7 @@ theorem changeOriginSeries_summable_aux₁ {r r' : ℝ≥0} (hr : (r + r' : ℝ Summable fun s : Σk l : ℕ, { s : Finset (Fin (k + l)) // s.card = l } => ‖p (s.1 + s.2.1)‖₊ * r ^ s.2.1 * r' ^ s.1 := by rw [← changeOriginIndexEquiv.symm.summable_iff] - dsimp only [(· ∘ ·), changeOriginIndexEquiv_symm_apply_fst, + dsimp only [Function.comp_def, changeOriginIndexEquiv_symm_apply_fst, changeOriginIndexEquiv_symm_apply_snd_fst] have : ∀ n : ℕ, HasSum (fun s : Finset (Fin n) => ‖p (n - s.card + s.card)‖₊ * r ^ s.card * r' ^ (n - s.card)) @@ -1342,7 +1342,7 @@ theorem changeOrigin_eval (h : (‖x‖₊ + ‖y‖₊ : ℝ≥0∞) < p.radius have : ∀ (m) (hm : n = m), p n (s.piecewise (fun _ => x) fun _ => y) = p m ((s.map (Fin.castIso hm).toEquiv.toEmbedding).piecewise (fun _ => x) fun _ => y) := by rintro m rfl - simp [Finset.piecewise] + simp (config := { unfoldPartialApp := true }) [Finset.piecewise] apply this #align formal_multilinear_series.change_origin_eval FormalMultilinearSeries.changeOrigin_eval diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index 0000e9c39a588..6a070f62ee43a 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -171,7 +171,7 @@ theorem applyComposition_update (p : FormalMultilinearSeries 𝕜 E F) {n : ℕ} theorem compContinuousLinearMap_applyComposition {n : ℕ} (p : FormalMultilinearSeries 𝕜 F G) (f : E →L[𝕜] F) (c : Composition n) (v : Fin n → E) : (p.compContinuousLinearMap f).applyComposition c v = p.applyComposition c (f ∘ v) := by - simp [applyComposition]; rfl + simp (config := {unfoldPartialApp := true}) [applyComposition]; rfl #align formal_multilinear_series.comp_continuous_linear_map_apply_composition FormalMultilinearSeries.compContinuousLinearMap_applyComposition end FormalMultilinearSeries @@ -1188,9 +1188,9 @@ theorem comp_assoc (r : FormalMultilinearSeries 𝕜 G H) (q : FormalMultilinear r c.1.length fun i : Fin c.1.length => q (c.2 i).length (applyComposition p (c.2 i) (v ∘ c.1.embedding i)) suffices ∑ c, f c = ∑ c, g c by - simpa only [FormalMultilinearSeries.comp, ContinuousMultilinearMap.sum_apply, - compAlongComposition_apply, Finset.sum_sigma', applyComposition, - ContinuousMultilinearMap.map_sum] + simpa (config := { unfoldPartialApp := true }) only [FormalMultilinearSeries.comp, + ContinuousMultilinearMap.sum_apply, compAlongComposition_apply, Finset.sum_sigma', + applyComposition, ContinuousMultilinearMap.map_sum] /- Now, we use `Composition.sigmaEquivSigmaPi n` to change variables in the second sum, and check that we get exactly the same sums. -/ rw [← (sigmaEquivSigmaPi n).sum_comp] diff --git a/Mathlib/Analysis/Analytic/RadiusLiminf.lean b/Mathlib/Analysis/Analytic/RadiusLiminf.lean index d93170f504ca4..97fe3033bab65 100644 --- a/Mathlib/Analysis/Analytic/RadiusLiminf.lean +++ b/Mathlib/Analysis/Analytic/RadiusLiminf.lean @@ -27,8 +27,6 @@ open Filter Asymptotics namespace FormalMultilinearSeries -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable (p : FormalMultilinearSeries 𝕜 E F) /-- The radius of a formal multilinear series is equal to diff --git a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean index 5683d99bb4787..276a43a76d8da 100644 --- a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean +++ b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean @@ -139,7 +139,7 @@ set_option linter.uppercaseLean3 false in theorem isEquivalent_const_iff_tendsto {c : β} (h : c ≠ 0) : u ~[l] const _ c ↔ Tendsto u l (𝓝 c) := by - simp_rw [IsEquivalent, const, isLittleO_const_iff h] + simp (config := { unfoldPartialApp := true }) only [IsEquivalent, const, isLittleO_const_iff h] constructor <;> intro h · have := h.sub (tendsto_const_nhds (a := -c)) simp only [Pi.sub_apply, sub_neg_eq_add, sub_add_cancel, zero_add] at this diff --git a/Mathlib/Analysis/BoxIntegral/Basic.lean b/Mathlib/Analysis/BoxIntegral/Basic.lean index 2fc0e3a4a2cbf..f3c8fe8c90b1a 100644 --- a/Mathlib/Analysis/BoxIntegral/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Basic.lean @@ -712,8 +712,6 @@ theorem integrable_of_continuousOn [CompleteSpace E] {I : Box ι} {f : ℝⁿ variable {l} -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- This is an auxiliary lemma used to prove two statements at once. Use one of the next two lemmas instead. -/ theorem HasIntegral.of_bRiemann_eq_false_of_forall_isLittleO (hl : l.bRiemann = false) diff --git a/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean b/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean index 3d452b1cdeacd..839454f631505 100644 --- a/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean +++ b/Mathlib/Analysis/BoxIntegral/DivergenceTheorem.lean @@ -139,8 +139,6 @@ theorem norm_volume_sub_integral_face_upper_sub_lower_smul_le {f : (Fin (n + 1) ac_rfl #align box_integral.norm_volume_sub_integral_face_upper_sub_lower_smul_le BoxIntegral.norm_volume_sub_integral_face_upper_sub_lower_smul_le -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- If `f : ℝⁿ⁺¹ → E` is differentiable on a closed rectangular box `I` with derivative `f'`, then the partial derivative `fun x ↦ f' x (Pi.single i 1)` is Henstock-Kurzweil integrable with integral equal to the difference of integrals of `f` over the faces `x i = I.upper i` and `x i = I.lower i`. diff --git a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean index eb538b796d39d..7b765bbcebcde 100644 --- a/Mathlib/Analysis/BoxIntegral/Partition/Split.lean +++ b/Mathlib/Analysis/BoxIntegral/Partition/Split.lean @@ -89,7 +89,8 @@ theorem splitLower_def [DecidableEq ι] {i x} (h : x ∈ Ioo (I.lower i) (I.uppe (forall_update_iff I.upper fun j y => I.lower j < y).2 ⟨h.1, fun j _ => I.lower_lt_upper _⟩) : I.splitLower i x = (⟨I.lower, update I.upper i x, h'⟩ : Box ι) := by - simp only [splitLower, mk'_eq_coe, min_eq_left h.2.le, update] + simp (config := { unfoldPartialApp := true }) only [splitLower, mk'_eq_coe, min_eq_left h.2.le, + update, and_self] #align box_integral.box.split_lower_def BoxIntegral.Box.splitLower_def /-- Given a box `I` and `x ∈ (I.lower i, I.upper i)`, the hyperplane `{y : ι → ℝ | y i = x}` splits @@ -130,7 +131,8 @@ theorem splitUpper_def [DecidableEq ι] {i x} (h : x ∈ Ioo (I.lower i) (I.uppe (forall_update_iff I.lower fun j y => y < I.upper j).2 ⟨h.2, fun j _ => I.lower_lt_upper _⟩) : I.splitUpper i x = (⟨update I.lower i x, I.upper, h'⟩ : Box ι) := by - simp only [splitUpper, mk'_eq_coe, max_eq_left h.1.le, update] + simp (config := { unfoldPartialApp := true }) only [splitUpper, mk'_eq_coe, max_eq_left h.1.le, + update, and_self] #align box_integral.box.split_upper_def BoxIntegral.Box.splitUpper_def theorem disjoint_splitLower_splitUpper (I : Box ι) (i : ι) (x : ℝ) : diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean b/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean index f10b97172c82f..a16ef72f4bfc0 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Convolution.lean @@ -104,8 +104,6 @@ theorem convolution_tendsto_right_of_continuous {ι} {φ : ι → ContDiffBump ( ((hg.tendsto x₀).comp tendsto_snd) tendsto_const_nhds #align cont_diff_bump.convolution_tendsto_right_of_continuous ContDiffBump.convolution_tendsto_right_of_continuous -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- If a function `g` is locally integrable, then the convolution `φ i * g` converges almost everywhere to `g` if `φ i` is a sequence of bump functions with support tending to `0`, provided that the ratio between the inner and outer radii of `φ i` remains bounded. -/ diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean index 4554af89c4cdc..e2e42167bd61d 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Normed.lean @@ -132,8 +132,6 @@ theorem integral_le_measure_closedBall : ∫ x, f x ∂μ ≤ (μ (closedBall c simp [measure_closedBall_lt_top] _ = (μ (closedBall c f.rOut)).toReal := by simp -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - theorem measure_closedBall_div_le_integral [IsAddHaarMeasure μ] (K : ℝ) (h : f.rOut ≤ K * f.rIn) : (μ (closedBall c f.rOut)).toReal / K ^ finrank ℝ E ≤ ∫ x, f x ∂μ := by have K_pos : 0 < K := by diff --git a/Mathlib/Analysis/Calculus/Monotone.lean b/Mathlib/Analysis/Calculus/Monotone.lean index c9cce07405def..f7db63b63df0d 100644 --- a/Mathlib/Analysis/Calculus/Monotone.lean +++ b/Mathlib/Analysis/Calculus/Monotone.lean @@ -36,8 +36,6 @@ open Set Filter Function Metric MeasureTheory MeasureTheory.Measure IsUnifLocDou open scoped Topology -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- If `(f y - f x) / (y - x)` converges to a limit as `y` tends to `x`, then the same goes if `y` is shifted a little bit, i.e., `f (y + (y-x)^2) - f x) / (y - x)` converges to the same limit. This lemma contains a slightly more general version of this statement (where one considers diff --git a/Mathlib/Analysis/Complex/AbsMax.lean b/Mathlib/Analysis/Complex/AbsMax.lean index 6d39b58088948..94d6aa004782f 100644 --- a/Mathlib/Analysis/Complex/AbsMax.lean +++ b/Mathlib/Analysis/Complex/AbsMax.lean @@ -424,4 +424,3 @@ theorem eqOn_of_eqOn_frontier {f g : E → F} {U : Set E} (hU : IsBounded U) (hf #align complex.eq_on_of_eq_on_frontier Complex.eqOn_of_eqOn_frontier end Complex - diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index 33d9e2e038b46..3ab5311dcc41d 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -22,9 +22,6 @@ The proof is based on the Cauchy integral formula for the derivative of an analy `Complex.deriv_eq_smul_circleIntegral`. -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open TopologicalSpace Metric Set Filter Asymptotics Function MeasureTheory Bornology open scoped Topology Filter NNReal Real diff --git a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean index f87fdcfaf2c56..2f23ea2ac1b66 100644 --- a/Mathlib/Analysis/Complex/LocallyUniformLimit.lean +++ b/Mathlib/Analysis/Complex/LocallyUniformLimit.lean @@ -27,8 +27,6 @@ open Set Metric MeasureTheory Filter Complex intervalIntegral open scoped Real Topology -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {E ι : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E] {U K : Set ℂ} {z : ℂ} {M r δ : ℝ} {φ : Filter ι} {F : ι → ℂ → E} {f g : ℂ → E} diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index 529b4f8ddf8ac..339beff93efff 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -48,7 +48,6 @@ open Set Function Filter Asymptotics Metric Complex open scoped Topology Filter Real local notation "expR" => Real.exp -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 namespace PhragmenLindelof diff --git a/Mathlib/Analysis/Complex/RemovableSingularity.lean b/Mathlib/Analysis/Complex/RemovableSingularity.lean index 87b921e7ce4b9..d71f6064978d0 100644 --- a/Mathlib/Analysis/Complex/RemovableSingularity.lean +++ b/Mathlib/Analysis/Complex/RemovableSingularity.lean @@ -23,8 +23,6 @@ open TopologicalSpace Metric Set Filter Asymptotics Function open scoped Topology Filter NNReal Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - universe u variable {E : Type u} [NormedAddCommGroup E] [NormedSpace ℂ E] [CompleteSpace E] diff --git a/Mathlib/Analysis/ConstantSpeed.lean b/Mathlib/Analysis/ConstantSpeed.lean index c939c34fae203..46d6c05471fc5 100644 --- a/Mathlib/Analysis/ConstantSpeed.lean +++ b/Mathlib/Analysis/ConstantSpeed.lean @@ -227,7 +227,7 @@ theorem unique_unit_speed_on_Icc_zero {s t : ℝ} (hs : 0 ≤ s) (ht : 0 ≤ t) rw [← φst] at hf convert unique_unit_speed φm hfφ hf ⟨le_rfl, hs⟩ using 1 have : φ 0 = 0 := by - have hm : 0 ∈ φ '' Icc 0 s := by simp only [mem_Icc, le_refl, ht, φst] + have hm : 0 ∈ φ '' Icc 0 s := by simp only [φst, ht, mem_Icc, le_refl, and_self] obtain ⟨x, xs, hx⟩ := hm apply le_antisymm ((φm ⟨le_rfl, hs⟩ xs xs.1).trans_eq hx) _ have := φst ▸ mapsTo_image φ (Icc 0 s) diff --git a/Mathlib/Analysis/Convex/Between.lean b/Mathlib/Analysis/Convex/Between.lean index 0d580c813e0d9..682c9ae1caaa0 100644 --- a/Mathlib/Analysis/Convex/Between.lean +++ b/Mathlib/Analysis/Convex/Between.lean @@ -610,7 +610,8 @@ theorem sbtw_of_sbtw_of_sbtw_of_mem_affineSpan_pair [NoZeroSMulDivisors R V] have hu : (Finset.univ : Finset (Fin 3)) = {i₁, i₂, i₃} := by clear h₁ h₂ h₁' h₂' -- Porting note: Originally `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at h₁₂ h₁₃ h₂₃ ⊢ have hp : p ∈ affineSpan R (Set.range t.points) := by have hle : line[R, t.points i₁, p₁] ≤ affineSpan R (Set.range t.points) := by refine' affineSpan_pair_le_of_mem_of_mem (mem_affineSpan R (Set.mem_range_self _)) _ diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean index 3ca7e9e3e1587..a7f2fd7df0d40 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Basic.lean @@ -34,9 +34,6 @@ For `p : ℝ`, prove that `fun x ↦ x ^ p` is concave when `0 ≤ p ≤ 1` and `0 < p < 1`. -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Real Set BigOperators NNReal /-- `Real.exp` is strictly convex on the whole real line. diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean index b1767ce282203..dddf0f51b9255 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Deriv.lean @@ -34,8 +34,6 @@ open Real Set open scoped BigOperators NNReal -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- `x^n`, `n : ℕ` is strictly convex on `[0, +∞)` for all `n` greater than `2`. -/ theorem strictConvexOn_pow {n : ℕ} (hn : 2 ≤ n) : StrictConvexOn ℝ (Ici 0) fun x : ℝ => x ^ n := by apply StrictMonoOn.strictConvexOn_of_deriv (convex_Ici _) (continuousOn_pow _) @@ -105,7 +103,7 @@ theorem strictConvexOn_zpow {m : ℤ} (hm₀ : m ≠ 0) (hm₁ : m ≠ 1) : rw [iter_deriv_zpow] refine' mul_pos _ (zpow_pos_of_pos hx _) norm_cast - refine' int_prod_range_pos (by simp only) fun hm => _ + refine' int_prod_range_pos (by decide) fun hm => _ rw [← Finset.coe_Ico] at hm norm_cast at hm fin_cases hm <;> simp_all -- Porting note: `simp_all` was `cc` diff --git a/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean b/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean index 1d79ae3eaeeba..ff3014479bd3a 100644 --- a/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean +++ b/Mathlib/Analysis/Convex/SpecificFunctions/Pow.lean @@ -27,8 +27,6 @@ requires slightly less imports. * Prove convexity for negative powers. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Set namespace NNReal diff --git a/Mathlib/Analysis/Convex/Strong.lean b/Mathlib/Analysis/Convex/Strong.lean index 29196ea1f17b7..19a08901c13cc 100644 --- a/Mathlib/Analysis/Convex/Strong.lean +++ b/Mathlib/Analysis/Convex/Strong.lean @@ -22,8 +22,6 @@ If `E` is an inner product space, this is equivalent to `x ↦ f x - m / 2 * ‖ Prove derivative properties of strongly convex functions. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) - open Real variable {E : Type*} [NormedAddCommGroup E] diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index 56554a96931fa..698c702afe91b 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1013,7 +1013,8 @@ theorem convolution_assoc (hL : ∀ (x : E) (y : E') (z : E''), L₂ (L x y) z = rw [← h3] at this convert this.comp_measurable (measurable_sub.prod_mk measurable_snd) ext ⟨x, y⟩ - simp_rw [uncurry, Function.comp_apply, sub_sub_sub_cancel_right] + simp (config := { unfoldPartialApp := true }) only [uncurry, Function.comp_apply, + sub_sub_sub_cancel_right] simp_rw [integrable_prod_iff' h_meas] refine' ⟨((quasiMeasurePreserving_sub_left_of_right_invariant ν x₀).ae hgk).mono fun t ht => (L₃ (f t)).integrable_comp <| ht.ofNorm L₄ hg hk, _⟩ diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 69785c8e214d1..0c04e70116c94 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -60,8 +60,6 @@ The implementation of the seminorms is taken almost literally from `ContinuousLi Schwartz space, tempered distributions -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open scoped BigOperators Nat diff --git a/Mathlib/Analysis/Fourier/AddCircle.lean b/Mathlib/Analysis/Fourier/AddCircle.lean index 871f57d1dbe6f..920eba649c205 100644 --- a/Mathlib/Analysis/Fourier/AddCircle.lean +++ b/Mathlib/Analysis/Fourier/AddCircle.lean @@ -63,8 +63,6 @@ converges to `f` in the uniform-convergence topology of `C(AddCircle T, ℂ)`. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped ENNReal ComplexConjugate Real open TopologicalSpace ContinuousMap MeasureTheory MeasureTheory.Measure Algebra Submodule Set @@ -148,12 +146,12 @@ section Monomials /-- The family of exponential monomials `fun x => exp (2 π i n x / T)`, parametrized by `n : ℤ` and considered as bundled continuous maps from `ℝ / ℤ • T` to `ℂ`. -/ def fourier (n : ℤ) : C(AddCircle T, ℂ) where - toFun x := toCircle (n • x) + toFun x := toCircle (n • x :) continuous_toFun := continuous_induced_dom.comp <| continuous_toCircle.comp <| continuous_zsmul _ #align fourier fourier @[simp] -theorem fourier_apply {n : ℤ} {x : AddCircle T} : fourier n x = toCircle (n • x) := +theorem fourier_apply {n : ℤ} {x : AddCircle T} : fourier n x = toCircle (n • x :) := rfl #align fourier_apply fourier_apply @@ -169,7 +167,7 @@ theorem fourier_coe_apply {n : ℤ} {x : ℝ} : @[simp] theorem fourier_coe_apply' {n : ℤ} {x : ℝ} : - toCircle (n • (x : AddCircle T)) = Complex.exp (2 * π * Complex.I * n * x / T) := by + toCircle (n • (x : AddCircle T) :) = Complex.exp (2 * π * Complex.I * n * x / T) := by rw [← fourier_apply]; exact fourier_coe_apply -- @[simp] -- Porting note: simp normal form is `fourier_zero'` @@ -214,7 +212,7 @@ theorem fourier_add {m n : ℤ} {x : AddCircle T} : fourier (m+n) x = fourier m @[simp] theorem fourier_add' {m n : ℤ} {x : AddCircle T} : - toCircle ((m + n) • x) = fourier m x * fourier n x := by + toCircle ((m + n) • x :) = fourier m x * fourier n x := by rw [← fourier_apply]; exact fourier_add theorem fourier_norm [Fact (0 < T)] (n : ℤ) : ‖@fourier T n‖ = 1 := by @@ -367,7 +365,7 @@ theorem fourierCoeff_eq_intervalIntegral (f : AddCircle T → E) (n : ℤ) (a : #align fourier_coeff_eq_interval_integral fourierCoeff_eq_intervalIntegral theorem fourierCoeff.const_smul (f : AddCircle T → E) (c : ℂ) (n : ℤ) : - fourierCoeff (c • f) n = c • fourierCoeff f n := by + fourierCoeff (c • f :) n = c • fourierCoeff f n := by simp_rw [fourierCoeff, Pi.smul_apply, ← smul_assoc, smul_eq_mul, mul_comm, ← smul_eq_mul, smul_assoc, integral_smul] #align fourier_coeff.const_smul fourierCoeff.const_smul diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index ce581cdfc3f8c..8774809388597 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -44,8 +44,6 @@ equivalence to an inner-product space. reformulations explicitly using the Fourier integral. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open MeasureTheory Filter Complex Set FiniteDimensional @@ -72,9 +70,7 @@ theorem fourier_integrand_integrable (w : V) : variable [CompleteSpace E] --- Porting note: binary operations appear not to work? --- local notation "i" => fun w => (1 / (2 * ‖w‖ ^ 2)) • w -local notation "i" => fun (w : V) => HDiv.hDiv (1 : ℝ) (HMul.hMul (2 : ℝ) (HPow.hPow ‖w‖ 2)) • w +local notation3 "i" => fun (w : V) => (1 / (2 * ‖w‖ ^ 2) : ℝ) • w /-- Shifting `f` by `(1 / (2 * ‖w‖ ^ 2)) • w` negates the integral in the Riemann-Lebesgue lemma. -/ theorem fourier_integral_half_period_translate {w : V} (hw : w ≠ 0) : diff --git a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean index 60fd1ad380164..1283a1f2401f9 100644 --- a/Mathlib/Analysis/InnerProductSpace/Adjoint.lean +++ b/Mathlib/Analysis/InnerProductSpace/Adjoint.lean @@ -40,8 +40,6 @@ adjoint -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open IsROrC diff --git a/Mathlib/Analysis/InnerProductSpace/Calculus.lean b/Mathlib/Analysis/InnerProductSpace/Calculus.lean index 37023816b8118..edaac53116266 100644 --- a/Mathlib/Analysis/InnerProductSpace/Calculus.lean +++ b/Mathlib/Analysis/InnerProductSpace/Calculus.lean @@ -26,8 +26,6 @@ and from the equivalence of norms in finite dimensions. The last part of the file should be generalized to `PiLp`. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open IsROrC Real Filter @@ -402,7 +400,7 @@ open Metric hiding mem_nhds_iff variable {n : ℕ∞} {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] theorem LocalHomeomorph.contDiff_univUnitBall : ContDiff ℝ n (univUnitBall : E → E) := by - suffices ContDiff ℝ n fun x : E => ((1 : ℝ) + ‖x‖ ^ 2).sqrt⁻¹ from this.smul contDiff_id + suffices ContDiff ℝ n fun x : E => (1 + ‖x‖ ^ 2 : ℝ).sqrt⁻¹ from this.smul contDiff_id have h : ∀ x : E, (0 : ℝ) < (1 : ℝ) + ‖x‖ ^ 2 := fun x => by positivity refine' ContDiff.inv _ fun x => Real.sqrt_ne_zero'.mpr (h x) exact (contDiff_const.add <| contDiff_norm_sq ℝ).sqrt fun x => (h x).ne' @@ -410,7 +408,7 @@ theorem LocalHomeomorph.contDiff_univUnitBall : ContDiff ℝ n (univUnitBall : E theorem LocalHomeomorph.contDiffOn_univUnitBall_symm : ContDiffOn ℝ n univUnitBall.symm (ball (0 : E) 1) := fun y hy ↦ by apply ContDiffAt.contDiffWithinAt - suffices ContDiffAt ℝ n (fun y : E => ((1 : ℝ) - ‖y‖ ^ 2).sqrt⁻¹) y from this.smul contDiffAt_id + suffices ContDiffAt ℝ n (fun y : E => (1 - ‖y‖ ^ 2 : ℝ).sqrt⁻¹) y from this.smul contDiffAt_id have h : (0 : ℝ) < (1 : ℝ) - ‖(y : E)‖ ^ 2 := by rwa [mem_ball_zero_iff, ← _root_.abs_one, ← abs_norm, ← sq_lt_sq, one_pow, ← sub_pos] at hy refine' ContDiffAt.inv _ (Real.sqrt_ne_zero'.mpr h) diff --git a/Mathlib/Analysis/InnerProductSpace/PiL2.lean b/Mathlib/Analysis/InnerProductSpace/PiL2.lean index fa09d08078de4..46f5678874a71 100644 --- a/Mathlib/Analysis/InnerProductSpace/PiL2.lean +++ b/Mathlib/Analysis/InnerProductSpace/PiL2.lean @@ -53,8 +53,6 @@ For consequences in infinite dimension (Hilbert bases, etc.), see the file -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - set_option linter.uppercaseLean3 false open Real Set Filter IsROrC Submodule Function BigOperators Uniformity Topology NNReal ENNReal diff --git a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean index db7b8956d2f7e..d88886429c42b 100644 --- a/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean +++ b/Mathlib/Analysis/InnerProductSpace/Rayleigh.lean @@ -117,8 +117,6 @@ theorem _root_.LinearMap.IsSymmetric.hasStrictFDerivAt_reApplyInnerSelf {T : F variable [CompleteSpace F] {T : F →L[ℝ] F} -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - theorem linearly_dependent_of_isLocalExtrOn (hT : IsSelfAdjoint T) {x₀ : F} (hextr : IsLocalExtrOn T.reApplyInnerSelf (sphere (0 : F) ‖x₀‖) x₀) : ∃ a b : ℝ, (a, b) ≠ 0 ∧ a • x₀ + b • T x₀ = 0 := by @@ -242,7 +240,7 @@ namespace IsSymmetric /-- The supremum of the Rayleigh quotient of a symmetric operator `T` on a nontrivial finite-dimensional vector space is an eigenvalue for that operator. -/ theorem hasEigenvalue_iSup_of_finiteDimensional (hT : T.IsSymmetric) : - HasEigenvalue T ↑(⨆ x : { x : E // x ≠ 0 }, IsROrC.re ⟪T x, x⟫ / ‖(x : E)‖ ^ 2) := by + HasEigenvalue T ↑(⨆ x : { x : E // x ≠ 0 }, IsROrC.re ⟪T x, x⟫ / ‖(x : E)‖ ^ 2 : ℝ) := by haveI := FiniteDimensional.proper_isROrC 𝕜 E let T' := hT.toSelfAdjoint obtain ⟨x, hx⟩ : ∃ x : E, x ≠ 0 := exists_ne 0 @@ -262,7 +260,7 @@ theorem hasEigenvalue_iSup_of_finiteDimensional (hT : T.IsSymmetric) : /-- The infimum of the Rayleigh quotient of a symmetric operator `T` on a nontrivial finite-dimensional vector space is an eigenvalue for that operator. -/ theorem hasEigenvalue_iInf_of_finiteDimensional (hT : T.IsSymmetric) : - HasEigenvalue T ↑(⨅ x : { x : E // x ≠ 0 }, IsROrC.re ⟪T x, x⟫ / ‖(x : E)‖ ^ 2) := by + HasEigenvalue T ↑(⨅ x : { x : E // x ≠ 0 }, IsROrC.re ⟪T x, x⟫ / ‖(x : E)‖ ^ 2 : ℝ) := by haveI := FiniteDimensional.proper_isROrC 𝕜 E let T' := hT.toSelfAdjoint obtain ⟨x, hx⟩ : ∃ x : E, x ≠ 0 := exists_ne 0 diff --git a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean index 2e089c0552cec..107dfd1cb317c 100644 --- a/Mathlib/Analysis/InnerProductSpace/Spectrum.lean +++ b/Mathlib/Analysis/InnerProductSpace/Spectrum.lean @@ -276,8 +276,6 @@ end LinearMap section Nonneg -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - @[simp] theorem inner_product_apply_eigenvector {μ : 𝕜} {v : E} {T : E →ₗ[𝕜] E} (h : v ∈ Module.End.eigenspace T μ) : ⟪v, T v⟫ = μ * (‖v‖ : 𝕜) ^ 2 := by diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index c6a4653f6226e..fa1516ea40c93 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -74,8 +74,6 @@ open scoped RealInnerProductSpace ComplexConjugate open FiniteDimensional -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - lemma FiniteDimensional.finiteDimensional_of_fact_finrank_eq_two {K V : Type*} [DivisionRing K] [AddCommGroup V] [Module K V] [Fact (finrank K V = 2)] : FiniteDimensional K V := fact_finiteDimensional_of_finrank_eq_succ 1 @@ -527,7 +525,7 @@ theorem kahler_neg_orientation (x y : E) : (-o).kahler x y = conj (o.kahler x y) #align orientation.kahler_neg_orientation Orientation.kahler_neg_orientation theorem kahler_mul (a x y : E) : o.kahler x a * o.kahler a y = ‖a‖ ^ 2 * o.kahler x y := by - trans (↑(‖a‖ ^ 2) : ℂ) * o.kahler x y + trans ((‖a‖ ^ 2 :) : ℂ) * o.kahler x y · ext · simp only [o.kahler_apply_apply, Complex.add_im, Complex.add_re, Complex.I_im, Complex.I_re, Complex.mul_im, Complex.mul_re, Complex.ofReal_im, Complex.ofReal_re, Complex.real_smul] diff --git a/Mathlib/Analysis/InnerProductSpace/l2Space.lean b/Mathlib/Analysis/InnerProductSpace/l2Space.lean index b88da47ac2fea..d16a7abf0a8b3 100644 --- a/Mathlib/Analysis/InnerProductSpace/l2Space.lean +++ b/Mathlib/Analysis/InnerProductSpace/l2Space.lean @@ -88,8 +88,6 @@ open scoped BigOperators NNReal ENNReal Classical ComplexConjugate Topology noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {ι : Type*} variable {𝕜 : Type*} [IsROrC 𝕜] {E : Type*} diff --git a/Mathlib/Analysis/Matrix.lean b/Mathlib/Analysis/Matrix.lean index e027e59356c01..ba4298a004005 100644 --- a/Mathlib/Analysis/Matrix.lean +++ b/Mathlib/Analysis/Matrix.lean @@ -43,9 +43,6 @@ These are not declared as instances because there are several natural choices fo of a matrix. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - - noncomputable section open scoped BigOperators NNReal Matrix diff --git a/Mathlib/Analysis/MeanInequalities.lean b/Mathlib/Analysis/MeanInequalities.lean index fbb89e827f829..a43245b617b7d 100644 --- a/Mathlib/Analysis/MeanInequalities.lean +++ b/Mathlib/Analysis/MeanInequalities.lean @@ -98,8 +98,6 @@ set_option linter.uppercaseLean3 false noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {ι : Type u} (s : Finset ι) section GeomMeanLEArithMean diff --git a/Mathlib/Analysis/MeanInequalitiesPow.lean b/Mathlib/Analysis/MeanInequalitiesPow.lean index 81e54cc14d9b5..668db2fec9f11 100644 --- a/Mathlib/Analysis/MeanInequalitiesPow.lean +++ b/Mathlib/Analysis/MeanInequalitiesPow.lean @@ -51,8 +51,6 @@ open Classical BigOperators NNReal ENNReal noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {ι : Type u} (s : Finset ι) namespace Real diff --git a/Mathlib/Analysis/MellinTransform.lean b/Mathlib/Analysis/MellinTransform.lean index a0d87486204b9..1aefad2b113a7 100644 --- a/Mathlib/Analysis/MellinTransform.lean +++ b/Mathlib/Analysis/MellinTransform.lean @@ -150,7 +150,7 @@ theorem mellin_comp_mul_left (f : ℝ → E) (s : ℂ) {a : ℝ} (ha : 0 < a) : rw [Ne.def, cpow_eq_zero_iff, ofReal_eq_zero, not_and_or] exact Or.inl ha.ne' rw [set_integral_congr measurableSet_Ioi this, integral_smul, - integral_comp_mul_left_Ioi (fun u ↦ ↑u ^ (s - 1) • f u) _ ha, + integral_comp_mul_left_Ioi (fun u ↦ (u : ℂ) ^ (s - 1) • f u) _ ha, mul_zero, ← Complex.coe_smul, ← mul_smul, sub_eq_add_neg, cpow_add _ _ (ofReal_ne_zero.mpr ha.ne'), cpow_one, abs_of_pos (inv_pos.mpr ha), ofReal_inv, mul_assoc, mul_comm, inv_mul_cancel_right₀ (ofReal_ne_zero.mpr ha.ne')] @@ -386,7 +386,7 @@ theorem mellin_hasDerivAt_of_isBigO_rpow [CompleteSpace E] [NormedSpace ℂ E] { rwa [sub_re, sub_le_iff_le_add, ← sub_le_iff_le_add'] at hz' have h5 : IntegrableOn bound (Ioi 0) := by simp_rw [add_mul, mul_assoc] - suffices ∀ {j : ℝ} (hj : b < j) (hj' : j < a), + suffices ∀ {j : ℝ}, b < j → j < a → IntegrableOn (fun t : ℝ => t ^ (j - 1) * (|log t| * ‖f t‖)) (Ioi 0) volume by refine' Integrable.add (this _ _) (this _ _) all_goals linarith diff --git a/Mathlib/Analysis/NormedSpace/DualNumber.lean b/Mathlib/Analysis/NormedSpace/DualNumber.lean index 76d8d8fe8257c..2c89f6dff9153 100644 --- a/Mathlib/Analysis/NormedSpace/DualNumber.lean +++ b/Mathlib/Analysis/NormedSpace/DualNumber.lean @@ -37,7 +37,7 @@ theorem exp_eps : exp 𝕜 (eps : DualNumber R) = 1 + eps := @[simp] theorem exp_smul_eps (r : R) : exp 𝕜 (r • eps : DualNumber R) = 1 + r • eps := by - rw [eps, ← inr_smul, exp_inr, Nat.cast_one] + rw [eps, ← inr_smul, exp_inr] #align dual_number.exp_smul_eps DualNumber.exp_smul_eps end DualNumber diff --git a/Mathlib/Analysis/NormedSpace/Exponential.lean b/Mathlib/Analysis/NormedSpace/Exponential.lean index c9276ecab9c9e..96184f66ef662 100644 --- a/Mathlib/Analysis/NormedSpace/Exponential.lean +++ b/Mathlib/Analysis/NormedSpace/Exponential.lean @@ -312,7 +312,7 @@ theorem map_exp_of_mem_ball {F} [RingHomClass F 𝔸 𝔹] (f : F) (hf : Continu (hx : x ∈ EMetric.ball (0 : 𝔸) (expSeries 𝕂 𝔸).radius) : f (exp 𝕂 x) = exp 𝕂 (f x) := by rw [exp_eq_tsum, exp_eq_tsum] refine' ((expSeries_summable_of_mem_ball' _ hx).hasSum.map f hf).tsum_eq.symm.trans _ - dsimp only [Function.comp] + dsimp only [Function.comp_def] simp_rw [map_inv_nat_cast_smul f 𝕂 𝕂, map_pow] #align map_exp_of_mem_ball map_exp_of_mem_ball diff --git a/Mathlib/Analysis/NormedSpace/Multilinear.lean b/Mathlib/Analysis/NormedSpace/Multilinear.lean index 57e58c75ff132..963075e8972e7 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear.lean @@ -214,7 +214,7 @@ theorem norm_image_sub_le_of_bound {C : ℝ} (hC : 0 ≤ C) (H : ∀ m, ‖f m calc ‖f m₁ - f m₂‖ ≤ C * ∑ i, ∏ j, if j = i then ‖m₁ i - m₂ i‖ else max ‖m₁ j‖ ‖m₂ j‖ := f.norm_image_sub_le_of_bound' hC H m₁ m₂ - _ ≤ C * ∑ i, ‖m₁ - m₂‖ * max ‖m₁‖ ‖m₂‖ ^ (Fintype.card ι - 1) := by gcongr; apply A + _ ≤ C * ∑ _i, ‖m₁ - m₂‖ * max ‖m₁‖ ‖m₂‖ ^ (Fintype.card ι - 1) := by gcongr; apply A _ = C * Fintype.card ι * max ‖m₁‖ ‖m₂‖ ^ (Fintype.card ι - 1) * ‖m₁ - m₂‖ := by rw [sum_const, card_univ, nsmul_eq_mul] ring diff --git a/Mathlib/Analysis/NormedSpace/PiLp.lean b/Mathlib/Analysis/NormedSpace/PiLp.lean index 886c51fe944b8..18dd58c4b29fb 100644 --- a/Mathlib/Analysis/NormedSpace/PiLp.lean +++ b/Mathlib/Analysis/NormedSpace/PiLp.lean @@ -64,8 +64,6 @@ We also set up the theory for `PseudoEMetricSpace` and `PseudoMetricSpace`. set_option linter.uppercaseLean3 false -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal ENNReal noncomputable section @@ -574,7 +572,7 @@ theorem norm_eq_of_nat {p : ℝ≥0∞} [Fact (1 ≤ p)] {β : ι → Type*} #align pi_Lp.norm_eq_of_nat PiLp.norm_eq_of_nat theorem norm_eq_of_L2 {β : ι → Type*} [∀ i, SeminormedAddCommGroup (β i)] (x : PiLp 2 β) : - ‖x‖ = sqrt (∑ i : ι, ‖x i‖ ^ 2) := by + ‖x‖ = Real.sqrt (∑ i : ι, ‖x i‖ ^ 2) := by rw [norm_eq_of_nat 2 (by norm_cast) _] -- Porting note: was `convert` rw [Real.sqrt_eq_rpow] norm_cast diff --git a/Mathlib/Analysis/NormedSpace/ProdLp.lean b/Mathlib/Analysis/NormedSpace/ProdLp.lean index bda3574465225..8936e016d274f 100644 --- a/Mathlib/Analysis/NormedSpace/ProdLp.lean +++ b/Mathlib/Analysis/NormedSpace/ProdLp.lean @@ -36,8 +36,6 @@ This files is a straight-forward adaption of `Mathlib.Analysis.NormedSpace.PiLp` -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Real Set Filter IsROrC Bornology BigOperators Uniformity Topology NNReal ENNReal noncomputable section @@ -382,7 +380,7 @@ def prodPseudoMetricAux [PseudoMetricSpace α] [PseudoMetricSpace β] : ← PseudoMetricSpace.edist_dist] exact le_sup_right · refine ENNReal.toReal_le_of_le_ofReal ?_ ?_ - · simp only [ge_iff_le, le_sup_iff, dist_nonneg] + · simp only [ge_iff_le, le_sup_iff, dist_nonneg, or_self] · simp [edist, PseudoMetricSpace.edist_dist, ENNReal.ofReal_le_ofReal] · have h1 : edist f.fst g.fst ^ p.toReal ≠ ⊤ := ENNReal.rpow_ne_top_of_nonneg (zero_le_one.trans h) (edist_ne_top _ _) @@ -434,7 +432,7 @@ theorem prod_antilipschitzWith_equiv_aux [PseudoEMetricSpace α] [PseudoEMetricS simp [edist] · refine ENNReal.rpow_le_rpow ?_ (le_of_lt pos) simp [edist] - _ = (2 : ℝ≥0) ^ (1 / p.toReal) * edist (WithLp.equiv p _ x) (WithLp.equiv p _ y) := by + _ = (2 ^ (1 / p.toReal) : ℝ≥0) * edist (WithLp.equiv p _ x) (WithLp.equiv p _ y) := by simp only [← two_mul, ENNReal.mul_rpow_of_nonneg _ _ nonneg, ← ENNReal.rpow_mul, cancel, ENNReal.rpow_one, ← ENNReal.coe_rpow_of_nonneg _ nonneg, coe_ofNat] @@ -632,7 +630,8 @@ theorem prod_nnnorm_eq_sup (f : WithLp ∞ (α × β)) : ‖f‖₊ = ‖f.fst ext norm_cast -theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : ‖x‖ = sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by +theorem prod_norm_eq_of_L2 (x : WithLp 2 (α × β)) : + ‖x‖ = Real.sqrt (‖x.fst‖ ^ 2 + ‖x.snd‖ ^ 2) := by rw [prod_norm_eq_of_nat 2 (by norm_cast) _, Real.sqrt_eq_rpow] norm_cast diff --git a/Mathlib/Analysis/NormedSpace/QuaternionExponential.lean b/Mathlib/Analysis/NormedSpace/QuaternionExponential.lean index 3dc379305b4b4..be010371557ef 100644 --- a/Mathlib/Analysis/NormedSpace/QuaternionExponential.lean +++ b/Mathlib/Analysis/NormedSpace/QuaternionExponential.lean @@ -34,8 +34,6 @@ theorem exp_coe (r : ℝ) : exp ℝ (r : ℍ[ℝ]) = ↑(exp ℝ r) := (map_exp ℝ (algebraMap ℝ ℍ[ℝ]) (continuous_algebraMap _ _) _).symm #align quaternion.exp_coe Quaternion.exp_coe -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- Auxiliary result; if the power series corresponding to `Real.cos` and `Real.sin` evaluated at `‖q‖` tend to `c` and `s`, then the exponential series tends to `c + (s / ‖q‖)`. -/ theorem hasSum_expSeries_of_imaginary {q : Quaternion ℝ} (hq : q.re = 0) {c s : ℝ} @@ -61,7 +59,7 @@ theorem hasSum_expSeries_of_imaginary {q : Quaternion ℝ} (hq : q.re = 0) {c s ext n : 1 letI k : ℝ := ↑(2 * n)! calc - k⁻¹ • q ^ (2 * n) = k⁻¹ • (-normSq q) ^ n := by rw [pow_mul, hq2]; norm_cast + k⁻¹ • q ^ (2 * n) = k⁻¹ • (-normSq q) ^ n := by rw [pow_mul, hq2] _ = k⁻¹ • ↑((-1 : ℝ) ^ n * ‖q‖ ^ (2 * n)) := ?_ _ = ↑((-1 : ℝ) ^ n * ‖q‖ ^ (2 * n) / k) := ?_ · congr 1 @@ -77,11 +75,11 @@ theorem hasSum_expSeries_of_imaginary {q : Quaternion ℝ} (hq : q.re = 0) {c s calc k⁻¹ • q ^ (2 * n + 1) = k⁻¹ • ((-normSq q) ^ n * q) := by rw [pow_succ', pow_mul, hq2] - norm_cast _ = k⁻¹ • ((-1 : ℝ) ^ n * ‖q‖ ^ (2 * n)) • q := ?_ _ = ((-1 : ℝ) ^ n * ‖q‖ ^ (2 * n + 1) / k / ‖q‖) • q := ?_ · congr 1 rw [neg_pow, normSq_eq_norm_mul_self, pow_mul, sq, ← coe_mul_eq_smul] + norm_cast · rw [smul_smul] congr 1 simp_rw [pow_succ', mul_div_assoc, div_div_cancel_left' hqn] diff --git a/Mathlib/Analysis/NormedSpace/Spectrum.lean b/Mathlib/Analysis/NormedSpace/Spectrum.lean index b129a6bb06018..09c8ded4feb14 100644 --- a/Mathlib/Analysis/NormedSpace/Spectrum.lean +++ b/Mathlib/Analysis/NormedSpace/Spectrum.lean @@ -230,8 +230,9 @@ theorem eventually_isUnit_resolvent (a : A) : ∀ᶠ z in cobounded 𝕜, IsUnit theorem resolvent_isBigO_inv (a : A) : resolvent a =O[cobounded 𝕜] Inv.inv := have h : (fun z ↦ resolvent (z⁻¹ • a) (1 : 𝕜)) =O[cobounded 𝕜] (fun _ ↦ (1 : ℝ)) := by - simpa [Function.comp, resolvent] using (NormedRing.inverse_one_sub_norm (R := A)).comp_tendsto - (by simpa using (tendsto_inv₀_cobounded (α := 𝕜)).smul_const a) + simpa [Function.comp_def, resolvent] using + (NormedRing.inverse_one_sub_norm (R := A)).comp_tendsto + (by simpa using (tendsto_inv₀_cobounded (α := 𝕜)).smul_const a) calc resolvent a =ᶠ[cobounded 𝕜] fun z ↦ z⁻¹ • resolvent (z⁻¹ • a) (1 : 𝕜) := by filter_upwards [isBounded_singleton (x := 0)] with z hz diff --git a/Mathlib/Analysis/NormedSpace/Star/Matrix.lean b/Mathlib/Analysis/NormedSpace/Star/Matrix.lean index 46fc975d491c0..5480cdcb90ba5 100644 --- a/Mathlib/Analysis/NormedSpace/Star/Matrix.lean +++ b/Mathlib/Analysis/NormedSpace/Star/Matrix.lean @@ -19,8 +19,6 @@ This file collects facts about the unitary matrices over `𝕜` (either `ℝ` or open scoped BigOperators Matrix -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {𝕜 m n E : Type*} section EntrywiseSupNorm @@ -41,7 +39,7 @@ theorem entry_norm_bound_of_unitary {U : Matrix n n 𝕜} (hU : U ∈ Matrix.uni use j simp only [eq_self_iff_true, Finset.mem_univ_val, and_self_iff, sq_eq_sq] -- The L2 norm of a row is a diagonal entry of U * Uᴴ - have diag_eq_norm_sum : (U * Uᴴ) i i = ∑ x : n, ‖U i x‖ ^ 2 := by + have diag_eq_norm_sum : (U * Uᴴ) i i = (∑ x : n, ‖U i x‖ ^ 2 : ℝ) := by simp only [Matrix.mul_apply, Matrix.conjTranspose_apply, ← starRingEnd_apply, IsROrC.mul_conj, IsROrC.normSq_eq_def', IsROrC.ofReal_pow]; norm_cast -- The L2 norm of a row is a diagonal entry of U * Uᴴ, real part diff --git a/Mathlib/Analysis/NormedSpace/lpSpace.lean b/Mathlib/Analysis/NormedSpace/lpSpace.lean index 33f0d126ddd3b..a22c013ff14ab 100644 --- a/Mathlib/Analysis/NormedSpace/lpSpace.lean +++ b/Mathlib/Analysis/NormedSpace/lpSpace.lean @@ -57,9 +57,6 @@ say that `‖-f‖ = ‖f‖`, instead of the non-working `f.norm_neg`. set_option autoImplicit true - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open scoped NNReal ENNReal BigOperators Function diff --git a/Mathlib/Analysis/ODE/PicardLindelof.lean b/Mathlib/Analysis/ODE/PicardLindelof.lean index a5b36dfb0cf3f..3916ef0f0ec84 100644 --- a/Mathlib/Analysis/ODE/PicardLindelof.lean +++ b/Mathlib/Analysis/ODE/PicardLindelof.lean @@ -34,8 +34,6 @@ related theorems in `Mathlib/Analysis/ODE/Gronwall.lean`. differential equation -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Filter Function Set Metric TopologicalSpace intervalIntegral MeasureTheory open MeasureTheory.MeasureSpace (volume) open scoped Filter Topology NNReal ENNReal Nat Interval diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 2b47876fb71f3..ad67f66ee16ed 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -27,9 +27,6 @@ test once we need it. p-series, Cauchy condensation test -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Filter open BigOperators ENNReal NNReal Topology @@ -77,8 +74,11 @@ theorem sum_condensed_le' (hf : ∀ ⦃m n⦄, 1 < m → m ≤ n → f n ≤ f m exacts [add_le_add ihn this, (add_le_add_right n.one_le_two_pow _ : 1 + 1 ≤ 2 ^ n + 1), add_le_add_right (Nat.pow_le_pow_of_le_right zero_lt_two n.le_succ) _] - have : ∀ k ∈ Ico (2 ^ n + 1) (2 ^ (n + 1) + 1), f (2 ^ (n + 1)) ≤ f k := fun k hk => - hf (n.one_le_two_pow.trans_lt <| (Nat.lt_succ_of_le le_rfl).trans_le (mem_Ico.mp hk).1) + have : ∀ k ∈ Ico (2 ^ n + 1) (2 ^ (n + 1) + 1), f (2 ^ (n + 1)) ≤ f k := by + -- Note(kmill): was `fun k hk => ...` but `mem_Ico.mp hk` was elaborating with some + -- delayed assignment metavariables that weren't resolved in time. `intro` fixes this. + intro k hk + exact hf (n.one_le_two_pow.trans_lt <| (Nat.lt_succ_of_le le_rfl).trans_le (mem_Ico.mp hk).1) (Nat.le_of_lt_succ <| (mem_Ico.mp hk).2) convert sum_le_sum this simp [pow_succ, two_mul] @@ -313,7 +313,7 @@ theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : positivity #align sum_Ioc_inv_sq_le_sub sum_Ioc_inv_sq_le_sub -theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i in Ioo k n, ((i : α) ^ 2)⁻¹) ≤ 2 / (k + 1) := +theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i in Ioo k n, (i ^ 2 : α)⁻¹) ≤ 2 / (k + 1) := calc (∑ i in Ioo k n, ((i : α) ^ 2)⁻¹) ≤ ∑ i in Ioc k (max (k + 1) n), ((i : α) ^ 2)⁻¹ := by apply sum_le_sum_of_subset_of_nonneg diff --git a/Mathlib/Analysis/SpecialFunctions/CompareExp.lean b/Mathlib/Analysis/SpecialFunctions/CompareExp.lean index bfdf07ab63ca4..c1e33c8b30af2 100644 --- a/Mathlib/Analysis/SpecialFunctions/CompareExp.lean +++ b/Mathlib/Analysis/SpecialFunctions/CompareExp.lean @@ -32,8 +32,6 @@ stronger assumptions (e.g., `im z` is bounded from below and from above) are not open Asymptotics Filter Function open scoped Topology -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Complex /-- We say that `l : Filter ℂ` is an *exponential comparison filter* if the real part tends to diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean index c2eb5253bcac3..c799a05737608 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Basic.lean @@ -209,7 +209,7 @@ theorem partialGamma_add_one {s : ℂ} (hs : 0 < s.re) {X : ℝ} (hX : 0 ≤ X) intro x hx have d1 : HasDerivAt (fun y : ℝ => (-y).exp) (-(-x).exp) x := by simpa using (hasDerivAt_neg x).exp - have d2 : HasDerivAt (fun y : ℝ => ↑y ^ s) (s * x ^ (s - 1)) x := by + have d2 : HasDerivAt (fun y : ℝ => (y : ℂ) ^ s) (s * x ^ (s - 1)) x := by have t := @HasDerivAt.cpow_const _ _ _ s (hasDerivAt_id ↑x) ?_ simpa only [mul_one] using t.comp_ofReal simpa only [id.def, ofReal_re, ofReal_im, Ne.def, eq_self_iff_true, not_true, or_false_iff, diff --git a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean index ce1ee25526b0e..3698ae430b74b 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gamma/Beta.lean @@ -40,8 +40,6 @@ refined properties of the Gamma function using these relations. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - set_option linter.uppercaseLean3 false open Filter intervalIntegral Set Real MeasureTheory @@ -288,13 +286,13 @@ theorem GammaSeq_eq_approx_Gamma_integral {s : ℂ} (hs : 0 < re s) {n : ℕ} (h /-- The main techical lemma for `GammaSeq_tendsto_Gamma`, expressing the integral defining the Gamma function for `0 < re s` as the limit of a sequence of integrals over finite intervals. -/ theorem approx_Gamma_integral_tendsto_Gamma_integral {s : ℂ} (hs : 0 < re s) : - Tendsto (fun n : ℕ => ∫ x : ℝ in (0)..n, ↑((1 - x / n) ^ n) * (x : ℂ) ^ (s - 1)) atTop + Tendsto (fun n : ℕ => ∫ x : ℝ in (0)..n, ((1 - x / n) ^ n : ℝ) * (x : ℂ) ^ (s - 1)) atTop (𝓝 <| Gamma s) := by rw [Gamma_eq_integral hs] -- We apply dominated convergence to the following function, which we will show is uniformly -- bounded above by the Gamma integrand `exp (-x) * x ^ (re s - 1)`. let f : ℕ → ℝ → ℂ := fun n => - indicator (Ioc 0 (n : ℝ)) fun x : ℝ => ↑((1 - x / n) ^ n) * (x : ℂ) ^ (s - 1) + indicator (Ioc 0 (n : ℝ)) fun x : ℝ => ((1 - x / n) ^ n : ℝ) * (x : ℂ) ^ (s - 1) -- integrability of f have f_ible : ∀ n : ℕ, Integrable (f n) (volume.restrict (Ioi 0)) := by intro n @@ -322,7 +320,7 @@ theorem approx_Gamma_integral_tendsto_Gamma_integral {s : ℂ} (hs : 0 < re s) : refine' (Tendsto.comp (continuous_ofReal.tendsto _) _).const_mul _ convert tendsto_one_plus_div_pow_exp (-x) using 1 ext1 n - rw [neg_div, ← sub_eq_add_neg]; norm_cast + rw [neg_div, ← sub_eq_add_neg] -- let `convert` identify the remaining goals convert tendsto_integral_of_dominated_convergence _ (fun n => (f_ible n).1) (Real.GammaIntegral_convergent hs) _ @@ -403,7 +401,7 @@ theorem GammaSeq_mul (z : ℂ) {n : ℕ} (hn : n ≠ 0) : rw [this, Finset.prod_range_succ', Finset.prod_range_succ, aux, ← Finset.prod_mul_distrib, Nat.cast_zero, add_zero, add_comm (1 - z) n, ← add_sub_assoc] have : ∀ j : ℕ, (z + ↑(j + 1)) * (↑1 - z + ↑j) = - ↑((j + 1) ^ 2) * (↑1 - z ^ 2 / ((j : ℂ) + 1) ^ 2) := by + ((j + 1) ^ 2 :) * (↑1 - z ^ 2 / ((j : ℂ) + 1) ^ 2) := by intro j push_cast have : (j : ℂ) + 1 ≠ 0 := by rw [← Nat.cast_succ, Nat.cast_ne_zero]; exact Nat.succ_ne_zero j diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian.lean index 648e69abf1301..a57ef34d22ef9 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian.lean @@ -37,8 +37,6 @@ for positive real `a`, or complex `a` with positive real part. (See also `NumberTheory.ModularForms.JacobiTheta`.) -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open Real Set MeasureTheory Filter Asymptotics diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 32799dfd0dbea..4f52e49c8a07c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -37,8 +37,6 @@ open Real Nat Set Finset open scoped Real BigOperators Interval -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {a b : ℝ} (n : ℕ) namespace intervalIntegral @@ -561,10 +559,9 @@ theorem integral_mul_cpow_one_add_sq {t : ℂ} (ht : t ≠ -1) : have : t + 1 ≠ 0 := by contrapose! ht; rwa [add_eq_zero_iff_eq_neg] at ht apply integral_eq_sub_of_hasDerivAt · intro x _ - have f : HasDerivAt (fun y : ℂ => 1 + y ^ 2) (2 * x) x := by + have f : HasDerivAt (fun y : ℂ => 1 + y ^ 2) (2 * x : ℂ) x := by convert (hasDerivAt_pow 2 (x : ℂ)).const_add 1 - · norm_cast - · simp + simp have g : ∀ {z : ℂ}, 0 < z.re → HasDerivAt (fun z => z ^ (t + 1) / (2 * (t + 1))) (z ^ t / 2) z := by intro z hz @@ -573,7 +570,6 @@ theorem integral_mul_cpow_one_add_sq {t : ℂ} (ht : t ≠ -1) : field_simp ring convert (HasDerivAt.comp (↑x) (g _) f).comp_ofReal using 1 - · simp · field_simp; ring · exact_mod_cast add_pos_of_pos_of_nonneg zero_lt_one (sq_nonneg x) · apply Continuous.intervalIntegrable @@ -599,7 +595,6 @@ theorem integral_mul_rpow_one_add_sq {t : ℝ} (ht : t ≠ -1) : · rw [← intervalIntegral.integral_ofReal] congr with x : 1 rw [ofReal_mul, this x t] - norm_cast · simp_rw [ofReal_sub, ofReal_div, this a (t + 1), this b (t + 1)] push_cast; rfl · rw [← ofReal_one, ← ofReal_neg, Ne.def, ofReal_inj] @@ -755,8 +750,9 @@ theorem integral_sin_pow_mul_cos_pow_odd (m n : ℕ) : simp only [_root_.pow_zero, _root_.pow_succ', mul_assoc, pow_mul, one_mul] congr! 5 rw [← sq, ← sq, cos_sq'] - _ = ∫ u in sin a..sin b, u ^ m * (↑1 - u ^ 2) ^ n := - integral_comp_mul_deriv (fun x _ => hasDerivAt_sin x) continuousOn_cos hc + _ = ∫ u in sin a..sin b, u ^ m * (1 - u ^ 2) ^ n := by + -- Note(kmill): Didn't need `by exact`, but elaboration order seems to matter here. + exact integral_comp_mul_deriv (fun x _ => hasDerivAt_sin x) continuousOn_cos hc #align integral_sin_pow_mul_cos_pow_odd integral_sin_pow_mul_cos_pow_odd /-- The integral of `sin x * cos x`, given in terms of sin². @@ -847,10 +843,10 @@ theorem integral_sin_sq_mul_cos_sq : /-! ### Integral of misc. functions -/ -theorem integral_sqrt_one_sub_sq : ∫ x in (-1 : ℝ)..1, sqrt ((1 : ℝ) - x ^ 2) = π / 2 := +theorem integral_sqrt_one_sub_sq : ∫ x in (-1 : ℝ)..1, sqrt (1 - x ^ 2 : ℝ) = π / 2 := calc - _ = ∫ x in sin (-(π / 2)).. sin (π / 2), sqrt (↑1 - x ^ 2) := by rw [sin_neg, sin_pi_div_two] - _ = ∫ x in (-(π / 2))..(π / 2), sqrt (↑1 - sin x ^ 2) * cos x := + _ = ∫ x in sin (-(π / 2)).. sin (π / 2), sqrt (1 - x ^ 2 : ℝ) := by rw [sin_neg, sin_pi_div_two] + _ = ∫ x in (-(π / 2))..(π / 2), sqrt (1 - sin x ^ 2 : ℝ) * cos x := (integral_comp_mul_deriv (fun x _ => hasDerivAt_sin x) continuousOn_cos (by continuity)).symm _ = ∫ x in (-(π / 2))..(π / 2), cos x ^ 2 := by diff --git a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean index 480202a6c3271..b73d3edc39161 100644 --- a/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean +++ b/Mathlib/Analysis/SpecialFunctions/JapaneseBracket.lean @@ -26,7 +26,6 @@ than the dimension. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 open scoped BigOperators NNReal Filter Topology ENNReal open Asymptotics Filter Set Real MeasureTheory FiniteDimensional diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean index a304c3bb83a86..870fa88ec0597 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Base.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Base.lean @@ -454,14 +454,14 @@ lemma Real.induction_Ico_mul {P : ℝ → Prop} (x₀ r : ℝ) (hr : 1 < r) (hx · intro x hx have hx' : 0 < x / x₀ := div_pos (hx₀.trans_le hx) hx₀ refine this ⌊logb r (x / x₀)⌋₊ x ?_ - rw [mem_Ico, ←div_lt_iff hx₀, ←logb_lt_iff_lt_rpow hr hx'] + rw [mem_Ico, ←div_lt_iff hx₀, ← rpow_nat_cast, ←logb_lt_iff_lt_rpow hr hx', Nat.cast_add, + Nat.cast_one] exact ⟨hx, Nat.lt_floor_add_one _⟩ intro n induction n case zero => simpa using base case succ n ih => specialize step (n + 1) (by simp) - simp only [Nat.cast_add_one] at step ⊢ exact fun x hx => (Ico_subset_Ico_union_Ico hx).elim (ih x) (step ih _) end Induction diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean index 77f793bf586fb..54afe6bbd09e7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Basic.lean @@ -348,7 +348,8 @@ theorem tendsto_log_nhdsWithin_zero : Tendsto log (𝓝[≠] 0) atBot := by #align real.tendsto_log_nhds_within_zero Real.tendsto_log_nhdsWithin_zero theorem continuousOn_log : ContinuousOn log {0}ᶜ := by - simp only [continuousOn_iff_continuous_restrict, restrict] + simp (config := { unfoldPartialApp := true }) only [continuousOn_iff_continuous_restrict, + restrict] conv in log _ => rw [log_of_ne_zero (show (x : ℝ) ≠ 0 from x.2)] exact expOrderIso.symm.continuous.comp (continuous_subtype_val.norm.subtype_mk _) #align real.continuous_on_log Real.continuousOn_log diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index 643b5bbefcd92..a2cf16f746e64 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -19,8 +19,6 @@ It satisfies the following change of variables formula (see `integral_comp_polar -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section Real open Real Set MeasureTheory diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean index bac158beefdce..e34623929ee4c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Complex.lean @@ -13,8 +13,6 @@ import Mathlib.Analysis.SpecialFunctions.Complex.Log We construct the power functions `x ^ y`, where `x` and `y` are complex numbers. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue #2220 - open Classical Real Topology Filter ComplexConjugate Finset Set namespace Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean index bdd10ff27c78a..3e679b1e0810f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Continuity.lean @@ -21,8 +21,6 @@ open Classical Real Topology NNReal ENNReal Filter BigOperators ComplexConjugate open Filter Finset Set -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - section CpowLimits /-! diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index fdc0edd04d244..41bded1ec1b11 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -236,9 +236,9 @@ theorem hasDerivAt_ofReal_cpow {x : ℝ} (hx : x ≠ 0) {r : ℂ} (hr : r ≠ -1 rw [mul_add ((π : ℂ) * _), mul_one, exp_add, exp_pi_mul_I, mul_comm (_ : ℂ) (-1 : ℂ), neg_one_mul] simp_rw [mul_neg, ← neg_mul, ← ofReal_neg] - suffices HasDerivAt (fun y : ℝ => ↑(-y) ^ (r + 1)) (-(r + 1) * ↑(-x) ^ r) x by + suffices HasDerivAt (fun y : ℝ => (↑(-y) : ℂ) ^ (r + 1)) (-(r + 1) * ↑(-x) ^ r) x by convert this.neg.mul_const _ using 1; ring - suffices HasDerivAt (fun y : ℝ => ↑y ^ (r + 1)) ((r + 1) * ↑(-x) ^ r) (-x) by + suffices HasDerivAt (fun y : ℝ => (y : ℂ) ^ (r + 1)) ((r + 1) * ↑(-x) ^ r) (-x) by convert @HasDerivAt.scomp ℝ _ ℂ _ _ x ℝ _ _ _ _ _ _ _ _ this (hasDerivAt_neg x) using 1 rw [real_smul, ofReal_neg 1, ofReal_one]; ring suffices HasDerivAt (fun y : ℂ => y ^ (r + 1)) ((r + 1) * ↑(-x) ^ r) ↑(-x) by @@ -329,7 +329,7 @@ lemma differentiableAt_rpow_const_of_ne (p : ℝ) {x : ℝ} (hx : x ≠ 0) : (hasStrictDerivAt_rpow_const_of_ne hx p).differentiableAt lemma differentiableOn_rpow_const (p : ℝ) : - DifferentiableOn ℝ (fun x => x ^ p) {0}ᶜ := + DifferentiableOn ℝ (fun x => (x : ℝ) ^ p) {0}ᶜ := fun _ hx => (Real.differentiableAt_rpow_const_of_ne p hx).differentiableWithinAt /-- This lemma says that `fun x => a ^ x` is strictly differentiable for `a < 0`. Note that these diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean index 2700b6c4e4895..9b6430b80920e 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean @@ -18,8 +18,6 @@ We construct the power functions `x ^ y` where We also prove basic properties of these functions. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open Classical Real NNReal ENNReal BigOperators ComplexConjugate @@ -444,7 +442,7 @@ theorem coe_rpow_of_nonneg (x : ℝ≥0) {y : ℝ} (h : 0 ≤ y) : (x : ℝ≥0 #align ennreal.coe_rpow_of_nonneg ENNReal.coe_rpow_of_nonneg theorem coe_rpow_def (x : ℝ≥0) (y : ℝ) : - (x : ℝ≥0∞) ^ y = if x = 0 ∧ y < 0 then ⊤ else (x ^ y : ℝ≥0∞) := + (x : ℝ≥0∞) ^ y = if x = 0 ∧ y < 0 then ⊤ else ↑(x ^ y) := rfl #align ennreal.coe_rpow_def ENNReal.coe_rpow_def diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index 0b5f1c153d95c..4caf7a61e230f 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -268,7 +268,7 @@ lemma cpow_ofReal (x : ℂ) (y : ℝ) : lemma cpow_ofReal_re (x : ℂ) (y : ℝ) : (x ^ (y : ℂ)).re = (abs x) ^ y * Real.cos (arg x * y) := by rw [cpow_ofReal]; generalize arg x * y = z; simp [Real.cos] -lemma cpow_ofReal_im (x : ℂ) (y : ℝ) : (x ^ y).im = (abs x) ^ y * Real.sin (arg x * y) := by +lemma cpow_ofReal_im (x : ℂ) (y : ℝ) : (x ^ (y : ℂ)).im = (abs x) ^ y * Real.sin (arg x * y) := by rw [cpow_ofReal]; generalize arg x * y = z; simp [Real.sin] theorem abs_cpow_of_ne_zero {z : ℂ} (hz : z ≠ 0) (w : ℂ) : @@ -322,8 +322,6 @@ end Complex namespace Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {x y z : ℝ} theorem rpow_mul {x : ℝ} (hx : 0 ≤ x) (y z : ℝ) : x ^ (y * z) = (x ^ y) ^ z := by @@ -745,8 +743,6 @@ end Sqrt variable {n : ℕ} -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - theorem exists_rat_pow_btwn_rat_aux (hn : n ≠ 0) (x y : ℝ) (h : x < y) (hy : 0 < y) : ∃ q : ℚ, 0 < q ∧ x < (q : ℝ) ^ n ∧ (q : ℝ) ^ n < y := by have hn' : 0 < (n : ℝ) := by exact_mod_cast hn.bot_lt diff --git a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean index 6799e71dec91e..ac3ec7e95c8cf 100644 --- a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean +++ b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean @@ -27,7 +27,6 @@ cannot have: set_option autoImplicit true -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 noncomputable section open scoped Classical Topology diff --git a/Mathlib/Analysis/SpecialFunctions/Stirling.lean b/Mathlib/Analysis/SpecialFunctions/Stirling.lean index 5ca7c9998f48f..0f092d8efd5ab 100644 --- a/Mathlib/Analysis/SpecialFunctions/Stirling.lean +++ b/Mathlib/Analysis/SpecialFunctions/Stirling.lean @@ -36,8 +36,6 @@ open scoped Topology Real BigOperators Nat open Finset Filter Nat Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Stirling /-! @@ -80,7 +78,7 @@ theorem log_stirlingSeq_diff_hasSum (m : ℕ) : HasSum (fun k : ℕ => ↑1 / (↑2 * ↑(k + 1) + ↑1) * (((1:ℝ)/(↑2 * ↑(m + 1) + ↑1)) ^ 2) ^ ↑(k + 1)) (log (stirlingSeq (m + 1)) - log (stirlingSeq (m + 2))) := by change HasSum - ((fun b : ℕ => (1:ℝ) / ((2:ℝ) * b + (1:ℝ)) * (((1:ℝ) / (2 * ↑(m + 1) + 1)) ^ 2) ^ b) ∘ succ) _ + ((fun b : ℕ => (1:ℝ) / ((2:ℝ) * b + (1:ℝ)) * (((1:ℝ) / (2 * (m + 1 :) + 1)) ^ 2) ^ b) ∘ succ) _ refine' (hasSum_nat_add_iff (a := _ - _) -- Porting note: must give implicit arguments (f := (fun b : ℕ => ↑1 / (↑2 * b + ↑1) * (((1:ℝ) / (2 * ↑(m + 1) + 1)) ^ 2) ^ b)) 1).mpr _ convert (hasSum_log_one_add_inv <| diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index b36f0f1b8c80a..0687a34849fa9 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -184,7 +184,7 @@ theorem two_zsmul_eq_iff {ψ θ : Angle} : (2 : ℤ) • ψ = (2 : ℤ) • θ -- Porting note: no `Int.natAbs_bit0` anymore have : Int.natAbs 2 = 2 := rfl rw [zsmul_eq_iff two_ne_zero, this, Fin.exists_fin_two, Fin.val_zero, - Fin.val_one, zero_smul, coe_zero, add_zero, one_smul, Int.cast_two, + Fin.val_one, zero_smul, add_zero, one_smul, Int.cast_two, mul_div_cancel_left (_ : ℝ) two_ne_zero] #align real.angle.two_zsmul_eq_iff Real.Angle.two_zsmul_eq_iff @@ -224,7 +224,7 @@ theorem neg_ne_self_iff {θ : Angle} : -θ ≠ θ ↔ θ ≠ 0 ∧ θ ≠ π := #align real.angle.neg_ne_self_iff Real.Angle.neg_ne_self_iff theorem two_nsmul_eq_pi_iff {θ : Angle} : (2 : ℕ) • θ = π ↔ θ = (π / 2 : ℝ) ∨ θ = (-π / 2 : ℝ) := by - have h : (π : Angle) = (2 : ℕ) • (π / 2 : ℝ) := by rw [two_nsmul, add_halves] + have h : (π : Angle) = ((2 : ℕ) • (π / 2 : ℝ) :) := by rw [two_nsmul, add_halves] nth_rw 1 [h] rw [coe_nsmul, two_nsmul_eq_iff] -- Porting note: `congr` didn't simplify the goal of iff of `Or`s diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean index 34233bc6a0c1f..14ef64ab59143 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Arctan.lean @@ -17,8 +17,6 @@ and `Real.tan` as a `LocalHomeomorph` between `(-(π / 2), π / 2)` and the whol noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Real open Set Filter diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean index 3fe76c99a7630..d396d8eef2484 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ArctanDeriv.lean @@ -17,8 +17,6 @@ Continuity and derivatives of the tangent and arctangent functions. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Real open Set Filter diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean index 3b09875f83f53..f14557d73a791 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Bounds.lean @@ -33,8 +33,6 @@ sin, cos, tan, angle noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Set namespace Real @@ -108,7 +106,7 @@ theorem lt_tan {x : ℝ} (h1 : 0 < x) (h2 : x < π / 2) : x < tan x := by apply lt_of_le_of_ne y.cos_sq_le_one rw [cos_sq'] simpa only [Ne.def, sub_eq_self, pow_eq_zero_iff, Nat.succ_pos'] using (sin_pos hy).ne' - rwa [one_div, lt_inv, inv_one] + rwa [lt_inv, inv_one] · exact zero_lt_one simpa only [sq, mul_self_pos] using this.ne' have mono := Convex.strictMonoOn_of_deriv_pos (convex_Ico 0 (π / 2)) tan_minus_id_cts deriv_pos diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Complex.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Complex.lean index e776100272383..5329e884426ca 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Complex.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Complex.lean @@ -121,8 +121,6 @@ theorem tan_add' {x y : ℂ} tan_add (Or.inl h) #align complex.tan_add' Complex.tan_add' -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - theorem tan_two_mul {z : ℂ} : tan (2 * z) = (2 : ℂ) * tan z / ((1 : ℂ) - tan z ^ 2) := by by_cases h : ∀ k : ℤ, z ≠ (2 * k + 1) * π / 2 · rw [two_mul, two_mul, sq, tan_add (Or.inl ⟨h, h⟩)] diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean index 2678cb0e3bcd9..1af7485b9511d 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean @@ -16,8 +16,6 @@ Basic facts and derivatives for the complex trigonometric functions. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Complex open Set Filter diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean index 743f6583cf95a..409c92704e6b2 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/EulerSineProd.lean @@ -20,9 +20,6 @@ is to prove a recurrence relation for the integrals `∫ x in 0..π/2, cos 2 z x generalising the arguments used to prove Wallis' limit formula for `π`. -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped Real Topology BigOperators open Real Set Filter intervalIntegral MeasureTheory.MeasureSpace @@ -336,8 +333,8 @@ theorem _root_.Real.tendsto_euler_sin_prod (x : ℝ) : · ext1 n rw [Function.comp_apply, ← Complex.ofReal_mul, Complex.ofReal_mul_re] suffices - (∏ j : ℕ in Finset.range n, ((1 : ℂ) - (x : ℂ) ^ 2 / ((j : ℂ) + 1) ^ 2)) = - (∏ j : ℕ in Finset.range n, ((1 : ℝ) - x ^ 2 / ((j : ℝ) + 1) ^ 2) : ℝ) by + (∏ j : ℕ in Finset.range n, (1 - x ^ 2 / (j + 1) ^ 2) : ℂ) = + (∏ j : ℕ in Finset.range n, (1 - x ^ 2 / (j + 1) ^ 2) : ℝ) by rw [this, Complex.ofReal_re] rw [Complex.ofReal_prod] refine' Finset.prod_congr (by rfl) fun n _ => _ diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean index 1776aa4a6924c..15e915ebaf98c 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Series.lean @@ -34,7 +34,7 @@ theorem Complex.hasSum_cos' (z : ℂ) : have := ((expSeries_div_hasSum_exp ℂ (z * Complex.I)).add (expSeries_div_hasSum_exp ℂ (-z * Complex.I))).div_const 2 replace := (Nat.divModEquiv 2).symm.hasSum_iff.mpr this - dsimp [Function.comp] at this + dsimp [Function.comp_def] at this simp_rw [← mul_comm 2 _] at this refine' this.prod_fiberwise fun k => _ dsimp only @@ -52,7 +52,7 @@ theorem Complex.hasSum_sin' (z : ℂ) : have := (((expSeries_div_hasSum_exp ℂ (-z * Complex.I)).sub (expSeries_div_hasSum_exp ℂ (z * Complex.I))).mul_right Complex.I).div_const 2 replace := (Nat.divModEquiv 2).symm.hasSum_iff.mpr this - dsimp [Function.comp] at this + dsimp [Function.comp_def] at this simp_rw [← mul_comm 2 _] at this refine' this.prod_fiberwise fun k => _ dsimp only diff --git a/Mathlib/Analysis/SpecificLimits/FloorPow.lean b/Mathlib/Analysis/SpecificLimits/FloorPow.lean index 5473de4519869..82ea4049d0a68 100644 --- a/Mathlib/Analysis/SpecificLimits/FloorPow.lean +++ b/Mathlib/Analysis/SpecificLimits/FloorPow.lean @@ -19,8 +19,6 @@ We state several auxiliary results pertaining to sequences of the form `⌊c^n to `1/j^2`, up to a multiplicative constant. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Filter Finset open Topology BigOperators diff --git a/Mathlib/CategoryTheory/Action.lean b/Mathlib/CategoryTheory/Action.lean index 99cccd566b98e..80c5b00afa6c0 100644 --- a/Mathlib/CategoryTheory/Action.lean +++ b/Mathlib/CategoryTheory/Action.lean @@ -160,7 +160,7 @@ def endMulEquivSubgroup (H : Subgroup G) : End (objEquiv G (G ⧸ H) ↑(1 : G)) #align category_theory.action_category.End_mul_equiv_subgroup CategoryTheory.ActionCategory.endMulEquivSubgroup /-- A target vertex `t` and a scalar `g` determine a morphism in the action groupoid. -/ -def homOfPair (t : X) (g : G) : @Quiver.Hom (ActionCategory G X) _ (g⁻¹ • t) t := +def homOfPair (t : X) (g : G) : @Quiver.Hom (ActionCategory G X) _ (g⁻¹ • t :) t := Subtype.mk g (smul_inv_smul g t) #align category_theory.action_category.hom_of_pair CategoryTheory.ActionCategory.homOfPair diff --git a/Mathlib/CategoryTheory/Category/Basic.lean b/Mathlib/CategoryTheory/Category/Basic.lean index 3710a98d7fa7b..394b27e4acfc4 100644 --- a/Mathlib/CategoryTheory/Category/Basic.lean +++ b/Mathlib/CategoryTheory/Category/Basic.lean @@ -117,6 +117,7 @@ use in auto-params. macro (name := aesop_cat) "aesop_cat" c:Aesop.tactic_clause* : tactic => `(tactic| aesop $c* (options := { introsTransparency? := some .default, terminal := true }) + (simp_options := { decide := true }) (rule_sets [$(Lean.mkIdent `CategoryTheory):ident])) /-- diff --git a/Mathlib/CategoryTheory/Category/KleisliCat.lean b/Mathlib/CategoryTheory/Category/KleisliCat.lean index 21f09a2d4e7cd..f217842150a4e 100644 --- a/Mathlib/CategoryTheory/Category/KleisliCat.lean +++ b/Mathlib/CategoryTheory/Category/KleisliCat.lean @@ -52,7 +52,7 @@ instance KleisliCat.category {m} [Monad.{u, v} m] [LawfulMonad m] : Category (Kl -- refine' { id_comp' := _, comp_id' := _, assoc' := _ } <;> intros <;> ext <;> unfold_projs <;> -- simp only [(· >=> ·), functor_norm] refine' { id_comp := _, comp_id := _, assoc := _ } <;> intros <;> refine funext (fun x => ?_) <;> - simp [CategoryStruct.id, CategoryStruct.comp, (· >=> ·)] + simp (config := { unfoldPartialApp := true }) [CategoryStruct.id, CategoryStruct.comp, (· >=> ·)] #align category_theory.Kleisli.category CategoryTheory.KleisliCat.category @[simp] diff --git a/Mathlib/CategoryTheory/CofilteredSystem.lean b/Mathlib/CategoryTheory/CofilteredSystem.lean index c16747f986b4c..44481ea1071ee 100644 --- a/Mathlib/CategoryTheory/CofilteredSystem.lean +++ b/Mathlib/CategoryTheory/CofilteredSystem.lean @@ -184,11 +184,11 @@ def toPreimages : J ⥤ Type v where rw [← mem_preimage, preimage_preimage, mem_preimage] convert h (g ≫ f); rw [F.map_comp]; rfl map_id j := by - simp_rw [MapsTo.restrict, Subtype.map, F.map_id] + simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_id] ext rfl map_comp f g := by - simp_rw [MapsTo.restrict, Subtype.map, F.map_comp] + simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_comp] rfl #align category_theory.functor.to_preimages CategoryTheory.Functor.toPreimages @@ -271,11 +271,11 @@ def toEventualRanges : J ⥤ Type v where obj j := F.eventualRange j map f := (F.eventualRange_mapsTo f).restrict _ _ _ map_id i := by - simp_rw [MapsTo.restrict, Subtype.map, F.map_id] + simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_id] ext rfl map_comp _ _ := by - simp_rw [MapsTo.restrict, Subtype.map, F.map_comp] + simp (config := { unfoldPartialApp := true }) only [MapsTo.restrict, Subtype.map, F.map_comp] rfl #align category_theory.functor.to_eventual_ranges CategoryTheory.Functor.toEventualRanges diff --git a/Mathlib/CategoryTheory/DifferentialObject.lean b/Mathlib/CategoryTheory/DifferentialObject.lean index e52f2713ca2d9..1e2da407c290b 100644 --- a/Mathlib/CategoryTheory/DifferentialObject.lean +++ b/Mathlib/CategoryTheory/DifferentialObject.lean @@ -200,8 +200,8 @@ def mapDifferentialObject (F : C ⥤ D) slice_lhs 2 3 => rw [← Functor.comp_map F (shiftFunctor D (1 : S)), ← η.naturality f.f] slice_lhs 1 2 => rw [Functor.comp_map, ← F.map_comp, f.comm, F.map_comp] rw [Category.assoc] } - map_id := by intros; ext; simp - map_comp := by intros; ext; simp + map_id := by intros; ext; simp [autoParam] + map_comp := by intros; ext; simp [autoParam] #align category_theory.functor.map_differential_object CategoryTheory.Functor.mapDifferentialObject end Functor diff --git a/Mathlib/CategoryTheory/Elements.lean b/Mathlib/CategoryTheory/Elements.lean index 1b2b79c8817d4..e75472b543473 100644 --- a/Mathlib/CategoryTheory/Elements.lean +++ b/Mathlib/CategoryTheory/Elements.lean @@ -273,7 +273,7 @@ theorem costructuredArrow_yoneda_equivalence_naturality {F₁ F₂ : Cᵒᵖ ⥤ congr ext _ f simpa using congr_fun (α.naturality f.op).symm (unop X).snd - · aesop + · simp [autoParam] #align category_theory.category_of_elements.costructured_arrow_yoneda_equivalence_naturality CategoryTheory.CategoryOfElements.costructuredArrow_yoneda_equivalence_naturality end CategoryOfElements diff --git a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean index f2a2357610d3e..46a945db9e513 100644 --- a/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean +++ b/Mathlib/CategoryTheory/Groupoid/Subgroupoid.lean @@ -688,7 +688,7 @@ theorem full_empty : full ∅ = (⊥ : Subgroupoid C) := by @[simp] theorem full_univ : full Set.univ = (⊤ : Subgroupoid C) := by ext - simp only [mem_full_iff, mem_univ, mem_top] + simp only [mem_full_iff, mem_univ, and_self, mem_top] #align category_theory.subgroupoid.full_univ CategoryTheory.Subgroupoid.full_univ theorem full_mono {D E : Set C} (h : D ≤ E) : full D ≤ full E := by diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean index 92ed064f07cb8..3f49afa673d47 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Biproducts.lean @@ -419,7 +419,7 @@ theorem biproduct.map_lift_mapBiprod (g : ∀ j, W ⟶ f j) : haveI : HasBiproduct fun j => F.obj (f j) := hasBiproduct_of_preserves F f F.map (biproduct.lift g) ≫ (F.mapBiproduct f).hom = biproduct.lift fun j => F.map (g j) := by ext j - dsimp only [Function.comp] + dsimp only [Function.comp_def] haveI : HasBiproduct fun j => F.obj (f j) := hasBiproduct_of_preserves F f simp only [mapBiproduct_hom, Category.assoc, biproduct.lift_π, ← F.map_comp] #align category_theory.limits.biproduct.map_lift_map_biprod CategoryTheory.Limits.biproduct.map_lift_mapBiprod @@ -429,7 +429,7 @@ theorem biproduct.mapBiproduct_inv_map_desc (g : ∀ j, f j ⟶ W) : haveI : HasBiproduct fun j => F.obj (f j) := hasBiproduct_of_preserves F f (F.mapBiproduct f).inv ≫ F.map (biproduct.desc g) = biproduct.desc fun j => F.map (g j) := by ext j - dsimp only [Function.comp] + dsimp only [Function.comp_def] haveI : HasBiproduct fun j => F.obj (f j) := hasBiproduct_of_preserves F f simp only [mapBiproduct_inv, ← Category.assoc, biproduct.ι_desc ,← F.map_comp] #align category_theory.limits.biproduct.map_biproduct_inv_map_desc CategoryTheory.Limits.biproduct.mapBiproduct_inv_map_desc diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 50f18610473e9..64d714910ec42 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -107,8 +107,8 @@ instance : Category.{v₁} (Mat_ C) where Hom := Hom id := Hom.id comp f g := f.comp g - id_comp f := by simp [dite_comp] - comp_id f := by simp [comp_dite] + id_comp f := by simp (config := { unfoldPartialApp := true }) [dite_comp] + comp_id f := by simp (config := { unfoldPartialApp := true }) [comp_dite] assoc f g h := by apply DMatrix.ext intros diff --git a/Mathlib/CategoryTheory/Sites/Types.lean b/Mathlib/CategoryTheory/Sites/Types.lean index b7d89512addcb..8bda080f4e657 100644 --- a/Mathlib/CategoryTheory/Sites/Types.lean +++ b/Mathlib/CategoryTheory/Sites/Types.lean @@ -191,7 +191,7 @@ theorem typesGrothendieckTopology_eq_canonical : have : (fun _ => ULift.up true) = fun _ => ULift.up false := (hs PUnit fun _ => x).isSeparatedFor.ext fun β f hf => funext fun y => hsx.elim <| S.2 hf fun _ => y - simp at this + simp [Function.funext_iff] at this #align category_theory.types_grothendieck_topology_eq_canonical CategoryTheory.typesGrothendieckTopology_eq_canonical end CategoryTheory diff --git a/Mathlib/Combinatorics/Additive/Behrend.lean b/Mathlib/Combinatorics/Additive/Behrend.lean index fcd3a060500a2..ce364878bceb6 100644 --- a/Mathlib/Combinatorics/Additive/Behrend.lean +++ b/Mathlib/Combinatorics/Additive/Behrend.lean @@ -213,7 +213,7 @@ theorem sum_lt : (∑ i : Fin n, d * (2 * d + 1) ^ (i : ℕ)) < (2 * d + 1) ^ n theorem card_sphere_le_rothNumberNat (n d k : ℕ) : (sphere n d k).card ≤ rothNumberNat ((2 * d - 1) ^ n) := by cases n - · dsimp; refine' (card_le_univ _).trans_eq _; simp + · dsimp; refine' (card_le_univ _).trans_eq _; rfl cases d · simp refine' addSalemSpencer_image_sphere.le_rothNumberNat _ _ (card_image_of_injOn _) @@ -239,7 +239,7 @@ that we then optimize by tweaking the parameters. The (almost) optimal parameter theorem exists_large_sphere_aux (n d : ℕ) : ∃ k ∈ range (n * (d - 1) ^ 2 + 1), - (↑(d ^ n) / (↑(n * (d - 1) ^ 2) + 1) : ℝ) ≤ (sphere n d k).card := by + (↑(d ^ n) / ((n * (d - 1) ^ 2 :) + 1) : ℝ) ≤ (sphere n d k).card := by refine' exists_le_card_fiber_of_nsmul_le_card_of_maps_to (fun x hx => _) nonempty_range_succ _ · rw [mem_range, lt_succ_iff] exact sum_sq_le_of_mem_box hx @@ -248,14 +248,14 @@ theorem exists_large_sphere_aux (n d : ℕ) : ∃ k ∈ range (n * (d - 1) ^ 2 + exact (cast_add_one_pos _).ne' #align behrend.exists_large_sphere_aux Behrend.exists_large_sphere_aux -theorem exists_large_sphere (n d : ℕ) : ∃ k, (d ^ n / ↑(n * d ^ 2) : ℝ) ≤ (sphere n d k).card := by +theorem exists_large_sphere (n d : ℕ) : + ∃ k, ((d ^ n :) / (n * d ^ 2 :) : ℝ) ≤ (sphere n d k).card := by obtain ⟨k, -, hk⟩ := exists_large_sphere_aux n d refine' ⟨k, _⟩ obtain rfl | hn := n.eq_zero_or_pos · simp obtain rfl | hd := d.eq_zero_or_pos · simp - rw [rpow_nat_cast, ← cast_pow] refine' (div_le_div_of_le_left _ _ _).trans hk · exact cast_nonneg _ · exact cast_add_one_pos _ @@ -268,16 +268,16 @@ theorem exists_large_sphere (n d : ℕ) : ∃ k, (d ^ n / ↑(n * d ^ 2) : ℝ) exact one_le_cast.2 hd #align behrend.exists_large_sphere Behrend.exists_large_sphere -theorem bound_aux' (n d : ℕ) : (d ^ n / ↑(n * d ^ 2) : ℝ) ≤ rothNumberNat ((2 * d - 1) ^ n) := +theorem bound_aux' (n d : ℕ) : ((d ^ n :) / (n * d ^ 2 :) : ℝ) ≤ rothNumberNat ((2 * d - 1) ^ n) := let ⟨_, h⟩ := exists_large_sphere n d h.trans <| cast_le.2 <| card_sphere_le_rothNumberNat _ _ _ #align behrend.bound_aux' Behrend.bound_aux' theorem bound_aux (hd : d ≠ 0) (hn : 2 ≤ n) : - (d ^ (n - 2) / n : ℝ) ≤ rothNumberNat ((2 * d - 1) ^ n) := by + (d ^ (n - 2 :) / n : ℝ) ≤ rothNumberNat ((2 * d - 1) ^ n) := by convert bound_aux' n d using 1 - rw [cast_mul, cast_pow, mul_comm, ← div_div, ← cast_two, ← cast_sub hn, rpow_nat_cast, - rpow_nat_cast, pow_sub₀ _ (cast_ne_zero.2 hd) hn, ← div_eq_mul_inv] + rw [cast_mul, cast_pow, mul_comm, ← div_div, pow_sub₀ _ _ hn, ← div_eq_mul_inv, cast_pow] + rwa [cast_ne_zero] #align behrend.bound_aux Behrend.bound_aux open scoped Filter Topology @@ -432,7 +432,8 @@ theorem le_N (hN : 2 ≤ N) : (2 * dValue N - 1) ^ nValue N ≤ N := by apply this.trans suffices ((2 * dValue N) ^ nValue N : ℝ) ≤ N by exact_mod_cast this suffices i : (2 * dValue N : ℝ) ≤ (N : ℝ) ^ (1 / nValue N : ℝ) - · apply (rpow_le_rpow (mul_nonneg zero_le_two (cast_nonneg _)) i (cast_nonneg _)).trans + · rw [← rpow_nat_cast] + apply (rpow_le_rpow (mul_nonneg zero_le_two (cast_nonneg _)) i (cast_nonneg _)).trans rw [← rpow_mul (cast_nonneg _), one_div_mul_cancel, rpow_one] rw [cast_ne_zero] apply (nValue_pos hN).ne' @@ -487,7 +488,6 @@ theorem roth_lower_bound_explicit (hN : 4096 ≤ N) : have hn₂ : 2 ≤ n := two_le_nValue (hN.trans' <| by norm_num1) have : (2 * dValue N - 1) ^ n ≤ N := le_N (hN.trans' <| by norm_num1) refine' ((bound_aux hd.ne' hn₂).trans <| cast_le.2 <| rothNumberNat.mono this).trans_lt' _ - conv_rhs => rw [← cast_two, ← cast_sub hn₂, rpow_nat_cast] refine' (div_lt_div_of_lt hn <| pow_lt_pow_of_lt_left (bound hN) _ _).trans_le' _ · exact div_nonneg (rpow_nonneg_of_nonneg (cast_nonneg _) _) (exp_pos _).le · exact tsub_pos_of_lt (three_le_nValue <| hN.trans' <| by norm_num1) @@ -521,7 +521,7 @@ theorem exp_four_lt : exp 4 < 64 := by theorem four_zero_nine_six_lt_exp_sixteen : 4096 < exp 16 := by rw [← log_lt_iff_lt_exp (show (0 : ℝ) < 4096 by norm_num), show (4096 : ℝ) = 2 ^ 12 by norm_cast, - log_rpow zero_lt_two] + ← rpow_nat_cast, log_rpow zero_lt_two, cast_ofNat] have : 12 * (0.6931471808 : ℝ) < 16 := by norm_num linarith [log_two_lt_d9] #align behrend.four_zero_nine_six_lt_exp_sixteen Behrend.four_zero_nine_six_lt_exp_sixteen diff --git a/Mathlib/Combinatorics/Quiver/Covering.lean b/Mathlib/Combinatorics/Quiver/Covering.lean index f385dbe801218..98ce09b60e344 100644 --- a/Mathlib/Combinatorics/Quiver/Covering.lean +++ b/Mathlib/Combinatorics/Quiver/Covering.lean @@ -216,7 +216,7 @@ theorem Prefunctor.pathStar_apply {u v : U} (p : Path u v) : theorem Prefunctor.pathStar_injective (hφ : ∀ u, Injective (φ.star u)) (u : U) : Injective (φ.pathStar u) := by - dsimp [Prefunctor.pathStar, Quiver.PathStar.mk] + dsimp (config := { unfoldPartialApp := true }) [Prefunctor.pathStar, Quiver.PathStar.mk] rintro ⟨v₁, p₁⟩ induction' p₁ with x₁ y₁ p₁ e₁ ih <;> rintro ⟨y₂, p₂⟩ <;> @@ -251,7 +251,7 @@ theorem Prefunctor.pathStar_injective (hφ : ∀ u, Injective (φ.star u)) (u : theorem Prefunctor.pathStar_surjective (hφ : ∀ u, Surjective (φ.star u)) (u : U) : Surjective (φ.pathStar u) := by - dsimp [Prefunctor.pathStar, Quiver.PathStar.mk] + dsimp (config := { unfoldPartialApp := true }) [Prefunctor.pathStar, Quiver.PathStar.mk] rintro ⟨v, p⟩ induction' p with v' v'' p' ev ih · use ⟨u, Path.nil⟩ diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index b7f2cbd34d6ae..df56819e27edf 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -238,7 +238,8 @@ protected lemma Finset.four_functions_theorem (u : Finset α) (collapse_modular hu h₁ h₂ h₃ h₄ h 𝒜 ℬ) Subset.rfl Subset.rfl have : 𝒜 ⊼ ℬ ⊆ powerset (insert a u) := by simpa using infs_subset h𝒜 hℬ have : 𝒜 ⊻ ℬ ⊆ powerset (insert a u) := by simpa using sups_subset h𝒜 hℬ - simpa only [powerset_sups_powerset_self, powerset_infs_powerset_self, sum_collapse, *] using ih + simpa only [powerset_sups_powerset_self, powerset_infs_powerset_self, sum_collapse, + not_false_eq_true, *] using ih variable (f₁ f₂ f₃ f₄) [Fintype α] diff --git a/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean b/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean index bd302168245f8..0edfa6f513ebf 100644 --- a/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean +++ b/Mathlib/Combinatorics/SetFamily/HarrisKleitman.lean @@ -59,10 +59,11 @@ theorem IsLowerSet.le_card_inter_finset' (h𝒜 : IsLowerSet (𝒜 : Set (Finset induction' s using Finset.induction with a s hs ih generalizing 𝒜 ℬ · simp_rw [subset_empty, ← subset_singleton_iff', subset_singleton_iff] at h𝒜s hℬs obtain rfl | rfl := h𝒜s - · simp only [card_empty, empty_inter, mul_zero, zero_mul] + · simp only [card_empty, zero_mul, empty_inter, mul_zero, le_refl] obtain rfl | rfl := hℬs - · simp only [card_empty, inter_empty, mul_zero, zero_mul] - · simp only [card_empty, pow_zero, inter_singleton_of_mem, mem_singleton, card_singleton] + · simp only [card_empty, inter_empty, mul_zero, zero_mul, le_refl] + · simp only [card_empty, pow_zero, inter_singleton_of_mem, mem_singleton, card_singleton, + le_refl] rw [card_insert_of_not_mem hs, ← card_memberSubfamily_add_card_nonMemberSubfamily a 𝒜, ← card_memberSubfamily_add_card_nonMemberSubfamily a ℬ, add_mul, mul_add, mul_add, add_comm (_ * _), add_add_add_comm] diff --git a/Mathlib/Combinatorics/SimpleGraph/Clique.lean b/Mathlib/Combinatorics/SimpleGraph/Clique.lean index d0010ff8b71b0..cd054df13128e 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Clique.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Clique.lean @@ -210,7 +210,7 @@ theorem not_cliqueFree_card_of_top_embedding [Fintype α] (f : (⊤ : SimpleGrap theorem cliqueFree_bot (h : 2 ≤ n) : (⊥ : SimpleGraph α).CliqueFree n := by intro t ht have := le_trans h (isNClique_bot_iff.1 ht).1 - simp only at this + contradiction #align simple_graph.clique_free_bot SimpleGraph.cliqueFree_bot theorem CliqueFree.mono (h : m ≤ n) : G.CliqueFree m → G.CliqueFree n := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean index 2757bd39389c2..0725f50222aac 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Bound.lean @@ -32,8 +32,6 @@ This entire file is internal to the proof of Szemerédi Regularity Lemma. open Finset Fintype Function Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open BigOperators namespace SzemerediRegularity @@ -72,7 +70,7 @@ local notation3 "a" => (card α / P.parts.card - m * 4 ^ P.parts.card : ℕ) namespace SzemerediRegularity.Positivity private theorem eps_pos {ε : ℝ} {n : ℕ} (h : 100 ≤ (4 : ℝ) ^ n * ε ^ 5) : 0 < ε := - (Odd.pow_pos_iff (by norm_num)).mp + (Odd.pow_pos_iff (by decide)).mp (pos_of_mul_pos_right ((show 0 < (100 : ℝ) by norm_num).trans_le h) (by positivity)) private theorem m_pos [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) : 0 < m := @@ -124,7 +122,7 @@ theorem eps_pow_five_pos (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : #align szemeredi_regularity.eps_pow_five_pos SzemerediRegularity.eps_pow_five_pos theorem eps_pos (hPε : 100 ≤ (4 : ℝ) ^ P.parts.card * ε ^ 5) : 0 < ε := - (Odd.pow_pos_iff (by norm_num)).mp (eps_pow_five_pos hPε) + (Odd.pow_pos_iff (by decide)).mp (eps_pow_five_pos hPε) #align szemeredi_regularity.eps_pos SzemerediRegularity.eps_pos theorem hundred_div_ε_pow_five_le_m [Nonempty α] (hPα : P.parts.card * 16 ^ P.parts.card ≤ card α) diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean index 0ec26e2331238..3cf2eb73c6258 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Chunk.lean @@ -39,8 +39,6 @@ Once ported to mathlib4, this file will be a great golfing ground for Heather's open Finpartition Finset Fintype Rel Nat -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped BigOperators Classical SzemerediRegularity.Positivity namespace SzemerediRegularity diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean index d43461750de92..bb0328a90275a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Energy.lean @@ -32,8 +32,6 @@ open BigOperators variable {α : Type*} [DecidableEq α] {s : Finset α} (P : Finpartition s) (G : SimpleGraph α) [DecidableRel G.Adj] -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace Finpartition /-- The energy of a partition, also known as index. Auxiliary quantity for Szemerédi's regularity diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean index d16579fa938ef..6c121642788fe 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Increment.lean @@ -38,8 +38,6 @@ Once ported to mathlib4, this file will be a great golfing ground for Heather's open Finset Fintype SimpleGraph SzemerediRegularity -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped BigOperators Classical SzemerediRegularity.Positivity variable {α : Type*} [Fintype α] {P : Finpartition (univ : Finset α)} (hP : P.IsEquipartition) @@ -116,8 +114,8 @@ theorem offDiag_pairs_le_increment_energy : refine' div_le_div_of_le_of_nonneg (α := ℚ) _ (sq_nonneg _) rw [← sum_biUnion] · exact sum_le_sum_of_subset_of_nonneg distinct_pairs_increment fun i _ _ => sq_nonneg _ - simp only [Set.PairwiseDisjoint, Function.onFun, disjoint_left, inf_eq_inter, mem_inter, - mem_product] + simp (config := { unfoldPartialApp := true }) only [Set.PairwiseDisjoint, Function.onFun, + disjoint_left, inf_eq_inter, mem_inter, mem_product] rintro ⟨⟨s₁, s₂⟩, hs⟩ _ ⟨⟨t₁, t₂⟩, ht⟩ _ hst ⟨u, v⟩ huv₁ huv₂ rw [mem_offDiag] at hs ht obtain ⟨a, ha⟩ := Finpartition.nonempty_of_mem_parts _ huv₁.1 diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean index 8a4ccf0de75c3..abceb6ec375e3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Lemma.lean @@ -68,8 +68,6 @@ open Finpartition Finset Fintype Function SzemerediRegularity open scoped Classical -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {α : Type*} [Fintype α] (G : SimpleGraph α) {ε : ℝ} {l : ℕ} /-- Effective **Szemerédi Regularity Lemma**: For any sufficiently large graph, there is an diff --git a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean index 8c39f7aa9a2b0..656e03e16d171 100644 --- a/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean +++ b/Mathlib/Combinatorics/SimpleGraph/StronglyRegular.lean @@ -225,7 +225,8 @@ theorem IsSRGWith.matrix_eq {α : Type*} [Semiring α] (h : G.IsSRGWith n k ℓ simp [commonNeighbors, ← neighborFinset_def, h.regular v] · simp only [Matrix.one_apply_ne' hn.symm, ne_eq, hn] by_cases ha : G.Adj v w <;> - simp only [ha, ite_true, ite_false, add_zero, zero_add, nsmul_eq_mul, smul_zero, mul_one] + simp only [ha, ite_true, ite_false, add_zero, zero_add, nsmul_eq_mul, smul_zero, mul_one, + not_true_eq_false, not_false_eq_true, and_false, and_self] · rw [h.of_adj v w ha] · rw [h.of_not_adj v w hn ha] diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 608d67ba074e7..ec26bfaaa6eda 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -499,7 +499,7 @@ theorem rowLens_length_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (hpo (ofRowLens w hw).rowLens.length = w.length := by simp only [length_rowLens, colLen, Nat.find_eq_iff, mem_cells, mem_ofRowLens, lt_self_iff_false, IsEmpty.exists_iff, Classical.not_not] - refine' ⟨True.intro, fun n hn => ⟨hn, hpos _ (List.get_mem _ _ hn)⟩⟩ + refine' ⟨not_false, fun n hn => ⟨hn, hpos _ (List.get_mem _ _ hn)⟩⟩ #align young_diagram.row_lens_length_of_row_lens YoungDiagram.rowLens_length_ofRowLens -- Porting note: use `List.get` instead of `List.nthLe` because it has been deprecated diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 6c2fdde4e3475..2a4625465eeec 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -101,7 +101,7 @@ theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by Nat.mul_sub_left_distrib, ← Nat.sub_add_comm, two_mul 3, Nat.add_sub_add_right] have H : 2 * 3 ≤ 2 * 2 ^ 3 := by norm_num apply H.trans - simp [pow_le_pow] + simp [pow_le_pow (show 1 ≤ 2 by norm_num)] #align ack_three ack_three theorem ack_pos : ∀ m n, 0 < ack m n diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 2a513d73756e4..9f60ef77c8412 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -378,7 +378,7 @@ theorem rfindOpt {n} {f : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) exists_congr fun a => (and_congr (iff_of_eq _) Iff.rfl).trans (and_congr_right fun h => _) · congr funext n - cases f (n ::ᵥ v) <;> simp [Nat.succ_le_succ]; rfl + cases f (n ::ᵥ v) <;> simp [Nat.succ_le_succ] <;> rfl · have := Nat.rfind_spec h simp only [Part.coe_some, Part.mem_some_iff] at this revert this; cases' f (a ::ᵥ v) with c <;> intro this diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index fd573c762d535..5edab43f57ebe 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -837,7 +837,7 @@ theorem nat_mod : Primrec₂ ((· % ·) : ℕ → ℕ → ℕ) := theorem nat_bodd : Primrec Nat.bodd := (Primrec.beq.comp (nat_mod.comp .id (const 2)) (const 1)).of_eq fun n => by - cases H : n.bodd <;> simp [Nat.mod_two_of_bodd, H] + cases H : n.bodd <;> simp [Nat.mod_two_of_bodd, H]; rfl #align primrec.nat_bodd Primrec.nat_bodd theorem nat_div2 : Primrec Nat.div2 := diff --git a/Mathlib/Computability/TMComputable.lean b/Mathlib/Computability/TMComputable.lean index d414d6d2a6bf7..a79be2891b674 100644 --- a/Mathlib/Computability/TMComputable.lean +++ b/Mathlib/Computability/TMComputable.lean @@ -260,7 +260,7 @@ def idComputableInPolyTime {α : Type} (ea : FinEncoding α) : outputsFun _ := { steps := 1 evals_in_steps := rfl - steps_le_m := by simp only [Polynomial.eval_one] } + steps_le_m := by simp only [Polynomial.eval_one, le_refl] } #align turing.id_computable_in_poly_time Turing.idComputableInPolyTime instance inhabitedTM2ComputableInPolyTime : diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 2c3872c83fc5f..d16613bcbfc9c 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -1536,7 +1536,7 @@ theorem succ_ok {q s n} {c d : List Γ'} : simp only [Option.mem_def, TM2.stepAux, elim_main, decide_False, elim_update_main, ne_eq, Function.update_noteq, elim_rev, elim_update_rev, decide_True, Function.update_same, cond_true, cond_false] - convert unrev_ok using 2 + convert unrev_ok using 1 simp only [elim_update_rev, elim_rev, elim_main, List.reverseAux_nil, elim_update_main] rfl simp only [trNum, Num.succ, Num.succ'] diff --git a/Mathlib/Control/Applicative.lean b/Mathlib/Control/Applicative.lean index 81241e776ecc4..b13edea4a0159 100644 --- a/Mathlib/Control/Applicative.lean +++ b/Mathlib/Control/Applicative.lean @@ -31,7 +31,8 @@ variable [Applicative F] [LawfulApplicative F] variable {α β γ σ : Type u} theorem Applicative.map_seq_map (f : α → β → γ) (g : σ → β) (x : F α) (y : F σ) : - f <$> x <*> g <$> y = (flip (· ∘ ·) g ∘ f) <$> x <*> y := by simp [flip, functor_norm] + f <$> x <*> g <$> y = (flip (· ∘ ·) g ∘ f) <$> x <*> y := by + simp [flip, functor_norm] #align applicative.map_seq_map Applicative.map_seq_map theorem Applicative.pure_seq_eq_map' (f : α → β) : (· <*> ·) (pure f : F (α → β)) = (· <$> ·) f := diff --git a/Mathlib/Control/Basic.lean b/Mathlib/Control/Basic.lean index 1818361be1f14..99f1c499ad700 100644 --- a/Mathlib/Control/Basic.lean +++ b/Mathlib/Control/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Init.Control.Combinators +import Mathlib.Init.Function import Mathlib.Tactic.CasesM import Mathlib.Tactic.Attr.Core import Std.Data.List.Basic @@ -87,26 +88,29 @@ theorem map_bind (x : m α) {g : α → m β} {f : β → γ} : theorem seq_bind_eq (x : m α) {g : β → m γ} {f : α → β} : f <$> x >>= g = x >>= g ∘ f := - show bind (f <$> x) g = bind x (g ∘ f) - by rw [← bind_pure_comp, bind_assoc]; simp [pure_bind, (· ∘ ·)] + show bind (f <$> x) g = bind x (g ∘ f) by + rw [← bind_pure_comp, bind_assoc] + simp [pure_bind, Function.comp_def] #align seq_bind_eq seq_bind_eq #align seq_eq_bind_map seq_eq_bind_mapₓ -- order of implicits and `Seq.seq` has a lazily evaluated second argument using `Unit` @[functor_norm] -theorem fish_pure {α β} (f : α → m β) : f >=> pure = f := by simp only [(· >=> ·), functor_norm] +theorem fish_pure {α β} (f : α → m β) : f >=> pure = f := by + simp (config := { unfoldPartialApp := true }) only [(· >=> ·), functor_norm] #align fish_pure fish_pure @[functor_norm] -theorem fish_pipe {α β} (f : α → m β) : pure >=> f = f := by simp only [(· >=> ·), functor_norm] +theorem fish_pipe {α β} (f : α → m β) : pure >=> f = f := by + simp (config := { unfoldPartialApp := true }) only [(· >=> ·), functor_norm] #align fish_pipe fish_pipe -- note: in Lean 3 `>=>` is left-associative, but in Lean 4 it is right-associative. @[functor_norm] theorem fish_assoc {α β γ φ} (f : α → m β) (g : β → m γ) (h : γ → m φ) : (f >=> g) >=> h = f >=> g >=> h := by - simp only [(· >=> ·), functor_norm] + simp (config := { unfoldPartialApp := true }) only [(· >=> ·), functor_norm] #align fish_assoc fish_assoc variable {β' γ' : Type v} diff --git a/Mathlib/Control/Functor/Multivariate.lean b/Mathlib/Control/Functor/Multivariate.lean index 6337351fee9b7..e093beb4c4626 100644 --- a/Mathlib/Control/Functor/Multivariate.lean +++ b/Mathlib/Control/Functor/Multivariate.lean @@ -144,8 +144,9 @@ theorem LiftR_def (x y : F α) : ∃ u : F (Subtype_ R), (TypeVec.prod.fst ⊚ subtypeVal R) <$$> u = x ∧ (TypeVec.prod.snd ⊚ subtypeVal R) <$$> u = y := - exists_iff_exists_of_mono _ _ _ (toSubtype'_of_subtype' R) - (by simp only [map_map, comp_assoc, subtypeVal_toSubtype']; simp [comp]) + exists_iff_exists_of_mono _ _ _ (toSubtype'_of_subtype' R) (by + simp only [map_map, comp_assoc, subtypeVal_toSubtype'] + simp (config := { unfoldPartialApp := true }) [comp]) #align mvfunctor.liftr_def MvFunctor.LiftR_def end LiftP' @@ -192,7 +193,7 @@ theorem LiftP_PredLast_iff {β} (P : β → Prop) (x : F (α ::: β)) : cases i <;> rfl · intros rw [MvFunctor.map_map] - dsimp [(· ⊚ ·)] + dsimp (config := { unfoldPartialApp := true }) [(· ⊚ ·)] suffices (fun i => Subtype.val) = (fun i x => (MvFunctor.f P n α i x).val) by rw[this]; ext i ⟨x, _⟩ @@ -228,7 +229,7 @@ theorem LiftR_RelLast_iff (x y : F (α ::: β)) : · ext i ⟨x, _⟩ : 2 cases i <;> rfl · intros - simp [MvFunctor.map_map, (· ⊚ ·)] + simp (config := { unfoldPartialApp := true }) [MvFunctor.map_map, (· ⊚ ·)] -- porting note: proof was -- rw [MvFunctor.map_map, MvFunctor.map_map, (· ⊚ ·), (· ⊚ ·)] -- congr <;> ext i ⟨x, _⟩ <;> cases i <;> rfl diff --git a/Mathlib/Control/Traversable/Equiv.lean b/Mathlib/Control/Traversable/Equiv.lean index 765559adb0946..8605e1f6c68de 100644 --- a/Mathlib/Control/Traversable/Equiv.lean +++ b/Mathlib/Control/Traversable/Equiv.lean @@ -143,7 +143,10 @@ protected theorem traverse_eq_map_id (f : α → β) (x : t' α) : protected theorem comp_traverse (f : β → F γ) (g : α → G β) (x : t' α) : Equiv.traverse eqv (Comp.mk ∘ Functor.map f ∘ g) x = Comp.mk (Equiv.traverse eqv f <$> Equiv.traverse eqv g x) := by - simp [Equiv.traverse, comp_traverse, functor_norm]; congr; ext; simp + simp (config := { unfoldPartialApp := true }) [Equiv.traverse, comp_traverse, functor_norm] + congr + ext + simp #align equiv.comp_traverse Equiv.comp_traverse protected theorem naturality (f : α → F β) (x : t' α) : diff --git a/Mathlib/Data/Bitvec/Lemmas.lean b/Mathlib/Data/Bitvec/Lemmas.lean index bc995d403f452..9e961fabc48f9 100644 --- a/Mathlib/Data/Bitvec/Lemmas.lean +++ b/Mathlib/Data/Bitvec/Lemmas.lean @@ -102,7 +102,7 @@ theorem toNat_lt {n : ℕ} (v : Bitvec n) : v.toNat < 2 ^ n := by -- Porting note: removed `ac_mono`, `mono` calls · rw [add_assoc] apply Nat.add_le_add_left - cases head <;> simp only + cases head <;> decide · rw [← left_distrib] rw [mul_comm _ 2] apply Nat.mul_le_mul_left @@ -112,14 +112,14 @@ theorem toNat_lt {n : ℕ} (v : Bitvec n) : v.toNat < 2 ^ n := by theorem addLsb_div_two {x b} : addLsb x b / 2 = x := by cases b <;> simp only [Nat.add_mul_div_left, addLsb, ← two_mul, add_comm, Nat.succ_pos', - Nat.mul_div_right, gt_iff_lt, zero_add, cond] + Nat.mul_div_right, gt_iff_lt, zero_add, zero_lt_two, cond] norm_num #align bitvec.add_lsb_div_two Bitvec.addLsb_div_two theorem decide_addLsb_mod_two {x b} : decide (addLsb x b % 2 = 1) = b := by cases b <;> simp only [Bool.decide_iff, Nat.add_mul_mod_self_left, addLsb, ← two_mul, add_comm, - Bool.decide_False, Nat.mul_mod_right, zero_add, cond, zero_ne_one] + Bool.decide_False, Nat.mul_mod_right, zero_add, cond, zero_ne_one]; rfl #align bitvec.to_bool_add_lsb_mod_two Bitvec.decide_addLsb_mod_two theorem ofNat_toNat {n : ℕ} (v : Bitvec n) : Bitvec.ofNat n v.toNat = v := by diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 2d8c69323c4f2..da6887305cd23 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -308,7 +308,7 @@ def toNat (b : Bool) : Nat := #align bool.to_nat Bool.toNat lemma toNat_le_one (b : Bool) : b.toNat ≤ 1 := by - cases b <;> simp only + cases b <;> decide /-- convert a `ℕ` to a `Bool`, `0 -> false`, everything else -> `true` -/ def ofNat (n : Nat) : Bool := @@ -334,7 +334,7 @@ theorem ofNat_le_ofNat {n m : Nat} (h : n ≤ m) : ofNat n ≤ ofNat m := by #align bool.of_nat_le_of_nat Bool.ofNat_le_ofNat theorem toNat_le_toNat {b₀ b₁ : Bool} (h : b₀ ≤ b₁) : toNat b₀ ≤ toNat b₁ := by - cases b₀ <;> cases b₁ <;> simp_all + cases b₀ <;> cases b₁ <;> simp_all (config := { decide := true }) #align bool.to_nat_le_to_nat Bool.toNat_le_toNat theorem ofNat_toNat (b : Bool) : ofNat (toNat b) = b := by diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 911de876c0eeb..0548d2211f081 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -569,7 +569,7 @@ theorem exp_conj : exp (conj x) = conj (exp x) := by dsimp [exp] rw [← lim_conj] refine' congr_arg CauSeq.lim (CauSeq.ext fun _ => _) - dsimp [exp', Function.comp, isCauSeq_conj, cauSeqConj] + dsimp [exp', Function.comp_def, isCauSeq_conj, cauSeqConj] rw [(starRingEnd _).map_sum] refine' sum_congr rfl fun n _ => _ rw [map_div₀, map_pow, ← ofReal_nat_cast, conj_ofReal] diff --git a/Mathlib/Data/DFinsupp/WellFounded.lean b/Mathlib/Data/DFinsupp/WellFounded.lean index 390889f8478fb..66fdc04af129a 100644 --- a/Mathlib/Data/DFinsupp/WellFounded.lean +++ b/Mathlib/Data/DFinsupp/WellFounded.lean @@ -224,7 +224,7 @@ protected theorem DFinsupp.wellFoundedLT [∀ i, Zero (α i)] [∀ i, Preorder ( refine Lex.wellFounded' ?_ (fun i ↦ IsWellFounded.wf) ?_ · rintro i ⟨a⟩ apply hbot - · simp only [Function.swap] + · simp (config := { unfoldPartialApp := true }) only [Function.swap] exact IsWellFounded.wf refine Subrelation.wf (fun h => ?_) <| InvImage.wf (mapRange (fun i ↦ e i) fun _ ↦ rfl) this have := IsStrictOrder.swap (@WellOrderingRel ι) diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index fcb273d3dc9bb..bcb95ce0da54e 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -149,13 +149,17 @@ theorem coe_ne_top (a : ℕ) : (a : ℕ∞) ≠ ⊤ := theorem top_sub_coe (a : ℕ) : (⊤ : ℕ∞) - a = ⊤ := WithTop.top_sub_coe +@[simp] +theorem zero_lt_top : (0 : ℕ∞) < ⊤ := + WithTop.zero_lt_top + --Porting note: new theorem copied from `WithTop` theorem sub_top (a : ℕ∞) : a - ⊤ = 0 := WithTop.sub_top @[simp] theorem coe_toNat_eq_self : ENat.toNat (n : ℕ∞) = n ↔ n ≠ ⊤ := - ENat.recTopCoe (by simp) (fun _ => by simp [toNat_coe]) n + ENat.recTopCoe (by decide) (fun _ => by simp [toNat_coe]) n #align enat.coe_to_nat_eq_self ENat.coe_toNat_eq_self alias ⟨_, coe_toNat⟩ := coe_toNat_eq_self diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 1f59a407d1e4d..7bf9f8b131ac9 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -506,7 +506,7 @@ theorem val_one'' {n : ℕ} : ((1 : Fin (n + 1)) : ℕ) = 1 % (n + 1) := #align fin.mk_one Fin.mk_one instance nontrivial {n : ℕ} : Nontrivial (Fin (n + 2)) where - exists_pair_ne := ⟨0, 1, (ne_iff_vne 0 1).mpr (by simp only [val_one, val_zero])⟩ + exists_pair_ne := ⟨0, 1, (ne_iff_vne 0 1).mpr (by simp [val_one, val_zero])⟩ theorem nontrivial_iff_two_le : Nontrivial (Fin n) ↔ 2 ≤ n := by rcases n with (_ | _ | n) <;> @@ -1231,7 +1231,7 @@ theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a #align fin.coe_fin_one Fin.coe_fin_one lemma eq_one_of_neq_zero (i : Fin 2) (hi : i ≠ 0) : i = 1 := - fin_two_eq_of_eq_zero_iff (by simpa only [iff_false] using hi) + fin_two_eq_of_eq_zero_iff (by simpa only [one_eq_zero_iff, succ.injEq, iff_false] using hi) @[simp] theorem coe_neg_one : ↑(-1 : Fin (n + 1)) = n := by @@ -1280,7 +1280,7 @@ theorem coe_sub_iff_lt {n : ℕ} {a b : Fin n} : (↑(a - b) : ℕ) = n + a - b theorem lt_sub_one_iff {n : ℕ} {k : Fin (n + 2)} : k < k - 1 ↔ k = 0 := by rcases k with ⟨_ | k, hk⟩ simp only [zero_eq, zero_eta, zero_sub, lt_iff_val_lt_val, val_zero, coe_neg_one, add_pos_iff, - or_true] + _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] simp [lt_iff_val_lt_val, ext_iff, Fin.coe_sub, succ_eq_add_one, this, diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 754a9aca44800..1065b88d7032e 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -67,7 +67,8 @@ def cons (x : α 0) (p : ∀ i : Fin n, α i.succ) : ∀ i, α i := fun j ↦ Fi #align fin.cons Fin.cons @[simp] -theorem tail_cons : tail (cons x p) = p := by simp [tail, cons] +theorem tail_cons : tail (cons x p) = p := by + simp (config := { unfoldPartialApp := true }) [tail, cons] #align fin.tail_cons Fin.tail_cons @[simp] @@ -373,8 +374,9 @@ theorem cons_eq_append {α : Type*} (x : α) (xs : Fin n → α) : Fin.append (xs ∘ Fin.cast h) ys = Fin.append xs ys ∘ (Fin.cast <| by rw[h]) := by subst h funext i - simp only [Fin.append, Fin.addCases, comp, Fin.cast, Fin.coe_castLT, Fin.subNat_mk, Fin.natAdd_mk, - ge_iff_le, eq_rec_constant, Fin.eta, Eq.ndrec, id_eq, eq_mpr_eq_cast, cast_eq] + simp (config := {unfoldPartialApp := true}) only [Fin.append, Fin.addCases, comp_def, Fin.cast, + Fin.coe_castLT, Fin.subNat_mk, Fin.natAdd_mk, ge_iff_le, eq_rec_constant, Fin.eta, Eq.ndrec, + id_eq, eq_mpr_eq_cast, cast_eq] @[simp] lemma append_cast_right {n m} {α : Type*} (xs : Fin n → α) (ys : Fin m → α) (m' : ℕ) (h : m' = m) : diff --git a/Mathlib/Data/Fin/Tuple/Curry.lean b/Mathlib/Data/Fin/Tuple/Curry.lean index f243c1ef2bd26..bd3861aacecf1 100644 --- a/Mathlib/Data/Fin/Tuple/Curry.lean +++ b/Mathlib/Data/Fin/Tuple/Curry.lean @@ -46,7 +46,7 @@ theorem curry_uncurry {n} (f : Function.OfArity α β n) : match n with | 0 => rfl | n + 1 => funext fun a => by - dsimp [curry, uncurry, Function.comp] + dsimp [curry, uncurry, Function.comp_def] simp only [Fin.cons_zero, Fin.cons_succ] rw [curry_uncurry] @@ -57,7 +57,7 @@ theorem uncurry_curry {n} (f : (Fin n → α) → β) : induction args using Fin.consInduction with | h0 => rfl | h a as ih => - dsimp [curry, uncurry, Function.comp] + dsimp [curry, uncurry, Function.comp_def] simp only [Fin.cons_zero, Fin.cons_succ, ih (f <| Fin.cons a ·)] /-- `Equiv.curry` for n-ary functions. -/ diff --git a/Mathlib/Data/Fin/Tuple/Monotone.lean b/Mathlib/Data/Fin/Tuple/Monotone.lean index 7c48630f7e14b..51d4c6d892045 100644 --- a/Mathlib/Data/Fin/Tuple/Monotone.lean +++ b/Mathlib/Data/Fin/Tuple/Monotone.lean @@ -38,12 +38,12 @@ theorem monotone_vecCons : Monotone (vecCons a f) ↔ a ≤ f 0 ∧ Monotone f : --Porting note: new lemma, in Lean3 would be proven by `Subsingleton.monotone` @[simp] -theorem monotone_vecEmpty : Monotone (vecCons a vecEmpty) +theorem monotone_vecEmpty : Monotone ![a] | ⟨0, _⟩, ⟨0, _⟩, _ => le_refl _ --Porting note: new lemma, in Lean3 would be proven by `Subsingleton.strictMono` @[simp] -theorem strictMono_vecEmpty : StrictMono (vecCons a vecEmpty) +theorem strictMono_vecEmpty : StrictMono ![a] | ⟨0, _⟩, ⟨0, _⟩, h => (irrefl _ h).elim @[simp] @@ -82,4 +82,5 @@ theorem Antitone.vecCons (hf : Antitone f) (ha : f 0 ≤ a) : Antitone (vecCons antitone_vecCons.2 ⟨ha, hf⟩ #align antitone.vec_cons Antitone.vecCons -example : Monotone ![1, 2, 2, 3] := by simp +-- NOTE: was "by simp", but simp lemmas were not being used +example : Monotone ![1, 2, 2, 3] := by decide diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index eabbc1fd9527a..499a3fa7d53f3 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -83,7 +83,7 @@ theorem mem_antidiagonalTuple {n : ℕ} {k : ℕ} {x : Fin k → ℕ} : induction x using Fin.consInduction generalizing n with | h0 => cases n - · simp + · decide · simp [eq_comm] | h x₀ x ih => simp_rw [Fin.sum_cons] diff --git a/Mathlib/Data/Finset/NAry.lean b/Mathlib/Data/Finset/NAry.lean index 9a70d9a8ab3da..0737cfda6822e 100644 --- a/Mathlib/Data/Finset/NAry.lean +++ b/Mathlib/Data/Finset/NAry.lean @@ -336,7 +336,9 @@ theorem image₂_mk_eq_product [DecidableEq α] [DecidableEq β] (s : Finset α) @[simp] theorem image₂_curry (f : α × β → γ) (s : Finset α) (t : Finset β) : image₂ (curry f) s t = (s ×ˢ t).image f := by - classical rw [← image₂_mk_eq_product, image_image₂]; dsimp [curry] + classical + rw [← image₂_mk_eq_product, image_image₂] + dsimp (config := { unfoldPartialApp := true }) [curry] #align finset.image₂_curry Finset.image₂_curry @[simp] diff --git a/Mathlib/Data/Finsupp/BigOperators.lean b/Mathlib/Data/Finsupp/BigOperators.lean index c27ae4e6e795f..a9a918ec5dcd5 100644 --- a/Mathlib/Data/Finsupp/BigOperators.lean +++ b/Mathlib/Data/Finsupp/BigOperators.lean @@ -103,7 +103,7 @@ theorem Multiset.support_sum_eq [AddCommMonoid M] (s : Multiset (ι →₀ M)) suffices : a.Pairwise (_root_.Disjoint on Finsupp.support) · convert List.support_sum_eq a this · simp only [Multiset.quot_mk_to_coe'', Multiset.coe_sum] - · dsimp only [Function.comp] + · dsimp only [Function.comp_def] simp only [quot_mk_to_coe'', coe_map, sup_coe, ge_iff_le, Finset.le_eq_subset, Finset.sup_eq_union, Finset.bot_eq_empty, List.foldr_map] · simp only [Multiset.quot_mk_to_coe'', Multiset.coe_map, Multiset.coe_eq_coe] at hl diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index f2d0ebf8d5a48..ea8a6c17d39ff 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -835,7 +835,7 @@ def embDomain (f : α ↪ β) (v : α →₀ M) : β →₀ M where · simp only [h, true_iff_iff, Ne.def] rw [← not_mem_support_iff, not_not] classical apply Finset.choose_mem - · simp only [h, Ne.def, ne_self_iff_false] + · simp only [h, Ne.def, ne_self_iff_false, not_true_eq_false] #align finsupp.emb_domain Finsupp.embDomain @[simp] diff --git a/Mathlib/Data/Holor.lean b/Mathlib/Data/Holor.lean index 27d884661c3f4..eed50ec61d359 100644 --- a/Mathlib/Data/Holor.lean +++ b/Mathlib/Data/Holor.lean @@ -225,7 +225,8 @@ nonrec theorem mul_zero {α : Type} [Ring α] (x : Holor α ds₁) : x ⊗ (0 : theorem mul_scalar_mul [Monoid α] (x : Holor α []) (y : Holor α ds) : x ⊗ y = x ⟨[], Forall₂.nil⟩ • y := by - simp [mul, SMul.smul, HolorIndex.take, HolorIndex.drop, HSMul.hSMul] + simp (config := { unfoldPartialApp := true }) [mul, SMul.smul, HolorIndex.take, HolorIndex.drop, + HSMul.hSMul] #align holor.mul_scalar_mul Holor.mul_scalar_mul -- holor slices diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index 5df6f41d88edc..f776c5598e108 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -53,7 +53,7 @@ theorem bodd_subNatNat (m n : ℕ) : bodd (subNatNat m n) = xor m.bodd n.bodd := @[simp] theorem bodd_negOfNat (n : ℕ) : bodd (negOfNat n) = n.bodd := by - cases n <;> simp + cases n <;> simp (config := {decide := true}) rfl #align int.bodd_neg_of_nat Int.bodd_negOfNat @@ -164,7 +164,7 @@ theorem bit_negSucc (b) (n : ℕ) : bit b -[n+1] = -[Nat.bit (not b) n+1] := by @[simp] theorem bodd_bit (b n) : bodd (bit b n) = b := by rw [bit_val] - cases b <;> cases bodd n <;> simp + cases b <;> cases bodd n <;> simp [(show bodd 2 = false by rfl)] #align int.bodd_bit Int.bodd_bit @[simp, deprecated] @@ -293,12 +293,12 @@ theorem bitwise_xor : bitwise xor = Int.xor := by theorem bitwise_bit (f : Bool → Bool → Bool) (a m b n) : bitwise f (bit a m) (bit b n) = bit (f a b) (bitwise f m n) := by cases' m with m m <;> cases' n with n n <;> - simp only [bitwise, ofNat_eq_coe, bit_coe_nat, natBitwise, Bool.not_false, Bool.not_eq_false', + simp [bitwise, ofNat_eq_coe, bit_coe_nat, natBitwise, Bool.not_false, Bool.not_eq_false', bit_negSucc] - · by_cases h : f false false <;> simp [h] - · by_cases h : f false true <;> simp [h] - · by_cases h : f true false <;> simp [h] - · by_cases h : f true true <;> simp [h] + · by_cases h : f false false <;> simp (config := {decide := true}) [h] + · by_cases h : f false true <;> simp (config := {decide := true}) [h] + · by_cases h : f true false <;> simp (config := {decide := true}) [h] + · by_cases h : f true true <;> simp (config := {decide := true}) [h] #align int.bitwise_bit Int.bitwise_bit @[simp] @@ -432,12 +432,12 @@ theorem shiftLeft_sub (m : ℤ) (n : ℕ) (k : ℤ) : m <<< (n - k) = (m <<< (n shiftLeft_add _ _ _ #align int.shiftl_sub Int.shiftLeft_sub -theorem shiftLeft_eq_mul_pow : ∀ (m : ℤ) (n : ℕ), m <<< (n : ℤ) = m * ↑(2 ^ n) +theorem shiftLeft_eq_mul_pow : ∀ (m : ℤ) (n : ℕ), m <<< (n : ℤ) = m * (2 ^ n : ℕ) | (m : ℕ), _ => congr_arg ((↑) : ℕ → ℤ) (by simp) | -[_+1], _ => @congr_arg ℕ ℤ _ _ (fun i => -i) (Nat.shiftLeft'_tt_eq_mul_pow _ _) #align int.shiftl_eq_mul_pow Int.shiftLeft_eq_mul_pow -theorem shiftRight_eq_div_pow : ∀ (m : ℤ) (n : ℕ), m >>> (n : ℤ) = m / ↑(2 ^ n) +theorem shiftRight_eq_div_pow : ∀ (m : ℤ) (n : ℕ), m >>> (n : ℤ) = m / (2 ^ n : ℕ) | (m : ℕ), n => by rw [shiftRight_coe_nat, Nat.shiftRight_eq_div_pow _ _]; simp | -[m+1], n => by rw [shiftRight_negSucc, negSucc_ediv, Nat.shiftRight_eq_div_pow]; rfl diff --git a/Mathlib/Data/Int/Lemmas.lean b/Mathlib/Data/Int/Lemmas.lean index 55d5ca346a1af..5fa5f5943d0b9 100644 --- a/Mathlib/Data/Int/Lemmas.lean +++ b/Mathlib/Data/Int/Lemmas.lean @@ -122,7 +122,7 @@ attribute [local simp] Int.zero_div theorem div2_bit (b n) : div2 (bit b n) = n := by rw [bit_val, div2_val, add_comm, Int.add_mul_ediv_left, (_ : (_ / 2 : ℤ) = 0), zero_add] cases b - · simp + · decide · show ofNat _ = _ rw [Nat.div_eq_of_lt] <;> simp · decide diff --git a/Mathlib/Data/Int/ModEq.lean b/Mathlib/Data/Int/ModEq.lean index a446a520cd44c..b0b3171818481 100644 --- a/Mathlib/Data/Int/ModEq.lean +++ b/Mathlib/Data/Int/ModEq.lean @@ -129,7 +129,7 @@ protected theorem mul_left' (h : a ≡ b [ZMOD n]) : c * a ≡ c * b [ZMOD c * n obtain hc | rfl | hc := lt_trichotomy c 0 · rw [← neg_modEq_neg, ← modEq_neg, ← neg_mul, ← neg_mul, ← neg_mul] simp only [ModEq, mul_emod_mul_of_pos _ _ (neg_pos.2 hc), h.eq] - · simp + · simp only [zero_mul]; rfl · simp only [ModEq, mul_emod_mul_of_pos _ _ hc, h.eq] #align int.modeq.mul_left' Int.ModEq.mul_left' diff --git a/Mathlib/Data/Int/Parity.lean b/Mathlib/Data/Int/Parity.lean index 6965c6a9c0470..d1961d01dd35f 100644 --- a/Mathlib/Data/Int/Parity.lean +++ b/Mathlib/Data/Int/Parity.lean @@ -27,6 +27,9 @@ theorem emod_two_ne_one : ¬n % 2 = 1 ↔ n % 2 = 0 := by cases' emod_two_eq_zero_or_one n with h h <;> simp [h] #align int.mod_two_ne_one Int.emod_two_ne_one +@[simp] +theorem one_emod_two : (1 : Int) % 2 = 1 := rfl + -- `EuclideanDomain.mod_eq_zero` uses (2 ∣ n) as normal form @[local simp] theorem emod_two_ne_zero : ¬n % 2 = 0 ↔ n % 2 = 1 := by @@ -102,6 +105,7 @@ theorem even_add : Even (m + n) ↔ (Even m ↔ Even n) := by cases' emod_two_eq_zero_or_one m with h₁ h₁ <;> cases' emod_two_eq_zero_or_one n with h₂ h₂ <;> simp [even_iff, h₁, h₂, Int.add_emod] + rfl #align int.even_add Int.even_add theorem even_add' : Even (m + n) ↔ (Odd m ↔ Odd n) := by @@ -279,8 +283,8 @@ theorem two_mul_ediv_two_of_odd (h : Odd n) : 2 * (n / 2) = n - 1 := -- Here are examples of how `parity_simps` can be used with `Int`. example (m n : ℤ) (h : Even m) : ¬Even (n + 3) ↔ Even (m ^ 2 + m + n) := by - simp [*, (by decide : ¬2 = 0), parity_simps] + simp (config := {decide := true}) [*, (by decide : ¬2 = 0), parity_simps] -example : ¬Even (25394535 : ℤ) := by simp +example : ¬Even (25394535 : ℤ) := by decide end Int diff --git a/Mathlib/Data/Int/Range.lean b/Mathlib/Data/Int/Range.lean index 91905365d8e9e..cc7e2b9ac47e1 100644 --- a/Mathlib/Data/Int/Range.lean +++ b/Mathlib/Data/Int/Range.lean @@ -49,7 +49,7 @@ instance decidableLELE (P : Int → Prop) [DecidablePred P] (m n : ℤ) : apply decidable_of_iff (∀ r ∈ range m (n + 1), P r) apply Iff.intro <;> intros h _ _ · intro _; apply h - simp_all only [mem_range_iff, and_imp, lt_add_one_iff] + simp_all only [mem_range_iff, and_imp, and_self, lt_add_one_iff] · simp_all only [mem_range_iff, and_imp, lt_add_one_iff] #align int.decidable_le_le Int.decidableLELE diff --git a/Mathlib/Data/Int/Units.lean b/Mathlib/Data/Int/Units.lean index 8667113994db0..5a69a99b1033f 100644 --- a/Mathlib/Data/Int/Units.lean +++ b/Mathlib/Data/Int/Units.lean @@ -122,7 +122,7 @@ theorem isUnit_add_isUnit_eq_isUnit_add_isUnit {a b c d : ℤ} (ha : IsUnit a) ( rw [isUnit_iff] at ha hb hc hd cases ha <;> cases hb <;> cases hc <;> cases hd <;> subst a <;> subst b <;> subst c <;> subst d <;> - simp + simp (config := {decide := true}) #align int.is_unit_add_is_unit_eq_is_unit_add_is_unit Int.isUnit_add_isUnit_eq_isUnit_add_isUnit theorem eq_one_or_neg_one_of_mul_eq_neg_one {z w : ℤ} (h : z * w = -1) : z = 1 ∨ z = -1 := diff --git a/Mathlib/Data/KVMap.lean b/Mathlib/Data/KVMap.lean deleted file mode 100644 index a5f89e420df48..0000000000000 --- a/Mathlib/Data/KVMap.lean +++ /dev/null @@ -1,27 +0,0 @@ -/- -Copyright (c) 2022 Alex J. Best. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Alex J. Best --/ -import Lean.Data.KVMap - -/-! -# Additional functionality for `KVMap` --/ - -namespace Lean.KVMap - -/-- erase pairs whose names match the second argument from a list of -`Name × DataValue` pairs-/ -def eraseCore : List (Name × DataValue) → Name → List (Name × DataValue) - | l, n => List.filter (fun a ↦ a.1 ≠ n) l - -/-- erase an entry from the map -/ -def erase : KVMap → Name → KVMap - | ⟨m⟩, k => ⟨eraseCore m k⟩ - -/-- update a Boolean entry based on its current value. -/ -def updateBool (m : KVMap) (k : Name) (f : Bool → Bool) : KVMap := - m.insert k <| DataValue.ofBool <| f <| m.getBool k - -end Lean.KVMap diff --git a/Mathlib/Data/LazyList/Basic.lean b/Mathlib/Data/LazyList/Basic.lean index 0b10c7c2cf1b0..d8c9a4658c973 100644 --- a/Mathlib/Data/LazyList/Basic.lean +++ b/Mathlib/Data/LazyList/Basic.lean @@ -82,7 +82,7 @@ instance : LawfulTraversable LazyList := by · simp only [traverse, Equiv.traverse, listEquivLazyList, Equiv.coe_fn_mk, Equiv.coe_fn_symm_mk] induction' xs using LazyList.rec with _ tl ih _ ih · simp only [List.traverse, map_pure]; rfl - · have : tl.get.traverse f = ofList <$> tl.get.toList.traverse f := ih + · replace ih : tl.get.traverse f = ofList <$> tl.get.toList.traverse f := ih simp only [traverse._eq_2, ih, Functor.map_map, seq_map_assoc, toList, List.traverse, map_seq] · rfl · apply ih diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 2bef777009544..e6d1e8526e154 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -1752,7 +1752,9 @@ theorem map_join (f : α → β) (L : List (List α)) : map f (join L) = join (m theorem bind_ret_eq_map (f : α → β) (l : List α) : l.bind (List.ret ∘ f) = map f l := by unfold List.bind - induction l <;> simp [map, join, List.ret, cons_append, nil_append, *] at * + induction l <;> + simp (config := { unfoldPartialApp := true }) + [map, join, List.ret, cons_append, nil_append, *] at * assumption #align list.bind_ret_eq_map List.bind_ret_eq_map @@ -2576,7 +2578,7 @@ theorem nthLe_succ_scanl {i : ℕ} {h : i + 1 < (scanl f b l).length} : induction i generalizing b l with | zero => cases l - · simp only [length, zero_add, scanl_nil] at h + · simp only [length, zero_eq, lt_self_iff_false] at h · simp [scanl_cons, singleton_append, nthLe_zero_scanl, nthLe_cons] | succ i hi => cases l @@ -3334,7 +3336,7 @@ theorem reduceOption_map {l : List (Option α)} {f : α → β} : reduceOption (map (Option.map f) l) = map f (reduceOption l) := by induction' l with hd tl hl · simp only [reduceOption_nil, map_nil] - ·cases hd <;> + · cases hd <;> simpa [true_and_iff, Option.map_some', map, eq_self_iff_true, reduceOption_cons_of_some] using hl #align list.reduce_option_map List.reduceOption_map @@ -3346,7 +3348,7 @@ theorem reduceOption_append (l l' : List (Option α)) : theorem reduceOption_length_le (l : List (Option α)) : l.reduceOption.length ≤ l.length := by induction' l with hd tl hl - · simp only [reduceOption_nil, length] + · simp [reduceOption_nil, length] · cases hd · exact Nat.le_succ_of_le hl · simpa only [length, add_le_add_iff_right, reduceOption_cons_of_some] using hl diff --git a/Mathlib/Data/List/BigOperators/Basic.lean b/Mathlib/Data/List/BigOperators/Basic.lean index 0f434a689325e..0e16729b630e4 100644 --- a/Mathlib/Data/List/BigOperators/Basic.lean +++ b/Mathlib/Data/List/BigOperators/Basic.lean @@ -610,7 +610,7 @@ If desired, we could add a class stating that `default = 0`. /-- This relies on `default ℕ = 0`. -/ theorem headI_add_tail_sum (L : List ℕ) : L.headI + L.tail.sum = L.sum := by - cases L <;> simp + cases L <;> simp; rfl #align list.head_add_tail_sum List.headI_add_tail_sum /-- This relies on `default ℕ = 0`. -/ diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 24ac55928b814..368b6cdf504f9 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -354,7 +354,7 @@ theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ rw [← and_forall_succ, chain'_cons, chain'_iff_get] simp only [length_cons, get_cons_succ, Fin.zero_eta, get_cons_zero, zero_add, Fin.mk_one, get_cons_cons_one, succ_sub_succ_eq_sub, nonpos_iff_eq_zero, add_eq_zero_iff, and_false, - tsub_zero, add_pos_iff, or_true, forall_true_left, and_congr_right_iff] + tsub_zero, add_pos_iff, zero_lt_one, or_true, forall_true_left, and_congr_right_iff] dsimp [succ_sub_one] exact fun _ => ⟨fun h i hi => h i (Nat.lt_of_succ_lt_succ hi), fun h i hi => h i (Nat.succ_lt_succ hi)⟩ diff --git a/Mathlib/Data/List/DropRight.lean b/Mathlib/Data/List/DropRight.lean index 894eda788227d..ab4856e998b08 100644 --- a/Mathlib/Data/List/DropRight.lean +++ b/Mathlib/Data/List/DropRight.lean @@ -15,12 +15,12 @@ Taking or removing element from the tail end of a list ## Main definitions -- `dropRight n`: drop `n : ℕ` elements from the tail -- `takeRight n`: take `n : ℕ` elements from the tail -- `dropRightWhile p`: remove all the elements from the tail of a list until it finds the first - element for which `p : α → Bool` returns false. This element and everything before is returned. -- `takeRightWhile p`: Returns the longest terminal segment of a list for which `p : α → Bool` - returns true. +- `rdrop n`: drop `n : ℕ` elements from the tail +- `rtake n`: take `n : ℕ` elements from the tail +- `rdropWhile p`: remove all the elements from the tail of a list until it finds the first element + for which `p : α → Bool` returns false. This element and everything before is returned. +- `rtakeWhile p`: Returns the longest terminal segment of a list for which `p : α → Bool` returns + true. ## Implementation detail @@ -38,113 +38,109 @@ variable {α : Type*} (p : α → Bool) (l : List α) (n : ℕ) namespace List /-- Drop `n` elements from the tail end of a list. -/ -def dropRight : List α := +def rdrop : List α := l.take (l.length - n) -#align list.rdrop List.dropRight +#align list.rdrop List.rdrop @[simp] -theorem dropRight_nil : dropRight ([] : List α) n = [] := by simp [dropRight] -#align list.rdrop_nil List.dropRight_nil +theorem rdrop_nil : rdrop ([] : List α) n = [] := by simp [rdrop] +#align list.rdrop_nil List.rdrop_nil @[simp] -theorem dropRight_zero : dropRight l 0 = l := by simp [dropRight] -#align list.rdrop_zero List.dropRight_zero +theorem rdrop_zero : rdrop l 0 = l := by simp [rdrop] +#align list.rdrop_zero List.rdrop_zero -theorem dropRight_eq_reverse_drop_reverse : l.dropRight n = reverse (l.reverse.drop n) := by - rw [dropRight] +theorem rdrop_eq_reverse_drop_reverse : l.rdrop n = reverse (l.reverse.drop n) := by + rw [rdrop] induction' l using List.reverseRecOn with xs x IH generalizing n · simp · cases n · simp [take_append] · simp [take_append_eq_append_take, IH] -#align list.rdrop_eq_reverse_drop_reverse List.dropRight_eq_reverse_drop_reverse +#align list.rdrop_eq_reverse_drop_reverse List.rdrop_eq_reverse_drop_reverse @[simp] -theorem dropRight_concat_succ (x : α) : dropRight (l ++ [x]) (n + 1) = dropRight l n := by - simp [dropRight_eq_reverse_drop_reverse] -#align list.rdrop_concat_succ List.dropRight_concat_succ +theorem rdrop_concat_succ (x : α) : rdrop (l ++ [x]) (n + 1) = rdrop l n := by + simp [rdrop_eq_reverse_drop_reverse] +#align list.rdrop_concat_succ List.rdrop_concat_succ /-- Take `n` elements from the tail end of a list. -/ -def takeRight : List α := +def rtake : List α := l.drop (l.length - n) -#align list.rtake List.takeRight +#align list.rtake List.rtake @[simp] -theorem takeRight_nil : takeRight ([] : List α) n = [] := by simp [takeRight] -#align list.rtake_nil List.takeRight_nil +theorem rtake_nil : rtake ([] : List α) n = [] := by simp [rtake] +#align list.rtake_nil List.rtake_nil @[simp] -theorem takeRight_zero : takeRight l 0 = [] := by simp [takeRight] -#align list.rtake_zero List.takeRight_zero +theorem rtake_zero : rtake l 0 = [] := by simp [rtake] +#align list.rtake_zero List.rtake_zero -theorem takeRight_eq_reverse_take_reverse : l.takeRight n = reverse (l.reverse.take n) := by - rw [takeRight] +theorem rtake_eq_reverse_take_reverse : l.rtake n = reverse (l.reverse.take n) := by + rw [rtake] induction' l using List.reverseRecOn with xs x IH generalizing n · simp · cases n · exact drop_length _ · simp [drop_append_eq_append_drop, IH] -#align list.rtake_eq_reverse_take_reverse List.takeRight_eq_reverse_take_reverse +#align list.rtake_eq_reverse_take_reverse List.rtake_eq_reverse_take_reverse @[simp] -theorem takeRight_concat_succ (x : α) : takeRight (l ++ [x]) (n + 1) = takeRight l n ++ [x] := by - simp [takeRight_eq_reverse_take_reverse] -#align list.rtake_concat_succ List.takeRight_concat_succ +theorem rtake_concat_succ (x : α) : rtake (l ++ [x]) (n + 1) = rtake l n ++ [x] := by + simp [rtake_eq_reverse_take_reverse] +#align list.rtake_concat_succ List.rtake_concat_succ /-- Drop elements from the tail end of a list that satisfy `p : α → Bool`. Implemented naively via `List.reverse` -/ -def dropRightWhile : List α := +def rdropWhile : List α := reverse (l.reverse.dropWhile p) -#align list.rdrop_while List.dropRightWhile +#align list.rdrop_while List.rdropWhile @[simp] -theorem dropRightWhile_nil : dropRightWhile p ([] : List α) = [] := by - simp [dropRightWhile, dropWhile] -#align list.rdrop_while_nil List.dropRightWhile_nil +theorem rdropWhile_nil : rdropWhile p ([] : List α) = [] := by simp [rdropWhile, dropWhile] +#align list.rdrop_while_nil List.rdropWhile_nil -theorem dropRightWhile_concat (x : α) : - dropRightWhile p (l ++ [x]) = if p x then dropRightWhile p l else l ++ [x] := by - simp only [dropRightWhile, dropWhile, reverse_append, reverse_singleton, singleton_append] +theorem rdropWhile_concat (x : α) : + rdropWhile p (l ++ [x]) = if p x then rdropWhile p l else l ++ [x] := by + simp only [rdropWhile, dropWhile, reverse_append, reverse_singleton, singleton_append] split_ifs with h <;> simp [h] -#align list.rdrop_while_concat List.dropRightWhile_concat +#align list.rdrop_while_concat List.rdropWhile_concat @[simp] -theorem dropRightWhile_concat_pos (x : α) (h : p x) : - dropRightWhile p (l ++ [x]) = dropRightWhile p l := by - rw [dropRightWhile_concat, if_pos h] -#align list.rdrop_while_concat_pos List.dropRightWhile_concat_pos +theorem rdropWhile_concat_pos (x : α) (h : p x) : rdropWhile p (l ++ [x]) = rdropWhile p l := by + rw [rdropWhile_concat, if_pos h] +#align list.rdrop_while_concat_pos List.rdropWhile_concat_pos @[simp] -theorem dropRightWhile_concat_neg (x : α) (h : ¬p x) : dropRightWhile p (l ++ [x]) = l ++ [x] := by - rw [dropRightWhile_concat, if_neg h] -#align list.rdrop_while_concat_neg List.dropRightWhile_concat_neg +theorem rdropWhile_concat_neg (x : α) (h : ¬p x) : rdropWhile p (l ++ [x]) = l ++ [x] := by + rw [rdropWhile_concat, if_neg h] +#align list.rdrop_while_concat_neg List.rdropWhile_concat_neg -theorem dropRightWhile_singleton (x : α) : dropRightWhile p [x] = if p x then [] else [x] := by - rw [← nil_append [x], dropRightWhile_concat, dropRightWhile_nil] -#align list.rdrop_while_singleton List.dropRightWhile_singleton +theorem rdropWhile_singleton (x : α) : rdropWhile p [x] = if p x then [] else [x] := by + rw [← nil_append [x], rdropWhile_concat, rdropWhile_nil] +#align list.rdrop_while_singleton List.rdropWhile_singleton -theorem dropRightWhile_last_not (hl : l.dropRightWhile p ≠ []) : - ¬p ((dropRightWhile p l).getLast hl) := by - simp_rw [dropRightWhile] +theorem rdropWhile_last_not (hl : l.rdropWhile p ≠ []) : ¬p ((rdropWhile p l).getLast hl) := by + simp_rw [rdropWhile] rw [getLast_reverse] exact dropWhile_nthLe_zero_not _ _ _ -#align list.rdrop_while_last_not List.dropRightWhile_last_not +#align list.rdrop_while_last_not List.rdropWhile_last_not -theorem dropRightWhile_prefix : l.dropRightWhile p <+: l := by - rw [← reverse_suffix, dropRightWhile, reverse_reverse] +theorem rdropWhile_prefix : l.rdropWhile p <+: l := by + rw [← reverse_suffix, rdropWhile, reverse_reverse] exact dropWhile_suffix _ -#align list.rdrop_while_prefix List.dropRightWhile_prefix +#align list.rdrop_while_prefix List.rdropWhile_prefix variable {p} {l} @[simp] -theorem dropRightWhile_eq_nil_iff : dropRightWhile p l = [] ↔ ∀ x ∈ l, p x := by - simp [dropRightWhile] -#align list.rdrop_while_eq_nil_iff List.dropRightWhile_eq_nil_iff +theorem rdropWhile_eq_nil_iff : rdropWhile p l = [] ↔ ∀ x ∈ l, p x := by simp [rdropWhile] +#align list.rdrop_while_eq_nil_iff List.rdropWhile_eq_nil_iff -- it is in this file because it requires `List.Infix` @[simp] -theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p (l.get ⟨0, hl⟩) := by +theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p (l.nthLe 0 hl) := by cases' l with hd tl · simp only [dropWhile, true_iff] intro h @@ -153,28 +149,29 @@ theorem dropWhile_eq_self_iff : dropWhile p l = l ↔ ∀ hl : 0 < l.length, ¬p · rw [dropWhile] refine' ⟨fun h => _, fun h => _⟩ · intro _ H - rw [get] at H + rw [nthLe, get] at H refine' (cons_ne_self hd tl) (Sublist.antisymm _ (sublist_cons _ _)) rw [← h] simp only [H] exact List.IsSuffix.sublist (dropWhile_suffix p) · have := h (by simp only [length, Nat.succ_pos]) - rw [get] at this + rw [nthLe, get] at this simp_rw [this] #align list.drop_while_eq_self_iff List.dropWhile_eq_self_iff /- porting note: This proof is longer than it used to be because `simp` refuses to rewrite the `l ≠ []` condition if `hl` is not `intro`'d yet -/ @[simp] -theorem dropRightWhile_eq_self_iff : dropRightWhile p l = l ↔ ∀ hl : l ≠ [], ¬p (l.getLast hl) := by - simp only [dropRightWhile._eq_1, reverse_eq_iff, dropWhile_eq_self_iff, length_reverse, - ←length_pos, ne_eq, Bool.not_eq_true, getLast_eq_get, ge_iff_le] - refine forall_congr' ?_ - intro h - rw [get_reverse'] - · simp - · exact Nat.pred_lt h.ne' -#align list.rdrop_while_eq_self_iff List.dropRightWhile_eq_self_iff +theorem rdropWhile_eq_self_iff : rdropWhile p l = l ↔ ∀ hl : l ≠ [], ¬p (l.getLast hl) := by + simp only [rdropWhile, reverse_eq_iff, dropWhile_eq_self_iff, getLast_eq_get] + refine' ⟨fun h hl => _, fun h hl => _⟩ + · rw [← length_pos, ← length_reverse] at hl + have := h hl + rwa [nthLe, get_reverse'] at this + · rw [length_reverse, length_pos] at hl + have := h hl + rwa [nthLe, get_reverse'] +#align list.rdrop_while_eq_self_iff List.rdropWhile_eq_self_iff variable (p) (l) @@ -183,72 +180,70 @@ theorem dropWhile_idempotent : dropWhile p (dropWhile p l) = dropWhile p l := by exact fun h => dropWhile_nthLe_zero_not p l h #align list.drop_while_idempotent List.dropWhile_idempotent -theorem dropRightWhile_idempotent : dropRightWhile p (dropRightWhile p l) = dropRightWhile p l := - dropRightWhile_eq_self_iff.mpr (dropRightWhile_last_not _ _) -#align list.rdrop_while_idempotent List.dropRightWhile_idempotent +theorem rdropWhile_idempotent : rdropWhile p (rdropWhile p l) = rdropWhile p l := + rdropWhile_eq_self_iff.mpr (rdropWhile_last_not _ _) +#align list.rdrop_while_idempotent List.rdropWhile_idempotent /-- Take elements from the tail end of a list that satisfy `p : α → Bool`. Implemented naively via `List.reverse` -/ -def takeRightWhile : List α := +def rtakeWhile : List α := reverse (l.reverse.takeWhile p) -#align list.rtake_while List.takeRightWhile +#align list.rtake_while List.rtakeWhile @[simp] -theorem takeRightWhile_nil : takeRightWhile p ([] : List α) = [] := by - simp [takeRightWhile, takeWhile] -#align list.rtake_while_nil List.takeRightWhile_nil +theorem rtakeWhile_nil : rtakeWhile p ([] : List α) = [] := by simp [rtakeWhile, takeWhile] +#align list.rtake_while_nil List.rtakeWhile_nil -theorem takeRightWhile_concat (x : α) : - takeRightWhile p (l ++ [x]) = if p x then takeRightWhile p l ++ [x] else [] := by - simp only [takeRightWhile, takeWhile, reverse_append, reverse_singleton, singleton_append] +theorem rtakeWhile_concat (x : α) : + rtakeWhile p (l ++ [x]) = if p x then rtakeWhile p l ++ [x] else [] := by + simp only [rtakeWhile, takeWhile, reverse_append, reverse_singleton, singleton_append] split_ifs with h <;> simp [h] -#align list.rtake_while_concat List.takeRightWhile_concat +#align list.rtake_while_concat List.rtakeWhile_concat @[simp] -theorem takeRightWhile_concat_pos (x : α) (h : p x) : - takeRightWhile p (l ++ [x]) = takeRightWhile p l ++ [x] := by - rw [takeRightWhile_concat, if_pos h] -#align list.rtake_while_concat_pos List.takeRightWhile_concat_pos +theorem rtakeWhile_concat_pos (x : α) (h : p x) : + rtakeWhile p (l ++ [x]) = rtakeWhile p l ++ [x] := by rw [rtakeWhile_concat, if_pos h] +#align list.rtake_while_concat_pos List.rtakeWhile_concat_pos @[simp] -theorem takeRightWhile_concat_neg (x : α) (h : ¬p x) : takeRightWhile p (l ++ [x]) = [] := by - rw [takeRightWhile_concat, if_neg h] -#align list.rtake_while_concat_neg List.takeRightWhile_concat_neg +theorem rtakeWhile_concat_neg (x : α) (h : ¬p x) : rtakeWhile p (l ++ [x]) = [] := by + rw [rtakeWhile_concat, if_neg h] +#align list.rtake_while_concat_neg List.rtakeWhile_concat_neg -theorem takeRightWhile_suffix : l.takeRightWhile p <:+ l := by - rw [← reverse_prefix, takeRightWhile, reverse_reverse] +theorem rtakeWhile_suffix : l.rtakeWhile p <:+ l := by + rw [← reverse_prefix, rtakeWhile, reverse_reverse] exact takeWhile_prefix _ -#align list.rtake_while_suffix List.takeRightWhile_suffix +#align list.rtake_while_suffix List.rtakeWhile_suffix variable {p} {l} @[simp] -theorem takeRightWhile_eq_self_iff : takeRightWhile p l = l ↔ ∀ x ∈ l, p x := by - simp [takeRightWhile, reverse_eq_iff] -#align list.rtake_while_eq_self_iff List.takeRightWhile_eq_self_iff +theorem rtakeWhile_eq_self_iff : rtakeWhile p l = l ↔ ∀ x ∈ l, p x := by + simp [rtakeWhile, reverse_eq_iff] +#align list.rtake_while_eq_self_iff List.rtakeWhile_eq_self_iff -- Porting note: This needed a lot of rewriting. @[simp] -theorem takeRightWhile_eq_nil_iff : takeRightWhile p l = [] ↔ ∀ hl : l ≠ [], ¬p (l.getLast hl) := by +theorem rtakeWhile_eq_nil_iff : rtakeWhile p l = [] ↔ ∀ hl : l ≠ [], ¬p (l.getLast hl) := by induction' l using List.reverseRecOn with l a - · simp only [takeRightWhile, takeWhile, reverse_nil, true_iff] + · simp only [rtakeWhile, takeWhile, reverse_nil, true_iff] intro f; contradiction - · simp only [takeRightWhile, reverse_append, takeWhile, reverse_eq_nil, getLast_append, ne_eq, - append_eq_nil, and_false, forall_true_left] + · simp only [rtakeWhile, reverse_append, takeWhile, reverse_eq_nil, getLast_append, ne_eq, + append_eq_nil, and_false, not_false_eq_true, forall_true_left] refine' ⟨fun h => _ , fun h => _⟩ - · intro pa; simp only [pa] at h - · simp only [h] -#align list.rtake_while_eq_nil_iff List.takeRightWhile_eq_nil_iff + · intro pa; simp [pa] at h + · simp [h] +#align list.rtake_while_eq_nil_iff List.rtakeWhile_eq_nil_iff -theorem mem_takeRightWhile_imp {x : α} (hx : x ∈ takeRightWhile p l) : p x := by - rw [takeRightWhile, mem_reverse] at hx +theorem mem_rtakeWhile_imp {x : α} (hx : x ∈ rtakeWhile p l) : p x := by + rw [rtakeWhile, mem_reverse] at hx exact mem_takeWhile_imp hx -#align list.mem_rtake_while_imp List.mem_takeRightWhile_imp +#align list.mem_rtake_while_imp List.mem_rtakeWhile_imp variable (p) (l) -theorem takeRightWhile_idempotent : takeRightWhile p (takeRightWhile p l) = takeRightWhile p l := - takeRightWhile_eq_self_iff.mpr fun _ => mem_takeRightWhile_imp -#align list.rtake_while_idempotent List.takeRightWhile_idempotent +theorem rtakeWhile_idempotent : rtakeWhile p (rtakeWhile p l) = rtakeWhile p l := + rtakeWhile_eq_self_iff.mpr fun _ => mem_rtakeWhile_imp +#align list.rtake_while_idempotent List.rtakeWhile_idempotent end List diff --git a/Mathlib/Data/List/EditDistance/Defs.lean b/Mathlib/Data/List/EditDistance/Defs.lean index 02186dc68b252..bf4bd6a00415f 100644 --- a/Mathlib/Data/List/EditDistance/Defs.lean +++ b/Mathlib/Data/List/EditDistance/Defs.lean @@ -278,7 +278,7 @@ theorem levenshtein_nil_nil : levenshtein C [] [] = 0 := by @[simp] theorem levenshtein_nil_cons (y) (ys) : levenshtein C [] (y :: ys) = C.insert y + levenshtein C [] ys := by - dsimp [levenshtein, suffixLevenshtein, impl] + dsimp (config := { unfoldPartialApp := true }) [levenshtein, suffixLevenshtein, impl] congr rw [List.getLast_eq_get] congr diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index c6d7b3aa8c342..9a2873706d4d7 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -238,13 +238,15 @@ end FoldrIdx theorem indexesValues_eq_filter_enum (p : α → Prop) [DecidablePred p] (as : List α) : indexesValues p as = filter (p ∘ Prod.snd) (enum as) := by - simp [indexesValues, foldrIdx_eq_foldr_enum, uncurry, filter_eq_foldr] + simp (config := { unfoldPartialApp := true }) [indexesValues, foldrIdx_eq_foldr_enum, uncurry, + filter_eq_foldr] #align list.indexes_values_eq_filter_enum List.indexesValues_eq_filter_enum theorem findIdxs_eq_map_indexesValues (p : α → Prop) [DecidablePred p] (as : List α) : findIdxs p as = map Prod.fst (indexesValues p as) := by - simp only [indexesValues_eq_filter_enum, map_filter_eq_foldr, findIdxs, uncurry, - foldrIdx_eq_foldr_enum, decide_eq_true_eq, comp_apply, Bool.cond_decide] + simp (config := { unfoldPartialApp := true }) only [indexesValues_eq_filter_enum, + map_filter_eq_foldr, findIdxs, uncurry, foldrIdx_eq_foldr_enum, decide_eq_true_eq, comp_apply, + Bool.cond_decide] #align list.find_indexes_eq_map_indexes_values List.findIdxs_eq_map_indexesValues section FoldlIdx @@ -281,7 +283,8 @@ variable {m : Type u → Type v} [Monad m] theorem foldrIdxM_eq_foldrM_enum {α β} (f : ℕ → α → β → m β) (b : β) (as : List α) [LawfulMonad m] : foldrIdxM f b as = foldrM (uncurry f) b (enum as) := by - simp only [foldrIdxM, foldrM_eq_foldr, foldrIdx_eq_foldr_enum, uncurry] + simp (config := { unfoldPartialApp := true }) only [foldrIdxM, foldrM_eq_foldr, + foldrIdx_eq_foldr_enum, uncurry] #align list.mfoldr_with_index_eq_mfoldr_enum List.foldrIdxM_eq_foldrM_enum theorem foldlIdxM_eq_foldlM_enum [LawfulMonad m] {α β} (f : ℕ → β → α → m β) (b : β) (as : List α) : diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index b7614d1d253b6..d6e882e1aa7df 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -259,7 +259,7 @@ theorem prefix_take_le_iff {L : List (List (Option α))} (hm : m < L.length) : | zero => refine' iff_of_false _ (zero_lt_succ _).not_le rw [take_zero, take_nil] - simp only [take] + simp only [take, not_false_eq_true] | succ n => simp only [length] at hm have specializedIH := @IH n ls (Nat.lt_of_succ_lt_succ hm) diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index 25136a1b3a160..a8fb8984e4e69 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -45,17 +45,17 @@ theorem zero_bot (n : ℕ) : Ico 0 n = range n := by rw [Ico, tsub_zero, range_e @[simp] theorem length (n m : ℕ) : length (Ico n m) = m - n := by dsimp [Ico] - simp only [length_range'] + simp [length_range', autoParam] #align list.Ico.length List.Ico.length theorem pairwise_lt (n m : ℕ) : Pairwise (· < ·) (Ico n m) := by dsimp [Ico] - simp only [pairwise_lt_range'] + simp [pairwise_lt_range', autoParam] #align list.Ico.pairwise_lt List.Ico.pairwise_lt theorem nodup (n m : ℕ) : Nodup (Ico n m) := by dsimp [Ico] - simp only [nodup_range'] + simp [nodup_range', autoParam] #align list.Ico.nodup List.Ico.nodup @[simp] diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 5902bcaca37ad..b442e56bd2580 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -134,12 +134,12 @@ theorem inter_nil (l : List α) : [] ∩ l = [] := @[simp] theorem inter_cons_of_mem (l₁ : List α) (h : a ∈ l₂) : (a :: l₁) ∩ l₂ = a :: l₁ ∩ l₂ := by - simp only [Inter.inter, List.inter, filter_cons_of_pos, h] + simp [Inter.inter, List.inter, h] #align list.inter_cons_of_mem List.inter_cons_of_mem @[simp] theorem inter_cons_of_not_mem (l₁ : List α) (h : a ∉ l₂) : (a :: l₁) ∩ l₂ = l₁ ∩ l₂ := by - simp only [Inter.inter, List.inter, filter_cons_of_neg, h] + simp [Inter.inter, List.inter, h] #align list.inter_cons_of_not_mem List.inter_cons_of_not_mem theorem mem_of_mem_inter_left : a ∈ l₁ ∩ l₂ → a ∈ l₁ := diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 795505b6527c3..551f7847b6f2e 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -278,9 +278,11 @@ theorem filter_append_perm (p : α → Bool) (l : List α) : induction' l with x l ih · rfl · by_cases h : p x - · simp only [h, filter_cons_of_pos, filter_cons_of_neg, not_true, not_false_iff, cons_append] + · simp only [h, filter_cons_of_pos, filter_cons_of_neg, not_true, not_false_iff, cons_append, + decide_False] exact ih.cons x - · simp only [h, filter_cons_of_neg, not_false_iff, filter_cons_of_pos] + · simp only [h, filter_cons_of_neg, not_false_iff, filter_cons_of_pos, cons_append, + not_false_eq_true, decide_True] refine' Perm.trans _ (ih.cons x) exact perm_append_comm.trans (perm_append_comm.cons _) #align list.filter_append_perm List.filter_append_perm diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 3825116af5511..16942639d2398 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -76,14 +76,15 @@ theorem nthLe_range'_1 {n m} (i) (H : i < (range' n m).length) : #align list.range_eq_nil List.range_eq_nil theorem pairwise_lt_range (n : ℕ) : Pairwise (· < ·) (range n) := by - simp only [range_eq_range', pairwise_lt_range'] + simp (config := {decide := true}) only [range_eq_range', pairwise_lt_range'] #align list.pairwise_lt_range List.pairwise_lt_range theorem pairwise_le_range (n : ℕ) : Pairwise (· ≤ ·) (range n) := Pairwise.imp (@le_of_lt ℕ _) (pairwise_lt_range _) #align list.pairwise_le_range List.pairwise_le_range -theorem nodup_range (n : ℕ) : Nodup (range n) := by simp only [range_eq_range', nodup_range'] +theorem nodup_range (n : ℕ) : Nodup (range n) := by + simp (config := {decide := true}) only [range_eq_range', nodup_range'] #align list.nodup_range List.nodup_range #align list.range_sublist List.range_sublist #align list.range_subset List.range_subset diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index ec87ec5e18bc9..808f618bdfa94 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -453,7 +453,7 @@ theorem mem_keys_kerase_of_ne {a₁ a₂} {l : List (Sigma β)} (h : a₁ ≠ a #align list.mem_keys_kerase_of_ne List.mem_keys_kerase_of_ne theorem keys_kerase {a} {l : List (Sigma β)} : (kerase a l).keys = l.keys.erase a := by - rw [keys, kerase, erase_eq_eraseP, eraseP_map]; dsimp [Function.comp] + rw [keys, kerase, erase_eq_eraseP, eraseP_map, Function.comp] #align list.keys_kerase List.keys_kerase theorem kerase_kerase {a a'} {l : List (Sigma β)} : diff --git a/Mathlib/Data/Matrix/Basis.lean b/Mathlib/Data/Matrix/Basis.lean index b779b313a5c12..10cdda78715e1 100644 --- a/Mathlib/Data/Matrix/Basis.lean +++ b/Mathlib/Data/Matrix/Basis.lean @@ -65,12 +65,12 @@ theorem matrix_eq_sum_std_basis [Fintype m] [Fintype n] (x : Matrix m n α) : -- Porting note: was `convert` refine (Fintype.sum_eq_single i ?_).trans ?_; swap · -- Porting note: `simp` seems unwilling to apply `Fintype.sum_apply` - simp only [stdBasisMatrix] + simp (config := { unfoldPartialApp := true }) only [stdBasisMatrix] rw [Fintype.sum_apply, Fintype.sum_apply] simp · intro j' hj' -- Porting note: `simp` seems unwilling to apply `Fintype.sum_apply` - simp only [stdBasisMatrix] + simp (config := { unfoldPartialApp := true }) only [stdBasisMatrix] rw [Fintype.sum_apply, Fintype.sum_apply] simp [hj'] #align matrix.matrix_eq_sum_std_basis Matrix.matrix_eq_sum_std_basis diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 90aa0f7396a06..3c5805cc6345f 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -2683,9 +2683,7 @@ for more discussion. theorem map_count_True_eq_filter_card (s : Multiset α) (p : α → Prop) [DecidablePred p] : (s.map p).count True = card (s.filter p) := by simp only [count_eq_card_filter_eq, map_filter, card_map, Function.comp.left_id, - eq_true_eq_id, Function.comp] - congr; funext _ - simp only [eq_iff_iff, true_iff] + eq_true_eq_id, Function.comp_apply] #align multiset.map_count_true_eq_filter_card Multiset.map_count_True_eq_filter_card /-! ### Lift a relation to `Multiset`s -/ diff --git a/Mathlib/Data/MvPolynomial/Division.lean b/Mathlib/Data/MvPolynomial/Division.lean index b7ddacdc2f8bd..634ae6990b14a 100644 --- a/Mathlib/Data/MvPolynomial/Division.lean +++ b/Mathlib/Data/MvPolynomial/Division.lean @@ -232,8 +232,8 @@ theorem monomial_dvd_monomial {r s : R} {i j : σ →₀ ℕ} : · exact ⟨Or.inr hi, _, hj⟩ · exact ⟨Or.inl hj, hj.symm ▸ dvd_zero _⟩ -- Porting note: two goals remain at this point in Lean 4 - · simp_all only [or_true, dvd_mul_right] - · simp_all only [ite_self, le_refl, ite_true, dvd_mul_right] + · simp_all only [or_true, dvd_mul_right, and_self] + · simp_all only [ite_self, le_refl, ite_true, dvd_mul_right, or_false, and_self] · rintro ⟨h | hij, d, rfl⟩ · simp_rw [h, monomial_zero, dvd_zero] · refine' ⟨monomial (j - i) d, _⟩ @@ -252,7 +252,7 @@ theorem X_dvd_X [Nontrivial R] {i j : σ} : (X i : MvPolynomial σ R) ∣ (X j : MvPolynomial σ R) ↔ i = j := by refine' monomial_one_dvd_monomial_one.trans _ simp_rw [Finsupp.single_le_iff, Nat.one_le_iff_ne_zero, Finsupp.single_apply_ne_zero, - and_true_iff] + ne_eq, not_false_eq_true, and_true] set_option linter.uppercaseLean3 false in #align mv_polynomial.X_dvd_X MvPolynomial.X_dvd_X diff --git a/Mathlib/Data/MvPolynomial/Variables.lean b/Mathlib/Data/MvPolynomial/Variables.lean index b63a5e3b788b1..d6c217d4fe3c1 100644 --- a/Mathlib/Data/MvPolynomial/Variables.lean +++ b/Mathlib/Data/MvPolynomial/Variables.lean @@ -686,7 +686,7 @@ theorem totalDegree_smul_le [CommSemiring S] [DistribMulAction R S] (a : R) (f : theorem totalDegree_pow (a : MvPolynomial σ R) (n : ℕ) : (a ^ n).totalDegree ≤ n * a.totalDegree := by induction' n with n ih - · simp only [Nat.zero_eq, zero_mul, pow_zero, totalDegree_one] + · simp only [Nat.zero_eq, pow_zero, totalDegree_one, zero_mul, le_refl] rw [pow_succ] calc totalDegree (a * a ^ n) ≤ a.totalDegree + (a ^ n).totalDegree := totalDegree_mul _ _ diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index e42b6dab81436..54789c59e8254 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -82,17 +82,18 @@ theorem binaryRec_of_ne_zero {C : Nat → Sort*} (z : C 0) (f : ∀ b n, C n → lemma bitwise_bit {f : Bool → Bool → Bool} (h : f false false = false := by rfl) (a m b n) : bitwise f (bit a m) (bit b n) = bit (f a b) (bitwise f m n) := by conv_lhs => unfold bitwise - simp only [bit, bit1, bit0, Bool.cond_eq_ite] + simp (config := { unfoldPartialApp := true }) only [bit, bit1, bit0, Bool.cond_eq_ite] have h1 x : (x + x) % 2 = 0 := by rw [← two_mul, mul_comm]; apply mul_mod_left have h2 x : (x + x + 1) % 2 = 1 := by rw [← two_mul, add_comm]; apply add_mul_mod_self_left have h3 x : (x + x) / 2 = x := by rw [← two_mul, mul_comm]; apply mul_div_left _ zero_lt_two - have h4 x : (x + x + 1) / 2 = x := by rw [← two_mul, add_comm]; simp [add_mul_div_left] - cases a <;> cases b <;> simp [h1, h2, h3, h4] <;> split_ifs <;> simp_all + have h4 x : (x + x + 1) / 2 = x := by rw [← two_mul, add_comm]; simp [add_mul_div_left]; rfl + cases a <;> cases b <;> simp [h1, h2, h3, h4] <;> split_ifs + <;> simp_all (config := {decide := true}) #align nat.bitwise_bit Nat.bitwise_bit lemma bit_mod_two (a : Bool) (x : ℕ) : bit a x % 2 = if a then 1 else 0 := by - simp [bit, bit0, bit1, Bool.cond_eq_ite, ←mul_two] + simp (config := { unfoldPartialApp := true }) [bit, bit0, bit1, Bool.cond_eq_ite, ←mul_two] split_ifs <;> simp [Nat.add_mod] @[simp] @@ -304,6 +305,7 @@ theorem testBit_two_pow_of_ne {n m : ℕ} (hm : n ≠ m) : testBit (2 ^ n) m = f rw [(rfl : succ 0 = 1)] simp only [ge_iff_le, tsub_le_iff_right, pow_succ, bodd_mul, Bool.and_eq_false_eq_eq_false_or_eq_false, or_true] + exact Or.inr rfl #align nat.test_bit_two_pow_of_ne Nat.testBit_two_pow_of_ne theorem testBit_two_pow (n m : ℕ) : testBit (2 ^ n) m = (n = m) := by @@ -349,8 +351,7 @@ theorem xor_comm (n m : ℕ) : n ^^^ m = m ^^^ n := #align nat.lxor_comm Nat.xor_comm @[simp] -theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by - simp only [HXor.hXor, Xor.xor, xor, bitwise_zero_left, ite_true] +theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by simp [HXor.hXor, Xor.xor, xor] #align nat.zero_lxor Nat.zero_xor @[simp] @@ -359,22 +360,22 @@ theorem xor_zero (n : ℕ) : n ^^^ 0 = n := by simp [HXor.hXor, Xor.xor, xor] @[simp] theorem zero_land (n : ℕ) : 0 &&& n = 0 := by - simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_left, ite_false]; + simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_left, ite_false, Bool.and_true]; #align nat.zero_land Nat.zero_land @[simp] theorem land_zero (n : ℕ) : n &&& 0 = 0 := by - simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_right, ite_false] + simp only [HAnd.hAnd, AndOp.and, land, bitwise_zero_right, ite_false, Bool.and_false] #align nat.land_zero Nat.land_zero @[simp] theorem zero_lor (n : ℕ) : 0 ||| n = n := by - simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_left, ite_true] + simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_left, ite_true, Bool.or_true] #align nat.zero_lor Nat.zero_lor @[simp] theorem lor_zero (n : ℕ) : n ||| 0 = n := by - simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_right, ite_true] + simp only [HOr.hOr, OrOp.or, lor, bitwise_zero_right, ite_true, Bool.or_false] #align nat.lor_zero Nat.lor_zero diff --git a/Mathlib/Data/Nat/Choose/Multinomial.lean b/Mathlib/Data/Nat/Choose/Multinomial.lean index 2aba264a718d6..9090528dd9be5 100644 --- a/Mathlib/Data/Nat/Choose/Multinomial.lean +++ b/Mathlib/Data/Nat/Choose/Multinomial.lean @@ -126,7 +126,7 @@ theorem binomial_succ_succ [DecidableEq α] (h : a ≠ b) : multinomial {a, b} (Function.update f a (f a).succ) + multinomial {a, b} (Function.update f b (f b).succ) := by simp only [binomial_eq_choose, Function.update_apply, - h, Ne.def, ite_true, ite_false] + h, Ne.def, ite_true, ite_false, not_false_eq_true] rw [if_neg h.symm] rw [add_succ, choose_succ_succ, succ_add_eq_succ_add] ring diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 391552e443058..8e293c6ffbfdf 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -357,6 +357,7 @@ theorem getLast_digit_ne_zero (b : ℕ) {m : ℕ} (hm : m ≠ 0) : · cases hm rfl rename ℕ => m simp only [digits_one, List.getLast_replicate_succ m 1] + exact Nat.one_ne_zero revert hm apply Nat.strongInductionOn m intro n IH hn @@ -569,7 +570,7 @@ theorem sub_one_mul_sum_div_pow_eq_sub_sum_digits · simp [ofDigits_one] · simp [lt_one_iff.mp h] cases L - · simp + · rfl · simp [ofDigits] theorem sub_one_mul_sum_log_div_pow_eq_sub_sum_digits (n : ℕ) : @@ -592,7 +593,7 @@ theorem sub_one_mul_sum_log_div_pow_eq_sub_sum_digits (n : ℕ) : theorem digits_two_eq_bits (n : ℕ) : digits 2 n = n.bits.map fun b => cond b 1 0 := by induction' n using Nat.binaryRecFromOne with b n h ih · simp - · simp + · rfl rw [bits_append_bit _ _ fun hn => absurd hn h] cases b · rw [digits_def' one_lt_two] @@ -689,7 +690,7 @@ theorem ofDigits_neg_one : theorem modEq_eleven_digits_sum (n : ℕ) : n ≡ ((digits 10 n).map fun n : ℕ => (n : ℤ)).alternatingSum [ZMOD 11] := by - have t := zmodeq_ofDigits_digits 11 10 (-1 : ℤ) (by unfold Int.ModEq; norm_num) n + have t := zmodeq_ofDigits_digits 11 10 (-1 : ℤ) (by unfold Int.ModEq; rfl) n rwa [ofDigits_neg_one] at t #align nat.modeq_eleven_digits_sum Nat.modEq_eleven_digits_sum diff --git a/Mathlib/Data/Nat/Factorial/BigOperators.lean b/Mathlib/Data/Nat/Factorial/BigOperators.lean index 64f65882c7a4f..642ba3f562abb 100644 --- a/Mathlib/Data/Nat/Factorial/BigOperators.lean +++ b/Mathlib/Data/Nat/Factorial/BigOperators.lean @@ -32,7 +32,7 @@ theorem prod_factorial_pos : 0 < ∏ i in s, (f i)! := theorem prod_factorial_dvd_factorial_sum : (∏ i in s, (f i)!) ∣ (∑ i in s, f i)! := by classical induction' s using Finset.induction with a' s' has ih - · simp only [Finset.sum_empty, Finset.prod_empty, factorial] + · simp only [prod_empty, factorial, dvd_refl] · simp only [Finset.prod_insert has, Finset.sum_insert has] refine' dvd_trans (mul_dvd_mul_left (f a')! ih) _ apply Nat.factorial_mul_factorial_dvd_factorial_add diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 3f785a85f4e20..dc024be1680c1 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -112,11 +112,11 @@ theorem factorization_inj : Set.InjOn factorization { x : ℕ | x ≠ 0 } := fun #align nat.factorization_inj Nat.factorization_inj @[simp] -theorem factorization_zero : factorization 0 = 0 := by simp [factorization] +theorem factorization_zero : factorization 0 = 0 := by decide #align nat.factorization_zero Nat.factorization_zero @[simp] -theorem factorization_one : factorization 1 = 0 := by simp [factorization] +theorem factorization_one : factorization 1 = 0 := by decide #align nat.factorization_one Nat.factorization_one #noalign nat.support_factorization @@ -430,6 +430,7 @@ theorem factorization_prime_le_iff_dvd {d n : ℕ} (hd : d ≠ 0) (hn : n ≠ 0) rw [← factorization_le_iff_dvd hd hn] refine' ⟨fun h p => (em p.Prime).elim (h p) fun hp => _, fun h p _ => h p⟩ simp_rw [factorization_eq_zero_of_non_prime _ hp] + rfl #align nat.factorization_prime_le_iff_dvd Nat.factorization_prime_le_iff_dvd theorem pow_succ_factorization_not_dvd {n p : ℕ} (hn : n ≠ 0) (hp : p.Prime) : diff --git a/Mathlib/Data/Nat/Hyperoperation.lean b/Mathlib/Data/Nat/Hyperoperation.lean index ad7d9bfcbf240..ad4daea7483e7 100644 --- a/Mathlib/Data/Nat/Hyperoperation.lean +++ b/Mathlib/Data/Nat/Hyperoperation.lean @@ -119,7 +119,7 @@ theorem hyperoperation_ge_four_zero (n k : ℕ) : hyperoperation (n + 4) 0 k = if Even k then 1 else 0 := by induction' k with kk kih · rw [hyperoperation_ge_three_eq_one] - simp only [even_zero, if_true] + simp only [Nat.zero_eq, even_zero, if_true] · rw [hyperoperation_recursion] rw [kih] simp_rw [Nat.even_add_one] diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index 66bd58ec8ce52..003252fd98157 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -142,7 +142,7 @@ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : · simpa only [log_of_left_le_one hb, hm.symm, false_iff_iff, not_and, not_lt] using le_trans (pow_le_pow_of_le_one' hb m.le_succ) · simpa only [log_zero_right, hm.symm, nonpos_iff_eq_zero, false_iff, not_and, not_lt, - add_pos_iff, or_true, pow_eq_zero_iff] using pow_eq_zero + add_pos_iff, zero_lt_one, or_true, pow_eq_zero_iff] using pow_eq_zero #align nat.log_eq_iff Nat.log_eq_iff theorem log_eq_of_pow_le_of_lt_pow {b m n : ℕ} (h₁ : b ^ m ≤ n) (h₂ : n < b ^ (m + 1)) : diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index da8a9bbea7bd7..4f0d0831b943d 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -526,7 +526,7 @@ theorem odd_of_mod_four_eq_three {n : ℕ} : n % 4 = 3 → n % 2 = 1 := by theorem odd_mod_four_iff {n : ℕ} : n % 2 = 1 ↔ n % 4 = 1 ∨ n % 4 = 3 := have help : ∀ m : ℕ, m < 4 → m % 2 = 1 → m = 1 ∨ m = 3 := by decide ⟨fun hn => - help (n % 4) (mod_lt n (by norm_num)) <| (mod_mod_of_dvd n (by norm_num : 2 ∣ 4)).trans hn, + help (n % 4) (mod_lt n (by norm_num)) <| (mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn, fun h => Or.elim h odd_of_mod_four_eq_one odd_of_mod_four_eq_three⟩ #align nat.odd_mod_four_iff Nat.odd_mod_four_iff diff --git a/Mathlib/Data/Nat/Parity.lean b/Mathlib/Data/Nat/Parity.lean index 7806c7fd8ad5a..e8329207d53b3 100644 --- a/Mathlib/Data/Nat/Parity.lean +++ b/Mathlib/Data/Nat/Parity.lean @@ -284,9 +284,8 @@ end example (m n : ℕ) (h : Even m) : ¬Even (n + 3) ↔ Even (m ^ 2 + m + n) := by simp [*, two_ne_zero, parity_simps] -/- Porting note: the `simp` lemmas about `bit*` no longer apply, but `simp` in Lean 4 currently -simplifies decidable propositions. This may change in the future. -/ -example : ¬Even 25394535 := by simp only +/- Porting note: the `simp` lemmas about `bit*` no longer apply. -/ +example : ¬Even 25394535 := by decide end Nat diff --git a/Mathlib/Data/Nat/PartENat.lean b/Mathlib/Data/Nat/PartENat.lean index cd6c16918db89..3d959a474bd46 100644 --- a/Mathlib/Data/Nat/PartENat.lean +++ b/Mathlib/Data/Nat/PartENat.lean @@ -561,7 +561,7 @@ theorem add_eq_top_iff {a b : PartENat} : a + b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := refine PartENat.casesOn a ?_ ?_ <;> refine PartENat.casesOn b ?_ ?_ <;> simp [top_add, add_top] - simp only [←Nat.cast_add, PartENat.natCast_ne_top, forall_const] + simp only [←Nat.cast_add, PartENat.natCast_ne_top, forall_const, not_false_eq_true] #align part_enat.add_eq_top_iff PartENat.add_eq_top_iff protected theorem add_right_cancel_iff {a b c : PartENat} (hc : c ≠ ⊤) : a + c = b + c ↔ a = b := by @@ -697,7 +697,7 @@ lemma ofENat_some (n : ℕ) : ofENat (Option.some n) = ↑n := rfl @[simp, norm_cast] theorem toWithTop_ofENat (n : ℕ∞) {_ : Decidable (n : PartENat).Dom} : toWithTop (↑n) = n := by induction n with - | none => simp + | none => simp; rfl | some n => simp only [toWithTop_natCast', ofENat_some] rfl diff --git a/Mathlib/Data/Nat/Sqrt.lean b/Mathlib/Data/Nat/Sqrt.lean index b59613067c3da..cbd3e6973d366 100644 --- a/Mathlib/Data/Nat/Sqrt.lean +++ b/Mathlib/Data/Nat/Sqrt.lean @@ -47,8 +47,8 @@ To turn this into a lean proof we need to manipulate, use properties of natural -/ private theorem sqrt_isSqrt (n : ℕ) : IsSqrt n (sqrt n) := by match n with - | 0 => simp [IsSqrt] - | 1 => simp [IsSqrt] + | 0 => simp [IsSqrt, sqrt] + | 1 => simp [IsSqrt, sqrt] | n + 2 => have h : ¬ (n + 2) ≤ 1 := by simp simp only [IsSqrt, sqrt, h, ite_false] diff --git a/Mathlib/Data/Nat/Squarefree.lean b/Mathlib/Data/Nat/Squarefree.lean index d2a46b5e8f685..29628a804b196 100644 --- a/Mathlib/Data/Nat/Squarefree.lean +++ b/Mathlib/Data/Nat/Squarefree.lean @@ -254,7 +254,7 @@ instance : DecidablePred (Squarefree : ℕ → Prop) := fun _ => decidable_of_iff' _ squarefree_iff_minSqFac theorem squarefree_two : Squarefree 2 := by - rw [squarefree_iff_nodup_factors] <;> norm_num + rw [squarefree_iff_nodup_factors] <;> decide #align nat.squarefree_two Nat.squarefree_two theorem divisors_filter_squarefree_of_squarefree {n : ℕ} (hn : Squarefree n) : diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index f523f068f8d40..7de807e457a8a 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -34,7 +34,8 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop + any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + aesop (simp_options := { decide := true }) repeat' erw [WithBot.coe_eq_coe] exact Nat.add_eq_one_iff #align nat.with_bot.add_eq_one_iff Nat.WithBot.add_eq_one_iff @@ -42,7 +43,8 @@ theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n theorem add_eq_two_iff {n m : WithBot ℕ} : n + m = 2 ↔ n = 0 ∧ m = 2 ∨ n = 1 ∧ m = 1 ∨ n = 2 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop + any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + aesop (simp_options := { decide := true }) repeat' erw [WithBot.coe_eq_coe] exact Nat.add_eq_two_iff #align nat.with_bot.add_eq_two_iff Nat.WithBot.add_eq_two_iff @@ -50,7 +52,8 @@ theorem add_eq_two_iff {n m : WithBot ℕ} : theorem add_eq_three_iff {n m : WithBot ℕ} : n + m = 3 ↔ n = 0 ∧ m = 3 ∨ n = 1 ∧ m = 2 ∨ n = 2 ∧ m = 1 ∨ n = 3 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop + any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + aesop (simp_options := { decide := true }) repeat' erw [WithBot.coe_eq_coe] exact Nat.add_eq_three_iff #align nat.with_bot.add_eq_three_iff Nat.WithBot.add_eq_three_iff diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index a765c68c599f0..dc9ced5fbb7cf 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -250,7 +250,7 @@ theorem bit1_succ : ∀ n : Num, n.bit1.succ = n.succ.bit0 #align num.bit1_succ Num.bit1_succ theorem ofNat'_succ : ∀ {n}, ofNat' (n + 1) = ofNat' n + 1 := - @(Nat.binaryRec (by simp) fun b n ih => by + @(Nat.binaryRec (by simp; rfl) fun b n ih => by cases b · erw [ofNat'_bit true n, ofNat'_bit] simp only [← bit1_of_bit1, ← bit0_of_bit0, cond, _root_.bit1] diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index dc631e06da0ee..c9609459f7871 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -328,7 +328,7 @@ theorem getD_default_eq_iget [Inhabited α] (o : Option α) : @[simp] theorem guard_eq_some' {p : Prop} [Decidable p] (u) : _root_.guard p = some u ↔ p := by cases u - by_cases h : p <;> simp [_root_.guard, h] + by_cases h : p <;> simp [_root_.guard, h]; rfl #align option.guard_eq_some' Option.guard_eq_some' theorem liftOrGet_choice {f : α → α → α} (h : ∀ a b, f a b = a ∨ f a b = b) : diff --git a/Mathlib/Data/PFunctor/Multivariate/W.lean b/Mathlib/Data/PFunctor/Multivariate/W.lean index bd4d14cd77342..d225c02b69d0b 100644 --- a/Mathlib/Data/PFunctor/Multivariate/W.lean +++ b/Mathlib/Data/PFunctor/Multivariate/W.lean @@ -251,7 +251,7 @@ theorem w_map_wMk {α β : TypeVec n} (g : α ⟹ β) (a : P.A) (f' : P.drop.B a show _ = P.wMk a (g ⊚ f') (MvFunctor.map g ∘ f) have : MvFunctor.map g ∘ f = fun i => ⟨(f i).fst, g ⊚ (f i).snd⟩ := by ext i : 1 - dsimp [Function.comp] + dsimp [Function.comp_def] cases f i rfl rw [this] diff --git a/Mathlib/Data/PNat/Basic.lean b/Mathlib/Data/PNat/Basic.lean index 5b6b27c6cca81..2116a40ab5c4d 100644 --- a/Mathlib/Data/PNat/Basic.lean +++ b/Mathlib/Data/PNat/Basic.lean @@ -20,8 +20,6 @@ It is defined in `Data.PNat.Defs`, but most of the development is deferred to he that `Data.PNat.Defs` can have very few imports. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - deriving instance AddLeftCancelSemigroup, AddRightCancelSemigroup, AddCommSemigroup, LinearOrderedCancelCommMonoid, Add, Mul, Distrib for PNat @@ -272,7 +270,7 @@ theorem coe_bit1 (a : ℕ+) : ((bit1 a : ℕ+) : ℕ) = bit1 (a : ℕ) := end deprecated @[simp, norm_cast] -theorem pow_coe (m : ℕ+) (n : ℕ) : (m ^ n : ℕ) = (m : ℕ) ^ n := +theorem pow_coe (m : ℕ+) (n : ℕ) : ↑(m ^ n) = (m : ℕ) ^ n := rfl #align pnat.pow_coe PNat.pow_coe diff --git a/Mathlib/Data/PNat/Prime.lean b/Mathlib/Data/PNat/Prime.lean index 17ef5e3afe5dd..ed2740dd7d6c8 100644 --- a/Mathlib/Data/PNat/Prime.lean +++ b/Mathlib/Data/PNat/Prime.lean @@ -295,7 +295,7 @@ theorem gcd_eq_left {m n : ℕ+} : m ∣ n → m.gcd n = m := by apply Nat.gcd_eq_left h #align pnat.gcd_eq_left PNat.gcd_eq_left -theorem Coprime.pow {m n : ℕ+} (k l : ℕ) (h : m.Coprime n) : (m ^ k).Coprime (n ^ l) := by +theorem Coprime.pow {m n : ℕ+} (k l : ℕ) (h : m.Coprime n) : (m ^ k : ℕ).Coprime (n ^ l) := by rw [← coprime_coe] at *; apply Nat.Coprime.pow; apply h #align pnat.coprime.pow PNat.Coprime.pow diff --git a/Mathlib/Data/Polynomial/Coeff.lean b/Mathlib/Data/Polynomial/Coeff.lean index 87dd07274ed48..568430883ea4b 100644 --- a/Mathlib/Data/Polynomial/Coeff.lean +++ b/Mathlib/Data/Polynomial/Coeff.lean @@ -213,7 +213,7 @@ theorem support_binomial {k m : ℕ} (hkm : k ≠ m) {x y : R} (hx : x ≠ 0) (h apply subset_antisymm (support_binomial' k m x y) simp_rw [insert_subset_iff, singleton_subset_iff, mem_support_iff, coeff_add, coeff_C_mul, coeff_X_pow_self, mul_one, coeff_X_pow, if_neg hkm, if_neg hkm.symm, mul_zero, zero_add, - add_zero, Ne.def, hx, hy] + add_zero, Ne.def, hx, hy, not_false_eq_true, and_true] #align polynomial.support_binomial Polynomial.support_binomial theorem support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z : R} (hx : x ≠ 0) @@ -223,7 +223,7 @@ theorem support_trinomial {k m n : ℕ} (hkm : k < m) (hmn : m < n) {x y z : R} simp_rw [insert_subset_iff, singleton_subset_iff, mem_support_iff, coeff_add, coeff_C_mul, coeff_X_pow_self, mul_one, coeff_X_pow, if_neg hkm.ne, if_neg hkm.ne', if_neg hmn.ne, if_neg hmn.ne', if_neg (hkm.trans hmn).ne, if_neg (hkm.trans hmn).ne', mul_zero, add_zero, - zero_add, Ne.def, hx, hy, hz] + zero_add, Ne.def, hx, hy, hz, not_false_eq_true, and_true] #align polynomial.support_trinomial Polynomial.support_trinomial theorem card_support_binomial {k m : ℕ} (h : k ≠ m) {x y : R} (hx : x ≠ 0) (hy : y ≠ 0) : diff --git a/Mathlib/Data/Polynomial/Div.lean b/Mathlib/Data/Polynomial/Div.lean index fbc3e930be0c5..0a474f27048eb 100644 --- a/Mathlib/Data/Polynomial/Div.lean +++ b/Mathlib/Data/Polynomial/Div.lean @@ -74,7 +74,7 @@ theorem multiplicity_finite_of_degree_pos_of_monic (hp : (0 : WithBot ℕ) < deg have hpn0' : leadingCoeff p ^ (natDegree q + 1) ≠ 0 := hpn1.symm ▸ zn0.symm have hpnr0 : leadingCoeff (p ^ (natDegree q + 1)) * leadingCoeff r ≠ 0 := by simp only [leadingCoeff_pow' hpn0', leadingCoeff_eq_zero, hpn1, one_pow, one_mul, Ne.def, - hr0] + hr0, not_false_eq_true] have hnp : 0 < natDegree p := Nat.cast_lt.1 <| by rw [← degree_eq_natDegree hp0]; exact hp have := congr_arg natDegree hr diff --git a/Mathlib/Data/Polynomial/FieldDivision.lean b/Mathlib/Data/Polynomial/FieldDivision.lean index f0a7fc6b6eda8..5aff519ea1db8 100644 --- a/Mathlib/Data/Polynomial/FieldDivision.lean +++ b/Mathlib/Data/Polynomial/FieldDivision.lean @@ -176,7 +176,7 @@ theorem isUnit_iff_degree_eq_zero : IsUnit p ↔ degree p = 0 := ⟨degree_eq_zero_of_isUnit, fun h => have : degree p ≤ 0 := by simp [*, le_refl] have hc : coeff p 0 ≠ 0 := fun hc => by - rw [eq_C_of_degree_le_zero this, hc] at h; simp at h + rw [eq_C_of_degree_le_zero this, hc] at h; simp only [map_zero] at h; contradiction isUnit_iff_dvd_one.2 ⟨C (coeff p 0)⁻¹, by conv in p => rw [eq_C_of_degree_le_zero this] @@ -403,7 +403,7 @@ theorem exists_root_of_degree_eq_one (h : degree p = 1) : ∃ x, IsRoot p x := have : p.coeff 1 ≠ 0 := by have h' := natDegree_eq_of_degree_eq_some h change natDegree p = 1 at h'; rw [←h'] - exact mt leadingCoeff_eq_zero.1 fun h0 => by simp [h0] at h + exact mt leadingCoeff_eq_zero.1 fun h0 => by simp [h0] at h; contradiction conv in p => rw [eq_X_add_C_of_degree_le_one (show degree p ≤ 1 by rw [h])] simp [IsRoot, mul_div_cancel' _ this]⟩ #align polynomial.exists_root_of_degree_eq_one Polynomial.exists_root_of_degree_eq_one @@ -486,7 +486,7 @@ theorem prime_of_degree_eq_one (hp1 : degree p = 1) : Prime p := by classical have : Prime (normalize p) := Monic.prime_of_degree_eq_one (hp1 ▸ degree_normalize) - (monic_normalize fun hp0 => absurd hp1 (hp0.symm ▸ by simp)) + (monic_normalize fun hp0 => absurd hp1 (hp0.symm ▸ by simp only [degree_zero]; decide)) exact (normalize_associated _).prime this #align polynomial.prime_of_degree_eq_one Polynomial.prime_of_degree_eq_one diff --git a/Mathlib/Data/Polynomial/RingDivision.lean b/Mathlib/Data/Polynomial/RingDivision.lean index 6d21fe59d68db..1af8098cc72be 100644 --- a/Mathlib/Data/Polynomial/RingDivision.lean +++ b/Mathlib/Data/Polynomial/RingDivision.lean @@ -1300,7 +1300,7 @@ theorem count_map_roots_of_injective [IsDomain A] [DecidableEq B] (p : A[X]) {f (p.roots.map f).count b ≤ rootMultiplicity b (p.map f) := by by_cases hp0 : p = 0 · simp only [hp0, roots_zero, Multiset.map_zero, Multiset.count_zero, Polynomial.map_zero, - rootMultiplicity_zero] + rootMultiplicity_zero, le_refl] · exact count_map_roots ((Polynomial.map_ne_zero_iff hf).mpr hp0) b #align polynomial.count_map_roots_of_injective Polynomial.count_map_roots_of_injective @@ -1328,7 +1328,7 @@ theorem card_roots_le_map [IsDomain A] [IsDomain B] {p : A[X]} {f : A →+* B} ( theorem card_roots_le_map_of_injective [IsDomain A] [IsDomain B] {p : A[X]} {f : A →+* B} (hf : Function.Injective f) : Multiset.card p.roots ≤ Multiset.card (p.map f).roots := by by_cases hp0 : p = 0 - · simp only [hp0, roots_zero, Polynomial.map_zero, Multiset.card_zero] + · simp only [hp0, roots_zero, Polynomial.map_zero, Multiset.card_zero]; rfl exact card_roots_le_map ((Polynomial.map_ne_zero_iff hf).mpr hp0) #align polynomial.card_roots_le_map_of_injective Polynomial.card_roots_le_map_of_injective diff --git a/Mathlib/Data/Polynomial/UnitTrinomial.lean b/Mathlib/Data/Polynomial/UnitTrinomial.lean index f5a3dba04d03c..e91b47e0ff1dd 100644 --- a/Mathlib/Data/Polynomial/UnitTrinomial.lean +++ b/Mathlib/Data/Polynomial/UnitTrinomial.lean @@ -203,6 +203,7 @@ theorem isUnitTrinomial_iff' : sum_insert (mt mem_singleton.mp hmn.ne), sum_singleton, trinomial_leading_coeff' hkm hmn, trinomial_middle_coeff hkm hmn, trinomial_trailing_coeff' hkm hmn] simp_rw [← Units.val_pow_eq_pow_val, Int.units_sq] + decide · have key : ∀ k ∈ p.support, p.coeff k ^ 2 = 1 := fun k hk => Int.sq_eq_one_of_sq_le_three ((single_le_sum (fun k _ => sq_nonneg (p.coeff k)) hk).trans hp.le) (mem_support_iff.mp hk) diff --git a/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean b/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean index b2ec710527574..3862cc4a8fb65 100644 --- a/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean +++ b/Mathlib/Data/QPF/Multivariate/Constructions/Cofix.lean @@ -534,10 +534,10 @@ theorem Cofix.dest_corec' {α : TypeVec.{u} n} {β : Type u} congr!; ext (i | i) <;> erw [corec_roll] <;> dsimp [Cofix.corec'] · mv_bisim i with R a b x Ha Hb rw [Ha, Hb, Cofix.dest_corec] - dsimp [Function.comp] + dsimp [Function.comp_def] repeat rw [MvFunctor.map_map, ← appendFun_comp_id] apply liftR_map_last' - dsimp [Function.comp] + dsimp [Function.comp_def] intros exact ⟨_, rfl, rfl⟩ · congr with y diff --git a/Mathlib/Data/QPF/Multivariate/Constructions/Comp.lean b/Mathlib/Data/QPF/Multivariate/Constructions/Comp.lean index b92d9449894db..2d04fb805eea5 100644 --- a/Mathlib/Data/QPF/Multivariate/Constructions/Comp.lean +++ b/Mathlib/Data/QPF/Multivariate/Constructions/Comp.lean @@ -79,10 +79,16 @@ instance : MvQPF (Comp F G) where abs := Comp.mk ∘ (map fun i ↦ abs) ∘ abs ∘ MvPFunctor.comp.get repr {α} := MvPFunctor.comp.mk ∘ repr ∘ (map fun i ↦ (repr : G i α → (fun i : Fin2 n ↦ Obj (P (G i)) α) i)) ∘ Comp.get - abs_repr := by intros; simp only [(· ∘ ·), comp.get_mk, abs_repr, map_map, - TypeVec.comp, MvFunctor.id_map', Comp.mk_get] - abs_map := by intros; simp only [(· ∘ ·)]; rw [← abs_map] - simp only [comp.get_map, map_map, TypeVec.comp, abs_map, map_mk] + abs_repr := by + intros + simp (config := { unfoldPartialApp := true }) only [Function.comp_def, comp.get_mk, abs_repr, + map_map, TypeVec.comp, MvFunctor.id_map', Comp.mk_get] + abs_map := by + intros + simp only [(· ∘ ·)] + rw [← abs_map] + simp (config := { unfoldPartialApp := true }) only [comp.get_map, map_map, TypeVec.comp, + abs_map, map_mk] end Comp diff --git a/Mathlib/Data/QPF/Univariate/Basic.lean b/Mathlib/Data/QPF/Univariate/Basic.lean index d3bee0e10762a..ee47afa0ec54e 100644 --- a/Mathlib/Data/QPF/Univariate/Basic.lean +++ b/Mathlib/Data/QPF/Univariate/Basic.lean @@ -547,7 +547,7 @@ def comp : QPF (Functor.Comp F₂ F₁) where cases' h' : repr (f x) with b g dsimp; rw [← h', abs_repr] abs_map {α β} f := by - dsimp [Functor.Comp, PFunctor.comp] + dsimp (config := { unfoldPartialApp := true }) [Functor.Comp, PFunctor.comp] intro p cases' p with a g; dsimp cases' a with b h; dsimp @@ -557,7 +557,7 @@ def comp : QPF (Functor.Comp F₂ F₁) where apply abs_map congr rw [PFunctor.map_eq] - dsimp [Function.comp] + dsimp [Function.comp_def] congr ext x rw [← abs_map] diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index 8469f4b7101fd..baee0240d5f37 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -804,7 +804,7 @@ section variable [s : Setoid α] -protected theorem mk''_eq_mk (x : α) : Quotient.mk'' x = Quotient.mk s x := +protected theorem mk''_eq_mk : Quotient.mk'' = Quotient.mk s := rfl @[simp] diff --git a/Mathlib/Data/Rat/Cast/CharZero.lean b/Mathlib/Data/Rat/Cast/CharZero.lean index a7ea66b826325..22be3631afb0f 100644 --- a/Mathlib/Data/Rat/Cast/CharZero.lean +++ b/Mathlib/Data/Rat/Cast/CharZero.lean @@ -121,7 +121,7 @@ theorem cast_mk (a b : ℤ) : (a /. b : α) = a / b := by #align rat.cast_mk Rat.cast_mk @[simp, norm_cast] -theorem cast_pow (q) (k : ℕ) : ((q : ℚ) ^ k : α) = (q : α) ^ k := +theorem cast_pow (q : ℚ) (k : ℕ) : (↑(q ^ k) : α) = (q : α) ^ k := (castHom α).map_pow q k #align rat.cast_pow Rat.cast_pow diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 785026a02536e..10e5c3c238590 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -240,14 +240,16 @@ theorem divInt_zero_one : 0 /. 1 = 0 := theorem divInt_one_one : 1 /. 1 = 1 := show divInt _ _ = _ by rw [divInt] - simp + simp [mkRat, normalize] + rfl #align rat.mk_one_one Rat.divInt_one_one @[simp] theorem divInt_neg_one_one : -1 /. 1 = -1 := show divInt _ _ = _ by rw [divInt] - simp + simp [mkRat, normalize] + rfl #align rat.mk_neg_one_one Rat.divInt_neg_one_one theorem divInt_one (n : ℤ) : n /. 1 = n := diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index 01ab28303b775..1a09e6d8b1a18 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -233,6 +233,7 @@ theorem coe_int_div_self (n : ℤ) : ((n / n : ℤ) : ℚ) = n / n := by by_cases hn : n = 0 · subst hn simp only [Int.cast_zero, Int.zero_div, zero_div] + rfl · have : (n : ℚ) ≠ 0 := by rwa [← coe_int_inj] at hn simp only [Int.ediv_self hn, Int.cast_one, Ne.def, not_false_iff, div_self this] #align rat.coe_int_div_self Rat.coe_int_div_self diff --git a/Mathlib/Data/Real/ENNReal.lean b/Mathlib/Data/Real/ENNReal.lean index f017fbbd1c812..7564da9aec680 100644 --- a/Mathlib/Data/Real/ENNReal.lean +++ b/Mathlib/Data/Real/ENNReal.lean @@ -343,6 +343,8 @@ theorem toReal_ofReal_eq_iff {a : ℝ} : (ENNReal.ofReal a).toReal = a ↔ 0 ≤ @[simp] theorem top_ne_one : ∞ ≠ 1 := top_ne_coe #align ennreal.top_ne_one ENNReal.top_ne_one +@[simp] theorem zero_lt_top : 0 < ∞ := coe_lt_top + @[simp, norm_cast] theorem coe_eq_coe : (↑r : ℝ≥0∞) = ↑q ↔ r = q := WithTop.coe_eq_coe #align ennreal.coe_eq_coe ENNReal.coe_eq_coe @@ -429,9 +431,11 @@ nonrec theorem one_lt_two : (1 : ℝ≥0∞) < 2 := coe_one ▸ coe_two ▸ by exact_mod_cast (one_lt_two : 1 < 2) #align ennreal.one_lt_two ENNReal.one_lt_two -theorem two_ne_top : (2 : ℝ≥0∞) ≠ ∞ := coe_ne_top +@[simp] theorem two_ne_top : (2 : ℝ≥0∞) ≠ ∞ := coe_ne_top #align ennreal.two_ne_top ENNReal.two_ne_top +@[simp] theorem two_lt_top : (2 : ℝ≥0∞) < ∞ := coe_lt_top + /-- `(1 : ℝ≥0∞) ≤ 1`, recorded as a `Fact` for use with `Lp` spaces. -/ instance _root_.fact_one_le_one_ennreal : Fact ((1 : ℝ≥0∞) ≤ 1) := ⟨le_rfl⟩ @@ -636,7 +640,7 @@ theorem mul_lt_top_iff {a b : ℝ≥0∞} : a * b < ∞ ↔ a < ∞ ∧ b < ∞ theorem mul_self_lt_top_iff {a : ℝ≥0∞} : a * a < ⊤ ↔ a < ⊤ := by rw [ENNReal.mul_lt_top_iff, and_self, or_self, or_iff_left_iff_imp] rintro rfl - norm_num + decide #align ennreal.mul_self_lt_top_iff ENNReal.mul_self_lt_top_iff theorem mul_pos_iff : 0 < a * b ↔ 0 < a ∧ 0 < b := @@ -652,7 +656,7 @@ theorem mul_pos (ha : a ≠ 0) (hb : b ≠ 0) : 0 < a * b := rcases n.eq_zero_or_pos with rfl | (hn : 0 < n) · simp · induction a using recTopCoe - · simp only [Ne.def, hn.ne', top_pow hn] + · simp only [Ne.def, hn.ne', top_pow hn, not_false_eq_true, and_self] · simp only [← coe_pow, coe_ne_top, false_and] #align ennreal.pow_eq_top_iff ENNReal.pow_eq_top_iff @@ -1796,7 +1800,8 @@ theorem add_thirds (a : ℝ≥0∞) : a / 3 + a / 3 + a / 3 = a := by @[simp] theorem div_pos_iff : 0 < a / b ↔ a ≠ 0 ∧ b ≠ ∞ := by simp [pos_iff_ne_zero, not_or] #align ennreal.div_pos_iff ENNReal.div_pos_iff -protected theorem half_pos (h : a ≠ 0) : 0 < a / 2 := by simp [h] +protected theorem half_pos (h : a ≠ 0) : 0 < a / 2 := by + simp only [div_pos_iff, ne_eq, h, not_false_eq_true]; decide #align ennreal.half_pos ENNReal.half_pos protected theorem one_half_lt_one : (2⁻¹ : ℝ≥0∞) < 1 := @@ -1917,7 +1922,8 @@ theorem coe_zpow (hr : r ≠ 0) (n : ℤ) : (↑(r ^ n) : ℝ≥0∞) = (r : ℝ theorem zpow_pos (ha : a ≠ 0) (h'a : a ≠ ∞) (n : ℤ) : 0 < a ^ n := by cases n · exact ENNReal.pow_pos ha.bot_lt _ - · simp only [h'a, pow_eq_top_iff, zpow_negSucc, Ne.def, not_false, ENNReal.inv_pos, false_and] + · simp only [h'a, pow_eq_top_iff, zpow_negSucc, Ne.def, not_false, ENNReal.inv_pos, false_and, + not_false_eq_true] #align ennreal.zpow_pos ENNReal.zpow_pos theorem zpow_lt_top (ha : a ≠ 0) (h'a : a ≠ ∞) (n : ℤ) : a ^ n < ∞ := by diff --git a/Mathlib/Data/Real/Pi/Bounds.lean b/Mathlib/Data/Real/Pi/Bounds.lean index 34d0be3eeddae..11e758316ab7a 100644 --- a/Mathlib/Data/Real/Pi/Bounds.lean +++ b/Mathlib/Data/Real/Pi/Bounds.lean @@ -19,8 +19,6 @@ See also `Mathlib/Data/Real/Pi/Leibniz.lean` and `Mathlib/Data/Real/Pi/Wallis.le formulas for `π`. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - -- Porting note: needed to add a lot of type ascriptions for lean to interpret numbers as reals. open scoped Real diff --git a/Mathlib/Data/Real/Pi/Leibniz.lean b/Mathlib/Data/Real/Pi/Leibniz.lean index b2119ea2a569c..644b510307692 100644 --- a/Mathlib/Data/Real/Pi/Leibniz.lean +++ b/Mathlib/Data/Real/Pi/Leibniz.lean @@ -12,8 +12,6 @@ import Mathlib.Analysis.SpecialFunctions.Trigonometric.ArctanDeriv namespace Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Filter Set open scoped Classical BigOperators Topology Real diff --git a/Mathlib/Data/Real/Pi/Wallis.lean b/Mathlib/Data/Real/Pi/Wallis.lean index 72c1ced1777b3..0fd4d8a5e2324 100644 --- a/Mathlib/Data/Real/Pi/Wallis.lean +++ b/Mathlib/Data/Real/Pi/Wallis.lean @@ -40,8 +40,6 @@ namespace Real namespace Wallis -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - set_option linter.uppercaseLean3 false /-- The product of the first `k` terms in Wallis' formula for `π`. -/ @@ -95,8 +93,8 @@ theorem le_W (k : ℕ) : ((2 : ℝ) * k + 1) / (2 * k + 2) * (π / 2) ≤ W k := rw [W_eq_integral_sin_pow_div_integral_sin_pow, le_div_iff (integral_sin_pow_pos _)] convert integral_sin_pow_succ_le (2 * k + 1) rw [integral_sin_pow (2 * k)] - simp only [sin_zero, zero_pow', Ne.def, Nat.succ_ne_zero, zero_mul, sin_pi, tsub_zero, zero_div, - zero_add] + simp only [sin_zero, ne_eq, add_eq_zero, and_false, not_false_eq_true, zero_pow', cos_zero, + mul_one, sin_pi, cos_pi, mul_neg, neg_zero, sub_self, zero_div, zero_add] norm_cast #align real.wallis.le_W Real.Wallis.le_W diff --git a/Mathlib/Data/Rel.lean b/Mathlib/Data/Rel.lean index 347800e2c5121..e004ce8a9d1d1 100644 --- a/Mathlib/Data/Rel.lean +++ b/Mathlib/Data/Rel.lean @@ -148,10 +148,12 @@ theorem inv_comp (r : Rel α β) (s : Rel β γ) : inv (r • s) = inv s • inv #align rel.inv_comp Rel.inv_comp @[simp] -theorem inv_bot : (⊥ : Rel α β).inv = (⊥ : Rel β α) := by simp [Bot.bot, inv, flip] +theorem inv_bot : (⊥ : Rel α β).inv = (⊥ : Rel β α) := by + simp [Bot.bot, inv, flip] @[simp] -theorem inv_top : (⊤ : Rel α β).inv = (⊤ : Rel β α) := by simp [Top.top, inv, flip] +theorem inv_top : (⊤ : Rel α β).inv = (⊤ : Rel β α) := by + simp [Top.top, inv, flip] /-- Image of a set under a relation -/ def image (s : Set α) : Set β := { y | ∃ x ∈ s, r x y } @@ -363,7 +365,7 @@ def graph (f : α → β) : Rel α β := fun x y => f x = y @[simp] lemma graph_def (f : α → β) (x y) : f.graph x y ↔ (f x = y) := Iff.rfl -theorem graph_id : graph id = @Eq α := by simp [graph] +theorem graph_id : graph id = @Eq α := by simp (config := { unfoldPartialApp := true }) [graph] theorem graph_comp {f : β → γ} {g : α → β} : graph (f ∘ g) = Rel.comp (graph g) (graph f) := by ext x y diff --git a/Mathlib/Data/Seq/Parallel.lean b/Mathlib/Data/Seq/Parallel.lean index cf4e2808a7859..5042278714e3e 100644 --- a/Mathlib/Data/Seq/Parallel.lean +++ b/Mathlib/Data/Seq/Parallel.lean @@ -328,7 +328,7 @@ def parallelRec {S : WSeq (Computation α)} (C : α → Sort v) (H : ∀ s ∈ S rw [← WSeq.map_comp] refine' (WSeq.map_id _).symm.trans (congr_arg (fun f => WSeq.map f S) _) funext c - dsimp [id, Function.comp] + dsimp [id, Function.comp_def] rw [← map_comp] exact (map_id _).symm have pe := congr_arg parallel this diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 01f631701632c..657fe4156daf1 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -1616,7 +1616,7 @@ theorem liftRel_join.lem (R : α → β → Prop) {S T} {U : WSeq α → WSeq β apply Nat.lt_succ_of_le (Nat.le_add_right _ _) let ⟨ob, mb, rob⟩ := IH _ this ST' rs5' refine' ⟨ob, _, rob⟩ - · simp [destruct_join] + · simp (config := { unfoldPartialApp := true }) [destruct_join] apply mem_bind mT simp only [destruct_append, destruct_append.aux] apply think_mem @@ -1625,7 +1625,7 @@ theorem liftRel_join.lem (R : α → β → Prop) {S T} {U : WSeq α → WSeq β | some (a, s'), some (b, t'), ⟨ab, st'⟩, _, rs5, mt => by simp at rs5 refine' ⟨some (b, append t' (join T')), _, _⟩ - · simp [destruct_join] + · simp (config := { unfoldPartialApp := true }) [destruct_join] apply mem_bind mT simp only [destruct_append, destruct_append.aux] apply think_mem @@ -1697,7 +1697,8 @@ theorem join_map_ret (s : WSeq α) : join (map ret s) ~ʷ s := by (map (fun a => cons a nil) s).join.destruct = (map (fun a => cons a nil) s').join.destruct ∧ destruct s = s'.destruct := fun s => ⟨s, rfl, rfl⟩ - induction' s using WSeq.recOn with a s s <;> simp [ret, ret_mem, this, Option.exists] + induction' s using WSeq.recOn with a s s <;> + simp (config := { unfoldPartialApp := true }) [ret, ret_mem, this, Option.exists] · exact ⟨s, rfl, rfl⟩ #align stream.wseq.join_map_ret Stream'.WSeq.join_map_ret diff --git a/Mathlib/Data/Set/Card.lean b/Mathlib/Data/Set/Card.lean index 8fa273b37547a..12f604840ca64 100644 --- a/Mathlib/Data/Set/Card.lean +++ b/Mathlib/Data/Set/Card.lean @@ -555,7 +555,7 @@ theorem nonempty_of_ncard_ne_zero (hs : s.ncard ≠ 0) : s.Nonempty := by #align set.nonempty_of_ncard_ne_zero Set.nonempty_of_ncard_ne_zero @[simp] theorem ncard_singleton (a : α) : ({a} : Set α).ncard = 1 := by - simp [ncard_eq_toFinset_card] + simp [ncard, ncard_eq_toFinset_card] #align set.ncard_singleton Set.ncard_singleton theorem ncard_singleton_inter (a : α) (s : Set α) : ({a} ∩ s).ncard ≤ 1 := by diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 25f04211d4ac4..8d976c1b5e019 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -1444,7 +1444,7 @@ instance Compl.decidableMem (j : α) : Decidable (j ∈ sᶜ) := theorem piecewise_insert [DecidableEq α] (j : α) [∀ i, Decidable (i ∈ insert j s)] : (insert j s).piecewise f g = Function.update (s.piecewise f g) j (f j) := by - simp only [piecewise, mem_insert_iff] + simp (config := { unfoldPartialApp := true }) only [piecewise, mem_insert_iff] ext i by_cases h : i = j · rw [h] diff --git a/Mathlib/Data/Set/Intervals/Group.lean b/Mathlib/Data/Set/Intervals/Group.lean index e192454669410..021b799751759 100644 --- a/Mathlib/Data/Set/Intervals/Group.lean +++ b/Mathlib/Data/Set/Intervals/Group.lean @@ -167,7 +167,8 @@ variable [OrderedCommGroup α] (a b : α) @[to_additive] theorem pairwise_disjoint_Ioc_mul_zpow : Pairwise (Disjoint on fun n : ℤ => Ioc (a * b ^ n) (a * b ^ (n + 1))) := by - simp_rw [Function.onFun, Set.disjoint_iff] + simp (config := { unfoldPartialApp := true }) only [Function.onFun] + simp_rw [Set.disjoint_iff] intro m n hmn x hx apply hmn have hb : 1 < b := by @@ -183,7 +184,8 @@ theorem pairwise_disjoint_Ioc_mul_zpow : @[to_additive] theorem pairwise_disjoint_Ico_mul_zpow : Pairwise (Disjoint on fun n : ℤ => Ico (a * b ^ n) (a * b ^ (n + 1))) := by - simp_rw [Function.onFun, Set.disjoint_iff] + simp (config := { unfoldPartialApp := true }) only [Function.onFun] + simp_rw [Set.disjoint_iff] intro m n hmn x hx apply hmn have hb : 1 < b := by diff --git a/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean b/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean index 292632f6f1339..0bee1b87f3725 100644 --- a/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean +++ b/Mathlib/Data/Set/Intervals/OrdConnectedComponent.lean @@ -140,7 +140,8 @@ def ordConnectedSection (s : Set α) : Set α := theorem dual_ordConnectedSection (s : Set α) : ordConnectedSection (ofDual ⁻¹' s) = ofDual ⁻¹' ordConnectedSection s := by - simp_rw [ordConnectedSection, ordConnectedProj] + simp only [ordConnectedSection] + simp (config := { unfoldPartialApp := true }) only [ordConnectedProj] ext x simp only [mem_range, Subtype.exists, mem_preimage, OrderDual.exists, dual_ordConnectedComponent, ofDual_toDual] diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 73b7d72e731ba..96b275784725c 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -316,49 +316,49 @@ theorem image_neg_Ioo : Neg.neg '' Ioo a b = Ioo (-b) (-a) := by simp @[simp] theorem image_const_sub_Ici : (fun x => a - x) '' Ici b = Iic (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Ici Set.image_const_sub_Ici @[simp] theorem image_const_sub_Iic : (fun x => a - x) '' Iic b = Ici (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Iic Set.image_const_sub_Iic @[simp] theorem image_const_sub_Ioi : (fun x => a - x) '' Ioi b = Iio (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Ioi Set.image_const_sub_Ioi @[simp] theorem image_const_sub_Iio : (fun x => a - x) '' Iio b = Ioi (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Iio Set.image_const_sub_Iio @[simp] theorem image_const_sub_Icc : (fun x => a - x) '' Icc b c = Icc (a - c) (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Icc Set.image_const_sub_Icc @[simp] theorem image_const_sub_Ico : (fun x => a - x) '' Ico b c = Ioc (a - c) (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Ico Set.image_const_sub_Ico @[simp] theorem image_const_sub_Ioc : (fun x => a - x) '' Ioc b c = Ico (a - c) (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Ioc Set.image_const_sub_Ioc @[simp] theorem image_const_sub_Ioo : (fun x => a - x) '' Ioo b c = Ioo (a - c) (a - b) := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_Ioo Set.image_const_sub_Ioo @@ -459,7 +459,7 @@ theorem image_add_const_uIcc : (fun x => x + a) '' [[b, c]] = [[b + a, c + a]] : @[simp] theorem image_const_sub_uIcc : (fun x => a - x) '' [[b, c]] = [[a - b, a - c]] := by - have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp] at this + have := image_comp (fun x => a + x) fun x => -x; dsimp [Function.comp_def] at this simp [sub_eq_add_neg, this, add_comm] #align set.image_const_sub_uIcc Set.image_const_sub_uIcc diff --git a/Mathlib/Data/Sign.lean b/Mathlib/Data/Sign.lean index d51cd58790179..ae98a5158f3eb 100644 --- a/Mathlib/Data/Sign.lean +++ b/Mathlib/Data/Sign.lean @@ -34,7 +34,7 @@ namespace SignType -- Porting note: Added Fintype SignType manually instance : Fintype SignType := - Fintype.ofMultiset (zero :: neg :: pos :: List.nil) (fun x ↦ by cases x <;> simp only) + Fintype.ofMultiset (zero :: neg :: pos :: List.nil) (fun x ↦ by cases x <;> decide) instance : Zero SignType := ⟨zero⟩ @@ -155,7 +155,7 @@ def fin3Equiv : SignType ≃* Fin 3 where | ⟨2, _⟩ => by simp | ⟨n + 3, h⟩ => by simp at h map_mul' a b := by - cases a <;> cases b <;> simp + cases a <;> cases b <;> rfl #align sign_type.fin3_equiv SignType.fin3Equiv section CaseBashing diff --git a/Mathlib/Data/String/Basic.lean b/Mathlib/Data/String/Basic.lean index 633bb00b8c2a9..e1e2b6da63107 100644 --- a/Mathlib/Data/String/Basic.lean +++ b/Mathlib/Data/String/Basic.lean @@ -75,13 +75,13 @@ theorem ltb_cons_addChar (c : Char) (cs₁ cs₂ : List Char) (i₁ i₂ : Pos) theorem lt_iff_toList_lt : ∀ {s₁ s₂ : String}, s₁ < s₂ ↔ s₁.toList < s₂.toList | ⟨s₁⟩, ⟨s₂⟩ => show ltb ⟨⟨s₁⟩, 0⟩ ⟨⟨s₂⟩, 0⟩ ↔ s₁ < s₂ by induction s₁ generalizing s₂ <;> cases s₂ - · simp + · decide · rename_i c₂ cs₂; apply iff_of_true - · rw [ltb]; simp; apply ne_false_of_eq_true; apply decide_eq_true + · rw [ltb]; simp only [Iterator.hasNext, Iterator.curr] simp [endPos, utf8ByteSize, utf8ByteSize.go, csize_pos] · apply List.nil_lt_cons · rename_i c₁ cs₁ ih; apply iff_of_false - · rw [ltb]; simp + · rw [ltb]; simp [Iterator.hasNext, Iterator.curr] · apply not_lt_of_lt; apply List.nil_lt_cons · rename_i c₁ cs₁ ih c₂ cs₂; rw [ltb] simp [Iterator.hasNext, endPos, utf8ByteSize, utf8ByteSize.go, csize_pos, Iterator.curr, get, @@ -134,7 +134,7 @@ theorem asString_inv_toList (s : String) : s.toList.asString = s := theorem toList_nonempty : ∀ {s : String}, s ≠ "" → s.toList = s.head :: (s.drop 1).toList | ⟨s⟩, h => by cases s - · simp only at h + · simp only [ne_eq, not_true_eq_false] at h · rename_i c cs simp only [toList, List.cons.injEq] constructor <;> [rfl; simp [drop_eq]] diff --git a/Mathlib/Data/Sym/Card.lean b/Mathlib/Data/Sym/Card.lean index d02afefd6229f..5f0b0c6c769fc 100644 --- a/Mathlib/Data/Sym/Card.lean +++ b/Mathlib/Data/Sym/Card.lean @@ -171,7 +171,7 @@ theorem card_image_offDiag (s : Finset α) : theorem card_subtype_diag [Fintype α] : card { a : Sym2 α // a.IsDiag } = card α := by convert card_image_diag (univ : Finset α) - simp_rw [Quotient.mk', ← Quotient.mk''_eq_mk] -- Porting note: Added `simp_rw` + simp_rw [Quotient.mk'_eq_mk, ← Quotient.mk''_eq_mk] rw [Fintype.card_of_subtype, ← filter_image_quotient_mk''_isDiag] rintro x rw [mem_filter, univ_product_univ, mem_image] @@ -182,7 +182,7 @@ theorem card_subtype_diag [Fintype α] : card { a : Sym2 α // a.IsDiag } = card theorem card_subtype_not_diag [Fintype α] : card { a : Sym2 α // ¬a.IsDiag } = (card α).choose 2 := by convert card_image_offDiag (univ : Finset α) - simp_rw [Quotient.mk', ← Quotient.mk''_eq_mk] -- Porting note: Added `simp_rw` + simp_rw [Quotient.mk'_eq_mk, ← Quotient.mk''_eq_mk] -- Porting note: Added `simp_rw` rw [Fintype.card_of_subtype, ← filter_image_quotient_mk''_not_isDiag] rintro x rw [mem_filter, univ_product_univ, mem_image] diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index 5de6d7b454cd7..91a4ff303b891 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -853,7 +853,7 @@ theorem neg_one_ne_one {n : ℕ} [Fact (2 < n)] : (-1 : ZMod n) ≠ 1 := #align zmod.neg_one_ne_one ZMod.neg_one_ne_one theorem neg_eq_self_mod_two (a : ZMod 2) : -a = a := by - fin_cases a <;> apply Fin.ext <;> simp [Fin.coe_neg, Int.natMod] + fin_cases a <;> apply Fin.ext <;> simp [Fin.coe_neg, Int.natMod]; rfl #align zmod.neg_eq_self_mod_two ZMod.neg_eq_self_mod_two @[simp] diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 79ff46b5386a0..b32f06053f921 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -24,8 +24,6 @@ by using `Module R (Additive M)` in its place, especially since this already has `R = ℕ` and `R = ℤ`. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - instance : SMul (ZMod 2) (Additive ℤˣ) where smul z au := .ofMul <| Additive.toMul au ^ z.val diff --git a/Mathlib/FieldTheory/AbelRuffini.lean b/Mathlib/FieldTheory/AbelRuffini.lean index 6f9a8a2d764d1..c5651d218db69 100644 --- a/Mathlib/FieldTheory/AbelRuffini.lean +++ b/Mathlib/FieldTheory/AbelRuffini.lean @@ -28,8 +28,6 @@ that is solvable by radicals has a solvable Galois group. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped Classical Polynomial IntermediateField open Polynomial IntermediateField diff --git a/Mathlib/FieldTheory/IsAlgClosed/Spectrum.lean b/Mathlib/FieldTheory/IsAlgClosed/Spectrum.lean index 923e92883b264..be981014ed60d 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Spectrum.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Spectrum.lean @@ -35,8 +35,6 @@ eigenvalue. * `σ a` : `spectrum R a` of `a : A` -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace spectrum open Set Polynomial diff --git a/Mathlib/FieldTheory/PrimitiveElement.lean b/Mathlib/FieldTheory/PrimitiveElement.lean index 0a81327e53788..7996039dc1e99 100644 --- a/Mathlib/FieldTheory/PrimitiveElement.lean +++ b/Mathlib/FieldTheory/PrimitiveElement.lean @@ -65,7 +65,7 @@ theorem exists_primitive_element_of_finite_top [Finite E] : ∃ α : E, F⟮α exact F⟮α.val⟯.zero_mem · obtain ⟨n, hn⟩ := Set.mem_range.mp (hα (Units.mk0 x hx)) simp only at hn - rw [show x = α ^ n by norm_cast; rw [hn, Units.val_mk0], Units.val_zpow_eq_zpow_val] + rw [show x = α ^ n by norm_cast; rw [hn, Units.val_mk0]] exact zpow_mem (mem_adjoin_simple_self F (E := E) ↑α) n #align field.exists_primitive_element_of_finite_top Field.exists_primitive_element_of_finite_top @@ -261,7 +261,7 @@ theorem isAlgebraic_of_adjoin_eq_adjoin {α : E} {m n : ℕ} (hneq : m ≠ n) · simp only [hzero, div_zero, pow_eq_zero_iff hm] at h exact h.symm ▸ isAlgebraic_zero · rw [eq_div_iff hzero, ← sub_eq_zero] at h - replace hzero : s ≠ 0 := by rintro rfl; simp only [map_zero] at hzero + replace hzero : s ≠ 0 := by rintro rfl; simp only [map_zero, not_true_eq_false] at hzero let f : F[X] := X ^ m * expand F n s - expand F n r refine ⟨f, ?_, ?_⟩ · have : f.coeff (n * s.natDegree + m) ≠ 0 := by @@ -273,7 +273,7 @@ theorem isAlgebraic_of_adjoin_eq_adjoin {α : E} {m n : ℕ} (hneq : m ≠ n) coeff_expand hn r, hndvd, ite_false, sub_zero] exact leadingCoeff_ne_zero.2 hzero intro h - simp only [h, coeff_zero, ne_eq] at this + simp only [h, coeff_zero, ne_eq, not_true_eq_false] at this · simp only [map_sub, map_mul, map_pow, aeval_X, expand_aeval, h] theorem isAlgebraic_of_finite_intermediateField diff --git a/Mathlib/FieldTheory/RatFunc.lean b/Mathlib/FieldTheory/RatFunc.lean index 0de15cd44e640..643ca0dd850d0 100644 --- a/Mathlib/FieldTheory/RatFunc.lean +++ b/Mathlib/FieldTheory/RatFunc.lean @@ -689,6 +689,7 @@ theorem coe_mapRingHom_eq_coe_map [RingHomClass F R[X] S[X]] (φ : F) (hφ : R[X rfl #align ratfunc.coe_map_ring_hom_eq_coe_map RatFunc.coe_mapRingHom_eq_coe_map +set_option maxHeartbeats 300000 in -- TODO: Generalize to `FunLike` classes, /-- Lift a monoid with zero homomorphism `R[X] →*₀ G₀` to a `RatFunc R →*₀ G₀` on the condition that `φ` maps non zero divisors to non zero divisors, @@ -703,7 +704,7 @@ def liftMonoidWithZeroHom (φ : R[X] →*₀ G₀) (hφ : R[X]⁰ ≤ G₀⁰.co map_one' := by dsimp only -- porting note: force the function to be applied rw [← ofFractionRing_one, ← Localization.mk_one, liftOn_ofFractionRing_mk] - simp only [map_one, Submonoid.coe_one, div_one] + simp only [map_one, OneMemClass.coe_one, div_one] map_mul' x y := by cases' x with x cases' y with y diff --git a/Mathlib/FieldTheory/Subfield.lean b/Mathlib/FieldTheory/Subfield.lean index d704a6749261e..dd2fd753336da 100644 --- a/Mathlib/FieldTheory/Subfield.lean +++ b/Mathlib/FieldTheory/Subfield.lean @@ -116,7 +116,7 @@ instance (s : S) : SMul ℚ s := ⟨fun a x => ⟨a • (x : K), rat_smul_mem s a x⟩⟩ @[simp] -theorem coe_rat_smul (s : S) (a : ℚ) (x : s) : (a • x : K) = a • (x : K) := +theorem coe_rat_smul (s : S) (a : ℚ) (x : s) : ↑(a • x) = a • (x : K) := rfl #align subfield_class.coe_rat_smul SubfieldClass.coe_rat_smul diff --git a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean index e56f6c784262a..9facc73f8ca09 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Oriented/Basic.lean @@ -79,7 +79,7 @@ theorem oangle_zero_right (x : V) : o.oangle x 0 = 0 := by simp [oangle] /-- If the two vectors passed to `oangle` are the same, the result is 0. -/ @[simp] theorem oangle_self (x : V) : o.oangle x x = 0 := by - rw [oangle, kahler_apply_self]; norm_cast + rw [oangle, kahler_apply_self, ← ofReal_pow]; norm_cast convert QuotientAddGroup.mk_zero (AddSubgroup.zmultiples (2 * π)) apply arg_ofReal_of_nonneg positivity diff --git a/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean b/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean index 2f81705007ce4..a27f45ae1fad7 100644 --- a/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean +++ b/Mathlib/Geometry/Euclidean/Angle/Unoriented/RightAngle.lean @@ -29,8 +29,6 @@ triangle unnecessarily. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped BigOperators open scoped EuclideanGeometry diff --git a/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean b/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean index 6520cf3a25b19..c2a4d1240abff 100644 --- a/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean +++ b/Mathlib/Geometry/Euclidean/Inversion/Calculus.lean @@ -22,8 +22,6 @@ space in this file. inversion, derivative -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Metric Function AffineMap Set AffineSubspace open scoped Topology RealInnerProductSpace @@ -90,7 +88,8 @@ theorem hasFDerivAt_inversion (hx : x ≠ c) : ((R / dist x c) ^ 2 • (reflection (ℝ ∙ (x - c))ᗮ : F →L[ℝ] F)) x := by rcases add_left_surjective c x with ⟨x, rfl⟩ have : HasFDerivAt (inversion c R) (_ : F →L[ℝ] F) (c + x) - · simp_rw [inversion, dist_eq_norm, div_pow, div_eq_mul_inv] + · simp (config := { unfoldPartialApp := true }) only [inversion] + simp_rw [dist_eq_norm, div_pow, div_eq_mul_inv] have A := (hasFDerivAt_id (𝕜 := ℝ) (c + x)).sub_const c have B := ((hasDerivAt_inv <| by simpa using hx).comp_hasFDerivAt _ A.norm_sq).const_mul (R ^ 2) diff --git a/Mathlib/Geometry/Euclidean/MongePoint.lean b/Mathlib/Geometry/Euclidean/MongePoint.lean index 9432a3a393167..ed39ed20f5d5f 100644 --- a/Mathlib/Geometry/Euclidean/MongePoint.lean +++ b/Mathlib/Geometry/Euclidean/MongePoint.lean @@ -474,16 +474,18 @@ theorem altitude_eq_mongePlane (t : Triangle ℝ P) {i₁ i₂ i₃ : Fin 3} (h (h₂₃ : i₂ ≠ i₃) : t.altitude i₁ = t.mongePlane i₂ i₃ := by have hs : ({i₂, i₃}ᶜ : Finset (Fin 3)) = {i₁} := by -- porting note: was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at h₁₂ h₁₃ h₂₃ ⊢ have he : univ.erase i₁ = {i₂, i₃} := by -- porting note: was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at h₁₂ h₁₃ h₂₃ ⊢ rw [mongePlane_def, altitude_def, direction_affineSpan, hs, he, centroid_singleton, coe_insert, coe_singleton, vectorSpan_image_eq_span_vsub_set_left_ne ℝ _ (Set.mem_insert i₂ _)] simp [h₂₃, Submodule.span_insert_eq_span] -- porting note: this didn't need the `congr` and the `fin_cases` congr - fin_cases i₂ <;> fin_cases i₃ <;> simp at h₂₃ ⊢ + fin_cases i₂ <;> fin_cases i₃ <;> simp (config := {decide := true}) at h₂₃ ⊢ #align affine.triangle.altitude_eq_monge_plane Affine.Triangle.altitude_eq_mongePlane /-- The orthocenter lies in the altitudes. -/ @@ -589,7 +591,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} · have hu : (Finset.univ : Finset (Fin 3)) = {j₁, j₂, j₃} := by clear h₁ h₂ h₃ -- porting note: was `decide!` - fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ <;> simp at hj₁₂ hj₁₃ hj₂₃ ⊢ + fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ + <;> simp (config := {decide := true}) at hj₁₂ hj₁₃ hj₂₃ ⊢ rw [← Set.image_univ, ← Finset.coe_univ, hu, Finset.coe_insert, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_insert_eq, Set.image_singleton, h₁, h₂, h₃, Set.insert_subset_iff, Set.insert_subset_iff, Set.singleton_subset_iff] @@ -604,7 +607,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} have hu : Finset.univ.erase j₂ = {j₁, j₃} := by clear h₁ h₂ h₃ -- porting note: was `decide!` - fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ <;> simp at hj₁₂ hj₁₃ hj₂₃ ⊢ + fin_cases j₁ <;> fin_cases j₂ <;> fin_cases j₃ + <;> simp (config := {decide := true}) at hj₁₂ hj₁₃ hj₂₃ ⊢ rw [hu, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_singleton, h₁, h₃] have hle : (t₁.altitude i₃).directionᗮ ≤ line[ℝ, t₁.orthocenter, t₁.points i₃].directionᗮ := Submodule.orthogonal_le (direction_le (affineSpan_orthocenter_point_le_altitude _ _)) @@ -612,7 +616,8 @@ theorem altitude_replace_orthocenter_eq_affineSpan {t₁ t₂ : Triangle ℝ P} have hui : Finset.univ.erase i₃ = {i₁, i₂} := by clear hle h₂ h₃ -- porting note: was `decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp at hi₁₂ hi₁₃ hi₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) at hi₁₂ hi₁₃ hi₂₃ ⊢ rw [hui, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_singleton] refine' vsub_mem_vectorSpan ℝ (Set.mem_insert _ _) (Set.mem_insert_of_mem _ (Set.mem_singleton _)) #align affine.triangle.altitude_replace_orthocenter_eq_affine_span Affine.Triangle.altitude_replace_orthocenter_eq_affineSpan @@ -783,7 +788,7 @@ theorem OrthocentricSystem.eq_insert_orthocenter {s : Set P} (ho : OrthocentricS ∃ j₁ : Fin 3, j₁ ≠ j₂ ∧ j₁ ≠ j₃ ∧ ∀ j : Fin 3, j = j₁ ∨ j = j₂ ∨ j = j₃ := by clear h₂ h₃ -- porting note: was `decide!` - fin_cases j₂ <;> fin_cases j₃ <;> simp at hj₂₃ ⊢ + fin_cases j₂ <;> fin_cases j₃ <;> simp (config := {decide := true}) at hj₂₃ ⊢ suffices h : t₀.points j₁ = t.orthocenter · have hui : (Set.univ : Set (Fin 3)) = {i₁, i₂, i₃} := by ext x; simpa using h₁₂₃ x have huj : (Set.univ : Set (Fin 3)) = {j₁, j₂, j₃} := by ext x; simpa using hj₁₂₃ x diff --git a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean index 8017fc00e9cc0..cb67bbb64f409 100644 --- a/Mathlib/Geometry/Euclidean/Sphere/Basic.lean +++ b/Mathlib/Geometry/Euclidean/Sphere/Basic.lean @@ -280,7 +280,7 @@ theorem Cospherical.affineIndependent_of_mem_of_ne {s : Set P} (hs : Cospherical refine' hs.affineIndependent _ _ · simp [h₁, h₂, h₃, Set.insert_subset_iff] · erw [Fin.cons_injective_iff, Fin.cons_injective_iff] - simp [h₁₂, h₁₃, h₂₃, Function.Injective] + simp [h₁₂, h₁₃, h₂₃, Function.Injective, eq_iff_true_of_subsingleton] #align euclidean_geometry.cospherical.affine_independent_of_mem_of_ne EuclideanGeometry.Cospherical.affineIndependent_of_mem_of_ne /-- The three points of a cospherical set are affinely independent. -/ diff --git a/Mathlib/Geometry/Euclidean/Triangle.lean b/Mathlib/Geometry/Euclidean/Triangle.lean index afe5d98760c97..940d0fe174958 100644 --- a/Mathlib/Geometry/Euclidean/Triangle.lean +++ b/Mathlib/Geometry/Euclidean/Triangle.lean @@ -38,8 +38,6 @@ unnecessarily. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped BigOperators open scoped Classical diff --git a/Mathlib/Geometry/Manifold/ChartedSpace.lean b/Mathlib/Geometry/Manifold/ChartedSpace.lean index a1de6c986ee61..5db698d70957e 100644 --- a/Mathlib/Geometry/Manifold/ChartedSpace.lean +++ b/Mathlib/Geometry/Manifold/ChartedSpace.lean @@ -1153,7 +1153,6 @@ protected instance instHasGroupoid [ClosedUnderRestriction G] : HasGroupoid s G compatible := by rintro e e' ⟨_, ⟨x, hc⟩, he⟩ ⟨_, ⟨x', hc'⟩, he'⟩ haveI : Nonempty s := ⟨x⟩ - have asdf := he rw [hc.symm, mem_singleton_iff] at he rw [hc'.symm, mem_singleton_iff] at he' rw [he, he'] diff --git a/Mathlib/Geometry/Manifold/ContMDiff.lean b/Mathlib/Geometry/Manifold/ContMDiff.lean index 1ca709132f713..3ee854cfa5962 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff.lean @@ -2114,7 +2114,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : LocalHomeomorph M M') : simp only [hx, h1, mfld_simps] have h2X : c' X = e (c (f.symm X)) := by rw [← hef hex] - dsimp only [Function.comp] + dsimp only [Function.comp_def] have hfX : f.symm X ∈ c.source := by simp only [hX, mfld_simps] rw [c.left_inv hfX, f.right_inv hX] have h3e : EqOn (c ∘ f.symm ∘ c'.symm) e.symm (c'.symm ⁻¹' f.target ∩ e.target) := by diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index 44477b6d96ce0..f12513c1cf22c 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -507,7 +507,7 @@ theorem ContMDiffOn.contMDiffOn_tangentMapWithin (hf : ContMDiffOn I I' n f s) ( refine' tangentMapWithin_subset (by mfld_set_tac) U'q _ apply hf.mdifferentiableOn one_le_n simp only [hq, mfld_simps] - dsimp only [(· ∘ ·)] at A B C D E ⊢ + dsimp only [Function.comp_def] at A B C D E ⊢ simp only [A, B, C, D, ← E] exact diff_DrirrflilDl.congr eq_comp #align cont_mdiff_on.cont_mdiff_on_tangent_map_within ContMDiffOn.contMDiffOn_tangentMapWithin @@ -581,9 +581,9 @@ theorem tangentMap_tangentBundle_pure (p : TangentBundle I M) : · exact differentiableAt_const _ · exact ModelWithCorners.unique_diff_at_image I · exact differentiableAt_id'.prod (differentiableAt_const _) - simp only [Bundle.zeroSection, tangentMap, mfderiv, A, if_pos, chartAt, - FiberBundle.chartedSpace_chartAt, TangentBundle.trivializationAt_apply, tangentBundleCore, - Function.comp, ContinuousLinearMap.map_zero, mfld_simps] + simp (config := { unfoldPartialApp := true }) only [Bundle.zeroSection, tangentMap, mfderiv, A, + if_pos, chartAt, FiberBundle.chartedSpace_chartAt, TangentBundle.trivializationAt_apply, + tangentBundleCore, Function.comp_def, ContinuousLinearMap.map_zero, mfld_simps] rw [← fderivWithin_inter N] at B rw [← fderivWithin_inter N, ← B] congr 1 diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 3154c94650cb3..c45388ea58b2b 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -72,8 +72,6 @@ open Metric FiniteDimensional Function open scoped Manifold -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - section StereographicProjection variable (v : E) @@ -426,8 +424,6 @@ instance smoothMfldWithCorners {n : ℕ} [Fact (finrank ℝ E = n + 1)] : (-- Removed type ascription, and this helped for some reason with timeout issues? OrthonormalBasis.fromOrthogonalSpanSingleton (𝕜 := ℝ) n (ne_zero_of_mem_unit_sphere v')).repr - -- Porting note: trouble synth instances - have := findim (E := E) n have H₁ := U'.contDiff.comp_contDiffOn contDiffOn_stereoToFun -- Porting note: need to help with implicit variables again have H₂ := (contDiff_stereoInvFunAux (v := v.val)|>.comp diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index f397b6e842224..aeb31731a39da 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -423,7 +423,8 @@ def inTangentCoordinates (f : N → M) (g : N → M') (ϕ : N → E →L[𝕜] E theorem inTangentCoordinates_model_space (f : N → H) (g : N → H') (ϕ : N → E →L[𝕜] E') (x₀ : N) : inTangentCoordinates I I' f g ϕ x₀ = ϕ := by - simp_rw [inTangentCoordinates, inCoordinates_tangent_bundle_core_model_space] + simp (config := { unfoldPartialApp := true }) only [inTangentCoordinates, + inCoordinates_tangent_bundle_core_model_space] #align in_tangent_coordinates_model_space inTangentCoordinates_model_space theorem inTangentCoordinates_eq (f : N → M) (g : N → M') (ϕ : N → E →L[𝕜] E') {x₀ x : N} diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index 26b9363e0ad6d..11b4a6ef32cc0 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -209,13 +209,13 @@ theorem commProb_reciprocal (n : ℕ) : rcases Nat.even_or_odd n with h2 | h2 · have := div_two_lt h0 rw [reciprocalFactors_even h0 h2, commProb_cons, commProb_reciprocal (n / 2), - commProb_odd (by norm_num)] + commProb_odd (by decide)] field_simp [h0, h2.two_dvd] norm_num · have := div_four_lt h0 h1 rw [reciprocalFactors_odd h1 h2, commProb_cons, commProb_reciprocal (n / 4 + 1)] have key : n % 4 = 1 ∨ n % 4 = 3 := Nat.odd_mod_four_iff.mp (Nat.odd_iff.mp h2) - have hn : Odd (n % 4) := by rcases key with h | h <;> rw [h] <;> norm_num + have hn : Odd (n % 4) := by rcases key with h | h <;> rw [h] <;> decide rw [commProb_odd (hn.mul h2), div_mul_div_comm, mul_one, div_eq_div_iff, one_mul] <;> norm_cast · have h0 : (n % 4) ^ 2 + 3 = n % 4 * 4 := by rcases key with h | h <;> rw [h] <;> norm_num have h1 := (Nat.div_add_mod n 4).symm diff --git a/Mathlib/GroupTheory/FiniteAbelian.lean b/Mathlib/GroupTheory/FiniteAbelian.lean index 9ab087f7b8a49..64243cd968c62 100644 --- a/Mathlib/GroupTheory/FiniteAbelian.lean +++ b/Mathlib/GroupTheory/FiniteAbelian.lean @@ -75,7 +75,7 @@ theorem equiv_directSum_zmod_of_fintype [Finite G] : obtain ⟨n, ι, fι, p, hp, e, ⟨f⟩⟩ := equiv_free_prod_directSum_zmod G cases' n with n · have : Unique (Fin Nat.zero →₀ ℤ) := - { uniq := by simp only [Nat.zero_eq, eq_iff_true_of_subsingleton] } + { uniq := by simp only [Nat.zero_eq, eq_iff_true_of_subsingleton]; trivial } exact ⟨ι, fι, p, hp, e, ⟨f.trans AddEquiv.uniqueProd⟩⟩ · haveI := @Fintype.prodLeft _ _ _ (Fintype.ofEquiv G f.toEquiv) _ exact diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 17a3333557acb..743d222b58bb4 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -101,7 +101,7 @@ theorem Quotient.smul_mk [QuotientAction β H] (b : β) (a : α) : @[to_additive (attr := simp)] theorem Quotient.smul_coe [QuotientAction β H] (b : β) (a : α) : - b • (a : α ⧸ H) = (b • a : α ⧸ H) := + b • (a : α ⧸ H) = (↑(b • a) : α ⧸ H) := rfl #align mul_action.quotient.smul_coe MulAction.Quotient.smul_coe #align add_action.quotient.vadd_coe AddAction.Quotient.vadd_coe diff --git a/Mathlib/GroupTheory/GroupAction/SubMulAction/Pointwise.lean b/Mathlib/GroupTheory/GroupAction/SubMulAction/Pointwise.lean index 370a8390d7d8d..4b9a8d2134463 100644 --- a/Mathlib/GroupTheory/GroupAction/SubMulAction/Pointwise.lean +++ b/Mathlib/GroupTheory/GroupAction/SubMulAction/Pointwise.lean @@ -115,14 +115,14 @@ instance : Monoid (SubMulAction R M) := { SubMulAction.semiGroup, SubMulAction.mulOneClass with } -theorem coe_pow (p : SubMulAction R M) : ∀ {n : ℕ} (_ : n ≠ 0), (p ^ n : Set M) = ((p : Set M) ^ n) +theorem coe_pow (p : SubMulAction R M) : ∀ {n : ℕ} (_ : n ≠ 0), ↑(p ^ n) = ((p : Set M) ^ n) | 0, hn => (hn rfl).elim | 1, _ => by rw [pow_one, pow_one] | n + 2, _ => by rw [pow_succ _ (n + 1), pow_succ _ (n + 1), coe_mul, coe_pow _ n.succ_ne_zero] #align sub_mul_action.coe_pow SubMulAction.coe_pow -theorem subset_coe_pow (p : SubMulAction R M) : ∀ {n : ℕ}, ((p : Set M) ^ n) ⊆ (p ^ n : Set M) +theorem subset_coe_pow (p : SubMulAction R M) : ∀ {n : ℕ}, ((p : Set M) ^ n) ⊆ ↑(p ^ n) | 0 => by rw [pow_zero, pow_zero] exact subset_coe_one diff --git a/Mathlib/GroupTheory/HNNExtension.lean b/Mathlib/GroupTheory/HNNExtension.lean index 58ac00238b7eb..0350fceacee4d 100644 --- a/Mathlib/GroupTheory/HNNExtension.lean +++ b/Mathlib/GroupTheory/HNNExtension.lean @@ -400,7 +400,7 @@ theorem not_cancels_of_cons_hyp (u : ℤˣ) (w : NormalWord d) theorem unitsSMul_cancels_iff (u : ℤˣ) (w : NormalWord d) : Cancels (-u) (unitsSMul φ u w) ↔ ¬ Cancels u w := by by_cases h : Cancels u w - · simp only [unitsSMul, dif_pos trivial, h, iff_false] + · simp only [unitsSMul, h, dite_true, not_true_eq_false, iff_false] induction w using consRecOn with | ofGroup => simp [Cancels, unitsSMulWithCancel] | cons g u' w h1 h2 _ => diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index d4e9634237a35..ef1854de09246 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -1175,7 +1175,7 @@ theorem LinearOrderedRing.orderOf_le_two : orderOf x ≤ 2 := by cases' ne_or_eq |x| 1 with h h · simp [orderOf_abs_ne_one h] rcases eq_or_eq_neg_of_abs_eq h with (rfl | rfl) - · simp + · simp; decide apply orderOf_le_of_pow_eq_one <;> norm_num #align linear_ordered_ring.order_of_le_two LinearOrderedRing.orderOf_le_two diff --git a/Mathlib/GroupTheory/Perm/Fin.lean b/Mathlib/GroupTheory/Perm/Fin.lean index 7a73e7bfe41be..4a47675d6ae71 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 + · simp; rfl · rw [finRotate_succ_eq_decomposeFin] simp [ih, pow_succ] #align sign_fin_rotate sign_finRotate diff --git a/Mathlib/GroupTheory/PushoutI.lean b/Mathlib/GroupTheory/PushoutI.lean index 83f730156d860..c138fe2640d8b 100644 --- a/Mathlib/GroupTheory/PushoutI.lean +++ b/Mathlib/GroupTheory/PushoutI.lean @@ -669,7 +669,7 @@ theorem inf_of_range_eq_base_range (hφ : ∀ i, Injective (φ i)) {i j : ι} (h (by intro x ⟨⟨g₁, hg₁⟩, ⟨g₂, hg₂⟩⟩ by_contra hx - have hx1 : x ≠ 1 := by rintro rfl; simp_all only [map_one, one_mem] + have hx1 : x ≠ 1 := by rintro rfl; simp_all only [ne_eq, one_mem, not_true_eq_false] have hg₁1 : g₁ ≠ 1 := ne_of_apply_ne (of (φ := φ) i) (by simp_all) have hg₂1 : g₂ ≠ 1 := diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 4c8d57c2689ea..575b8375e1d63 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -262,6 +262,7 @@ theorem normalClosure_swap_mul_swap_five : ⟨finRotate 5, finRotate_bit1_mem_alternatingGroup (n := 2)⟩ := by rw [Subtype.ext_iff] simp only [Fin.val_mk, Subgroup.coe_mul, Subgroup.coe_inv, Fin.val_mk] + decide rw [eq_top_iff, ← normalClosure_finRotate_five] refine' normalClosure_le_normal _ rw [Set.singleton_subset_iff, SetLike.mem_coe, ← h5] @@ -280,7 +281,7 @@ theorem isConj_swap_mul_swap_of_cycleType_two {g : Perm (Fin 5)} (ha : g ∈ alt rw [← Multiset.eq_replicate_card] at h2 rw [← sum_cycleType, h2, Multiset.sum_replicate, smul_eq_mul] at h have h : Multiset.card g.cycleType ≤ 3 := - le_of_mul_le_mul_right (le_trans h (by simp only [card_fin])) (by simp) + le_of_mul_le_mul_right (le_trans h (by simp only [card_fin]; ring_nf; decide)) (by simp) rw [mem_alternatingGroup, sign_of_cycleType, h2] at ha norm_num at ha rw [pow_add, pow_mul, Int.units_pow_two, one_mul, Units.ext_iff, Units.val_one, @@ -299,7 +300,7 @@ theorem isConj_swap_mul_swap_of_cycleType_two {g : Perm (Fin 5)} (ha : g ∈ alt · rw [disjoint_iff_disjoint_support, support_swap h04, support_swap h13] decide · contrapose! ha - simp [h_1] + decide #align alternating_group.is_conj_swap_mul_swap_of_cycle_type_two alternatingGroup.isConj_swap_mul_swap_of_cycleType_two /-- Shows that $A_5$ is simple by taking an arbitrary non-identity element and showing by casework diff --git a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean index 962092e3c2155..df50ccc519fcc 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Dihedral.lean @@ -118,7 +118,7 @@ instance : Infinite (DihedralGroup 0) := DihedralGroup.fintypeHelper.infinite_iff.mp inferInstance instance : Nontrivial (DihedralGroup n) := - ⟨⟨r 0, sr 0, by simp_rw [ne_eq]⟩⟩ + ⟨⟨r 0, sr 0, by simp_rw [ne_eq, not_false_eq_true]⟩⟩ /-- If `0 < n`, then `DihedralGroup n` has `2n` elements. -/ @@ -161,7 +161,7 @@ theorem orderOf_sr (i : ZMod n) : orderOf (sr i) = 2 := by · rw [sq, sr_mul_self] · -- Porting note: Previous proof was `decide` revert n - simp_rw [one_def, ne_eq, forall_const] + simp_rw [one_def, ne_eq, forall_const, not_false_eq_true] #align dihedral_group.order_of_sr DihedralGroup.orderOf_sr /-- If `0 < n`, then `r 1` has order `n`. diff --git a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean index 99b8715b61d80..9202e43f248a4 100644 --- a/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean +++ b/Mathlib/GroupTheory/SpecificGroups/KleinFour.lean @@ -61,7 +61,7 @@ theorem card : Fintype.card KleinFour = 4 := /-- Klein four-group is a group of order 4. -/ theorem nat_card : Nat.card KleinFour = 4 := by - simp only [Nat.card_eq_fintype_card] + simp (config := {decide := true}) only [Nat.card_eq_fintype_card] @[simp] theorem a_sq : a ^ 2 = 1 := rfl @@ -107,6 +107,6 @@ def mulEquivDihedralGroupTwo : KleinFour ≃* DihedralGroup 2 where | sr 1 * r 1 => (1, 1) left_inv := by decide right_inv := by decide - map_mul' := by simp + map_mul' := by simp (config := {decide := true}) end KleinFour diff --git a/Mathlib/GroupTheory/Submonoid/Operations.lean b/Mathlib/GroupTheory/Submonoid/Operations.lean index 05a7a3b1bf32a..2b466fc491a2d 100644 --- a/Mathlib/GroupTheory/Submonoid/Operations.lean +++ b/Mathlib/GroupTheory/Submonoid/Operations.lean @@ -560,7 +560,7 @@ attribute [to_additive existing nSMul] nPow @[to_additive (attr := simp, norm_cast)] theorem coe_pow {M} [Monoid M] {A : Type*} [SetLike A M] [SubmonoidClass A M] {S : A} (x : S) - (n : ℕ) : (x ^ n : M) = (x : M) ^ n := + (n : ℕ) : ↑(x ^ n) = (x : M) ^ n := rfl #align submonoid_class.coe_pow SubmonoidClass.coe_pow #align add_submonoid_class.coe_nsmul AddSubmonoidClass.coe_nsmul diff --git a/Mathlib/Init/Data/Bool/Lemmas.lean b/Mathlib/Init/Data/Bool/Lemmas.lean index bc1078774c302..ced7ce8edc03e 100644 --- a/Mathlib/Init/Data/Bool/Lemmas.lean +++ b/Mathlib/Init/Data/Bool/Lemmas.lean @@ -132,7 +132,7 @@ theorem bool_eq_false {b : Bool} : ¬b → b = false := bool_iff_false.1 #align bool_eq_false Bool.bool_eq_false -theorem decide_false_iff (p : Prop) [Decidable p] : decide p = false ↔ ¬p := +theorem decide_false_iff (p : Prop) {_ : Decidable p} : decide p = false ↔ ¬p := bool_iff_false.symm.trans (not_congr (decide_iff _)) #align to_bool_ff_iff Bool.decide_false_iff diff --git a/Mathlib/Init/Data/Nat/Lemmas.lean b/Mathlib/Init/Data/Nat/Lemmas.lean index 3c51e21e550e5..eef918bc45fa2 100644 --- a/Mathlib/Init/Data/Nat/Lemmas.lean +++ b/Mathlib/Init/Data/Nat/Lemmas.lean @@ -788,7 +788,7 @@ lemma to_digits_core_length (b : Nat) (h : 2 <= b) (f n e : Nat) rw [‹e = 0›] have _ : b ^ 1 = b := by simp only [pow_succ, pow_zero, Nat.one_mul] have _ : n < b := ‹b ^ 1 = b› ▸ (‹e = 0› ▸ hlt : n < b ^ Nat.succ 0) - simp only [(@Nat.div_eq_of_lt n b ‹n < b› : n / b = 0), if_true, List.length] + simp [(@Nat.div_eq_of_lt n b ‹n < b› : n / b = 0)] /-- The core implementation of `Nat.repr` returns a String with length less than or equal to the number of digits in the decimal number (represented by `e`). For example, the decimal string diff --git a/Mathlib/Init/Function.lean b/Mathlib/Init/Function.lean index b6a83ee09f6e0..6ca9bd07a87e1 100644 --- a/Mathlib/Init/Function.lean +++ b/Mathlib/Init/Function.lean @@ -24,8 +24,6 @@ variable {α : Sort u₁} {β : Sort u₂} {φ : Sort u₃} {δ : Sort u₄} {ζ #align function.comp Function.comp -lemma comp_def {α β δ : Sort _} (f : β → δ) (g : α → β) : f ∘ g = fun x ↦ f (g x) := rfl - attribute [eqns comp_def] comp lemma flip_def {f : α → β → φ} : flip f = fun b a => f a b := rfl diff --git a/Mathlib/Init/Order/Defs.lean b/Mathlib/Init/Order/Defs.lean index 71354059a3a5b..b823526c32e79 100644 --- a/Mathlib/Init/Order/Defs.lean +++ b/Mathlib/Init/Order/Defs.lean @@ -262,7 +262,7 @@ implicit arguments, requires us to unfold the defs and split the `if`s in the de macro "compareOfLessAndEq_rfl" : tactic => `(tactic| (intros a b; first | rfl | (simp only [compare, compareOfLessAndEq]; split_ifs <;> rfl) | - (induction a <;> induction b <;> simp only []))) + (induction a <;> induction b <;> simp (config := {decide := true}) only []))) /-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ @@ -436,16 +436,16 @@ theorem compare_eq_iff_eq {a b : α} : (compare a b = .eq) ↔ a = b := by case _ _ h => exact false_iff_iff.2 h theorem compare_le_iff_le {a b : α} : (compare a b ≠ .gt) ↔ a ≤ b := by - cases h : compare a b <;> simp only [] - · exact true_iff_iff.2 <| le_of_lt <| compare_lt_iff_lt.1 h - · exact true_iff_iff.2 <| le_of_eq <| compare_eq_iff_eq.1 h - · exact false_iff_iff.2 <| not_le_of_gt <| compare_gt_iff_gt.1 h + cases h : compare a b <;> simp + · exact le_of_lt <| compare_lt_iff_lt.1 h + · exact le_of_eq <| compare_eq_iff_eq.1 h + · exact compare_gt_iff_gt.1 h theorem compare_ge_iff_ge {a b : α} : (compare a b ≠ .lt) ↔ a ≥ b := by - cases h : compare a b <;> simp only [] - · exact false_iff_iff.2 <| (lt_iff_not_ge a b).1 <| compare_lt_iff_lt.1 h - · exact true_iff_iff.2 <| le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h - · exact true_iff_iff.2 <| le_of_lt <| compare_gt_iff_gt.1 h + cases h : compare a b <;> simp + · exact compare_lt_iff_lt.1 h + · exact le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h + · exact le_of_lt <| compare_gt_iff_gt.1 h theorem compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.toRel a b := by cases o <;> simp only [Ordering.toRel] diff --git a/Mathlib/Init/Quot.lean b/Mathlib/Init/Quot.lean index e854882518837..b417b5d29eecb 100644 --- a/Mathlib/Init/Quot.lean +++ b/Mathlib/Init/Quot.lean @@ -39,3 +39,5 @@ protected abbrev Quot.recOnSubsingleton' apply f apply Subsingleton.elim #align quot.rec_on_subsingleton Quot.recOnSubsingleton' + +theorem Quotient.mk'_eq_mk [s : Setoid α] : Quotient.mk' = Quotient.mk s := rfl diff --git a/Mathlib/Lean/Linter.lean b/Mathlib/Lean/Linter.lean deleted file mode 100644 index 848b4d646d051..0000000000000 --- a/Mathlib/Lean/Linter.lean +++ /dev/null @@ -1,24 +0,0 @@ -/- -Copyright (c) 2023 Floris van Doorn. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Floris van Doorn --/ -import Lean - -/-! -# Additional declarations related to linters --/ - -set_option autoImplicit true - -open Lean Elab - -namespace Lean.Linter - -/-- If `linterOption` is true, print a linter warning message at the position determined by `stx`. --/ -def logLintIf [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] - (linterOption : Lean.Option Bool) (stx : Syntax) (msg : MessageData) : m Unit := do - if linterOption.get (← getOptions) then logLint linterOption stx msg - -end Lean.Linter diff --git a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean index 14819d2553db9..381e4f20e2fe0 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/FiniteDimensional.lean @@ -479,7 +479,8 @@ theorem affineIndependent_iff_not_collinear_of_ne {p : Fin 3 → P} {i₁ i₂ i AffineIndependent k p ↔ ¬Collinear k ({p i₁, p i₂, p i₃} : Set P) := by have hu : (Finset.univ : Finset (Fin 3)) = {i₁, i₂, i₃} := by -- Porting note: Originally `by decide!` - fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ <;> simp only at h₁₂ h₁₃ h₂₃ ⊢ + fin_cases i₁ <;> fin_cases i₂ <;> fin_cases i₃ + <;> simp (config := {decide := true}) only at h₁₂ h₁₃ h₂₃ ⊢ rw [affineIndependent_iff_not_collinear, ← Set.image_univ, ← Finset.coe_univ, hu, Finset.coe_insert, Finset.coe_insert, Finset.coe_singleton, Set.image_insert_eq, Set.image_pair] #align affine_independent_iff_not_collinear_of_ne affineIndependent_iff_not_collinear_of_ne diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 3b32e603f4928..f162ad0dbda54 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -688,7 +688,7 @@ theorem affineIndependent_of_ne_of_mem_of_mem_of_not_mem {s : AffineSubspace k P refine' hp₃ ((AffineSubspace.le_def' _ s).1 _ p₃ h) simp_rw [affineSpan_le, Set.image_subset_iff, Set.subset_def, Set.mem_preimage] intro x - fin_cases x <;> simp [hp₁, hp₂] + fin_cases x <;> simp (config := {decide := true}) [hp₁, hp₂] #align affine_independent_of_ne_of_mem_of_mem_of_not_mem affineIndependent_of_ne_of_mem_of_mem_of_not_mem /-- If distinct points `p₁` and `p₃` lie in `s` but `p₂` does not, the three points are affinely diff --git a/Mathlib/LinearAlgebra/Alternating/Basic.lean b/Mathlib/LinearAlgebra/Alternating/Basic.lean index 279e2c85141b9..b4315413757df 100644 --- a/Mathlib/LinearAlgebra/Alternating/Basic.lean +++ b/Mathlib/LinearAlgebra/Alternating/Basic.lean @@ -255,8 +255,7 @@ theorem smul_apply (c : S) (m : ι → M) : (c • f) m = c • f m := #align alternating_map.smul_apply AlternatingMap.smul_apply @[norm_cast] -theorem coe_smul (c : S) : (c • f : MultilinearMap R (fun _ : ι => M) N) = - c • (f : MultilinearMap R (fun _ : ι => M) N) := +theorem coe_smul (c : S) : ↑(c • f) = c • (f : MultilinearMap R (fun _ : ι => M) N) := rfl #align alternating_map.coe_smul AlternatingMap.coe_smul diff --git a/Mathlib/LinearAlgebra/Basic.lean b/Mathlib/LinearAlgebra/Basic.lean index 98b7f17b864cf..f787fbed5d765 100644 --- a/Mathlib/LinearAlgebra/Basic.lean +++ b/Mathlib/LinearAlgebra/Basic.lean @@ -55,8 +55,6 @@ linear algebra, vector space, module -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Function open BigOperators Pointwise diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index 7bec1909a15bc..017440dbe7239 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -322,7 +322,7 @@ theorem det_conj {N : Type*} [AddCommGroup N] [Module A N] (f : M →ₗ[A] M) ( contrapose! H rcases H with ⟨s, ⟨b⟩⟩ exact ⟨_, ⟨(b.map e.symm).reindexFinsetRange⟩⟩ - simp only [coe_det, H, H', MonoidHom.one_apply, dif_neg] + simp only [coe_det, H, H', MonoidHom.one_apply, dif_neg, not_false_eq_true] #align linear_map.det_conj LinearMap.det_conj /-- If a linear map is invertible, so is its determinant. -/ diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean index 76204f904d8c0..87526cfc31a4d 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean @@ -37,7 +37,7 @@ protected def GradedAlgebra.ι : theorem GradedAlgebra.ι_apply (m : M) : GradedAlgebra.ι R M m = - DirectSum.of (fun i : ℕ => LinearMap.range (ι R : M →ₗ[R] ExteriorAlgebra R M) ^ i) 1 + DirectSum.of (fun i : ℕ => ↥(LinearMap.range (ι R : M →ₗ[R] ExteriorAlgebra R M) ^ i)) 1 ⟨ι R m, by simpa only [pow_one] using LinearMap.mem_range_self _ m⟩ := rfl #align exterior_algebra.graded_algebra.ι_apply ExteriorAlgebra.GradedAlgebra.ι_apply diff --git a/Mathlib/LinearAlgebra/Finrank.lean b/Mathlib/LinearAlgebra/Finrank.lean index fd170deeb6b8e..85060e948a8aa 100644 --- a/Mathlib/LinearAlgebra/Finrank.lean +++ b/Mathlib/LinearAlgebra/Finrank.lean @@ -436,7 +436,7 @@ theorem linearIndependent_of_top_le_span_of_card_eq_finrank {ι : Type*} [Fintyp · refine' neg_mem (smul_mem _ _ (sum_mem fun k hk => _)) obtain ⟨k_ne_i, _⟩ := Finset.mem_erase.mp hk refine' smul_mem _ _ (subset_span ⟨k, _, rfl⟩) - simp_all only [Set.mem_univ, Set.mem_diff, Set.mem_singleton_iff] + simp_all only [Set.mem_univ, Set.mem_diff, Set.mem_singleton_iff, and_self, not_false_eq_true] -- To show `b i` is a weighted sum of the other `b j`s, we'll rewrite this sum -- to have the form of the assumption `dependent`. apply eq_neg_of_add_eq_zero_left diff --git a/Mathlib/LinearAlgebra/LinearPMap.lean b/Mathlib/LinearAlgebra/LinearPMap.lean index a6ed09becfc8d..41e97ff7f2722 100644 --- a/Mathlib/LinearAlgebra/LinearPMap.lean +++ b/Mathlib/LinearAlgebra/LinearPMap.lean @@ -791,7 +791,7 @@ theorem graph_map_fst_eq_domain (f : E →ₗ.[R] F) : · rcases h with ⟨x, hx, _⟩ exact hx · use f ⟨x, h⟩ - simp only [h, exists_prop] + simp only [h, exists_const] theorem graph_map_snd_eq_range (f : E →ₗ.[R] F) : f.graph.map (LinearMap.snd R E F) = LinearMap.range f.toFun := by ext; simp diff --git a/Mathlib/LinearAlgebra/Matrix/Polynomial.lean b/Mathlib/LinearAlgebra/Matrix/Polynomial.lean index 6f9668c408ca2..94750ff201f43 100644 --- a/Mathlib/LinearAlgebra/Matrix/Polynomial.lean +++ b/Mathlib/LinearAlgebra/Matrix/Polynomial.lean @@ -37,7 +37,7 @@ open Polynomial Matrix Equiv.Perm namespace Polynomial theorem natDegree_det_X_add_C_le (A B : Matrix n n α) : - natDegree (det ((X : α[X]) • A.map C + B.map C)) ≤ Fintype.card n := by + natDegree (det ((X : α[X]) • A.map C + B.map C : Matrix n n α[X])) ≤ Fintype.card n := by rw [det_apply] refine' (natDegree_sum_le _ _).trans _ refine' Multiset.max_nat_le_of_forall_le _ _ _ @@ -45,13 +45,14 @@ theorem natDegree_det_X_add_C_le (A B : Matrix n n α) : Multiset.mem_map, exists_imp, Finset.mem_univ_val] intro g calc - natDegree (sign g • ∏ i : n, (X • A.map C + B.map C) (g i) i) ≤ - natDegree (∏ i : n, (X • A.map C + B.map C) (g i) i) := by + natDegree (sign g • ∏ i : n, (X • A.map C + B.map C : Matrix n n α[X]) (g i) i) ≤ + natDegree (∏ i : n, (X • A.map C + B.map C : Matrix n n α[X]) (g i) i) := by cases' Int.units_eq_one_or (sign g) with sg sg · rw [sg, one_smul] · rw [sg, Units.neg_smul, one_smul, natDegree_neg] - _ ≤ ∑ i : n, natDegree (((X : α[X]) • A.map C + B.map C) (g i) i) := - (natDegree_prod_le (Finset.univ : Finset n) fun i : n => (X • A.map C + B.map C) (g i) i) + _ ≤ ∑ i : n, natDegree (((X : α[X]) • A.map C + B.map C : Matrix n n α[X]) (g i) i) := + (natDegree_prod_le (Finset.univ : Finset n) fun i : n => + (X • A.map C + B.map C : Matrix n n α[X]) (g i) i) _ ≤ Finset.univ.card • 1 := (Finset.sum_le_card_nsmul _ _ 1 fun (i : n) _ => ?_) _ ≤ Fintype.card n := by simp [mul_one, Algebra.id.smul_eq_mul, Finset.card_univ] dsimp only [add_apply, smul_apply, map_apply, smul_eq_mul] diff --git a/Mathlib/LinearAlgebra/Matrix/Transvection.lean b/Mathlib/LinearAlgebra/Matrix/Transvection.lean index 48ba5b08e9e15..21d209adf9735 100644 --- a/Mathlib/LinearAlgebra/Matrix/Transvection.lean +++ b/Mathlib/LinearAlgebra/Matrix/Transvection.lean @@ -484,7 +484,8 @@ theorem mul_listTransvecRow_last_row (hM : M (inr unit) (inr unit) ≠ 0) (i : F rintro rfl cases i tauto - simp only [IH hnr.le, Ne.def, mul_transvection_apply_of_ne, Ne.symm h, inl.injEq] + simp only [IH hnr.le, Ne.def, mul_transvection_apply_of_ne, Ne.symm h, inl.injEq, + not_false_eq_true] rcases le_or_lt (n + 1) i with (hi | hi) · simp [hi, n.le_succ.trans hi, if_true] · rw [if_neg, if_neg] diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 82cde9bfdce87..6a74425cb4d4c 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -281,7 +281,8 @@ theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b := not_not.symm.trans h.not /-! ### Declarations about `Xor'` -/ -@[simp] theorem xor_true : Xor' True = Not := by simp [Xor'] +@[simp] theorem xor_true : Xor' True = Not := by + simp (config := { unfoldPartialApp := true }) [Xor'] #align xor_true xor_true @[simp] theorem xor_false : Xor' False = id := by ext; simp [Xor'] diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index 991a0baab95d0..9b92aa59f028c 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -192,7 +192,7 @@ def setValue {α β} (f : α ↪ β) (a : α) (b : β) [∀ a', Decidable (a' = -- split_ifs at h <;> (try subst b) <;> (try simp only [f.injective.eq_iff] at *) <;> cc split_ifs at h with h₁ h₂ _ _ h₅ h₆ <;> (try subst b) <;> - (try simp only [f.injective.eq_iff] at *) + (try simp only [f.injective.eq_iff, not_true_eq_false] at *) · rw[h₁,h₂] · rw[h₁,h] · rw[h₅,←h] diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index 9c5acf1782b6d..34c9e367e7d14 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -1632,10 +1632,17 @@ theorem measure_eq_measure_preimage_add_measure_tsum_Ico_zpow [MeasurableSpace · exact hs.inter (hf measurableSet_Ioi) have B : μ (s ∩ f ⁻¹' Ioi 0) = μ (s ∩ f ⁻¹' {∞}) + μ (s ∩ f ⁻¹' Ioo 0 ∞) := by rw [← measure_union] - · rw [← inter_union_distrib_left, ← preimage_union, - ← Ioo_union_Ici_eq_Ioi ENNReal.zero_ne_top.lt_top, Ici_top, union_comm] - · refine disjoint_singleton_left.mpr right_mem_Ioo.mp - |>.preimage f |>.inter_right' s |>.inter_left' s + · rw [← inter_union_distrib_left] + congr + ext x + simp only [mem_singleton_iff, mem_union, mem_Ioo, mem_Ioi, mem_preimage] + have H : f x = ∞ ∨ f x < ∞ := eq_or_lt_of_le le_top + cases' H with H H + · simp only [H, eq_self_iff_true, or_false_iff, zero_lt_top, not_top_lt, and_false_iff] + · simp only [H, H.ne, and_true_iff, false_or_iff] + · refine disjoint_left.2 fun x hx h'x => ?_ + have : f x < ∞ := h'x.2.2 + exact lt_irrefl _ (this.trans_le (le_of_eq hx.2.symm)) · exact hs.inter (hf measurableSet_Ioo) have C : μ (s ∩ f ⁻¹' Ioo 0 ∞) = ∑' n : ℤ, μ (s ∩ f ⁻¹' Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1))) := by diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 7eb0f36a514c2..9d146825158a1 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -274,8 +274,7 @@ def pi' : Measure (∀ i, α i) := theorem pi'_pi [∀ i, SigmaFinite (μ i)] (s : ∀ i, Set (α i)) : pi' μ (pi univ s) = ∏ i, μ i (s i) := by rw [pi'] - simp only [TProd.elim'] -- Porting note: new step - erw [← MeasurableEquiv.piMeasurableEquivTProd_symm_apply, MeasurableEquiv.map_apply, + rw [← MeasurableEquiv.piMeasurableEquivTProd_symm_apply, MeasurableEquiv.map_apply, MeasurableEquiv.piMeasurableEquivTProd_symm_apply, elim_preimage_pi, tprod_tprod _ μ, ← List.prod_toFinset, sortedUniv_toFinset] <;> exact sortedUniv_nodup ι diff --git a/Mathlib/MeasureTheory/Covering/Besicovitch.lean b/Mathlib/MeasureTheory/Covering/Besicovitch.lean index 4044536644405..31b309490ea9e 100644 --- a/Mathlib/MeasureTheory/Covering/Besicovitch.lean +++ b/Mathlib/MeasureTheory/Covering/Besicovitch.lean @@ -928,7 +928,7 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SigmaFinit obtain ⟨u, su, u_open, μu⟩ : ∃ U, U ⊇ s ∧ IsOpen U ∧ μ U ≤ μ s + ε / 2 := Set.exists_isOpen_le_add _ _ (by - simpa only [or_false_iff, Ne.def, ENNReal.div_eq_zero_iff, ENNReal.one_ne_top] using hε) + simpa only [or_false, Ne.def, ENNReal.div_eq_zero_iff, ENNReal.two_ne_top] using hε) have : ∀ x ∈ s, ∃ R > 0, ball x R ⊆ u := fun x hx => Metric.mem_nhds_iff.1 (u_open.mem_nhds (su hx)) choose! R hR using this @@ -946,8 +946,8 @@ theorem exists_closedBall_covering_tsum_measure_le (μ : Measure α) [SigmaFinit obtain ⟨v, s'v, v_open, μv⟩ : ∃ v, v ⊇ s' ∧ IsOpen v ∧ μ v ≤ μ s' + ε / 2 / N := Set.exists_isOpen_le_add _ _ (by - simp only [hε, ENNReal.nat_ne_top, WithTop.mul_eq_top_iff, Ne.def, ENNReal.div_eq_zero_iff, - ENNReal.one_ne_top, not_false_iff, and_false_iff, false_and_iff, or_self_iff]) + simp only [ne_eq, ENNReal.div_eq_zero_iff, hε, ENNReal.two_ne_top, or_self, + ENNReal.nat_ne_top, not_false_eq_true]) have : ∀ x ∈ s', ∃ r1 ∈ f x ∩ Ioo (0 : ℝ) 1, closedBall x r1 ⊆ v := by intro x hx rcases Metric.mem_nhds_iff.1 (v_open.mem_nhds (s'v hx)) with ⟨r, rpos, hr⟩ diff --git a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean index e6498ba76c448..d9a02f2a4eeee 100644 --- a/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean +++ b/Mathlib/MeasureTheory/Covering/BesicovitchVectorSpace.lean @@ -44,8 +44,6 @@ In particular, this number is bounded by `5 ^ dim` by a straightforward measure universe u -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Metric Set FiniteDimensional MeasureTheory Filter Fin open scoped ENNReal Topology diff --git a/Mathlib/MeasureTheory/Covering/Differentiation.lean b/Mathlib/MeasureTheory/Covering/Differentiation.lean index a5d1a9dda468c..612ae9c1e71c6 100644 --- a/Mathlib/MeasureTheory/Covering/Differentiation.lean +++ b/Mathlib/MeasureTheory/Covering/Differentiation.lean @@ -76,9 +76,6 @@ make no sense. However, the measure is not globally zero if the space is big eno * [Herbert Federer, Geometric Measure Theory, Chapter 2.9][Federer1996] -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open MeasureTheory Metric Set Filter TopologicalSpace MeasureTheory.Measure open scoped Filter ENNReal MeasureTheory NNReal Topology @@ -552,13 +549,16 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht let ν := μ.withDensity (v.limRatioMeas hρ) let f := v.limRatioMeas hρ have f_meas : Measurable f := v.limRatioMeas_measurable hρ - have A : ν (s ∩ f ⁻¹' {0}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ) (s ∩ f ⁻¹' {0}) := by + -- Note(kmill): smul elaborator when used for CoeFun fails to get CoeFun instance to trigger + -- unless you use the `(... :)` notation. Another fix is using `(2 : Nat)`, so this appears + -- to be an unpleasant interaction with default instances. + have A : ν (s ∩ f ⁻¹' {0}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {0}) := by apply le_trans _ (zero_le _) have M : MeasurableSet (s ∩ f ⁻¹' {0}) := hs.inter (f_meas (measurableSet_singleton _)) simp only [nonpos_iff_eq_zero, M, withDensity_apply, lintegral_eq_zero_iff f_meas] apply (ae_restrict_iff' M).2 exact eventually_of_forall fun x hx => hx.2 - have B : ν (s ∩ f ⁻¹' {∞}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ) (s ∩ f ⁻¹' {∞}) := by + have B : ν (s ∩ f ⁻¹' {∞}) ≤ ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {∞}) := by apply le_trans (le_of_eq _) (zero_le _) apply withDensity_absolutelyContinuous μ _ rw [← nonpos_iff_eq_zero] @@ -566,13 +566,13 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht have C : ∀ n : ℤ, ν (s ∩ f ⁻¹' Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1))) ≤ - ((t : ℝ≥0∞) ^ 2 • ρ) (s ∩ f ⁻¹' Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1))) := by + ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1))) := by intro n let I := Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1)) have M : MeasurableSet (s ∩ f ⁻¹' I) := hs.inter (f_meas measurableSet_Ico) simp only [M, withDensity_apply, coe_nnreal_smul_apply] calc - (∫⁻ x in s ∩ f ⁻¹' I, f x ∂μ) ≤ ∫⁻ x in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := + (∫⁻ x in s ∩ f ⁻¹' I, f x ∂μ) ≤ ∫⁻ _ in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := lintegral_mono_ae ((ae_restrict_iff' M).2 (eventually_of_forall fun x hx => hx.2.2.le)) _ = (t : ℝ≥0∞) ^ (n + 1) * μ (s ∩ f ⁻¹' I) := by simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] @@ -598,10 +598,10 @@ theorem withDensity_le_mul {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht ∑' n : ℤ, ν (s ∩ f ⁻¹' Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1))) := measure_eq_measure_preimage_add_measure_tsum_Ico_zpow ν f_meas hs ht _ ≤ - ((t : ℝ≥0∞) ^ 2 • ρ) (s ∩ f ⁻¹' {0}) + ((t : ℝ≥0∞) ^ 2 • ρ) (s ∩ f ⁻¹' {∞}) + - ∑' n : ℤ, ((t : ℝ≥0∞) ^ 2 • ρ) (s ∩ f ⁻¹' Ico ((t : ℝ≥0∞) ^ n) ((t : ℝ≥0∞) ^ (n + 1))) := + ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {0}) + ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' {∞}) + + ∑' n : ℤ, ((t : ℝ≥0∞) ^ 2 • ρ :) (s ∩ f ⁻¹' Ico (t ^ n) (t ^ (n + 1))) := (add_le_add (add_le_add A B) (ENNReal.tsum_le_tsum C)) - _ = ((t : ℝ≥0∞) ^ 2 • ρ) s := + _ = ((t : ℝ≥0∞) ^ 2 • ρ :) s := (measure_eq_measure_preimage_add_measure_tsum_Ico_zpow ((t : ℝ≥0∞) ^ 2 • ρ) f_meas hs ht).symm #align vitali_family.with_density_le_mul VitaliFamily.withDensity_le_mul @@ -642,7 +642,7 @@ theorem le_mul_withDensity {s : Set α} (hs : MeasurableSet s) {t : ℝ≥0} (ht intro x hx apply hx.2.2.trans_le (le_of_eq _) rw [ENNReal.coe_zpow t_ne_zero'] - _ = ∫⁻ x in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := by + _ = ∫⁻ _ in s ∩ f ⁻¹' I, (t : ℝ≥0∞) ^ (n + 1) ∂μ := by simp only [lintegral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter] _ ≤ ∫⁻ x in s ∩ f ⁻¹' I, t * f x ∂μ := by apply lintegral_mono_ae ((ae_restrict_iff' M).2 (eventually_of_forall fun x hx => ?_)) diff --git a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean index d024172dedf8e..8c7c11add456d 100644 --- a/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/Lebesgue.lean @@ -240,7 +240,8 @@ lemma MutuallySingular.rnDeriv_ae_eq_zero {μ ν : Measure α} (hμν : μ ⟂ exact hμν · rw [rnDeriv_of_not_haveLebesgueDecomposition h] -lemma rnDeriv_singularPart (μ ν : Measure α) : (μ.singularPart ν).rnDeriv ν =ᵐ[ν] 0 := by +lemma rnDeriv_singularPart (μ ν : Measure α) : + (μ.singularPart ν).rnDeriv ν =ᵐ[ν] 0 := by rw [rnDeriv_eq_zero] exact mutuallySingular_singularPart μ ν diff --git a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean index cf9189f17055a..03eb8fbdac493 100644 --- a/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean +++ b/Mathlib/MeasureTheory/Function/ConvergenceInMeasure.lean @@ -139,8 +139,6 @@ theorem tendstoInMeasure_of_tendsto_ae [IsFiniteMeasure μ] (hf : ∀ n, AEStron exact hxfg #align measure_theory.tendsto_in_measure_of_tendsto_ae MeasureTheory.tendstoInMeasure_of_tendsto_ae -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace ExistsSeqTendstoAe theorem exists_nat_measure_lt_two_inv (hfg : TendstoInMeasure μ f atTop g) (n : ℕ) : @@ -211,7 +209,7 @@ theorem TendstoInMeasure.exists_seq_tendsto_ae (hfg : TendstoInMeasure μ f atTo set s := Filter.atTop.limsup S with hs have hμs : μ s = 0 := by refine' measure_limsup_eq_zero (ne_of_lt <| lt_of_le_of_lt (ENNReal.tsum_le_tsum hμS_le) _) - simp only [ENNReal.tsum_geometric, ENNReal.one_sub_inv_two, inv_inv] + simp only [ENNReal.tsum_geometric, ENNReal.one_sub_inv_two, ENNReal.two_lt_top, inv_inv] have h_tendsto : ∀ x ∈ sᶜ, Tendsto (fun i => f (ns i) x) atTop (𝓝 (g x)) := by refine' fun x hx => Metric.tendsto_atTop.mpr fun ε hε => _ rw [hs, limsup_eq_iInf_iSup_of_nat] at hx diff --git a/Mathlib/MeasureTheory/Function/Jacobian.lean b/Mathlib/MeasureTheory/Function/Jacobian.lean index d2f3721665fee..0d15e8586c59c 100644 --- a/Mathlib/MeasureTheory/Function/Jacobian.lean +++ b/Mathlib/MeasureTheory/Function/Jacobian.lean @@ -88,9 +88,6 @@ Change of variables in integrals [Fremlin, *Measure Theory* (volume 2)][fremlin_vol2] -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open MeasureTheory MeasureTheory.Measure Metric Filter Set FiniteDimensional Asymptotics TopologicalSpace diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 80ed01056da59..d09b5e8d76307 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -229,7 +229,7 @@ theorem HasFiniteIntegral.smul_measure {f : α → β} (h : HasFiniteIntegral f @[simp] theorem hasFiniteIntegral_zero_measure {m : MeasurableSpace α} (f : α → β) : HasFiniteIntegral f (0 : Measure α) := by - simp only [HasFiniteIntegral, lintegral_zero_measure, WithTop.zero_lt_top] + simp only [HasFiniteIntegral, lintegral_zero_measure, zero_lt_top] #align measure_theory.has_finite_integral_zero_measure MeasureTheory.hasFiniteIntegral_zero_measure variable (α β μ) diff --git a/Mathlib/MeasureTheory/Function/L2Space.lean b/Mathlib/MeasureTheory/Function/L2Space.lean index 0e82a13ad9748..b880a7b3dba5c 100644 --- a/Mathlib/MeasureTheory/Function/L2Space.lean +++ b/Mathlib/MeasureTheory/Function/L2Space.lean @@ -24,9 +24,6 @@ is also an inner product space, with inner product defined as `inner f g = ∫ a -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - set_option linter.uppercaseLean3 false noncomputable section @@ -134,7 +131,8 @@ theorem snorm_inner_lt_top (f g : α →₂[μ] E) : snorm (fun x : α => ⟪f x _ ≤ 2 * ‖f x‖ * ‖g x‖ := (mul_le_mul_of_nonneg_right (le_mul_of_one_le_left (norm_nonneg _) one_le_two) (norm_nonneg _)) - _ ≤ ‖‖f x‖ ^ 2 + ‖g x‖ ^ 2‖ := (two_mul_le_add_sq _ _).trans (le_abs_self _) + -- TODO(kmill): the type ascription is getting around an elaboration error + _ ≤ ‖(‖f x‖ ^ 2 + ‖g x‖ ^ 2 : ℝ)‖ := (two_mul_le_add_sq _ _).trans (le_abs_self _) refine' (snorm_mono_ae (ae_of_all _ h)).trans_lt ((snorm_add_le _ _ le_rfl).trans_lt _) · exact ((Lp.aestronglyMeasurable f).norm.aemeasurable.pow_const _).aestronglyMeasurable diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm.lean b/Mathlib/MeasureTheory/Function/LpSeminorm.lean index e01f810dc21f5..387d19670469a 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm.lean @@ -327,10 +327,10 @@ theorem snorm_const_lt_top_iff {p : ℝ≥0∞} {c : F} (hp_ne_zero : p ≠ 0) ( snorm (fun _ : α => c) p μ < ∞ ↔ c = 0 ∨ μ Set.univ < ∞ := by have hp : 0 < p.toReal := ENNReal.toReal_pos hp_ne_zero hp_ne_top by_cases hμ : μ = 0 - · simp only [hμ, Measure.coe_zero, Pi.zero_apply, or_true_iff, WithTop.zero_lt_top, + · simp only [hμ, Measure.coe_zero, Pi.zero_apply, or_true_iff, zero_lt_top, snorm_measure_zero] by_cases hc : c = 0 - · simp only [hc, true_or_iff, eq_self_iff_true, WithTop.zero_lt_top, snorm_zero'] + · simp only [hc, true_or_iff, eq_self_iff_true, zero_lt_top, snorm_zero'] rw [snorm_const' c hp_ne_zero hp_ne_top] by_cases hμ_top : μ Set.univ = ∞ · simp [hc, hμ_top, hp] @@ -357,7 +357,7 @@ theorem memℒp_const (c : E) [IsFiniteMeasure μ] : Memℒp (fun _ : α => c) p theorem memℒp_top_const (c : E) : Memℒp (fun _ : α => c) ∞ μ := by refine' ⟨aestronglyMeasurable_const, _⟩ by_cases h : μ = 0 - · simp only [h, snorm_measure_zero, WithTop.zero_lt_top] + · simp only [h, snorm_measure_zero, zero_lt_top] · rw [snorm_const _ ENNReal.top_ne_zero h] simp only [ENNReal.top_toReal, _root_.div_zero, ENNReal.rpow_zero, mul_one, ENNReal.coe_lt_top] #align measure_theory.mem_ℒp_top_const MeasureTheory.memℒp_top_const diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index b9f773238415c..68cbeabbf3656 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -68,8 +68,6 @@ open TopologicalSpace MeasureTheory Filter open scoped NNReal ENNReal BigOperators Topology MeasureTheory Uniformity -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {α E F G : Type*} {m m0 : MeasurableSpace α} {p : ℝ≥0∞} {q : ℝ} {μ ν : Measure α} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup G] diff --git a/Mathlib/MeasureTheory/Function/SimpleFunc.lean b/Mathlib/MeasureTheory/Function/SimpleFunc.lean index 4a79b034f928b..020ddfb162648 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFunc.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFunc.lean @@ -161,7 +161,7 @@ theorem coe_const (b : β) : ⇑(const α b) = Function.const α b := @[simp] theorem range_const (α) [MeasurableSpace α] [Nonempty α] (b : β) : (const α b).range = {b} := - Finset.coe_injective <| by simp [Function.const] + Finset.coe_injective <| by simp (config := { unfoldPartialApp := true }) [Function.const] #align measure_theory.simple_func.range_const MeasureTheory.SimpleFunc.range_const theorem range_const_subset (α) [MeasurableSpace α] (b : β) : (const α b).range ⊆ {b} := diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean index d6354f5cdac21..cde88cb4818a3 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDense.lean @@ -168,7 +168,7 @@ theorem tendsto_approxOn {f : β → α} (hf : Measurable f) {s : Set α} {y₀ theorem edist_approxOn_mono {f : β → α} (hf : Measurable f) {s : Set α} {y₀ : α} (h₀ : y₀ ∈ s) [SeparableSpace s] (x : β) {m n : ℕ} (h : m ≤ n) : edist (approxOn f hf s y₀ h₀ n x) (f x) ≤ edist (approxOn f hf s y₀ h₀ m x) (f x) := by - dsimp only [approxOn, coe_comp, Function.comp] + dsimp only [approxOn, coe_comp, Function.comp_def] exact edist_nearestPt_le _ _ ((nearestPtInd_le _ _ _).trans h) #align measure_theory.simple_func.edist_approx_on_mono MeasureTheory.SimpleFunc.edist_approxOn_mono diff --git a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean index a6a68b4cc15bf..3b1595ea6cbaf 100644 --- a/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean +++ b/Mathlib/MeasureTheory/Function/SimpleFuncDenseLp.lean @@ -986,7 +986,7 @@ theorem Memℒp.induction_dense (hp_ne_top : p ≠ ∞) (P : (α → E) → Prop (h1P : ∀ f g, P f → P g → P (f + g)) (h2P : ∀ f, P f → AEStronglyMeasurable f μ) {f : α → E} (hf : Memℒp f p μ) {ε : ℝ≥0∞} (hε : ε ≠ 0) : ∃ g : α → E, snorm (f - g) p μ ≤ ε ∧ P g := by rcases eq_or_ne p 0 with (rfl | hp_pos) - · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, WithTop.zero_lt_top]) + · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, zero_lt_top]) hε with ⟨g, _, Pg⟩ exact ⟨g, by simp only [snorm_exponent_zero, zero_le'], Pg⟩ suffices H : @@ -1001,7 +1001,7 @@ theorem Memℒp.induction_dense (hp_ne_top : p ≠ ∞) (P : (α → E) → Prop apply SimpleFunc.induction · intro c s hs ε εpos Hs rcases eq_or_ne c 0 with (rfl | hc) - · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, WithTop.zero_lt_top]) + · rcases h0P (0 : E) MeasurableSet.empty (by simp only [measure_empty, zero_lt_top]) εpos with ⟨g, hg, Pg⟩ rw [← snorm_neg, neg_sub] at hg refine' ⟨g, _, Pg⟩ diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index c97893d13b3a3..58b8060a254cf 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -987,7 +987,7 @@ theorem finStronglyMeasurable_zero {α β} {m : MeasurableSpace α} {μ : Measur [TopologicalSpace β] : FinStronglyMeasurable (0 : α → β) μ := ⟨0, by simp only [Pi.zero_apply, SimpleFunc.coe_zero, support_zero', measure_empty, - WithTop.zero_lt_top, forall_const], + zero_lt_top, forall_const], fun _ => tendsto_const_nhds⟩ #align measure_theory.fin_strongly_measurable_zero MeasureTheory.finStronglyMeasurable_zero diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 9c824b3cb486a..dc5fe00f73994 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -950,11 +950,9 @@ theorem uniformIntegrable_average refine' le_trans _ (_ : ↑(↑n : ℝ≥0)⁻¹ * (n • C : ℝ≥0∞) ≤ C) · refine' (ENNReal.mul_le_mul_left hn ENNReal.coe_ne_top).2 _ conv_rhs => rw [← Finset.card_range n] - -- Porting note: Originally `exact Finset.sum_le_card_nsmul _ _ _ fun i hi => hC i` - convert Finset.sum_le_card_nsmul _ _ _ fun i _ => hC i - rw [ENNReal.coe_smul] + exact Finset.sum_le_card_nsmul _ _ _ fun i _ => hC i · simp only [ENNReal.coe_eq_zero, inv_eq_zero, Nat.cast_eq_zero] at hn - rw [ENNReal.coe_smul, nsmul_eq_mul, ← mul_assoc, ENNReal.coe_inv, ENNReal.coe_nat, + rw [nsmul_eq_mul, ← mul_assoc, ENNReal.coe_inv, ENNReal.coe_nat, ENNReal.inv_mul_cancel _ (ENNReal.nat_ne_top _), one_mul] all_goals simpa only [Ne.def, Nat.cast_eq_zero] diff --git a/Mathlib/MeasureTheory/Group/Measure.lean b/Mathlib/MeasureTheory/Group/Measure.lean index 4f79d9acdb66d..72fa3fb16dba6 100644 --- a/Mathlib/MeasureTheory/Group/Measure.lean +++ b/Mathlib/MeasureTheory/Group/Measure.lean @@ -548,16 +548,6 @@ section TopologicalGroup variable [TopologicalSpace G] [BorelSpace G] {μ : Measure G} [Group G] -@[to_additive] -instance Measure.IsFiniteMeasureOnCompacts.inv [ContinuousInv G] [IsFiniteMeasureOnCompacts μ] : - IsFiniteMeasureOnCompacts μ.inv := - IsFiniteMeasureOnCompacts.map μ (Homeomorph.inv G) - -@[to_additive] -instance Measure.IsOpenPosMeasure.inv [ContinuousInv G] [IsOpenPosMeasure μ] : - IsOpenPosMeasure μ.inv := - (Homeomorph.inv G).continuous.isOpenPosMeasure_map (Homeomorph.inv G).surjective - @[to_additive] instance Measure.Regular.inv [ContinuousInv G] [Regular μ] : Regular μ.inv := Regular.map (Homeomorph.inv G) @@ -571,8 +561,10 @@ instance Measure.InnerRegular.inv [ContinuousInv G] [InnerRegular μ] : InnerReg variable [TopologicalGroup G] @[to_additive] -theorem regular_inv_iff : μ.inv.Regular ↔ μ.Regular := - Regular.map_iff (Homeomorph.inv G) +theorem regular_inv_iff : μ.inv.Regular ↔ μ.Regular := by + constructor + · intro h; rw [← μ.inv_inv]; exact Measure.Regular.inv + · intro h; exact Measure.Regular.inv #align measure_theory.regular_inv_iff MeasureTheory.regular_inv_iff #align measure_theory.regular_neg_iff MeasureTheory.regular_neg_iff diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 7dfa7c53300d6..bd16215c20a8b 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1265,7 +1265,7 @@ theorem integral_nonpos {f : α → ℝ} (hf : f ≤ 0) : ∫ a, f a ∂μ ≤ 0 theorem integral_eq_zero_iff_of_nonneg_ae {f : α → ℝ} (hf : 0 ≤ᵐ[μ] f) (hfi : Integrable f μ) : ∫ x, f x ∂μ = 0 ↔ f =ᵐ[μ] 0 := by simp_rw [integral_eq_lintegral_of_nonneg_ae hf hfi.1, ENNReal.toReal_eq_zero_iff, - ← ENNReal.not_lt_top, ← hasFiniteIntegral_iff_ofReal hf, hfi.2, or_false_iff] + ← ENNReal.not_lt_top, ← hasFiniteIntegral_iff_ofReal hf, hfi.2, not_true_eq_false, or_false_iff] -- Porting note: split into parts, to make `rw` and `simp` work rw [lintegral_eq_zero_iff'] · rw [← hf.le_iff_eq, Filter.EventuallyEq, Filter.EventuallyLE] @@ -1324,17 +1324,13 @@ theorem Memℒp.snorm_eq_integral_rpow_norm {f : α → H} {p : ℝ≥0∞} (hp1 have A : ∫⁻ a : α, ENNReal.ofReal (‖f a‖ ^ p.toReal) ∂μ = ∫⁻ a : α, ‖f a‖₊ ^ p.toReal ∂μ := by apply lintegral_congr intro x - rw [← ofReal_rpow_of_nonneg (norm_nonneg _) toReal_nonneg, ofReal_norm_eq_coe_nnnorm, - -- Porting note: Here and below `ENNReal.coe_rpow_of_nonneg` was not needed - ← ENNReal.coe_rpow_of_nonneg _ toReal_nonneg] + rw [← ofReal_rpow_of_nonneg (norm_nonneg _) toReal_nonneg, ofReal_norm_eq_coe_nnnorm] simp only [snorm_eq_lintegral_rpow_nnnorm hp1 hp2, one_div] rw [integral_eq_lintegral_of_nonneg_ae]; rotate_left · exact eventually_of_forall fun x => Real.rpow_nonneg_of_nonneg (norm_nonneg _) _ · exact (hf.aestronglyMeasurable.norm.aemeasurable.pow_const _).aestronglyMeasurable rw [A, ← ofReal_rpow_of_nonneg toReal_nonneg (inv_nonneg.2 toReal_nonneg), ofReal_toReal] - · simp_rw [← ENNReal.coe_rpow_of_nonneg _ toReal_nonneg] - · simp_rw [← ENNReal.coe_rpow_of_nonneg _ toReal_nonneg] - exact (lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp1 hp2 hf.2).ne + exact (lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp1 hp2 hf.2).ne #align measure_theory.mem_ℒp.snorm_eq_integral_rpow_norm MeasureTheory.Memℒp.snorm_eq_integral_rpow_norm end NormedAddCommGroup diff --git a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean index 1b3204cd2edb2..6ca6a9d32f16c 100644 --- a/Mathlib/MeasureTheory/Integral/CircleIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/CircleIntegral.lean @@ -71,8 +71,6 @@ integral, circle, Cauchy integral variable {E : Type*} [NormedAddCommGroup E] -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open scoped Real NNReal Interval Pointwise Topology @@ -142,8 +140,8 @@ theorem circleMap_ne_mem_ball {c : ℂ} {R : ℝ} {w : ℂ} (hw : w ∈ ball c R theorem range_circleMap (c : ℂ) (R : ℝ) : range (circleMap c R) = sphere c |R| := calc range (circleMap c R) = c +ᵥ R • range fun θ : ℝ => exp (θ * I) := by - simp only [← image_vadd, ← image_smul, ← range_comp, vadd_eq_add, circleMap, (· ∘ ·), - real_smul] + simp (config := { unfoldPartialApp := true }) only [← image_vadd, ← image_smul, ← range_comp, + vadd_eq_add, circleMap, Function.comp_def, real_smul] _ = sphere c |R| := by rw [Complex.range_exp_mul_I, smul_sphere R 0 zero_le_one] simp @@ -274,7 +272,7 @@ theorem circleIntegrable_iff [NormedSpace ℂ E] {f : ℂ → E} {c : ℂ} (R : CircleIntegrable f c R ↔ IntervalIntegrable (fun θ : ℝ => deriv (circleMap c R) θ • f (circleMap c R θ)) volume 0 (2 * π) := by by_cases h₀ : R = 0 - · simp [h₀, const] + · simp (config := { unfoldPartialApp := true }) [h₀, const] refine' ⟨fun h => h.out, fun h => _⟩ simp only [CircleIntegrable, intervalIntegrable_iff, deriv_circleMap] at h ⊢ refine' (h.norm.const_mul |R|⁻¹).mono' _ _ @@ -357,7 +355,7 @@ namespace circleIntegral @[simp] theorem integral_radius_zero (f : ℂ → E) (c : ℂ) : (∮ z in C(c, 0), f z) = 0 := by - simp [circleIntegral, const] + simp (config := { unfoldPartialApp := true }) [circleIntegral, const] #align circle_integral.integral_radius_zero circleIntegral.integral_radius_zero theorem integral_congr {f g : ℂ → E} {c : ℂ} {R : ℝ} (hR : 0 ≤ R) (h : EqOn f g (sphere c R)) : diff --git a/Mathlib/MeasureTheory/Integral/CircleTransform.lean b/Mathlib/MeasureTheory/Integral/CircleTransform.lean index 4996f048a5c32..ce63978006ba4 100644 --- a/Mathlib/MeasureTheory/Integral/CircleTransform.lean +++ b/Mathlib/MeasureTheory/Integral/CircleTransform.lean @@ -25,8 +25,6 @@ open Set MeasureTheory Metric Filter Function open scoped Interval Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] (R : ℝ) (z w : ℂ) diff --git a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean index b5a9582183f7b..3fd8394c65129 100644 --- a/Mathlib/MeasureTheory/Integral/MeanInequalities.lean +++ b/Mathlib/MeasureTheory/Integral/MeanInequalities.lean @@ -58,8 +58,6 @@ open Classical BigOperators NNReal ENNReal MeasureTheory Finset set_option linter.uppercaseLean3 false -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {α : Type*} [MeasurableSpace α] {μ : Measure α} namespace ENNReal diff --git a/Mathlib/MeasureTheory/Integral/PeakFunction.lean b/Mathlib/MeasureTheory/Integral/PeakFunction.lean index 8f268aa102add..864de31622f76 100644 --- a/Mathlib/MeasureTheory/Integral/PeakFunction.lean +++ b/Mathlib/MeasureTheory/Integral/PeakFunction.lean @@ -30,9 +30,6 @@ Note that there are related results about convolution with respect to peak funct `Analysis.Convolution`, such as `convolution_tendsto_right` there. -/ - -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Set Filter MeasureTheory MeasureTheory.Measure TopologicalSpace Metric open scoped Topology ENNReal diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 58d4b17771ee9..f7e2403a64610 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1291,10 +1291,10 @@ theorem integral_smul_const {𝕜 : Type*} [IsROrC 𝕜] [NormedSpace 𝕜 E] [C by_cases hf : Integrable f μ · exact ((1 : 𝕜 →L[𝕜] 𝕜).smulRight c).integral_comp_comm hf · by_cases hc : c = 0 - · simp only [hc, integral_zero, smul_zero] + · simp [hc, integral_zero, smul_zero] rw [integral_undef hf, integral_undef, zero_smul] rw [integrable_smul_const hc] - simp_rw [hf] + simp_rw [hf, not_false_eq_true] #align integral_smul_const integral_smul_const theorem integral_withDensity_eq_integral_smul {f : α → ℝ≥0} (f_meas : Measurable f) (g : α → E) : diff --git a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean index 5e51e0229c33e..509c63200481e 100644 --- a/Mathlib/MeasureTheory/Integral/TorusIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/TorusIntegral.lean @@ -71,7 +71,6 @@ local macro:arg t:term:max noWs "ⁿ⁺¹" : term => `(Fin (n + 1) → $t) local macro:arg t:term:max noWs "ⁿ" : term => `(Fin n → $t) local macro:arg t:term:max noWs "⁰" : term => `(Fin 0 → $t) local macro:arg t:term:max noWs "¹" : term => `(Fin 1 → $t) -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 /-! ### `torusMap`, a parametrization of a torus @@ -244,8 +243,8 @@ theorem torusIntegral_succAbove {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : rw [torusIntegral, ← hem.map_eq, set_integral_map_equiv, heπ, Measure.volume_eq_prod, set_integral_prod, circleIntegral_def_Icc] · refine' set_integral_congr measurableSet_Icc fun θ _ => _ - simp only [torusIntegral, ← integral_smul, deriv_circleMap, i.prod_univ_succAbove _, smul_smul, - torusMap, circleMap_zero] + simp (config := { unfoldPartialApp := true }) only [torusIntegral, ← integral_smul, + deriv_circleMap, i.prod_univ_succAbove _, smul_smul, torusMap, circleMap_zero] refine' set_integral_congr measurableSet_Icc fun Θ _ => _ simp only [MeasurableEquiv.piFinSuccAboveEquiv_symm_apply, i.insertNth_apply_same, i.insertNth_apply_succAbove, (· ∘ ·)] diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 5735fb7d7c8df..ba5a37cd8ba25 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -928,7 +928,7 @@ theorem measurable_uniqueElim [Unique δ] [∀ i, MeasurableSpace (π i)] : theorem measurable_updateFinset [DecidableEq δ] {s : Finset δ} {x : ∀ i, π i} : Measurable (updateFinset x s) := by - simp_rw [updateFinset, measurable_pi_iff] + simp (config := { unfoldPartialApp := true }) only [updateFinset, measurable_pi_iff] intro i by_cases h : i ∈ s <;> simp [h, measurable_pi_apply] diff --git a/Mathlib/MeasureTheory/Measure/Doubling.lean b/Mathlib/MeasureTheory/Measure/Doubling.lean index 0318c1fadd650..8e059bccb4200 100644 --- a/Mathlib/MeasureTheory/Measure/Doubling.lean +++ b/Mathlib/MeasureTheory/Measure/Doubling.lean @@ -25,9 +25,6 @@ This file records basic facts about uniformly locally doubling measures. appearing in the definition of a uniformly locally doubling measure. -/ --- Porting note: for 2 ^ n in exists_eventually_forall_measure_closedBall_le_mul -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open Set Filter Metric MeasureTheory TopologicalSpace ENNReal NNReal Topology diff --git a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean index 49ee1df879122..6858344d54368 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/Basic.lean @@ -665,7 +665,7 @@ instance isHaarMeasure_haarMeasure (K₀ : PositiveCompacts G) : IsHaarMeasure ( isHaarMeasure_of_isCompact_nonempty_interior (haarMeasure K₀) K₀ K₀.isCompact K₀.interior_nonempty · simp only [haarMeasure_self]; exact one_ne_zero - · simp only [haarMeasure_self] + · simp only [haarMeasure_self, ne_eq, ENNReal.one_ne_top, not_false_eq_true] #align measure_theory.measure.is_haar_measure_haar_measure MeasureTheory.Measure.isHaarMeasure_haarMeasure #align measure_theory.measure.is_add_haar_measure_add_haar_measure MeasureTheory.Measure.isAddHaarMeasure_addHaarMeasure diff --git a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean index a4f75ce89df44..60241aca84ee1 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/InnerProductSpace.lean @@ -20,8 +20,6 @@ measure `1` to the parallelepiped spanned by any orthonormal basis, and that it the canonical `volume` from the `MeasureSpace` instance. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open FiniteDimensional MeasureTheory MeasureTheory.Measure Set variable {ι F : Type*} @@ -134,15 +132,14 @@ theorem volume_ball (x : EuclideanSpace ℝ (Fin 2)) (r : ℝ) : exact lt_of_add_lt_of_nonneg_left h (sq_nonneg _) _ = volume (regionBetween (fun x => - Real.sqrt (1 - x ^ 2)) (fun x => Real.sqrt (1 - x ^ 2)) (Set.Ioc (-1) 1)) := by - simp_rw [regionBetween, Set.mem_Ioo, Set.mem_Ioc, ← Real.sq_lt, lt_tsub_iff_left, - Nat.cast_one] + simp_rw [regionBetween, Set.mem_Ioo, Set.mem_Ioc, ← Real.sq_lt, lt_tsub_iff_left] _ = ENNReal.ofReal ((2 : ℝ) * ∫ (a : ℝ) in Set.Ioc (-1) 1, Real.sqrt (1 - a ^ 2)) := by rw [volume_eq_prod, volume_regionBetween_eq_integral (Continuous.integrableOn_Ioc (by continuity)) (Continuous.integrableOn_Ioc (by continuity)) measurableSet_Ioc (fun _ _ => neg_le_self (Real.sqrt_nonneg _))] simp_rw [Pi.sub_apply, sub_neg_eq_add, ← two_mul, integral_mul_left] _ = NNReal.pi := by - rw [← intervalIntegral.integral_of_le (by norm_num : (-1 : ℝ) ≤ 1), Nat.cast_one, + rw [← intervalIntegral.integral_of_le (by norm_num : (-1 : ℝ) ≤ 1), integral_sqrt_one_sub_sq, two_mul, add_halves, ← NNReal.coe_real_pi, ofReal_coe_nnreal] diff --git a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean index f4f059be045a2..7ed4718d1723a 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/NormedSpace.lean @@ -13,8 +13,6 @@ import Mathlib.MeasureTheory.Integral.Bochner -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open scoped NNReal ENNReal Pointwise BigOperators Topology @@ -137,7 +135,7 @@ theorem integrable_comp_smul_iff {E : Type*} [NormedAddCommGroup E] [NormedSpace (f : E → F) {R : ℝ} (hR : R ≠ 0) : Integrable (fun x => f (R • x)) μ ↔ Integrable f μ := by -- reduce to one-way implication suffices - ∀ {g : E → F} (hg : Integrable g μ) {S : ℝ} (hS : S ≠ 0), Integrable (fun x => g (S • x)) μ by + ∀ {g : E → F} (_ : Integrable g μ) {S : ℝ} (_ : S ≠ 0), Integrable (fun x => g (S • x)) μ by refine' ⟨fun hf => _, fun hf => this hf hR⟩ convert this hf (inv_ne_zero hR) rw [← mul_smul, mul_inv_cancel hR, one_smul] diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean index a4dc0e5d2c72d..87bf58908cee9 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Basic.lean @@ -34,8 +34,6 @@ assert_not_exists MeasureTheory.integral noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Classical Set Filter MeasureTheory MeasureTheory.Measure TopologicalSpace open ENNReal (ofReal) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean index d331adad640e8..ad72b5e54bad8 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean @@ -18,8 +18,6 @@ used ways to represent `ℝ²` in `mathlib`: `ℝ × ℝ` and `Fin 2 → ℝ`, d of `MeasureTheory.measurePreserving`). -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open MeasureTheory noncomputable section diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean index 4c1a06ff60ef5..07835e9bbe327 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean @@ -43,8 +43,6 @@ density one for the rescaled copies `{x} + r • t` of a given set `t` with posi small `r`, see `eventually_nonempty_inter_smul_of_density_one`. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - assert_not_exists MeasureTheory.integral open TopologicalSpace Set Filter Metric Bornology diff --git a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean index f722e92f77b3b..b730b3cb142d7 100644 --- a/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/ProbabilityMeasure.lean @@ -326,7 +326,7 @@ total mass. -/ def normalize : ProbabilityMeasure Ω := if zero : μ.mass = 0 then ⟨Measure.dirac ‹Nonempty Ω›.some, Measure.dirac.isProbabilityMeasure⟩ else - { val := μ.mass⁻¹ • μ + { val := ↑(μ.mass⁻¹ • μ) property := by refine' ⟨_⟩ -- porting note: paying the price that this isn't `simp` lemma now. diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 3abc275638f77..70c8060fd2c9d 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -454,7 +454,7 @@ theorem FiniteAtFilter.exists_mem_basis {f : Filter α} (hμ : FiniteAtFilter μ #align measure_theory.measure.finite_at_filter.exists_mem_basis MeasureTheory.Measure.FiniteAtFilter.exists_mem_basis theorem finiteAtBot {m0 : MeasurableSpace α} (μ : Measure α) : μ.FiniteAtFilter ⊥ := - ⟨∅, mem_bot, by simp only [measure_empty, WithTop.zero_lt_top]⟩ + ⟨∅, mem_bot, by simp only [OuterMeasure.empty', zero_lt_top]⟩ #align measure_theory.measure.finite_at_bot MeasureTheory.Measure.finiteAtBot /-- `μ` has finite spanning sets in `C` if there is a countable sequence of sets in `C` that have diff --git a/Mathlib/NumberTheory/Basic.lean b/Mathlib/NumberTheory/Basic.lean index 3d5afd3fc4c1a..878ef79b3cf88 100644 --- a/Mathlib/NumberTheory/Basic.lean +++ b/Mathlib/NumberTheory/Basic.lean @@ -30,7 +30,7 @@ theorem dvd_sub_pow_of_dvd_sub {R : Type*} [CommRing R] {p : ℕ} {a b : R} (h : (k : ℕ) : (p ^ (k + 1) : R) ∣ a ^ p ^ k - b ^ p ^ k := by induction' k with k ih · rwa [pow_one, pow_zero, pow_one, pow_one] - rw [pow_succ' p k, pow_mul, pow_mul, ← geom_sum₂_mul, pow_succ, Nat.cast_mul] + rw [pow_succ' p k, pow_mul, pow_mul, ← geom_sum₂_mul, pow_succ] refine' mul_dvd_mul _ ih let f : R →+* R ⧸ span {(p : R)} := mk (span {(p : R)}) have hf : ∀ r : R, (p : R) ∣ r ↔ f r = 0 := fun r ↦ by rw [eq_zero_iff_mem, mem_span_singleton] diff --git a/Mathlib/NumberTheory/Bernoulli.lean b/Mathlib/NumberTheory/Bernoulli.lean index 4184598955f54..73259be69f824 100644 --- a/Mathlib/NumberTheory/Bernoulli.lean +++ b/Mathlib/NumberTheory/Bernoulli.lean @@ -125,7 +125,7 @@ theorem bernoulli'_three : bernoulli' 3 = 0 := by @[simp] theorem bernoulli'_four : bernoulli' 4 = -1 / 30 := by - have : Nat.choose 4 2 = 6 := by norm_num -- shrug + have : Nat.choose 4 2 = 6 := by decide -- shrug rw [bernoulli'_def] norm_num [sum_range_succ, sum_range_succ, sum_range_zero, this] #align bernoulli'_four bernoulli'_four @@ -185,7 +185,7 @@ theorem bernoulli'_odd_eq_zero {n : ℕ} (h_odd : Odd n) (hlt : 1 < n) : bernoul · apply eq_zero_of_neg_eq specialize h n split_ifs at h <;> simp_all [h_odd.neg_one_pow, factorial_ne_zero] - · simpa [Nat.factorial] using h 1 + · simpa (config := {decide := true}) [Nat.factorial] using h 1 have h : B * (exp ℚ - 1) = X * exp ℚ := by simpa [bernoulli'PowerSeries] using bernoulli'PowerSeries_mul_exp_sub_one ℚ rw [sub_mul, h, mul_sub X, sub_right_inj, ← neg_sub, mul_neg, neg_eq_iff_eq_neg] @@ -349,16 +349,15 @@ theorem sum_range_pow (n p : ℕ) : have hexp : exp ℚ - 1 ≠ 0 := by simp only [exp, PowerSeries.ext_iff, Ne, not_forall] use 1 - simp + simp [factorial_ne_zero] have h_r : exp ℚ ^ n - 1 = X * mk fun p => coeff ℚ (p + 1) (exp ℚ ^ n) := by have h_const : C ℚ (constantCoeff ℚ (exp ℚ ^ n)) = 1 := by simp rw [← h_const, sub_const_eq_X_mul_shift] -- key step: a chain of equalities of power series -- porting note: altered proof slightly rw [← mul_right_inj' hexp, mul_comm] - simp only [← cast_pow] rw [←exp_pow_sum, geom_sum_mul, h_r, ← bernoulliPowerSeries_mul_exp_sub_one, - bernoulliPowerSeries, mul_right_comm] + bernoulliPowerSeries, mul_right_comm] simp only [mul_comm, mul_eq_mul_left_iff, hexp, or_false] refine' Eq.trans (mul_eq_mul_right_iff.mpr _) (Eq.trans h_cauchy _) · left diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 48ee97fc0d7b5..5f5c16c57c6b8 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -99,17 +99,15 @@ theorem real_main_inequality {x : ℝ} (n_large : (512 : ℝ) ≤ x) : · have : sqrt (2 * 512) = 32 := (sqrt_eq_iff_mul_self_eq_of_pos (by norm_num1)).mpr (by norm_num1) rw [hf, log_nonpos_iff (hf' _ _), this, div_le_one] <;> norm_num1 - have : (512 : ℝ) = 2 ^ (9 : ℕ) - · rw [rpow_nat_cast 2 9]; norm_num1 - conv_lhs => rw [this] - have : (1024 : ℝ) = 2 ^ (10 : ℕ) - · rw [rpow_nat_cast 2 10]; norm_num1 - rw [this, ← rpow_mul, ← rpow_add] <;> norm_num1 - have : (4 : ℝ) = 2 ^ (2 : ℕ) - · rw [rpow_nat_cast 2 2]; norm_num1 - rw [this, ← rpow_mul] <;> norm_num1 - apply rpow_le_rpow_of_exponent_le <;> norm_num1 - apply rpow_pos_of_pos four_pos + · conv in 512 => equals 2 ^ 9 => norm_num1 + conv in 1024 => equals 2 ^ 10 => norm_num1 + conv in 32 => rw [← Nat.cast_ofNat] + rw [rpow_nat_cast, ← pow_mul, ← pow_add] + conv in 4 => equals 2 ^ (2 : ℝ) => rw [rpow_two]; norm_num1 + rw [← rpow_mul, ← rpow_nat_cast] + apply rpow_le_rpow_of_exponent_le + all_goals norm_num1 + · apply rpow_pos_of_pos four_pos #align bertrand.real_main_inequality Bertrand.real_main_inequality end Bertrand diff --git a/Mathlib/NumberTheory/ClassNumber/Finite.lean b/Mathlib/NumberTheory/ClassNumber/Finite.lean index 9d565d1bead91..3cd3f9c15a728 100644 --- a/Mathlib/NumberTheory/ClassNumber/Finite.lean +++ b/Mathlib/NumberTheory/ClassNumber/Finite.lean @@ -215,7 +215,8 @@ theorem exists_mem_finsetApprox (a : S) {b} (hb : b ≠ (0 : R)) : have dim_pos := Fintype.card_pos_iff.mpr bS.index_nonempty set ε : ℝ := normBound abv bS ^ (-1 / Fintype.card ι : ℝ) with ε_eq have hε : 0 < ε := Real.rpow_pos_of_pos (Int.cast_pos.mpr (normBound_pos abv bS)) _ - have ε_le : (normBound abv bS : ℝ) * (abv b • ε) ^ Fintype.card ι ≤ abv b ^ Fintype.card ι := by + have ε_le : (normBound abv bS : ℝ) * (abv b • ε) ^ (Fintype.card ι : ℝ) + ≤ abv b ^ (Fintype.card ι : ℝ) := by have := normBound_pos abv bS have := abv.nonneg b rw [ε_eq, Algebra.smul_def, eq_intCast, mul_rpow, ← rpow_mul, div_mul_cancel, rpow_neg_one, diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index 57dcb6facf33c..cde79426c23c4 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -23,8 +23,6 @@ We compute the discriminant of a `p ^ n`-th cyclotomic extension. universe u v -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open Algebra Polynomial Nat IsPrimitiveRoot PowerBasis open scoped Polynomial Cyclotomic @@ -92,7 +90,7 @@ theorem discr_prime_pow_ne_two [IsCyclotomicExtension {p ^ (k + 1)} K L] [hp : F obtain ⟨a, ha⟩ := (hp.out.even_sub_one hp2).two_dvd rw [ha, mul_left_comm, mul_assoc, Nat.mul_div_cancel_left _ two_pos, Nat.mul_div_cancel_left _ two_pos, mul_right_comm, pow_mul, (hpo.pow.mul _).neg_one_pow, - pow_mul, hpo.pow.neg_one_pow]; · norm_cast + pow_mul, hpo.pow.neg_one_pow] refine' Nat.Even.sub_odd _ (even_two_mul _) odd_one rw [mul_left_comm, ← ha] exact one_le_mul (one_le_pow _ _ hp.1.pos) (succ_le_iff.2 <| tsub_pos_of_lt hp.1.one_lt) @@ -176,7 +174,7 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( simp only [discr, traceMatrix_apply, Matrix.det_unique, Fin.default_eq_zero, Fin.val_zero, _root_.pow_zero, traceForm_apply, mul_one] rw [← (algebraMap K L).map_one, trace_algebraMap, finrank _ hirr, hp, hk]; norm_num - simp [← coe_two] + simp [← coe_two, Even.neg_pow (by decide : Even (1 / 2))] · exact discr_prime_pow_ne_two hζ hirr hk #align is_cyclotomic_extension.discr_prime_pow IsCyclotomicExtension.discr_prime_pow @@ -203,7 +201,7 @@ theorem discr_odd_prime [IsCyclotomicExtension {p} K L] [hp : Fact (p : ℕ).Pri have : IsCyclotomicExtension {p ^ (0 + 1)} K L := by rw [zero_add, pow_one] infer_instance - have hζ' : IsPrimitiveRoot ζ ↑(p ^ (0 + 1)) := by simpa using hζ + have hζ' : IsPrimitiveRoot ζ (p ^ (0 + 1) :) := by simpa using hζ convert discr_prime_pow_ne_two hζ' (by simpa [hirr]) (by simp [hodd]) using 2 · rw [zero_add, pow_one, totient_prime hp.out] · rw [_root_.pow_zero, one_mul, zero_add, mul_one, Nat.sub_sub] diff --git a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean index a97b0a4225734..ab26287bcee10 100644 --- a/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean +++ b/Mathlib/NumberTheory/Cyclotomic/PrimitiveRoots.lean @@ -72,7 +72,6 @@ variable {p n : ℕ+} (A : Type w) (B : Type z) (K : Type u) {L : Type v} (C : T variable [CommRing A] [CommRing B] [Algebra A B] [IsCyclotomicExtension {n} A B] -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 section Zeta namespace IsCyclotomicExtension @@ -413,8 +412,8 @@ then the norm of `ζ - 1` is `p`. -/ theorem sub_one_norm_prime [hpri : Fact (p : ℕ).Prime] [hcyc : IsCyclotomicExtension {p} K L] (hζ : IsPrimitiveRoot ζ p) (hirr : Irreducible (cyclotomic p K)) (h : p ≠ 2) : norm K (ζ - 1) = p := by - replace hirr : Irreducible (cyclotomic (↑(p ^ (0 + 1)) : ℕ) K) := by simp [hirr] - replace hζ : IsPrimitiveRoot ζ (↑(p ^ (0 + 1)) : ℕ) := by simp [hζ] + replace hirr : Irreducible (cyclotomic (p ^ (0 + 1) : ℕ) K) := by simp [hirr] + replace hζ : IsPrimitiveRoot ζ (p ^ (0 + 1) : ℕ) := by simp [hζ] haveI : IsCyclotomicExtension {p ^ (0 + 1)} K L := by simp [hcyc] simpa using sub_one_norm_prime_ne_two hζ hirr h #align is_primitive_root.sub_one_norm_prime IsPrimitiveRoot.sub_one_norm_prime @@ -452,7 +451,7 @@ theorem sub_one_norm_two {k : ℕ} (hζ : IsPrimitiveRoot ζ (2 ^ k)) (hk : 2 -- Porting note: `simpa using hirr` was `simp [hirr]`_ replace hirr : Irreducible (cyclotomic ((2 : ℕ+) ^ k : ℕ+) K) := by simpa using hirr -- Porting note: `simpa using hζ` was `simp [hζ]`_ - replace hζ : IsPrimitiveRoot ζ ((2 : ℕ+) ^ k) := by simpa using hζ + replace hζ : IsPrimitiveRoot ζ (2 ^ k : ℕ+) := by simpa using hζ obtain ⟨k₁, hk₁⟩ := exists_eq_succ_of_ne_zero (lt_of_lt_of_le zero_lt_two hk).ne.symm -- Porting note: the proof is slightly different because of coercions. simpa [hk₁, show ((2 : ℕ+) : ℕ) = 2 from rfl] using sub_one_norm_eq_eval_cyclotomic hζ this hirr @@ -469,7 +468,7 @@ theorem pow_sub_one_norm_prime_pow_of_ne_zero {k s : ℕ} (hζ : IsPrimitiveRoot rw [← PNat.coe_inj, PNat.pow_coe, ← pow_one 2] at htwo replace htwo := eq_of_prime_pow_eq (prime_iff.1 hpri.out) (prime_iff.1 Nat.prime_two) (succ_pos _) htwo - rwa [show 2 = ((2 : ℕ+) : ℕ) by simp, PNat.coe_inj] at htwo + rwa [show 2 = ((2 : ℕ+) : ℕ) by decide, PNat.coe_inj] at htwo replace hs : s = k · rw [hp, ← PNat.coe_inj, PNat.pow_coe] at htwo nth_rw 2 [← pow_one 2] at htwo diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 51464c8a80fd1..24dcf4cdeb883 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -31,8 +31,6 @@ open scoped Cyclotomic NumberField Nat variable {p : ℕ+} {k : ℕ} {K : Type u} [Field K] [CharZero K] {ζ : K} [hp : Fact (p : ℕ).Prime] -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace IsCyclotomicExtension.Rat /-- The discriminant of the power basis given by `ζ - 1`. -/ @@ -92,7 +90,7 @@ theorem isIntegralClosure_adjoin_singleton_of_prime_pow [hcycl : IsCyclotomicExt replace H := Subalgebra.smul_mem _ H u.inv -- Porting note: the proof is slightly different because of coercions. rw [← smul_assoc, ← smul_mul_assoc, Units.inv_eq_val_inv, zsmul_eq_mul, ← Int.cast_mul, - Units.inv_mul, Int.cast_one, one_mul, PNat.pow_coe, Nat.cast_pow, smul_def, map_pow] at H + Units.inv_mul, Int.cast_one, one_mul, smul_def, map_pow] at H cases k · haveI : IsCyclotomicExtension {1} ℚ K := by simpa using hcycl have : x ∈ (⊥ : Subalgebra ℚ K) := by @@ -303,7 +301,7 @@ theorem zeta_sub_one_prime_of_two_pow [IsCyclotomicExtension {(2 : ℕ+) ^ (k + Prime (hζ.toInteger - 1) := by letI := IsCyclotomicExtension.numberField {(2 : ℕ+) ^ (k + 1)} ℚ K refine Ideal.prime_of_irreducible_absNorm_span (fun h ↦ ?_) ?_ - · apply hζ.pow_ne_one_of_pos_of_lt zero_lt_one (one_lt_pow (by norm_num) (by simp)) + · apply hζ.pow_ne_one_of_pos_of_lt zero_lt_one (one_lt_pow (by decide) (by simp)) rw [← Subalgebra.coe_eq_zero] at h simpa [sub_eq_zero] using h rw [Nat.irreducible_iff_prime, Ideal.absNorm_span_singleton, ← Nat.prime_iff, diff --git a/Mathlib/NumberTheory/Dioph.lean b/Mathlib/NumberTheory/Dioph.lean index f370519127211..f92d8dc0a97e8 100644 --- a/Mathlib/NumberTheory/Dioph.lean +++ b/Mathlib/NumberTheory/Dioph.lean @@ -372,7 +372,7 @@ theorem DiophList.forall (l : List (Set <| α → ℕ)) (d : l.Forall Dioph) : ⟨m ⊗ n, by rw [show (v ⊗ m ⊗ n) ∘ (inl ⊗ inr ∘ inl) = v ⊗ m from funext fun s => by cases' s with a b <;> rfl]; exact hm, by - refine List.Forall.imp (fun q hq => ?_) hn; dsimp [(· ∘ ·)] + refine List.Forall.imp (fun q hq => ?_) hn; dsimp [Function.comp_def] rw [show (fun x : Sum α γ => (v ⊗ m ⊗ n) ((inl ⊗ fun x : γ => inr (inr x)) x)) = v ⊗ n from funext fun s => by cases' s with a b <;> rfl]; exact hq⟩, @@ -381,7 +381,7 @@ theorem DiophList.forall (l : List (Set <| α → ℕ)) (d : l.Forall Dioph) : rwa [show (v ⊗ t) ∘ (inl ⊗ inr ∘ inl) = v ⊗ t ∘ inl from funext fun s => by cases' s with a b <;> rfl] at hl⟩, ⟨t ∘ inr, by - refine List.Forall.imp (fun q hq => ?_) hr; dsimp [(· ∘ ·)] at hq + refine List.Forall.imp (fun q hq => ?_) hr; dsimp [Function.comp_def] at hq rwa [show (fun x : Sum α γ => (v ⊗ t) ((inl ⊗ fun x : γ => inr (inr x)) x)) = v ⊗ t ∘ inr diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 6ab57f89f54c3..7a90aaed0a596 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -404,7 +404,8 @@ theorem properDivisors_eq_singleton_one_iff_prime : n.properDivisors = {1} ↔ n · simp [hdvd, this] exact (le_iff_eq_or_lt.mp this).symm · by_contra' - simp [nonpos_iff_eq_zero.mp this, this] at h + simp only [nonpos_iff_eq_zero.mp this, this] at h + contradiction · exact fun h => Prime.properDivisors h #align nat.proper_divisors_eq_singleton_one_iff_prime Nat.properDivisors_eq_singleton_one_iff_prime diff --git a/Mathlib/NumberTheory/FLT/Four.lean b/Mathlib/NumberTheory/FLT/Four.lean index cb8144742a703..20cde9384f650 100644 --- a/Mathlib/NumberTheory/FLT/Four.lean +++ b/Mathlib/NumberTheory/FLT/Four.lean @@ -131,7 +131,7 @@ theorem exists_odd_minimal {a b c : ℤ} (h : Fermat42 a b c) : Int.dvd_gcd (Int.dvd_of_emod_eq_zero hap) (Int.dvd_of_emod_eq_zero hbp) rw [Int.gcd_eq_one_iff_coprime.mpr (coprime_of_minimal hf)] at h1 revert h1 - norm_num + decide · exact ⟨b0, ⟨a0, ⟨c0, minimal_comm hf, hbp⟩⟩⟩ exact ⟨a0, ⟨b0, ⟨c0, hf, hap⟩⟩⟩ #align fermat_42.exists_odd_minimal Fermat42.exists_odd_minimal @@ -179,7 +179,7 @@ theorem not_minimal {a b c : ℤ} (h : Minimal a b c) (ha2 : a % 2 = 1) (hc : 0 -- it helps if we know the parity of a ^ 2 (and the sign of c): have ha22 : a ^ 2 % 2 = 1 := by rw [sq, Int.mul_emod, ha2] - norm_num + decide obtain ⟨m, n, ht1, ht2, ht3, ht4, ht5, ht6⟩ := ht.coprime_classification' h2 ha22 hc -- Now a, n, m form a pythagorean triple and so we can obtain r and s such that -- a = r ^ 2 - s ^ 2, n = 2 * r * s and m = r ^ 2 + s ^ 2 diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index ffd1d3f8e681d..5be06f3e6cced 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -19,8 +19,6 @@ open Finset Nat open scoped BigOperators Nat -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - section GaussEisenstein namespace ZMod @@ -95,7 +93,7 @@ private theorem gauss_lemma_aux₁ (p : ℕ) [Fact p.Prime] {a : ℤ} theorem gauss_lemma_aux (p : ℕ) [hp : Fact p.Prime] {a : ℤ} (hap : (a : ZMod p) ≠ 0) : (↑a ^ (p / 2) : ZMod p) = - (-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card := + ((-1) ^ ((Ico 1 (p / 2).succ).filter fun x : ℕ => p / 2 < (a * x : ZMod p).val).card :) := (mul_left_inj' (show ((p / 2)! : ZMod p) ≠ 0 by rw [Ne.def, CharP.cast_eq_zero_iff (ZMod p) p, hp.1.dvd_factorial, not_le] exact Nat.div_lt_self hp.1.pos (by decide))).1 <| by diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index 29e832b2e7e19..04552a4060b64 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -468,7 +468,7 @@ theorem quadratic_reciprocity_three_mod_four {a b : ℕ} (ha : a % 4 = 3) (hb : theorem mod_right' (a : ℕ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * a)) := by rcases eq_or_ne a 0 with (rfl | ha₀) · rw [mul_zero, mod_zero] - have hb' : Odd (b % (4 * a)) := hb.mod_even (Even.mul_right (by norm_num) _) + have hb' : Odd (b % (4 * a)) := hb.mod_even (Even.mul_right (by decide) _) rcases exists_eq_pow_mul_and_not_dvd ha₀ 2 (by norm_num) with ⟨e, a', ha₁', ha₂⟩ have ha₁ := odd_iff.mpr (two_dvd_ne_zero.mp ha₁') nth_rw 2 [ha₂]; nth_rw 1 [ha₂] @@ -480,7 +480,7 @@ theorem mod_right' (a : ℕ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * rw [χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * a)), mod_mod_of_dvd b (dvd_mul_right 4 a)] · rw [mod_left ↑(b % _), mod_left b, Int.coe_nat_mod, Int.emod_emod_of_dvd b] simp only [ha₂, Nat.cast_mul, ← mul_assoc] - exact dvd_mul_left (a' : ℤ) (↑4 * ↑(2 ^ e)) + apply dvd_mul_left -- porting note: In mathlib3, it was written `cases' e`. In Lean 4, this resulted in the choice -- of a name other than e (for the case distinction of line 482) so we indicate the name -- to use explicitly. @@ -493,7 +493,7 @@ theorem mod_right' (a : ℕ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * theorem mod_right (a : ℤ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * a.natAbs)) := by cases' Int.natAbs_eq a with ha ha <;> nth_rw 2 [ha] <;> nth_rw 1 [ha] · exact mod_right' a.natAbs hb - · have hb' : Odd (b % (4 * a.natAbs)) := hb.mod_even (Even.mul_right (by norm_num) _) + · have hb' : Odd (b % (4 * a.natAbs)) := hb.mod_even (Even.mul_right (by decide) _) rw [jacobiSym.neg _ hb, jacobiSym.neg _ hb', mod_right' _ hb, χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * _)), mod_mod_of_dvd b (dvd_mul_right 4 _)] #align jacobi_sym.mod_right jacobiSym.mod_right diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean index 4a608d0f1874e..aa50a12f53dbb 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/Basic.lean @@ -85,7 +85,7 @@ theorem quadraticCharFun_one : quadraticCharFun F 1 = 1 := by theorem quadraticCharFun_eq_one_of_char_two (hF : ringChar F = 2) {a : F} (ha : a ≠ 0) : quadraticCharFun F a = 1 := by simp only [quadraticCharFun, ha, if_false, ite_eq_left_iff] - exact fun h => h (FiniteField.isSquare_of_char_two hF a) + exact fun h => (h (FiniteField.isSquare_of_char_two hF a)).elim #align quadratic_char_fun_eq_one_of_char_two quadraticCharFun_eq_one_of_char_two /-- If `ringChar F` is odd, then `quadraticCharFun F a` can be computed in diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean index f7d3bddc79562..3fe7b1f41e066 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticChar/GaussSum.lean @@ -54,7 +54,7 @@ theorem FiniteField.isSquare_two_iff : simp only [h, Nat.one_ne_zero, if_false, ite_eq_left_iff, Ne.def, (by decide : (-1 : ℤ) ≠ 1), imp_false, Classical.not_not] all_goals - rw [← Nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h + rw [← Nat.mod_mod_of_dvd _ (by decide : 2 ∣ 8)] at h have h₁ := Nat.mod_lt (Fintype.card F) (by decide : 0 < 8) revert h₁ h generalize Fintype.card F % 8 = n @@ -65,7 +65,7 @@ theorem FiniteField.isSquare_two_iff : theorem quadraticChar_neg_two [DecidableEq F] (hF : ringChar F ≠ 2) : quadraticChar F (-2) = χ₈' (Fintype.card F) := by rw [(by norm_num : (-2 : F) = -1 * 2), map_mul, χ₈'_eq_χ₄_mul_χ₈, quadraticChar_neg_one hF, - quadraticChar_two hF, @cast_nat_cast _ (ZMod 4) _ _ _ (by norm_num : 4 ∣ 8)] + quadraticChar_two hF, @cast_nat_cast _ (ZMod 4) _ _ _ (by decide : 4 ∣ 8)] #align quadratic_char_neg_two quadraticChar_neg_two /-- `-2` is a square in `F` iff `#F` is not congruent to `5` or `7` mod `8`. -/ @@ -84,7 +84,7 @@ theorem FiniteField.isSquare_neg_two_iff : simp only [h, Nat.one_ne_zero, if_false, ite_eq_left_iff, Ne.def, (by decide : (-1 : ℤ) ≠ 1), imp_false, Classical.not_not] all_goals - rw [← Nat.mod_mod_of_dvd _ (by norm_num : 2 ∣ 8)] at h + rw [← Nat.mod_mod_of_dvd _ (by decide : 2 ∣ 8)] at h have h₁ := Nat.mod_lt (Fintype.card F) (by decide : 0 < 8) revert h₁ h generalize Fintype.card F % 8 = n diff --git a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean index 52d412108318c..36b611aecfb82 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/QuadraticReciprocity.lean @@ -78,7 +78,7 @@ variable (hp : p ≠ 2) theorem exists_sq_eq_two_iff : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 := by rw [FiniteField.isSquare_two_iff, card p] have h₁ := Prime.mod_two_eq_one_iff_ne_two.mpr hp - rw [← mod_mod_of_dvd p (by norm_num : 2 ∣ 8)] at h₁ + rw [← mod_mod_of_dvd p (by decide : 2 ∣ 8)] at h₁ have h₂ := mod_lt p (by norm_num : 0 < 8) revert h₂ h₁ generalize p % 8 = m; clear! p @@ -89,7 +89,7 @@ theorem exists_sq_eq_two_iff : IsSquare (2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 7 theorem exists_sq_eq_neg_two_iff : IsSquare (-2 : ZMod p) ↔ p % 8 = 1 ∨ p % 8 = 3 := by rw [FiniteField.isSquare_neg_two_iff, card p] have h₁ := Prime.mod_two_eq_one_iff_ne_two.mpr hp - rw [← mod_mod_of_dvd p (by norm_num : 2 ∣ 8)] at h₁ + rw [← mod_mod_of_dvd p (by decide : 2 ∣ 8)] at h₁ have h₂ := mod_lt p (by norm_num : 0 < 8) revert h₂ h₁ generalize p % 8 = m; clear! p diff --git a/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean b/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean index 7b7731b54bbf6..61dc2f08096cd 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/ZModChar.lean @@ -67,7 +67,7 @@ theorem χ₄_int_eq_if_mod_four (n : ℤ) : χ₄ n = if n % 2 = 0 then 0 else if n % 4 = 1 then 1 else -1 := by have help : ∀ m : ℤ, 0 ≤ m → m < 4 → χ₄ m = if m % 2 = 0 then 0 else if m = 1 then 1 else -1 := by decide - rw [← Int.emod_emod_of_dvd n (by norm_num : (2 : ℤ) ∣ 4), ← ZMod.int_cast_mod n 4] + rw [← Int.emod_emod_of_dvd n (by decide : (2 : ℤ) ∣ 4), ← ZMod.int_cast_mod n 4] exact help (n % 4) (Int.emod_nonneg n (by norm_num)) (Int.emod_lt n (by norm_num)) #align zmod.χ₄_int_eq_if_mod_four ZMod.χ₄_int_eq_if_mod_four @@ -88,7 +88,7 @@ theorem χ₄_eq_neg_one_pow {n : ℕ} (hn : n % 2 = 1) : χ₄ n = (-1) ^ (n / have help : ∀ m : ℕ, m < 4 → m % 2 = 1 → ite (m = 1) (1 : ℤ) (-1) = (-1) ^ (m / 2) := by decide exact help (n % 4) (Nat.mod_lt n (by norm_num)) - ((Nat.mod_mod_of_dvd n (by norm_num : 2 ∣ 4)).trans hn) + ((Nat.mod_mod_of_dvd n (by decide : 2 ∣ 4)).trans hn) #align zmod.χ₄_eq_neg_one_pow ZMod.χ₄_eq_neg_one_pow /-- If `n % 4 = 1`, then `χ₄ n = 1`. -/ @@ -162,7 +162,7 @@ theorem χ₈_int_eq_if_mod_eight (n : ℤ) : have help : ∀ m : ℤ, 0 ≤ m → m < 8 → χ₈ m = if m % 2 = 0 then 0 else if m = 1 ∨ m = 7 then 1 else -1 := by decide - rw [← Int.emod_emod_of_dvd n (by norm_num : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] + rw [← Int.emod_emod_of_dvd n (by decide : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] exact help (n % 8) (Int.emod_nonneg n (by norm_num)) (Int.emod_lt n (by norm_num)) #align zmod.χ₈_int_eq_if_mod_eight ZMod.χ₈_int_eq_if_mod_eight @@ -195,7 +195,7 @@ theorem χ₈'_int_eq_if_mod_eight (n : ℤ) : have help : ∀ m : ℤ, 0 ≤ m → m < 8 → χ₈' m = if m % 2 = 0 then 0 else if m = 1 ∨ m = 3 then 1 else -1 := by decide - rw [← Int.emod_emod_of_dvd n (by norm_num : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] + rw [← Int.emod_emod_of_dvd n (by decide : (2 : ℤ) ∣ 8), ← ZMod.int_cast_mod n 8] exact help (n % 8) (Int.emod_nonneg n (by norm_num)) (Int.emod_lt n (by norm_num)) #align zmod.χ₈'_int_eq_if_mod_eight ZMod.χ₈'_int_eq_if_mod_eight @@ -212,7 +212,7 @@ theorem χ₈'_eq_χ₄_mul_χ₈ (a : ZMod 8) : χ₈' a = χ₄ a * χ₈ a := #align zmod.χ₈'_eq_χ₄_mul_χ₈ ZMod.χ₈'_eq_χ₄_mul_χ₈ theorem χ₈'_int_eq_χ₄_mul_χ₈ (a : ℤ) : χ₈' a = χ₄ a * χ₈ a := by - rw [← @cast_int_cast 8 (ZMod 4) _ 4 _ (by norm_num) a] + rw [← @cast_int_cast 8 (ZMod 4) _ 4 _ (by decide) a] exact χ₈'_eq_χ₄_mul_χ₈ a #align zmod.χ₈'_int_eq_χ₄_mul_χ₈ ZMod.χ₈'_int_eq_χ₄_mul_χ₈ diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean index 115358469645a..0c1230384b3c6 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleNumber.lean @@ -172,7 +172,7 @@ theorem remainder_lt (n : ℕ) {m : ℝ} (m2 : 2 ≤ m) : remainder m n < 1 / (m /-- The sum of the `k` initial terms of the Liouville number to base `m` is a ratio of natural numbers where the denominator is `m ^ k!`. -/ theorem partialSum_eq_rat {m : ℕ} (hm : 0 < m) (k : ℕ) : - ∃ p : ℕ, partialSum m k = p / (m ^ k ! : ℝ) := by + ∃ p : ℕ, partialSum m k = p / ((m ^ k ! :) : ℝ) := by induction' k with k h · exact ⟨1, by rw [partialSum, range_one, sum_singleton, Nat.cast_one, Nat.factorial, pow_one, pow_one]⟩ @@ -196,7 +196,7 @@ theorem liouville_liouvilleNumber {m : ℕ} (hm : 2 ≤ m) : Liouville (liouvill intro n -- the first `n` terms sum to `p / m ^ k!` rcases partialSum_eq_rat (zero_lt_two.trans_le hm) n with ⟨p, hp⟩ - refine' ⟨p, m ^ n !, by rw [Nat.cast_pow]; exact one_lt_pow mZ1 n.factorial_ne_zero, _⟩ + refine' ⟨p, m ^ n !, one_lt_pow mZ1 n.factorial_ne_zero, _⟩ push_cast rw [Nat.cast_pow] at hp -- separate out the sum of the first `n` terms and the rest diff --git a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean index dabb8385867ce..005ffa818ccd2 100644 --- a/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean +++ b/Mathlib/NumberTheory/Liouville/LiouvilleWith.lean @@ -337,8 +337,6 @@ namespace Liouville variable {x : ℝ} -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- If `x` is a Liouville number, then for any `n`, for infinitely many denominators `b` there exists a numerator `a` such that `x ≠ a / b` and `|x - a / b| < 1 / b ^ n`. -/ theorem frequently_exists_num (hx : Liouville x) (n : ℕ) : diff --git a/Mathlib/NumberTheory/Liouville/Measure.lean b/Mathlib/NumberTheory/Liouville/Measure.lean index fbfb674f40a19..70318cdd9405d 100644 --- a/Mathlib/NumberTheory/Liouville/Measure.lean +++ b/Mathlib/NumberTheory/Liouville/Measure.lean @@ -31,8 +31,6 @@ open scoped Filter BigOperators ENNReal Topology NNReal open Filter Set Metric MeasureTheory Real -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - theorem setOf_liouvilleWith_subset_aux : { x : ℝ | ∃ p > 2, LiouvilleWith p x } ⊆ ⋃ m : ℤ, (· + (m : ℝ)) ⁻¹' ⋃ n > (0 : ℕ), diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index b9bced6e19a27..e534be43359e7 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -132,7 +132,7 @@ theorem tendsto_normSq_coprime_pair : normSq ∘ f ∘ fun p : Fin 2 → ℤ => ((↑) : ℤ → ℝ) ∘ p := by ext1 rw [f_def] - dsimp only [Function.comp] + dsimp only [Function.comp_def] rw [ofReal_int_cast, ofReal_int_cast] rw [this] have hf : LinearMap.ker f = ⊥ := by diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Basic.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Basic.lean index de2cf9c2a3f26..14fc769abb191 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Basic.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Basic.lean @@ -26,8 +26,6 @@ open Complex Real Asymptotics Filter open scoped Real BigOperators UpperHalfPlane -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- Jacobi's theta function `∑' (n : ℤ), exp (π * I * n ^ 2 * τ)`. -/ noncomputable def jacobiTheta (z : ℂ) : ℂ := ∑' n : ℤ, cexp (π * I * (n : ℂ) ^ 2 * z) @@ -39,13 +37,13 @@ theorem norm_exp_mul_sq_le {z : ℂ} (hz : 0 < z.im) (n : ℤ) : have h : y < 1 := exp_lt_one_iff.mpr (mul_neg_of_neg_of_pos (neg_lt_zero.mpr pi_pos) hz) refine' (le_of_eq _).trans (_ : y ^ n ^ 2 ≤ _) · rw [Complex.norm_eq_abs, Complex.abs_exp] - have : (↑π * I * (n : ℂ) ^ 2 * z).re = -π * z.im * (n : ℝ) ^ 2 := by - rw [(by push_cast; ring : ↑π * I * (n : ℂ) ^ 2 * z = ↑(π * (n : ℝ) ^ 2) * (z * I)), + have : (π * I * n ^ 2 * z : ℂ).re = -π * z.im * (n : ℝ) ^ 2 := by + rw [(by push_cast; ring : (π * I * n ^ 2 * z : ℂ) = (π * n ^ 2 : ℝ) * (z * I)), ofReal_mul_re, mul_I_re] ring obtain ⟨m, hm⟩ := Int.eq_ofNat_of_zero_le (sq_nonneg n) rw [this, exp_mul, ← Int.cast_pow, rpow_int_cast, hm, zpow_ofNat] - · have : n ^ 2 = ↑(n.natAbs ^ 2) := by rw [Nat.cast_pow, Int.natAbs_sq] + · have : n ^ 2 = (n.natAbs ^ 2 :) := by rw [Nat.cast_pow, Int.natAbs_sq] rw [this, zpow_ofNat] exact pow_le_pow_of_le_one (exp_pos _).le h.le ((sq n.natAbs).symm ▸ n.natAbs.le_mul_self) #align norm_exp_mul_sq_le norm_exp_mul_sq_le @@ -72,13 +70,13 @@ theorem summable_exp_mul_sq {z : ℂ} (hz : 0 < z.im) : theorem jacobiTheta_two_add (z : ℂ) : jacobiTheta (2 + z) = jacobiTheta z := by refine' tsum_congr fun n => _ - suffices cexp (↑π * I * (n : ℂ) ^ 2 * 2) = 1 by rw [mul_add, Complex.exp_add, this, one_mul] - rw [(by push_cast; ring : ↑π * I * ↑n ^ 2 * 2 = ↑(n ^ 2) * (2 * π * I)), Complex.exp_int_mul, + suffices cexp (π * I * n ^ 2 * 2 : ℂ) = 1 by rw [mul_add, Complex.exp_add, this, one_mul] + rw [(by push_cast; ring : (π * I * n ^ 2 * 2 : ℂ) = (n ^ 2 :) * (2 * π * I)), Complex.exp_int_mul, Complex.exp_two_pi_mul_I, one_zpow] #align jacobi_theta_two_add jacobiTheta_two_add -theorem jacobiTheta_T_sq_smul (τ : ℍ) : jacobiTheta ↑(ModularGroup.T ^ 2 • τ) = jacobiTheta τ := by - suffices ↑(ModularGroup.T ^ 2 • τ) = (2 : ℂ) + ↑τ by simp_rw [this, jacobiTheta_two_add] +theorem jacobiTheta_T_sq_smul (τ : ℍ) : jacobiTheta (ModularGroup.T ^ 2 • τ :) = jacobiTheta τ := by + suffices (ModularGroup.T ^ 2 • τ :) = (2 : ℂ) + ↑τ by simp_rw [this, jacobiTheta_two_add] have : ModularGroup.T ^ (2 : ℕ) = ModularGroup.T ^ (2 : ℤ) := rfl simp_rw [this, UpperHalfPlane.modular_T_zpow_smul, UpperHalfPlane.coe_vadd] norm_cast diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index 2a417e3e8ab96..50c6c2484f6da 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -29,8 +29,6 @@ open Complex UpperHalfPlane open scoped UpperHalfPlane -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - local notation "GL(" n ", " R ")" "⁺" => Matrix.GLPos (Fin n) R local notation "SL(" n ", " R ")" => Matrix.SpecialLinearGroup (Fin n) R diff --git a/Mathlib/NumberTheory/Multiplicity.lean b/Mathlib/NumberTheory/Multiplicity.lean index cae7fdf3231f7..175f78c69e15b 100644 --- a/Mathlib/NumberTheory/Multiplicity.lean +++ b/Mathlib/NumberTheory/Multiplicity.lean @@ -109,7 +109,7 @@ theorem odd_sq_dvd_geom_sum₂_sub (hp : Odd p) : _ = mk (span {s}) (∑ x : ℕ in Finset.range p, a ^ (x - 1) * (a ^ (p - 1 - x) * (↑p * (b * ↑x)))) + - mk (span {s}) (∑ x : ℕ in Finset.range p, a ^ (p - 1)) := by + mk (span {s}) (∑ _x : ℕ in Finset.range p, a ^ (p - 1)) := by rw [add_right_inj] have : ∀ (x : ℕ), (hx : x ∈ range p) → a ^ (x + (p - 1 - x)) = a ^ (p - 1) := by intro x hx @@ -269,13 +269,13 @@ theorem Int.sq_mod_four_eq_one_of_odd {x : ℤ} : Odd x → x ^ 2 % 4 = 1 := by rcases hx with ⟨_, rfl⟩ ring_nf rw [add_assoc, ← add_mul, Int.add_mul_emod_self] - norm_num + decide #align int.sq_mod_four_eq_one_of_odd Int.sq_mod_four_eq_one_of_odd theorem Int.two_pow_two_pow_add_two_pow_two_pow {x y : ℤ} (hx : ¬2 ∣ x) (hxy : 4 ∣ x - y) (i : ℕ) : multiplicity 2 (x ^ 2 ^ i + y ^ 2 ^ i) = ↑(1 : ℕ) := by have hx_odd : Odd x := by rwa [Int.odd_iff_not_even, even_iff_two_dvd] - have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by norm_num) hxy) + have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by decide) hxy) have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even refine' multiplicity.eq_coe_iff.mpr ⟨_, _⟩ · rw [pow_one, ← even_iff_two_dvd] @@ -292,7 +292,7 @@ theorem Int.two_pow_two_pow_add_two_pow_two_pow {x y : ℤ} (hx : ¬2 ∣ x) (hx suffices ∀ x : ℤ, Odd x → x ^ 2 ^ (i + 1) % 4 = 1 by rw [show (2 ^ (1 + 1) : ℤ) = 4 by norm_num, Int.dvd_iff_emod_eq_zero, Int.add_emod, this _ hx_odd, this _ hy_odd] - norm_num + decide intro x hx rw [pow_succ, mul_comm, pow_mul, Int.sq_mod_four_eq_one_of_odd hx.pow] #align int.two_pow_two_pow_add_two_pow_two_pow Int.two_pow_two_pow_add_two_pow_two_pow @@ -307,7 +307,7 @@ theorem Int.two_pow_two_pow_sub_pow_two_pow {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x theorem Int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬2 ∣ x) : multiplicity 2 (x ^ n - y ^ n) = multiplicity 2 (x - y) + multiplicity (2 : ℤ) n := by have hx_odd : Odd x := by rwa [Int.odd_iff_not_even, even_iff_two_dvd] - have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by norm_num) hxy) + have hxy_even : Even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by decide) hxy) have hy_odd : Odd y := by simpa using hx_odd.sub_even hxy_even cases' n with n · simp only [pow_zero, sub_self, multiplicity.zero, Int.ofNat_zero, Nat.zero_eq, add_top] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index eb1309bb492c5..d5c322f7263bb 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -343,10 +343,10 @@ def matrixToStdBasis : Matrix (index K) (index K) ℂ := theorem det_matrixToStdBasis : (matrixToStdBasis K).det = (2⁻¹ * I) ^ NrComplexPlaces K := calc - _ = ∏ k : { w : InfinitePlace K // IsComplex w }, det ((2 : ℂ)⁻¹ • !![1, 1; -I, I]) := by + _ = ∏ _k : { w : InfinitePlace K // IsComplex w }, det ((2 : ℂ)⁻¹ • !![1, 1; -I, I]) := by rw [matrixToStdBasis, det_fromBlocks_zero₂₁, det_diagonal, Finset.prod_const_one, one_mul, det_reindex_self, det_blockDiagonal] - _ = ∏ k : { w : InfinitePlace K // IsComplex w }, (2⁻¹ * Complex.I) := by + _ = ∏ _k : { w : InfinitePlace K // IsComplex w }, (2⁻¹ * Complex.I) := by refine Finset.prod_congr (Eq.refl _) (fun _ _ => ?_) field_simp; ring _ = (2⁻¹ * Complex.I) ^ Fintype.card {w : InfinitePlace K // IsComplex w} := by diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 388a3f33a4703..722fb0b0fe4ea 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -32,8 +32,6 @@ for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where the product is over th number field, embeddings, places, infinite places -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped Classical namespace NumberField.Embeddings @@ -496,7 +494,7 @@ theorem card_complex_embeddings : · rwa [Subtype.mk_eq_mk, ← Subtype.ext_iff, ← Subtype.ext_iff] at h · refine ⟨⟨⟨φ, not_isReal_of_mk_isComplex (hφ.symm ▸ hw)⟩, ?_⟩, rfl⟩ rwa [Subtype.ext_iff, mkComplex_coe] - · simp_rw [mult, not_isReal_iff_isComplex.mpr hw] + · simp_rw [mult, not_isReal_iff_isComplex.mpr hw, ite_false] #align number_field.infinite_place.card_complex_embeddings NumberField.InfinitePlace.card_complex_embeddings theorem card_add_two_mul_card_eq_rank : diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 13c5bc1c48b18..23869926dec23 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -88,10 +88,10 @@ variable {K} theorem coe_mul (x y : (𝓞 K)ˣ) : ((x * y : (𝓞 K)ˣ) : K) = (x : K) * (y : K) := rfl -theorem coe_pow (x : (𝓞 K)ˣ) (n : ℕ) : (x ^ n : K) = (x : K) ^ n := by +theorem coe_pow (x : (𝓞 K)ˣ) (n : ℕ) : (↑(x ^ n) : K) = (x : K) ^ n := by rw [← SubmonoidClass.coe_pow, ← val_pow_eq_pow_val] -theorem coe_zpow (x : (𝓞 K)ˣ) (n : ℤ) : (x ^ n : K) = (x : K) ^ n := by +theorem coe_zpow (x : (𝓞 K)ˣ) (n : ℤ) : (↑(x ^ n) : K) = (x : K) ^ n := by change ((Units.coeHom K).comp (map (algebraMap (𝓞 K) K))) (x ^ n) = _ exact map_zpow _ x n @@ -227,7 +227,7 @@ theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : mult w * Real.log (w x) = 0 ↔ w x = 1 := by rw [mul_eq_zero, or_iff_right, Real.log_eq_zero, or_iff_right, or_iff_left] · linarith [(map_nonneg _ _ : 0 ≤ w x)] - · simp only [ne_eq, map_eq_zero, coe_ne_zero x] + · simp only [ne_eq, map_eq_zero, coe_ne_zero x, not_false_eq_true] · refine (ne_of_gt ?_) rw [mult]; split_ifs <;> norm_num @@ -345,7 +345,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] exact le_of_eq (congr_arg toReal h_gprod) · refine div_lt_self ?_ (by norm_num) - simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx] + simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx, not_false_eq_true] intro _ _ rw [ne_eq, Nonneg.mk_eq_zero, div_eq_zero_iff, map_eq_zero, not_or, ZeroMemClass.coe_eq_zero] exact ⟨hx, by norm_num⟩ @@ -410,7 +410,7 @@ theorem exists_unit (w₁ : InfinitePlace K) : (Ideal.span ({ (seq K w₁ hB n : 𝓞 K) }) = Ideal.span ({ (seq K w₁ hB m : 𝓞 K) })) · have hu := Ideal.span_singleton_eq_span_singleton.mp h refine ⟨hu.choose, fun w hw => Real.log_neg ?_ ?_⟩ - · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero] + · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero, not_false_eq_true] · calc _ = w ((seq K w₁ hB m : K) * (seq K w₁ hB n : K)⁻¹) := by rw [← congr_arg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 88f02656c02ab..11224b03c8c7c 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -357,7 +357,8 @@ theorem norm_int_lt_one_iff_dvd (k : ℤ) : ‖(k : ℤ_[p])‖ < 1 ↔ (p : ℤ theorem norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ‖(k : ℤ_[p])‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ (p ^ n : ℤ) ∣ k := - suffices ‖(k : ℚ_[p])‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k by simpa [norm_int_cast_eq_padic_norm] + suffices ‖(k : ℚ_[p])‖ ≤ (p : ℝ) ^ (-n : ℤ) ↔ (p ^ n : ℤ) ∣ k by + simpa [norm_int_cast_eq_padic_norm] padicNormE.norm_int_le_pow_iff_dvd _ _ #align padic_int.norm_int_le_pow_iff_dvd PadicInt.norm_int_le_pow_iff_dvd diff --git a/Mathlib/NumberTheory/Padics/PadicVal.lean b/Mathlib/NumberTheory/Padics/PadicVal.lean index 703c9a3dd3719..8227fd8184cd4 100644 --- a/Mathlib/NumberTheory/Padics/PadicVal.lean +++ b/Mathlib/NumberTheory/Padics/PadicVal.lean @@ -607,7 +607,8 @@ lemma Nat.log_ne_padicValNat_succ {n : ℕ} (hn : n ≠ 0) : log 2 n ≠ padicVa rw [← lt_add_one_iff, ← mul_one (2 ^ _)] at h1 rw [← add_one_le_iff, pow_succ] at h2 refine' not_dvd_of_between_consec_multiples h1 (lt_of_le_of_ne' h2 _) pow_padicValNat_dvd - exact pow_succ_padicValNat_not_dvd n.succ_ne_zero ∘ dvd_of_eq + -- TODO(kmill): Why is this `p := 2` necessary? + exact pow_succ_padicValNat_not_dvd (p := 2) n.succ_ne_zero ∘ dvd_of_eq lemma Nat.max_log_padicValNat_succ_eq_log_succ (n : ℕ) : max (log 2 n) (padicValNat 2 (n + 1)) = log 2 (n + 1) := by diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index c5980b3b24eea..addb01df1228f 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -146,7 +146,7 @@ theorem zmod_congr_of_sub_mem_span_aux (n : ℕ) (x : ℤ_[p]) (a b : ℤ) rw [← dvd_neg, neg_sub] at ha have := dvd_add ha hb rwa [sub_eq_add_neg, sub_eq_add_neg, add_assoc, neg_add_cancel_left, ← sub_eq_add_neg, ← - Int.cast_sub, pow_p_dvd_int_iff, Nat.cast_pow] at this + Int.cast_sub, pow_p_dvd_int_iff] at this #align padic_int.zmod_congr_of_sub_mem_span_aux PadicInt.zmod_congr_of_sub_mem_span_aux theorem zmod_congr_of_sub_mem_span (n : ℕ) (x : ℤ_[p]) (a b : ℕ) @@ -488,7 +488,9 @@ def nthHom (r : R) : ℕ → ℤ := fun n => (f n r : ZMod (p ^ n)).val #align padic_int.nth_hom PadicInt.nthHom @[simp] -theorem nthHom_zero : nthHom f 0 = 0 := by simp [nthHom]; rfl +theorem nthHom_zero : nthHom f 0 = 0 := by + simp (config := { unfoldPartialApp := true }) [nthHom] + rfl #align padic_int.nth_hom_zero PadicInt.nthHom_zero variable {f} @@ -671,7 +673,7 @@ theorem ext_of_toZModPow {x y : ℤ_[p]} : (∀ n, toZModPow n x = toZModPow n y constructor · intro h rw [← lift_self x, ← lift_self y] - simp [lift, limNthHom, nthHom, h] + simp (config := { unfoldPartialApp := true }) [lift, limNthHom, nthHom, h] · rintro rfl _ rfl #align padic_int.ext_of_to_zmod_pow PadicInt.ext_of_toZModPow diff --git a/Mathlib/NumberTheory/Pell.lean b/Mathlib/NumberTheory/Pell.lean index a68df8142d5d3..bba6c74e643fb 100644 --- a/Mathlib/NumberTheory/Pell.lean +++ b/Mathlib/NumberTheory/Pell.lean @@ -551,7 +551,7 @@ theorem y_strictMono {a : Solution₁ d} (h : IsFundamental a) : add_pos_of_pos_of_nonneg (mul_pos (x_zpow_pos h.x_pos _) h.2.1) (mul_nonneg _ (by rw [sub_nonneg]; exact h.1.le)) rcases hn.eq_or_lt with (rfl | hn) - · simp only [zpow_zero, y_one] + · simp only [zpow_zero, y_one, le_refl] · exact (y_zpow_pos h.x_pos h.2.1 hn).le refine' strictMono_int_of_lt_succ fun n => _ cases' le_or_lt 0 n with hn hn diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index 78a8edff6b0cd..f108fc103fc1c 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -551,8 +551,8 @@ theorem yn_modEq_a_sub_one : ∀ n, yn a1 n ≡ n [MOD a - 1] #align pell.yn_modeq_a_sub_one Pell.yn_modEq_a_sub_one theorem yn_modEq_two : ∀ n, yn a1 n ≡ n [MOD 2] - | 0 => by simp - | 1 => by simp + | 0 => by rfl + | 1 => by simp; rfl | n + 2 => (yn_modEq_two n).add_right_cancel <| by rw [yn_succ_succ, mul_assoc, (by ring : n + 2 + n = 2 * (n + 1))] diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 33b39e9744483..e909fe28d2c6a 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -29,7 +29,7 @@ the bulk of the proof below. theorem sq_ne_two_fin_zmod_four (z : ZMod 4) : z * z ≠ 2 := by change Fin 4 at z - fin_cases z <;> norm_num [Fin.ext_iff] + fin_cases z <;> decide #align sq_ne_two_fin_zmod_four sq_ne_two_fin_zmod_four theorem Int.sq_ne_two_mod_four (z : ℤ) : z * z % 4 ≠ 2 := by @@ -155,7 +155,8 @@ theorem even_odd_of_coprime (hc : Int.gcd x y = 1) : rw [show z * z = 4 * (x0 * x0 + x0 + y0 * y0 + y0) + 2 by rw [← h.eq] ring] - norm_num [Int.add_emod] + simp only [Int.add_emod, Int.mul_emod_right, zero_add] + decide #align pythagorean_triple.even_odd_of_coprime PythagoreanTriple.even_odd_of_coprime theorem gcd_dvd : (Int.gcd x y : ℤ) ∣ z := by @@ -243,10 +244,10 @@ theorem isPrimitiveClassified_of_coprime_of_zero_left (hc : Int.gcd x y = 1) (hx cases' Int.natAbs_eq y with hy hy · use 1, 0 rw [hy, hc, Int.gcd_zero_right] - norm_num + decide · use 0, 1 rw [hy, hc, Int.gcd_zero_left] - norm_num + decide #align pythagorean_triple.is_primitive_classified_of_coprime_of_zero_left PythagoreanTriple.isPrimitiveClassified_of_coprime_of_zero_left theorem coprime_of_coprime (hc : Int.gcd x y = 1) : Int.gcd y z = 1 := by @@ -339,7 +340,9 @@ private theorem coprime_sq_sub_sq_add_of_even_odd {m n : ℤ} (h : Int.gcd m n = have hnc : p = 2 ∨ p ∣ Int.natAbs n := prime_two_or_dvd_of_dvd_two_mul_pow_self_two hp h2n by_cases h2 : p = 2 -- Porting note: norm_num is not enough to close h3 - · have h3 : (m ^ 2 + n ^ 2) % 2 = 1 := by field_simp [sq, Int.add_emod, Int.mul_emod, hm, hn] + · have h3 : (m ^ 2 + n ^ 2) % 2 = 1 := by + simp only [sq, Int.add_emod, Int.mul_emod, hm, hn, dvd_refl, Int.emod_emod_of_dvd] + decide have h4 : (m ^ 2 + n ^ 2) % 2 = 0 := by apply Int.emod_eq_zero_of_dvd rwa [h2] at hp2 @@ -372,7 +375,9 @@ private theorem coprime_sq_sub_mul_of_even_odd {m n : ℤ} (h : Int.gcd m n = 1) rw [hp2'] apply mt Int.emod_eq_zero_of_dvd -- Porting note: norm_num is not enough to close this - field_simp [sq, Int.sub_emod, Int.mul_emod, hm, hn] + simp only [sq, Nat.cast_ofNat, Int.sub_emod, Int.mul_emod, hm, hn, + mul_zero, EuclideanDomain.zero_mod, mul_one, zero_sub] + decide apply mt (Int.dvd_gcd (Int.coe_nat_dvd_left.mpr hpm)) hnp apply or_self_iff.mp apply Int.Prime.dvd_mul' hp @@ -528,7 +533,7 @@ theorem isPrimitiveClassified_of_coprime_of_odd_of_pos (hc : Int.gcd x y = 1) (h Int.dvd_gcd (Int.dvd_of_emod_eq_zero hn2) (Int.dvd_of_emod_eq_zero hm2) rw [hnmcp] at h1 revert h1 - norm_num + decide · -- m even, n odd apply h.isPrimitiveClassified_aux hc hzpos hm2n2 hv2 hw2 _ hmncp · apply Or.intro_left diff --git a/Mathlib/NumberTheory/Rayleigh.lean b/Mathlib/NumberTheory/Rayleigh.lean index 9b511a02cc709..6730de57895ac 100644 --- a/Mathlib/NumberTheory/Rayleigh.lean +++ b/Mathlib/NumberTheory/Rayleigh.lean @@ -117,8 +117,8 @@ theorem compl_beattySeq {r s : ℝ} (hrs : r.IsConjugateExponent s) : ext j by_cases h₁ : j ∈ {beattySeq r k | k} <;> by_cases h₂ : j ∈ {beattySeq' s k | k} · exact (Set.not_disjoint_iff.2 ⟨j, h₁, h₂⟩ (Beatty.no_collision hrs)).elim - · simp only [Set.mem_compl_iff, h₁, h₂] - · simp only [Set.mem_compl_iff, h₁, h₂] + · simp only [Set.mem_compl_iff, h₁, h₂, not_true_eq_false] + · simp only [Set.mem_compl_iff, h₁, h₂, not_false_eq_true] · have ⟨k, h₁₁, h₁₂⟩ := (Beatty.hit_or_miss hrs.pos).resolve_left h₁ have ⟨m, h₂₁, h₂₂⟩ := (Beatty.hit_or_miss' hrs.symm.pos).resolve_left h₂ exact (Beatty.no_anticollision hrs ⟨j, k, m, h₁₁, h₁₂, h₂₁, h₂₂⟩).elim diff --git a/Mathlib/NumberTheory/SmoothNumbers.lean b/Mathlib/NumberTheory/SmoothNumbers.lean index 0318cd35feb98..6f423703c156b 100644 --- a/Mathlib/NumberTheory/SmoothNumbers.lean +++ b/Mathlib/NumberTheory/SmoothNumbers.lean @@ -81,7 +81,8 @@ lemma smoothNumbers_succ {N : ℕ} (hN : ¬ N.Prime) : N.succ.smoothNumbers = N. fun hm ↦ ⟨hm.1, fun p hp ↦ (hm.2 p hp).trans <| lt.base N⟩⟩ exact lt_of_le_of_ne (lt_succ.mp <| hm.2 p hp) fun h ↦ hN <| h ▸ prime_of_mem_factors hp -@[simp] lemma smoothNumbers_one : smoothNumbers 1 = {1} := by simp [smoothNumbers_succ] +@[simp] lemma smoothNumbers_one : smoothNumbers 1 = {1} := by + simp (config := {decide := true}) [smoothNumbers_succ] @[gcongr] lemma smoothNumbers_mono {N M : ℕ} (hNM : N ≤ M) : N.smoothNumbers ⊆ M.smoothNumbers := fun _ hx ↦ ⟨hx.1, fun p hp => (hx.2 p hp).trans_le hNM⟩ diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index 05e44b1be541c..f03d351084a61 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -27,8 +27,6 @@ open Finset Polynomial FiniteField Equiv open scoped BigOperators -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- **Euler's four-square identity**. -/ theorem euler_four_squares {R : Type*} [CommRing R] (a b c d x y z w : R) : (a * x - b * y - c * z - d * w) ^ 2 + (a * y + b * x + c * w - d * z) ^ 2 + @@ -187,7 +185,7 @@ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : rcases (zero_le r).eq_or_gt with rfl | hr₀ · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0; · simpa [and_assoc] using hr obtain ⟨⟨a, rfl⟩, ⟨b, rfl⟩, ⟨c, rfl⟩, ⟨d, rfl⟩⟩ : m ∣ a ∧ m ∣ b ∧ m ∣ c ∧ m ∣ d - · simp only [← ZMod.nat_cast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero] + · simp only [← ZMod.nat_cast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, and_self] have : m * m ∣ m * p := habcd ▸ ⟨a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2, by ring⟩ rw [mul_dvd_mul_iff_left hm₀.ne'] at this exact (hp.eq_one_or_self_of_dvd _ this).elim hm₁.ne' hmp.ne diff --git a/Mathlib/NumberTheory/SumTwoSquares.lean b/Mathlib/NumberTheory/SumTwoSquares.lean index 4936e6c381629..4829dcc69b667 100644 --- a/Mathlib/NumberTheory/SumTwoSquares.lean +++ b/Mathlib/NumberTheory/SumTwoSquares.lean @@ -109,7 +109,7 @@ theorem ZMod.isSquare_neg_one_iff {n : ℕ} (hn : Squarefree n) : refine' ⟨fun H q hqp hqd => hqp.mod_four_ne_three_of_dvd_isSquare_neg_one hqd H, fun H => _⟩ induction' n using induction_on_primes with p n hpp ih · exact False.elim (hn.ne_zero rfl) - · exact ⟨0, by simp only [Fin.zero_mul, neg_eq_zero, Fin.one_eq_zero_iff]⟩ + · exact ⟨0, by simp only [mul_zero, eq_iff_true_of_subsingleton]⟩ · haveI : Fact p.Prime := ⟨hpp⟩ have hcp : p.Coprime n := by by_contra hc diff --git a/Mathlib/NumberTheory/ZetaFunction.lean b/Mathlib/NumberTheory/ZetaFunction.lean index d0b6784d140eb..b792119a64068 100644 --- a/Mathlib/NumberTheory/ZetaFunction.lean +++ b/Mathlib/NumberTheory/ZetaFunction.lean @@ -68,8 +68,6 @@ noncomputable section ## Definition of the Riemann zeta function and related functions -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- Function whose Mellin transform is `π ^ (-s) * Γ(s) * zeta (2 * s)`, for `1 / 2 < Re s`. -/ def zetaKernel₁ (t : ℝ) : ℂ := ∑' n : ℕ, rexp (-π * t * ((n : ℝ) + 1) ^ 2) @@ -493,7 +491,7 @@ theorem integral_cpow_mul_exp_neg_pi_mul_sq {s : ℂ} (hs : 0 < s.re) (n : ℕ) conv_rhs => congr rw [← smul_eq_mul, ← mellin_comp_mul_left _ _ pi_pos] - have : 1 / ((n : ℂ) + 1) ^ (2 * s) = (((n : ℝ) + 1) ^ (2 : ℝ) : ℂ) ^ (-s) := by + have : 1 / ((n : ℂ) + 1) ^ (2 * s) = (((n + 1) ^ (2 : ℝ) : ℝ) : ℂ) ^ (-s) := by rw [(by norm_num : (n : ℂ) + 1 = ↑((n : ℝ) + 1)), (by norm_num : 2 * s = ↑(2 : ℝ) * s), cpow_mul_ofReal_nonneg, one_div, cpow_neg] rw [← Nat.cast_succ] diff --git a/Mathlib/NumberTheory/ZetaValues.lean b/Mathlib/NumberTheory/ZetaValues.lean index e5ecafd7639ed..086b0dc4ea3dd 100644 --- a/Mathlib/NumberTheory/ZetaValues.lean +++ b/Mathlib/NumberTheory/ZetaValues.lean @@ -28,8 +28,6 @@ zeta functions, in terms of Bernoulli polynomials. an explicit multiple of `Bₖ(x)`, for any `x ∈ [0, 1]` and `k ≥ 3` odd. -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - noncomputable section open scoped Nat Real Interval diff --git a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean index 93e1830904d1b..403f40dba084b 100644 --- a/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean +++ b/Mathlib/NumberTheory/Zsqrtd/QuadraticReciprocity.lean @@ -23,8 +23,6 @@ open Zsqrtd Complex open scoped ComplexConjugate -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - local notation "ℤ[i]" => GaussianInt namespace GaussianInt @@ -89,7 +87,7 @@ theorem prime_of_nat_prime_of_mod_four_eq_three (p : ℕ) [hp : Fact p.Prime] (h irreducible_iff_prime.1 <| by_contradiction fun hpi => let ⟨a, b, hab⟩ := sq_add_sq_of_nat_prime_of_not_irreducible p hpi - have : ∀ a b : ZMod 4, a ^ 2 + b ^ 2 ≠ p := by + have : ∀ a b : ZMod 4, a ^ 2 + b ^ 2 ≠ ↑p := by erw [← ZMod.nat_cast_mod p 4, hp3]; exact by decide this a b (hab ▸ by simp) #align gaussian_int.prime_of_nat_prime_of_mod_four_eq_three GaussianInt.prime_of_nat_prime_of_mod_four_eq_three diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index db241ed968de9..f831f2f24efc2 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -2058,7 +2058,8 @@ theorem Monotone.piecewise_eventually_eq_iUnion {β : α → Type*} [Preorder ι · refine (eventually_ge_atTop i).mono fun j hij ↦ ?_ simp only [Set.piecewise_eq_of_mem, hs hij hi, subset_iUnion _ _ hi] · refine eventually_of_forall fun i ↦ ?_ - simp only [Set.piecewise_eq_of_not_mem, not_exists.1 ha i, mt mem_iUnion.1 ha] + simp only [Set.piecewise_eq_of_not_mem, not_exists.1 ha i, mt mem_iUnion.1 ha, + not_false_eq_true, exists_false] theorem Antitone.piecewise_eventually_eq_iInter {β : α → Type*} [Preorder ι] {s : ι → Set α} [∀ i, DecidablePred (· ∈ s i)] [DecidablePred (· ∈ ⋂ i, s i)] diff --git a/Mathlib/Order/Filter/ENNReal.lean b/Mathlib/Order/Filter/ENNReal.lean index 82cf5da48a02f..b43fbd981c410 100644 --- a/Mathlib/Order/Filter/ENNReal.lean +++ b/Mathlib/Order/Filter/ENNReal.lean @@ -65,7 +65,7 @@ theorem limsup_const_mul [CountableInterFilter f] {u : α → ℝ≥0∞} {a : have h_top_le : (f.limsup fun x : α => ite (u x = 0) (0 : ℝ≥0∞) ⊤) = ⊤ := eq_top_iff.mpr (le_limsup_of_frequently_le hu_mul) have hfu : f.limsup u ≠ 0 := mt limsup_eq_zero_iff.1 hu - simp only [ha_top, top_mul', hfu, h_top_le] + simp only [ha_top, top_mul', h_top_le, hfu, ite_false] #align ennreal.limsup_const_mul ENNReal.limsup_const_mul theorem limsup_mul_le [CountableInterFilter f] (u v : α → ℝ≥0∞) : diff --git a/Mathlib/Order/Filter/IndicatorFunction.lean b/Mathlib/Order/Filter/IndicatorFunction.lean index 5e405dfff62a7..e3b129df67696 100644 --- a/Mathlib/Order/Filter/IndicatorFunction.lean +++ b/Mathlib/Order/Filter/IndicatorFunction.lean @@ -31,7 +31,7 @@ theorem mulIndicator_eventuallyEq (hf : f =ᶠ[l ⊓ 𝓟 s] g) (hs : s =ᶠ[l] (eventually_inf_principal.1 hf).mp <| hs.mem_iff.mono fun x hst hfg => by_cases (fun hxs : x ∈ s => by simp only [*, hst.1 hxs, mulIndicator_of_mem]) - (fun hxs => by simp only [mulIndicator_of_not_mem, hxs, mt hst.2 hxs]) + (fun hxs => by simp only [mulIndicator_of_not_mem, hxs, mt hst.2 hxs, not_false_eq_true]) #align indicator_eventually_eq indicator_eventuallyEq end One diff --git a/Mathlib/Order/LocallyFinite.lean b/Mathlib/Order/LocallyFinite.lean index 6912e8eeb1b6e..8156f0d79f8f5 100644 --- a/Mathlib/Order/LocallyFinite.lean +++ b/Mathlib/Order/LocallyFinite.lean @@ -1081,8 +1081,8 @@ instance locallyFiniteOrder : LocallyFiniteOrder (WithTop α) where rw [some_mem_insertNone] simp | (a : α), (b : α), ⊤ => by - simp only [some, le_eq_subset, mem_map, mem_Icc, le_top, top_le_iff, and_false, iff_false, - not_exists, not_and, and_imp, Embedding.some, forall_const] + simp only [Embedding.some, mem_map, mem_Icc, and_false, exists_const, some, le_top, + top_le_iff] | (a : α), (b : α), (x : α) => by simp only [some, le_eq_subset, Embedding.some, mem_map, mem_Icc, Embedding.coeFn_mk, some_le_some] diff --git a/Mathlib/Probability/Distributions/Gaussian.lean b/Mathlib/Probability/Distributions/Gaussian.lean index 64d043b893cc9..672395b66b59d 100644 --- a/Mathlib/Probability/Distributions/Gaussian.lean +++ b/Mathlib/Probability/Distributions/Gaussian.lean @@ -34,8 +34,6 @@ open scoped ENNReal NNReal Real open MeasureTheory -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace ProbabilityTheory section GaussianPdf @@ -285,7 +283,7 @@ lemma gaussianReal_map_const_mul (c : ℝ) : rw [Measure.map_const] simp only [ne_eq, measure_univ, one_smul, mul_eq_zero] convert (gaussianReal_zero_var 0).symm - simp only [ne_eq, zero_pow', mul_eq_zero, hv, or_false] + simp only [ne_eq, zero_pow', mul_eq_zero, hv, or_false, not_false_eq_true] rfl let e : ℝ ≃ᵐ ℝ := (Homeomorph.mulLeft₀ c hc).symm.toMeasurableEquiv have he' : ∀ x, HasDerivAt e ((fun _ ↦ c⁻¹) x) x := by diff --git a/Mathlib/Probability/Integration.lean b/Mathlib/Probability/Integration.lean index 10faa875ee53e..164def57b5de1 100644 --- a/Mathlib/Probability/Integration.lean +++ b/Mathlib/Probability/Integration.lean @@ -172,7 +172,7 @@ theorem IndepFun.integrable_left_of_integrable_mul {β : Type*} [MeasurableSpace have A : (∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ) < ∞ := h'XY.2 simp only [nnnorm_mul, ENNReal.coe_mul] at A rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.ennnorm hY.ennnorm J, H] at A - simp only [ENNReal.top_mul I] at A + simp only [ENNReal.top_mul I, lt_self_iff_false] at A #align probability_theory.indep_fun.integrable_left_of_integrable_mul ProbabilityTheory.IndepFun.integrable_left_of_integrable_mul /-- If the product of two independent real-valued random variables is integrable and the @@ -194,7 +194,7 @@ theorem IndepFun.integrable_right_of_integrable_mul {β : Type*} [MeasurableSpac have A : (∫⁻ ω, ‖X ω * Y ω‖₊ ∂μ) < ∞ := h'XY.2 simp only [nnnorm_mul, ENNReal.coe_mul] at A rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indepFun'' hX.ennnorm hY.ennnorm J, H] at A - simp only [ENNReal.mul_top I] at A + simp only [ENNReal.mul_top I, lt_self_iff_false] at A #align probability_theory.indep_fun.integrable_right_of_integrable_mul ProbabilityTheory.IndepFun.integrable_right_of_integrable_mul /-- The (Bochner) integral of the product of two independent, nonnegative random diff --git a/Mathlib/Probability/Kernel/Composition.lean b/Mathlib/Probability/Kernel/Composition.lean index 27ef08a67dd46..94c70d3aa13eb 100644 --- a/Mathlib/Probability/Kernel/Composition.lean +++ b/Mathlib/Probability/Kernel/Composition.lean @@ -405,8 +405,9 @@ theorem lintegral_compProd' (κ : kernel α β) [IsSFiniteKernel κ] (η : kerne ∫⁻ (a_1 : β), ∫⁻ (c : γ), f (a_1, c) ∂η (a, a_1) ∂κ a)) _ _ (F n) · intro c s hs classical -- Porting note: Added `classical` for `Set.piecewise_eq_indicator` - simp only [SimpleFunc.const_zero, SimpleFunc.coe_piecewise, SimpleFunc.coe_const, - SimpleFunc.coe_zero, Set.piecewise_eq_indicator, Function.const, lintegral_indicator_const hs] + simp (config := { unfoldPartialApp := true }) only [SimpleFunc.const_zero, + SimpleFunc.coe_piecewise, SimpleFunc.coe_const, SimpleFunc.coe_zero, + Set.piecewise_eq_indicator, Function.const, lintegral_indicator_const hs] rw [compProd_apply κ η _ hs, ← lintegral_const_mul c _] swap · exact (measurable_kernel_prod_mk_left ((measurable_fst.snd.prod_mk measurable_snd) hs)).comp diff --git a/Mathlib/Probability/Martingale/BorelCantelli.lean b/Mathlib/Probability/Martingale/BorelCantelli.lean index 360faf71b2f4b..1994fe0bc7c6a 100644 --- a/Mathlib/Probability/Martingale/BorelCantelli.lean +++ b/Mathlib/Probability/Martingale/BorelCantelli.lean @@ -89,7 +89,9 @@ theorem leastGE_eq_min (π : Ω → ℕ) (r : ℝ) (ω : Ω) {n : ℕ} (hπn : theorem stoppedValue_stoppedValue_leastGE (f : ℕ → Ω → ℝ) (π : Ω → ℕ) (r : ℝ) {n : ℕ} (hπn : ∀ ω, π ω ≤ n) : stoppedValue (fun i => stoppedValue f (leastGE f r i)) π = stoppedValue (stoppedProcess f (leastGE f r n)) π := by - ext1 ω; simp_rw [stoppedProcess, stoppedValue]; rw [leastGE_eq_min _ _ _ hπn] + ext1 ω + simp (config := { unfoldPartialApp := true }) only [stoppedProcess, stoppedValue] + rw [leastGE_eq_min _ _ _ hπn] #align measure_theory.stopped_value_stopped_value_least_ge MeasureTheory.stoppedValue_stoppedValue_leastGE theorem Submartingale.stoppedValue_leastGE [IsFiniteMeasure μ] (hf : Submartingale f ℱ μ) (r : ℝ) : diff --git a/Mathlib/Probability/Martingale/Convergence.lean b/Mathlib/Probability/Martingale/Convergence.lean index 63ea1272dbcee..d4383382f162c 100644 --- a/Mathlib/Probability/Martingale/Convergence.lean +++ b/Mathlib/Probability/Martingale/Convergence.lean @@ -399,7 +399,7 @@ theorem Integrable.tendsto_ae_condexp (hg : Integrable g μ) by_cases hnm : n ≤ m · exact ⟨m, (ℱ.mono hnm _ hs).inter ht⟩ · exact ⟨n, hs.inter (ℱ.mono (not_le.1 hnm).le _ ht)⟩ - · simp only [measure_empty, WithTop.zero_lt_top, Measure.restrict_empty, integral_zero_measure, + · simp only [measure_empty, ENNReal.zero_lt_top, Measure.restrict_empty, integral_zero_measure, forall_true_left] · rintro t ⟨n, ht⟩ - exact this n _ ht diff --git a/Mathlib/Probability/Martingale/Upcrossing.lean b/Mathlib/Probability/Martingale/Upcrossing.lean index 5f45784e6c532..1adef8da72489 100644 --- a/Mathlib/Probability/Martingale/Upcrossing.lean +++ b/Mathlib/Probability/Martingale/Upcrossing.lean @@ -683,7 +683,8 @@ theorem crossing_pos_eq (hab : a < b) : rw [LatticeOrderedGroup.pos_nonpos_iff, sub_nonpos] induction' n with k ih · refine' ⟨rfl, _⟩ - simp only [lowerCrossingTime_zero, hitting, Set.mem_Icc, Set.mem_Iic, Nat.zero_eq] + simp (config := { unfoldPartialApp := true }) only [lowerCrossingTime_zero, hitting, + Set.mem_Icc, Set.mem_Iic, Nat.zero_eq] ext ω split_ifs with h₁ h₂ h₂ · simp_rw [hf'] @@ -699,7 +700,7 @@ theorem crossing_pos_eq (hab : a < b) : split_ifs with h₁ h₂ h₂ · simp_rw [← sub_le_iff_le_add, hf ω] · refine' False.elim (h₂ _) - simp_all only [Set.mem_Ici] + simp_all only [Set.mem_Ici, not_true_eq_false] · refine' False.elim (h₁ _) simp_all only [Set.mem_Ici] · rfl @@ -709,7 +710,7 @@ theorem crossing_pos_eq (hab : a < b) : split_ifs with h₁ h₂ h₂ · simp_rw [hf' ω] · refine' False.elim (h₂ _) - simp_all only [Set.mem_Iic] + simp_all only [Set.mem_Iic, not_true_eq_false] · refine' False.elim (h₁ _) simp_all only [Set.mem_Iic] · rfl diff --git a/Mathlib/Probability/Notation.lean b/Mathlib/Probability/Notation.lean index 63beed340bc7b..596ebcc961833 100644 --- a/Mathlib/Probability/Notation.lean +++ b/Mathlib/Probability/Notation.lean @@ -38,10 +38,11 @@ open scoped MeasureTheory scoped[ProbabilityTheory] notation "𝔼[" X "|" m "]" => MeasureTheory.condexp m MeasureTheory.MeasureSpace.volume X +-- Note(kmill): this notation tends to lead to ambiguity with GetElem notation. set_option quotPrecheck false in scoped[ProbabilityTheory] notation P "[" X "]" => ∫ x, ↑(X x) ∂P -scoped[ProbabilityTheory] notation "𝔼[" X "]" => ∫ a, X a +scoped[ProbabilityTheory] notation "𝔼[" X "]" => ∫ a, (X : _ → _) a scoped[ProbabilityTheory] notation P "⟦" s "|" m "⟧" => MeasureTheory.condexp m P (Set.indicator s fun ω => (1 : ℝ)) diff --git a/Mathlib/Probability/Process/Stopping.lean b/Mathlib/Probability/Process/Stopping.lean index 00a5c033f3822..02a643943ff06 100644 --- a/Mathlib/Probability/Process/Stopping.lean +++ b/Mathlib/Probability/Process/Stopping.lean @@ -832,7 +832,7 @@ theorem progMeasurable_min_stopping_time [MetrizableSpace ι] (hτ : IsStoppingT suffices h_min_eq_left : (fun x : sc => min (↑(x : Set.Iic i × Ω).fst) (τ (x : Set.Iic i × Ω).snd)) = fun x : sc => ↑(x : Set.Iic i × Ω).fst - · simp_rw [Set.restrict, h_min_eq_left] + · simp (config := { unfoldPartialApp := true }) only [Set.restrict, h_min_eq_left] exact h_meas_fst _ ext1 ω rw [min_eq_left] diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 57580ccf3cf08..2cd01ce9b2b7b 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -57,8 +57,6 @@ random variables. Let `Yₙ` be the truncation of `Xₙ` up to `n`. We claim tha noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open MeasureTheory Filter Finset Asymptotics open Set (indicator) @@ -366,7 +364,7 @@ theorem sum_variance_truncation_le {X : Ω → ℝ} (hint : Integrable X) (hnonn simp only [hij, mem_sigma, mem_range, and_self_iff] · rintro ⟨i, j⟩ hij; rfl · rintro ⟨i, j⟩ hij; rfl - _ ≤ ∑ k in range K, ↑2 / (k + ↑1) * ∫ x in k..(k + 1 : ℕ), x ^ 2 ∂ρ := by + _ ≤ ∑ k in range K, 2 / (k + 1 : ℝ) * ∫ x in k..(k + 1 : ℕ), x ^ 2 ∂ρ := by apply sum_le_sum fun k _ => ?_ refine' mul_le_mul_of_nonneg_right (sum_Ioo_inv_sq_le _ _) _ refine' intervalIntegral.integral_nonneg_of_forall _ fun u => sq_nonneg _ diff --git a/Mathlib/Probability/Variance.lean b/Mathlib/Probability/Variance.lean index 278a22de2e5e8..82f515535cfe0 100644 --- a/Mathlib/Probability/Variance.lean +++ b/Mathlib/Probability/Variance.lean @@ -43,8 +43,6 @@ open scoped BigOperators MeasureTheory ProbabilityTheory ENNReal NNReal namespace ProbabilityTheory -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - -- Porting note: this lemma replaces `ENNReal.toReal_bit0`, which does not exist in Lean 4 private lemma coe_two : ENNReal.toReal 2 = (2 : ℝ) := rfl @@ -116,7 +114,7 @@ theorem evariance_eq_lintegral_ofReal (X : Ω → ℝ) (μ : Measure Ω) : #align probability_theory.evariance_eq_lintegral_of_real ProbabilityTheory.evariance_eq_lintegral_ofReal theorem _root_.MeasureTheory.Memℒp.variance_eq_of_integral_eq_zero (hX : Memℒp X 2 μ) - (hXint : μ[X] = 0) : variance X μ = μ[X ^ 2] := by + (hXint : μ[X] = 0) : variance X μ = μ[X ^ (2 : Nat)] := by rw [variance, evariance_eq_lintegral_ofReal, ← ofReal_integral_eq_lintegral_ofReal, ENNReal.toReal_ofReal] <;> simp_rw [hXint, sub_zero] @@ -129,7 +127,7 @@ theorem _root_.MeasureTheory.Memℒp.variance_eq_of_integral_eq_zero (hX : Mem #align measure_theory.mem_ℒp.variance_eq_of_integral_eq_zero MeasureTheory.Memℒp.variance_eq_of_integral_eq_zero theorem _root_.MeasureTheory.Memℒp.variance_eq [IsFiniteMeasure μ] (hX : Memℒp X 2 μ) : - variance X μ = μ[(X - fun _ => μ[X]) ^ 2] := by + variance X μ = μ[(X - fun _ => μ[X] :) ^ (2 : Nat)] := by rw [variance, evariance_eq_lintegral_ofReal, ← ofReal_integral_eq_lintegral_ofReal, ENNReal.toReal_ofReal] · rfl @@ -243,7 +241,7 @@ theorem variance_le_expectation_sq [@IsProbabilityMeasure Ω _ ℙ] {X : Ω → #align probability_theory.variance_le_expectation_sq ProbabilityTheory.variance_le_expectation_sq theorem evariance_def' [@IsProbabilityMeasure Ω _ ℙ] {X : Ω → ℝ} (hX : AEStronglyMeasurable X ℙ) : - eVar[X] = (∫⁻ ω, ‖X ω‖₊ ^ 2) - ENNReal.ofReal (𝔼[X] ^ 2) := by + eVar[X] = (∫⁻ ω, (‖X ω‖₊ ^ 2 :)) - ENNReal.ofReal (𝔼[X] ^ 2) := by by_cases hℒ : Memℒp X 2 · rw [← hℒ.ofReal_variance_eq, variance_def' hℒ, ENNReal.ofReal_sub _ (sq_nonneg _)] congr @@ -287,7 +285,7 @@ theorem meas_ge_le_variance_div_sq [@IsFiniteMeasure Ω _ ℙ] {X : Ω → ℝ} rw [ENNReal.ofReal_div_of_pos (sq_pos_of_ne_zero _ hc.ne.symm), hX.ofReal_variance_eq] convert @meas_ge_le_evariance_div_sq _ _ _ hX.1 c.toNNReal (by simp [hc]) using 1 · simp only [Real.coe_toNNReal', max_le_iff, abs_nonneg, and_true_iff] - · rw [ENNReal.ofReal_pow hc.le, ENNReal.coe_pow] + · rw [ENNReal.ofReal_pow hc.le] rfl #align probability_theory.meas_ge_le_variance_div_sq ProbabilityTheory.meas_ge_le_variance_div_sq diff --git a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean index 3c9812d907b17..0b51992e29c79 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/Resolution.lean @@ -640,7 +640,8 @@ theorem forget₂ToModuleCatHomotopyEquiv_f_0_eq : rfl · congr · ext x - dsimp [HomotopyEquiv.ofIso, Finsupp.LinearEquiv.finsuppUnique] + dsimp (config := { unfoldPartialApp := true }) [HomotopyEquiv.ofIso, + Finsupp.LinearEquiv.finsuppUnique] rw [Finsupp.total_single, one_smul, @Unique.eq_default _ Types.terminalIso.toEquiv.unique x, ChainComplex.single₀_map_f_zero, LinearMap.coe_mk, AddHom.coe_mk, Function.comp_apply, Finsupp.equivFunOnFinite_apply, Finsupp.single_eq_same] diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index 08330c9a3773d..25b8320fc8529 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -342,7 +342,7 @@ instance : SMul α (HomogeneousLocalization 𝒜 x) where smul m := Quotient.map' (m • ·) fun c1 c2 (h : Localization.mk _ _ = Localization.mk _ _) => by change Localization.mk _ _ = Localization.mk _ _ simp only [num_smul, den_smul] - convert congr_arg (fun z : at x => m • z) h <;> rw [Localization.smul_mk] <;> rfl + convert congr_arg (fun z : at x => m • z) h <;> rw [Localization.smul_mk] @[simp] theorem smul_val (y : HomogeneousLocalization 𝒜 x) (n : α) : (n • y).val = n • y.val := by diff --git a/Mathlib/RingTheory/LocalProperties.lean b/Mathlib/RingTheory/LocalProperties.lean index f14c4681546fb..3863cf39aa40c 100644 --- a/Mathlib/RingTheory/LocalProperties.lean +++ b/Mathlib/RingTheory/LocalProperties.lean @@ -38,8 +38,6 @@ The following properties are covered: -/ -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - open scoped Pointwise Classical BigOperators universe u diff --git a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean index 0a7d5ac0d0261..f93639b52657e 100644 --- a/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/NewtonIdentities.lean @@ -93,6 +93,7 @@ private theorem pairMap_mem_pairs {k : ℕ} (t : Finset σ × σ) (h : t ∈ pai simp only [card_erase_of_mem h1, tsub_le_iff_right, mem_erase, ne_eq, h1] refine ⟨le_step h, ?_⟩ by_contra h2 + simp only [not_true_eq_false, and_true, not_forall, not_false_eq_true, exists_prop] at h2 rw [← h2] at h exact not_le_of_lt (sub_lt (card_pos.mpr ⟨t.snd, h1⟩) zero_lt_one) h · rw [pairMap_of_snd_nmem_fst σ h1] @@ -184,8 +185,9 @@ private theorem sum_filter_pairs_eq_sum_filter_antidiagonal_powersetCard_sum (k have hnk' : n.fst ≤ k := by apply le_of_lt; aesop aesop · simp_all only [mem_antidiagonal, mem_filter, mem_pairs, disjiUnion_eq_biUnion, - add_tsub_cancel_of_le] - · simp_all only [mem_antidiagonal, mem_filter, mem_pairs, disjiUnion_eq_biUnion, implies_true] + add_tsub_cancel_of_le, and_true] + · simp_all only [mem_antidiagonal, mem_filter, mem_pairs, disjiUnion_eq_biUnion, + implies_true, and_true] simp only [← hdisj, disj_equiv] private theorem disjoint_filter_pairs_lt_filter_pairs_eq (k : ℕ) : diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean index c2528bfbaa470..53eaf0060c3c2 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Basic.lean @@ -339,7 +339,7 @@ theorem degree_cyclotomic (n : ℕ) (R : Type*) [Ring R] [Nontrivial R] : rw [← map_cyclotomic_int] rw [degree_map_eq_of_leadingCoeff_ne_zero (Int.castRingHom R) _] · cases' n with k - · simp only [cyclotomic, degree_one, dif_pos, Nat.totient_zero, WithTop.coe_zero] + · simp only [cyclotomic, degree_one, dif_pos, Nat.totient_zero, CharP.cast_eq_zero] rw [← degree_cyclotomic' (Complex.isPrimitiveRoot_exp k.succ (Nat.succ_ne_zero k))] exact (int_cyclotomic_spec k.succ).2.1 simp only [(int_cyclotomic_spec n).right.right, eq_intCast, Monic.leadingCoeff, Int.cast_one, diff --git a/Mathlib/RingTheory/Polynomial/Selmer.lean b/Mathlib/RingTheory/Polynomial/Selmer.lean index 7d54d57060f47..7a6453a002f82 100644 --- a/Mathlib/RingTheory/Polynomial/Selmer.lean +++ b/Mathlib/RingTheory/Polynomial/Selmer.lean @@ -26,8 +26,6 @@ namespace Polynomial open scoped Polynomial -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - variable {n : ℕ} theorem X_pow_sub_X_sub_one_irreducible_aux (z : ℂ) : ¬(z ^ n = z + 1 ∧ z ^ n + z ^ 2 = 0) := by diff --git a/Mathlib/RingTheory/RootsOfUnity/Basic.lean b/Mathlib/RingTheory/RootsOfUnity/Basic.lean index ff5f7d134c8f2..3ff13c4853b6d 100644 --- a/Mathlib/RingTheory/RootsOfUnity/Basic.lean +++ b/Mathlib/RingTheory/RootsOfUnity/Basic.lean @@ -129,7 +129,7 @@ theorem map_rootsOfUnity (f : Mˣ →* Nˣ) (k : ℕ+) : (rootsOfUnity k M).map @[norm_cast] theorem rootsOfUnity.coe_pow [CommMonoid R] (ζ : rootsOfUnity k R) (m : ℕ) : - ((ζ ^ m : Rˣ) : R) = ((ζ : Rˣ) : R) ^ m := by + (((ζ ^ m :) : Rˣ) : R) = ((ζ : Rˣ) : R) ^ m := by rw [Subgroup.coe_pow, Units.val_pow_eq_pow_val] #align roots_of_unity.coe_pow rootsOfUnity.coe_pow @@ -265,8 +265,6 @@ section Reduced variable (R) [CommRing R] [IsReduced R] -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - -- @[simp] -- Porting note: simp normal form is `mem_rootsOfUnity_prime_pow_mul_iff'` theorem mem_rootsOfUnity_prime_pow_mul_iff (p k : ℕ) (m : ℕ+) [hp : Fact p.Prime] [CharP R p] {ζ : Rˣ} : ζ ∈ rootsOfUnity (⟨p, hp.1.pos⟩ ^ k * m) R ↔ ζ ∈ rootsOfUnity m R := by diff --git a/Mathlib/RingTheory/WittVector/Basic.lean b/Mathlib/RingTheory/WittVector/Basic.lean index ccd6aed2eeaf9..b6b4eee3c8396 100644 --- a/Mathlib/RingTheory/WittVector/Basic.lean +++ b/Mathlib/RingTheory/WittVector/Basic.lean @@ -166,7 +166,8 @@ elab "ghost_fun_tac" φ:term "," fn:term : tactic => do simp only [wittZero, OfNat.ofNat, Zero.zero, wittOne, One.one, HAdd.hAdd, Add.add, HSub.hSub, Sub.sub, Neg.neg, HMul.hMul, Mul.mul,HPow.hPow, Pow.pow, wittNSMul, wittZSMul, HSMul.hSMul, SMul.smul] - simpa [WittVector.ghostFun, aeval_rename, aeval_bind₁, comp, uncurry, peval, eval] using this + simpa (config := { unfoldPartialApp := true }) [WittVector.ghostFun, aeval_rename, aeval_bind₁, + comp, uncurry, peval, eval] using this ))) end Tactic @@ -236,7 +237,8 @@ private def ghostEquiv' [Invertible (p : R)] : 𝕎 R ≃ (ℕ → R) where ext n have := bind₁_wittPolynomial_xInTermsOfW p R n apply_fun aeval x.coeff at this - simpa only [aeval_bind₁, aeval_X, ghostFun, aeval_wittPolynomial] + simpa (config := { unfoldPartialApp := true }) only [aeval_bind₁, aeval_X, ghostFun, + aeval_wittPolynomial] right_inv := by intro x ext n diff --git a/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean b/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean index 511e717d8c5ac..554e2051aead4 100644 --- a/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean +++ b/Mathlib/RingTheory/WittVector/DiscreteValuationRing.lean @@ -30,8 +30,6 @@ When `k` is also a field, this `b` can be chosen to be a unit of `𝕎 k`. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace WittVector variable {p : ℕ} [hp : Fact p.Prime] diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index 3c6fbd43a22b6..f1d6de691e67f 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -190,7 +190,7 @@ theorem map_frobeniusPoly (n : ℕ) : have aux : ∀ k : ℕ, (p : ℚ)^ k ≠ 0 := by intro; apply pow_ne_zero; exact_mod_cast hp.1.ne_zero simpa [aux, -one_div, field_simps] using this.symm - rw [mul_comm _ (p : ℚ), mul_assoc, Nat.cast_pow, mul_assoc, ← pow_add, + rw [mul_comm _ (p : ℚ), mul_assoc, mul_assoc, ← pow_add, map_frobeniusPoly.key₂ p hi.le hj, Nat.cast_mul, Nat.cast_pow] ring #align witt_vector.map_frobenius_poly WittVector.map_frobeniusPoly diff --git a/Mathlib/RingTheory/WittVector/IsPoly.lean b/Mathlib/RingTheory/WittVector/IsPoly.lean index 0d70990e0a39a..1d46e11ae9613 100644 --- a/Mathlib/RingTheory/WittVector/IsPoly.lean +++ b/Mathlib/RingTheory/WittVector/IsPoly.lean @@ -240,7 +240,8 @@ instance IsPoly₂.comp {h f g} [hh : IsPoly₂ p h] [hf : IsPoly p f] [hg : IsP fun k ↦ rename (Prod.mk (1 : Fin 2)) (ψ k)]) (χ n), _⟩⟩ intros funext n - simp only [peval, aeval_bind₁, Function.comp, hh, hf, hg, uncurry] + simp (config := { unfoldPartialApp := true }) only [peval, aeval_bind₁, Function.comp, hh, hf, hg, + uncurry] apply eval₂Hom_congr rfl _ rfl ext ⟨i, n⟩ fin_cases i <;> @@ -271,7 +272,7 @@ instance IsPoly₂.diag {f} [hf : IsPoly₂ p f] : IsPoly p fun R _Rcr x => f x obtain ⟨φ, hf⟩ := hf refine' ⟨⟨fun n => bind₁ (uncurry ![X, X]) (φ n), _⟩⟩ intros; funext n - simp only [hf, peval, uncurry, aeval_bind₁] + simp (config := { unfoldPartialApp := true }) only [hf, peval, uncurry, aeval_bind₁] apply eval₂Hom_congr rfl _ rfl ext ⟨i, k⟩; fin_cases i <;> @@ -328,8 +329,8 @@ theorem bind₁_onePoly_wittPolynomial [hp : Fact p.Prime] (n : ℕ) : AlgHom.map_pow, bind₁_X_right, AlgHom.map_mul] · rw [Finset.mem_range] -- porting note: was `decide` - simp only [add_pos_iff, or_true, not_true, pow_zero, map_one, ge_iff_le, nonpos_iff_eq_zero, - tsub_zero, one_mul, gt_iff_lt, IsEmpty.forall_iff] + intro h + simp only [add_pos_iff, zero_lt_one, or_true, not_true_eq_false] at h #align witt_vector.bind₁_one_poly_witt_polynomial WittVector.bind₁_onePoly_wittPolynomial /-- The function that is constantly one on Witt vectors is a polynomial function. -/ @@ -434,7 +435,7 @@ theorem map [Fact p.Prime] {f} (hf : IsPoly₂ p f) (g : R →+* S) (x y : 𝕎 -- so that applications do not have to worry about the universe issue obtain ⟨φ, hf⟩ := hf ext n - simp only [map_coeff, hf, map_aeval, peval, uncurry] + simp (config := { unfoldPartialApp := true }) only [map_coeff, hf, map_aeval, peval, uncurry] apply eval₂Hom_congr (RingHom.ext_int _ _) _ rfl try ext ⟨i, k⟩; fin_cases i all_goals simp only [map_coeff, Matrix.cons_val_zero, Matrix.head_cons, Matrix.cons_val_one] diff --git a/Mathlib/RingTheory/WittVector/MulCoeff.lean b/Mathlib/RingTheory/WittVector/MulCoeff.lean index 9300e40ee7489..7a893e71b3818 100644 --- a/Mathlib/RingTheory/WittVector/MulCoeff.lean +++ b/Mathlib/RingTheory/WittVector/MulCoeff.lean @@ -30,8 +30,6 @@ that needs to happen in characteristic 0. noncomputable section -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - namespace WittVector variable (p : ℕ) [hp : Fact p.Prime] @@ -243,7 +241,8 @@ theorem peval_polyOfInterest' (n : ℕ) (x y : 𝕎 k) : x.coeff (n + 1) * y.coeff 0 ^ p ^ (n + 1) := by rw [peval_polyOfInterest] have : (p : k) = 0 := CharP.cast_eq_zero k p - simp only [this, Nat.cast_pow, ne_eq, add_eq_zero, and_false, zero_pow', zero_mul, add_zero] + simp only [this, Nat.cast_pow, ne_eq, add_eq_zero, and_false, zero_pow', zero_mul, add_zero, + not_false_eq_true] have sum_zero_pow_mul_pow_p : ∀ y : 𝕎 k, ∑ x : ℕ in range (n + 1 + 1), (0 : k) ^ x * y.coeff x ^ p ^ (n + 1 - x) = y.coeff 0 ^ p ^ (n + 1) := by intro y diff --git a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean index cee7b2b2a2c88..074909e3e71ca 100644 --- a/Mathlib/RingTheory/WittVector/StructurePolynomial.lean +++ b/Mathlib/RingTheory/WittVector/StructurePolynomial.lean @@ -232,7 +232,7 @@ theorem C_p_pow_dvd_bind₁_rename_wittPolynomial_sub_sum (Φ : MvPolynomial idx m < n → map (Int.castRingHom ℚ) (wittStructureInt p Φ m) = wittStructureRat p (map (Int.castRingHom ℚ) Φ) m) : - (C (p ^ n : ℤ) : MvPolynomial (idx × ℕ) ℤ) ∣ + (C ((p ^ n :) : ℤ) : MvPolynomial (idx × ℕ) ℤ) ∣ bind₁ (fun b : idx => rename (fun i => (b, i)) (wittPolynomial p ℤ n)) Φ - ∑ i in range n, C ((p : ℤ) ^ i) * wittStructureInt p Φ i ^ p ^ (n - i) := by cases' n with n @@ -261,7 +261,6 @@ theorem C_p_pow_dvd_bind₁_rename_wittPolynomial_sub_sum (Φ : MvPolynomial idx apply mul_dvd_mul_left ((p : MvPolynomial (idx × ℕ) ℤ) ^ k) rw [show p ^ (n + 1 - k) = p * p ^ (n - k) by rw [← pow_succ, ← tsub_add_eq_add_tsub hk]] rw [pow_mul] - rw [← Nat.cast_pow] -- Porting note: added -- the machine! apply dvd_sub_pow_of_dvd_sub rw [← C_eq_coe_nat, C_dvd_iff_zmod, RingHom.map_sub, sub_eq_zero, map_expand, RingHom.map_pow, diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 8b423b8686ec5..2c78cf41c28fa 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -493,7 +493,7 @@ local infixr:80 " ^' " => @HPow.hPow Cardinal Cardinal Cardinal _ -- -- mathport name: cardinal.pow.nat local infixr:80 " ^ℕ " => @HPow.hPow Cardinal ℕ Cardinal instHPow -theorem power_def (α β) : #α ^ #β = #(β → α) := +theorem power_def (α β : Type u) : #α ^ #β = #(β → α) := rfl #align cardinal.power_def Cardinal.power_def @@ -508,12 +508,12 @@ theorem lift_power (a b : Cardinal.{u}) : lift.{v} (a ^ b) = lift.{v} a ^ lift.{ #align cardinal.lift_power Cardinal.lift_power @[simp] -theorem power_zero {a : Cardinal} : a ^ 0 = 1 := +theorem power_zero {a : Cardinal} : a ^ (0 : Cardinal) = 1 := inductionOn a fun _ => mk_eq_one _ #align cardinal.power_zero Cardinal.power_zero @[simp] -theorem power_one {a : Cardinal.{u}} : a ^ 1 = a := +theorem power_one {a : Cardinal.{u}} : a ^ (1 : Cardinal) = a := inductionOn a fun α => mk_congr (Equiv.funUnique (ULift.{u} (Fin 1)) α) #align cardinal.power_one Cardinal.power_one @@ -538,9 +538,9 @@ instance commSemiring : CommSemiring Cardinal.{u} where mul_comm := mul_comm' left_distrib a b c := inductionOn₃ a b c fun α β γ => mk_congr <| Equiv.prodSumDistrib α β γ right_distrib a b c := inductionOn₃ a b c fun α β γ => mk_congr <| Equiv.sumProdDistrib α β γ - npow n c := c ^ n + npow n c := c ^ (n : Cardinal) npow_zero := @power_zero - npow_succ n c := show c ^ (n + 1 : ℕ) = c * (c ^ n) + npow_succ n c := show c ^ (↑(n + 1) : Cardinal) = c * c ^ (↑n : Cardinal) by rw [Cardinal.cast_succ, power_add, power_one, mul_comm'] natCast := (fun n => lift.{u} #(Fin n) : ℕ → Cardinal.{u}) natCast_zero := rfl @@ -563,7 +563,7 @@ theorem power_bit1 (a b : Cardinal) : a ^ bit1 b = a ^ b * a ^ b * a := by end deprecated @[simp] -theorem one_power {a : Cardinal} : 1 ^ a = 1 := +theorem one_power {a : Cardinal} : (1 : Cardinal) ^ a = 1 := inductionOn a fun _ => mk_eq_one _ #align cardinal.one_power Cardinal.one_power @@ -578,7 +578,7 @@ theorem mk_Prop : #Prop = 2 := by simp #align cardinal.mk_Prop Cardinal.mk_Prop @[simp] -theorem zero_power {a : Cardinal} : a ≠ 0 → 0 ^ a = 0 := +theorem zero_power {a : Cardinal} : a ≠ 0 → (0 : Cardinal) ^ a = 0 := inductionOn a fun _ heq => mk_eq_zero_iff.2 <| isEmpty_pi.2 <| @@ -1347,7 +1347,7 @@ theorem card_le_of_finset {α} (s : Finset α) : (s.card : Cardinal) ≤ #α := -- Porting note: was `simp`. LHS is not normal form. -- @[simp, norm_cast] @[norm_cast] -theorem natCast_pow {m n : ℕ} : (↑(m ^ n) : Cardinal) = m ^ n := by +theorem natCast_pow {m n : ℕ} : (↑(m ^ n) : Cardinal) = (↑m : Cardinal) ^ (↑n : Cardinal) := by induction n <;> simp [pow_succ', power_add, *, Pow.pow] #align cardinal.nat_cast_pow Cardinal.natCast_pow @@ -1898,13 +1898,13 @@ def toPartENat : Cardinal →+ PartENat where rw [← Nat.cast_add, toNat_cast, Nat.cast_add] · simp_rw [if_neg hy, PartENat.add_top] contrapose! hy - simp only [ne_eq, ite_eq_right_iff, - PartENat.natCast_ne_top, not_forall, exists_prop, and_true] at hy + simp only [ne_eq, ite_eq_right_iff, PartENat.natCast_ne_top, not_forall, exists_prop, + and_true, not_false_eq_true] at hy exact le_add_self.trans_lt hy · simp_rw [if_neg hx, PartENat.top_add] contrapose! hx - simp only [ne_eq, ite_eq_right_iff, - PartENat.natCast_ne_top, not_forall, exists_prop, and_true] at hx + simp only [ne_eq, ite_eq_right_iff, PartENat.natCast_ne_top, not_forall, exists_prop, + and_true, not_false_eq_true] at hx exact le_self_add.trans_lt hx #align cardinal.to_part_enat Cardinal.toPartENat @@ -2422,14 +2422,14 @@ theorem three_le {α : Type*} (h : 3 ≤ #α) (x : α) (y : α) : ∃ z : α, z /-- The function `a ^< b`, defined as the supremum of `a ^ c` for `c < b`. -/ def powerlt (a b : Cardinal.{u}) : Cardinal.{u} := - ⨆ c : Iio b, a^c + ⨆ c : Iio b, a ^ (c : Cardinal) #align cardinal.powerlt Cardinal.powerlt @[inherit_doc] infixl:80 " ^< " => powerlt -theorem le_powerlt {b c : Cardinal.{u}} (a) (h : c < b) : a ^ c ≤ a ^< b := by - refine le_ciSup (f := fun y : Iio b => a ^ y) ?_ ⟨c, h⟩ +theorem le_powerlt {b c : Cardinal.{u}} (a) (h : c < b) : (a^c) ≤ a ^< b := by + refine le_ciSup (f := fun y : Iio b => a ^ (y : Cardinal)) ?_ ⟨c, h⟩ rw [← image_eq_range] exact bddAbove_image.{u, u} _ bddAbove_Iio #align cardinal.le_powerlt Cardinal.le_powerlt diff --git a/Mathlib/SetTheory/Cardinal/Divisibility.lean b/Mathlib/SetTheory/Cardinal/Divisibility.lean index 0cb127333ca0f..aa01cfde5faab 100644 --- a/Mathlib/SetTheory/Cardinal/Divisibility.lean +++ b/Mathlib/SetTheory/Cardinal/Divisibility.lean @@ -150,7 +150,8 @@ theorem isPrimePow_iff {a : Cardinal} : IsPrimePow a ↔ ℵ₀ ≤ a ∨ ∃ n ⟨_, fun ⟨n, han, p, k, hp, hk, h⟩ => ⟨p, k, nat_is_prime_iff.2 hp, hk, by rw [han]; exact_mod_cast h⟩⟩ rintro ⟨p, k, hp, hk, hpk⟩ - have key : p ^ 1 ≤ ↑a := by rw [←hpk]; apply power_le_power_left hp.ne_zero; exact_mod_cast hk + have key : p ^ (1 : Cardinal) ≤ ↑a := by + rw [←hpk]; apply power_le_power_left hp.ne_zero; exact_mod_cast hk rw [power_one] at key lift p to ℕ using key.trans_lt (nat_lt_aleph0 a) exact ⟨a, rfl, p, k, nat_is_prime_iff.mp hp, hk, by exact_mod_cast hpk⟩ diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 6d64907eab926..01af4796a066e 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -92,7 +92,7 @@ an equivalence between `α` and `Fin (Nat.card α)`. See also `Finite.equivFin`. def equivFinOfCardPos {α : Type*} (h : Nat.card α ≠ 0) : α ≃ Fin (Nat.card α) := by cases fintypeOrInfinite α · simpa only [card_eq_fintype_card] using Fintype.equivFin α - · simp only [card_eq_zero_of_infinite, ne_eq] at h + · simp only [card_eq_zero_of_infinite, ne_eq, not_true_eq_false] at h #align nat.equiv_fin_of_card_pos Nat.equivFinOfCardPos theorem card_of_subsingleton (a : α) [Subsingleton α] : Nat.card α = 1 := by diff --git a/Mathlib/SetTheory/Lists.lean b/Mathlib/SetTheory/Lists.lean index 9306d7ecadcd2..3b811c7d62078 100644 --- a/Mathlib/SetTheory/Lists.lean +++ b/Mathlib/SetTheory/Lists.lean @@ -369,7 +369,7 @@ def Equiv.decidableMeas : theorem sizeof_pos {b} (l : Lists' α b) : 0 < SizeOf.sizeOf l := by cases l <;> simp only [Lists'.atom.sizeOf_spec, Lists'.nil.sizeOf_spec, Lists'.cons'.sizeOf_spec, - true_or, add_pos_iff] + true_or, add_pos_iff, zero_lt_one] #align lists.sizeof_pos Lists.sizeof_pos theorem lt_sizeof_cons' {b} (a : Lists' α b) (l) : diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index fef126c93c80a..f8708e3d7f39f 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -1003,7 +1003,8 @@ theorem le_of_dvd : ∀ {a b : Ordinal}, b ≠ 0 → a ∣ b → a ≤ b -- Porting note: `Ne` is required. simpa only [mul_one] using mul_le_mul_left' - (one_le_iff_ne_zero.2 fun h : b = 0 => by simp only [h, mul_zero, Ne] at b0) a + (one_le_iff_ne_zero.2 fun h : b = 0 => by + simp only [h, mul_zero, Ne, not_true_eq_false] at b0) a #align ordinal.le_of_dvd Ordinal.le_of_dvd theorem dvd_antisymm {a b : Ordinal} (h₁ : a ∣ b) (h₂ : b ∣ a) : a = b := @@ -1459,7 +1460,8 @@ theorem sSup_eq_bsup {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : @[simp] theorem bsup_eq_sup' {ι : Type u} (r : ι → ι → Prop) [IsWellOrder ι r] (f : ι → Ordinal.{max u v}) : bsup.{_, v} _ (bfamilyOfFamily' r f) = sup.{_, v} f := by - simp only [← sup_eq_bsup' r, enum_typein, familyOfBFamily', bfamilyOfFamily'] + simp (config := { unfoldPartialApp := true }) only [← sup_eq_bsup' r, enum_typein, + familyOfBFamily', bfamilyOfFamily'] #align ordinal.bsup_eq_sup' Ordinal.bsup_eq_sup' theorem bsup_eq_bsup {ι : Type u} (r r' : ι → ι → Prop) [IsWellOrder ι r] [IsWellOrder ι r'] diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 6cc16b2266bd9..af02590d43ff3 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -648,7 +648,7 @@ def lift (o : Ordinal.{v}) : Ordinal.{max v u} := -- @[simp] -- Porting note: Not in simpnf, added aux lemma below theorem type_uLift (r : α → α → Prop) [IsWellOrder α r] : type (ULift.down.{v,u} ⁻¹'o r) = lift.{v} (type r) := by - simp + simp (config := { unfoldPartialApp := true }) rfl #align ordinal.type_ulift Ordinal.type_uLift @@ -921,7 +921,6 @@ instance add_covariantClass_le : CovariantClass Ordinal.{u} Ordinal.{u} (· + · refine inductionOn b (fun α₂ r₂ _ ↦ ?_) rintro c ⟨⟨⟨f, fo⟩, fi⟩⟩ refine inductionOn c (fun β s _ ↦ ?_) - have := (Embedding.refl β).sumMap f refine ⟨⟨⟨(Embedding.refl.{u+1} _).sumMap f, ?_⟩, ?_⟩⟩ · intros a b match a, b with diff --git a/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean b/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean index 0ce02e0b9ead4..bed9105c56c5b 100644 --- a/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean +++ b/Mathlib/SetTheory/Ordinal/CantorNormalForm.lean @@ -162,11 +162,11 @@ set_option linter.uppercaseLean3 false in /-- The exponents of the Cantor normal form are decreasing. -/ theorem CNF_sorted (b o : Ordinal) : ((CNF b o).map Prod.fst).Sorted (· > ·) := by refine' CNFRec b _ (fun o ho IH ↦ _) o - · simp only [CNF_zero] + · simp only [gt_iff_lt, CNF_zero, map_nil, sorted_nil] · cases' le_or_lt b 1 with hb hb - · simp only [CNF_of_le_one hb ho, map] + · simp only [CNF_of_le_one hb ho, gt_iff_lt, map_cons, map, sorted_singleton] · cases' lt_or_le o b with hob hbo - · simp only [CNF_of_lt ho hob, map] + · simp only [CNF_of_lt ho hob, gt_iff_lt, map_cons, map, sorted_singleton] · rw [CNF_ne_zero ho, map_cons, sorted_cons] refine' ⟨fun a H ↦ _, IH⟩ rw [mem_map] at H diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index c8e20f4ca2d0c..1001103c7694c 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -42,12 +42,12 @@ theorem zero_opow' (a : Ordinal) : 0 ^ a = 1 - a := by simp only [opow_def, if_t #align ordinal.zero_opow' Ordinal.zero_opow' @[simp] -theorem zero_opow {a : Ordinal} (a0 : a ≠ 0) : 0 ^ a = 0 := by +theorem zero_opow {a : Ordinal} (a0 : a ≠ 0) : (0 : Ordinal) ^ a = 0 := by rwa [zero_opow', Ordinal.sub_eq_zero_iff_le, one_le_iff_ne_zero] #align ordinal.zero_opow Ordinal.zero_opow @[simp] -theorem opow_zero (a : Ordinal) : a ^ 0 = 1 := by +theorem opow_zero (a : Ordinal) : a ^ (0 : Ordinal) = 1 := by by_cases h : a = 0 · simp only [opow_def, if_pos h, sub_zero] · simp only [opow_def, if_neg h, limitRecOn_zero] @@ -74,12 +74,12 @@ theorem lt_opow_of_limit {a b c : Ordinal} (b0 : b ≠ 0) (h : IsLimit c) : #align ordinal.lt_opow_of_limit Ordinal.lt_opow_of_limit @[simp] -theorem opow_one (a : Ordinal) : a ^ 1 = a := by +theorem opow_one (a : Ordinal) : a ^ (1 : Ordinal) = a := by rw [← succ_zero, opow_succ]; simp only [opow_zero, one_mul] #align ordinal.opow_one Ordinal.opow_one @[simp] -theorem one_opow (a : Ordinal) : 1 ^ a = 1 := by +theorem one_opow (a : Ordinal) : (1 : Ordinal) ^ a = 1 := by induction a using limitRecOn with | H₁ => simp only [opow_zero] | H₂ _ ih => @@ -91,7 +91,7 @@ theorem one_opow (a : Ordinal) : 1 ^ a = 1 := by #align ordinal.one_opow Ordinal.one_opow theorem opow_pos {a : Ordinal} (b : Ordinal) (a0 : 0 < a) : 0 < a ^ b := by - have h0 : 0 < a ^ 0 := by simp only [opow_zero, zero_lt_one] + have h0 : 0 < a ^ (0 : Ordinal) := by simp only [opow_zero, zero_lt_one] induction b using limitRecOn with | H₁ => exact h0 | H₂ b IH => @@ -199,7 +199,7 @@ theorem opow_add (a b c : Ordinal) : a ^ (b + c) = a ^ b * a ^ c := by refine' eq_of_forall_ge_iff fun d => (((opow_isNormal a1).trans (add_isNormal b)).limit_le l).trans _ - dsimp only [Function.comp] + dsimp only [Function.comp_def] simp (config := { contextual := true }) only [IH] exact (((mul_isNormal <| opow_pos b (Ordinal.pos_iff_ne_zero.2 a0)).trans @@ -242,7 +242,7 @@ theorem opow_mul (a b c : Ordinal) : a ^ (b * c) = (a ^ b) ^ c := by (((opow_isNormal a1).trans (mul_isNormal (Ordinal.pos_iff_ne_zero.2 b0))).limit_le l).trans _ - dsimp only [Function.comp] + dsimp only [Function.comp_def] simp (config := { contextual := true }) only [IH] exact (opow_le_of_limit (opow_ne_zero _ a0) l).symm #align ordinal.opow_mul Ordinal.opow_mul @@ -258,7 +258,7 @@ def log (b : Ordinal) (x : Ordinal) : Ordinal := #align ordinal.log Ordinal.log /-- The set in the definition of `log` is nonempty. -/ -theorem log_nonempty {b x : Ordinal} (h : 1 < b) : { o | x < b ^ o }.Nonempty := +theorem log_nonempty {b x : Ordinal} (h : 1 < b) : { o : Ordinal | x < b ^ o }.Nonempty := ⟨_, succ_le_iff.1 (right_le_opow _ h)⟩ #align ordinal.log_nonempty Ordinal.log_nonempty @@ -296,9 +296,9 @@ theorem log_one_left : ∀ b, log 1 b = 0 := #align ordinal.log_one_left Ordinal.log_one_left theorem succ_log_def {b x : Ordinal} (hb : 1 < b) (hx : x ≠ 0) : - succ (log b x) = sInf { o | x < b ^ o } := by - let t := sInf { o | x < b ^ o } - have : x < b ^ t := csInf_mem (log_nonempty hb) + succ (log b x) = sInf { o : Ordinal | x < b ^ o } := by + let t := sInf { o : Ordinal | x < b ^ o } + have : x < (b^t) := csInf_mem (log_nonempty hb) rcases zero_or_succ_or_limit t with (h | h | h) · refine' ((one_le_iff_ne_zero.2 hx).not_lt _).elim simpa only [h, opow_zero] using this @@ -452,18 +452,18 @@ theorem add_log_le_log_mul {x y : Ordinal} (b : Ordinal) (hx : x ≠ 0) (hy : y /-! ### Interaction with `Nat.cast` -/ @[simp, norm_cast] -theorem nat_cast_opow (m : ℕ) : ∀ n : ℕ, ((m ^ n : ℕ) : Ordinal) = m ^ n +theorem nat_cast_opow (m : ℕ) : ∀ n : ℕ, ↑(m ^ n : ℕ) = (m : Ordinal) ^ (n : Ordinal) | 0 => by simp | n + 1 => by rw [pow_succ', nat_cast_mul, nat_cast_opow m n, Nat.cast_succ, add_one_eq_succ, opow_succ] #align ordinal.nat_cast_opow Ordinal.nat_cast_opow -theorem sup_opow_nat {o : Ordinal} (ho : 0 < o) : (sup fun n : ℕ => o ^ n) = o ^ ω := by +theorem sup_opow_nat {o : Ordinal} (ho : 0 < o) : (sup fun n : ℕ => o ^ (n : Ordinal)) = o ^ ω := by rcases lt_or_eq_of_le (one_le_iff_pos.2 ho) with (ho₁ | rfl) · exact (opow_isNormal ho₁).apply_omega · rw [one_opow] refine' le_antisymm (sup_le fun n => by rw [one_opow]) _ - convert le_sup (fun n : ℕ => 1 ^ n) 0 + convert le_sup (fun n : ℕ => 1 ^ (n : Ordinal)) 0 rw [Nat.cast_zero, opow_zero] #align ordinal.sup_opow_nat Ordinal.sup_opow_nat diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index 0ce5aceecc77e..6de1fb79e4e8b 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -189,7 +189,8 @@ theorem eq_of_cmp_eq : ∀ {o₁ o₂}, cmp o₁ o₂ = Ordering.eq → o₁ = o #align onote.eq_of_cmp_eq ONote.eq_of_cmp_eq protected theorem zero_lt_one : (0 : ONote) < 1 := by - simp only [lt_def, repr, repr_one, opow_zero, one_mul, add_zero, nat_cast_pos] + simp only [lt_def, repr, opow_zero, Nat.succPNat_coe, Nat.cast_one, mul_one, add_zero, + zero_lt_one] #align onote.zero_lt_one ONote.zero_lt_one /-- `NFBelow o b` says that `o` is a normal form ordinal notation @@ -830,9 +831,10 @@ instance nf_opow (o₁ o₂) [NF o₁] [NF o₂] : NF (o₁ ^ o₂) := by haveI := (nf_repr_split' e₂).1 cases' a with a0 n a' · cases' m with m - · by_cases o₂ = 0 <;> simp [(· ^ ·), Pow.pow, pow, opow, opowAux2, *] + · by_cases o₂ = 0 <;> simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, *] <;> decide · by_cases m = 0 · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, *, zero_def] + decide · simp only [(· ^ ·), Pow.pow, pow, opow, opowAux2, mulNat_eq_mul, ofNat, *] infer_instance · simp [(· ^ ·),Pow.pow,pow, opow, opowAux2, e₁, e₂, split_eq_scale_split' e₂] @@ -887,9 +889,9 @@ set_option linter.unusedVariables false in theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω ∣ repr a') (e0 : repr a0 ≠ 0) (h : repr a' + m < (ω ^ repr a0)) (n : ℕ+) (k : ℕ) : let R := repr (opowAux 0 a0 (oadd a0 n a' * ofNat m) k m) - (k ≠ 0 → R < ((ω ^ repr a0) ^ succ ↑k)) ∧ - ((ω ^ repr a0) ^ k) * ((ω ^ repr a0) * (n : ℕ) + repr a') + R = - ((ω ^ repr a0) * (n : ℕ) + repr a' + m) ^ succ ↑k := by + (k ≠ 0 → R < ((ω ^ repr a0) ^ succ (k : Ordinal))) ∧ + ((ω ^ repr a0) ^ (k : Ordinal)) * ((ω ^ repr a0) * (n : ℕ) + repr a') + R = + ((ω ^ repr a0) * (n : ℕ) + repr a' + m) ^ succ (k : Ordinal) := by intro R' haveI No : NF (oadd a0 n a') := N0.oadd n (Na'.below_of_lt' <| lt_of_le_of_lt (le_add_right _ _) h) @@ -899,15 +901,16 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω let R := repr (opowAux 0 a0 (oadd a0 n a' * ofNat m) k m) let ω0 := ω ^ repr a0 let α' := ω0 * n + repr a' - change (k ≠ 0 → R < (ω0 ^ succ ↑k)) ∧ (ω0 ^ k) * α' + R = (α' + m) ^ succ ↑k at IH - have RR : R' = (ω0 ^ k) * (α' * m) + R := by + change (k ≠ 0 → R < (ω0 ^ succ (k : Ordinal))) ∧ (ω0 ^ (k : Ordinal)) * α' + R + = (α' + m) ^ (succ ↑k : Ordinal) at IH + have RR : R' = ω0 ^ (k : Ordinal) * (α' * m) + R := by by_cases h : m = 0 · simp only [h, ONote.ofNat, Nat.cast_zero, zero_add, ONote.repr, mul_zero, ONote.opowAux, add_zero] · simp only [ONote.repr_scale, ONote.repr, ONote.mulNat_eq_mul, ONote.opowAux, ONote.repr_ofNat, ONote.repr_mul, ONote.repr_add, Ordinal.opow_mul, ONote.zero_add] have α0 : 0 < α' := by simpa [lt_def, repr] using oadd_pos a0 n a' - have ω00 : 0 < (ω0 ^ k) := opow_pos _ (opow_pos _ omega_pos) + have ω00 : 0 < ω0 ^ (k : Ordinal) := opow_pos _ (opow_pos _ omega_pos) have Rl : R < ω ^ (repr a0 * succ ↑k) := by by_cases k0 : k = 0 · simp [k0] @@ -932,21 +935,22 @@ theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω (opow_le_opow_right omega_pos <| mul_le_mul_left' (succ_le_succ_iff.2 (nat_cast_le.2 (le_of_lt k.lt_succ_self))) _) calc - (ω0 ^ k.succ) * α' + R' - _ = (ω0 ^ succ ↑k) * α' + ((ω0 ^ k) * α' * m + R) := by rw [nat_cast_succ, RR, ← mul_assoc] - _ = ((ω0 ^ k) * α' + R) * α' + ((ω0 ^ k) * α' + R) * m := ?_ - _ = (α' + m) ^ succ ↑k.succ := by rw [← mul_add, nat_cast_succ, opow_succ, IH.2] + (ω0 ^ (k.succ : Ordinal)) * α' + R' + _ = (ω0 ^ succ (k : Ordinal)) * α' + ((ω0 ^ (k : Ordinal)) * α' * m + R) := + by rw [nat_cast_succ, RR, ← mul_assoc] + _ = ((ω0 ^ (k : Ordinal)) * α' + R) * α' + ((ω0 ^ (k : Ordinal)) * α' + R) * m := ?_ + _ = (α' + m) ^ succ (k.succ : Ordinal) := by rw [← mul_add, nat_cast_succ, opow_succ, IH.2] congr 1 · have αd : ω ∣ α' := dvd_add (dvd_mul_of_dvd_left (by simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 e0)) _) d - rw [mul_add (ω0^k), add_assoc, ← mul_assoc, ← opow_succ, + rw [mul_add (ω0 ^ (k : Ordinal)), add_assoc, ← mul_assoc, ← opow_succ, add_mul_limit _ (isLimit_iff_omega_dvd.2 ⟨ne_of_gt α0, αd⟩), mul_assoc, @mul_omega_dvd n (nat_cast_pos.2 n.pos) (nat_lt_omega _) _ αd] apply @add_absorp _ (repr a0 * succ ↑k) · refine' principal_add_omega_opow _ _ Rl rw [opow_mul, opow_succ, Ordinal.mul_lt_mul_iff_left ω00] exact No.snd'.repr_lt - · have := mul_le_mul_left' (one_le_iff_pos.2 <| nat_cast_pos.2 n.pos) (ω0^succ k) + · have := mul_le_mul_left' (one_le_iff_pos.2 <| nat_cast_pos.2 n.pos) (ω0 ^ succ (k : Ordinal)) rw [opow_mul] simpa [-opow_succ] · cases m @@ -1098,6 +1102,7 @@ theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamenta 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' diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 818dcae351052..53af177bf8aa8 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -202,7 +202,7 @@ theorem principal_add_omega_opow (o : Ordinal) : Principal (· + ·) (omega^o) : /-- The main characterization theorem for additive principal ordinals. -/ theorem principal_add_iff_zero_or_omega_opow {o : Ordinal} : - Principal (· + ·) o ↔ o = 0 ∨ ∃ a, o = (omega^a) := by + Principal (· + ·) o ↔ o = 0 ∨ ∃ a : Ordinal, o = (omega^a) := by rcases eq_or_ne o 0 with (rfl | ho) · simp only [principal_zero, Or.inl] · rw [principal_add_iff_add_left_eq_self] @@ -371,7 +371,7 @@ theorem principal_add_of_principal_mul_opow {o b : Ordinal} (hb : 1 < b) /-- The main characterization theorem for multiplicative principal ordinals. -/ theorem principal_mul_iff_le_two_or_omega_opow_opow {o : Ordinal} : - Principal (· * ·) o ↔ o ≤ 2 ∨ ∃ a, o = (omega^omega^a) := by + Principal (· * ·) o ↔ o ≤ 2 ∨ ∃ a : Ordinal, o = (omega^omega^a) := by refine' ⟨fun ho => _, _⟩ · cases' le_or_lt o 2 with ho₂ ho₂ · exact Or.inl ho₂ diff --git a/Mathlib/Tactic/ComputeDegree.lean b/Mathlib/Tactic/ComputeDegree.lean index d523afc63969e..5ab9de971c149 100644 --- a/Mathlib/Tactic/ComputeDegree.lean +++ b/Mathlib/Tactic/ComputeDegree.lean @@ -448,12 +448,14 @@ elab_rules : tactic | `(tactic| compute_degree $[!%$bang]?) => focus <| withMain -- simplify the left-hand sides, since this is where the degree computations leave -- expressions such as `max (0 * 1) (max (1 + 0 + 3 * 4) (7 * 0))` evalTactic - (← `(tactic| try any_goals conv_lhs => (simp only [Nat.cast_withBot]; norm_num))) + (← `(tactic| try any_goals conv_lhs => + (simp (config := {decide := true}) only [Nat.cast_withBot]; norm_num))) if bang.isSome then let mut false_goals : Array MVarId := #[] let mut new_goals : Array MVarId := #[] for g in ← getGoals do - let gs' ← run g do evalTactic (← `(tactic| try (any_goals norm_num <;> try assumption))) + let gs' ← run g do evalTactic (← + `(tactic| try (any_goals norm_num <;> norm_cast <;> try assumption))) new_goals := new_goals ++ gs'.toArray if ← gs'.anyM fun g' => g'.withContext do return (← g'.getType'').isConstOf ``False then false_goals := false_goals.push g diff --git a/Mathlib/Tactic/DeriveTraversable.lean b/Mathlib/Tactic/DeriveTraversable.lean index 88f5712fd8693..2bd7d81ad59f3 100644 --- a/Mathlib/Tactic/DeriveTraversable.lean +++ b/Mathlib/Tactic/DeriveTraversable.lean @@ -463,7 +463,9 @@ def deriveLawfulTraversable (m : MVarId) : TermElabM Unit := do if b then let hs ← getPropHyps s ← hs.foldlM (fun s f => f.getDecl >>= fun d => s.add (.fvar f) #[] d.toExpr) s - return { config := { failIfUnchanged := false }, simpTheorems := #[s] } + pure <| + { config := { failIfUnchanged := false, unfoldPartialApp := true }, + simpTheorems := #[s] } let .app (.app (.const ``LawfulTraversable _) F) _ ← m.getType >>= instantiateMVars | failure let some n := F.getAppFn.constName? | failure let [mit, mct, mtmi, mn] ← m.applyConst ``LawfulTraversable.mk | failure diff --git a/Mathlib/Tactic/LibrarySearch.lean b/Mathlib/Tactic/LibrarySearch.lean index 676af95e5ed19..e81c187e65522 100644 --- a/Mathlib/Tactic/LibrarySearch.lean +++ b/Mathlib/Tactic/LibrarySearch.lean @@ -94,7 +94,7 @@ def cachePath : IO FilePath := try return (← findOLean `MathlibExtras.LibrarySearch).withExtension "extra" catch _ => - return "build" / "lib" / "MathlibExtras" / "LibrarySearch.extra" + return ".lake" / "build" / "lib" / "MathlibExtras" / "LibrarySearch.extra" /-- Retrieve the current current of lemmas. diff --git a/Mathlib/Tactic/NormNum/LegendreSymbol.lean b/Mathlib/Tactic/NormNum/LegendreSymbol.lean index db5269773b3c2..8417a02e44a34 100644 --- a/Mathlib/Tactic/NormNum/LegendreSymbol.lean +++ b/Mathlib/Tactic/NormNum/LegendreSymbol.lean @@ -124,6 +124,7 @@ theorem jacobiSymNat.odd_even (a b c : ℕ) (r : ℤ) (ha : a % 2 = 1) (hb : b % jacobiSymNat a b = r := by have ha' : legendreSym 2 a = 1 := by simp only [legendreSym.mod 2 a, Int.ofNat_mod_ofNat, ha] + decide rcases eq_or_ne c 0 with (rfl | hc') · rw [← hr, Nat.eq_zero_of_dvd_of_div_eq_zero (Nat.dvd_of_mod_eq_zero hb) hc] · haveI : NeZero c := ⟨hc'⟩ diff --git a/Mathlib/Tactic/Rewrites.lean b/Mathlib/Tactic/Rewrites.lean index a7c2ad334165a..5a6161ebeb7ff 100644 --- a/Mathlib/Tactic/Rewrites.lean +++ b/Mathlib/Tactic/Rewrites.lean @@ -126,7 +126,7 @@ def cachePath : IO FilePath := try return (← findOLean `MathlibExtras.Rewrites).withExtension "extra" catch _ => - return "build" / "lib" / "MathlibExtras" / "Rewrites.extra" + return ".lake" / "build" / "lib" / "MathlibExtras" / "Rewrites.extra" /-- Retrieve the current cache of lemmas. diff --git a/Mathlib/Tactic/Simps/Basic.lean b/Mathlib/Tactic/Simps/Basic.lean index 2e033e74ec179..7cdceb6df4dc7 100644 --- a/Mathlib/Tactic/Simps/Basic.lean +++ b/Mathlib/Tactic/Simps/Basic.lean @@ -4,12 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import Mathlib.Data.KVMap import Mathlib.Tactic.Simps.NotationClass import Std.Classes.Dvd import Std.Data.String.Basic import Std.Util.LibraryNote -import Mathlib.Lean.Linter import Std.Data.List.Count import Mathlib.Lean.Expr.Basic diff --git a/Mathlib/Tactic/ToAdditive.lean b/Mathlib/Tactic/ToAdditive.lean index 0217a19889152..ed6555f95bdaa 100644 --- a/Mathlib/Tactic/ToAdditive.lean +++ b/Mathlib/Tactic/ToAdditive.lean @@ -6,7 +6,6 @@ Authors: Mario Carneiro, Yury Kudryashov, Floris van Doorn, Jon Eugster import Mathlib.Init.Data.Nat.Notation import Mathlib.Data.String.Defs import Mathlib.Data.Array.Defs -import Mathlib.Data.KVMap import Mathlib.Lean.Expr.ReplaceRec import Mathlib.Lean.EnvExtension import Mathlib.Lean.Meta.Simp diff --git a/Mathlib/Tactic/UnsetOption.lean b/Mathlib/Tactic/UnsetOption.lean index 1ed84c401ccea..b0e1446df15ab 100644 --- a/Mathlib/Tactic/UnsetOption.lean +++ b/Mathlib/Tactic/UnsetOption.lean @@ -6,7 +6,6 @@ Authors: Alex J. Best import Lean.Parser.Term import Lean.Parser.Do import Lean.Elab.Command -import Mathlib.Data.KVMap /-! # The `unset_option` command diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index 414d8eb56f7ec..c10aaa08e3332 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -1135,8 +1135,9 @@ theorem cauchySeq_finset_iff_vanishing : (CauchySeq fun s : Finset β ↦ ∑ b in s, f b) ↔ ∀ e ∈ 𝓝 (0 : α), ∃ s : Finset β, ∀ t, Disjoint t s → (∑ b in t, f b) ∈ e := by classical - simp_rw [CauchySeq, cauchy_map_iff, and_iff_right atTop_neBot, prod_atTop_atTop_eq, - uniformity_eq_comap_nhds_zero α, tendsto_comap_iff, (· ∘ ·), tendsto_atTop'] + simp only [CauchySeq, cauchy_map_iff, and_iff_right atTop_neBot, prod_atTop_atTop_eq, + uniformity_eq_comap_nhds_zero α, tendsto_comap_iff, (· ∘ ·), atTop_neBot, true_and] + rw [tendsto_atTop'] constructor · intro h e he obtain ⟨⟨s₁, s₂⟩, h⟩ := h e he diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index fdd9b9dab65a3..0216ed67b218f 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -605,13 +605,13 @@ theorem smul_apply (c : S₂) (f : M₁ →SL[σ₁₂] M₂) (x : M₁) : (c @[simp, norm_cast] theorem coe_smul (c : S₂) (f : M₁ →SL[σ₁₂] M₂) : - (c • f : M₁ →ₛₗ[σ₁₂] M₂) = c • (f : M₁ →ₛₗ[σ₁₂] M₂) := + ↑(c • f) = c • (f : M₁ →ₛₗ[σ₁₂] M₂) := rfl #align continuous_linear_map.coe_smul ContinuousLinearMap.coe_smul @[simp, norm_cast] theorem coe_smul' (c : S₂) (f : M₁ →SL[σ₁₂] M₂) : - (c • f : M₁ → M₂) = c • (f : M₁ → M₂) := + ↑(c • f) = c • (f : M₁ → M₂) := rfl #align continuous_linear_map.coe_smul' ContinuousLinearMap.coe_smul' diff --git a/Mathlib/Topology/Algebra/Module/Cardinality.lean b/Mathlib/Topology/Algebra/Module/Cardinality.lean index 86dce76fac8dc..d87a1032f5c94 100644 --- a/Mathlib/Topology/Algebra/Module/Cardinality.lean +++ b/Mathlib/Topology/Algebra/Module/Cardinality.lean @@ -81,9 +81,9 @@ lemma cardinal_eq_of_mem_nhds_zero rw [zero_smul] at this filter_upwards [this hs] with n (hn : (c ^ n)⁻¹ • x ∈ s) exact (mem_smul_set_iff_inv_smul_mem₀ (cn_ne n) _ _).2 hn - have B : ∀ n, #(c^n • s) = #s := by + have B : ∀ n, #(c^n • s :) = #s := by intro n - have : c^n • s ≃ s := + have : (c^n • s :) ≃ s := { toFun := fun x ↦ ⟨(c^n)⁻¹ • x.1, (mem_smul_set_iff_inv_smul_mem₀ (cn_ne n) _ _).1 x.2⟩ invFun := fun x ↦ ⟨(c^n) • x.1, smul_mem_smul_set x.2⟩ left_inv := fun x ↦ by simp [smul_smul, mul_inv_cancel (cn_ne n)] diff --git a/Mathlib/Topology/Algebra/Monoid.lean b/Mathlib/Topology/Algebra/Monoid.lean index 164e8c477920e..fddb9fa902076 100644 --- a/Mathlib/Topology/Algebra/Monoid.lean +++ b/Mathlib/Topology/Algebra/Monoid.lean @@ -276,16 +276,17 @@ theorem ContinuousMul.of_nhds_one {M : Type u} [Monoid M] [TopologicalSpace M] ext x simp [mul_assoc] calc - map (uncurry (· * ·)) (𝓝 (x₀, y₀)) = map (uncurry (· * ·)) (𝓝 x₀ ×ˢ 𝓝 y₀) := by - rw [nhds_prod_eq] + map (uncurry (· * ·)) (𝓝 (x₀, y₀)) = map (uncurry (· * ·)) (𝓝 x₀ ×ˢ 𝓝 y₀) := + by rw [nhds_prod_eq] _ = map (fun p : M × M => x₀ * p.1 * (p.2 * y₀)) (𝓝 1 ×ˢ 𝓝 1) := by -- Porting note: `rw` was able to prove this -- Now it fails with `failed to rewrite using equation theorems for 'Function.uncurry'` -- and `failed to rewrite using equation theorems for 'Function.comp'`. -- Removing those two lemmas, the `rw` would succeed, but then needs a `rfl`. - simp_rw [uncurry, hleft x₀, hright y₀, prod_map_map_eq, Filter.map_map, Function.comp] - _ = map ((fun x => x₀ * x) ∘ fun x => x * y₀) (map (uncurry (· * ·)) (𝓝 1 ×ˢ 𝓝 1)) := by - rw [key, ← Filter.map_map] + simp (config := { unfoldPartialApp := true }) only [uncurry] + simp_rw [hleft x₀, hright y₀, prod_map_map_eq, Filter.map_map, Function.comp_def] + _ = map ((fun x => x₀ * x) ∘ fun x => x * y₀) (map (uncurry (· * ·)) (𝓝 1 ×ˢ 𝓝 1)) := + by rw [key, ← Filter.map_map] _ ≤ map ((fun x : M => x₀ * x) ∘ fun x => x * y₀) (𝓝 1) := map_mono hmul _ = 𝓝 (x₀ * y₀) := by rw [← Filter.map_map, ← hright, hleft y₀, Filter.map_map, key₂, ← hleft]⟩ diff --git a/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean b/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean index f1139d116aa17..9a23271b177f5 100644 --- a/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean +++ b/Mathlib/Topology/Algebra/Nonarchimedean/AdicTopology.lean @@ -66,7 +66,7 @@ theorem adic_basis (I : Ideal R) : SubmodulesRingBasis fun n : ℕ => (I ^ n • rintro a ⟨x, hx, rfl⟩ exact (I ^ n).smul_mem r hx mul := by - suffices ∀ i : ℕ, ∃ j : ℕ, (I ^ j: Set R) * (I ^ j : Set R) ⊆ (I ^ i : Set R) by + suffices ∀ i : ℕ, ∃ j : ℕ, (↑(I ^ j) * ↑(I ^ j) : Set R) ⊆ (↑(I ^ i) : Set R) by simpa only [smul_top_eq_map, Algebra.id.map_eq_id, map_id] using this intro n use n @@ -140,7 +140,7 @@ def openAddSubgroup (n : ℕ) : @OpenAddSubgroup R _ I.adicTopology := by letI := I.adicTopology refine ⟨(I ^ n).toAddSubgroup, ?_⟩ convert (I.adic_basis.toRing_subgroups_basis.openAddSubgroup n).isOpen - change (I ^ n : Set R) = (I ^ n • (⊤ : Ideal R) : Set R) + change (↑(I ^ n) : Set R) = ↑(I ^ n • (⊤ : Ideal R)) simp [smul_top_eq_map, Algebra.id.map_eq_id, map_id, restrictScalars_self] #align ideal.open_add_subgroup Ideal.openAddSubgroup diff --git a/Mathlib/Topology/Algebra/ValuedField.lean b/Mathlib/Topology/Algebra/ValuedField.lean index 036e7ffabda59..c1c14cf3b87b5 100644 --- a/Mathlib/Topology/Algebra/ValuedField.lean +++ b/Mathlib/Topology/Algebra/ValuedField.lean @@ -167,11 +167,10 @@ instance (priority := 100) completable : CompletableTopField K := · rw [mem_map] apply mem_of_superset (Filter.inter_mem M₀_in M₁_in) exact subset_preimage_image _ _ - · rintro _ ⟨x, ⟨x_in₀, x_in₁⟩, rfl⟩ _ ⟨y, ⟨y_in₀, y_in₁⟩, rfl⟩ + · rintro _ ⟨x, ⟨x_in₀, x_in₁⟩, rfl⟩ _ ⟨y, ⟨_, y_in₁⟩, rfl⟩ simp only [mem_setOf_eq] specialize H₁ x x_in₁ y y_in₁ replace x_in₀ := H₀ x x_in₀ - replace := H₀ y y_in₀ clear H₀ apply Valuation.inversion_estimate · have : (v x : Γ₀) ≠ 0 := by diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 6b1c90477a8cd..d2b36293f56bc 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -848,7 +848,7 @@ scoped[Topology] notation "𝓝" => nhds scoped[Topology] notation "𝓝[" s "] " x:100 => nhdsWithin x s /-- Notation for the filter of punctured neighborhoods of a point. -/ -scoped[Topology] notation "𝓝[≠] " x:100 => nhdsWithin x {x}ᶜ +scoped[Topology] notation "𝓝[≠] " x:100 => nhdsWithin x (@singleton _ (Set _) instSingletonSet x)ᶜ /-- Notation for the filter of right neighborhoods of a point. -/ scoped[Topology] notation "𝓝[≥] " x:100 => nhdsWithin x (Set.Ici x) diff --git a/Mathlib/Topology/Category/Profinite/Nobeling.lean b/Mathlib/Topology/Category/Profinite/Nobeling.lean index 982314c312b73..acc37c3ac0177 100644 --- a/Mathlib/Topology/Category/Profinite/Nobeling.lean +++ b/Mathlib/Topology/Category/Profinite/Nobeling.lean @@ -90,7 +90,7 @@ def Proj : (I → Bool) → (I → Bool) := @[simp] theorem continuous_proj : Continuous (Proj J : (I → Bool) → (I → Bool)) := by - dsimp [Proj] + dsimp (config := { unfoldPartialApp := true }) [Proj] apply continuous_pi intro i split @@ -387,7 +387,7 @@ theorem eval_eq (l : Products I) (x : C) : push_neg at h convert h with i dsimp [LocallyConstant.evalMonoidHom, e] - simp only [ite_eq_right_iff] + simp only [ite_eq_right_iff, one_ne_zero] theorem evalFacProp {l : Products I} (J : I → Prop) (h : ∀ a, a ∈ l.val → J a) [∀ j, Decidable (J j)] : @@ -883,7 +883,7 @@ theorem isClosed_proj (o : Ordinal) (hC : IsClosed C) : IsClosed (π C (ord I · theorem contained_proj (o : Ordinal) : contained (π C (ord I · < o)) o := by intro x ⟨_, ⟨_, h⟩⟩ j hj - dsimp [Proj] at h + dsimp (config := { unfoldPartialApp := true }) [Proj] at h simp only [← congr_fun h j, Bool.ite_eq_true_distrib, if_false_right_eq_and] at hj exact hj.1 @@ -1024,7 +1024,7 @@ def range_equiv_smaller_toFun (o : Ordinal) (x : range (π C (ord I · < o))) : theorem range_equiv_smaller_toFun_bijective (o : Ordinal) : Function.Bijective (range_equiv_smaller_toFun C o) := by - dsimp [range_equiv_smaller_toFun] + dsimp (config := { unfoldPartialApp := true }) [range_equiv_smaller_toFun] refine ⟨fun a b hab ↦ ?_, fun ⟨a, b, hb⟩ ↦ ?_⟩ · ext1 simp only [Subtype.mk.injEq] at hab @@ -1232,7 +1232,7 @@ def SwapTrue : (I → Bool) → I → Bool := theorem continuous_swapTrue : Continuous (SwapTrue o : (I → Bool) → I → Bool) := by - dsimp [SwapTrue] + dsimp (config := { unfoldPartialApp := true }) [SwapTrue] apply continuous_pi intro i apply Continuous.comp' @@ -1245,7 +1245,7 @@ theorem swapTrue_mem_C1 (f : π (C1 C ho) (ord I · < o)) : SwapTrue o f.val ∈ C1 C ho := by obtain ⟨f, g, hg, rfl⟩ := f convert hg - dsimp [SwapTrue] + dsimp (config := { unfoldPartialApp := true }) [SwapTrue] ext i split_ifs with h · rw [ord_term ho] at h @@ -1797,7 +1797,7 @@ open Classical in /-- The map `Nobeling.ι` is a closed embedding. -/ theorem Nobeling.embedding : ClosedEmbedding (Nobeling.ι S) := by apply Continuous.closedEmbedding - · dsimp [ι] + · dsimp (config := { unfoldPartialApp := true }) [ι] refine continuous_pi ?_ intro C rw [← IsLocallyConstant.iff_continuous] @@ -1816,7 +1816,7 @@ theorem Nobeling.embedding : ClosedEmbedding (Nobeling.ι S) := by by_contra hn obtain ⟨C, hC, hh⟩ := exists_clopen_of_totally_separated hn apply hh.2 ∘ of_decide_eq_true - dsimp [ι] at h + dsimp (config := { unfoldPartialApp := true }) [ι] at h rw [← congr_fun h ⟨C, hC⟩] exact decide_eq_true hh.1 diff --git a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean index 8ce5d253e3d9f..f173aa6e68e30 100644 --- a/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean +++ b/Mathlib/Topology/Category/TopCat/Limits/Pullbacks.lean @@ -51,7 +51,7 @@ abbrev pullbackSnd (f : X ⟶ Z) (g : Y ⟶ Z) : TopCat.of { p : X × Y // f p.1 def pullbackCone (f : X ⟶ Z) (g : Y ⟶ Z) : PullbackCone f g := PullbackCone.mk (pullbackFst f g) (pullbackSnd f g) (by - dsimp [pullbackFst, pullbackSnd, Function.comp] + dsimp [pullbackFst, pullbackSnd, Function.comp_def] ext ⟨x, h⟩ -- Next 2 lines were -- `rw [comp_apply, ContinuousMap.coe_mk, comp_apply, ContinuousMap.coe_mk]` diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 24b62d858a2cf..5ce81050e5fba 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -802,7 +802,7 @@ theorem DenseRange.prod_map {ι : Type*} {κ : Type*} {f : ι → β} {g : κ theorem Inducing.prod_map {f : α → β} {g : γ → δ} (hf : Inducing f) (hg : Inducing g) : Inducing (Prod.map f g) := - inducing_iff_nhds.2 fun (a, b) => by simp_rw [Prod.map, nhds_prod_eq, hf.nhds_eq_comap, + inducing_iff_nhds.2 fun (a, b) => by simp_rw [Prod.map_def, nhds_prod_eq, hf.nhds_eq_comap, hg.nhds_eq_comap, prod_comap_comap_eq] #align inducing.prod_mk Inducing.prod_map @@ -1283,7 +1283,8 @@ lemma Pi.continuous_restrict (S : Set ι) : lemma Pi.induced_restrict (S : Set ι) : induced (S.restrict) Pi.topologicalSpace = ⨅ i ∈ S, induced (eval i) (T i) := by - simp [← iInf_subtype'', ← induced_precomp' ((↑) : S → ι), Set.restrict] + simp (config := { unfoldPartialApp := true }) [← iInf_subtype'', ← induced_precomp' ((↑) : S → ι), + Set.restrict] theorem Filter.Tendsto.update [DecidableEq ι] {l : Filter β} {f : β → ∀ i, π i} {x : ∀ i, π i} (hf : Tendsto f l (𝓝 x)) (i : ι) {g : β → π i} {xi : π i} (hg : Tendsto g l (𝓝 xi)) : diff --git a/Mathlib/Topology/ContinuousFunction/Basic.lean b/Mathlib/Topology/ContinuousFunction/Basic.lean index ee83919adae28..b46778e7984b0 100644 --- a/Mathlib/Topology/ContinuousFunction/Basic.lean +++ b/Mathlib/Topology/ContinuousFunction/Basic.lean @@ -452,7 +452,8 @@ noncomputable def liftCover : C(α, β) := Set.iUnion_eq_univ_iff.2 fun x ↦ (hS x).imp fun _ ↦ mem_of_mem_nhds mk (Set.liftCover S (fun i ↦ φ i) hφ H) <| continuous_of_cover_nhds hS fun i ↦ by rw [continuousOn_iff_continuous_restrict] - simpa only [Set.restrict, Set.liftCover_coe] using (φ i).continuous + simpa (config := { unfoldPartialApp := true }) only [Set.restrict, Set.liftCover_coe] using + (φ i).continuous #align continuous_map.lift_cover ContinuousMap.liftCover variable {S φ hφ hS} diff --git a/Mathlib/Topology/EMetricSpace/Paracompact.lean b/Mathlib/Topology/EMetricSpace/Paracompact.lean index 1dfa424744a91..b4494c9f093de 100644 --- a/Mathlib/Topology/EMetricSpace/Paracompact.lean +++ b/Mathlib/Topology/EMetricSpace/Paracompact.lean @@ -45,7 +45,7 @@ instance (priority := 100) instParacompactSpace [PseudoEMetricSpace α] : Paraco have hpow_le : ∀ {m n : ℕ}, m ≤ n → (2⁻¹ : ℝ≥0∞) ^ n ≤ 2⁻¹ ^ m := @fun m n h => pow_le_pow_of_le_one' (ENNReal.inv_le_one.2 ENNReal.one_lt_two.le) h have h2pow : ∀ n : ℕ, 2 * (2⁻¹ : ℝ≥0∞) ^ (n + 1) = 2⁻¹ ^ n := fun n => by - simp [pow_succ, ← mul_assoc, ENNReal.mul_inv_cancel] + simp [pow_succ, ← mul_assoc, ENNReal.mul_inv_cancel two_ne_zero two_ne_top] -- Consider an open covering `S : Set (Set α)` refine' ⟨fun ι s ho hcov => _⟩ simp only [iUnion_eq_univ_iff] at hcov diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 9e4e8d5ef2391..2ba78c3935530 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -309,7 +309,7 @@ protected theorem tendsto_atTop_zero [Nonempty β] [SemilatticeSup β] {f : β theorem tendsto_sub : ∀ {a b : ℝ≥0∞}, (a ≠ ∞ ∨ b ≠ ∞) → Tendsto (fun p : ℝ≥0∞ × ℝ≥0∞ => p.1 - p.2) (𝓝 (a, b)) (𝓝 (a - b)) - | ⊤, ⊤, h => by simp only at h + | ⊤, ⊤, h => by simp only [ne_eq, not_true_eq_false, or_self] at h | ⊤, (b : ℝ≥0), _ => by rw [top_sub_coe, tendsto_nhds_top_iff_nnreal] refine fun x => ((lt_mem_nhds <| @coe_lt_top (b + 1 + x)).prod_nhds @@ -346,11 +346,11 @@ protected theorem tendsto_mul (ha : a ≠ 0 ∨ b ≠ ⊤) (hb : b ≠ 0 ∨ a refine' this.mono fun c hc => _ exact (ENNReal.div_mul_cancel hε.ne' coe_ne_top).symm.trans_lt (mul_lt_mul hc.1 hc.2) induction a using recTopCoe with - | top => simp only [ne_eq, or_false] at hb; simp [ht b hb, top_mul hb] + | top => simp only [ne_eq, or_false, not_true_eq_false] at hb; simp [ht b hb, top_mul hb] | coe a => induction b using recTopCoe with | top => - simp only [ne_eq, or_false] at ha + simp only [ne_eq, or_false, not_true_eq_false] at ha simpa [(· ∘ ·), mul_comm, mul_top ha] using (ht a ha).comp (continuous_swap.tendsto (ofNNReal a, ⊤)) | coe b => @@ -1459,7 +1459,7 @@ theorem continuous_of_le_add_edist {f : α → ℝ≥0∞} (C : ℝ≥0∞) (hC #align continuous_of_le_add_edist continuous_of_le_add_edist theorem continuous_edist : Continuous fun p : α × α => edist p.1 p.2 := by - apply continuous_of_le_add_edist 2 (by norm_num) + apply continuous_of_le_add_edist 2 (by decide) rintro ⟨x, y⟩ ⟨x', y'⟩ calc edist x y ≤ edist x x' + edist x' y' + edist y' y := edist_triangle4 _ _ _ _ diff --git a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean index 8e2df25c6f8fd..47dff7ad50d9d 100644 --- a/Mathlib/Topology/MetricSpace/GromovHausdorff.lean +++ b/Mathlib/Topology/MetricSpace/GromovHausdorff.lean @@ -1002,8 +1002,6 @@ def auxGluing (n : ℕ) : AuxGluingStruct (X n) := isom := (toGlueR_isometry _ _).comp (isometry_optimalGHInjr (X n) (X (n + 1))) } #align Gromov_Hausdorff.aux_gluing GromovHausdorff.auxGluing -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue lean4#2220 - /-- The Gromov-Hausdorff space is complete. -/ instance : CompleteSpace GHSpace := by set d := fun n : ℕ ↦ ((1 : ℝ) / 2) ^ n diff --git a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean index 1a0d27b39f5be..b9925100f71eb 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDimension.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDimension.lean @@ -519,7 +519,7 @@ theorem dimH_ball_pi (x : ι → ℝ) {r : ℝ} (hr : 0 < r) : exact fun x _ y _ => Subsingleton.elim x y · rw [← ENNReal.coe_nat] have : μH[Fintype.card ι] (Metric.ball x r) = ENNReal.ofReal ((2 * r) ^ Fintype.card ι) := by - rw [hausdorffMeasure_pi_real, Real.volume_pi_ball _ hr, rpow_nat_cast] + rw [hausdorffMeasure_pi_real, Real.volume_pi_ball _ hr] refine dimH_of_hausdorffMeasure_ne_zero_ne_top ?_ ?_ <;> rw [NNReal.coe_nat_cast, this] · simp [pow_pos (mul_pos (zero_lt_two' ℝ) hr)] · exact ENNReal.ofReal_ne_top diff --git a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean index 4820d1b3b9a9c..7bbf2eabe93e4 100644 --- a/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean +++ b/Mathlib/Topology/MetricSpace/ThickenedIndicator.lean @@ -78,7 +78,7 @@ theorem thickenedIndicatorAux_lt_top {δ : ℝ} {E : Set α} {x : α} : theorem thickenedIndicatorAux_closure_eq (δ : ℝ) (E : Set α) : thickenedIndicatorAux δ (closure E) = thickenedIndicatorAux δ E := by - simp_rw [thickenedIndicatorAux, infEdist_closure] + simp (config := { unfoldPartialApp := true }) only [thickenedIndicatorAux, infEdist_closure] #align thickened_indicator_aux_closure_eq thickenedIndicatorAux_closure_eq theorem thickenedIndicatorAux_one (δ : ℝ) (E : Set α) {x : α} (x_in_E : x ∈ E) : diff --git a/Mathlib/Topology/Sequences.lean b/Mathlib/Topology/Sequences.lean index 485c7a537b8be..89c87c6a4a56c 100644 --- a/Mathlib/Topology/Sequences.lean +++ b/Mathlib/Topology/Sequences.lean @@ -163,7 +163,8 @@ theorem FrechetUrysohnSpace.of_seq_tendsto_imp_tendsto · exact subset_seqClosure hx · obtain ⟨u, hux, hus⟩ : ∃ u : ℕ → X, Tendsto u atTop (𝓝 x) ∧ ∃ᶠ x in atTop, u x ∈ s · simpa only [ContinuousAt, hx, tendsto_nhds_true, (· ∘ ·), ← not_frequently, exists_prop, - ← mem_closure_iff_frequently, hcx, imp_false, not_forall, not_not] using h (· ∉ s) x + ← mem_closure_iff_frequently, hcx, imp_false, not_forall, not_not, not_false_eq_true, + not_true_eq_false] using h (· ∉ s) x rcases extraction_of_frequently_atTop hus with ⟨φ, φ_mono, hφ⟩ exact ⟨u ∘ φ, hφ, hux.comp φ_mono.tendsto_atTop⟩ #align frechet_urysohn_space.of_seq_tendsto_imp_tendsto FrechetUrysohnSpace.of_seq_tendsto_imp_tendsto diff --git a/Mathlib/Topology/Sheaves/Presheaf.lean b/Mathlib/Topology/Sheaves/Presheaf.lean index 335b2751453c5..9fc3a4158b33e 100644 --- a/Mathlib/Topology/Sheaves/Presheaf.lean +++ b/Mathlib/Topology/Sheaves/Presheaf.lean @@ -343,7 +343,7 @@ def pullbackObjObjOfImageOpen {X Y : TopCat.{v}} (f : X ⟶ Y) (ℱ : Y.Presheaf · refine' (homOfLE _).op apply (Set.image_subset f s.pt.hom.unop.le).trans exact Set.image_preimage.l_u_le (SetLike.coe s.pt.left.unop) - · simp [eq_iff_true_of_subsingleton] + · simp [autoParam, eq_iff_true_of_subsingleton] -- porting note : add `fac`, `uniq` manually fac := fun _ _ => by ext; simp [eq_iff_true_of_subsingleton] uniq := fun _ _ _ => by ext; simp [eq_iff_true_of_subsingleton] } diff --git a/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean b/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean index a3500ad9b6068..4f13510ceb9df 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean @@ -283,7 +283,7 @@ theorem objSupIsoProdEqLocus_inv_eq_iff {X : TopCat.{u}} (F : X.Sheaf CommRingCa constructor · rintro rfl rw [← TopCat.Sheaf.objSupIsoProdEqLocus_inv_fst, ← TopCat.Sheaf.objSupIsoProdEqLocus_inv_snd] - simp only [← comp_apply, ← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp] + simp only [← comp_apply, ← Functor.map_comp, ← op_comp, Category.assoc, homOfLE_comp, and_self] · rintro ⟨e₁, e₂⟩ refine' F.eq_of_locally_eq₂ (homOfLE (inf_le_right : U ⊓ W ≤ W)) (homOfLE (inf_le_right : V ⊓ W ≤ W)) _ _ _ _ _ diff --git a/Mathlib/Topology/UniformSpace/Pi.lean b/Mathlib/Topology/UniformSpace/Pi.lean index 45bb942f3be67..8b3ff531a0c7e 100644 --- a/Mathlib/Topology/UniformSpace/Pi.lean +++ b/Mathlib/Topology/UniformSpace/Pi.lean @@ -88,7 +88,8 @@ lemma Pi.uniformContinuous_restrict (S : Set ι) : lemma Pi.uniformSpace_comap_restrict (S : Set ι) : UniformSpace.comap (S.restrict) (Pi.uniformSpace (fun i : S ↦ α i)) = ⨅ i ∈ S, UniformSpace.comap (eval i) (U i) := by - simp [← iInf_subtype'', ← uniformSpace_comap_precomp' _ ((↑) : S → ι), Set.restrict] + simp (config := { unfoldPartialApp := true }) + [← iInf_subtype'', ← uniformSpace_comap_precomp' _ ((↑) : S → ι), Set.restrict] lemma cauchy_pi_iff [Nonempty ι] {l : Filter (∀ i, α i)} : Cauchy l ↔ ∀ i, Cauchy (map (eval i) l) := by diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index 7941694614d8a..de9ecf145514d 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -914,7 +914,7 @@ def toFiberPrebundle (a : VectorPrebundle R F E) : FiberPrebundle F E := rw [e.target_inter_preimage_symm_source_eq e', inter_comm] refine' (continuousOn_fst.prod this).congr _ rintro ⟨b, f⟩ ⟨hb, -⟩ - dsimp only [Function.comp, Prod.map] + dsimp only [Function.comp_def, Prod.map] rw [a.mk_coordChange _ _ hb, e'.mk_symm hb.1] } #align vector_prebundle.to_fiber_prebundle VectorPrebundle.toFiberPrebundle diff --git a/lake-manifest.json b/lake-manifest.json index d92d661054bba..ecc867233c1ed 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -1,44 +1,50 @@ -{"version": 6, - "packagesDir": "lake-packages", +{"version": 7, + "packagesDir": ".lake/packages", "packages": - [{"git": - {"url": "https://github.com/leanprover-community/quote4", - "subDir?": null, - "rev": "d3a1d25f3eba0d93a58d5d3d027ffa78ece07755", - "opts": {}, - "name": "Qq", - "inputRev?": "master", - "inherited": false}}, - {"git": - {"url": "https://github.com/leanprover-community/aesop", - "subDir?": null, - "rev": "bf5ab42a58e71de7ebad399ce3f90d29aae7fca9", - "opts": {}, - "name": "aesop", - "inputRev?": "master", - "inherited": false}}, - {"git": - {"url": "https://github.com/leanprover/lean4-cli", - "subDir?": null, - "rev": "a751d21d4b68c999accb6fc5d960538af26ad5ec", - "opts": {}, - "name": "Cli", - "inputRev?": "main", - "inherited": false}}, - {"git": - {"url": "https://github.com/leanprover-community/ProofWidgets4", - "subDir?": null, - "rev": "c3b9f0d4ebedc43635d3f7e764e277b1010844b7", - "opts": {}, - "name": "proofwidgets", - "inputRev?": "v0.0.22", - "inherited": false}}, - {"git": - {"url": "https://github.com/leanprover/std4", - "subDir?": null, - "rev": "c14f6a65b2c08caa38e1ab5db83451460d6cde3e", - "opts": {}, - "name": "std", - "inputRev?": "main", - "inherited": false}}], - "name": "mathlib"} + [{"url": "https://github.com/leanprover/std4", + "type": "git", + "subDir": null, + "rev": "01f7ec0b0c5470f5c72b93decd8011262fd82f40", + "name": "std", + "manifestFile": "lake-manifest.json", + "inputRev": "main", + "inherited": false, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/quote4", + "type": "git", + "subDir": null, + "rev": "d3a1d25f3eba0d93a58d5d3d027ffa78ece07755", + "name": "Qq", + "manifestFile": "lake-manifest.json", + "inputRev": "master", + "inherited": false, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/aesop", + "type": "git", + "subDir": null, + "rev": "a76881668efab6e48a4f3333ad7ba894e8095b04", + "name": "aesop", + "manifestFile": "lake-manifest.json", + "inputRev": "nightly-testing", + "inherited": false, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover/lean4-cli", + "type": "git", + "subDir": null, + "rev": "39229f3630d734af7d9cfb5937ddc6b41d3aa6aa", + "name": "Cli", + "manifestFile": "lake-manifest.json", + "inputRev": "nightly", + "inherited": false, + "configFile": "lakefile.lean"}, + {"url": "https://github.com/leanprover-community/ProofWidgets4", + "type": "git", + "subDir": null, + "rev": "c3b9f0d4ebedc43635d3f7e764e277b1010844b7", + "name": "proofwidgets", + "manifestFile": "lake-manifest.json", + "inputRev": "v0.0.22", + "inherited": false, + "configFile": "lakefile.lean"}], + "name": "mathlib", + "lakeDir": ".lake"} diff --git a/lakefile.lean b/lakefile.lean index e3bee10517980..7551c32b86a8c 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -48,8 +48,8 @@ require «doc-gen4» from git "https://github.com/leanprover/doc-gen4" @ "main" require std from git "https://github.com/leanprover/std4" @ "main" require Qq from git "https://github.com/leanprover-community/quote4" @ "master" -require aesop from git "https://github.com/leanprover-community/aesop" @ "master" -require Cli from git "https://github.com/leanprover/lean4-cli" @ "main" +require aesop from git "https://github.com/leanprover-community/aesop" @ "nightly-testing" +require Cli from git "https://github.com/leanprover/lean4-cli" @ "nightly" require proofwidgets from git "https://github.com/leanprover-community/ProofWidgets4" @ "v0.0.22" lean_lib Cache where diff --git a/lean-toolchain b/lean-toolchain index e8560170ab7d7..24a3cdb8ecae2 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.3.0-rc1 +leanprover/lean4:v4.3.0-rc2 diff --git a/test/Clean.lean b/test/Clean.lean index 6742f46bb9bd5..c93c77f7a781f 100644 --- a/test/Clean.lean +++ b/test/Clean.lean @@ -19,22 +19,22 @@ info: def Tests.x' : Id Nat := #guard_msgs in #print x' -- def x : Id Nat := 1 -def withClean : 2 + 2 = 4 := clean% by exact id rfl -def withoutClean : 2 + 2 = 4 := by exact id rfl +theorem withClean : 2 + 2 = 4 := clean% by exact id rfl +theorem withoutClean : 2 + 2 = 4 := by exact id rfl /-- -info: theorem Tests.withClean.proof_1 : 2 + 2 = 2 + 2 := +info: theorem Tests.withClean : 2 + 2 = 4 := rfl -/ #guard_msgs in -#print Tests.withClean.proof_1 +#print Tests.withClean /-- -info: theorem Tests.withoutClean.proof_1 : 2 + 2 = 4 := +info: theorem Tests.withoutClean : 2 + 2 = 4 := id rfl -/ #guard_msgs in -#print Tests.withoutClean.proof_1 +#print Tests.withoutClean example : True := by let x : id Nat := by dsimp; exact 1 diff --git a/test/RewriteSearch/Basic.lean b/test/RewriteSearch/Basic.lean index 7ca2eacd390d4..d06ecc1d270af 100644 --- a/test/RewriteSearch/Basic.lean +++ b/test/RewriteSearch/Basic.lean @@ -50,6 +50,6 @@ def makeSingleton : Nat → List Nat #guard_msgs in example (n : Nat) : makeSingleton n = [0] := by induction' n with n' ih - · simp + · simp only [makeSingleton] · -- At one point, this failed with: unknown free variable '_uniq.62770' rw_search diff --git a/test/hint.lean b/test/hint.lean index ee177feaa1f0f..047874f9325e3 100644 --- a/test/hint.lean +++ b/test/hint.lean @@ -39,7 +39,7 @@ example : 37^2 - 35^2 = 72 * 2 := by hint /-- info: Try these: -• simp_all only +• decide -/ #guard_msgs in example : Nat.Prime 37 := by hint diff --git a/test/linear_combination.lean b/test/linear_combination.lean index 55e5adab860ed..05d7b32149fd0 100644 --- a/test/linear_combination.lean +++ b/test/linear_combination.lean @@ -140,7 +140,7 @@ example (a b : ℝ) (ha : 2 * a = 4) (hab : 2 * b = a - b) : b = 2 / 3 := by example (x y : ℤ) (h1 : x = -3) (_h2 : y = 10) : 2 * x = -6 := by linear_combination (norm := skip) 2 * h1 - simp + simp (config := {decide := true}) /-! ### Cases without any arguments provided -/ diff --git a/test/matrix.lean b/test/matrix.lean index e13ff85cceab2..deac900999aea 100644 --- a/test/matrix.lean +++ b/test/matrix.lean @@ -139,9 +139,10 @@ 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, Nat.odd_iff_not_even, 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, ne_eq, - 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, + cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, submatrix_apply, + Fin.succ_zero_eq_one, ne_eq, cons_val_one, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, + pow_zero, one_mul, not_true_eq_false, 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 @@ -150,10 +151,10 @@ example {α : Type _} [CommRing α] {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, Nat.odd_iff_not_even, 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, ne_eq, 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, + 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, + ne_eq, cons_val_two, tail_cons, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, + one_mul, not_true_eq_false, 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, even_add_self, Even.neg_pow, one_pow, Finset.sum_const, Finset.card_singleton, one_smul] diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index 172f10c5435b0..bf31496c7dfe6 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -336,18 +336,24 @@ variable {α : Type _} [CommRing α] open BigOperators -- Lists: +-- `by decide` closes the three goals below. +/- example : ([1, 2, 1, 3]).sum = 7 := by norm_num only --- example : (([1, 2, 1, 3] : List ℚ).map (fun i => i^2)).sum = 15 := by norm_num [-List.map] --TODO example : (List.range 10).sum = 45 := by norm_num only example : (List.finRange 10).sum = 45 := by norm_num only +-/ +-- example : (([1, 2, 1, 3] : List ℚ).map (fun i => i^2)).sum = 15 := by norm_num [-List.map] --TODO -- Multisets: +-- `by decide` closes the three goals below. +/- example : (1 ::ₘ 2 ::ₘ 1 ::ₘ 3 ::ₘ {}).sum = 7 := by norm_num only example : ((1 ::ₘ 2 ::ₘ 1 ::ₘ 3 ::ₘ {}).map (fun i => i^2)).sum = 15 := by norm_num only --- example : (({1, 2, 1, 3} : Multiset ℚ).map (fun i => i^2)).sum = 15 := by -- TODO --- norm_num [-Multiset.map_cons] example : (Multiset.range 10).sum = 45 := by norm_num only example : (↑[1, 2, 1, 3] : Multiset ℕ).sum = 7 := by norm_num only +-/ +-- example : (({1, 2, 1, 3} : Multiset ℚ).map (fun i => i^2)).sum = 15 := by -- TODO +-- norm_num [-Multiset.map_cons] -- Finsets: example : Finset.prod (Finset.cons 2 ∅ (Finset.not_mem_empty _)) (λ x => x) = 2 := by norm_num1 diff --git a/test/positivity.lean b/test/positivity.lean index 27e19c43e5726..028cdba4cb6eb 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -188,7 +188,7 @@ example [LinearOrderedSemifield α] (a : α) : 0 < a ^ (0 : ℤ) := by positivit example {a b : ℝ} (ha : 0 ≤ a) : 0 ≤ a ^ b := by positivity example {a b : ℝ} (ha : 0 < a) : 0 < a ^ b := by positivity -example {a : ℝ≥0} {b : ℝ} (ha : 0 < a) : 0 < a ^ b := by positivity +example {a : ℝ≥0} {b : ℝ} (ha : 0 < a) : 0 < (a : ℝ) ^ b := by positivity -- example {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a ^ b := by positivity -- example {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := by positivity example {a : ℝ} : 0 < a ^ 0 := by positivity