From f641187bde40fff19d4da5406fd0c707bc8be96b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 3 Oct 2024 22:46:45 +0000 Subject: [PATCH] chore: bump toolchain to v4.13.0-rc1 (#17377) This merges the already reviewed `bump/v4.13.0` branch. Co-authored-by: Johan Commelin Co-authored-by: leanprover-community-mathlib4-bot --- .../IfNormalization/WithoutAesop.lean | 10 +- Cache/IO.lean | 2 +- Cache/Requests.lean | 2 +- Mathlib/Algebra/AddTorsor.lean | 1 - Mathlib/Algebra/BigOperators/Group/List.lean | 13 +-- Mathlib/Algebra/CharZero/Defs.lean | 1 - .../ContinuedFractions/ConvergentsEquiv.lean | 3 - Mathlib/Algebra/Field/Defs.lean | 4 +- Mathlib/Algebra/Free.lean | 1 - Mathlib/Algebra/Group/Hom/Basic.lean | 2 +- Mathlib/Algebra/Group/Hom/Defs.lean | 5 +- Mathlib/Algebra/Group/Subgroup/Basic.lean | 12 +-- Mathlib/Algebra/Group/Subgroup/Finite.lean | 8 +- .../Algebra/Group/Submonoid/Operations.lean | 2 +- .../Group/Subsemigroup/Operations.lean | 2 +- Mathlib/Algebra/Group/ZeroOne.lean | 13 +-- .../GroupWithZero/NonZeroDivisors.lean | 2 +- .../Homology/HomologicalBicomplex.lean | 2 +- Mathlib/Algebra/Homology/TotalComplex.lean | 4 +- .../Algebra/Homology/TotalComplexShift.lean | 4 +- Mathlib/Algebra/Module/Submodule/Lattice.lean | 2 +- Mathlib/Algebra/NeZero.lean | 26 +----- .../Order/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Order/Floor.lean | 22 ++++- Mathlib/Algebra/Order/Group/Abs.lean | 4 +- Mathlib/Algebra/Order/Group/Cone.lean | 4 +- Mathlib/Algebra/Order/Group/Defs.lean | 10 +- Mathlib/Algebra/Order/Hom/Basic.lean | 40 +++++--- Mathlib/Algebra/Order/Monoid/Prod.lean | 2 +- Mathlib/Algebra/Order/Ring/Cone.lean | 2 +- Mathlib/Algebra/Order/ZeroLEOne.lean | 1 - Mathlib/Algebra/Polynomial/BigOperators.lean | 2 +- .../Polynomial/Degree/CardPowDegree.lean | 2 +- Mathlib/Algebra/Ring/CentroidHom.lean | 4 +- Mathlib/Algebra/Ring/Parity.lean | 6 +- Mathlib/Algebra/Ring/Subring/Basic.lean | 2 +- Mathlib/Algebra/Ring/Subsemiring/Basic.lean | 4 +- Mathlib/Algebra/Star/Free.lean | 2 +- Mathlib/AlgebraicGeometry/Modules/Tilde.lean | 9 +- .../Morphisms/QuasiCompact.lean | 10 +- Mathlib/AlgebraicGeometry/OpenImmersion.lean | 8 +- .../ProjectiveSpectrum/Scheme.lean | 9 +- Mathlib/AlgebraicGeometry/Pullbacks.lean | 5 +- Mathlib/Analysis/Analytic/Composition.lean | 4 +- .../Instances.lean | 12 +-- .../NonUnital.lean | 2 +- .../ContinuousFunctionalCalculus/Unital.lean | 4 +- Mathlib/Analysis/Calculus/Deriv/Basic.lean | 2 +- Mathlib/Analysis/Calculus/FDeriv/Mul.lean | 10 +- .../Analysis/InnerProductSpace/OfNorm.lean | 2 +- .../NormedSpace/Multilinear/Basic.lean | 2 +- .../Analysis/SpecialFunctions/Integrals.lean | 5 +- .../SpecialFunctions/Trigonometric/Basic.lean | 2 +- .../CategoryTheory/ConnectedComponents.lean | 1 - Mathlib/CategoryTheory/Extensive.lean | 4 +- .../Functor/KanExtension/Adjunction.lean | 2 +- .../CategoryTheory/GradedObject/Monoidal.lean | 2 +- .../LiftingProperties/Basic.lean | 2 +- .../CategoryTheory/Limits/Shapes/Images.lean | 2 - .../Limits/Shapes/Pullback/CommSq.lean | 16 ++-- .../MorphismProperty/Basic.lean | 16 ++-- .../Sites/Coherent/SheafComparison.lean | 20 +++- .../CategoryTheory/Sites/Grothendieck.lean | 8 +- .../CategoryTheory/Triangulated/Functor.lean | 2 +- .../Combinatorics/Additive/AP/Three/Defs.lean | 4 +- Mathlib/Combinatorics/Configuration.lean | 2 +- .../Combinatorics/Enumerative/DyckWord.lean | 5 +- .../SetFamily/KruskalKatona.lean | 12 ++- .../SimpleGraph/Hamiltonian.lean | 6 +- Mathlib/Combinatorics/SimpleGraph/Path.lean | 2 +- .../SimpleGraph/Regularity/Equitabilise.lean | 8 +- .../SimpleGraph/Triangle/Basic.lean | 6 +- Mathlib/Combinatorics/SimpleGraph/Walk.lean | 3 +- Mathlib/Combinatorics/Young/YoungDiagram.lean | 2 +- Mathlib/Computability/Ackermann.lean | 6 +- Mathlib/Computability/Halting.lean | 4 +- Mathlib/Computability/Language.lean | 2 +- Mathlib/Computability/PartrecCode.lean | 8 +- Mathlib/Computability/TuringMachine.lean | 2 +- Mathlib/Control/Applicative.lean | 3 +- Mathlib/Control/Basic.lean | 11 +-- Mathlib/Control/Functor.lean | 4 +- Mathlib/Control/Traversable/Basic.lean | 1 + Mathlib/Control/Traversable/Equiv.lean | 2 +- Mathlib/Control/Traversable/Instances.lean | 4 +- Mathlib/Data/Array/ExtractLemmas.lean | 8 +- Mathlib/Data/DFinsupp/Order.lean | 9 +- Mathlib/Data/Fin/Basic.lean | 31 +------ Mathlib/Data/Fin/Tuple/Basic.lean | 2 +- Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean | 2 +- Mathlib/Data/Finset/Basic.lean | 4 +- Mathlib/Data/Finset/Functor.lean | 9 +- Mathlib/Data/Finset/NatDivisors.lean | 6 +- Mathlib/Data/Fintype/Card.lean | 2 - Mathlib/Data/Fintype/Prod.lean | 2 - Mathlib/Data/Int/Defs.lean | 8 +- Mathlib/Data/List/Basic.lean | 92 +++++-------------- Mathlib/Data/List/Chain.lean | 24 ++++- Mathlib/Data/List/Cycle.lean | 4 +- Mathlib/Data/List/Dedup.lean | 2 +- Mathlib/Data/List/Defs.lean | 1 + Mathlib/Data/List/Forall2.lean | 8 +- Mathlib/Data/List/GetD.lean | 7 +- Mathlib/Data/List/Indexes.lean | 23 ++--- Mathlib/Data/List/Infix.lean | 15 ++- Mathlib/Data/List/InsertNth.lean | 6 +- Mathlib/Data/List/Intervals.lean | 4 +- Mathlib/Data/List/Lattice.lean | 2 +- Mathlib/Data/List/Lemmas.lean | 8 -- Mathlib/Data/List/MinMax.lean | 22 +++-- Mathlib/Data/List/Nodup.lean | 35 ++++--- Mathlib/Data/List/NodupEquivFin.lean | 2 +- Mathlib/Data/List/Perm.lean | 25 +---- Mathlib/Data/List/Permutation.lean | 6 +- Mathlib/Data/List/Rotate.lean | 10 +- Mathlib/Data/List/Sort.lean | 36 ++------ Mathlib/Data/List/Sublists.lean | 12 +-- Mathlib/Data/List/Sym.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 29 ++++-- Mathlib/Data/Multiset/Functor.lean | 4 +- Mathlib/Data/Nat/BitIndices.lean | 4 +- Mathlib/Data/Nat/Bits.lean | 15 ++- Mathlib/Data/Nat/Bitwise.lean | 23 +---- Mathlib/Data/Nat/Cast/NeZero.lean | 1 - Mathlib/Data/Nat/Defs.lean | 18 ++-- Mathlib/Data/Nat/Digits.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 8 +- Mathlib/Data/Nat/Factorization/Defs.lean | 2 +- Mathlib/Data/Nat/Log.lean | 8 +- Mathlib/Data/Nat/WithBot.lean | 6 +- Mathlib/Data/Num/Lemmas.lean | 2 +- Mathlib/Data/Option/Basic.lean | 35 ------- Mathlib/Data/Option/Defs.lean | 4 - Mathlib/Data/Ordmap/Ordset.lean | 4 +- Mathlib/Data/PFunctor/Univariate/M.lean | 4 + Mathlib/Data/PNat/Defs.lean | 1 - Mathlib/Data/Prod/Basic.lean | 4 +- Mathlib/Data/Prod/Lex.lean | 4 +- Mathlib/Data/Rat/Lemmas.lean | 12 +-- Mathlib/Data/Seq/WSeq.lean | 7 +- Mathlib/Data/Set/MemPartition.lean | 1 - Mathlib/Data/Setoid/Basic.lean | 2 +- Mathlib/Data/Vector/Defs.lean | 2 +- Mathlib/Data/ZMod/Defs.lean | 2 +- Mathlib/Dynamics/PeriodicPts.lean | 4 +- .../IsAlgClosed/AlgebraicClosure.lean | 2 +- Mathlib/FieldTheory/KummerExtension.lean | 5 + Mathlib/FieldTheory/PurelyInseparable.lean | 2 +- Mathlib/FieldTheory/Separable.lean | 2 +- Mathlib/FieldTheory/SeparableDegree.lean | 2 +- .../Geometry/RingedSpace/OpenImmersion.lean | 44 ++++----- .../RingedSpace/PresheafedSpace/Gluing.lean | 31 +++---- Mathlib/GroupTheory/CommutingProbability.lean | 1 + Mathlib/GroupTheory/FreeGroup/Basic.lean | 5 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 6 ++ Mathlib/GroupTheory/OrderOfElement.lean | 2 +- Mathlib/GroupTheory/Perm/Cycle/Concrete.lean | 2 +- Mathlib/GroupTheory/Perm/Sign.lean | 2 +- .../LinearAlgebra/Alternating/DomCoprod.lean | 2 +- .../CliffordAlgebra/Grading.lean | 6 +- Mathlib/LinearAlgebra/Dual.lean | 2 +- .../LinearAlgebra/FiniteDimensional/Defs.lean | 2 +- Mathlib/LinearAlgebra/Matrix/Adjugate.lean | 4 +- .../LinearAlgebra/Matrix/Charpoly/Coeff.lean | 2 +- .../Matrix/Determinant/Basic.lean | 4 +- Mathlib/LinearAlgebra/Semisimple.lean | 2 +- .../TensorProduct/Graded/Internal.lean | 2 +- Mathlib/Logic/Basic.lean | 2 +- Mathlib/Logic/Encodable/Basic.lean | 6 +- Mathlib/Logic/Equiv/Array.lean | 2 +- Mathlib/Logic/Nonempty.lean | 7 -- .../Constructions/BorelSpace/Order.lean | 12 +-- Mathlib/MeasureTheory/Measure/Content.lean | 2 +- Mathlib/ModelTheory/Algebra/Ring/Basic.lean | 2 +- Mathlib/ModelTheory/Basic.lean | 4 +- Mathlib/ModelTheory/Encoding.lean | 4 +- Mathlib/ModelTheory/Syntax.lean | 4 +- Mathlib/NumberTheory/ADEInequality.lean | 2 +- Mathlib/NumberTheory/Bertrand.lean | 6 +- .../NumberTheory/Cyclotomic/Discriminant.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 24 ++++- .../LegendreSymbol/GaussEisensteinLemmas.lean | 6 +- .../LegendreSymbol/JacobiSymbol.lean | 10 +- .../NumberField/Units/DirichletTheorem.lean | 2 +- Mathlib/NumberTheory/PythagoreanTriples.lean | 2 +- Mathlib/NumberTheory/SmoothNumbers.lean | 2 +- Mathlib/NumberTheory/SumFourSquares.lean | 9 +- Mathlib/Order/Directed.lean | 2 +- Mathlib/Order/Filter/Bases.lean | 4 +- Mathlib/Order/Hom/Basic.lean | 3 +- Mathlib/Order/Hom/Bounded.lean | 6 +- Mathlib/Order/Interval/Finset/Nat.lean | 8 +- Mathlib/Order/RelIso/Basic.lean | 2 +- .../Order/SuccPred/LinearLocallyFinite.lean | 9 +- Mathlib/Order/SymmDiff.lean | 2 +- .../ProbabilityMassFunction/Basic.lean | 3 +- .../ProbabilityMassFunction/Monad.lean | 6 +- .../GroupCohomology/LowDegree.lean | 2 +- Mathlib/RingTheory/AdjoinRoot.lean | 4 +- .../IsIntegralClosure/Basic.lean | 6 +- Mathlib/RingTheory/LocalRing/Module.lean | 26 +++--- .../RingTheory/Localization/FractionRing.lean | 2 +- .../Symmetric/NewtonIdentities.lean | 6 +- .../NonUnitalSubsemiring/Basic.lean | 2 +- .../Polynomial/Cyclotomic/Eval.lean | 9 +- Mathlib/RingTheory/RingHomProperties.lean | 6 +- Mathlib/RingTheory/SimpleRing/Basic.lean | 2 +- Mathlib/RingTheory/Smooth/StandardSmooth.lean | 4 +- Mathlib/SetTheory/Cardinal/Finite.lean | 8 +- Mathlib/Tactic/CC/Addition.lean | 8 +- Mathlib/Tactic/IntervalCases.lean | 1 + Mathlib/Tactic/NormNum/DivMod.lean | 6 +- Mathlib/Tactic/Ring/Basic.lean | 2 +- Mathlib/Tactic/WLOG.lean | 3 +- Mathlib/Testing/SlimCheck/Functions.lean | 2 +- .../Topology/Algebra/Valued/NormedValued.lean | 2 +- .../Category/LightProfinite/Extend.lean | 4 +- .../Topology/Category/Profinite/Extend.lean | 4 +- Mathlib/Topology/ContinuousMap/Basic.lean | 4 +- Mathlib/Topology/MetricSpace/Dilation.lean | 2 +- Mathlib/Topology/Metrizable/Uniformity.lean | 4 +- Mathlib/Topology/Order/LawsonTopology.lean | 2 +- Mathlib/Topology/UniformSpace/Cauchy.lean | 2 +- Mathlib/Util/CountHeartbeats.lean | 4 +- lake-manifest.json | 10 +- lakefile.lean | 5 +- lean-toolchain | 2 +- scripts/noshake.json | 3 +- test/ValuedCSP.lean | 2 + test/aesop_cat.lean | 2 + test/matrix.lean | 4 +- test/says.lean | 3 +- 232 files changed, 809 insertions(+), 761 deletions(-) diff --git a/Archive/Examples/IfNormalization/WithoutAesop.lean b/Archive/Examples/IfNormalization/WithoutAesop.lean index 1ac14973a51065..b3e94e0a5653e5 100644 --- a/Archive/Examples/IfNormalization/WithoutAesop.lean +++ b/Archive/Examples/IfNormalization/WithoutAesop.lean @@ -92,8 +92,8 @@ def normalize' (l : AList (fun _ : ℕ => Bool)) : · simp_all · have := ht₃ v have := he₃ v - simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_true', - AList.lookup_insert_eq_none, ne_eq, AList.lookup_insert] + simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_eq_eq_not, + Bool.not_true, AList.lookup_insert_eq_none, ne_eq, AList.lookup_insert] obtain ⟨⟨⟨tn, tc⟩, tr⟩, td⟩ := ht₂ split <;> rename_i h' · subst h' @@ -103,9 +103,9 @@ def normalize' (l : AList (fun _ : ℕ => Bool)) : have := he₃ w by_cases h : w = v · subst h; simp_all - · simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_true', - AList.lookup_insert_eq_none, ne_eq, not_false_eq_true, AList.lookup_insert_ne, - implies_true] + · simp_all? says simp_all only [normalized, Bool.and_eq_true, Bool.not_eq_eq_eq_not, + Bool.not_true, AList.lookup_insert_eq_none, ne_eq, not_false_eq_true, + AList.lookup_insert_ne, implies_true] obtain ⟨⟨⟨en, ec⟩, er⟩, ed⟩ := he₂ split at b <;> rename_i h' · subst h'; simp_all diff --git a/Cache/IO.lean b/Cache/IO.lean index 6cbd295c310316..43e3f70652623c 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -338,7 +338,7 @@ def packCache (hashMap : HashMap) (overwrite verbose unpackedOnly : Bool) /-- Gets the set of all cached files -/ def getLocalCacheSet : IO <| Lean.RBTree String compare := do let paths ← getFilesWithExtension CACHEDIR "ltar" - return .fromList (paths.data.map (·.withoutParent CACHEDIR |>.toString)) _ + return .fromList (paths.toList.map (·.withoutParent CACHEDIR |>.toString)) _ def isPathFromMathlib (path : FilePath) : Bool := match path.components with diff --git a/Cache/Requests.lean b/Cache/Requests.lean index 5c4c6038aecaa9..71b74cb4b051dd 100644 --- a/Cache/Requests.lean +++ b/Cache/Requests.lean @@ -184,7 +184,7 @@ def UPLOAD_URL : String := /-- Formats the config file for `curl`, containing the list of files to be uploaded -/ def mkPutConfigContent (fileNames : Array String) (token : String) : IO String := do let token := if useFROCache then "" else s!"?{token}" -- the FRO cache doesn't pass the token here - let l ← fileNames.data.mapM fun fileName : String => do + let l ← fileNames.toList.mapM fun fileName : String => do pure s!"-T {(IO.CACHEDIR / fileName).toString}\nurl = {mkFileURL UPLOAD_URL fileName}{token}" return "\n".intercalate l diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 9d3b4b23f9e67e..ed65a54d0f15f9 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -248,7 +248,6 @@ instance instAddTorsor : AddTorsor (G × G') (P × P') where zero_vadd _ := Prod.ext (zero_vadd _ _) (zero_vadd _ _) add_vadd _ _ _ := Prod.ext (add_vadd _ _ _) (add_vadd _ _ _) vsub p₁ p₂ := (p₁.1 -ᵥ p₂.1, p₁.2 -ᵥ p₂.2) - nonempty := Prod.instNonempty vsub_vadd' _ _ := Prod.ext (vsub_vadd _ _) (vsub_vadd _ _) vadd_vsub' _ _ := Prod.ext (vadd_vsub _ _) (vadd_vsub _ _) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 59be51355e078b..6a1a5fb8af68e6 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -125,7 +125,7 @@ theorem prod_replicate (n : ℕ) (a : M) : (replicate n a).prod = a ^ n := by @[to_additive sum_eq_card_nsmul] theorem prod_eq_pow_card (l : List M) (m : M) (h : ∀ x ∈ l, x = m) : l.prod = m ^ l.length := by - rw [← prod_replicate, ← List.eq_replicate.mpr ⟨rfl, h⟩] + rw [← prod_replicate, ← List.eq_replicate_iff.mpr ⟨rfl, h⟩] @[to_additive] theorem prod_hom_rel (l : List ι) {r : M → N → Prop} {f : ι → M} {g : ι → N} (h₁ : r 1 1) @@ -639,20 +639,15 @@ lemma ranges_join (l : List ℕ) : l.ranges.join = range l.sum := by simp [range lemma mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by simp [mem_mem_ranges_iff_lt_natSum] -lemma countP_join (p : α → Bool) : ∀ L : List (List α), countP p L.join = (L.map (countP p)).sum - | [] => rfl - | a :: l => by rw [join, countP_append, map_cons, sum_cons, countP_join _ l] - -lemma count_join [BEq α] (L : List (List α)) (a : α) : L.join.count a = (L.map (count a)).sum := - countP_join _ _ - @[simp] theorem length_bind (l : List α) (f : α → List β) : length (List.bind l f) = sum (map (length ∘ f) l) := by rw [List.bind, length_join, map_map, Nat.sum_eq_listSum] lemma countP_bind (p : β → Bool) (l : List α) (f : α → List β) : - countP p (l.bind f) = sum (map (countP p ∘ f) l) := by rw [List.bind, countP_join, map_map] + countP p (l.bind f) = sum (map (countP p ∘ f) l) := by + rw [List.bind, countP_join, map_map] + simp lemma count_bind [BEq β] (l : List α) (f : α → List β) (x : β) : count x (l.bind f) = sum (map (count x ∘ f) l) := countP_bind _ _ _ diff --git a/Mathlib/Algebra/CharZero/Defs.lean b/Mathlib/Algebra/CharZero/Defs.lean index 8f2cf74d4cef22..c82356e6fa6e7a 100644 --- a/Mathlib/Algebra/CharZero/Defs.lean +++ b/Mathlib/Algebra/CharZero/Defs.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Int.Cast.Defs -import Mathlib.Algebra.NeZero import Mathlib.Logic.Function.Basic /-! diff --git a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean index 2adf5829097015..99e76fa64a5f58 100644 --- a/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean +++ b/Mathlib/Algebra/ContinuedFractions/ConvergentsEquiv.lean @@ -166,9 +166,6 @@ theorem succ_succ_nth_conv'Aux_eq_succ_nth_conv'Aux_squashSeq : gp_head.a / (gp_head.b + convs'Aux s.tail (m + 2)) = convs'Aux (squashSeq s (m + 1)) (m + 2) by simpa only [convs'Aux, s_head_eq] - have : convs'Aux s.tail (m + 2) = convs'Aux (squashSeq s.tail m) (m + 1) := by - refine IH gp_succ_n ?_ - simpa [Stream'.Seq.get?_tail] using s_succ_nth_eq have : (squashSeq s (m + 1)).head = some gp_head := (squashSeq_nth_of_lt m.succ_pos).trans s_head_eq simp_all [convs'Aux, squashSeq_succ_n_tail_eq_squashSeq_tail_n] diff --git a/Mathlib/Algebra/Field/Defs.lean b/Mathlib/Algebra/Field/Defs.lean index 834f25d5fc450a..46e889528e3e65 100644 --- a/Mathlib/Algebra/Field/Defs.lean +++ b/Mathlib/Algebra/Field/Defs.lean @@ -45,8 +45,8 @@ field, division ring, skew field, skew-field, skewfield assert_not_imported Mathlib.Tactic.Common --- `NeZero` should not be needed in the basic algebraic hierarchy. -assert_not_exists NeZero +-- `NeZero` theory should not be needed in the basic algebraic hierarchy +assert_not_imported Mathlib.Algebra.NeZero assert_not_exists MonoidHom diff --git a/Mathlib/Algebra/Free.lean b/Mathlib/Algebra/Free.lean index 1b0f927525cca6..134a4e81ed9ff9 100644 --- a/Mathlib/Algebra/Free.lean +++ b/Mathlib/Algebra/Free.lean @@ -6,7 +6,6 @@ Authors: Kenny Lau import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Control.Applicative import Mathlib.Control.Traversable.Basic -import Mathlib.Data.List.Basic import Mathlib.Logic.Equiv.Defs import Mathlib.Tactic.AdaptationNote diff --git a/Mathlib/Algebra/Group/Hom/Basic.lean b/Mathlib/Algebra/Group/Hom/Basic.lean index 148098efacb87c..094fd70d45cc5b 100644 --- a/Mathlib/Algebra/Group/Hom/Basic.lean +++ b/Mathlib/Algebra/Group/Hom/Basic.lean @@ -14,7 +14,7 @@ import Mathlib.Algebra.Group.Hom.Defs -- `NeZero` cannot be additivised, hence its theory should be developed outside of the -- `Algebra.Group` folder. -assert_not_exists NeZero +assert_not_imported Mathlib.Algebra.NeZero variable {α β M N P : Type*} diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index febcba6ed781c2..2279253985c322 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -143,8 +143,9 @@ homomorphisms. You should also extend this typeclass when you extend `AddMonoidHom`. -/ -class AddMonoidHomClass (F M N : Type*) [AddZeroClass M] [AddZeroClass N] [FunLike F M N] - extends AddHomClass F M N, ZeroHomClass F M N : Prop +class AddMonoidHomClass (F : Type*) (M N : outParam Type*) + [AddZeroClass M] [AddZeroClass N] [FunLike F M N] + extends AddHomClass F M N, ZeroHomClass F M N : Prop -- Instances and lemmas are defined below through `@[to_additive]`. end add_zero diff --git a/Mathlib/Algebra/Group/Subgroup/Basic.lean b/Mathlib/Algebra/Group/Subgroup/Basic.lean index 2df0f53da5529f..556df038ddca9b 100644 --- a/Mathlib/Algebra/Group/Subgroup/Basic.lean +++ b/Mathlib/Algebra/Group/Subgroup/Basic.lean @@ -94,27 +94,27 @@ variable {A : Type*} [AddGroup A] section SubgroupClass /-- `InvMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under inverses. -/ -class InvMemClass (S G : Type*) [Inv G] [SetLike S G] : Prop where +class InvMemClass (S : Type*) (G : outParam Type*) [Inv G] [SetLike S G] : Prop where /-- `s` is closed under inverses -/ inv_mem : ∀ {s : S} {x}, x ∈ s → x⁻¹ ∈ s export InvMemClass (inv_mem) /-- `NegMemClass S G` states `S` is a type of subsets `s ⊆ G` closed under negation. -/ -class NegMemClass (S G : Type*) [Neg G] [SetLike S G] : Prop where +class NegMemClass (S : Type*) (G : outParam Type*) [Neg G] [SetLike S G] : Prop where /-- `s` is closed under negation -/ neg_mem : ∀ {s : S} {x}, x ∈ s → -x ∈ s export NegMemClass (neg_mem) /-- `SubgroupClass S G` states `S` is a type of subsets `s ⊆ G` that are subgroups of `G`. -/ -class SubgroupClass (S G : Type*) [DivInvMonoid G] [SetLike S G] extends SubmonoidClass S G, - InvMemClass S G : Prop +class SubgroupClass (S : Type*) (G : outParam Type*) [DivInvMonoid G] [SetLike S G] + extends SubmonoidClass S G, InvMemClass S G : Prop /-- `AddSubgroupClass S G` states `S` is a type of subsets `s ⊆ G` that are additive subgroups of `G`. -/ -class AddSubgroupClass (S G : Type*) [SubNegMonoid G] [SetLike S G] extends AddSubmonoidClass S G, - NegMemClass S G : Prop +class AddSubgroupClass (S : Type*) (G : outParam Type*) [SubNegMonoid G] [SetLike S G] + extends AddSubmonoidClass S G, NegMemClass S G : Prop attribute [to_additive] InvMemClass SubgroupClass diff --git a/Mathlib/Algebra/Group/Subgroup/Finite.lean b/Mathlib/Algebra/Group/Subgroup/Finite.lean index b4dc8b267698dd..6e5342a357f4ee 100644 --- a/Mathlib/Algebra/Group/Subgroup/Finite.lean +++ b/Mathlib/Algebra/Group/Subgroup/Finite.lean @@ -165,9 +165,11 @@ theorem pi_mem_of_mulSingle_mem_aux [DecidableEq η] (I : Finset η) {H : Subgro x ∈ H := by induction I using Finset.induction_on generalizing x with | empty => - convert one_mem H - ext i - exact h1 i (Finset.not_mem_empty i) + have : x = 1 := by + ext i + exact h1 i (Finset.not_mem_empty i) + rw [this] + exact one_mem H | insert hnmem ih => rename_i i I have : x = Function.update x i 1 * Pi.mulSingle i (x i) := by diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 1bf4b67f902588..f84160423a5b5c 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -591,7 +591,7 @@ theorem closure_closure_coe_preimage {s : Set M} : closure (((↑) : closure s Subtype.recOn x fun x hx _ => by refine closure_induction' (p := fun y hy ↦ (⟨y, hy⟩ : closure s) ∈ closure (((↑) : closure s → M) ⁻¹' s)) - (fun g hg => subset_closure hg) ?_ (fun g₁ g₂ hg₁ hg₂ => ?_) hx + _ (fun g hg => subset_closure hg) ?_ (fun g₁ g₂ hg₁ hg₂ => ?_) hx · exact Submonoid.one_mem _ · exact Submonoid.mul_mem _ diff --git a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean index 05de353785319d..85f523619d7ca3 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean @@ -501,7 +501,7 @@ theorem closure_closure_coe_preimage {s : Set M} : eq_top_iff.2 fun x => Subtype.recOn x fun _ hx' _ => closure_induction' (p := fun y hy ↦ (⟨y, hy⟩ : closure s) ∈ closure (((↑) : closure s → M) ⁻¹' s)) - (fun _ hg => subset_closure hg) (fun _ _ _ _ => Subsemigroup.mul_mem _) hx' + _ (fun _ hg => subset_closure hg) (fun _ _ _ _ => Subsemigroup.mul_mem _) hx' /-- Given `Subsemigroup`s `s`, `t` of semigroups `M`, `N` respectively, `s × t` as a subsemigroup of `M × N`. -/ diff --git a/Mathlib/Algebra/Group/ZeroOne.lean b/Mathlib/Algebra/Group/ZeroOne.lean index 3e729680956219..8822f97d1d17ed 100644 --- a/Mathlib/Algebra/Group/ZeroOne.lean +++ b/Mathlib/Algebra/Group/ZeroOne.lean @@ -6,17 +6,10 @@ Authors: Gabriel Ebner, Mario Carneiro import Mathlib.Tactic.ToAdditive /-! -## Classes for `Zero` and `One` --/ - -class Zero.{u} (α : Type u) where - zero : α +## Typeclass `One` -instance (priority := 300) Zero.toOfNat0 {α} [Zero α] : OfNat α (nat_lit 0) where - ofNat := ‹Zero α›.1 - -instance (priority := 200) Zero.ofOfNat0 {α} [OfNat α (nat_lit 0)] : Zero α where - zero := 0 +`Zero` has already been defined in Lean. +-/ universe u diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index 3c09328f414ca3..2c4e1832cb5a37 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -304,7 +304,7 @@ theorem mk_mem_nonZeroDivisors_associates : Associates.mk a ∈ (Associates M₀ /-- The non-zero divisors of associates of a monoid with zero `M₀` are isomorphic to the associates of the non-zero divisors of `M₀` under the map `⟨⟦a⟧, _⟩ ↦ ⟦⟨a, _⟩⟧`. -/ def associatesNonZeroDivisorsEquiv : (Associates M₀)⁰ ≃* Associates M₀⁰ where - toEquiv := .subtypeQuotientEquivQuotientSubtype (s₂ := Associated.setoid _) + toEquiv := .subtypeQuotientEquivQuotientSubtype _ (s₂ := Associated.setoid _) (· ∈ nonZeroDivisors _) (by simp [mem_nonZeroDivisors_iff, Quotient.forall, Associates.mk_mul_mk]) (by simp [Associated.setoid]) diff --git a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean index b02aad6a648b84..dbd85f9d8e7100 100644 --- a/Mathlib/Algebra/Homology/HomologicalBicomplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalBicomplex.lean @@ -205,7 +205,7 @@ def XXIsoOfEq {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : I₂} (h₂ : @[simp] lemma XXIsoOfEq_rfl (i₁ : I₁) (i₂ : I₂) : - K.XXIsoOfEq (rfl : i₁ = i₁) (rfl : i₂ = i₂) = Iso.refl _ := rfl + K.XXIsoOfEq _ _ _ (rfl : i₁ = i₁) (rfl : i₂ = i₂) = Iso.refl _ := rfl end HomologicalComplex₂ diff --git a/Mathlib/Algebra/Homology/TotalComplex.lean b/Mathlib/Algebra/Homology/TotalComplex.lean index 73fb3959181c0e..645d747de064fd 100644 --- a/Mathlib/Algebra/Homology/TotalComplex.lean +++ b/Mathlib/Algebra/Homology/TotalComplex.lean @@ -260,7 +260,7 @@ noncomputable def ιTotal (i₁ : I₁) (i₂ : I₂) (i₁₂ : I₁₂) @[reassoc (attr := simp)] lemma XXIsoOfEq_hom_ιTotal {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : I₂} (h₂ : x₂ = y₂) (i₁₂ : I₁₂) (h : ComplexShape.π c₁ c₂ c₁₂ (y₁, y₂) = i₁₂) : - (K.XXIsoOfEq h₁ h₂).hom ≫ K.ιTotal c₁₂ y₁ y₂ i₁₂ h = + (K.XXIsoOfEq _ _ _ h₁ h₂).hom ≫ K.ιTotal c₁₂ y₁ y₂ i₁₂ h = K.ιTotal c₁₂ x₁ x₂ i₁₂ (by rw [h₁, h₂, h]) := by subst h₁ h₂ simp @@ -268,7 +268,7 @@ lemma XXIsoOfEq_hom_ιTotal {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : @[reassoc (attr := simp)] lemma XXIsoOfEq_inv_ιTotal {x₁ y₁ : I₁} (h₁ : x₁ = y₁) {x₂ y₂ : I₂} (h₂ : x₂ = y₂) (i₁₂ : I₁₂) (h : ComplexShape.π c₁ c₂ c₁₂ (x₁, x₂) = i₁₂) : - (K.XXIsoOfEq h₁ h₂).inv ≫ K.ιTotal c₁₂ x₁ x₂ i₁₂ h = + (K.XXIsoOfEq _ _ _ h₁ h₂).inv ≫ K.ιTotal c₁₂ x₁ x₂ i₁₂ h = K.ιTotal c₁₂ y₁ y₂ i₁₂ (by rw [← h, h₁, h₂]) := by subst h₁ h₂ simp diff --git a/Mathlib/Algebra/Homology/TotalComplexShift.lean b/Mathlib/Algebra/Homology/TotalComplexShift.lean index 8d32945310df22..345677aa579340 100644 --- a/Mathlib/Algebra/Homology/TotalComplexShift.lean +++ b/Mathlib/Algebra/Homology/TotalComplexShift.lean @@ -129,7 +129,7 @@ noncomputable def totalShift₁XIso (n n' : ℤ) (h : n + x = n') : (((shiftFunctor₁ C x).obj K).total (up ℤ)).X n ≅ (K.total (up ℤ)).X n' where hom := totalDesc _ (fun p q hpq => K.ιTotal (up ℤ) (p + x) q n' (by dsimp at hpq ⊢; omega)) inv := totalDesc _ (fun p q hpq => - (K.XXIsoOfEq (Int.sub_add_cancel p x) rfl).inv ≫ + (K.XXIsoOfEq _ _ _ (Int.sub_add_cancel p x) rfl).inv ≫ ((shiftFunctor₁ C x).obj K).ιTotal (up ℤ) (p - x) q n (by dsimp at hpq ⊢; omega)) hom_inv_id := by @@ -235,7 +235,7 @@ noncomputable def totalShift₂XIso (n n' : ℤ) (h : n + y = n') : hom := totalDesc _ (fun p q hpq => (p * y).negOnePow • K.ιTotal (up ℤ) p (q + y) n' (by dsimp at hpq ⊢; omega)) inv := totalDesc _ (fun p q hpq => (p * y).negOnePow • - (K.XXIsoOfEq rfl (Int.sub_add_cancel q y)).inv ≫ + (K.XXIsoOfEq _ _ _ rfl (Int.sub_add_cancel q y)).inv ≫ ((shiftFunctor₂ C y).obj K).ιTotal (up ℤ) p (q - y) n (by dsimp at hpq ⊢; omega)) hom_inv_id := by ext p q h diff --git a/Mathlib/Algebra/Module/Submodule/Lattice.lean b/Mathlib/Algebra/Module/Submodule/Lattice.lean index ce617ef5b2a6b0..3e15b4721bc188 100644 --- a/Mathlib/Algebra/Module/Submodule/Lattice.lean +++ b/Mathlib/Algebra/Module/Submodule/Lattice.lean @@ -291,7 +291,7 @@ theorem toAddSubmonoid_sSup (s : Set (Submodule R M)) : { toAddSubmonoid := sSup (toAddSubmonoid '' s) smul_mem' := fun t {m} h ↦ by simp_rw [AddSubsemigroup.mem_carrier, AddSubmonoid.mem_toSubsemigroup, sSup_eq_iSup'] at h ⊢ - refine AddSubmonoid.iSup_induction' + refine AddSubmonoid.iSup_induction' _ (C := fun x _ ↦ t • x ∈ ⨆ p : toAddSubmonoid '' s, (p : AddSubmonoid M)) ?_ ?_ (fun x y _ _ ↦ ?_) h · rintro ⟨-, ⟨p : Submodule R M, hp : p ∈ s, rfl⟩⟩ x (hx : x ∈ p) diff --git a/Mathlib/Algebra/NeZero.lean b/Mathlib/Algebra/NeZero.lean index aaa15d8a6209ed..3fdf3e370cb387 100644 --- a/Mathlib/Algebra/NeZero.lean +++ b/Mathlib/Algebra/NeZero.lean @@ -10,32 +10,12 @@ import Mathlib.Order.Defs /-! # `NeZero` typeclass -We create a typeclass `NeZero n` which carries around the fact that `(n : R) ≠ 0`. +We give basic facts about the `NeZero n` typeclass. -## Main declarations - -* `NeZero`: `n ≠ 0` as a typeclass. -/ variable {R : Type*} [Zero R] -/-- A type-class version of `n ≠ 0`. -/ -class NeZero (n : R) : Prop where - /-- The proposition that `n` is not zero. -/ - out : n ≠ 0 - -theorem NeZero.ne (n : R) [h : NeZero n] : n ≠ 0 := - h.out - -theorem NeZero.ne' (n : R) [h : NeZero n] : 0 ≠ n := - h.out.symm - -theorem neZero_iff {n : R} : NeZero n ↔ n ≠ 0 := - ⟨fun h ↦ h.out, NeZero.mk⟩ - -@[simp] lemma neZero_zero_iff_false {α : Type*} [Zero α] : NeZero (0 : α) ↔ False := - ⟨fun h ↦ h.ne rfl, fun h ↦ h.elim⟩ - theorem not_neZero {n : R} : ¬NeZero n ↔ n = 0 := by simp [neZero_iff] theorem eq_zero_or_neZero (a : R) : a = 0 ∨ NeZero a := @@ -77,10 +57,6 @@ namespace NeZero variable {M : Type*} {x : M} -instance succ {n : ℕ} : NeZero (n + 1) := ⟨n.succ_ne_zero⟩ - theorem of_pos [Preorder M] [Zero M] (h : 0 < x) : NeZero x := ⟨ne_of_gt h⟩ end NeZero - -lemma Nat.pos_of_neZero (n : ℕ) [NeZero n] : 0 < n := Nat.pos_of_ne_zero (NeZero.ne _) diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 3d5a671d2a7f8f..900c60857b3b13 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -164,7 +164,7 @@ variable [CanonicallyOrderedCommMonoid M] {l : List M} @[to_additive] lemma prod_eq_one_iff : l.prod = 1 ↔ ∀ x ∈ l, x = (1 : M) := ⟨all_one_of_le_one_le_of_prod_eq_one fun _ _ => one_le _, fun h => by - rw [List.eq_replicate.2 ⟨_, h⟩, prod_replicate, one_pow] + rw [List.eq_replicate_iff.2 ⟨_, h⟩, prod_replicate, one_pow] · exact (length l) · rfl⟩ diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index 18797cc6170d3f..6bd51a8b46a96e 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -222,13 +222,21 @@ theorem floor_eq_zero : ⌊a⌋₊ = 0 ↔ a < 1 := by rw [← lt_one_iff, ← @cast_one α] exact floor_lt' Nat.one_ne_zero +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem floor_eq_iff (ha : 0 ≤ a) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff ha, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt ha, Nat.lt_add_one_iff, - le_antisymm_iff, and_comm] + le_antisymm_iff, _root_.and_comm] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem floor_eq_iff' (hn : n ≠ 0) : ⌊a⌋₊ = n ↔ ↑n ≤ a ∧ a < ↑n + 1 := by rw [← le_floor_iff' hn, ← Nat.cast_one, ← Nat.cast_add, ← floor_lt' (Nat.add_one_ne_zero n), - Nat.lt_add_one_iff, le_antisymm_iff, and_comm] + Nat.lt_add_one_iff, le_antisymm_iff, _root_.and_comm] theorem floor_eq_on_Ico (n : ℕ) : ∀ a ∈ (Set.Ico n (n + 1) : Set α), ⌊a⌋₊ = n := fun _ ⟨h₀, h₁⟩ => (floor_eq_iff <| n.cast_nonneg.trans h₀).mpr ⟨h₀, h₁⟩ @@ -320,10 +328,14 @@ theorem floor_lt_ceil_of_lt_of_pos {a b : α} (h : a < b) (h' : 0 < b) : ⌊a⌋ exact h.trans_le (le_ceil _) · rwa [floor_of_nonpos ha.le, lt_ceil, Nat.cast_zero] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem ceil_eq_iff (hn : n ≠ 0) : ⌈a⌉₊ = n ↔ ↑(n - 1) < a ∧ a ≤ n := by rw [← ceil_le, ← not_le, ← ceil_le, not_le, tsub_lt_iff_right (Nat.add_one_le_iff.2 (pos_iff_ne_zero.2 hn)), Nat.lt_add_one_iff, - le_antisymm_iff, and_comm] + le_antisymm_iff, _root_.and_comm] @[simp] theorem preimage_ceil_zero : (Nat.ceil : α → ℕ) ⁻¹' {0} = Iic 0 := @@ -1054,7 +1066,7 @@ theorem fract_div_intCast_eq_div_intCast_mod {m : ℤ} {n : ℕ} : obtain ⟨l₀, rfl | rfl⟩ := l.eq_nat_or_neg · rw [cast_natCast, ← natCast_mod, cast_natCast, fract_div_natCast_eq_div_natCast_mod] · rw [Right.nonneg_neg_iff, natCast_nonpos_iff] at hl - simp [hl, zero_mod] + simp [hl] obtain ⟨m₀, rfl | rfl⟩ := m.eq_nat_or_neg · exact this (ofNat_nonneg m₀) let q := ⌈↑m₀ / (n : k)⌉ @@ -1694,4 +1706,4 @@ def evalIntCeil : PositivityExt where eval {u α} _zα _pα e := do end Mathlib.Meta.Positivity -set_option linter.style.longFile 1700 +set_option linter.style.longFile 1800 diff --git a/Mathlib/Algebra/Order/Group/Abs.lean b/Mathlib/Algebra/Order/Group/Abs.lean index 3d6763c037cfa2..9693c189cfb6c7 100644 --- a/Mathlib/Algebra/Order/Group/Abs.lean +++ b/Mathlib/Algebra/Order/Group/Abs.lean @@ -94,7 +94,7 @@ theorem apply_abs_le_mul_of_one_le {β : Type*} [MulOneClass β] [Preorder β] theorem abs_add (a b : α) : |a + b| ≤ |a| + |b| := abs_le.2 ⟨(neg_add |a| |b|).symm ▸ - add_le_add ((@neg_le α ..).2 <| neg_le_abs _) ((@neg_le α ..).2 <| neg_le_abs _), + add_le_add (neg_le.2 <| neg_le_abs _) (neg_le.2 <| neg_le_abs _), add_le_add (le_abs_self _) (le_abs_self _)⟩ theorem abs_add' (a b : α) : |a| ≤ |b| + |b + a| := by simpa using abs_add (-b) (b + a) @@ -122,7 +122,7 @@ theorem sub_lt_of_abs_sub_lt_right (h : |a - b| < c) : a - c < b := sub_lt_of_abs_sub_lt_left (abs_sub_comm a b ▸ h) theorem abs_sub_abs_le_abs_sub (a b : α) : |a| - |b| ≤ |a - b| := - (@sub_le_iff_le_add α ..).2 <| + sub_le_iff_le_add.2 <| calc |a| = |a - b + b| := by rw [sub_add_cancel] _ ≤ |a - b| + |b| := abs_add _ _ diff --git a/Mathlib/Algebra/Order/Group/Cone.lean b/Mathlib/Algebra/Order/Group/Cone.lean index a49307108ac8ec..59f273599f45a4 100644 --- a/Mathlib/Algebra/Order/Group/Cone.lean +++ b/Mathlib/Algebra/Order/Group/Cone.lean @@ -18,8 +18,8 @@ cones in groups and the corresponding ordered groups. -/ /-- `AddGroupConeClass S G` says that `S` is a type of cones in `G`. -/ -class AddGroupConeClass (S G : Type*) [AddCommGroup G] [SetLike S G] extends - AddSubmonoidClass S G : Prop where +class AddGroupConeClass (S : Type*) (G : outParam Type*) [AddCommGroup G] [SetLike S G] + extends AddSubmonoidClass S G : Prop where eq_zero_of_mem_of_neg_mem {C : S} {a : G} : a ∈ C → -a ∈ C → a = 0 /-- `GroupConeClass S G` says that `S` is a type of cones in `G`. -/ diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index a47229cb3489fa..53b53337c45d3d 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -21,9 +21,9 @@ The reason is that we did not want to change existing names in the library. -/ /- -`NeZero` should not be needed at this point in the ordered algebraic hierarchy. +`NeZero` theory should not be needed at this point in the ordered algebraic hierarchy. -/ -assert_not_exists NeZero +assert_not_imported Mathlib.Algebra.NeZero open Function @@ -175,13 +175,11 @@ variable [OrderedCommGroup α] {a b : α} @[to_additive (attr := gcongr) neg_le_neg] theorem inv_le_inv' : a ≤ b → b⁻¹ ≤ a⁻¹ := - -- Porting note: explicit type annotation was not needed before. - (@inv_le_inv_iff α ..).mpr + inv_le_inv_iff.mpr @[to_additive (attr := gcongr) neg_lt_neg] theorem inv_lt_inv' : a < b → b⁻¹ < a⁻¹ := - -- Porting note: explicit type annotation was not needed before. - (@inv_lt_inv_iff α ..).mpr + inv_lt_inv_iff.mpr -- The additive version is also a `linarith` lemma. @[to_additive] diff --git a/Mathlib/Algebra/Order/Hom/Basic.lean b/Mathlib/Algebra/Order/Hom/Basic.lean index dbcb880887ff30..b97b19b03d22d4 100644 --- a/Mathlib/Algebra/Order/Hom/Basic.lean +++ b/Mathlib/Algebra/Order/Hom/Basic.lean @@ -73,29 +73,33 @@ variable {ι F α β γ δ : Type*} /-! ### Basics -/ /-- `NonnegHomClass F α β` states that `F` is a type of nonnegative morphisms. -/ -class NonnegHomClass (F α β : Type*) [Zero β] [LE β] [FunLike F α β] : Prop where +class NonnegHomClass (F : Type*) (α β : outParam Type*) [Zero β] [LE β] [FunLike F α β] : Prop where /-- the image of any element is non negative. -/ apply_nonneg (f : F) : ∀ a, 0 ≤ f a /-- `SubadditiveHomClass F α β` states that `F` is a type of subadditive morphisms. -/ -class SubadditiveHomClass (F α β : Type*) [Add α] [Add β] [LE β] [FunLike F α β] : Prop where +class SubadditiveHomClass (F : Type*) (α β : outParam Type*) + [Add α] [Add β] [LE β] [FunLike F α β] : Prop where /-- the image of a sum is less or equal than the sum of the images. -/ map_add_le_add (f : F) : ∀ a b, f (a + b) ≤ f a + f b /-- `SubmultiplicativeHomClass F α β` states that `F` is a type of submultiplicative morphisms. -/ @[to_additive SubadditiveHomClass] -class SubmultiplicativeHomClass (F α β : Type*) [Mul α] [Mul β] [LE β] [FunLike F α β] : Prop where +class SubmultiplicativeHomClass (F : Type*) (α β : outParam (Type*)) [Mul α] [Mul β] [LE β] + [FunLike F α β] : Prop where /-- the image of a product is less or equal than the product of the images. -/ map_mul_le_mul (f : F) : ∀ a b, f (a * b) ≤ f a * f b /-- `MulLEAddHomClass F α β` states that `F` is a type of subadditive morphisms. -/ @[to_additive SubadditiveHomClass] -class MulLEAddHomClass (F α β : Type*) [Mul α] [Add β] [LE β] [FunLike F α β] : Prop where +class MulLEAddHomClass (F : Type*) (α β : outParam Type*) [Mul α] [Add β] [LE β] [FunLike F α β] : + Prop where /-- the image of a product is less or equal than the sum of the images. -/ map_mul_le_add (f : F) : ∀ a b, f (a * b) ≤ f a + f b /-- `NonarchimedeanHomClass F α β` states that `F` is a type of non-archimedean morphisms. -/ -class NonarchimedeanHomClass (F α β : Type*) [Add α] [LinearOrder β] [FunLike F α β] : Prop where +class NonarchimedeanHomClass (F : Type*) (α β : outParam Type*) + [Add α] [LinearOrder β] [FunLike F α β] : Prop where /-- the image of a sum is less or equal than the maximum of the images. -/ map_add_le_max (f : F) : ∀ a b, f (a + b) ≤ max (f a) (f b) @@ -154,7 +158,8 @@ end Mathlib.Meta.Positivity group `α`. You should extend this class when you extend `AddGroupSeminorm`. -/ -class AddGroupSeminormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] +class AddGroupSeminormClass (F : Type*) (α β : outParam Type*) + [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] extends SubadditiveHomClass F α β : Prop where /-- The image of zero is zero. -/ map_zero (f : F) : f 0 = 0 @@ -165,7 +170,8 @@ class AddGroupSeminormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoi You should extend this class when you extend `GroupSeminorm`. -/ @[to_additive] -class GroupSeminormClass (F α β : Type*) [Group α] [OrderedAddCommMonoid β] [FunLike F α β] +class GroupSeminormClass (F : Type*) (α β : outParam Type*) + [Group α] [OrderedAddCommMonoid β] [FunLike F α β] extends MulLEAddHomClass F α β : Prop where /-- The image of one is zero. -/ map_one_eq_zero (f : F) : f 1 = 0 @@ -176,7 +182,8 @@ class GroupSeminormClass (F α β : Type*) [Group α] [OrderedAddCommMonoid β] `α`. You should extend this class when you extend `AddGroupNorm`. -/ -class AddGroupNormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] +class AddGroupNormClass (F : Type*) (α β : outParam Type*) + [AddGroup α] [OrderedAddCommMonoid β] [FunLike F α β] extends AddGroupSeminormClass F α β : Prop where /-- The argument is zero if its image under the map is zero. -/ eq_zero_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 0 @@ -185,7 +192,8 @@ class AddGroupNormClass (F α β : Type*) [AddGroup α] [OrderedAddCommMonoid β You should extend this class when you extend `GroupNorm`. -/ @[to_additive] -class GroupNormClass (F α β : Type*) [Group α] [OrderedAddCommMonoid β] [FunLike F α β] +class GroupNormClass (F : Type*) (α β : outParam Type*) + [Group α] [OrderedAddCommMonoid β] [FunLike F α β] extends GroupSeminormClass F α β : Prop where /-- The argument is one if its image under the map is zero. -/ eq_one_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 1 @@ -275,20 +283,23 @@ theorem map_pos_of_ne_one [Group α] [LinearOrderedAddCommMonoid β] [GroupNormC /-- `RingSeminormClass F α` states that `F` is a type of `β`-valued seminorms on the ring `α`. You should extend this class when you extend `RingSeminorm`. -/ -class RingSeminormClass (F α β : Type*) [NonUnitalNonAssocRing α] [OrderedSemiring β] - [FunLike F α β] extends AddGroupSeminormClass F α β, SubmultiplicativeHomClass F α β : Prop +class RingSeminormClass (F : Type*) (α β : outParam Type*) + [NonUnitalNonAssocRing α] [OrderedSemiring β] [FunLike F α β] + extends AddGroupSeminormClass F α β, SubmultiplicativeHomClass F α β : Prop /-- `RingNormClass F α` states that `F` is a type of `β`-valued norms on the ring `α`. You should extend this class when you extend `RingNorm`. -/ -class RingNormClass (F α β : Type*) [NonUnitalNonAssocRing α] [OrderedSemiring β] [FunLike F α β] +class RingNormClass (F : Type*) (α β : outParam Type*) + [NonUnitalNonAssocRing α] [OrderedSemiring β] [FunLike F α β] extends RingSeminormClass F α β, AddGroupNormClass F α β : Prop /-- `MulRingSeminormClass F α` states that `F` is a type of `β`-valued multiplicative seminorms on the ring `α`. You should extend this class when you extend `MulRingSeminorm`. -/ -class MulRingSeminormClass (F α β : Type*) [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] +class MulRingSeminormClass (F : Type*) (α β : outParam Type*) + [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] extends AddGroupSeminormClass F α β, MonoidWithZeroHomClass F α β : Prop -- Lower the priority of these instances since they require synthesizing an order structure. @@ -299,7 +310,8 @@ attribute [instance 50] ring `α`. You should extend this class when you extend `MulRingNorm`. -/ -class MulRingNormClass (F α β : Type*) [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] +class MulRingNormClass (F : Type*) (α β : outParam Type*) + [NonAssocRing α] [OrderedSemiring β] [FunLike F α β] extends MulRingSeminormClass F α β, AddGroupNormClass F α β : Prop -- See note [out-param inheritance] diff --git a/Mathlib/Algebra/Order/Monoid/Prod.lean b/Mathlib/Algebra/Order/Monoid/Prod.lean index c6998f304353cb..e3d9ab5369849c 100644 --- a/Mathlib/Algebra/Order/Monoid/Prod.lean +++ b/Mathlib/Algebra/Order/Monoid/Prod.lean @@ -40,7 +40,7 @@ instance [CanonicallyOrderedCommMonoid α] [CanonicallyOrderedCommMonoid β] : CanonicallyOrderedCommMonoid (α × β) := { (inferInstance : OrderedCommMonoid _), (inferInstance : OrderBot _), (inferInstance : ExistsMulOfLE _) with - le_self_mul := fun _ _ ↦ ⟨le_self_mul, le_self_mul⟩ } + le_self_mul := fun _ _ ↦ le_def.mpr ⟨le_self_mul, le_self_mul⟩ } namespace Lex diff --git a/Mathlib/Algebra/Order/Ring/Cone.lean b/Mathlib/Algebra/Order/Ring/Cone.lean index 4ddd18e2f67d0b..71c21076f00933 100644 --- a/Mathlib/Algebra/Order/Ring/Cone.lean +++ b/Mathlib/Algebra/Order/Ring/Cone.lean @@ -19,7 +19,7 @@ cones in rings and the corresponding ordered rings. -/ /-- `RingConeClass S R` says that `S` is a type of cones in `R`. -/ -class RingConeClass (S R : Type*) [Ring R] [SetLike S R] +class RingConeClass (S : Type*) (R : outParam Type*) [Ring R] [SetLike S R] extends AddGroupConeClass S R, SubsemiringClass S R : Prop /-- A (positive) cone in a ring is a subsemiring that diff --git a/Mathlib/Algebra/Order/ZeroLEOne.lean b/Mathlib/Algebra/Order/ZeroLEOne.lean index ecd0e356d73ada..538c8f6985b425 100644 --- a/Mathlib/Algebra/Order/ZeroLEOne.lean +++ b/Mathlib/Algebra/Order/ZeroLEOne.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import Mathlib.Order.Basic -import Mathlib.Algebra.NeZero /-! # Typeclass expressing `0 ≤ 1`. diff --git a/Mathlib/Algebra/Polynomial/BigOperators.lean b/Mathlib/Algebra/Polynomial/BigOperators.lean index 32db3b30a3a973..cef12c507490ae 100644 --- a/Mathlib/Algebra/Polynomial/BigOperators.lean +++ b/Mathlib/Algebra/Polynomial/BigOperators.lean @@ -68,7 +68,7 @@ theorem degree_list_sum_le (l : List S[X]) : degree l.sum ≤ (l.map natDegree). rw [← List.foldr_max_of_ne_nil] · congr contrapose! h - rw [List.map_eq_nil] at h + rw [List.map_eq_nil_iff] at h simp [h] theorem natDegree_list_prod_le (l : List S[X]) : natDegree l.prod ≤ (l.map natDegree).sum := by diff --git a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean index 1a1dd67ce6467c..01b1c8ec0561ec 100644 --- a/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean +++ b/Mathlib/Algebra/Polynomial/Degree/CardPowDegree.lean @@ -50,7 +50,7 @@ noncomputable def cardPowDegree : AbsoluteValue Fq[X] ℤ := · rfl exact pow_nonneg (Int.ofNat_zero_le _) _ eq_zero' := fun p => - ite_eq_left_iff.trans <| + ite_eq_left_iff.trans ⟨fun h => by contrapose! h exact ⟨h, (pow_pos _).ne'⟩, absurd⟩ diff --git a/Mathlib/Algebra/Ring/CentroidHom.lean b/Mathlib/Algebra/Ring/CentroidHom.lean index 25857d52638e5e..37a70431efb16b 100644 --- a/Mathlib/Algebra/Ring/CentroidHom.lean +++ b/Mathlib/Algebra/Ring/CentroidHom.lean @@ -61,8 +61,8 @@ attribute [nolint docBlame] CentroidHom.toAddMonoidHom /-- `CentroidHomClass F α` states that `F` is a type of centroid homomorphisms. You should extend this class when you extend `CentroidHom`. -/ -class CentroidHomClass (F α : Type*) [NonUnitalNonAssocSemiring α] [FunLike F α α] extends - AddMonoidHomClass F α α : Prop where +class CentroidHomClass (F : Type*) (α : outParam Type*) + [NonUnitalNonAssocSemiring α] [FunLike F α α] extends AddMonoidHomClass F α α : Prop where /-- Commutativity of centroid homomorphims with left multiplication. -/ map_mul_left (f : F) (a b : α) : f (a * b) = a * f b /-- Commutativity of centroid homomorphims with right multiplication. -/ diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 8d540524abfb47..088d5ce242d13c 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -231,10 +231,14 @@ lemma even_or_odd (n : ℕ) : Even n ∨ Odd n := (even_xor_odd n).or lemma even_or_odd' (n : ℕ) : ∃ k, n = 2 * k ∨ n = 2 * k + 1 := by simpa only [← two_mul, exists_or, Odd, Even] using even_or_odd n +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma even_xor_odd' (n : ℕ) : ∃ k, Xor' (n = 2 * k) (n = 2 * k + 1) := by obtain ⟨k, rfl⟩ | ⟨k, rfl⟩ := even_or_odd n <;> use k · simpa only [← two_mul, eq_self_iff_true, xor_true] using (succ_ne_self (2 * k)).symm - · simpa only [xor_true, xor_comm] using (succ_ne_self _) + · simpa only [xor_true, _root_.xor_comm] using (succ_ne_self _) lemma mod_two_add_add_odd_mod_two (m : ℕ) {n : ℕ} (hn : Odd n) : m % 2 + (m + n) % 2 = 1 := ((even_or_odd m).elim fun hm ↦ by rw [even_iff.1 hm, odd_iff.1 (hm.add_odd hn)]) fun hm ↦ by diff --git a/Mathlib/Algebra/Ring/Subring/Basic.lean b/Mathlib/Algebra/Ring/Subring/Basic.lean index dece9ec6923b97..4d2d0317d2fcfb 100644 --- a/Mathlib/Algebra/Ring/Subring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subring/Basic.lean @@ -72,7 +72,7 @@ section SubringClass /-- `SubringClass S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative submonoid and an additive subgroup. -/ -class SubringClass (S : Type*) (R : Type u) [Ring R] [SetLike S R] extends +class SubringClass (S : Type*) (R : outParam (Type u)) [Ring R] [SetLike S R] extends SubsemiringClass S R, NegMemClass S R : Prop -- See note [lower instance priority] diff --git a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean index 2c20870f699012..f775d6a04fdda0 100644 --- a/Mathlib/Algebra/Ring/Subsemiring/Basic.lean +++ b/Mathlib/Algebra/Ring/Subsemiring/Basic.lean @@ -59,12 +59,12 @@ section SubsemiringClass /-- `SubsemiringClass S R` states that `S` is a type of subsets `s ⊆ R` that are both a multiplicative and an additive submonoid. -/ -class SubsemiringClass (S : Type*) (R : Type u) [NonAssocSemiring R] +class SubsemiringClass (S : Type*) (R : outParam (Type u)) [NonAssocSemiring R] [SetLike S R] extends SubmonoidClass S R, AddSubmonoidClass S R : Prop -- See note [lower instance priority] instance (priority := 100) SubsemiringClass.addSubmonoidWithOneClass (S : Type*) - (R : Type u) [NonAssocSemiring R] [SetLike S R] [h : SubsemiringClass S R] : + (R : Type u) {_ : NonAssocSemiring R} [SetLike S R] [h : SubsemiringClass S R] : AddSubmonoidWithOneClass S R := { h with } diff --git a/Mathlib/Algebra/Star/Free.lean b/Mathlib/Algebra/Star/Free.lean index eefbd4ba0c2088..b29ed2137d81e2 100644 --- a/Mathlib/Algebra/Star/Free.lean +++ b/Mathlib/Algebra/Star/Free.lean @@ -48,7 +48,7 @@ instance : StarRing (FreeAlgebra R X) where unfold Star.star simp only [Function.comp_apply] let y := lift R (X := X) (MulOpposite.op ∘ ι R) - apply induction (C := fun x ↦ (y (y x).unop).unop = x) _ _ _ _ x + refine induction (C := fun x ↦ (y (y x).unop).unop = x) _ _ ?_ ?_ ?_ ?_ x · intros simp only [AlgHom.commutes, MulOpposite.algebraMap_apply, MulOpposite.unop_op] · intros diff --git a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean index b81f47d19763c3..b3aa7de64c8527 100644 --- a/Mathlib/AlgebraicGeometry/Modules/Tilde.lean +++ b/Mathlib/AlgebraicGeometry/Modules/Tilde.lean @@ -187,10 +187,11 @@ lemma smul_section_apply (r : R) (U : Opens (PrimeSpectrum.Top R)) lemma smul_stalk_no_nonzero_divisor {x : PrimeSpectrum R} (r : x.asIdeal.primeCompl) (st : (tildeInModuleCat M).stalk x) (hst : r.1 • st = 0) : st = 0 := by - refine Limits.Concrete.colimit_no_zero_smul_divisor (hx := hst) - ⟨op ⟨PrimeSpectrum.basicOpen r.1, r.2⟩, fun U i s hs ↦ Subtype.eq <| funext fun pt ↦ ?_⟩ - exact LocalizedModule.eq_zero_of_smul_eq_zero (hx := congr_fun (Subtype.ext_iff.1 hs) pt) <| - i.unop pt |>.2 + refine Limits.Concrete.colimit_no_zero_smul_divisor + _ _ _ ⟨op ⟨PrimeSpectrum.basicOpen r.1, r.2⟩, fun U i s hs ↦ Subtype.eq <| funext fun pt ↦ ?_⟩ + _ hst + apply LocalizedModule.eq_zero_of_smul_eq_zero _ (i.unop pt).2 _ + (congr_fun (Subtype.ext_iff.1 hs) pt) /-- If `U` is an open subset of `Spec R`, this is the morphism of `R`-modules from `M` to diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index c85901ed45bc93..cf3a6124cc51f5 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -198,7 +198,15 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isAffineOpen (X : Sch {U : X.Opens} (hU : IsAffineOpen U) (x f : Γ(X, U)) (H : x |_ X.basicOpen f = 0) : ∃ n : ℕ, f ^ n * x = 0 := by rw [← map_zero (X.presheaf.map (homOfLE <| X.basicOpen_le f : X.basicOpen f ⟶ U).op)] at H - obtain ⟨n, e⟩ := (hU.isLocalization_basicOpen f).exists_of_eq H + #adaptation_note + /-- + Prior to nightly-2024-09-29, we could use dot notation here: + `(hU.isLocalization_basicOpen f).exists_of_eq H` + This is no longer possible; + likely changing the signature of `IsLocalization.Away.exists_of_eq` is in order. + -/ + obtain ⟨n, e⟩ := + @IsLocalization.Away.exists_of_eq _ _ _ _ _ _ (hU.isLocalization_basicOpen f) _ _ H exact ⟨n, by simpa [mul_comm x] using e⟩ /-- If `x : Γ(X, U)` is zero on `D(f)` for some `f : Γ(X, U)`, and `U` is quasi-compact, then diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/OpenImmersion.lean index c65f7aa086fbc0..5ace362109bfbe 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion.lean @@ -431,7 +431,7 @@ instance pullback_fst_of_right : IsOpenImmersion (pullback.fst g f) := by rw [← pullbackSymmetry_hom_comp_snd] -- Porting note: was just `infer_instance`, it is a bit weird that no explicit class instance is -- provided but still class inference fail to find this - exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ + exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ _ instance pullback_to_base [IsOpenImmersion g] : IsOpenImmersion (limit.π (cospan f g) WalkingCospan.one) := by @@ -439,14 +439,14 @@ instance pullback_to_base [IsOpenImmersion g] : change IsOpenImmersion (_ ≫ f) -- Porting note: was just `infer_instance`, it is a bit weird that no explicit class instance is -- provided but still class inference fail to find this - exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ + exact LocallyRingedSpace.IsOpenImmersion.comp (H := inferInstance) _ _ instance forgetToTopPreservesOfLeft : PreservesLimit (cospan f g) Scheme.forgetToTop := by delta Scheme.forgetToTop - apply @Limits.compPreservesLimit (K := cospan f g) (F := forget) + refine @Limits.compPreservesLimit _ _ _ _ _ _ (K := cospan f g) _ _ (F := forget) (G := LocallyRingedSpace.forgetToTop) ?_ ?_ · infer_instance - apply @preservesLimitOfIsoDiagram (F := _) _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ + refine @preservesLimitOfIsoDiagram _ _ _ _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ dsimp [LocallyRingedSpace.forgetToTop] infer_instance diff --git a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean index 139077e8ace13b..938d2b0845645b 100644 --- a/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Scheme.lean @@ -294,7 +294,12 @@ theorem mem_carrier_iff_of_mem (hm : 0 < m) (q : Spec.T A⁰_ f) (a : A) {n} (hn trans (HomogeneousLocalization.mk ⟨m * n, ⟨proj 𝒜 n a ^ m, by rw [← smul_eq_mul]; mem_tac⟩, ⟨f ^ n, by rw [mul_comm]; mem_tac⟩, ⟨_, rfl⟩⟩ : A⁰_ f) ∈ q.asIdeal · refine ⟨fun h ↦ h n, fun h i ↦ if hi : i = n then hi ▸ h else ?_⟩ - convert zero_mem q.asIdeal + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/5376 + we need to specify the implicit arguments by `inferInstance`. + -/ + convert @zero_mem _ _ inferInstance inferInstance _ q.asIdeal apply HomogeneousLocalization.val_injective simp only [proj_apply, decompose_of_mem_ne _ hn (Ne.symm hi), zero_pow hm.ne', HomogeneousLocalization.val_mk, Localization.mk_zero, HomogeneousLocalization.val_zero] @@ -806,7 +811,7 @@ If `f ∈ A` is a homogeneous element of positive degree, then the projective sp -/ def projIsoSpec (f) {m} (f_deg : f ∈ 𝒜 m) (hm : 0 < m) : (Proj| pbo f) ≅ (Spec (A⁰_ f)) := - @asIso (f := toSpec 𝒜 f) (isIso_toSpec 𝒜 f f_deg hm) + @asIso _ _ _ _ (f := toSpec 𝒜 f) (isIso_toSpec 𝒜 f f_deg hm) /-- This is the scheme `Proj(A)` for any `ℕ`-graded ring `A`. diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index 2dcfa434d2017c..1b43cf9c18bfdf 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -507,7 +507,7 @@ def openCoverOfBase' (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCove pasteVertIsPullback rfl (pullbackIsPullback g (𝒰.map i)) (pullbackIsPullback (pullback.snd g (𝒰.map i)) (pullback.snd f (𝒰.map i))) refine - @openCoverOfIsIso + @openCoverOfIsIso _ _ (f := (pullbackSymmetry _ _).hom ≫ (limit.isoLimitCone ⟨_, this⟩).inv ≫ pullback.map _ _ _ _ (𝟙 _) (𝟙 _) (𝟙 _) ?_ ?_) inferInstance · simp [← pullback.condition] @@ -583,7 +583,8 @@ the morphism `Spec (S ⊗[R] T) ⟶ Spec T` obtained by applying `Spec.map` to t -/ @[reassoc (attr := simp)] lemma pullbackSpecIso_inv_snd : - (pullbackSpecIso R S T).inv ≫ pullback.snd _ _ = Spec.map (ofHom (toRingHom includeRight)) := + (pullbackSpecIso R S T).inv ≫ pullback.snd _ _ = + Spec.map (ofHom (R := T) (S := S ⊗[R] T) (toRingHom includeRight)) := limit.isoLimitCone_inv_π _ _ /-- The composition of the isomorphism `pullbackSepcIso R S T` (from the pullback of diff --git a/Mathlib/Analysis/Analytic/Composition.lean b/Mathlib/Analysis/Analytic/Composition.lean index d9bd84e10d143e..49ef36b0ddbd9a 100644 --- a/Mathlib/Analysis/Analytic/Composition.lean +++ b/Mathlib/Analysis/Analytic/Composition.lean @@ -1043,7 +1043,7 @@ def sigmaCompositionAux (a : Composition n) (b : Composition a.length) a.blocks_pos (by rw [← a.blocks.join_splitWrtComposition b] - exact mem_join_of_mem (List.getElem_mem _ _ _) hi) + exact mem_join_of_mem (List.getElem_mem _) hi) blocks_sum := by simp [Composition.blocksFun, getElem_map, Composition.gather] theorem length_sigmaCompositionAux (a : Composition n) (b : Composition a.length) @@ -1109,7 +1109,7 @@ theorem sizeUpTo_sizeUpTo_add (a : Composition n) (b : Composition a.length) {i have : sizeUpTo b i + Nat.succ j = (sizeUpTo b i + j).succ := rfl rw [this, sizeUpTo_succ _ D, IHj A, sizeUpTo_succ _ B] simp only [sigmaCompositionAux, add_assoc, add_left_inj, Fin.val_mk] - rw [getElem_of_eq (getElem_splitWrtComposition _ _ _ _), getElem_drop, getElem_take _ _ C] + rw [getElem_of_eq (getElem_splitWrtComposition _ _ _ _), getElem_drop, getElem_take' _ _ C] /-- Natural equivalence between `(Σ (a : Composition n), Composition a.length)` and `(Σ (c : Composition n), Π (i : Fin c.length), Composition (c.blocksFun i))`, that shows up as a diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean index a085109299f03e..180ca90ca52884 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Instances.lean @@ -605,7 +605,7 @@ variable {A : Type*} [TopologicalSpace A] [Ring A] [StarRing A] [Algebra ℂ A] lemma cfcHom_real_eq_restrict {a : A} (ha : IsSelfAdjoint a) : cfcHom ha = ha.spectrumRestricts.starAlgHom (cfcHom ha.isStarNormal) (f := Complex.reCLM) := - ha.spectrumRestricts.cfcHom_eq_restrict Complex.isometry_ofReal.isUniformEmbedding + ha.spectrumRestricts.cfcHom_eq_restrict _ Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal lemma cfc_real_eq_complex {a : A} (f : ℝ → ℝ) (ha : IsSelfAdjoint a := by cfc_tac) : @@ -626,7 +626,7 @@ variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [StarRing A] [Module lemma cfcₙHom_real_eq_restrict {a : A} (ha : IsSelfAdjoint a) : cfcₙHom ha = (ha.quasispectrumRestricts.2).nonUnitalStarAlgHom (cfcₙHom ha.isStarNormal) (f := Complex.reCLM) := - ha.quasispectrumRestricts.2.cfcₙHom_eq_restrict Complex.isometry_ofReal.isUniformEmbedding + ha.quasispectrumRestricts.2.cfcₙHom_eq_restrict _ Complex.isometry_ofReal.isUniformEmbedding ha ha.isStarNormal lemma cfcₙ_real_eq_complex {a : A} (f : ℝ → ℝ) (ha : IsSelfAdjoint a := by cfc_tac) : @@ -650,12 +650,12 @@ variable {A : Type*} [TopologicalSpace A] [Ring A] [PartialOrder A] [StarRing A] lemma cfcHom_nnreal_eq_restrict {a : A} (ha : 0 ≤ a) : cfcHom ha = (SpectrumRestricts.nnreal_of_nonneg ha).starAlgHom (cfcHom (IsSelfAdjoint.of_nonneg ha)) := by - apply (SpectrumRestricts.nnreal_of_nonneg ha).cfcHom_eq_restrict isUniformEmbedding_subtype_val + apply (SpectrumRestricts.nnreal_of_nonneg ha).cfcHom_eq_restrict _ isUniformEmbedding_subtype_val lemma cfc_nnreal_eq_real {a : A} (f : ℝ≥0 → ℝ≥0) (ha : 0 ≤ a := by cfc_tac) : cfc f a = cfc (fun x ↦ f x.toNNReal : ℝ → ℝ) a := by replace ha : 0 ≤ a := ha -- hack to avoid issues caused by autoParam - apply (SpectrumRestricts.nnreal_of_nonneg ha).cfc_eq_restrict + apply (SpectrumRestricts.nnreal_of_nonneg ha).cfc_eq_restrict _ isUniformEmbedding_subtype_val ha (.of_nonneg ha) end NNRealEqReal @@ -674,13 +674,13 @@ variable {A : Type*} [TopologicalSpace A] [NonUnitalRing A] [PartialOrder A] [St lemma cfcₙHom_nnreal_eq_restrict {a : A} (ha : 0 ≤ a) : cfcₙHom ha = (QuasispectrumRestricts.nnreal_of_nonneg ha).nonUnitalStarAlgHom (cfcₙHom (IsSelfAdjoint.of_nonneg ha)) := by - apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙHom_eq_restrict + apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙHom_eq_restrict _ isUniformEmbedding_subtype_val lemma cfcₙ_nnreal_eq_real {a : A} (f : ℝ≥0 → ℝ≥0) (ha : 0 ≤ a := by cfc_tac) : cfcₙ f a = cfcₙ (fun x ↦ f x.toNNReal : ℝ → ℝ) a := by replace ha : 0 ≤ a := ha -- hack to avoid issues caused by autoParam - apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙ_eq_restrict + apply (QuasispectrumRestricts.nnreal_of_nonneg ha).cfcₙ_eq_restrict _ isUniformEmbedding_subtype_val ha (.of_nonneg ha) end NNRealEqRealNonUnital diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean index 3bcd0795b06d8e..a298ba04ac176f 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/NonUnital.lean @@ -421,7 +421,7 @@ lemma cfcₙ_comp (g f : R → R) (a : A) ext simp rw [cfcₙ_apply .., cfcₙ_apply f a, - cfcₙ_apply _ (by convert hg) (ha := cfcₙHom_predicate (show p a from ha) _) , + cfcₙ_apply _ _ (by convert hg) (ha := cfcₙHom_predicate (show p a from ha) _), ← cfcₙHom_comp _ _] swap · exact ⟨.mk _ <| hf.restrict.codRestrict fun x ↦ by rw [sp_eq]; use x.1; simp, Subtype.ext hf0⟩ diff --git a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean index 89017072beb7c9..8b6473cba0de3e 100644 --- a/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean +++ b/Mathlib/Analysis/CStarAlgebra/ContinuousFunctionalCalculus/Unital.lean @@ -723,7 +723,7 @@ noncomputable def cfcUnits (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) lemma cfcUnits_pow (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) (n : ℕ) (hf : ContinuousOn f (spectrum R a) := by cfc_cont_tac) (ha : p a := by cfc_tac) : (cfcUnits f a hf') ^ n = - cfcUnits (forall₂_imp (fun _ _ ↦ pow_ne_zero n) hf') (hf := hf.pow n) := by + cfcUnits _ _ (forall₂_imp (fun _ _ ↦ pow_ne_zero n) hf') (hf := hf.pow n) := by ext cases n with | zero => simp [cfc_const_one R a] @@ -778,7 +778,7 @@ lemma cfcUnits_zpow (hf' : ∀ x ∈ spectrum R a, f x ≠ 0) (n : ℤ) | negSucc n => simp only [zpow_negSucc, ← inv_pow] ext - exact cfc_pow (hf := hf.inv₀ hf') _ |>.symm + exact cfc_pow (hf := hf.inv₀ hf') .. |>.symm lemma cfc_zpow (a : Aˣ) (n : ℤ) (ha : p a := by cfc_tac) : cfc (fun x : R ↦ x ^ n) (a : A) = ↑(a ^ n) := by diff --git a/Mathlib/Analysis/Calculus/Deriv/Basic.lean b/Mathlib/Analysis/Calculus/Deriv/Basic.lean index edda887e12373a..5ed888191d0e0b 100644 --- a/Mathlib/Analysis/Calculus/Deriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/Deriv/Basic.lean @@ -416,7 +416,7 @@ theorem norm_deriv_eq_norm_fderiv : ‖deriv f x‖ = ‖fderiv 𝕜 f x‖ := b theorem DifferentiableAt.derivWithin (h : DifferentiableAt 𝕜 f x) (hxs : UniqueDiffWithinAt 𝕜 s x) : derivWithin f s x = deriv f x := by - unfold derivWithin deriv + unfold _root_.derivWithin deriv rw [h.fderivWithin hxs] theorem HasDerivWithinAt.deriv_eq_zero (hd : HasDerivWithinAt f 0 s x) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean index d867ef086d9c76..990bb02bd95999 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean @@ -573,10 +573,10 @@ theorem hasStrictFDerivAt_list_prod_finRange' {n : ℕ} {x : Fin n → 𝔸} : theorem hasStrictFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : {i // i ∈ l} → 𝔸} : HasStrictFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.attach.map x).prod) (∑ i : Fin l.length, ((l.attach.take i).map x).prod • - smulRight (proj l.attach[i.cast l.length_attach.symm]) + smulRight (proj l.attach[i.cast List.length_attach.symm]) ((l.attach.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod'.congr_fderiv <| Eq.symm <| - Finset.sum_equiv (finCongr l.length_attach.symm) (by simp) (by simp) + Finset.sum_equiv (finCongr List.length_attach.symm) (by simp) (by simp) @[fun_prop] theorem hasFDerivAt_list_prod' [Fintype ι] {l : List ι} {x : ι → 𝔸'} : @@ -596,7 +596,7 @@ theorem hasFDerivAt_list_prod_finRange' {n : ℕ} {x : Fin n → 𝔸} : theorem hasFDerivAt_list_prod_attach' [DecidableEq ι] {l : List ι} {x : {i // i ∈ l} → 𝔸} : HasFDerivAt (𝕜 := 𝕜) (fun x ↦ (l.attach.map x).prod) (∑ i : Fin l.length, ((l.attach.take i).map x).prod • - smulRight (proj l.attach[i.cast l.length_attach.symm]) + smulRight (proj l.attach[i.cast List.length_attach.symm]) ((l.attach.drop (.succ i)).map x).prod) x := hasStrictFDerivAt_list_prod_attach'.hasFDerivAt @@ -648,7 +648,7 @@ theorem HasStrictFDerivAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasStrictFDerivAt_list_prod_finRange'.comp x - (hasStrictFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem ..))) ?_ + (hasStrictFDerivAt_pi.mpr fun i ↦ h l[i] (List.getElem_mem ..))) ?_ ext m simp [← List.map_map] @@ -663,7 +663,7 @@ theorem HasFDerivAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasFDerivAt_list_prod_finRange'.comp x - (hasFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem i i.isLt))) ?_ + (hasFDerivAt_pi.mpr fun i ↦ h l[i] (List.getElem_mem i.isLt))) ?_ ext m simp [← List.map_map] diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index b29417e75a1af8..e2e23c0b46a9de 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -101,7 +101,7 @@ variable {E} theorem _root_.Continuous.inner_ {f g : ℝ → E} (hf : Continuous f) (hg : Continuous g) : Continuous fun x => inner_ 𝕜 (f x) (g x) := by - unfold inner_ + unfold _root_.inner_ fun_prop theorem inner_.norm_sq (x : E) : ‖x‖ ^ 2 = re (inner_ 𝕜 x x) := by diff --git a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean index c334bfd9e636b5..770630be496f36 100644 --- a/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean +++ b/Mathlib/Analysis/NormedSpace/Multilinear/Basic.lean @@ -747,7 +747,7 @@ theorem norm_mkPiAlgebraFin_succ_le : ‖ContinuousMultilinearMap.mkPiAlgebraFin simp only [ContinuousMultilinearMap.mkPiAlgebraFin_apply, one_mul, List.ofFn_eq_map, Fin.prod_univ_def, Multiset.map_coe, Multiset.prod_coe] refine (List.norm_prod_le' ?_).trans_eq ?_ - · rw [Ne, List.map_eq_nil, List.finRange_eq_nil] + · rw [Ne, List.map_eq_nil_iff, List.finRange_eq_nil] exact Nat.succ_ne_zero _ rw [List.map_map, Function.comp_def] diff --git a/Mathlib/Analysis/SpecialFunctions/Integrals.lean b/Mathlib/Analysis/SpecialFunctions/Integrals.lean index 506735693822ca..31f51791cdb312 100644 --- a/Mathlib/Analysis/SpecialFunctions/Integrals.lean +++ b/Mathlib/Analysis/SpecialFunctions/Integrals.lean @@ -33,7 +33,7 @@ integrate, integration, integrable, integrability -/ -open Real Nat Set Finset +open Real Set Finset open scoped Real Interval @@ -580,8 +580,9 @@ theorem integral_mul_rpow_one_add_sq {t : ℝ} (ht : t ≠ -1) : end RpowCpow -/-! ### Integral of `sin x ^ n` -/ +open Nat +/-! ### Integral of `sin x ^ n` -/ theorem integral_sin_pow_aux : (∫ x in a..b, sin x ^ (n + 2)) = diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index dfd9f1a9f4692b..64086037384505 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -517,7 +517,7 @@ theorem cos_eq_one_iff (x : ℝ) : cos x = 1 ↔ ∃ n : ℤ, (n : ℝ) * (2 * (Int.emod_two_eq_zero_or_one n).elim (fun hn0 => by rwa [← mul_assoc, ← @Int.cast_two ℝ, ← Int.cast_mul, - Int.ediv_mul_cancel ((Int.dvd_iff_emod_eq_zero _ _).2 hn0)]) + Int.ediv_mul_cancel (Int.dvd_iff_emod_eq_zero.2 hn0)]) fun hn1 => by rw [← Int.emod_add_ediv n 2, hn1, Int.cast_add, Int.cast_one, add_mul, one_mul, add_comm, mul_comm (2 : ℤ), Int.cast_mul, mul_assoc, Int.cast_two] at hn diff --git a/Mathlib/CategoryTheory/ConnectedComponents.lean b/Mathlib/CategoryTheory/ConnectedComponents.lean index cc86ae90a21671..0d205fcbb01705 100644 --- a/Mathlib/CategoryTheory/ConnectedComponents.lean +++ b/Mathlib/CategoryTheory/ConnectedComponents.lean @@ -7,7 +7,6 @@ import Mathlib.Data.List.Chain import Mathlib.CategoryTheory.IsConnected import Mathlib.CategoryTheory.Sigma.Basic import Mathlib.CategoryTheory.FullSubcategory -import Mathlib.Data.List.Infix /-! # Connected components of a category diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 887e41f0debb85..94619f08a21bbb 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -527,7 +527,7 @@ instance FinitaryPreExtensive.hasPullbacks_of_inclusions [FinitaryPreExtensive C {α : Type*} (f : X ⟶ Z) {Y : (a : α) → C} (i : (a : α) → Y a ⟶ Z) [Finite α] [hi : IsIso (Sigma.desc i)] (a : α) : HasPullback f (i a) := by apply FinitaryPreExtensive.hasPullbacks_of_is_coproduct (c := Cofan.mk Z i) - exact @IsColimit.ofPointIso (t := Cofan.mk Z i) (P := _) hi + exact @IsColimit.ofPointIso (t := Cofan.mk Z i) (P := _) (i := hi) lemma FinitaryPreExtensive.sigma_desc_iso [FinitaryPreExtensive C] {α : Type} [Finite α] {X : C} {Z : α → C} (π : (a : α) → Z a ⟶ X) {Y : C} (f : Y ⟶ X) (hπ : IsIso (Sigma.desc π)) : @@ -536,7 +536,7 @@ lemma FinitaryPreExtensive.sigma_desc_iso [FinitaryPreExtensive C] {α : Type} [ change IsIso (this.coconePointUniqueUpToIso (getColimitCocone _).2).inv infer_instance let this : IsColimit (Cofan.mk X π) := by - refine @IsColimit.ofPointIso (t := Cofan.mk X π) (P := coproductIsCoproduct Z) ?_ + refine @IsColimit.ofPointIso (t := Cofan.mk X π) (P := coproductIsCoproduct Z) (i := ?_) convert hπ simp [coproductIsCoproduct] refine (FinitaryPreExtensive.isUniversal_finiteCoproducts this diff --git a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean index 57c9c1c9353e06..fe8433c905fb7b 100644 --- a/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean +++ b/Mathlib/CategoryTheory/Functor/KanExtension/Adjunction.lean @@ -168,7 +168,7 @@ precomposition by `L`. -/ noncomputable def ranAdjunction : (whiskeringLeft C D H).obj L ⊣ L.ran := Adjunction.mkOfHomEquiv { homEquiv := fun F G => - (homEquivOfIsRightKanExtension (α := L.ranCounit.app G) F).symm + (homEquivOfIsRightKanExtension (α := L.ranCounit.app G) _ F).symm homEquiv_naturality_right := fun {F G₁ G₂} β f ↦ hom_ext_of_isRightKanExtension _ (L.ranCounit.app G₂) _ _ (by ext X diff --git a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean index 917484d4564657..b005a034c6b490 100644 --- a/Mathlib/CategoryTheory/GradedObject/Monoidal.lean +++ b/Mathlib/CategoryTheory/GradedObject/Monoidal.lean @@ -325,7 +325,7 @@ lemma left_tensor_tensorObj₃_ext {j : I} {A : C} (Z : C) (_ ◁ ιTensorObj₃ X₁ X₂ X₃ i₁ i₂ i₃ j h) ≫ f = (_ ◁ ιTensorObj₃ X₁ X₂ X₃ i₁ i₂ i₃ j h) ≫ g) : f = g := by refine (@isColimitOfPreserves C _ C _ _ _ _ ((curriedTensor C).obj Z) _ - (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj (H := H) j) hZ).hom_ext ?_ + (isColimitCofan₃MapBifunctorBifunctor₂₃MapObj (H := H) (j := j)) hZ).hom_ext ?_ intro ⟨⟨i₁, i₂, i₃⟩, hi⟩ exact h _ _ _ hi diff --git a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean index 034992348ba842..6b3ab2bd33ac24 100644 --- a/Mathlib/CategoryTheory/LiftingProperties/Basic.lean +++ b/Mathlib/CategoryTheory/LiftingProperties/Basic.lean @@ -41,7 +41,7 @@ class HasLiftingProperty : Prop where sq_hasLift : ∀ {f : A ⟶ X} {g : B ⟶ Y} (sq : CommSq f i p g), sq.HasLift instance (priority := 100) sq_hasLift_of_hasLiftingProperty {f : A ⟶ X} {g : B ⟶ Y} - (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := by apply hip.sq_hasLift + (sq : CommSq f i p g) [hip : HasLiftingProperty i p] : sq.HasLift := hip.sq_hasLift _ namespace HasLiftingProperty diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean index 22aa87a4a18cd5..10e2d68009de0e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Images.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Images.lean @@ -81,8 +81,6 @@ attribute [reassoc (attr := simp)] MonoFactorisation.fac attribute [instance] MonoFactorisation.m_mono -attribute [instance] MonoFactorisation.m_mono - namespace MonoFactorisation /-- The obvious factorisation of a monomorphism through itself. -/ diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean index e22ae9d2d8bed7..423fed53f4080e 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullback/CommSq.lean @@ -280,22 +280,22 @@ noncomputable def isoIsPullback (h : IsPullback fst snd f g) (h' : IsPullback fs @[reassoc (attr := simp)] theorem isoIsPullback_hom_fst (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').hom ≫ fst' = fst := + (h.isoIsPullback _ _ h').hom ≫ fst' = fst := IsLimit.conePointUniqueUpToIso_hom_comp h.isLimit h'.isLimit WalkingCospan.left @[reassoc (attr := simp)] theorem isoIsPullback_hom_snd (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').hom ≫ snd' = snd := + (h.isoIsPullback _ _ h').hom ≫ snd' = snd := IsLimit.conePointUniqueUpToIso_hom_comp h.isLimit h'.isLimit WalkingCospan.right @[reassoc (attr := simp)] theorem isoIsPullback_inv_fst (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').inv ≫ fst = fst' := by + (h.isoIsPullback _ _ h').inv ≫ fst = fst' := by simp only [Iso.inv_comp_eq, isoIsPullback_hom_fst] @[reassoc (attr := simp)] theorem isoIsPullback_inv_snd (h : IsPullback fst snd f g) (h' : IsPullback fst' snd' f g) : - (h.isoIsPullback h').inv ≫ snd = snd' := by + (h.isoIsPullback _ _ h').inv ≫ snd = snd' := by simp only [Iso.inv_comp_eq, isoIsPullback_hom_snd] end @@ -468,22 +468,22 @@ noncomputable def isoIsPushout (h : IsPushout f g inl inr) (h' : IsPushout f g i @[reassoc (attr := simp)] theorem inl_isoIsPushout_hom (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inl ≫ (h.isoIsPushout h').hom = inl' := + inl ≫ (h.isoIsPushout _ _ h').hom = inl' := IsColimit.comp_coconePointUniqueUpToIso_hom h.isColimit h'.isColimit WalkingSpan.left @[reassoc (attr := simp)] theorem inr_isoIsPushout_hom (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inr ≫ (h.isoIsPushout h').hom = inr' := + inr ≫ (h.isoIsPushout _ _ h').hom = inr' := IsColimit.comp_coconePointUniqueUpToIso_hom h.isColimit h'.isColimit WalkingSpan.right @[reassoc (attr := simp)] theorem inl_isoIsPushout_inv (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inl' ≫ (h.isoIsPushout h').inv = inl := by + inl' ≫ (h.isoIsPushout _ _ h').inv = inl := by simp only [Iso.comp_inv_eq, inl_isoIsPushout_hom] @[reassoc (attr := simp)] theorem inr_isoIsPushout_inv (h : IsPushout f g inl inr) (h' : IsPushout f g inl' inr') : - inr' ≫ (h.isoIsPushout h').inv = inr := by + inr' ≫ (h.isoIsPushout _ _ h').inv = inr := by simp only [Iso.comp_inv_eq, inr_isoIsPushout_hom] end diff --git a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean index e9d8b8fb9c0ffa..0a3a382da80167 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean @@ -162,13 +162,13 @@ lemma RespectsIso.postcomp (P : MorphismProperty C) [P.RespectsIso] {X Y Z : C} [IsIso e] (f : X ⟶ Y) (hf : P f) : P (f ≫ e) := RespectsRight.postcomp (Q := isomorphisms C) e ‹IsIso e› f hf -instance RespectsIso.op (P : MorphismProperty C) [h : RespectsIso P] : RespectsIso P.op where - precomp e (_ : IsIso e) f hf := h.postcomp e.unop f.unop hf - postcomp e (_ : IsIso e) f hf := h.precomp e.unop f.unop hf +instance RespectsIso.op (P : MorphismProperty C) [RespectsIso P] : RespectsIso P.op where + precomp e (_ : IsIso e) f hf := postcomp P e.unop f.unop hf + postcomp e (_ : IsIso e) f hf := precomp P e.unop f.unop hf -instance RespectsIso.unop (P : MorphismProperty Cᵒᵖ) [h : RespectsIso P] : RespectsIso P.unop where - precomp e (_ : IsIso e) f hf := h.postcomp e.op f.op hf - postcomp e (_ : IsIso e) f hf := h.precomp e.op f.op hf +instance RespectsIso.unop (P : MorphismProperty Cᵒᵖ) [RespectsIso P] : RespectsIso P.unop where + precomp e (_ : IsIso e) f hf := postcomp P e.op f.op hf + postcomp e (_ : IsIso e) f hf := precomp P e.op f.op hf /-- The closure by isomorphisms of a `MorphismProperty` -/ def isoClosure (P : MorphismProperty C) : MorphismProperty C := @@ -190,11 +190,11 @@ lemma monotone_isoClosure : Monotone (isoClosure (C := C)) := by theorem cancel_left_of_respectsIso (P : MorphismProperty C) [hP : RespectsIso P] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso f] : P (f ≫ g) ↔ P g := - ⟨fun h => by simpa using hP.precomp (inv f) (f ≫ g) h, hP.precomp f g⟩ + ⟨fun h => by simpa using RespectsIso.precomp P (inv f) (f ≫ g) h, RespectsIso.precomp P f g⟩ theorem cancel_right_of_respectsIso (P : MorphismProperty C) [hP : RespectsIso P] {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z) [IsIso g] : P (f ≫ g) ↔ P f := - ⟨fun h => by simpa using hP.postcomp (inv g) (f ≫ g) h, hP.postcomp g f⟩ + ⟨fun h => by simpa using RespectsIso.postcomp P (inv g) (f ≫ g) h, RespectsIso.postcomp P g f⟩ theorem arrow_iso_iff (P : MorphismProperty C) [RespectsIso P] {f g : Arrow C} (e : f ≅ g) : P f.hom ↔ P g.hom := by diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index 7f2817071b11a0..a20deb1ceff5dc 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -85,7 +85,15 @@ lemma eq_induced : haveI := F.reflects_precoherent instance : haveI := F.reflects_precoherent; F.IsDenseSubsite (coherentTopology C) (coherentTopology D) where - functorPushforward_mem_iff := by simp_rw [eq_induced F]; rfl + functorPushforward_mem_iff := by + rw [eq_induced F] + #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + exact Iff.rfl lemma coverPreserving : haveI := F.reflects_precoherent CoverPreserving (coherentTopology _) (coherentTopology _) F := @@ -181,7 +189,15 @@ lemma eq_induced : haveI := F.reflects_preregular instance : haveI := F.reflects_preregular; F.IsDenseSubsite (regularTopology C) (regularTopology D) where - functorPushforward_mem_iff := by simp_rw [eq_induced F]; rfl + functorPushforward_mem_iff := by + rw [eq_induced F] + #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + exact Iff.rfl lemma coverPreserving : haveI := F.reflects_preregular CoverPreserving (regularTopology _) (regularTopology _) F := diff --git a/Mathlib/CategoryTheory/Sites/Grothendieck.lean b/Mathlib/CategoryTheory/Sites/Grothendieck.lean index 3e453496f9a9d6..0d3552925cc54d 100644 --- a/Mathlib/CategoryTheory/Sites/Grothendieck.lean +++ b/Mathlib/CategoryTheory/Sites/Grothendieck.lean @@ -263,7 +263,13 @@ instance : InfSet (GrothendieckTopology C) where /-- See -/ theorem isGLB_sInf (s : Set (GrothendieckTopology C)) : IsGLB s (sInf s) := by refine @IsGLB.of_image _ _ _ _ sieves ?_ _ _ ?_ - · rfl + · #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + exact Iff.rfl · exact _root_.isGLB_sInf _ /-- Construct a complete lattice from the `Inf`, but make the trivial and discrete topologies diff --git a/Mathlib/CategoryTheory/Triangulated/Functor.lean b/Mathlib/CategoryTheory/Triangulated/Functor.lean index d78ec3c589a8f9..9507a7937e954e 100644 --- a/Mathlib/CategoryTheory/Triangulated/Functor.lean +++ b/Mathlib/CategoryTheory/Triangulated/Functor.lean @@ -310,6 +310,6 @@ lemma isTriangulated_of_essSurj_mapComposableArrows_two exact ⟨Octahedron.ofIso (e₁ := (e.app 0).symm) (e₂ := (e.app 1).symm) (e₃ := (e.app 2).symm) (comm₁₂ := ComposableArrows.naturality' e.inv 0 1) (comm₂₃ := ComposableArrows.naturality' e.inv 1 2) - (H := (someOctahedron rfl h₁₂' h₂₃' h₁₃').map F) _ _ _ _ _⟩ + (H := (someOctahedron rfl h₁₂' h₂₃' h₁₃').map F) ..⟩ end CategoryTheory diff --git a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean index 7c5f0dc5ba024f..213220d6ea2963 100644 --- a/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean +++ b/Mathlib/Combinatorics/Additive/AP/Three/Defs.lean @@ -43,7 +43,7 @@ the size of the biggest 3AP-free subset of `{0, ..., n - 1}`. 3AP-free, Salem-Spencer, Roth, arithmetic progression, average, three-free -/ -open Finset Function Nat +open Finset Function open scoped Pointwise variable {F α β 𝕜 E : Type*} @@ -273,7 +273,7 @@ variable {s t} {n : ℕ} @[to_additive] theorem ThreeGPFree.le_mulRothNumber (hs : ThreeGPFree (s : Set α)) (h : s ⊆ t) : s.card ≤ mulRothNumber t := - le_findGreatest (card_le_card h) ⟨s, h, rfl, hs⟩ + Nat.le_findGreatest (card_le_card h) ⟨s, h, rfl, hs⟩ @[to_additive] theorem ThreeGPFree.mulRothNumber_eq (hs : ThreeGPFree (s : Set α)) : diff --git a/Mathlib/Combinatorics/Configuration.lean b/Mathlib/Combinatorics/Configuration.lean index aabfcf9052b584..a703f011ed4da5 100644 --- a/Mathlib/Combinatorics/Configuration.lean +++ b/Mathlib/Combinatorics/Configuration.lean @@ -128,7 +128,7 @@ theorem Nondegenerate.exists_injective_of_card_le [Nondegenerate P L] [Fintype P by_cases hs₁ : s.card = 1 -- If `s = {l}`, then pick a point `p ∉ l` · obtain ⟨l, rfl⟩ := Finset.card_eq_one.mp hs₁ - obtain ⟨p, hl⟩ := exists_point l + obtain ⟨p, hl⟩ := exists_point (P := P) l rw [Finset.card_singleton, Finset.singleton_biUnion, Nat.one_le_iff_ne_zero] exact Finset.card_ne_zero_of_mem (Set.mem_toFinset.mpr hl) suffices (s.biUnion t)ᶜ.card ≤ sᶜ.card by diff --git a/Mathlib/Combinatorics/Enumerative/DyckWord.lean b/Mathlib/Combinatorics/Enumerative/DyckWord.lean index ab393d095b11c5..220e3d4c4f1414 100644 --- a/Mathlib/Combinatorics/Enumerative/DyckWord.lean +++ b/Mathlib/Combinatorics/Enumerative/DyckWord.lean @@ -291,7 +291,7 @@ lemma count_take_firstReturn_add_one : lemma count_D_lt_count_U_of_lt_firstReturn {i : ℕ} (hi : i < p.firstReturn) : (p.toList.take (i + 1)).count D < (p.toList.take (i + 1)).count U := by have ne := not_of_lt_findIdx hi - rw [decide_eq_true_eq, ← ne_eq, getElem_range] at ne + rw [decide_eq_false_iff_not, ← ne_eq, getElem_range] at ne exact lt_of_le_of_ne (p.count_D_le_count_U (i + 1)) ne.symm @[simp] @@ -307,7 +307,7 @@ lemma firstReturn_add : (p + q).firstReturn = if p = 0 then q.firstReturn else p · intro j hj rw [take_append_eq_append_take, show j + 1 - p.toList.length = 0 by omega, take_zero, append_nil] - exact (count_D_lt_count_U_of_lt_firstReturn hj).ne' + simpa using (count_D_lt_count_U_of_lt_firstReturn hj).ne' · rw [length_range, u, length_append] exact Nat.lt_add_right _ (firstReturn_lt_length h) @@ -323,6 +323,7 @@ lemma firstReturn_nest : p.nest.firstReturn = p.toList.length + 1 := by beq_iff_eq, reduceCtorEq, ite_false, take_append_eq_append_take, show j - p.toList.length = 0 by omega, take_zero, append_nil] have := p.count_D_le_count_U j + simp only [add_zero, decide_eq_false_iff_not, ne_eq] omega · simp_rw [length_range, u, length_append, length_cons] exact Nat.lt_add_one _ diff --git a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean index 8b5ce0135fec7a..29713f6399feb0 100644 --- a/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean +++ b/Mathlib/Combinatorics/SetFamily/KruskalKatona.lean @@ -52,6 +52,10 @@ namespace Finset namespace Colex variable {α : Type*} [LinearOrder α] {𝒜 𝒜₁ 𝒜₂ : Finset (Finset α)} {s t : Finset α} {r : ℕ} +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- This is important for iterating Kruskal-Katona: the shadow of an initial segment is also an initial segment. -/ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : @@ -67,7 +71,7 @@ lemma shadow_initSeg [Fintype α] (hs : s.Nonempty) : · simpa [ha] using erase_le_erase_min' hts hst.ge (mem_insert_self _ _) -- Now show that if t ≤ s - min s, there is j such that t ∪ j ≤ s -- We choose j as the smallest thing not in t - simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, and_assoc] + simp_rw [le_iff_eq_or_lt, lt_iff_exists_filter_lt, mem_sdiff, filter_inj, _root_.and_assoc] simp only [toColex_inj, ofColex_toColex, ne_eq, and_imp] rintro cards' (rfl | ⟨k, hks, hkt, z⟩) -- If t = s - min s, then use j = min s so t ∪ j = s @@ -125,13 +129,17 @@ variable {α : Type*} [LinearOrder α] {s U V : Finset α} {n : ℕ} namespace UV +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Applying the compression makes the set smaller in colex. This is intuitive since a portion of the set is being "shifted down" as `max U < max V`. -/ lemma toColex_compress_lt_toColex {hU : U.Nonempty} {hV : V.Nonempty} (h : max' U hU < max' V hV) (hA : compress U V s ≠ s) : toColex (compress U V s) < toColex s := by rw [compress, ite_ne_right_iff] at hA rw [compress, if_pos hA.1, lt_iff_exists_filter_lt] - simp_rw [mem_sdiff (s := s), filter_inj, and_assoc] + simp_rw [mem_sdiff (s := s), filter_inj, _root_.and_assoc] refine ⟨_, hA.1.2 <| max'_mem _ hV, not_mem_sdiff_of_mem_right <| max'_mem _ _, fun a ha ↦ ?_⟩ have : a ∉ V := fun H ↦ ha.not_le (le_max' _ _ H) have : a ∉ U := fun H ↦ ha.not_lt ((le_max' _ _ H).trans_lt h) diff --git a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean index 0cb730278bd98f..30662b99dcef95 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean @@ -36,7 +36,7 @@ lemma IsHamiltonian.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective f) ( /-- A hamiltonian path visits every vertex. -/ @[simp] lemma IsHamiltonian.mem_support (hp : p.IsHamiltonian) (c : α) : c ∈ p.support := by - simp only [← List.count_pos_iff_mem, hp _, Nat.zero_lt_one] + simp only [← List.count_pos_iff, hp _, Nat.zero_lt_one] /-- Hamiltonian paths are paths. -/ lemma IsHamiltonian.isPath (hp : p.IsHamiltonian) : p.IsPath := @@ -45,7 +45,7 @@ lemma IsHamiltonian.isPath (hp : p.IsHamiltonian) : p.IsPath := /-- A path whose support contains every vertex is hamiltonian. -/ lemma IsPath.isHamiltonian_of_mem (hp : p.IsPath) (hp' : ∀ w, w ∈ p.support) : p.IsHamiltonian := fun _ ↦ - le_antisymm (List.nodup_iff_count_le_one.1 hp.support_nodup _) (List.count_pos_iff_mem.2 (hp' _)) + le_antisymm (List.nodup_iff_count_le_one.1 hp.support_nodup _) (List.count_pos_iff.2 (hp' _)) lemma IsPath.isHamiltonian_iff (hp : p.IsPath) : p.IsHamiltonian ↔ ∀ w, w ∈ p.support := ⟨(·.mem_support), hp.isHamiltonian_of_mem⟩ @@ -78,7 +78,7 @@ lemma IsHamiltonianCycle.map {H : SimpleGraph β} (f : G →g H) (hf : Bijective (hp : p.IsHamiltonianCycle) : (p.map f).IsHamiltonianCycle where toIsCycle := hp.isCycle.map hf.injective isHamiltonian_tail := by - simp only [IsHamiltonian, support_tail, support_map, ne_eq, List.map_eq_nil, support_ne_nil, + simp only [IsHamiltonian, support_tail, support_map, ne_eq, List.map_eq_nil_iff, support_ne_nil, not_false_eq_true, List.count_tail, List.head_map, beq_iff_eq, hf.surjective.forall, hf.injective, List.count_map_of_injective] intro x diff --git a/Mathlib/Combinatorics/SimpleGraph/Path.lean b/Mathlib/Combinatorics/SimpleGraph/Path.lean index 9a5f20e639064c..f40c73152000a6 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Path.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Path.lean @@ -590,7 +590,7 @@ end Path namespace Walk variable {G} {p} {u v : V} {H : SimpleGraph V} -variable (p : G.Walk u v) +variable {p : G.Walk u v} protected theorem IsPath.transfer (hp) (pp : p.IsPath) : (p.transfer H hp).IsPath := by diff --git a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean index d0c4c37b647fa7..e74fad0ddff83a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Regularity/Equitabilise.lean @@ -32,6 +32,10 @@ namespace Finpartition variable {α : Type*} [DecidableEq α] {s t : Finset α} {m n a b : ℕ} {P : Finpartition s} +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Given a partition `P` of `s`, as well as a proof that `a * m + b * (m + 1) = s.card`, we can find a new partition `Q` of `s` where each part has size `m` or `m + 1`, every part of `P` is the union of parts of `Q` plus at most `m` extra elements, there are `b` parts of size `m + 1` and @@ -45,8 +49,8 @@ theorem equitabilise_aux (hs : a * m + b * (m + 1) = s.card) : -- Get rid of the easy case `m = 0` obtain rfl | m_pos := m.eq_zero_or_pos · refine ⟨⊥, by simp, ?_, by simpa [Finset.filter_true_of_mem] using hs.symm⟩ - simp only [le_zero_iff, card_eq_zero, mem_biUnion, exists_prop, mem_filter, id, and_assoc, - sdiff_eq_empty_iff_subset, subset_iff] + simp only [le_zero_iff, card_eq_zero, mem_biUnion, exists_prop, mem_filter, id, + _root_.and_assoc, sdiff_eq_empty_iff_subset, subset_iff] exact fun x hx a ha => ⟨{a}, mem_map_of_mem _ (P.le hx ha), singleton_subset_iff.2 ha, mem_singleton_self _⟩ -- Prove the case `m > 0` by strong induction on `s` diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index b05bb8804a63b2..6c7459ec02c9a4 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -83,6 +83,10 @@ lemma LocallyLinear.map (f : α ↪ β) (hG : G.LocallyLinear) : (G.map f).Local · rw [← Equiv.coe_toEmbedding, ← map_symm] exact LocallyLinear.map _ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : G.EdgeDisjointTriangles ↔ ∀ ⦃e : Sym2 α⦄, ¬ e.IsDiag → {s ∈ G.cliqueSet 3 | e ∈ (s : Finset α).sym2}.Subsingleton := by @@ -92,7 +96,7 @@ lemma edgeDisjointTriangles_iff_mem_sym2_subsingleton : ext s simp only [mem_sym2_iff, Sym2.mem_iff, forall_eq_or_imp, forall_eq, Set.sep_and, Set.mem_inter_iff, Set.mem_sep_iff, mem_cliqueSet_iff, Set.mem_setOf_eq, - and_and_and_comm (b := _ ∈ _), and_self, is3Clique_iff] + and_and_and_comm (b := _ ∈ _), _root_.and_self, is3Clique_iff] constructor · rintro ⟨⟨c, d, e, hcd, hce, hde, rfl⟩, hab⟩ simp only [mem_insert, mem_singleton] at hab diff --git a/Mathlib/Combinatorics/SimpleGraph/Walk.lean b/Mathlib/Combinatorics/SimpleGraph/Walk.lean index ccdc8da2c18eac..e2c574948f9c02 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Walk.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Walk.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kyle Miller -/ import Mathlib.Combinatorics.SimpleGraph.Maps -import Mathlib.Data.List.Lemmas /-! @@ -509,7 +508,7 @@ theorem getLast_support {G : SimpleGraph V} {a b : V} (p : G.Walk a b) : theorem tail_support_append {u v w : V} (p : G.Walk u v) (p' : G.Walk v w) : (p.append p').support.tail = p.support.tail ++ p'.support.tail := by - rw [support_append, List.tail_append_of_ne_nil _ _ (support_ne_nil _)] + rw [support_append, List.tail_append_of_ne_nil (support_ne_nil _)] theorem support_eq_cons {u v : V} (p : G.Walk u v) : p.support = u :: p.support.tail := by cases p <;> simp diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 79b3e2892c8915..4e120e2d502495 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -436,7 +436,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] - exact ⟨not_false, fun n hn => ⟨hn, hpos _ (List.getElem_mem _ _ hn)⟩⟩ + exact ⟨not_false, fun n hn => ⟨hn, hpos _ (List.getElem_mem hn)⟩⟩ /-- The length of the `i`th row in `ofRowLens w hw` is the `i`th entry of `w` -/ theorem rowLen_ofRowLens {w : List ℕ} {hw : w.Sorted (· ≥ ·)} (i : Fin w.length) : diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 1b60a716c05f3a..7017d99c1eb686 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -75,20 +75,20 @@ theorem ack_succ_succ (m n : ℕ) : ack (m + 1) (n + 1) = ack m (ack (m + 1) n) @[simp] theorem ack_one (n : ℕ) : ack 1 n = n + 2 := by induction' n with n IH - · rfl + · simp · simp [IH] @[simp] theorem ack_two (n : ℕ) : ack 2 n = 2 * n + 3 := by induction' n with n IH - · rfl + · simp · simpa [mul_succ] -- Porting note: re-written to get rid of ack_three_aux @[simp] theorem ack_three (n : ℕ) : ack 3 n = 2 ^ (n + 3) - 3 := by induction' n with n IH - · rfl + · simp · rw [ack_succ_succ, IH, ack_two, Nat.succ_add, Nat.pow_succ 2 (n + 3), mul_comm _ 2, 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 diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 084c5246a26c61..a2de5176d359ba 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -223,7 +223,7 @@ theorem rice₂ (C : Set Code) (H : ∀ cf cg, eval cf = eval cg → (cf ∈ C (Partrec.nat_iff.1 <| eval_part.comp (const cg) Computable.id) ((hC _).1 fC), fun h => by { obtain rfl | rfl := h <;> simpa [ComputablePred, Set.mem_empty_iff_false] using - ⟨⟨inferInstance⟩, Computable.const _⟩ }⟩ + Computable.const _}⟩ /-- The Halting problem is recursively enumerable -/ theorem halting_problem_re (n) : RePred fun c => (eval c n).Dom := @@ -281,8 +281,6 @@ namespace Nat.Partrec' open Mathlib.Vector Partrec Computable -open Nat (Partrec') - open Nat.Partrec' theorem to_part {n f} (pf : @Partrec' n f) : _root_.Partrec f := by diff --git a/Mathlib/Computability/Language.lean b/Mathlib/Computability/Language.lean index cc112e2e092668..a8779cd1c92831 100644 --- a/Mathlib/Computability/Language.lean +++ b/Mathlib/Computability/Language.lean @@ -159,7 +159,7 @@ lemma mem_kstar_iff_exists_nonempty {x : List α} : x ∈ l∗ ↔ ∃ S : List (List α), x = S.join ∧ ∀ y ∈ S, y ∈ l ∧ y ≠ [] := by constructor · rintro ⟨S, rfl, h⟩ - refine ⟨S.filter fun l ↦ !List.isEmpty l, by simp, fun y hy ↦ ?_⟩ + refine ⟨S.filter fun l ↦ !List.isEmpty l, by simp [List.join_filter_not_isEmpty], fun y hy ↦ ?_⟩ -- Porting note: The previous code was: -- rw [mem_filter, empty_iff_eq_nil] at hy rw [mem_filter, Bool.not_eq_true', ← Bool.bool_iff_false, List.isEmpty_iff] at hy diff --git a/Mathlib/Computability/PartrecCode.lean b/Mathlib/Computability/PartrecCode.lean index f563f6babf2c4b..74b64e82859abd 100644 --- a/Mathlib/Computability/PartrecCode.lean +++ b/Mathlib/Computability/PartrecCode.lean @@ -170,7 +170,7 @@ private theorem encode_ofNatCode : ∀ n, encodeCode (ofNatCode n) = n instance instDenumerable : Denumerable Code := mk' ⟨encodeCode, ofNatCode, fun c => by - induction c <;> try {rfl} <;> simp [encodeCode, ofNatCode, Nat.div2_val, *], + induction c <;> simp [encodeCode, ofNatCode, Nat.div2_val, *], encode_ofNatCode⟩ theorem encodeCode_eq : encode = encodeCode := @@ -900,7 +900,7 @@ private theorem hG : Primrec G := by Primrec.fst private theorem evaln_map (k c n) : - ((((List.range k)[n]?).map (evaln k c)).bind fun b => b) = evaln k c n := by + ((List.range k)[n]?.bind fun a ↦ evaln k c a) = evaln k c n := by by_cases kn : n < k · simp [List.getElem?_range kn] · rw [List.getElem?_len_le] @@ -937,7 +937,7 @@ theorem evaln_prim : Primrec fun a : (ℕ × Code) × ℕ => evaln a.1.1 a.1.2 a (List.range n.unpair.1).map (evaln n.unpair.1 (ofNat Code n.unpair.2))) (k', c') n = evaln k' c' n := by intro k₁ c₁ n₁ hl - simp [lup, List.getElem?_range hl, evaln_map, Bind.bind] + simp [lup, List.getElem?_range hl, evaln_map, Bind.bind, Option.bind_map] cases' c with cf cg cf cg cf cg cf <;> simp [evaln, nk, Bind.bind, Functor.map, Seq.seq, pure] · cases' encode_lt_pair cf cg with lf lg @@ -969,7 +969,7 @@ theorem evaln_prim : Primrec fun a : (ℕ × Code) × ℕ => evaln a.1.1 a.1.2 a (Primrec.option_bind (Primrec.list_get?.comp (this.comp (_root_.Primrec.const ()) (Primrec.encode_iff.2 Primrec.fst)) Primrec.snd) Primrec.snd.to₂).of_eq - fun ⟨⟨k, c⟩, n⟩ => by simp [evaln_map] + fun ⟨⟨k, c⟩, n⟩ => by simp [evaln_map, Option.bind_map] end diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index ea36cdcbb60532..73b365b5200ee5 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -263,7 +263,7 @@ def ListBlank.nth {Γ} [Inhabited Γ] (l : ListBlank Γ) (n : ℕ) : Γ := by rw [List.getI_eq_default _ h] rcases le_or_lt _ n with h₂ | h₂ · rw [List.getI_eq_default _ h₂] - rw [List.getI_eq_get _ h₂, List.get_eq_getElem, List.getElem_append_right' h, + rw [List.getI_eq_get _ h₂, List.get_eq_getElem, List.getElem_append_right h, List.getElem_replicate] @[simp] diff --git a/Mathlib/Control/Applicative.lean b/Mathlib/Control/Applicative.lean index 87acbfcaf58336..3b6bb33ac89511 100644 --- a/Mathlib/Control/Applicative.lean +++ b/Mathlib/Control/Applicative.lean @@ -5,6 +5,7 @@ Authors: Simon Hudon -/ import Mathlib.Algebra.Group.Defs import Mathlib.Control.Functor +import Mathlib.Control.Basic /-! # `applicative` instances @@ -28,7 +29,7 @@ variable {α β γ σ : Type u} theorem Applicative.map_seq_map (f : α → β → γ) (g : σ → β) (x : F α) (y : F σ) : f <$> x <*> g <$> y = ((· ∘ g) ∘ f) <$> x <*> y := by - simp [flip, functor_norm] + simp [flip, functor_norm, Function.comp_def] theorem Applicative.pure_seq_eq_map' (f : α → β) : ((pure f : F (α → β)) <*> ·) = (f <$> ·) := by ext; simp [functor_norm] diff --git a/Mathlib/Control/Basic.lean b/Mathlib/Control/Basic.lean index 9f03e34b31c1e6..e6e6aa94abeeeb 100644 --- a/Mathlib/Control/Basic.lean +++ b/Mathlib/Control/Basic.lean @@ -18,12 +18,7 @@ variable {α β γ : Type u} section Functor -variable {f : Type u → Type v} [Functor f] [LawfulFunctor f] -@[functor_norm] -theorem Functor.map_map (m : α → β) (g : β → γ) (x : f α) : g <$> m <$> x = (g ∘ m) <$> x := - (comp_map _ _ _).symm --- order of implicits --- order of implicits +attribute [functor_norm] Functor.map_map end Functor @@ -67,10 +62,6 @@ section Monad variable {m : Type u → Type v} [Monad m] [LawfulMonad m] -theorem map_bind (x : m α) {g : α → m β} {f : β → γ} : - f <$> (x >>= g) = x >>= fun a => f <$> g a := by - rw [← bind_pure_comp, bind_assoc]; simp [bind_pure_comp] - 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 diff --git a/Mathlib/Control/Functor.lean b/Mathlib/Control/Functor.lean index 59ebe2553d456a..305a88d2cc83f7 100644 --- a/Mathlib/Control/Functor.lean +++ b/Mathlib/Control/Functor.lean @@ -3,7 +3,7 @@ Copyright (c) 2017 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ -import Mathlib.Control.Basic +import Mathlib.Tactic.Attr.Register import Mathlib.Data.Set.Defs import Mathlib.Tactic.TypeStar import Batteries.Tactic.Lint @@ -181,7 +181,7 @@ protected theorem id_map : ∀ x : Comp F G α, Comp.map id x = x protected theorem comp_map (g' : α → β) (h : β → γ) : ∀ x : Comp F G α, Comp.map (h ∘ g') x = Comp.map h (Comp.map g' x) - | Comp.mk x => by simp [Comp.map, Comp.mk, Functor.map_comp_map, functor_norm] + | Comp.mk x => by simp [Comp.map, Comp.mk, Functor.map_comp_map, functor_norm, Function.comp_def] -- Porting note: `Comp.mk` wasn't needed in mathlib3 instance lawfulFunctor : LawfulFunctor (Comp F G) where diff --git a/Mathlib/Control/Traversable/Basic.lean b/Mathlib/Control/Traversable/Basic.lean index e9103c5c55aa4f..52d25c007718ac 100644 --- a/Mathlib/Control/Traversable/Basic.lean +++ b/Mathlib/Control/Traversable/Basic.lean @@ -6,6 +6,7 @@ Authors: Simon Hudon import Mathlib.Data.Option.Defs import Mathlib.Control.Functor import Batteries.Data.List.Basic +import Mathlib.Control.Basic /-! # Traversable type class diff --git a/Mathlib/Control/Traversable/Equiv.lean b/Mathlib/Control/Traversable/Equiv.lean index fbd1a7fcc724e6..db32b26b14bb20 100644 --- a/Mathlib/Control/Traversable/Equiv.lean +++ b/Mathlib/Control/Traversable/Equiv.lean @@ -53,7 +53,7 @@ protected theorem id_map {α : Type u} (x : t' α) : Equiv.map eqv id x = x := b protected theorem comp_map {α β γ : Type u} (g : α → β) (h : β → γ) (x : t' α) : Equiv.map eqv (h ∘ g) x = Equiv.map eqv h (Equiv.map eqv g x) := by - simpa [Equiv.map] using comp_map .. + simp [Equiv.map, Function.comp_def] protected theorem lawfulFunctor : @LawfulFunctor _ (Equiv.functor eqv) := -- Porting note: why is `_inst` required here? diff --git a/Mathlib/Control/Traversable/Instances.lean b/Mathlib/Control/Traversable/Instances.lean index 9e4afad6029472..a91d6a6202f43d 100644 --- a/Mathlib/Control/Traversable/Instances.lean +++ b/Mathlib/Control/Traversable/Instances.lean @@ -32,7 +32,7 @@ theorem Option.id_traverse {α} (x : Option α) : Option.traverse (pure : α → theorem Option.comp_traverse {α β γ} (f : β → F γ) (g : α → G β) (x : Option α) : Option.traverse (Comp.mk ∘ (f <$> ·) ∘ g) x = Comp.mk (Option.traverse f <$> Option.traverse g x) := by - cases x <;> simp! [functor_norm] <;> rfl + cases x <;> (simp! [functor_norm] <;> rfl) theorem Option.traverse_eq_map_id {α β} (f : α → β) (x : Option α) : Option.traverse ((pure : _ → Id _) ∘ f) x = (pure : _ → Id _) (f <$> x) := by cases x <;> rfl @@ -148,7 +148,7 @@ variable [LawfulApplicative G] protected theorem comp_traverse {α β γ : Type u} (f : β → F γ) (g : α → G β) (x : σ ⊕ α) : Sum.traverse (Comp.mk ∘ (f <$> ·) ∘ g) x = Comp.mk.{u} (Sum.traverse f <$> Sum.traverse g x) := by - cases x <;> simp! [Sum.traverse, map_id, functor_norm] <;> rfl + cases x <;> (simp! [Sum.traverse, map_id, functor_norm] <;> rfl) protected theorem traverse_eq_map_id {α β} (f : α → β) (x : σ ⊕ α) : Sum.traverse ((pure : _ → Id _) ∘ f) x = (pure : _ → Id _) (f <$> x) := by diff --git a/Mathlib/Data/Array/ExtractLemmas.lean b/Mathlib/Data/Array/ExtractLemmas.lean index bc66fc0660fbca..b27b5245af41d2 100644 --- a/Mathlib/Data/Array/ExtractLemmas.lean +++ b/Mathlib/Data/Array/ExtractLemmas.lean @@ -27,7 +27,7 @@ theorem extract_append_left {a b : Array α} {i j : Nat} (h : j ≤ a.size) : · simp only [size_extract, size_append] omega · intro h1 h2 h3 - rw [get_extract, get_append_left, get_extract] + rw [getElem_extract, getElem_append_left, getElem_extract] theorem extract_append_right {a b : Array α} {i j : Nat} (h : a.size ≤ i) : (a ++ b).extract i j = b.extract (i - a.size) (j - a.size) := by @@ -35,8 +35,8 @@ theorem extract_append_right {a b : Array α} {i j : Nat} (h : a.size ≤ i) : · rw [size_extract, size_extract, size_append] omega · intro k hi h2 - rw [get_extract, get_extract, - get_append_right (show size a ≤ i + k by omega)] + rw [getElem_extract, getElem_extract, + getElem_append_right (show size a ≤ i + k by omega)] congr omega @@ -50,6 +50,6 @@ theorem extract_extract {s1 e2 e1 s2 : Nat} {a : Array α} (h : s1 + e2 ≤ e1) · simp only [size_extract] omega · intro i h1 h2 - simp only [get_extract, Nat.add_assoc] + simp only [getElem_extract, Nat.add_assoc] end Array diff --git a/Mathlib/Data/DFinsupp/Order.lean b/Mathlib/Data/DFinsupp/Order.lean index db1c8bc2bdd5f3..ebb4970f7a3ceb 100644 --- a/Mathlib/Data/DFinsupp/Order.lean +++ b/Mathlib/Data/DFinsupp/Order.lean @@ -44,7 +44,14 @@ lemma le_def : f ≤ g ↔ ∀ i, f i ≤ g i := Iff.rfl def orderEmbeddingToFun : (Π₀ i, α i) ↪o ∀ i, α i where toFun := DFunLike.coe inj' := DFunLike.coe_injective - map_rel_iff' := by rfl + map_rel_iff' := + #adaptation_note + /-- + This proof used to be `rfl`, + but has been temporarily broken by https://github.com/leanprover/lean4/pull/5329. + It can hopefully be restored after https://github.com/leanprover/lean4/pull/5359 + -/ + Iff.rfl @[simp, norm_cast] lemma coe_orderEmbeddingToFun : ⇑(orderEmbeddingToFun (α := α)) = DFunLike.coe := rfl diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index d4c6eeb4329411..d875809cf1a4a5 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -414,16 +414,6 @@ section Monoid protected theorem add_zero [NeZero n] (k : Fin n) : k + 0 = k := by simp only [add_def, val_zero', Nat.add_zero, mod_eq_of_lt (is_lt k)] --- Porting note (#10618): removing `simp`, `simp` can prove it with AddCommMonoid instance -protected theorem zero_add [NeZero n] (k : Fin n) : 0 + k = k := by - simp [Fin.ext_iff, add_def, mod_eq_of_lt (is_lt k)] - -instance {a : ℕ} [NeZero n] : OfNat (Fin n) a where - ofNat := Fin.ofNat' a n.pos_of_neZero - -instance inhabited (n : ℕ) [NeZero n] : Inhabited (Fin n) := - ⟨0⟩ - instance inhabitedFinOneAdd (n : ℕ) : Inhabited (Fin (1 + n)) := haveI : NeZero (1 + n) := by rw [Nat.add_comm]; infer_instance inferInstance @@ -434,8 +424,8 @@ theorem default_eq_zero (n : ℕ) [NeZero n] : (default : Fin n) = 0 := section from_ad_hoc -@[simp] lemma ofNat'_zero {h : 0 < n} [NeZero n] : (Fin.ofNat' 0 h : Fin n) = 0 := rfl -@[simp] lemma ofNat'_one {h : 0 < n} [NeZero n] : (Fin.ofNat' 1 h : Fin n) = 1 := rfl +@[simp] lemma ofNat'_zero [NeZero n] : (Fin.ofNat' n 0) = 0 := rfl +@[simp] lemma ofNat'_one [NeZero n] : (Fin.ofNat' n 1) = 1 := rfl end from_ad_hoc @@ -516,13 +506,6 @@ lemma natCast_strictMono (hbn : b ≤ n) (hab : a < b) : (a : Fin (n + 1)) < b : end OfNatCoe -@[simp] -theorem one_eq_zero_iff [NeZero n] : (1 : Fin n) = 0 ↔ n = 1 := by - obtain _ | _ | n := n <;> simp [Fin.ext_iff] - -@[simp] -theorem zero_eq_one_iff [NeZero n] : (0 : Fin n) = 1 ↔ n = 1 := by rw [eq_comm, one_eq_zero_iff] - end Add section Succ @@ -578,10 +561,6 @@ This one instead uses a `NeZero n` typeclass hypothesis. theorem le_zero_iff' {n : ℕ} [NeZero n] {k : Fin n} : k ≤ 0 ↔ k = 0 := ⟨fun h => Fin.ext <| by rw [Nat.eq_zero_of_le_zero h]; rfl, by rintro rfl; exact Nat.le_refl _⟩ --- Move to Batteries? -@[simp] theorem cast_refl {n : Nat} (h : n = n) : - Fin.cast h = id := rfl - -- TODO: Move to Batteries @[simp] lemma castLE_inj {hmn : m ≤ n} {a b : Fin m} : castLE hmn a = castLE hmn b ↔ a = b := by simp [Fin.ext_iff] @@ -775,7 +754,7 @@ theorem castSucc_ne_zero_of_lt {p i : Fin n} (h : p < i) : castSucc i ≠ 0 := b exact ((zero_le _).trans_lt h).ne' theorem succ_ne_last_iff (a : Fin (n + 1)) : succ a ≠ last (n + 1) ↔ a ≠ last n := - not_iff_not.mpr <| succ_eq_last_succ a + not_iff_not.mpr <| succ_eq_last_succ theorem succ_ne_last_of_lt {p i : Fin n} (h : i < p) : succ i ≠ last n := by cases n @@ -838,7 +817,7 @@ theorem le_pred_iff {j : Fin n} {i : Fin (n + 1)} (hi : i ≠ 0) : j ≤ pred i rw [← succ_le_succ_iff, succ_pred] theorem castSucc_pred_eq_pred_castSucc {a : Fin (n + 1)} (ha : a ≠ 0) - (ha' := a.castSucc_ne_zero_iff.mpr ha) : + (ha' := castSucc_ne_zero_iff.mpr ha) : (a.pred ha).castSucc = (castSucc a).pred ha' := rfl theorem castSucc_pred_add_one_eq {a : Fin (n + 1)} (ha : a ≠ 0) : @@ -1491,7 +1470,7 @@ theorem eq_zero (n : Fin 1) : n = 0 := Subsingleton.elim _ _ instance uniqueFinOne : Unique (Fin 1) where uniq _ := Subsingleton.elim _ _ -@[simp] +@[deprecated val_eq_zero (since := "2024-09-18")] theorem coe_fin_one (a : Fin 1) : (a : ℕ) = 0 := by simp [Subsingleton.elim a 0] lemma eq_one_of_neq_zero (i : Fin 2) (hi : i ≠ 0) : i = 1 := by diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index 4d2f060f732a2b..e4cb2370272a7d 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -788,7 +788,7 @@ theorem insertNth_apply_succAbove (i : Fin (n + 1)) (x : α i) (p : ∀ j, α (i generalize hk : castPred ((succAbove i) j) H₁ = k rw [castPred_succAbove _ _ hlt] at hk; cases hk intro; rfl - · generalize_proofs H₁ H₂; revert H₂ + · generalize_proofs H₀ H₁ H₂; revert H₂ generalize hk : pred (succAbove i j) H₁ = k erw [pred_succAbove _ _ (Fin.not_lt.1 hlt)] at hk; cases hk intro; rfl diff --git a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean index 7f72be2b670cbd..7f311baa1744cd 100644 --- a/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean +++ b/Mathlib/Data/Fin/Tuple/NatAntidiagonal.lean @@ -123,7 +123,7 @@ theorem antidiagonalTuple_one (n : ℕ) : antidiagonalTuple 1 n = [![n]] := by Nat.sub_self, List.bind_append, List.bind_singleton, List.bind_map] conv_rhs => rw [← List.nil_append [![n]]] congr 1 - simp_rw [List.bind_eq_nil, List.mem_range, List.map_eq_nil] + simp_rw [List.bind_eq_nil_iff, List.mem_range, List.map_eq_nil_iff] intro x hx obtain ⟨m, rfl⟩ := Nat.exists_eq_add_of_lt hx rw [add_assoc, add_tsub_cancel_left, antidiagonalTuple_zero_succ] diff --git a/Mathlib/Data/Finset/Basic.lean b/Mathlib/Data/Finset/Basic.lean index 86d298bb40842b..b84e46a5f7a05a 100644 --- a/Mathlib/Data/Finset/Basic.lean +++ b/Mathlib/Data/Finset/Basic.lean @@ -125,7 +125,7 @@ assert_not_exists CompleteLattice assert_not_exists OrderedCommMonoid -open Multiset Subtype Nat Function +open Multiset Subtype Function universe u @@ -2518,6 +2518,8 @@ end Filter section Range +open Nat + variable {n m l : ℕ} /-- `range n` is the set of natural numbers less than `n`. -/ diff --git a/Mathlib/Data/Finset/Functor.lean b/Mathlib/Data/Finset/Functor.lean index 4319e4a6b1b3ab..1359446c44d7e9 100644 --- a/Mathlib/Data/Finset/Functor.lean +++ b/Mathlib/Data/Finset/Functor.lean @@ -198,11 +198,16 @@ theorem map_comp_coe (h : α → β) : Functor.map h ∘ Multiset.toFinset = Multiset.toFinset ∘ Functor.map h := funext fun _ => image_toFinset +@[simp] +theorem map_comp_coe_apply (h : α → β) (s : Multiset α) : + s.toFinset.image h = (h <$> s).toFinset := + congrFun (map_comp_coe h) s + theorem map_traverse (g : α → G β) (h : β → γ) (s : Finset α) : Functor.map h <$> traverse g s = traverse (Functor.map h ∘ g) s := by unfold traverse - simp only [map_comp_coe, functor_norm] - rw [LawfulFunctor.comp_map, Multiset.map_traverse] + simp only [Functor.map_map, fmap_def, map_comp_coe_apply, Multiset.fmap_def, ← + Multiset.map_traverse] end Traversable diff --git a/Mathlib/Data/Finset/NatDivisors.lean b/Mathlib/Data/Finset/NatDivisors.lean index cbc7aa896c2f54..03edf844dbd267 100644 --- a/Mathlib/Data/Finset/NatDivisors.lean +++ b/Mathlib/Data/Finset/NatDivisors.lean @@ -16,12 +16,16 @@ exhibiting `Nat.divisors` as a multiplicative homomorphism from `ℕ` to `Finset open Nat Finset open scoped Pointwise +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- The divisors of a product of natural numbers are the pointwise product of the divisors of the factors. -/ lemma Nat.divisors_mul (m n : ℕ) : divisors (m * n) = divisors m * divisors n := by ext k simp_rw [mem_mul, mem_divisors, dvd_mul, mul_ne_zero_iff, ← exists_and_left, ← exists_and_right] - simp only [and_assoc, and_comm, and_left_comm] + simp only [_root_.and_assoc, _root_.and_comm, and_left_comm] /-- `Nat.divisors` as a `MonoidHom`. -/ @[simps] diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index fd8ebbddcb92b5..45c0a90562ef59 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -48,8 +48,6 @@ assert_not_exists MulAction open Function -open Nat - universe u v variable {α β γ : Type*} diff --git a/Mathlib/Data/Fintype/Prod.lean b/Mathlib/Data/Fintype/Prod.lean index 0d91984495e7f1..2a5dfcdbbbf59e 100644 --- a/Mathlib/Data/Fintype/Prod.lean +++ b/Mathlib/Data/Fintype/Prod.lean @@ -14,8 +14,6 @@ import Mathlib.Data.Finset.Prod open Function -open Nat - universe u v variable {α β γ : Type*} diff --git a/Mathlib/Data/Int/Defs.lean b/Mathlib/Data/Int/Defs.lean index b5619f8bcea083..4cb4e2f0e0ba69 100644 --- a/Mathlib/Data/Int/Defs.lean +++ b/Mathlib/Data/Int/Defs.lean @@ -43,7 +43,13 @@ protected lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := ⟨fun h ↦ ⟨Int.le_of_eq h, Int.ge_of_eq h⟩, fun h ↦ Int.le_antisymm h.1 h.2⟩ protected lemma le_iff_eq_or_lt : a ≤ b ↔ a = b ∨ a < b := by rw [Int.le_antisymm_iff, Int.lt_iff_le_not_le, ← and_or_left]; simp [em] -protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, or_comm] + +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ +protected lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := by rw [Int.le_iff_eq_or_lt, + _root_.or_comm] end Order diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index b4562695ae7184..87990322c40655 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2014 Parikshit Khanna. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Parikshit Khanna, Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Mario Carneiro -/ +import Mathlib.Control.Basic import Mathlib.Data.Nat.Defs import Mathlib.Data.Option.Basic import Mathlib.Data.List.Defs @@ -35,8 +36,6 @@ variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : Lis @[deprecated (since := "2024-07-27")] theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl -@[deprecated (since := "2024-06-07")] alias toArray_data := Array.data_toArray - -- Porting note: Delete this attribute -- attribute [inline] List.head! @@ -59,9 +58,6 @@ instance : Std.Associative (α := List α) Append.append where theorem singleton_injective : Injective fun a : α => [a] := fun _ _ h => (cons_eq_cons.1 h).1 -theorem singleton_inj {a b : α} : [a] = [b] ↔ a = b := - singleton_injective.eq_iff - theorem set_of_mem_cons (l : List α) (a : α) : { x | x ∈ a :: l } = insert a { x | x ∈ l } := Set.ext fun _ => mem_cons @@ -199,10 +195,6 @@ theorem map_subset_iff {l₁ l₂ : List α} (f : α → β) (h : Injective f) : theorem append_eq_has_append {L₁ L₂ : List α} : List.append L₁ L₂ = L₁ ++ L₂ := rfl -@[deprecated (since := "2024-03-24")] alias append_eq_cons_iff := append_eq_cons - -@[deprecated (since := "2024-03-24")] alias cons_eq_append_iff := cons_eq_append - @[deprecated (since := "2024-01-18")] alias append_left_cancel := append_cancel_left @[deprecated (since := "2024-01-18")] alias append_right_cancel := append_cancel_right @@ -229,10 +221,10 @@ theorem replicate_subset_singleton (n) (a : α) : replicate n a ⊆ [a] := fun _ mem_singleton.2 (eq_of_mem_replicate h) theorem subset_singleton_iff {a : α} {L : List α} : L ⊆ [a] ↔ ∃ n, L = replicate n a := by - simp only [eq_replicate, subset_def, mem_singleton, exists_eq_left'] + simp only [eq_replicate_iff, subset_def, mem_singleton, exists_eq_left'] theorem replicate_right_injective {n : ℕ} (hn : n ≠ 0) : Injective (@replicate α n) := - fun _ _ h => (eq_replicate.1 h).2 _ <| mem_replicate.2 ⟨hn, rfl⟩ + fun _ _ h => (eq_replicate_iff.1 h).2 _ <| mem_replicate.2 ⟨hn, rfl⟩ theorem replicate_right_inj {a b : α} {n : ℕ} (hn : n ≠ 0) : replicate n a = replicate n b ↔ a = b := @@ -360,21 +352,10 @@ lemma getLast_filter {p : α → Bool} : /-! ### getLast? -/ --- This is a duplicate of `getLast?_eq_none_iff`. --- We should remove one of them. -theorem getLast?_eq_none : ∀ {l : List α}, getLast? l = none ↔ l = [] - | [] => by simp - | [a] => by simp - | a :: b :: l => by simp [@getLast?_eq_none (b :: l)] +@[deprecated (since := "2024-09-06")] alias getLast?_eq_none := getLast?_eq_none_iff @[deprecated (since := "2024-06-20")] alias getLast?_isNone := getLast?_eq_none -@[simp] -theorem getLast?_isSome : ∀ {l : List α}, l.getLast?.isSome ↔ l ≠ [] - | [] => by simp - | [a] => by simp - | a :: b :: l => by simp [@getLast?_isSome (b :: l)] - theorem mem_getLast?_eq_getLast : ∀ {l : List α} {x : α}, x ∈ l.getLast? → ∃ h, x = getLast l h | [], x, hx => False.elim <| by simp at hx | [a], x, hx => @@ -395,10 +376,6 @@ theorem mem_getLast?_cons {x y : α} : ∀ {l : List α}, x ∈ l.getLast? → x | [], _ => by contradiction | _ :: _, h => h -theorem mem_of_mem_getLast? {l : List α} {a : α} (ha : a ∈ l.getLast?) : a ∈ l := - let ⟨_, h₂⟩ := mem_getLast?_eq_getLast ha - h₂.symm ▸ getLast_mem _ - theorem dropLast_append_getLast? : ∀ {l : List α}, ∀ a ∈ l.getLast?, dropLast l ++ [a] = l | [], a, ha => (Option.not_mem_none a ha).elim | [a], _, rfl => rfl @@ -462,9 +439,6 @@ theorem eq_cons_of_mem_head? {x : α} : ∀ {l : List α}, x ∈ l.head? → l = simp only [head?, Option.mem_def, Option.some_inj] at h exact h ▸ rfl -theorem mem_of_mem_head? {x : α} {l : List α} (h : x ∈ l.head?) : x ∈ l := - (eq_cons_of_mem_head? h).symm ▸ mem_cons_self _ _ - @[simp] theorem head!_cons [Inhabited α] (a : α) (l : List α) : head! (a :: l) = a := rfl @[simp] @@ -960,9 +934,6 @@ theorem infix_bind_of_mem {a : α} {as : List α} (h : a ∈ as) (f : α → Lis theorem map_eq_map {α β} (f : α → β) (l : List α) : f <$> l = map f l := rfl -@[simp] -theorem map_tail (f : α → β) (l) : map f (tail l) = tail (map f l) := by cases l <;> rfl - /-- A single `List.map` of a composition of functions is equal to composing a `List.map` with another `List.map`, fully applied. This is the reverse direction of `List.map_map`. @@ -1198,7 +1169,7 @@ lemma append_cons_inj_of_not_mem {x₁ x₂ z₁ z₂ : List α} {a₁ a₂ : α (notin_x : a₂ ∉ x₁) (notin_z : a₂ ∉ z₁) : x₁ ++ a₁ :: z₁ = x₂ ++ a₂ :: z₂ ↔ x₁ = x₂ ∧ a₁ = a₂ ∧ z₁ = z₂ := by constructor - · simp only [append_eq_append_iff, cons_eq_append, cons_eq_cons] + · simp only [append_eq_append_iff, cons_eq_append_iff, cons_eq_cons] rintro (⟨c, rfl, ⟨rfl, rfl, rfl⟩ | ⟨d, rfl, rfl⟩⟩ | ⟨c, rfl, ⟨rfl, rfl, rfl⟩ | ⟨d, rfl, rfl⟩⟩) <;> simp_all · rintro ⟨rfl, rfl, rfl⟩ @@ -1267,7 +1238,7 @@ theorem getElem_succ_scanl {i : ℕ} (h : i + 1 < (scanl f b l).length) : · simp only [length] at h exact absurd h (by omega) · simp_rw [scanl_cons] - rw [getElem_append_right'] + rw [getElem_append_right] · simp only [length, Nat.zero_add 1, succ_add_sub_one, hi]; rfl · simp only [length_singleton]; omega @@ -1360,25 +1331,12 @@ local notation a " ⋆ " b => op a b /-- Notation for `foldl op a l`. -/ local notation l " <*> " a => foldl op a l -theorem foldl_assoc : ∀ {l : List α} {a₁ a₂}, (l <*> a₁ ⋆ a₂) = a₁ ⋆ l <*> a₂ - | [], a₁, a₂ => rfl - | a :: l, a₁, a₂ => - calc - ((a :: l) <*> a₁ ⋆ a₂) = l <*> a₁ ⋆ a₂ ⋆ a := by simp only [foldl_cons, ha.assoc] - _ = a₁ ⋆ (a :: l) <*> a₂ := by rw [foldl_assoc, foldl_cons] - theorem foldl_op_eq_op_foldr_assoc : ∀ {l : List α} {a₁ a₂}, ((l <*> a₁) ⋆ a₂) = a₁ ⋆ l.foldr (· ⋆ ·) a₂ | [], a₁, a₂ => rfl | a :: l, a₁, a₂ => by simp only [foldl_cons, foldr_cons, foldl_assoc, ha.assoc]; rw [foldl_op_eq_op_foldr_assoc] -theorem foldr_assoc : ∀ {l : List α} {a₁ a₂}, l.foldr op (op a₁ a₂) = op (l.foldr op a₁) a₂ - | [], a₁, a₂ => rfl - | a :: l, a₁, a₂ => by - simp only [foldr_cons, ha.assoc] - rw [foldr_assoc] - variable [hc : Std.Commutative op] theorem foldl_assoc_comm_cons {l : List α} {a₁ a₂} : ((a₁ :: l) <*> a₂) = a₁ ⋆ l <*> a₂ := by @@ -1555,18 +1513,19 @@ theorem modifyLast.go_append_one (f : α → α) (a : α) (tl : List α) (r : Ar rw [modifyLast.go, modifyLast.go] case x_3 | x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) rw [modifyLast.go_append_one _ _ tl _, modifyLast.go_append_one _ _ tl (Array.push #[] hd)] - simp only [Array.toListAppend_eq, Array.push_data, Array.data_toArray, nil_append, append_assoc] + simp only [Array.toListAppend_eq, Array.push_toList, Array.toList_toArray, nil_append, + append_assoc] theorem modifyLast_append_one (f : α → α) (a : α) (l : List α) : modifyLast f (l ++ [a]) = l ++ [f a] := by cases l with | nil => - simp only [nil_append, modifyLast, modifyLast.go, Array.toListAppend_eq, Array.data_toArray] + simp only [nil_append, modifyLast, modifyLast.go, Array.toListAppend_eq, Array.toList_toArray] | cons _ tl => simp only [cons_append, modifyLast] rw [modifyLast.go] case x_3 => exact append_ne_nil_of_right_ne_nil tl (cons_ne_nil a []) - rw [modifyLast.go_append_one, Array.toListAppend_eq, Array.push_data, Array.data_toArray, + rw [modifyLast.go_append_one, Array.toListAppend_eq, Array.push_toList, Array.toList_toArray, nil_append, cons_append, nil_append, cons_inj_right] exact modifyLast_append_one _ _ tl @@ -1617,11 +1576,13 @@ variable (f : α → Option α) theorem lookmap.go_append (l : List α) (acc : Array α) : lookmap.go f l acc = acc.toListAppend (lookmap f l) := by cases l with - | nil => rfl + | nil => simp [go, lookmap] | cons hd tl => rw [lookmap, go, go] cases f hd with - | none => simp only [go_append tl _, Array.toListAppend_eq, append_assoc, Array.push_data]; rfl + | none => + simp only [go_append tl _, Array.toListAppend_eq, append_assoc, Array.push_toList] + rfl | some a => rfl @[simp] @@ -1631,13 +1592,13 @@ theorem lookmap_nil : [].lookmap f = [] := @[simp] theorem lookmap_cons_none {a : α} (l : List α) (h : f a = none) : (a :: l).lookmap f = a :: l.lookmap f := by - simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.data_toArray, nil_append] + simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.toList_toArray, nil_append] rw [lookmap.go_append, h]; rfl @[simp] theorem lookmap_cons_some {a b : α} (l : List α) (h : f a = some b) : (a :: l).lookmap f = b :: l := by - simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.data_toArray, nil_append] + simp only [lookmap, lookmap.go, Array.toListAppend_eq, Array.toList_toArray, nil_append] rw [h] theorem lookmap_some : ∀ l : List α, l.lookmap some = l @@ -1798,7 +1759,7 @@ lemma filter_attach (l : List α) (p : α → Bool) : ← filter_map, attach_map_subtype_val] lemma filter_comm (q) (l : List α) : filter p (filter q l) = filter q (filter p l) := by - simp [and_comm] + simp [Bool.and_comm] @[simp] theorem filter_true (l : List α) : @@ -1920,7 +1881,7 @@ theorem erase_getElem [DecidableEq ι] {l : List ι} {i : ℕ} (hi : i < l.lengt | succ i => have hi' : i < l.length := by simpa using hi if ha : a = l[i] then - simpa [ha] using .trans (perm_cons_erase (l.getElem_mem i _)) (.cons _ (IH hi')) + simpa [ha] using .trans (perm_cons_erase (getElem_mem _)) (.cons _ (IH hi')) else simpa [ha] using IH hi' @@ -2240,24 +2201,21 @@ end Forall /-! ### Miscellaneous lemmas -/ -@[simp] -theorem getElem_attach (L : List α) (i : Nat) (h : i < L.attach.length) : - L.attach[i].1 = L[i]'(length_attach L ▸ h) := - calc - L.attach[i].1 = (L.attach.map Subtype.val)[i]'(by simpa using h) := by - rw [getElem_map] - _ = L[i]'_ := by congr 2; simp - theorem get_attach (L : List α) (i) : - (L.attach.get i).1 = L.get ⟨i, length_attach L ▸ i.2⟩ := by simp + (L.attach.get i).1 = L.get ⟨i, length_attach (L := L) ▸ i.2⟩ := by simp +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp 1100] theorem mem_map_swap (x : α) (y : β) (xs : List (α × β)) : (y, x) ∈ map Prod.swap xs ↔ (x, y) ∈ xs := by induction' xs with x xs xs_ih · simp only [not_mem_nil, map_nil] · cases' x with a b - simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, and_comm] + simp only [mem_cons, Prod.mk.inj_iff, map, Prod.swap_prod_mk, Prod.exists, xs_ih, + _root_.and_comm] theorem dropSlice_eq (xs : List α) (n m : ℕ) : dropSlice n m xs = xs.take n ++ xs.drop (n + m) := by induction n generalizing xs diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index a1953d05295181..1213160f03d0ae 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -47,6 +47,10 @@ theorem Chain.iff_mem {a : α} {l : List α} : theorem chain_singleton {a b : α} : Chain R a [b] ↔ R a b := by simp only [chain_cons, Chain.nil, and_true] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem chain_split {a b : α} {l₁ l₂ : List α} : Chain R a (l₁ ++ b :: l₂) ↔ Chain R a (l₁ ++ [b]) ∧ Chain R b l₂ := by induction' l₁ with x l₁ IH generalizing a <;> @@ -226,14 +230,18 @@ theorem Chain'.cons' {x} : ∀ {l : List α}, Chain' R l → (∀ y ∈ l.head?, theorem chain'_cons' {x l} : Chain' R (x :: l) ↔ (∀ y ∈ head? l, R x y) ∧ Chain' R l := ⟨fun h => ⟨h.rel_head?, h.tail⟩, fun ⟨h₁, h₂⟩ => h₂.cons' h₁⟩ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefixes below. +-/ theorem chain'_append : ∀ {l₁ l₂ : List α}, Chain' R (l₁ ++ l₂) ↔ Chain' R l₁ ∧ Chain' R l₂ ∧ ∀ x ∈ l₁.getLast?, ∀ y ∈ l₂.head?, R x y | [], l => by simp - | [a], l => by simp [chain'_cons', and_comm] + | [a], l => by simp [chain'_cons', _root_.and_comm] | a :: b :: l₁, l₂ => by rw [cons_append, cons_append, chain'_cons, chain'_cons, ← cons_append, chain'_append, - and_assoc] + _root_.and_assoc] simp theorem Chain'.append (h₁ : Chain' R l₁) (h₂ : Chain' R l₂) @@ -272,12 +280,16 @@ theorem Chain'.imp_head {x y} (h : ∀ {z}, R x z → R y z) {l} (hl : Chain' R Chain' R (y :: l) := hl.tail.cons' fun _ hz => h <| hl.rel_head? hz +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem chain'_reverse : ∀ {l}, Chain' R (reverse l) ↔ Chain' (flip R) l | [] => Iff.rfl | [a] => by simp only [chain'_singleton, reverse_singleton] | a :: b :: l => by rw [chain'_cons, reverse_cons, reverse_cons, append_assoc, cons_append, nil_append, - chain'_split, ← reverse_cons, @chain'_reverse (b :: l), and_comm, chain'_pair, flip] + chain'_split, ← reverse_cons, @chain'_reverse (b :: l), _root_.and_comm, chain'_pair, flip] theorem chain'_iff_get {R} : ∀ {l : List α}, Chain' R l ↔ ∀ (i : ℕ) (h : i < length l - 1), @@ -295,6 +307,10 @@ theorem Chain'.append_overlap {l₁ l₂ l₃ : List α} (h₁ : Chain' R (l₁ h₁.append h₂.right_of_append <| by simpa only [getLast?_append_of_ne_nil _ hn] using (chain'_append.1 h₂).2.2 +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → (Chain' R L.join ↔ (∀ l ∈ L, Chain' R l) ∧ L.Chain' (fun l₁ l₂ => ∀ᵉ (x ∈ l₁.getLast?) (y ∈ l₂.head?), R x y)) @@ -304,7 +320,7 @@ lemma chain'_join : ∀ {L : List (List α)}, [] ∉ L → rw [mem_cons, not_or, ← Ne] at hL rw [join, chain'_append, chain'_join hL.2, forall_mem_cons, chain'_cons] rw [mem_cons, not_or, ← Ne] at hL - simp only [forall_mem_cons, and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] + simp only [forall_mem_cons, _root_.and_assoc, join, head?_append_of_ne_nil _ hL.2.1.symm] exact Iff.rfl.and (Iff.rfl.and <| Iff.rfl.and and_comm) /-- If `a` and `b` are related by the reflexive transitive closure of `r`, then there is an diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 5f038e2ae1e94c..9c8cbaa1432099 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -334,7 +334,7 @@ theorem prev_next (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : obtain ⟨⟨n, hn⟩, rfl⟩ := get_of_mem hx simp only [next_get, prev_get, h, Nat.mod_add_mod] cases' l with hd tl - · simp at hx + · simp at hn · have : (n + 1 + length tl) % (length tl + 1) = n := by rw [length_cons] at hn rw [add_assoc, add_comm 1, Nat.add_mod_right, Nat.mod_eq_of_lt hn] @@ -345,7 +345,7 @@ theorem next_prev (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : obtain ⟨⟨n, hn⟩, rfl⟩ := get_of_mem hx simp only [next_get, prev_get, h, Nat.mod_add_mod] cases' l with hd tl - · simp at hx + · simp at hn · have : (n + length tl + 1) % (length tl + 1) = n := by rw [length_cons] at hn rw [add_assoc, Nat.add_mod_right, Nat.mod_eq_of_lt hn] diff --git a/Mathlib/Data/List/Dedup.lean b/Mathlib/Data/List/Dedup.lean index bb2bf2901cd893..55576850f682a2 100644 --- a/Mathlib/Data/List/Dedup.lean +++ b/Mathlib/Data/List/Dedup.lean @@ -82,7 +82,7 @@ theorem dedup_eq_cons (l : List α) (a : α) (l' : List α) : l.dedup = a :: l' ↔ a ∈ l ∧ a ∉ l' ∧ l.dedup.tail = l' := by refine ⟨fun h => ?_, fun h => ?_⟩ · refine ⟨mem_dedup.1 (h.symm ▸ mem_cons_self _ _), fun ha => ?_, by rw [h, tail_cons]⟩ - have := count_pos_iff_mem.2 ha + have := count_pos_iff.2 ha have : count a l.dedup ≤ 1 := nodup_iff_count_le_one.1 (nodup_dedup l) a rw [h, count_cons_self] at this omega diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 80750fd2c8b30a..f0aad1fcd0f005 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -10,6 +10,7 @@ import Mathlib.Util.CompileInductive import Batteries.Tactic.Lint.Basic import Batteries.Data.List.Lemmas import Batteries.Data.RBMap.Basic +import Batteries.Logic /-! ## Definitions on lists diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index 79f5bf6c1cb0a7..20808976c60ab9 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -84,14 +84,18 @@ theorem forall₂_cons_right_iff {b l u} : match u, h with | _, ⟨_, _, h₁, h₂, rfl⟩ => Forall₂.cons h₁ h₂ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefixes below. +-/ theorem forall₂_and_left {p : α → Prop} : ∀ l u, Forall₂ (fun a b => p a ∧ R a b) l u ↔ (∀ a ∈ l, p a) ∧ Forall₂ R l u | [], u => by simp only [forall₂_nil_left_iff, forall_prop_of_false (not_mem_nil _), imp_true_iff, true_and] | a :: l, u => by - simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, and_assoc, + simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, _root_.and_assoc, @and_comm _ (p a), @and_left_comm _ (p a), exists_and_left] - simp only [and_comm, and_assoc, and_left_comm, ← exists_and_right] + simp only [_root_.and_comm, _root_.and_assoc, and_left_comm, ← exists_and_right] @[simp] theorem forall₂_map_left_iff {f : γ → α} : diff --git a/Mathlib/Data/List/GetD.lean b/Mathlib/Data/List/GetD.lean index bc144f95b4785d..e16ce04be7e338 100644 --- a/Mathlib/Data/List/GetD.lean +++ b/Mathlib/Data/List/GetD.lean @@ -72,16 +72,13 @@ alias getD_replicate_default_eq := getElem?_getD_replicate_default_eq theorem getD_append (l l' : List α) (d : α) (n : ℕ) (h : n < l.length) : (l ++ l').getD n d = l.getD n d := by rw [getD_eq_getElem _ _ (Nat.lt_of_lt_of_le h (length_append _ _ ▸ Nat.le_add_right _ _)), - getElem_append _ h, getD_eq_getElem] + getElem_append_left h, getD_eq_getElem] theorem getD_append_right (l l' : List α) (d : α) (n : ℕ) (h : l.length ≤ n) : (l ++ l').getD n d = l'.getD (n - l.length) d := by cases Nat.lt_or_ge n (l ++ l').length with | inl h' => - rw [getD_eq_getElem (l ++ l') d h', getElem_append_right, getD_eq_getElem] - · rw [length_append] at h' - exact Nat.sub_lt_left_of_lt_add h h' - · exact Nat.not_lt_of_le h + rw [getD_eq_getElem (l ++ l') d h', getElem_append_right h, getD_eq_getElem] | inr h' => rw [getD_eq_default _ _ h', getD_eq_default] rwa [Nat.le_sub_iff_add_le' h, ← length_append] diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index b2351ad435a246..7e53018e89d4b4 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -55,9 +55,9 @@ theorem mapIdxGo_append : ∀ (f : ℕ → α → β) (l₁ l₂ : List α) (arr cases l₂ · rfl · rw [List.length_append] at h; contradiction - rw [l₁_nil, l₂_nil]; simp only [mapIdx.go, Array.toList_eq, Array.toArray_data] + rw [l₁_nil, l₂_nil]; simp only [mapIdx.go, List.toArray_toList] · cases' l₁ with head tail <;> simp only [mapIdx.go] - · simp only [nil_append, Array.toList_eq, Array.toArray_data] + · simp only [nil_append, List.toArray_toList] · simp only [List.append_eq] rw [ih] · simp only [cons_append, length_cons, length_append, Nat.succ.injEq] at h @@ -67,7 +67,7 @@ theorem mapIdxGo_length : ∀ (f : ℕ → α → β) (l : List α) (arr : Array length (mapIdx.go f l arr) = length l + arr.size := by intro f l induction' l with head tail ih - · intro; simp only [mapIdx.go, Array.toList_eq, length_nil, Nat.zero_add] + · intro; simp only [mapIdx.go, length_nil, Nat.zero_add] · intro; simp only [mapIdx.go]; rw [ih]; simp only [Array.size_push, length_cons] simp only [Nat.add_succ, Fin.add_zero, Nat.add_comm] @@ -77,7 +77,7 @@ theorem mapIdx_append_one : ∀ (f : ℕ → α → β) (l : List α) (e : α), unfold mapIdx rw [mapIdxGo_append f l [e]] simp only [mapIdx.go, Array.size_toArray, mapIdxGo_length, length_nil, Nat.add_zero, - Array.toList_eq, Array.push_data, Array.data_toArray] + Array.push_toList, Array.toList_toArray] @[local simp] theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α → β), @@ -119,15 +119,16 @@ theorem getElem?_mapIdx_go (f : ℕ → α → β) : ∀ (l : List α) (arr : Ar (mapIdx.go f l arr)[i]? = if h : i < arr.size then some arr[i] else Option.map (f i) l[i - arr.size]? | [], arr, i => by - simp [mapIdx.go, getElem?_eq, Array.getElem_eq_data_getElem] + simp only [mapIdx.go, Array.toListImpl_eq, getElem?_eq, Array.length_toList, + Array.getElem_eq_getElem_toList, length_nil, Nat.not_lt_zero, ↓reduceDIte, Option.map_none'] | a :: l, arr, i => by rw [mapIdx.go, getElem?_mapIdx_go] simp only [Array.size_push] split <;> split · simp only [Option.some.injEq] - rw [Array.getElem_eq_data_getElem] - simp only [Array.push_data] - rw [getElem_append_left, Array.getElem_eq_data_getElem] + rw [Array.getElem_eq_getElem_toList] + simp only [Array.push_toList] + rw [getElem_append_left, Array.getElem_eq_getElem_toList] · have : i = arr.size := by omega simp_all · omega @@ -158,7 +159,7 @@ theorem mapIdx_append (K L : List α) (f : ℕ → α → β) : @[simp] theorem mapIdx_eq_nil {f : ℕ → α → β} {l : List α} : List.mapIdx f l = [] ↔ l = [] := by - rw [List.mapIdx_eq_enum_map, List.map_eq_nil, List.enum_eq_nil] + rw [List.mapIdx_eq_enum_map, List.map_eq_nil_iff, List.enum_eq_nil] theorem get_mapIdx (l : List α) (f : ℕ → α → β) (i : ℕ) (h : i < l.length) (h' : i < (l.mapIdx f).length := h.trans_le (l.length_mapIdx f).ge) : @@ -356,12 +357,12 @@ theorem mapIdxMGo_eq_mapIdxMAuxSpec congr conv => { lhs; intro x; rw [ih _ _ h]; } funext x - simp only [Array.toList_eq, Array.push_data, append_assoc, singleton_append, Array.size_push, + simp only [Array.push_toList, append_assoc, singleton_append, Array.size_push, map_eq_pure_bind] theorem mapIdxM_eq_mmap_enum [LawfulMonad m] {β} (f : ℕ → α → m β) (as : List α) : as.mapIdxM f = List.traverse (uncurry f) (enum as) := by - simp only [mapIdxM, mapIdxMGo_eq_mapIdxMAuxSpec, Array.toList_eq, Array.data_toArray, + simp only [mapIdxM, mapIdxMGo_eq_mapIdxMAuxSpec, Array.toList_toArray, nil_append, mapIdxMAuxSpec, Array.size_toArray, length_nil, id_map', enum] end MapIdxM diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index 974beaac0fde74..7e655efa8abee1 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -69,8 +69,7 @@ theorem tail_subset (l : List α) : tail l ⊆ l := theorem mem_of_mem_dropLast (h : a ∈ l.dropLast) : a ∈ l := dropLast_subset l h -theorem mem_of_mem_tail (h : a ∈ l.tail) : a ∈ l := - tail_subset l h +attribute [gcongr] Sublist.drop theorem concat_get_prefix {x y : List α} (h : x <+: y) (hl : x.length < y.length) : x ++ [y.get ⟨x.length, hl⟩] <+: y := by @@ -150,7 +149,17 @@ theorem inits_cons (a : α) (l : List α) : inits (a :: l) = [] :: l.inits.map f theorem tails_cons (a : α) (l : List α) : tails (a :: l) = (a :: l) :: l.tails := by simp -@[simp] +#adaptation_note +/-- +This can be removed after nightly-2024-09-07. +-/ +attribute [-simp] map_tail + +#adaptation_note +/-- +`nolint simpNF` should be removed after nightly-2024-09-07. +-/ +@[simp, nolint simpNF] theorem inits_append : ∀ s t : List α, inits (s ++ t) = s.inits ++ t.inits.tail.map fun l => s ++ l | [], [] => by simp | [], a :: t => by simp diff --git a/Mathlib/Data/List/InsertNth.lean b/Mathlib/Data/List/InsertNth.lean index 99646ea33a4c51..b6b147595d8d88 100644 --- a/Mathlib/Data/List/InsertNth.lean +++ b/Mathlib/Data/List/InsertNth.lean @@ -83,13 +83,17 @@ theorem insertNth_comm (a b : α) : simp only [insertNth_succ_cons, cons.injEq, true_and] exact insertNth_comm a b i j l (Nat.le_of_succ_le_succ h₀) (Nat.le_of_succ_le_succ h₁) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefixes below. +-/ theorem mem_insertNth {a b : α} : ∀ {n : ℕ} {l : List α} (_ : n ≤ l.length), a ∈ l.insertNth n b ↔ a = b ∨ a ∈ l | 0, as, _ => by simp | n + 1, [], h => (Nat.not_succ_le_zero _ h).elim | n + 1, a' :: as, h => by rw [List.insertNth_succ_cons, mem_cons, mem_insertNth (Nat.le_of_succ_le_succ h), - ← or_assoc, @or_comm (a = a'), or_assoc, mem_cons] + ← _root_.or_assoc, @or_comm (a = a'), _root_.or_assoc, mem_cons] theorem insertNth_of_length_lt (l : List α) (x : α) (n : ℕ) (h : l.length < n) : insertNth n x l = l := by diff --git a/Mathlib/Data/List/Intervals.lean b/Mathlib/Data/List/Intervals.lean index c1419e19f2901e..870f434107df2a 100644 --- a/Mathlib/Data/List/Intervals.lean +++ b/Mathlib/Data/List/Intervals.lean @@ -135,7 +135,7 @@ theorem filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : simp only [(lt_of_lt_of_le (mem.1 hk).2 hml), decide_True] theorem filter_lt_of_le_bot {n m l : ℕ} (hln : l ≤ n) : ((Ico n m).filter fun x => x < l) = [] := - filter_eq_nil.2 fun k hk => by + filter_eq_nil_iff.2 fun k hk => by simp only [decide_eq_true_eq, not_lt] apply le_trans hln exact (mem.1 hk).1 @@ -161,7 +161,7 @@ theorem filter_le_of_le_bot {n m l : ℕ} (hln : l ≤ n) : exact le_trans hln (mem.1 hk).1 theorem filter_le_of_top_le {n m l : ℕ} (hml : m ≤ l) : ((Ico n m).filter fun x => l ≤ x) = [] := - filter_eq_nil.2 fun k hk => by + filter_eq_nil_iff.2 fun k hk => by rw [decide_eq_true_eq] exact not_le_of_gt (lt_of_lt_of_le (mem.1 hk).2 hml) diff --git a/Mathlib/Data/List/Lattice.lean b/Mathlib/Data/List/Lattice.lean index 7c3ce41155f9fa..75dc0c1e112bfa 100644 --- a/Mathlib/Data/List/Lattice.lean +++ b/Mathlib/Data/List/Lattice.lean @@ -203,7 +203,7 @@ theorem count_bagInter {a : α} : by_cases ba : b = a · simp only [beq_iff_eq] rw [if_pos ba, Nat.sub_add_cancel] - rwa [succ_le_iff, count_pos_iff_mem, ← ba] + rwa [succ_le_iff, count_pos_iff, ← ba] · simp only [beq_iff_eq] rw [if_neg ba, Nat.sub_zero, Nat.add_zero, Nat.add_zero] · rw [cons_bagInter_of_neg _ hb, count_bagInter] diff --git a/Mathlib/Data/List/Lemmas.lean b/Mathlib/Data/List/Lemmas.lean index ea2d1faa23b359..b055f1b1872225 100644 --- a/Mathlib/Data/List/Lemmas.lean +++ b/Mathlib/Data/List/Lemmas.lean @@ -30,14 +30,6 @@ theorem tail_reverse_eq_reverse_dropLast (l : List α) : · rw [getElem?_eq_none, getElem?_eq_none] all_goals (simp; omega) -theorem getLast_tail (l : List α) (hl : l.tail ≠ []) : - l.tail.getLast hl = l.getLast (by intro h; rw [h] at hl; simp at hl) := by - simp only [← drop_one, ne_eq, drop_eq_nil_iff_le, - not_le, getLast_eq_getElem, length_drop] at hl |- - rw [← getElem_drop'] - · simp [show 1 + (l.length - 1 - 1) = l.length - 1 by omega] - omega - @[deprecated (since := "2024-08-19")] alias nthLe_tail := getElem_tail theorem injOn_insertNth_index_of_not_mem (l : List α) (x : α) (hx : x ∉ l) : diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index c20935506c3717..52a8f59fb71674 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -433,25 +433,31 @@ theorem minimum_of_length_pos_le_getElem {i : ℕ} (w : i < l.length) (h := (Nat l.minimum_of_length_pos h ≤ l[i] := getElem_le_maximum_of_length_pos (α := αᵒᵈ) w -lemma getD_maximum?_eq_unbot'_maximum (l : List α) (d : α) : - l.maximum?.getD d = l.maximum.unbot' d := by +lemma getD_max?_eq_unbot'_maximum (l : List α) (d : α) : + l.max?.getD d = l.maximum.unbot' d := by cases hy : l.maximum with | bot => simp [List.maximum_eq_bot.mp hy] | coe y => rw [List.maximum_eq_coe_iff] at hy simp only [WithBot.unbot'_coe] - cases hz : l.maximum? with - | none => simp [List.maximum?_eq_none_iff.mp hz] at hy + cases hz : l.max? with + | none => simp [List.max?_eq_none_iff.mp hz] at hy | some z => have : Antisymm (α := α) (· ≤ ·) := ⟨_root_.le_antisymm⟩ - rw [List.maximum?_eq_some_iff] at hz + rw [List.max?_eq_some_iff] at hz · rw [Option.getD_some] exact _root_.le_antisymm (hy.right _ hz.left) (hz.right _ hy.left) all_goals simp [le_total] -lemma getD_minimum?_eq_untop'_minimum (l : List α) (d : α) : - l.minimum?.getD d = l.minimum.untop' d := - getD_maximum?_eq_unbot'_maximum (α := αᵒᵈ) _ _ +@[deprecated (since := "2024-09-29")] +alias getD_maximum?_eq_unbot'_maximum := getD_max?_eq_unbot'_maximum + +lemma getD_min?_eq_untop'_minimum (l : List α) (d : α) : + l.min?.getD d = l.minimum.untop' d := + getD_max?_eq_unbot'_maximum (α := αᵒᵈ) _ _ + +@[deprecated (since := "2024-09-29")] +alias getD_minimum?_eq_untop'_minimum := getD_min?_eq_untop'_minimum end LinearOrder diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index c4ede7b589049e..9d86abe857e93b 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -16,7 +16,7 @@ predicate. universe u v -open Nat Function +open Function variable {α : Type u} {β : Type v} {l l₁ l₂ : List α} {r : α → α → Prop} {a b : α} @@ -102,13 +102,17 @@ theorem nodup_iff_get?_ne_get? {l : List α} : l.Nodup ↔ ∀ i j : ℕ, i < j → j < l.length → l.get? i ≠ l.get? j := by simp [nodup_iff_getElem?_ne_getElem?] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem Nodup.ne_singleton_iff {l : List α} (h : Nodup l) (x : α) : l ≠ [x] ↔ l = [] ∨ ∃ y ∈ l, y ≠ x := by induction' l with hd tl hl · simp · specialize hl h.of_cons by_cases hx : tl = [x] - · simpa [hx, and_comm, and_or_left] using h + · simpa [hx, _root_.and_comm, and_or_left] using h · rw [← Ne, hl] at hx rcases hx with (rfl | ⟨y, hy, hx⟩) · simp @@ -140,14 +144,14 @@ theorem nodup_iff_count_le_one [DecidableEq α] {l : List α} : Nodup l ↔ ∀ theorem nodup_iff_count_eq_one [DecidableEq α] : Nodup l ↔ ∀ a ∈ l, count a l = 1 := nodup_iff_count_le_one.trans <| forall_congr' fun _ => - ⟨fun H h => H.antisymm (count_pos_iff_mem.mpr h), + ⟨fun H h => H.antisymm (count_pos_iff.mpr h), fun H => if h : _ then (H h).le else (count_eq_zero.mpr h).trans_le (Nat.zero_le 1)⟩ @[simp] theorem count_eq_one_of_mem [DecidableEq α] {a : α} {l : List α} (d : Nodup l) (h : a ∈ l) : count a l = 1 := - _root_.le_antisymm (nodup_iff_count_le_one.1 d a) (Nat.succ_le_of_lt (count_pos_iff_mem.2 h)) + _root_.le_antisymm (nodup_iff_count_le_one.1 d a) (Nat.succ_le_of_lt (count_pos_iff.2 h)) theorem count_eq_of_nodup [DecidableEq α] {a : α} {l : List α} (d : Nodup l) : count a l = if a ∈ l then 1 else 0 := by @@ -174,9 +178,13 @@ theorem Nodup.append (d₁ : Nodup l₁) (d₂ : Nodup l₂) (dj : Disjoint l₁ theorem nodup_append_comm {l₁ l₂ : List α} : Nodup (l₁ ++ l₂) ↔ Nodup (l₂ ++ l₁) := by simp only [nodup_append, and_left_comm, disjoint_comm] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem nodup_middle {a : α} {l₁ l₂ : List α} : Nodup (l₁ ++ a :: l₂) ↔ Nodup (a :: (l₁ ++ l₂)) := by - simp only [nodup_append, not_or, and_left_comm, and_assoc, nodup_cons, mem_append, + simp only [nodup_append, not_or, and_left_comm, _root_.and_assoc, nodup_cons, mem_append, disjoint_cons_right] theorem Nodup.of_map (f : α → β) {l : List α} : Nodup (map f l) → Nodup l := @@ -244,8 +252,8 @@ theorem Nodup.erase_getElem [DecidableEq α] {l : List α} (hl : l.Nodup) · simp [IH hl.2] · rw [beq_iff_eq] simp only [getElem_cons_succ] - simp only [length_cons, succ_eq_add_one, Nat.add_lt_add_iff_right] at h - exact mt (· ▸ l.getElem_mem i h) hl.1 + simp only [length_cons, Nat.succ_eq_add_one, Nat.add_lt_add_iff_right] at h + exact mt (· ▸ getElem_mem h) hl.1 theorem Nodup.erase_get [DecidableEq α] {l : List α} (hl : l.Nodup) (i : Fin l.length) : l.erase (l.get i) = l.eraseIdx ↑i := by @@ -259,11 +267,15 @@ theorem nodup_join {L : List (List α)} : Nodup (join L) ↔ (∀ l ∈ L, Nodup l) ∧ Pairwise Disjoint L := by simp only [Nodup, pairwise_join, disjoint_left.symm, forall_mem_ne] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem nodup_bind {l₁ : List α} {f : α → List β} : Nodup (l₁.bind f) ↔ (∀ x ∈ l₁, Nodup (f x)) ∧ Pairwise (fun a b : α => Disjoint (f a) (f b)) l₁ := by - simp only [List.bind, nodup_join, pairwise_map, and_comm, and_left_comm, mem_map, exists_imp, - and_imp] + simp only [List.bind, nodup_join, pairwise_map, _root_.and_comm, and_left_comm, mem_map, + exists_imp, and_imp] rw [show (∀ (l : List β) (x : α), f x = l → x ∈ l₁ → Nodup l) ↔ ∀ x : α, x ∈ l₁ → Nodup (f x) from forall_swap.trans <| forall_congr' fun _ => forall_eq'] @@ -304,13 +316,12 @@ theorem Nodup.union [DecidableEq α] (l₁ : List α) (h : Nodup l₂) : (l₁ theorem Nodup.inter [DecidableEq α] (l₂ : List α) : Nodup l₁ → Nodup (l₁ ∩ l₂) := Nodup.filter _ -theorem Nodup.diff_eq_filter [DecidableEq α] : +theorem Nodup.diff_eq_filter [BEq α] [LawfulBEq α] : ∀ {l₁ l₂ : List α} (_ : l₁.Nodup), l₁.diff l₂ = l₁.filter (· ∉ l₂) | l₁, [], _ => by simp | l₁, a :: l₂, hl₁ => by rw [diff_cons, (hl₁.erase _).diff_eq_filter, hl₁.erase_eq_filter, filter_filter] - simp only [decide_not, Bool.not_eq_true', decide_eq_false_iff_not, bne_iff_ne, ne_eq, and_comm, - Bool.decide_and, mem_cons, not_or] + simp only [decide_not, bne, Bool.and_comm, mem_cons, not_or, decide_mem_cons, Bool.not_or] theorem Nodup.mem_diff_iff [DecidableEq α] (hl₁ : l₁.Nodup) : a ∈ l₁.diff l₂ ↔ a ∈ l₁ ∧ a ∉ l₂ := by rw [hl₁.diff_eq_filter, mem_filter, decide_eq_true_iff] diff --git a/Mathlib/Data/List/NodupEquivFin.lean b/Mathlib/Data/List/NodupEquivFin.lean index 5266ae7b60164a..0dcf11280caae8 100644 --- a/Mathlib/Data/List/NodupEquivFin.lean +++ b/Mathlib/Data/List/NodupEquivFin.lean @@ -126,7 +126,7 @@ theorem sublist_of_orderEmbedding_get?_eq {l l' : List α} (f : ℕ ↪o ℕ) exact ix.succ_pos rw [← List.take_append_drop (f 0 + 1) l', ← List.singleton_append] apply List.Sublist.append _ (IH _ this) - rw [List.singleton_sublist, ← h, l'.getElem_take _ (Nat.lt_succ_self _)] + rw [List.singleton_sublist, ← h, l'.getElem_take' _ (Nat.lt_succ_self _)] apply List.get_mem /-- A `l : List α` is `Sublist l l'` for `l' : List α` iff diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index fdff37f5163e58..f0e74dcc0d2b3c 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -51,25 +51,6 @@ theorem Perm.subset_congr_left {l₁ l₂ l₃ : List α} (h : l₁ ~ l₂) : l theorem Perm.subset_congr_right {l₁ l₂ l₃ : List α} (h : l₁ ~ l₂) : l₃ ⊆ l₁ ↔ l₃ ⊆ l₂ := ⟨fun h' => h'.trans h.subset, fun h' => h'.trans h.symm.subset⟩ -/-- Variant of `Perm.foldr_eq` with explicit commutativity argument. -/ -theorem Perm.foldr_eq' {f : α → β → β} {l₁ l₂ : List α} (p : l₁ ~ l₂) - (comm : ∀ x ∈ l₁, ∀ y ∈ l₁, ∀ z, f y (f x z) = f x (f y z)) - (init : β) : foldr f init l₁ = foldr f init l₂ := by - induction p using recOnSwap' generalizing init with - | nil => simp - | cons x _p IH => - simp only [foldr] - congr 1 - apply IH; intros; apply comm <;> exact .tail _ ‹_› - | swap' x y _p IH => - simp only [foldr] - rw [comm x (.tail _ <| .head _) y (.head _)] - congr 2 - apply IH; intros; apply comm <;> exact .tail _ (.tail _ ‹_›) - | trans p₁ _p₂ IH₁ IH₂ => - refine (IH₁ comm init).trans (IH₂ ?_ _) - intros; apply comm <;> apply p₁.symm.subset <;> assumption - section Rel open Relator @@ -265,13 +246,17 @@ theorem Perm.bagInter {l₁ l₂ t₁ t₂ : List α} (hl : l₁ ~ l₂) (ht : t l₁.bagInter t₁ ~ l₂.bagInter t₂ := ht.bagInter_left l₂ ▸ hl.bagInter_right _ +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem perm_replicate_append_replicate {l : List α} {a b : α} {m n : ℕ} (h : a ≠ b) : l ~ replicate m a ++ replicate n b ↔ count a l = m ∧ count b l = n ∧ l ⊆ [a, b] := by rw [perm_iff_count, ← Decidable.and_forall_ne a, ← Decidable.and_forall_ne b] suffices l ⊆ [a, b] ↔ ∀ c, c ≠ b → c ≠ a → c ∉ l by simp (config := { contextual := true }) [count_replicate, h, this, count_eq_zero, Ne.symm] trans ∀ c, c ∈ l → c = b ∨ c = a - · simp [subset_def, or_comm] + · simp [subset_def, _root_.or_comm] · exact forall_congr' fun _ => by rw [← and_imp, ← not_or, not_imp_not] theorem Perm.dedup {l₁ l₂ : List α} (p : l₁ ~ l₂) : dedup l₁ ~ dedup l₂ := diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index dfa46749569338..37fb13a041412e 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -166,6 +166,10 @@ theorem foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : · rfl · simp_rw [foldr_cons, ih, bind_cons, append_assoc, permutationsAux2_append] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α)} {l' : List α} : l' ∈ foldr (fun y r => (permutationsAux2 t ts r y id).2) r L ↔ l' ∈ r ∨ ∃ l₁ l₂, l₁ ++ l₂ ∈ L ∧ l₂ ≠ [] ∧ l' = l₁ ++ t :: l₂ ++ ts := by @@ -176,7 +180,7 @@ theorem mem_foldr_permutationsAux2 {t : α} {ts : List α} {r L : List (List α) ⟨fun ⟨_, aL, l₁, l₂, l0, e, h⟩ => ⟨l₁, l₂, l0, e ▸ aL, h⟩, fun ⟨l₁, l₂, l0, aL, h⟩ => ⟨_, aL, l₁, l₂, l0, rfl, h⟩⟩ rw [foldr_permutationsAux2] - simp only [mem_permutationsAux2', ← this, or_comm, and_left_comm, mem_append, mem_bind, + simp only [mem_permutationsAux2', ← this, _root_.or_comm, and_left_comm, mem_append, mem_bind, append_assoc, cons_append, exists_prop] theorem length_foldr_permutationsAux2 (t : α) (ts : List α) (r L : List (List α)) : diff --git a/Mathlib/Data/List/Rotate.lean b/Mathlib/Data/List/Rotate.lean index 5cbd6337a56306..923a8ba69b0eda 100644 --- a/Mathlib/Data/List/Rotate.lean +++ b/Mathlib/Data/List/Rotate.lean @@ -104,11 +104,15 @@ theorem rotate_cons_succ (l : List α) (a : α) (n : ℕ) : (a :: l : List α).rotate (n + 1) = (l ++ [a]).rotate n := by rw [rotate_eq_rotate', rotate_eq_rotate', rotate'_cons_succ] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp] theorem mem_rotate : ∀ {l : List α} {a : α} {n : ℕ}, a ∈ l.rotate n ↔ a ∈ l | [], _, n => by simp | a :: l, _, 0 => by simp - | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, or_comm] + | a :: l, _, n + 1 => by simp [rotate_cons_succ, mem_rotate, _root_.or_comm] @[simp] theorem length_rotate (l : List α) (n : ℕ) : (l.rotate n).length = l.length := by @@ -116,7 +120,7 @@ theorem length_rotate (l : List α) (n : ℕ) : (l.rotate n).length = l.length : @[simp] theorem rotate_replicate (a : α) (n : ℕ) (k : ℕ) : (replicate n a).rotate k = replicate n a := - eq_replicate.2 ⟨by rw [length_rotate, length_replicate], fun b hb => + eq_replicate_iff.2 ⟨by rw [length_rotate, length_replicate], fun b hb => eq_of_mem_replicate <| mem_rotate.1 hb⟩ theorem rotate_eq_drop_append_take {l : List α} {n : ℕ} : @@ -473,7 +477,7 @@ theorem IsRotated.dropLast_tail {α} | [] => by simp | [_] => by simp | a :: b :: L => by - simp at hL' |- + simp only [head_cons, ne_eq, reduceCtorEq, not_false_eq_true, getLast_cons] at hL' simp [hL', IsRotated.cons_getLast_dropLast] /-- List of all cyclic permutations of `l`. diff --git a/Mathlib/Data/List/Sort.lean b/Mathlib/Data/List/Sort.lean index 9742e4370e93c4..162ca0b8a4138f 100644 --- a/Mathlib/Data/List/Sort.lean +++ b/Mathlib/Data/List/Sort.lean @@ -123,8 +123,8 @@ theorem eq_of_perm_of_sorted [IsAntisymm α r] {l₁ l₂ : List α} (hp : l₁ congr have : ∀ x ∈ u₂, x = a := fun x m => antisymm ((pairwise_append.1 hs₂).2.2 _ m a (mem_cons_self _ _)) (h₁ _ (by simp [m])) - rw [(@eq_replicate _ a (length u₂ + 1) (a :: u₂)).2, - (@eq_replicate _ a (length u₂ + 1) (u₂ ++ [a])).2] <;> + rw [(@eq_replicate_iff _ a (length u₂ + 1) (a :: u₂)).2, + (@eq_replicate_iff _ a (length u₂ + 1) (u₂ ++ [a])).2] <;> constructor <;> simp [iff_true_intro this, or_comm] @@ -149,7 +149,7 @@ theorem Sorted.rel_of_mem_take_of_mem_drop {l : List α} (h : List.Sorted r l) { (hx : x ∈ List.take k l) (hy : y ∈ List.drop k l) : r x y := by obtain ⟨iy, hiy, rfl⟩ := getElem_of_mem hy obtain ⟨ix, hix, rfl⟩ := getElem_of_mem hx - rw [getElem_take', getElem_drop] + rw [getElem_take, getElem_drop] rw [length_take] at hix exact h.rel_get_of_lt (Nat.lt_add_right _ (Nat.lt_min.mp hix).left) @@ -534,12 +534,12 @@ def mergeSort' : List α → List α let ls := (split (a :: b :: l)) have := length_split_fst_le l have := length_split_snd_le l - exact merge (r · ·) (mergeSort' ls.1) (mergeSort' ls.2) + exact merge (mergeSort' ls.1) (mergeSort' ls.2) (r · ·) termination_by l => length l @[nolint unusedHavesSuffices] -- Porting note: false positive theorem mergeSort'_cons_cons {a b} {l l₁ l₂ : List α} (h : split (a :: b :: l) = (l₁, l₂)) : - mergeSort' r (a :: b :: l) = merge (r · ·) (mergeSort' r l₁) (mergeSort' r l₂) := by + mergeSort' r (a :: b :: l) = merge (mergeSort' r l₁) (mergeSort' r l₂) (r · ·) := by simp only [mergeSort', h] section Correctness @@ -568,13 +568,13 @@ section TotalAndTransitive variable {r} [IsTotal α r] [IsTrans α r] -theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge (r · ·) l l') +theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sorted r (merge l l' (r · ·) ) | [], [], _, _ => by simp | [], b :: l', _, h₂ => by simpa using h₂ | a :: l, [], h₁, _ => by simpa using h₁ | a :: l, b :: l', h₁, h₂ => by by_cases h : a ≼ b - · suffices ∀ b' ∈ List.merge (r · ·) l (b :: l'), r a b' by + · suffices ∀ b' ∈ List.merge l (b :: l') (r · ·) , r a b' by simpa [h, h₁.of_cons.merge h₂] intro b' bm rcases show b' = b ∨ b' ∈ l ∨ b' ∈ l' by @@ -584,7 +584,7 @@ theorem Sorted.merge : ∀ {l l' : List α}, Sorted r l → Sorted r l' → Sort assumption · exact rel_of_sorted_cons h₁ _ bl · exact _root_.trans h (rel_of_sorted_cons h₂ _ bl') - · suffices ∀ b' ∈ List.merge (r · ·) (a :: l) l', r b b' by + · suffices ∀ b' ∈ List.merge (a :: l) l' (r · ·) , r b b' by simpa [h, h₁.merge h₂.of_cons] intro b' bm have ba : b ≼ a := (total_of r _ _).resolve_left h @@ -625,24 +625,6 @@ theorem mergeSort'_nil : [].mergeSort' r = [] := by rw [List.mergeSort'] @[simp] theorem mergeSort'_singleton (a : α) : [a].mergeSort' r = [a] := by rw [List.mergeSort'] -theorem map_merge (f : α → β) (r : α → α → Bool) (s : β → β → Bool) (l l' : List α) - (hl : ∀ a ∈ l, ∀ b ∈ l', r a b = s (f a) (f b)) : - (l.merge r l').map f = (l.map f).merge s (l'.map f) := by - match l, l' with - | [], x' => simp - | x, [] => simp - | x :: xs, x' :: xs' => - simp_rw [List.forall_mem_cons, forall_and] at hl - simp_rw [List.map, List.cons_merge_cons] - rw [← hl.1.1] - split - · rw [List.map, map_merge _ r s, List.map] - simp_rw [List.forall_mem_cons, forall_and] - exact ⟨hl.2.1, hl.2.2⟩ - · rw [List.map, map_merge _ r s, List.map] - simp_rw [List.forall_mem_cons] - exact ⟨hl.1.2, hl.2.2⟩ - theorem map_mergeSort' (f : α → β) (l : List α) (hl : ∀ a ∈ l, ∀ b ∈ l, a ≼ b ↔ f a ≼ f b) : (l.mergeSort' r).map f = (l.map f).mergeSort' s := match l with @@ -659,7 +641,7 @@ theorem map_mergeSort' (f : α → β) (l : List α) (hl : ∀ a ∈ l, ∀ b have := length_split_snd_le l simp_rw [List.map] rw [List.mergeSort'_cons_cons _ e, List.mergeSort'_cons_cons _ fe, - map_merge _ (r · ·) (s · ·), map_mergeSort' _ l₁ hl.1.1, map_mergeSort' _ l₂ hl.2.2] + map_merge, map_mergeSort' _ l₁ hl.1.1, map_mergeSort' _ l₂ hl.2.2] simp_rw [mem_mergeSort', decide_eq_decide] exact hl.1.2 termination_by length l diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 8eac342a35b419..3d8300bd2d98a8 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -44,7 +44,7 @@ theorem sublists'Aux_eq_array_foldl (a : α) : ∀ (r₁ r₂ : List (List α)), sublists'Aux a r₁ r₂ = ((r₁.toArray).foldl (init := r₂.toArray) (fun r l => r.push (a :: l))).toList := by intro r₁ r₂ - rw [sublists'Aux, Array.foldl_eq_foldl_data] + rw [sublists'Aux, Array.foldl_eq_foldl_toList] have := List.foldl_hom Array.toList (fun r l => r.push (a :: l)) (fun r l => r ++ [a :: l]) r₁ r₂.toArray (by simp) simpa using this @@ -53,8 +53,7 @@ theorem sublists'_eq_sublists'Aux (l : List α) : sublists' l = l.foldr (fun a r => sublists'Aux a r r) [[]] := by simp only [sublists', sublists'Aux_eq_array_foldl] rw [← List.foldr_hom Array.toList] - · rfl - · intros _ _; congr <;> simp + · intros _ _; congr theorem sublists'Aux_eq_map (a : α) (r₁ : List (List α)) : ∀ (r₂ : List (List α)), sublists'Aux a r₁ r₂ = r₂ ++ map (cons a) r₁ := @@ -107,7 +106,7 @@ theorem sublistsAux_eq_array_foldl : (r.toArray.foldl (init := #[]) fun r l => (r.push l).push (a :: l)).toList := by funext a r - simp only [sublistsAux, Array.foldl_eq_foldl_data, Array.mkEmpty] + simp only [sublistsAux, Array.foldl_eq_foldl_toList, Array.mkEmpty] have := foldl_hom Array.toList (fun r l => (r.push l).push (a :: l)) (fun (r : List (List α)) l => r ++ [l, a :: l]) r #[] (by simp) @@ -126,10 +125,9 @@ theorem sublistsAux_eq_bind : ext α l : 2 trans l.foldr sublistsAux [[]] · rw [sublistsAux_eq_bind, sublists] - · simp only [sublistsFast, sublistsAux_eq_array_foldl, Array.foldr_eq_foldr_data] + · simp only [sublistsFast, sublistsAux_eq_array_foldl, Array.foldr_eq_foldr_toList] rw [← foldr_hom Array.toList] - · rfl - · intros _ _; congr <;> simp + · intros _ _; congr theorem sublists_append (l₁ l₂ : List α) : sublists (l₁ ++ l₂) = (sublists l₂) >>= (fun x => (sublists l₁).map (· ++ x)) := by diff --git a/Mathlib/Data/List/Sym.lean b/Mathlib/Data/List/Sym.lean index 80954c1ad9dc7c..1d6228b9258e35 100644 --- a/Mathlib/Data/List/Sym.lean +++ b/Mathlib/Data/List/Sym.lean @@ -237,7 +237,7 @@ theorem sym_one_eq : xs.sym 1 = xs.map (· ::ₛ .nil) := by theorem sym2_eq_sym_two : xs.sym2.map (Sym2.equivSym α) = xs.sym 2 := by induction xs with - | nil => simp only [List.sym, map_eq_nil, sym2_eq_nil_iff] + | nil => simp only [List.sym, map_eq_nil_iff, sym2_eq_nil_iff] | cons x xs ih => rw [List.sym, ← ih, sym_one_eq, map_map, List.sym2, map_append, map_map] rfl diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index c124885555f3b0..6991b7dc20d725 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1345,8 +1345,7 @@ theorem pmap_eq_map (p : α → Prop) (f : α → β) (s : Multiset α) : theorem pmap_congr {p q : α → Prop} {f : ∀ a, p a → β} {g : ∀ a, q a → β} (s : Multiset α) : ∀ {H₁ H₂}, (∀ a ∈ s, ∀ (h₁ h₂), f a h₁ = g a h₂) → pmap f s H₁ = pmap g s H₂ := - @(Quot.inductionOn s (fun l _H₁ _H₂ h => congr_arg _ <| List.pmap_congr l h)) - + @(Quot.inductionOn s (fun l _H₁ _H₂ h => congr_arg _ <| List.pmap_congr_left l h)) theorem map_pmap {p : α → Prop} (g : β → γ) (f : ∀ a, p a → β) (s) : ∀ H, map g (pmap f s H) = pmap (fun a h => g (f a h)) s H := @@ -1391,7 +1390,7 @@ theorem attach_cons (a : α) (m : Multiset α) : Quotient.inductionOn m fun l => congr_arg _ <| congr_arg (List.cons _) <| by - rw [List.map_pmap]; exact List.pmap_congr _ fun _ _ _ _ => Subtype.eq rfl + rw [List.map_pmap]; exact List.pmap_congr_left _ fun _ _ _ _ => Subtype.eq rfl section DecidablePiExists @@ -1828,17 +1827,25 @@ theorem filter_filter (q) [DecidablePred q] (s : Multiset α) : filter p (filter q s) = filter (fun a => p a ∧ q a) s := Quot.inductionOn s fun l => by simp +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma filter_comm (q) [DecidablePred q] (s : Multiset α) : - filter p (filter q s) = filter q (filter p s) := by simp [and_comm] + filter p (filter q s) = filter q (filter p s) := by simp [_root_.and_comm] theorem filter_add_filter (q) [DecidablePred q] (s : Multiset α) : filter p s + filter q s = filter (fun a => p a ∨ q a) s + filter (fun a => p a ∧ q a) s := Multiset.induction_on s rfl fun a s IH => by by_cases p a <;> by_cases q a <;> simp [*] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem filter_add_not (s : Multiset α) : filter p s + filter (fun a => ¬p a) s = s := by rw [filter_add_filter, filter_eq_self.2, filter_eq_nil.2] · simp only [add_zero] - · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, or_comm] + · simp [Decidable.em, -Bool.not_eq_true, -not_and, not_and_or, _root_.or_comm] · simp only [Bool.not_eq_true, decide_eq_true_eq, Bool.eq_false_or_eq_true, decide_True, implies_true, Decidable.em] @@ -2017,11 +2024,11 @@ theorem countP_eq_countP_filter_add (s) (p q : α → Prop) [DecidablePred p] [D @[simp] theorem countP_True {s : Multiset α} : countP (fun _ => True) s = card s := - Quot.inductionOn s fun _l => List.countP_true + Quot.inductionOn s fun _l => congrFun List.countP_true _ @[simp] theorem countP_False {s : Multiset α} : countP (fun _ => False) s = 0 := - Quot.inductionOn s fun _l => List.countP_false + Quot.inductionOn s fun _l => congrFun List.countP_false _ theorem countP_map (f : α → β) (s : Multiset α) (p : β → Prop) [DecidablePred p] : countP p (map f s) = card (s.filter fun a => p (f a)) := by @@ -2050,7 +2057,7 @@ lemma filter_attach (s : Multiset α) (p : α → Prop) [DecidablePred p] : variable {p} theorem countP_pos {s} : 0 < countP p s ↔ ∃ a ∈ s, p a := - Quot.inductionOn s fun _l => by simpa using List.countP_pos (p ·) + Quot.inductionOn s fun _l => by simp theorem countP_eq_zero {s} : countP p s = 0 ↔ ∀ a ∈ s, ¬p a := Quot.inductionOn s fun _l => by simp [List.countP_eq_zero] @@ -2678,9 +2685,13 @@ lemma add_eq_union_left_of_le [DecidableEq α] {s t u : Multiset α} (h : t ≤ · rintro ⟨h0, rfl⟩ exact h0 +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ lemma add_eq_union_right_of_le [DecidableEq α] {x y z : Multiset α} (h : z ≤ y) : x + y = x ∪ z ↔ y = z ∧ x.Disjoint y := by - simpa only [and_comm] using add_eq_union_left_of_le h + simpa only [_root_.and_comm] using add_eq_union_left_of_le h theorem disjoint_map_map {f : α → γ} {g : β → γ} {s : Multiset α} {t : Multiset β} : Disjoint (s.map f) (t.map g) ↔ ∀ a ∈ s, ∀ b ∈ t, f a ≠ g b := by diff --git a/Mathlib/Data/Multiset/Functor.lean b/Mathlib/Data/Multiset/Functor.lean index 9c8f9ecbfb459a..81575f5f445c49 100644 --- a/Mathlib/Data/Multiset/Functor.lean +++ b/Mathlib/Data/Multiset/Functor.lean @@ -99,7 +99,6 @@ theorem comp_traverse {G H : Type _ → Type _} [Applicative G] [Applicative H] intro simp only [traverse, quot_mk_to_coe, lift_coe, Coe.coe, Function.comp_apply, Functor.map_map, functor_norm] - simp only [Function.comp_def, lift_coe] theorem map_traverse {G : Type* → Type _} [Applicative G] [CommApplicative G] {α β γ : Type _} (g : α → G β) (h : β → γ) (x : Multiset α) : @@ -107,7 +106,8 @@ theorem map_traverse {G : Type* → Type _} [Applicative G] [CommApplicative G] refine Quotient.inductionOn x ?_ intro simp only [traverse, quot_mk_to_coe, lift_coe, Function.comp_apply, Functor.map_map, map_comp_coe] - rw [LawfulFunctor.comp_map, Traversable.map_traverse'] + rw [Traversable.map_traverse'] + simp only [fmap_def, Function.comp_apply, Functor.map_map, List.map_eq_map] rfl theorem traverse_map {G : Type* → Type _} [Applicative G] [CommApplicative G] {α β γ : Type _} diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 8e257d61ead923..79bd2eb583ef4c 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -35,9 +35,9 @@ elements of `s` in increasing order. -/ def bitIndices (n : ℕ) : List ℕ := @binaryRec (fun _ ↦ List ℕ) [] (fun b _ s ↦ b.casesOn (s.map (· + 1)) (0 :: s.map (· + 1))) n -@[simp] theorem bitIndices_zero : bitIndices 0 = [] := by rfl +@[simp] theorem bitIndices_zero : bitIndices 0 = [] := by simp [bitIndices] -@[simp] theorem bitIndices_one : bitIndices 1 = [0] := by rfl +@[simp] theorem bitIndices_one : bitIndices 1 = [0] := by simp [bitIndices] theorem bitIndices_bit_true (n : ℕ) : bitIndices (bit true n) = 0 :: ((bitIndices n).map (· + 1)) := diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index feb4a5b86056f9..753381e7318d10 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -25,7 +25,7 @@ and `Nat.digits`. -- Once we're in the `Nat` namespace, `xor` will inconveniently resolve to `Nat.xor`. /-- `bxor` denotes the `xor` function i.e. the exclusive-or function on type `Bool`. -/ -local notation "bxor" => _root_.xor +local notation "bxor" => xor namespace Nat universe u @@ -48,7 +48,7 @@ def bodd (n : ℕ) : Bool := (boddDiv2 n).1 @[simp] lemma bodd_zero : bodd 0 = false := rfl -lemma bodd_one : bodd 1 = true := rfl +@[simp] lemma bodd_one : bodd 1 = true := rfl lemma bodd_two : bodd 2 = false := rfl @@ -88,12 +88,12 @@ lemma mod_two_of_bodd (n : ℕ) : n % 2 = cond (bodd n) 1 0 := by @[simp] lemma div2_zero : div2 0 = 0 := rfl -lemma div2_one : div2 1 = 0 := rfl +@[simp] lemma div2_one : div2 1 = 0 := rfl lemma div2_two : div2 2 = 1 := rfl @[simp] -lemma div2_succ (n : ℕ) : div2 (succ n) = cond (bodd n) (succ (div2 n)) (div2 n) := by +lemma div2_succ (n : ℕ) : div2 (n + 1) = cond (bodd n) (succ (div2 n)) (div2 n) := by simp only [bodd, boddDiv2, div2] rcases boddDiv2 n with ⟨_|_, _⟩ <;> simp @@ -195,6 +195,12 @@ lemma binaryRec_zero {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit rw [binaryRec] rfl +@[simp] +lemma binaryRec_one {C : Nat → Sort u} (z : C 0) (f : ∀ b n, C n → C (bit b n)) : + binaryRec z f 1 = f true 0 z := by + rw [binaryRec] + simp + /-! bitwise ops -/ lemma bodd_bit (b n) : bodd (bit b n) = b := by @@ -391,6 +397,7 @@ theorem bit1_bits (n : ℕ) : (2 * n + 1).bits = true :: n.bits := @[simp] theorem one_bits : Nat.bits 1 = [true] := by convert bit1_bits 0 + simp -- TODO Find somewhere this can live. -- example : bits 3423 = [true, true, true, true, true, false, true, false, true, false, true, true] diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index ea2eb8e119049b..98bc466c54c8db 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Ring.Nat import Mathlib.Order.Basic import Mathlib.Tactic.AdaptationNote import Mathlib.Tactic.Common +import Mathlib.Algebra.NeZero /-! # Bitwise operations on natural numbers @@ -272,9 +273,6 @@ theorem lor_comm (n m : ℕ) : n ||| m = m ||| n := theorem land_comm (n m : ℕ) : n &&& m = m &&& n := bitwise_comm Bool.and_comm n m -protected lemma xor_comm (n m : ℕ) : n ^^^ m = m ^^^ n := - bitwise_comm (Bool.bne_eq_xor ▸ Bool.xor_comm) n m - lemma and_two_pow (n i : ℕ) : n &&& 2 ^ i = (n.testBit i).toNat * 2 ^ i := by refine eq_of_testBit_eq fun j => ?_ obtain rfl | hij := Decidable.eq_or_ne i j <;> cases' h : n.testBit i @@ -286,13 +284,6 @@ lemma and_two_pow (n i : ℕ) : n &&& 2 ^ i = (n.testBit i).toNat * 2 ^ i := by lemma two_pow_and (n i : ℕ) : 2 ^ i &&& n = 2 ^ i * (n.testBit i).toNat := by rw [mul_comm, land_comm, and_two_pow] -@[simp] -theorem zero_xor (n : ℕ) : 0 ^^^ n = n := by simp [HXor.hXor, Xor.xor, xor] - -@[simp] -theorem xor_zero (n : ℕ) : n ^^^ 0 = n := by simp [HXor.hXor, Xor.xor, xor] - - /-- Proving associativity of bitwise operations in general essentially boils down to a huge case distinction, so it is shorter to use this tactic instead of proving it in the general case. -/ macro "bitwise_assoc_tac" : tactic => set_option hygiene false in `(tactic| ( @@ -305,22 +296,16 @@ macro "bitwise_assoc_tac" : tactic => set_option hygiene false in `(tactic| ( -- This is necessary because these are simp lemmas in mathlib <;> simp [hn, Bool.or_assoc, Bool.and_assoc, Bool.bne_eq_xor])) -protected lemma xor_assoc (n m k : ℕ) : (n ^^^ m) ^^^ k = n ^^^ (m ^^^ k) := by bitwise_assoc_tac - theorem land_assoc (n m k : ℕ) : (n &&& m) &&& k = n &&& (m &&& k) := by bitwise_assoc_tac theorem lor_assoc (n m k : ℕ) : (n ||| m) ||| k = n ||| (m ||| k) := by bitwise_assoc_tac -@[simp] -theorem xor_self (n : ℕ) : n ^^^ n = 0 := - zero_of_testBit_eq_false fun i => by simp - -- These lemmas match `mul_inv_cancel_right` and `mul_inv_cancel_left`. theorem xor_cancel_right (n m : ℕ) : (m ^^^ n) ^^^ n = m := by - rw [Nat.xor_assoc, xor_self, xor_zero] + rw [Nat.xor_assoc, Nat.xor_self, xor_zero] theorem xor_cancel_left (n m : ℕ) : n ^^^ (n ^^^ m) = m := by - rw [← Nat.xor_assoc, xor_self, zero_xor] + rw [← Nat.xor_assoc, Nat.xor_self, zero_xor] theorem xor_right_injective {n : ℕ} : Function.Injective (HXor.hXor n : ℕ → ℕ) := fun m m' h => by rw [← xor_cancel_left n m, ← xor_cancel_left n m', h] @@ -339,7 +324,7 @@ theorem xor_left_inj {n m m' : ℕ} : m ^^^ n = m' ^^^ n ↔ m = m' := @[simp] theorem xor_eq_zero {n m : ℕ} : n ^^^ m = 0 ↔ n = m := by - rw [← xor_self n, xor_right_inj, eq_comm] + rw [← Nat.xor_self n, xor_right_inj, eq_comm] theorem xor_ne_zero {n m : ℕ} : n ^^^ m ≠ 0 ↔ n ≠ m := xor_eq_zero.not diff --git a/Mathlib/Data/Nat/Cast/NeZero.lean b/Mathlib/Data/Nat/Cast/NeZero.lean index d49f78b3b30952..c3f00eb8b88209 100644 --- a/Mathlib/Data/Nat/Cast/NeZero.lean +++ b/Mathlib/Data/Nat/Cast/NeZero.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Gabriel Ebner -/ import Mathlib.Data.Nat.Cast.Defs -import Mathlib.Algebra.NeZero /-! # Lemmas about nonzero elements of an `AddMonoidWithOne` diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 81287ae2f16f5a..4ee7eb93975ba4 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -137,8 +137,6 @@ lemma one_lt_iff_ne_zero_and_ne_one : ∀ {n : ℕ}, 1 < n ↔ n ≠ 0 ∧ n ≠ lemma le_one_iff_eq_zero_or_eq_one : ∀ {n : ℕ}, n ≤ 1 ↔ n = 0 ∨ n = 1 := by simp [le_succ_iff] -@[simp] lemma lt_one_iff : n < 1 ↔ n = 0 := Nat.lt_succ_iff.trans <| by rw [le_zero_eq] - lemma one_le_of_lt (h : a < b) : 1 ≤ b := Nat.lt_of_le_of_lt (Nat.zero_le _) h protected lemma min_left_comm (a b c : ℕ) : min a (min b c) = min b (min a c) := by @@ -167,8 +165,13 @@ lemma pred_eq_of_eq_succ (H : m = n.succ) : m.pred = n := by simp [H] @[simp] lemma pred_eq_succ_iff : n - 1 = m + 1 ↔ n = m + 2 := by cases n <;> constructor <;> rintro ⟨⟩ <;> rfl +#adaptation_note +/-- +After nightly-2024-09-06 we can remove both the `_root_` prefixes below. +-/ lemma forall_lt_succ : (∀ m < n + 1, p m) ↔ (∀ m < n, p m) ∧ p n := by - simp only [Nat.lt_succ_iff, Nat.le_iff_lt_or_eq, or_comm, forall_eq_or_imp, and_comm] + simp only [Nat.lt_succ_iff, Nat.le_iff_lt_or_eq, _root_.or_comm, forall_eq_or_imp, + _root_.and_comm] lemma exists_lt_succ : (∃ m < n + 1, p m) ↔ (∃ m < n, p m) ∨ p n := by rw [← not_iff_not] @@ -310,11 +313,11 @@ lemma two_mul_ne_two_mul_add_one : 2 * n ≠ 2 * m + 1 := -- TODO: Replace `Nat.mul_right_cancel_iff` with `Nat.mul_left_inj` protected lemma mul_left_inj (ha : a ≠ 0) : b * a = c * a ↔ b = c := - Nat.mul_right_cancel_iff (Nat.pos_iff_ne_zero.2 ha) _ _ + Nat.mul_right_cancel_iff (Nat.pos_iff_ne_zero.2 ha) -- TODO: Replace `Nat.mul_left_cancel_iff` with `Nat.mul_right_inj` protected lemma mul_right_inj (ha : a ≠ 0) : a * b = a * c ↔ b = c := - Nat.mul_left_cancel_iff (Nat.pos_iff_ne_zero.2 ha) _ _ + Nat.mul_left_cancel_iff (Nat.pos_iff_ne_zero.2 ha) protected lemma mul_ne_mul_left (ha : a ≠ 0) : b * a ≠ c * a ↔ b ≠ c := not_congr (Nat.mul_left_inj ha) @@ -834,7 +837,7 @@ This is an alias of `Nat.leRec`, specialized to `Prop`. -/ @[elab_as_elim] lemma le_induction {m : ℕ} {P : ∀ n, m ≤ n → Prop} (base : P m m.le_refl) (succ : ∀ n hmn, P n hmn → P (n + 1) (le_succ_of_le hmn)) : ∀ n hmn, P n hmn := - @Nat.leRec (motive := P) base succ + @Nat.leRec (motive := P) _ base succ /-- Induction principle deriving the next case from the two previous ones. -/ def twoStepInduction {P : ℕ → Sort*} (zero : P 0) (one : P 1) @@ -1018,9 +1021,6 @@ lemma div_ne_zero_iff_of_dvd (hba : b ∣ a) : a / b ≠ 0 ↔ a ≠ 0 ∧ b ≠ @[simp] lemma mul_mod_mod (a b c : ℕ) : (a * (b % c)) % c = a * b % c := by rw [mul_mod, mod_mod, ← mul_mod] -@[simp] lemma mod_mul_mod (a b c : ℕ) : (a % c * b) % c = a * b % c := by - rw [mul_mod, mod_mod, ← mul_mod] - lemma pow_mod (a b n : ℕ) : a ^ b % n = (a % n) ^ b % n := by induction b with | zero => rfl diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index f7955d07682763..23cd40d657cba6 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -570,7 +570,7 @@ theorem sub_one_mul_sum_log_div_pow_eq_sub_sum_digits {p : ℕ} (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 - · rfl + · simp rw [bits_append_bit _ _ fun hn => absurd hn h] cases b · rw [digits_def' one_lt_two] diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index a00644b704354a..1285453521a3bd 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -13,7 +13,7 @@ import Mathlib.Tactic.IntervalCases # Basic lemmas on prime factorizations -/ -open Nat Finset List Finsupp +open Finset List Finsupp namespace Nat variable {a b m n p : ℕ} @@ -455,6 +455,10 @@ theorem setOf_pow_dvd_eq_Icc_factorization {n p : ℕ} (pp : p.Prime) (hn : n ext simp [Nat.lt_succ_iff, one_le_iff_ne_zero, pp.pow_dvd_iff_le_factorization hn] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- The set of positive powers of prime `p` that divide `n` is exactly the set of positive natural numbers up to `n.factorization p`. -/ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : @@ -462,7 +466,7 @@ theorem Icc_factorization_eq_pow_dvd (n : ℕ) {p : ℕ} (pp : Prime p) : rcases eq_or_ne n 0 with (rfl | hn) · simp ext x - simp only [mem_Icc, Finset.mem_filter, mem_Ico, and_assoc, and_congr_right_iff, + simp only [mem_Icc, Finset.mem_filter, mem_Ico, _root_.and_assoc, and_congr_right_iff, pp.pow_dvd_iff_le_factorization hn, iff_and_self] exact fun _ H => lt_of_le_of_lt H (factorization_lt p hn) diff --git a/Mathlib/Data/Nat/Factorization/Defs.lean b/Mathlib/Data/Nat/Factorization/Defs.lean index 4ff0ac2cf90f5f..fb64f3dd9d2d66 100644 --- a/Mathlib/Data/Nat/Factorization/Defs.lean +++ b/Mathlib/Data/Nat/Factorization/Defs.lean @@ -86,7 +86,7 @@ alias factorization_eq_factors_multiset := factorization_eq_primeFactorsList_mul theorem Prime.factorization_pos_of_dvd {n p : ℕ} (hp : p.Prime) (hn : n ≠ 0) (h : p ∣ n) : 0 < n.factorization p := by - rwa [← primeFactorsList_count_eq, count_pos_iff_mem, mem_primeFactorsList_iff_dvd hn hp] + rwa [← primeFactorsList_count_eq, count_pos_iff, mem_primeFactorsList_iff_dvd hn hp] theorem multiplicity_eq_factorization {n p : ℕ} (pp : p.Prime) (hn : n ≠ 0) : multiplicity p n = n.factorization p := by diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index eb2ae3bc58f525..744b67ca07c103 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -121,11 +121,15 @@ lemma log_le_self (b x : ℕ) : log b x ≤ x := theorem lt_pow_succ_log_self {b : ℕ} (hb : 1 < b) (x : ℕ) : x < b ^ (log b x).succ := lt_pow_of_log_lt hb (lt_succ_self _) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem log_eq_iff {b m n : ℕ} (h : m ≠ 0 ∨ 1 < b ∧ n ≠ 0) : log b n = m ↔ b ^ m ≤ n ∧ n < b ^ (m + 1) := by rcases em (1 < b ∧ n ≠ 0) with (⟨hb, hn⟩ | hbn) - · rw [le_antisymm_iff, ← Nat.lt_succ_iff, ← pow_le_iff_le_log, ← lt_pow_iff_log_lt, and_comm] <;> - assumption + · rw [le_antisymm_iff, ← Nat.lt_succ_iff, ← pow_le_iff_le_log, ← lt_pow_iff_log_lt, + _root_.and_comm] <;> assumption have hm : m ≠ 0 := h.resolve_right hbn rw [not_and_or, not_lt, Ne, not_not] at hbn rcases hbn with (hb | rfl) diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index 125eeae18ef2be..418d1347bbc662 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -30,9 +30,13 @@ theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := b · simp [WithBot.add_bot] simp [← WithBot.coe_add, add_eq_zero_iff_of_nonneg] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by cases n - · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, or_self] + · simp only [WithBot.bot_add, WithBot.bot_ne_one, WithBot.bot_ne_zero, false_and, _root_.or_self] cases m · simp [WithBot.add_bot] simp [← WithBot.coe_add, Nat.add_eq_one_iff] diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index ac946c8f455bc8..66c8cae9f974a9 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -867,7 +867,7 @@ theorem castNum_testBit (m n) : testBit m n = Nat.testBit m n := by · rfl · rw [PosNum.cast_bit1, ← two_mul, ← congr_fun Nat.bit_true, Nat.testBit_bit_zero] · rw [PosNum.cast_bit0, ← two_mul, ← congr_fun Nat.bit_false, Nat.testBit_bit_zero] - · simp + · simp [Nat.testBit_add_one] · rw [PosNum.cast_bit1, ← two_mul, ← congr_fun Nat.bit_true, Nat.testBit_bit_succ, IH] · rw [PosNum.cast_bit0, ← two_mul, ← congr_fun Nat.bit_false, Nat.testBit_bit_succ, IH] diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index 138ba4471da951..b0bab527487d94 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -137,29 +137,12 @@ variable {p : α → Prop} (f : ∀ a : α, p a → β) (x : Option α) theorem pbind_eq_bind (f : α → Option β) (x : Option α) : (x.pbind fun a _ ↦ f a) = x.bind f := by cases x <;> simp only [pbind, none_bind', some_bind'] -theorem map_bind {α β γ} (f : β → γ) (x : Option α) (g : α → Option β) : - Option.map f (x >>= g) = x >>= fun a ↦ Option.map f (g a) := by - simp only [← map_eq_map, ← bind_pure_comp, LawfulMonad.bind_assoc] - theorem map_bind' (f : β → γ) (x : Option α) (g : α → Option β) : Option.map f (x.bind g) = x.bind fun a ↦ Option.map f (g a) := by cases x <;> simp -theorem map_pbind (f : β → γ) (x : Option α) (g : ∀ a, a ∈ x → Option β) : - Option.map f (x.pbind g) = x.pbind fun a H ↦ Option.map f (g a H) := by - cases x <;> simp only [pbind, map_none'] - theorem pbind_map (f : α → β) (x : Option α) (g : ∀ b : β, b ∈ x.map f → Option γ) : pbind (Option.map f x) g = x.pbind fun a h ↦ g (f a) (mem_map_of_mem _ h) := by cases x <;> rfl -@[simp] -theorem pmap_none (f : ∀ a : α, p a → β) {H} : pmap f (@none α) H = none := - rfl - -@[simp] -theorem pmap_some (f : ∀ a : α, p a → β) {x : α} (h : p x) : - pmap f (some x) = fun _ ↦ some (f x h) := - rfl - theorem mem_pmem {a : α} (h : ∀ a ∈ x, p a) (ha : a ∈ x) : f a (h a ha) ∈ pmap f x h := by rw [mem_def] at ha ⊢ subst ha @@ -208,24 +191,6 @@ theorem pbind_eq_some {f : ∀ a : α, a ∈ x → Option β} {y : β} : simp only [mem_def, Option.some_inj] at H simpa [H] using hz --- Porting note: Can't simp tag this anymore because `pmap` simplifies --- @[simp] -theorem pmap_eq_none_iff {h} : pmap f x h = none ↔ x = none := by cases x <;> simp - --- Porting note: Can't simp tag this anymore because `pmap` simplifies --- @[simp] -theorem pmap_eq_some_iff {hf} {y : β} : - pmap f x hf = some y ↔ ∃ (a : α) (H : x = some a), f a (hf a H) = y := by - rcases x with (_|x) - · simp only [not_mem_none, exists_false, pmap, not_false_iff, exists_prop_of_false, reduceCtorEq] - · constructor - · intro h - simp only [pmap, Option.some_inj] at h - exact ⟨x, rfl, h⟩ - · rintro ⟨a, H, rfl⟩ - simp only [mem_def, Option.some_inj] at H - simp only [H, pmap] - -- Porting note: Can't simp tag this anymore because `join` and `pmap` simplify -- @[simp] theorem join_pmap_eq_pmap_join {f : ∀ a, p a → β} {x : Option (Option α)} (H) : diff --git a/Mathlib/Data/Option/Defs.lean b/Mathlib/Data/Option/Defs.lean index e140c43e996c97..c15bedf8531353 100644 --- a/Mathlib/Data/Option/Defs.lean +++ b/Mathlib/Data/Option/Defs.lean @@ -75,10 +75,6 @@ abbrev iget [Inhabited α] : Option α → α theorem iget_some [Inhabited α] {a : α} : (some a).iget = a := rfl -@[simp] -theorem mem_toList {a : α} {o : Option α} : a ∈ toList o ↔ a ∈ o := by - cases o <;> simp [toList, eq_comm] - instance liftOrGet_isCommutative (f : α → α → α) [Std.Commutative f] : Std.Commutative (liftOrGet f) := ⟨fun a b ↦ by cases a <;> cases b <;> simp [liftOrGet, Std.Commutative.comm]⟩ diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index 48a139e2ee3647..06b25c58fc9533 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -378,7 +378,7 @@ theorem Sized.rotateR_size {l x r} (hl : Sized l) : rw [← size_dual, dual_rotateR, hl.dual.rotateL_size, size_dual, size_dual, add_comm (size l)] theorem Sized.balance' {l x r} (hl : @Sized α l) (hr : Sized r) : Sized (balance' l x r) := by - unfold balance'; split_ifs + unfold Ordnode.balance'; split_ifs · exact hl.node' hr · exact hl.rotateL hr · exact hl.rotateR hr @@ -1258,7 +1258,7 @@ theorem Valid'.glue_aux {l r o₁ o₂} (hl : Valid' o₁ l o₂) (hr : Valid' o suffices H : _ by refine ⟨Valid'.balanceL (hl.of_lt ?_ ?_) v H, ?_⟩ · refine @findMin'_all (P := fun a : α => Bounded nil o₁ (a : WithBot α)) - rl rx (sep.2.1.1.imp ?_) hr.1.1.to_nil + _ rl rx (sep.2.1.1.imp ?_) hr.1.1.to_nil exact fun y h => hl.1.1.to_nil.mono_right (le_of_lt h) · exact @findMin'_all _ (fun a => All (· < a) (.node ls ll lx lr)) rl rx diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 73fe251a4dda5e..699b8b76a816fc 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -470,6 +470,10 @@ theorem corec_def {X} (f : X → F X) (x₀ : X) : M.corec f x₀ = M.mk (F.map dsimp only [PFunctor.map] congr +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem ext_aux [Inhabited (M F)] [DecidableEq F.A] {n : ℕ} (x y z : M F) (hx : Agree' n z x) (hy : Agree' n z y) (hrec : ∀ ps : Path F, n = ps.length → iselect ps x = iselect ps y) : x.approx (n + 1) = y.approx (n + 1) := by diff --git a/Mathlib/Data/PNat/Defs.lean b/Mathlib/Data/PNat/Defs.lean index 3004d33fc9df71..6eb644184ac56f 100644 --- a/Mathlib/Data/PNat/Defs.lean +++ b/Mathlib/Data/PNat/Defs.lean @@ -3,7 +3,6 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Neil Strickland -/ -import Mathlib.Algebra.NeZero import Mathlib.Data.Nat.Defs import Mathlib.Order.Basic import Mathlib.Order.TypeTags diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index f30532075537d6..125f987461805d 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -186,12 +186,12 @@ theorem snd_eq_iff : ∀ {p : α × β} {x : β}, p.2 = x ↔ p = (p.1, x) variable {r : α → α → Prop} {s : β → β → Prop} {x y : α × β} -lemma lex_iff : Prod.Lex r s x y ↔ r x.1 y.1 ∨ x.1 = y.1 ∧ s x.2 y.2 := lex_def _ _ +lemma lex_iff : Prod.Lex r s x y ↔ r x.1 y.1 ∨ x.1 = y.1 ∧ s x.2 y.2 := lex_def instance Lex.decidable [DecidableEq α] (r : α → α → Prop) (s : β → β → Prop) [DecidableRel r] [DecidableRel s] : DecidableRel (Prod.Lex r s) := - fun _ _ ↦ decidable_of_decidable_of_iff (lex_def r s).symm + fun _ _ ↦ decidable_of_decidable_of_iff lex_def.symm @[refl] theorem Lex.refl_left (r : α → α → Prop) (s : β → β → Prop) [IsRefl α r] : ∀ x, Prod.Lex r s x x diff --git a/Mathlib/Data/Prod/Lex.lean b/Mathlib/Data/Prod/Lex.lean index 1a8b1f220c0f88..f8116316921eeb 100644 --- a/Mathlib/Data/Prod/Lex.lean +++ b/Mathlib/Data/Prod/Lex.lean @@ -48,11 +48,11 @@ instance instLT (α β : Type*) [LT α] [LT β] : LT (α ×ₗ β) where lt := P theorem le_iff [LT α] [LE β] (a b : α × β) : toLex a ≤ toLex b ↔ a.1 < b.1 ∨ a.1 = b.1 ∧ a.2 ≤ b.2 := - Prod.lex_def (· < ·) (· ≤ ·) + Prod.lex_def theorem lt_iff [LT α] [LT β] (a b : α × β) : toLex a < toLex b ↔ a.1 < b.1 ∨ a.1 = b.1 ∧ a.2 < b.2 := - Prod.lex_def (· < ·) (· < ·) + Prod.lex_def example (x : α) (y : β) : toLex (x, y) = toLex (x, y) := rfl diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index f684d39ce96dec..0aa562b90bce1b 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -53,7 +53,7 @@ theorem num_mk (n d : ℤ) : (n /. d).num = d.sign * n / n.gcd d := by have (m : ℕ) : Int.natAbs (m + 1) = m + 1 := by rw [← Nat.cast_one, ← Nat.cast_add, Int.natAbs_cast] rcases d with ((_ | _) | _) <;> - rw [← Int.div_eq_ediv_of_dvd] <;> + rw [← Int.tdiv_eq_ediv_of_dvd] <;> simp [divInt, mkRat, Rat.normalize, Nat.succPNat, Int.sign, Int.gcd, Int.zero_ediv, Int.ofNat_dvd_left, Nat.gcd_dvd_left, this] @@ -195,7 +195,7 @@ theorem div_int_inj {a b c d : ℤ} (hb0 : 0 < b) (hd0 : 0 < d) (h1 : Nat.Coprim theorem intCast_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, Int.ediv_zero] + simp only [Int.cast_zero, Int.zero_tdiv, zero_div, Int.ediv_zero] · have : (n : ℚ) ≠ 0 := by rwa [← coe_int_inj] at hn simp only [Int.ediv_self hn, Int.cast_one, Ne, not_false_iff, div_self this] @@ -247,9 +247,9 @@ theorem inv_intCast_num (a : ℤ) : (a : ℚ)⁻¹.num = Int.sign a := by rcases lt_trichotomy a 0 with lt | rfl | gt · obtain ⟨a, rfl⟩ : ∃ b, -b = a := ⟨-a, a.neg_neg⟩ simp at lt - simp [Rat.inv_neg, inv_intCast_num_of_pos lt, (Int.sign_eq_one_iff_pos _).mpr lt] - · rfl - · simp [inv_intCast_num_of_pos gt, (Int.sign_eq_one_iff_pos _).mpr gt] + simp [Rat.inv_neg, inv_intCast_num_of_pos lt, Int.sign_eq_one_iff_pos.mpr lt] + · simp + · simp [inv_intCast_num_of_pos gt, Int.sign_eq_one_iff_pos.mpr gt] @[simp] theorem inv_natCast_num (a : ℕ) : (a : ℚ)⁻¹.num = Int.sign a := @@ -268,7 +268,7 @@ theorem inv_intCast_den (a : ℤ) : (a : ℚ)⁻¹.den = if a = 0 then 1 else a. rw [if_neg (by omega)] simp only [Int.cast_neg, Rat.inv_neg, neg_den, inv_intCast_den_of_pos lt, Int.natAbs_neg] exact Int.eq_natAbs_of_zero_le (by omega) - · rfl + · simp · rw [if_neg (by omega)] simp only [inv_intCast_den_of_pos gt] exact Int.eq_natAbs_of_zero_le (by omega) diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 8a167a30145c91..4817d9024ec51a 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -717,12 +717,9 @@ theorem head_terminates_of_head_tail_terminates (s : WSeq α) [T : Terminates (h Terminates (head s) := (head_terminates_iff _).2 <| by rcases (head_terminates_iff _).1 T with ⟨⟨a, h⟩⟩ - simp? [tail] at h says simp only [tail, destruct_flatten] at h + simp? [tail] at h says simp only [tail, destruct_flatten, bind_map_left] at h rcases exists_of_mem_bind h with ⟨s', h1, _⟩ - unfold Functor.map at h1 - exact - let ⟨t, h3, _⟩ := Computation.exists_of_mem_map h1 - Computation.terminates_of_mem h3 + exact terminates_of_mem h1 theorem destruct_some_of_destruct_tail_some {s : WSeq α} {a} (h : some a ∈ destruct (tail s)) : ∃ a', some a' ∈ destruct s := by diff --git a/Mathlib/Data/Set/MemPartition.lean b/Mathlib/Data/Set/MemPartition.lean index 25b578031b0f68..5a13d62ad7baf8 100644 --- a/Mathlib/Data/Set/MemPartition.lean +++ b/Mathlib/Data/Set/MemPartition.lean @@ -118,7 +118,6 @@ lemma memPartitionSet_succ (f : ℕ → Set α) (n : ℕ) (a : α) [Decidable (a memPartitionSet f (n + 1) a = if a ∈ f n then memPartitionSet f n a ∩ f n else memPartitionSet f n a \ f n := by simp [memPartitionSet] - congr lemma memPartitionSet_mem (f : ℕ → Set α) (n : ℕ) (a : α) : memPartitionSet f n a ∈ memPartition f n := by diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index ca1d0ac6650681..4f812be71f3cde 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -420,7 +420,7 @@ def sigmaQuotientEquivOfLe {r s : Setoid α} (hle : r ≤ s) : (Σ q : Quotient s, Quotient (r.comap (Subtype.val : Quotient.mk s ⁻¹' {q} → α))) ≃ Quotient r := .trans (.symm <| .sigmaCongrRight fun _ ↦ .subtypeQuotientEquivQuotientSubtype - (s₁ := r) (s₂ := r.comap Subtype.val) _ (fun _ ↦ Iff.rfl) fun _ _ ↦ Iff.rfl) + (s₁ := r) (s₂ := r.comap Subtype.val) _ _ (fun _ ↦ Iff.rfl) fun _ _ ↦ Iff.rfl) (.sigmaFiberEquiv fun a ↦ a.lift (Quotient.mk s) fun _ _ h ↦ Quotient.sound <| hle h) end Setoid diff --git a/Mathlib/Data/Vector/Defs.lean b/Mathlib/Data/Vector/Defs.lean index 81ec5ceb945702..8c08c0254b5026 100644 --- a/Mathlib/Data/Vector/Defs.lean +++ b/Mathlib/Data/Vector/Defs.lean @@ -119,7 +119,7 @@ def take (i : ℕ) : Vector α n → Vector α (min i n) /-- Remove the element at position `i` from a vector of length `n`. -/ def eraseIdx (i : Fin n) : Vector α n → Vector α (n - 1) - | ⟨l, p⟩ => ⟨List.eraseIdx l i.1, by rw [l.length_eraseIdx] <;> rw [p]; exact i.2⟩ + | ⟨l, p⟩ => ⟨List.eraseIdx l i.1, by rw [l.length_eraseIdx_of_lt] <;> rw [p]; exact i.2⟩ @[deprecated (since := "2024-05-04")] alias removeNth := eraseIdx diff --git a/Mathlib/Data/ZMod/Defs.lean b/Mathlib/Data/ZMod/Defs.lean index 6977197f083cc6..76e27012548a54 100644 --- a/Mathlib/Data/ZMod/Defs.lean +++ b/Mathlib/Data/ZMod/Defs.lean @@ -104,7 +104,7 @@ namespace ZMod instance instUnique : Unique (ZMod 1) := Fin.uniqueFinOne instance fintype : ∀ (n : ℕ) [NeZero n], Fintype (ZMod n) - | 0, h => (h.ne rfl).elim + | 0, h => (h.ne _ rfl).elim | n + 1, _ => Fin.fintype (n + 1) instance infinite : Infinite (ZMod 0) := diff --git a/Mathlib/Dynamics/PeriodicPts.lean b/Mathlib/Dynamics/PeriodicPts.lean index 3f47192c4cc47e..160a6606f5e0fa 100644 --- a/Mathlib/Dynamics/PeriodicPts.lean +++ b/Mathlib/Dynamics/PeriodicPts.lean @@ -338,7 +338,7 @@ theorem not_isPeriodicPt_of_pos_of_lt_minimalPeriod : theorem IsPeriodicPt.minimalPeriod_dvd (hx : IsPeriodicPt f n x) : minimalPeriod f x ∣ n := (eq_or_lt_of_le <| n.zero_le).elim (fun hn0 => hn0 ▸ dvd_zero _) fun hn0 => -- Porting note: `Nat.dvd_iff_mod_eq_zero` gained explicit arguments - (Nat.dvd_iff_mod_eq_zero _ _).2 <| + Nat.dvd_iff_mod_eq_zero.2 <| (hx.mod <| isPeriodicPt_minimalPeriod f x).eq_zero_of_lt_minimalPeriod <| Nat.mod_lt _ <| hx.minimalPeriod_pos hn0 @@ -433,7 +433,7 @@ theorem periodicOrbit_length : (periodicOrbit f x).length = minimalPeriod f x := @[simp] theorem periodicOrbit_eq_nil_iff_not_periodic_pt : periodicOrbit f x = Cycle.nil ↔ x ∉ periodicPts f := by - simp only [periodicOrbit.eq_1, Cycle.coe_eq_nil, List.map_eq_nil, List.range_eq_nil] + simp only [periodicOrbit.eq_1, Cycle.coe_eq_nil, List.map_eq_nil_iff, List.range_eq_nil] exact minimalPeriod_eq_zero_iff_nmem_periodicPts theorem periodicOrbit_eq_nil_of_not_periodic_pt (h : x ∉ periodicPts f) : diff --git a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean index 21d9418e249a5c..787429c5e3827f 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/AlgebraicClosure.lean @@ -79,7 +79,7 @@ theorem spanEval_ne_top : spanEval k ≠ ⊤ := by rw [map_one, Finsupp.linearCombination_apply, Finsupp.sum, map_sum, Finset.sum_eq_zero] at hv · exact zero_ne_one hv intro j hj - rw [smul_eq_mul, map_mul, toSplittingField_evalXSelf (s := v.support) hj, + rw [smul_eq_mul, map_mul, toSplittingField_evalXSelf _ (s := v.support) hj, mul_zero] /-- A random maximal ideal that contains `spanEval k` -/ diff --git a/Mathlib/FieldTheory/KummerExtension.lean b/Mathlib/FieldTheory/KummerExtension.lean index 939bc161ec3f87..671a4b1bb374a7 100644 --- a/Mathlib/FieldTheory/KummerExtension.lean +++ b/Mathlib/FieldTheory/KummerExtension.lean @@ -272,6 +272,11 @@ theorem Polynomial.separable_X_pow_sub_C_of_irreducible : (X ^ n - C a).Separabl AdjoinRoot.algebraMap_eq, X_pow_sub_C_eq_prod (hζ.map_of_injective (algebraMap K _).injective) hn (root_X_pow_sub_C_pow n a), separable_prod_X_sub_C_iff'] + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/5376 we need to provide this helper instance. + -/ + have : MonoidHomClass (K →+* K[n√a]) K K[n√a] := inferInstance exact (hζ.map_of_injective (algebraMap K K[n√a]).injective).injOn_pow_mul (root_X_pow_sub_C_ne_zero (lt_of_le_of_ne (show 1 ≤ n from hn) (Ne.symm hn')) _) diff --git a/Mathlib/FieldTheory/PurelyInseparable.lean b/Mathlib/FieldTheory/PurelyInseparable.lean index c880a81782d7fb..93c3226de53d77 100644 --- a/Mathlib/FieldTheory/PurelyInseparable.lean +++ b/Mathlib/FieldTheory/PurelyInseparable.lean @@ -166,7 +166,7 @@ variable {F K} theorem isPurelyInseparable_iff : IsPurelyInseparable F E ↔ ∀ x : E, IsIntegral F x ∧ (IsSeparable F x → x ∈ (algebraMap F E).range) := - ⟨fun h x ↦ ⟨h.isIntegral' x, h.inseparable' x⟩, fun h ↦ ⟨⟨fun x ↦ (h x).1⟩, fun x ↦ (h x).2⟩⟩ + ⟨fun h x ↦ ⟨h.isIntegral' _ x, h.inseparable' x⟩, fun h ↦ ⟨⟨fun x ↦ (h x).1⟩, fun x ↦ (h x).2⟩⟩ /-- Transfer `IsPurelyInseparable` across an `AlgEquiv`. -/ theorem AlgEquiv.isPurelyInseparable (e : K ≃ₐ[F] E) [IsPurelyInseparable F K] : diff --git a/Mathlib/FieldTheory/Separable.lean b/Mathlib/FieldTheory/Separable.lean index c6a87d0dbf1fbf..fd852c7cdf7d7c 100644 --- a/Mathlib/FieldTheory/Separable.lean +++ b/Mathlib/FieldTheory/Separable.lean @@ -652,7 +652,7 @@ theorem IsSeparable.tower_bot {x : K} (h : IsSeparable F (algebraMap K E x)) : I variable (K E) in theorem Algebra.isSeparable_tower_bot_of_isSeparable [h : Algebra.IsSeparable F E] : Algebra.IsSeparable F K := - ⟨fun _ ↦ IsSeparable.tower_bot (h.isSeparable _)⟩ + ⟨fun _ ↦ IsSeparable.tower_bot (h.isSeparable _ _)⟩ end IsScalarTower diff --git a/Mathlib/FieldTheory/SeparableDegree.lean b/Mathlib/FieldTheory/SeparableDegree.lean index 6ebcf74215644a..fb1ec8ab2843c4 100644 --- a/Mathlib/FieldTheory/SeparableDegree.lean +++ b/Mathlib/FieldTheory/SeparableDegree.lean @@ -208,7 +208,7 @@ def embEquivOfAdjoinSplits {S : Set E} (hS : adjoin F S = ⊤) (hS ▸ isAlgebraic_adjoin (S := S) fun x hx ↦ (hK x hx).1) have halg := (topEquiv (F := F) (E := E)).isAlgebraic Classical.choice <| Function.Embedding.antisymm - (halg.algHomEmbeddingOfSplits (fun _ ↦ splits_of_mem_adjoin F (S := S) hK (hS ▸ mem_top)) _) + (halg.algHomEmbeddingOfSplits (fun _ ↦ splits_of_mem_adjoin F E (S := S) hK (hS ▸ mem_top)) _) (halg.algHomEmbeddingOfSplits (fun _ ↦ IsAlgClosed.splits_codomain _) _) /-- The `Field.finSepDegree F E` is equal to the cardinality of `E →ₐ[F] K` diff --git a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean index 5c5349f4adc953..a966c23b3b9185 100644 --- a/Mathlib/Geometry/RingedSpace/OpenImmersion.lean +++ b/Mathlib/Geometry/RingedSpace/OpenImmersion.lean @@ -167,7 +167,7 @@ noncomputable def invApp (U : Opens X) : @[simp, reassoc] theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.invApp (unop V) = + X.presheaf.map i ≫ H.invApp _ (unop V) = invApp f (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := by simp only [invApp, ← Category.assoc] rw [IsIso.comp_inv_eq] @@ -179,11 +179,11 @@ theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : instance (U : Opens X) : IsIso (invApp f U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : - inv (H.invApp U) = + inv (H.invApp _ U) = f.c.app (op (opensFunctor f |>.obj U)) ≫ X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := by - rw [← cancel_epi (H.invApp U), IsIso.hom_inv_id] + rw [← cancel_epi (H.invApp _ U), IsIso.hom_inv_id] delta invApp simp [← Functor.map_comp] @@ -195,7 +195,7 @@ theorem invApp_app (U : Opens X) : @[simp, reassoc] theorem app_invApp (U : Opens Y) : - f.c.app (op U) ≫ H.invApp ((Opens.map f.base).obj U) = + f.c.app (op U) ≫ H.invApp _ ((Opens.map f.base).obj U) = Y.presheaf.map ((homOfLE (Set.image_preimage_subset f.base U.1)).op : op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.base).obj U))) := by @@ -244,7 +244,7 @@ instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} @[elementwise, simp] theorem ofRestrict_invApp {C : Type*} [Category C] (X : PresheafedSpace C) {Y : TopCat} {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : - (PresheafedSpace.IsOpenImmersion.ofRestrict X h).invApp U = 𝟙 _ := by + (PresheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := by delta invApp rw [IsIso.comp_inv_eq, Category.id_comp] change X.presheaf.map _ = X.presheaf.map _ @@ -290,7 +290,7 @@ def pullbackConeOfLeftFst : base := pullback.fst _ _ c := { app := fun U => - hf.invApp (unop U) ≫ + hf.invApp _ (unop U) ≫ g.c.app (op (hf.base_open.isOpenMap.functor.obj (unop U))) ≫ Y.presheaf.map (eqToHom @@ -772,21 +772,21 @@ noncomputable def invApp (U : Opens X) : @[reassoc (attr := simp)] theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.invApp (unop V) = - H.invApp (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := + X.presheaf.map i ≫ H.invApp _ (unop V) = + H.invApp _ (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := PresheafedSpace.IsOpenImmersion.inv_naturality f i -instance (U : Opens X) : IsIso (H.invApp U) := by delta invApp; infer_instance +instance (U : Opens X) : IsIso (H.invApp _ U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : - inv (H.invApp U) = + inv (H.invApp _ U) = f.c.app (op (opensFunctor f |>.obj U)) ≫ X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := PresheafedSpace.IsOpenImmersion.inv_invApp f U @[reassoc (attr := simp)] theorem invApp_app (U : Opens X) : - H.invApp U ≫ f.c.app (op (opensFunctor f |>.obj U)) = + H.invApp _ U ≫ f.c.app (op (opensFunctor f |>.obj U)) = X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := PresheafedSpace.IsOpenImmersion.invApp_app f U @@ -794,7 +794,7 @@ attribute [elementwise] invApp_app @[reassoc (attr := simp)] theorem app_invApp (U : Opens Y) : - f.c.app (op U) ≫ H.invApp ((Opens.map f.base).obj U) = + f.c.app (op U) ≫ H.invApp _ ((Opens.map f.base).obj U) = Y.presheaf.map ((homOfLE (Set.image_preimage_subset f.base U.1)).op : op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.base).obj U))) := @@ -818,7 +818,7 @@ instance ofRestrict {X : TopCat} (Y : SheafedSpace C) {f : X ⟶ Y.carrier} @[elementwise, simp] theorem ofRestrict_invApp {C : Type*} [Category C] (X : SheafedSpace C) {Y : TopCat} {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : - (SheafedSpace.IsOpenImmersion.ofRestrict X h).invApp U = 𝟙 _ := + (SheafedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U /-- An open immersion is an iso if the underlying continuous map is epi. -/ @@ -1154,7 +1154,7 @@ is an open immersion iff every stalk map is an iso. theorem of_stalk_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) (hf : OpenEmbedding f.1.base) [stalk_iso : ∀ x : X.1, IsIso (f.stalkMap x)] : LocallyRingedSpace.IsOpenImmersion f := - SheafedSpace.IsOpenImmersion.of_stalk_iso hf (H := stalk_iso) + SheafedSpace.IsOpenImmersion.of_stalk_iso _ hf (H := stalk_iso) end OfStalkIso @@ -1180,21 +1180,21 @@ noncomputable def invApp (U : Opens X) : @[reassoc (attr := simp)] theorem inv_naturality {U V : (Opens X)ᵒᵖ} (i : U ⟶ V) : - X.presheaf.map i ≫ H.invApp (unop V) = - H.invApp (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := + X.presheaf.map i ≫ H.invApp _ (unop V) = + H.invApp _ (unop U) ≫ Y.presheaf.map (opensFunctor f |>.op.map i) := PresheafedSpace.IsOpenImmersion.inv_naturality f.1 i -instance (U : Opens X) : IsIso (H.invApp U) := by delta invApp; infer_instance +instance (U : Opens X) : IsIso (H.invApp _ U) := by delta invApp; infer_instance theorem inv_invApp (U : Opens X) : - inv (H.invApp U) = + inv (H.invApp _ U) = f.1.c.app (op (opensFunctor f |>.obj U)) ≫ X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := PresheafedSpace.IsOpenImmersion.inv_invApp f.1 U @[reassoc (attr := simp)] theorem invApp_app (U : Opens X) : - H.invApp U ≫ f.1.c.app (op (opensFunctor f |>.obj U)) = + H.invApp _ U ≫ f.1.c.app (op (opensFunctor f |>.obj U)) = X.presheaf.map (eqToHom (by simp [Opens.map, Set.preimage_image_eq _ H.base_open.inj])) := PresheafedSpace.IsOpenImmersion.invApp_app f.1 U @@ -1202,7 +1202,7 @@ attribute [elementwise] invApp_app @[reassoc (attr := simp)] theorem app_invApp (U : Opens Y) : - f.1.c.app (op U) ≫ H.invApp ((Opens.map f.1.base).obj U) = + f.1.c.app (op U) ≫ H.invApp _ ((Opens.map f.1.base).obj U) = Y.presheaf.map ((homOfLE (Set.image_preimage_subset f.1.base U.1)).op : op U ⟶ op (opensFunctor f |>.obj ((Opens.map f.1.base).obj U))) := @@ -1211,7 +1211,7 @@ theorem app_invApp (U : Opens Y) : /-- A variant of `app_inv_app` that gives an `eqToHom` instead of `homOfLe`. -/ @[reassoc] theorem app_inv_app' (U : Opens Y) (hU : (U : Set Y) ⊆ Set.range f.1.base) : - f.1.c.app (op U) ≫ H.invApp ((Opens.map f.1.base).obj U) = + f.1.c.app (op U) ≫ H.invApp _ ((Opens.map f.1.base).obj U) = Y.presheaf.map (eqToHom <| le_antisymm (Set.image_preimage_subset f.1.base U.1) <| @@ -1226,7 +1226,7 @@ instance ofRestrict {X : TopCat} (Y : LocallyRingedSpace) {f : X ⟶ Y.carrier} @[elementwise, simp] theorem ofRestrict_invApp (X : LocallyRingedSpace) {Y : TopCat} {f : Y ⟶ TopCat.of X.carrier} (h : OpenEmbedding f) (U : Opens (X.restrict h).carrier) : - (LocallyRingedSpace.IsOpenImmersion.ofRestrict X h).invApp U = 𝟙 _ := + (LocallyRingedSpace.IsOpenImmersion.ofRestrict X h).invApp _ U = 𝟙 _ := PresheafedSpace.IsOpenImmersion.ofRestrict_invApp _ h U instance stalk_iso (x : X) : IsIso (f.stalkMap x) := diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index b5ac976467db82..4d9076090a0486 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -149,7 +149,7 @@ theorem pullback_base (i j k : D.J) (S : Set (D.V (i, j)).carrier) : /-- The red and the blue arrows in ![this diagram](https://i.imgur.com/0GiBUh6.png) commute. -/ @[simp, reassoc] theorem f_invApp_f_app (i j k : D.J) (U : Opens (D.V (i, j)).carrier) : - (D.f_open i j).invApp U ≫ (D.f i k).c.app _ = + (D.f_open i j).invApp _ U ≫ (D.f i k).c.app _ = (π₁ i, j, k).c.app (op U) ≫ (π₂⁻¹ i, j, k) (unop _) ≫ (D.V _).presheaf.map @@ -161,7 +161,7 @@ theorem f_invApp_f_app (i j k : D.J) (U : Opens (D.V (i, j)).carrier) : apply pullback_base)) := by have := PresheafedSpace.congr_app (@pullback.condition _ _ _ _ _ (D.f i j) (D.f i k) _) dsimp only [comp_c_app] at this - rw [← cancel_epi (inv ((D.f_open i j).invApp U)), IsIso.inv_hom_id_assoc, + rw [← cancel_epi (inv ((D.f_open i j).invApp _ U)), IsIso.inv_hom_id_assoc, IsOpenImmersion.inv_invApp] simp_rw [Category.assoc] erw [(π₁ i, j, k).c.naturality_assoc, reassoc_of% this, ← Functor.map_comp_assoc, @@ -271,7 +271,7 @@ def opensImagePreimageMap (i j : D.J) (U : Opens (D.U i).carrier) : (Opens.map (𝖣.ι j).base).obj ((D.ι_openEmbedding i).isOpenMap.functor.obj U)) := (D.f i j).c.app (op U) ≫ (D.t j i).c.app _ ≫ - (D.f_open j i).invApp (unop _) ≫ + (D.f_open j i).invApp _ (unop _) ≫ (𝖣.U j).presheaf.map (eqToHom (D.ι_image_preimage_eq i j U)).op theorem opensImagePreimageMap_app' (i j k : D.J) (U : Opens (D.U i).carrier) : @@ -436,19 +436,9 @@ abbrev ιInvAppπEqMap {i : D.J} (U : Opens (D.U i).carrier) := theorem π_ιInvApp_π (i j : D.J) (U : Opens (D.U i).carrier) : D.diagramOverOpenπ U i ≫ D.ιInvAppπEqMap U ≫ D.ιInvApp U ≫ D.diagramOverOpenπ U j = D.diagramOverOpenπ U j := by - -- Porting note: originally, the proof of monotonicity was left a blank and proved in the end - -- but Lean 4 doesn't like this any more, so the proof is restructured - rw [← @cancel_mono (f := (componentwiseDiagram 𝖣.diagram.multispan _).map - (Quiver.Hom.op (WalkingMultispan.Hom.snd (i, j))) ≫ 𝟙 _) _ _ (by - rw [Category.comp_id] - apply (config := { allowSynthFailures := true }) mono_comp - change Mono ((_ ≫ D.f j i).c.app _) - rw [comp_c_app] - apply (config := { allowSynthFailures := true }) mono_comp - · erw [D.ι_image_preimage_eq i j U] - infer_instance - · have : IsIso (D.t i j).c := by apply c_isIso_of_iso - infer_instance)] + rw [← @cancel_mono + (f := (componentwiseDiagram 𝖣.diagram.multispan _).map + (Quiver.Hom.op (WalkingMultispan.Hom.snd (i, j))) ≫ 𝟙 _) ..] simp_rw [Category.assoc] rw [limit.w_assoc] erw [limit.lift_π_assoc] @@ -466,6 +456,15 @@ theorem π_ιInvApp_π (i j : D.J) (U : Opens (D.U i).carrier) : convert limit.w (componentwiseDiagram 𝖣.diagram.multispan _) (Quiver.Hom.op (WalkingMultispan.Hom.fst (i, j))) + · rw [Category.comp_id] + apply (config := { allowSynthFailures := true }) mono_comp + change Mono ((_ ≫ D.f j i).c.app _) + rw [comp_c_app] + apply (config := { allowSynthFailures := true }) mono_comp + · erw [D.ι_image_preimage_eq i j U] + infer_instance + · have : IsIso (D.t i j).c := by apply c_isIso_of_iso + infer_instance /-- `ιInvApp` is the inverse of `D.ι i` on `U`. -/ theorem π_ιInvApp_eq_id (i : D.J) (U : Opens (D.U i).carrier) : diff --git a/Mathlib/GroupTheory/CommutingProbability.lean b/Mathlib/GroupTheory/CommutingProbability.lean index c02743ac71a509..40bcb398e5f676 100644 --- a/Mathlib/GroupTheory/CommutingProbability.lean +++ b/Mathlib/GroupTheory/CommutingProbability.lean @@ -131,6 +131,7 @@ lemma commProb_odd {n : ℕ} (hn : Odd n) : qify [show 2 ∣ n + 3 by rw [Nat.dvd_iff_mod_eq_zero, Nat.add_mod, Nat.odd_iff.mp hn]] rw [div_div, ← mul_assoc] congr + norm_num private lemma div_two_lt {n : ℕ} (h0 : n ≠ 0) : n / 2 < n := Nat.div_lt_self (Nat.pos_of_ne_zero h0) (lt_add_one 1) diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index e1096e68292116..4dfea2fbf13f63 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -128,7 +128,7 @@ theorem not_step_nil : ¬Step [] L := by generalize h' : [] = L' intro h cases' h with L₁ L₂ - simp [List.nil_eq_append] at h' + simp [List.nil_eq_append_iff] at h' @[to_additive] theorem Step.cons_left_iff {a : α} {b : Bool} : @@ -287,7 +287,8 @@ theorem red_iff_irreducible {x1 b1 x2 b2} (h : (x1, b1) ≠ (x2, b2)) : generalize eq : [(x1, not b1), (x2, b2)] = L' intro L h' cases h' - simp [List.cons_eq_append, List.nil_eq_append] at eq + simp only [List.cons_eq_append_iff, List.cons.injEq, Prod.mk.injEq, and_false, + List.nil_eq_append_iff, exists_const, or_self, or_false, List.cons_ne_nil] at eq rcases eq with ⟨rfl, ⟨rfl, rfl⟩, ⟨rfl, rfl⟩, rfl⟩ simp at h diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index f036e80ab22448..7bf00ac53f9d30 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -404,6 +404,12 @@ noncomputable def equivSubgroupOrbitsQuotientGroup [IsPretransitive α β] convert one_mem H rw [inv_mul_eq_one, eq_comm, ← inv_mul_eq_one, ← Subgroup.mem_bot, ← free (g⁻¹ • x), mem_stabilizer_iff, mul_smul, (exists_smul_eq α (g⁻¹ • x) x).choose_spec] + #adaptation_note + /-- + After https://github.com/leanprover/lean4/pull/5376 we need to search for this instance explicitly. + TODO: change `convert` to more agressively solve such goals with `infer_instance` itself. + -/ + infer_instance end MulAction diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index 87db896052260d..9e33d48b219962 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -717,7 +717,7 @@ automatic in the case of a finite cancellative monoid. -/ `addOrderOf_nsmul` but with one assumption less which is automatic in the case of a finite cancellative additive monoid."] theorem orderOf_pow (x : G) : orderOf (x ^ n) = orderOf x / gcd (orderOf x) n := - (isOfFinOrder_of_finite _).orderOf_pow _ + (isOfFinOrder_of_finite _).orderOf_pow .. @[to_additive] theorem mem_powers_iff_mem_range_orderOf [DecidableEq G] : diff --git a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean index 57afc2d245faf0..c16d73ee548bd5 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Concrete.lean @@ -316,7 +316,7 @@ theorem toList_formPerm_nontrivial (l : List α) (hl : 2 ≤ l.length) (hn : Nod · refine ext_getElem (by simp) fun k hk hk' => ?_ simp only [get_eq_getElem, formPerm_pow_apply_getElem _ hn, zero_add, getElem_map, getElem_range, Nat.mod_eq_of_lt hk'] - · simpa [hs] using get_mem _ _ _ + · simp [hs] theorem toList_formPerm_isRotated_self (l : List α) (hl : 2 ≤ l.length) (hn : Nodup l) (x : α) (hx : x ∈ l) : toList (formPerm l) x ~r l := by diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index a7d22e67b1e55d..269920ef0aab37 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -417,7 +417,7 @@ theorem sign_trans_trans_symm [DecidableEq β] [Fintype β] (f : Perm β) (e : theorem sign_prod_list_swap {l : List (Perm α)} (hl : ∀ g ∈ l, IsSwap g) : sign l.prod = (-1) ^ l.length := by have h₁ : l.map sign = List.replicate l.length (-1) := - List.eq_replicate.2 + List.eq_replicate_iff.2 ⟨by simp, fun u hu => let ⟨g, hg⟩ := List.mem_map.1 hu hg.2 ▸ (hl _ hg.1).sign_eq⟩ diff --git a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean index 7ae6931a6aa29d..9f90f39e9ee1f7 100644 --- a/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean +++ b/Mathlib/LinearAlgebra/Alternating/DomCoprod.lean @@ -226,7 +226,7 @@ theorem MultilinearMap.domCoprod_alternization [DecidableEq ιa] [DecidableEq ι -- unfold the quotient mess left by `Finset.sum_partition` -- Porting note: Was `conv in .. => ..`. erw - [@Finset.filter_congr _ _ (fun a => @Quotient.decidableEq _ _ + [@Finset.filter_congr _ _ _ (fun a => @Quotient.decidableEq _ _ (QuotientGroup.leftRelDecidable (MonoidHom.range (Perm.sumCongrHom ιa ιb))) (Quotient.mk (QuotientGroup.leftRel (MonoidHom.range (Perm.sumCongrHom ιa ιb))) a) (Quotient.mk'' σ)) _ (s := Finset.univ) diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index 4359ba5b328d06..e5f6c2be853838 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -154,7 +154,7 @@ theorem evenOdd_induction (n : ZMod 2) {motive : ∀ x, x ∈ evenOdd Q n → Pr motive (ι Q m₁ * ι Q m₂ * x) (zero_add n ▸ SetLike.mul_mem_graded (ι_mul_ι_mem_evenOdd_zero Q m₁ m₂) hx)) (x : CliffordAlgebra Q) (hx : x ∈ evenOdd Q n) : motive x hx := by - apply Submodule.iSup_induction' (C := motive) _ (range_ι_pow 0 (Submodule.zero_mem _)) add + apply Submodule.iSup_induction' (C := motive) _ _ (range_ι_pow 0 (Submodule.zero_mem _)) add refine Subtype.rec ?_ simp_rw [ZMod.natCast_eq_iff, add_comm n.val] rintro n' ⟨k, rfl⟩ xv @@ -197,7 +197,7 @@ theorem even_induction {motive : ∀ x, x ∈ evenOdd Q 0 → Prop} motive (ι Q m₁ * ι Q m₂ * x) (zero_add (0 : ZMod 2) ▸ SetLike.mul_mem_graded (ι_mul_ι_mem_evenOdd_zero Q m₁ m₂) hx)) (x : CliffordAlgebra Q) (hx : x ∈ evenOdd Q 0) : motive x hx := by - refine evenOdd_induction (motive := motive) (fun rx => ?_) add ι_mul_ι_mul x hx + refine evenOdd_induction _ _ (motive := motive) (fun rx => ?_) add ι_mul_ι_mul x hx rintro ⟨r, rfl⟩ exact algebraMap r @@ -213,7 +213,7 @@ theorem odd_induction {P : ∀ x, x ∈ evenOdd Q 1 → Prop} P (CliffordAlgebra.ι Q m₁ * CliffordAlgebra.ι Q m₂ * x) (zero_add (1 : ZMod 2) ▸ SetLike.mul_mem_graded (ι_mul_ι_mem_evenOdd_zero Q m₁ m₂) hx)) (x : CliffordAlgebra Q) (hx : x ∈ evenOdd Q 1) : P x hx := by - refine evenOdd_induction (motive := P) (fun ιv => ?_) add ι_mul_ι_mul x hx + refine evenOdd_induction _ _ (motive := P) (fun ιv => ?_) add ι_mul_ι_mul x hx -- Porting note: was `simp_rw [ZMod.val_one, pow_one]`, lean4#1926 intro h; rw [ZMod.val_one, pow_one] at h; revert h rintro ⟨v, rfl⟩ diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 1d2a2669cab759..2b1fa24c031b5a 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -183,7 +183,7 @@ def LinearMap.dualMap (f : M₁ →ₗ[R] M₂) : Dual R M₂ →ₗ[R] Dual R M -- Porting note: with reducible def need to specify some parameters to transpose explicitly Module.Dual.transpose (R := R) f -lemma LinearMap.dualMap_eq_lcomp (f : M₁ →ₗ[R] M₂) : f.dualMap = f.lcomp R := rfl +lemma LinearMap.dualMap_eq_lcomp (f : M₁ →ₗ[R] M₂) : f.dualMap = f.lcomp R R := rfl -- Porting note: with reducible def need to specify some parameters to transpose explicitly theorem LinearMap.dualMap_def (f : M₁ →ₗ[R] M₂) : f.dualMap = Module.Dual.transpose (R := R) f := diff --git a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean index acd30d6e430a97..0b10bbf352e99f 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional/Defs.lean @@ -671,7 +671,7 @@ noncomputable def divisionRingOfFiniteDimensional (F K : Type*) [Field F] [Ring inv x := letI := Classical.decEq K if H : x = 0 then 0 else Classical.choose <| FiniteDimensional.exists_mul_eq_one F H - mul_inv_cancel x hx := show x * dite _ (h := _) _ = _ by + mul_inv_cancel x hx := show x * dite _ (h := _) _ _ = _ by rw [dif_neg hx] exact (Classical.choose_spec (FiniteDimensional.exists_mul_eq_one F hx):) inv_zero := dif_pos rfl diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index b814a9748abac8..84ce3839bb7552 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -495,8 +495,8 @@ theorem adjugate_adjugate (A : Matrix n n α) (h : Fintype.card n ≠ 1) : let A' := mvPolynomialX n n ℤ suffices adjugate (adjugate A') = det A' ^ (Fintype.card n - 2) • A' by rw [← mvPolynomialX_mapMatrix_aeval ℤ A, ← AlgHom.map_adjugate, ← AlgHom.map_adjugate, this, - ← AlgHom.map_det, ← map_pow (MvPolynomial.aeval _), AlgHom.mapMatrix_apply, - AlgHom.mapMatrix_apply, Matrix.map_smul' _ _ _ (_root_.map_mul _)] + ← AlgHom.map_det, ← map_pow (MvPolynomial.aeval fun p : n × n ↦ A p.1 p.2), + AlgHom.mapMatrix_apply, AlgHom.mapMatrix_apply, Matrix.map_smul' _ _ _ (_root_.map_mul _)] have h_card' : Fintype.card n - 2 + 1 = Fintype.card n - 1 := by simp [h_card] have is_reg : IsSMulRegular (MvPolynomial (n × n) ℤ) (det A') := fun x y => mul_left_cancel₀ (det_mvPolynomialX_ne_zero n ℤ) diff --git a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean index cc04b3b64ac378..4a0d9f53997f48 100644 --- a/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean +++ b/Mathlib/LinearAlgebra/Matrix/Charpoly/Coeff.lean @@ -204,7 +204,7 @@ lemma derivative_det_one_add_X_smul_aux {n} (M : Matrix (Fin n) (Fin n) R) : rw [det_eq_zero_of_column_eq_zero 0, eval_zero, mul_zero] intro j rw [submatrix_apply, Fin.succAbove_of_castSucc_lt, one_apply_ne] - · exact (bne_iff_ne (Fin.succ j) (Fin.castSucc 0)).mp rfl + · exact (bne_iff_ne (a := Fin.succ j) (b := Fin.castSucc 0)).mp rfl · rw [Fin.castSucc_zero]; exact lt_of_le_of_ne (Fin.zero_le _) hi.symm · exact fun H ↦ (H <| Finset.mem_univ _).elim diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 9490aa658bb95e..c33327a8858aeb 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -747,7 +747,7 @@ theorem det_fin_one_of (a : R) : det !![a] = a := theorem det_fin_two (A : Matrix (Fin 2) (Fin 2) R) : det A = A 0 0 * A 1 1 - A 0 1 * A 1 0 := by simp only [det_succ_row_zero, det_unique, Fin.default_eq_zero, submatrix_apply, Fin.succ_zero_eq_one, Fin.sum_univ_succ, Fin.val_zero, Fin.zero_succAbove, univ_unique, - Fin.val_succ, Fin.coe_fin_one, Fin.succ_succAbove_zero, sum_singleton] + Fin.val_succ, Fin.val_eq_zero, Fin.succ_succAbove_zero, sum_singleton] ring @[simp] @@ -763,7 +763,7 @@ theorem det_fin_three (A : Matrix (Fin 3) (Fin 3) R) : simp only [det_succ_row_zero, ← Nat.not_even_iff_odd, submatrix_apply, Fin.succ_zero_eq_one, submatrix_submatrix, det_unique, Fin.default_eq_zero, comp_apply, Fin.succ_one_eq_two, Fin.sum_univ_succ, Fin.val_zero, Fin.zero_succAbove, univ_unique, Fin.val_succ, - Fin.coe_fin_one, Fin.succ_succAbove_zero, sum_singleton, Fin.succ_succAbove_one, even_add_self] + Fin.val_eq_zero, Fin.succ_succAbove_zero, sum_singleton, Fin.succ_succAbove_one, even_add_self] ring end Matrix diff --git a/Mathlib/LinearAlgebra/Semisimple.lean b/Mathlib/LinearAlgebra/Semisimple.lean index 0b3d984857b624..3aeb94be63ed63 100644 --- a/Mathlib/LinearAlgebra/Semisimple.lean +++ b/Mathlib/LinearAlgebra/Semisimple.lean @@ -80,7 +80,7 @@ lemma eq_zero_of_isNilpotent_isSemisimple (hn : IsNilpotent f) (hs : f.IsSemisim have ⟨n, h0⟩ := hn rw [← aeval_X (R := R) f]; rw [← aeval_X_pow (R := R) f] at h0 rw [← RingHom.mem_ker, ← AEval.annihilator_eq_ker_aeval (M := M)] at h0 ⊢ - exact hs.annihilator_isRadical ⟨n, h0⟩ + exact hs.annihilator_isRadical _ _ ⟨n, h0⟩ @[simp] lemma isSemisimple_sub_algebraMap_iff {μ : R} : diff --git a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean index d06c680fb68593..18c2623b098064 100644 --- a/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean +++ b/Mathlib/LinearAlgebra/TensorProduct/Graded/Internal.lean @@ -220,7 +220,7 @@ theorem tmul_coe_mul_one_tmul {j₁ : ι} (a₁ : A) (b₁ : ℬ j₁) (b₂ : B theorem tmul_one_mul_one_tmul (a₁ : A) (b₂ : B) : (a₁ ᵍ⊗ₜ[R] (1 : B) * (1 : A) ᵍ⊗ₜ[R] b₂ : 𝒜 ᵍ⊗[R] ℬ) = (a₁ : A) ᵍ⊗ₜ (b₂ : B) := by convert tmul_coe_mul_zero_coe_tmul 𝒜 ℬ - a₁ (@GradedMonoid.GOne.one _ (ℬ ·) _ _) (@GradedMonoid.GOne.one _ (𝒜 ·) _ _) b₂ + a₁ (GradedMonoid.GOne.one (A := (ℬ ·))) (GradedMonoid.GOne.one (A := (𝒜 ·))) b₂ · rw [SetLike.coe_gOne, mul_one] · rw [SetLike.coe_gOne, one_mul] diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 8fd66baa31e7d8..78b70b4f493080 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -975,7 +975,7 @@ theorem not_beq_of_ne {α : Type*} [BEq α] [LawfulBEq α] {a b : α} (ne : a fun h => ne (eq_of_beq h) theorem beq_eq_decide {α : Type*} [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by - rw [← beq_iff_eq a b] + rw [← beq_iff_eq (a := a) (b := b)] cases a == b <;> simp @[simp] lemma beq_eq_beq {α β : Type*} [BEq α] [LawfulBEq α] [BEq β] [LawfulBEq β] {a₁ a₂ : α} diff --git a/Mathlib/Logic/Encodable/Basic.lean b/Mathlib/Logic/Encodable/Basic.lean index 60412ec2012925..c111c09a267370 100644 --- a/Mathlib/Logic/Encodable/Basic.lean +++ b/Mathlib/Logic/Encodable/Basic.lean @@ -555,9 +555,13 @@ theorem sequence_mono_nat {r : β → β → Prop} {f : α → β} (hf : Directe · exact (Classical.choose_spec (hf p p)).1 · exact (Classical.choose_spec (hf p a)).1 +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem rel_sequence {r : β → β → Prop} {f : α → β} (hf : Directed r f) (a : α) : r (f a) (f (hf.sequence f (encode a + 1))) := by - simp only [Directed.sequence, add_eq, add_zero, encodek, and_self] + simp only [Directed.sequence, add_eq, add_zero, encodek, _root_.and_self] exact (Classical.choose_spec (hf _ a)).2 variable [Preorder β] {f : α → β} diff --git a/Mathlib/Logic/Equiv/Array.lean b/Mathlib/Logic/Equiv/Array.lean index cd9c6fbef7d1fc..eb98b0edd323b8 100644 --- a/Mathlib/Logic/Equiv/Array.lean +++ b/Mathlib/Logic/Equiv/Array.lean @@ -42,7 +42,7 @@ namespace Equiv /-- The natural equivalence between arrays and lists. -/ def arrayEquivList (α : Type*) : Array α ≃ List α := - ⟨Array.data, Array.mk, fun _ => rfl, fun _ => rfl⟩ + ⟨Array.toList, Array.mk, fun _ => rfl, fun _ => rfl⟩ end Equiv diff --git a/Mathlib/Logic/Nonempty.lean b/Mathlib/Logic/Nonempty.lean index 1a52ca02ab2d48..1c8198b9475d39 100644 --- a/Mathlib/Logic/Nonempty.lean +++ b/Mathlib/Logic/Nonempty.lean @@ -99,13 +99,6 @@ theorem Nonempty.elim_to_inhabited {α : Sort*} [h : Nonempty α] {p : Prop} (f p := h.elim <| f ∘ Inhabited.mk -protected instance Prod.instNonempty {α β} [h : Nonempty α] [h2 : Nonempty β] : Nonempty (α × β) := - h.elim fun g ↦ h2.elim fun g2 ↦ ⟨⟨g, g2⟩⟩ - -protected instance Pi.instNonempty {ι : Sort*} {α : ι → Sort*} [∀ i, Nonempty (α i)] : - Nonempty (∀ i, α i) := - ⟨fun _ ↦ Classical.arbitrary _⟩ - theorem Classical.nonempty_pi {ι} {α : ι → Sort*} : Nonempty (∀ i, α i) ↔ ∀ i, Nonempty (α i) := ⟨fun ⟨f⟩ a ↦ ⟨f a⟩, @Pi.instNonempty _ _⟩ diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean index e0029edb7097f6..4bd14a8a98d53a 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Order.lean @@ -568,7 +568,8 @@ theorem Measurable.isLUB_of_mem {ι} [Countable ι] {f : ι → δ → α} {g g' · simp [hb, hg' hb] rw [this] exact Measurable.piecewise hs measurable_const g'_meas - · let f' : ι → δ → α := fun i ↦ s.piecewise (f i) g' + · have : Nonempty ι := ⟨i⟩ + let f' : ι → δ → α := fun i ↦ s.piecewise (f i) g' suffices ∀ b, IsLUB { a | ∃ i, f' i b = a } (g b) from Measurable.isLUB (fun i ↦ Measurable.piecewise hs (hf i) g'_meas) this intro b @@ -576,14 +577,7 @@ theorem Measurable.isLUB_of_mem {ι} [Countable ι] {f : ι → δ → α} {g g' · have A : ∀ i, f' i b = f i b := fun i ↦ by simp [f', hb] simpa [A] using hg b hb · have A : ∀ i, f' i b = g' b := fun i ↦ by simp [f', hb] - have : {a | ∃ (_i : ι), g' b = a} = {g' b} := by - apply Subset.antisymm - · rintro - ⟨_j, rfl⟩ - simp only [mem_singleton_iff] - · rintro - rfl - exact ⟨i, rfl⟩ - simp only [exists_prop'] at this - simp [A, this, hg' hb, isLUB_singleton] + simp [A, hg' hb, isLUB_singleton] theorem AEMeasurable.isLUB {ι} {μ : Measure δ} [Countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, AEMeasurable (f i) μ) (hg : ∀ᵐ b ∂μ, IsLUB { a | ∃ i, f i b = a } (g b)) : diff --git a/Mathlib/MeasureTheory/Measure/Content.lean b/Mathlib/MeasureTheory/Measure/Content.lean index 5306eb0281091c..b07a09b9ad57e5 100644 --- a/Mathlib/MeasureTheory/Measure/Content.lean +++ b/Mathlib/MeasureTheory/Measure/Content.lean @@ -310,7 +310,7 @@ variable [S : MeasurableSpace G] [BorelSpace G] /-- For the outer measure coming from a content, all Borel sets are measurable. -/ theorem borel_le_caratheodory : S ≤ μ.outerMeasure.caratheodory := by - rw [@BorelSpace.measurable_eq G _ _] + rw [BorelSpace.measurable_eq (α := G)] refine MeasurableSpace.generateFrom_le ?_ intro U hU rw [μ.outerMeasure_caratheodory] diff --git a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean index 2a839ffd63d950..c2c75e666e4b2e 100644 --- a/Mathlib/ModelTheory/Algebra/Ring/Basic.lean +++ b/Mathlib/ModelTheory/Algebra/Ring/Basic.lean @@ -140,7 +140,7 @@ theorem card_ring : card Language.ring = 5 := by have : Fintype.card Language.ring.Symbols = 5 := rfl simp [Language.card, this] -open Language ring Structure +open Language Structure /-- A Type `R` is a `CompatibleRing` if it is a structure for the language of rings and this structure is the same as the structure already given on `R` by the classes `Add`, `Mul` etc. diff --git a/Mathlib/ModelTheory/Basic.lean b/Mathlib/ModelTheory/Basic.lean index e19707451b558b..baa0ada8d42712 100644 --- a/Mathlib/ModelTheory/Basic.lean +++ b/Mathlib/ModelTheory/Basic.lean @@ -247,14 +247,14 @@ theorem nonempty_of_nonempty_constants [h : Nonempty L.Constants] : Nonempty M : /-- `HomClass L F M N` states that `F` is a type of `L`-homomorphisms. You should extend this typeclass when you extend `FirstOrder.Language.Hom`. -/ -class HomClass (L : outParam Language) (F M N : Type*) +class HomClass (L : outParam Language) (F : Type*) (M N : outParam Type*) [FunLike F M N] [L.Structure M] [L.Structure N] : Prop where map_fun : ∀ (φ : F) {n} (f : L.Functions n) (x), φ (funMap f x) = funMap f (φ ∘ x) map_rel : ∀ (φ : F) {n} (r : L.Relations n) (x), RelMap r x → RelMap r (φ ∘ x) /-- `StrongHomClass L F M N` states that `F` is a type of `L`-homomorphisms which preserve relations in both directions. -/ -class StrongHomClass (L : outParam Language) (F M N : Type*) +class StrongHomClass (L : outParam Language) (F : Type*) (M N : outParam Type*) [FunLike F M N] [L.Structure M] [L.Structure N] : Prop where map_fun : ∀ (φ : F) {n} (f : L.Functions n) (x), φ (funMap f x) = funMap f (φ ∘ x) map_rel : ∀ (φ : F) {n} (r : L.Relations n) (x), RelMap r (φ ∘ x) ↔ RelMap r x diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index fab8363491989a..a9157aba577e84 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -83,7 +83,7 @@ theorem listDecode_encode_list (l : List (L.Term α)) : simp only [h, length_append, length_map, length_finRange, le_add_iff_nonneg_right, _root_.zero_le, ↓reduceDIte, getElem_fin, cons.injEq, func.injEq, heq_eq_eq, true_and] refine ⟨funext (fun i => ?_), ?_⟩ - · rw [List.getElem_append, List.getElem_map, List.getElem_finRange] + · rw [List.getElem_append_left, List.getElem_map, List.getElem_finRange] simp only [length_map, length_finRange, i.2] · simp only [length_map, length_finRange, drop_left'] @@ -244,7 +244,7 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : simp only [Option.join, map_append, map_map, Option.bind_eq_some, id, exists_eq_right, get?_eq_some, length_append, length_map, length_finRange] refine ⟨lt_of_lt_of_le i.2 le_self_add, ?_⟩ - rw [get_eq_getElem, getElem_append, getElem_map] + rw [get_eq_getElem, getElem_append_left, getElem_map] · simp only [getElem_finRange, Fin.eta, Function.comp_apply, Sum.getLeft?] · simp only [length_map, length_finRange, is_lt] rw [dif_pos] diff --git a/Mathlib/ModelTheory/Syntax.lean b/Mathlib/ModelTheory/Syntax.lean index fb5f3691aab1a2..7fa52ce76edb5b 100644 --- a/Mathlib/ModelTheory/Syntax.lean +++ b/Mathlib/ModelTheory/Syntax.lean @@ -576,8 +576,8 @@ theorem relabel_sum_inl (φ : L.BoundedFormula α n) : | falsum => rfl | equal => simp [Fin.natAdd_zero, castLE_of_eq, mapTermRel] | rel => simp [Fin.natAdd_zero, castLE_of_eq, mapTermRel]; rfl - | imp _ _ ih1 ih2 => simp [mapTermRel, ih1, ih2] - | all _ ih3 => simp [mapTermRel, ih3, castLE] + | imp _ _ ih1 ih2 => simp_all [mapTermRel] + | all _ ih3 => simp_all [mapTermRel] /-- Substitutes the variables in a given formula with terms. -/ def subst {n : ℕ} (φ : L.BoundedFormula α n) (f : α → L.Term β) : L.BoundedFormula β n := diff --git a/Mathlib/NumberTheory/ADEInequality.lean b/Mathlib/NumberTheory/ADEInequality.lean index 4c41bee0d6a332..aabf630647bb3d 100644 --- a/Mathlib/NumberTheory/ADEInequality.lean +++ b/Mathlib/NumberTheory/ADEInequality.lean @@ -148,7 +148,7 @@ theorem Admissible.one_lt_sumInv {pqr : Multiset ℕ+} : Admissible pqr → 1 < all_goals rw [← H, E', sumInv_pqr] conv_rhs => simp only [OfNat.ofNat, PNat.mk_coe] - rfl + norm_num theorem lt_three {p q r : ℕ+} (hpq : p ≤ q) (hqr : q ≤ r) (H : 1 < sumInv {p, q, r}) : p < 3 := by have h3 : (0 : ℚ) < 3 := by norm_num diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 772c5d12f2dff3..7f3a1d190e6f75 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -121,6 +121,10 @@ theorem bertrand_main_inequality {n : ℕ} (n_large : 512 ≤ n) : · norm_num1 · exact cast_div_le.trans (by norm_cast) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- A lemma that tells us that, in the case where Bertrand's postulate does not hold, the prime factorization of the central binomial coefficient only has factors at most `2 * n / 3 + 1`. -/ @@ -134,7 +138,7 @@ theorem centralBinom_factorization_small (n : ℕ) (n_large : 2 < n) rw [Finset.mem_range, Nat.lt_succ_iff] at hx h2x rw [not_le, div_lt_iff_lt_mul' three_pos, mul_comm x] at h2x replace no_prime := not_exists.mp no_prime x - rw [← and_assoc, not_and', not_and_or, not_lt] at no_prime + rw [← _root_.and_assoc, not_and', not_and_or, not_lt] at no_prime cases' no_prime hx with h h · rw [factorization_eq_zero_of_non_prime n.centralBinom h, Nat.pow_zero] · rw [factorization_centralBinom_of_two_mul_self_lt_three_mul n_large h h2x, Nat.pow_zero] diff --git a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean index e526d1422118e9..c4b31c98a929a3 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Discriminant.lean @@ -168,7 +168,7 @@ theorem discr_prime_pow [hcycl : IsCyclotomicExtension {p ^ k} K L] [hp : Fact ( convert_to (discr K fun i : Fin 1 ↦ (algebraMap K L) (-1) ^ ↑i) = _ · congr ext i - simp only [map_neg, map_one, Function.comp_apply, Fin.coe_fin_one, _root_.pow_zero] + simp only [map_neg, map_one, Function.comp_apply, Fin.val_eq_zero, _root_.pow_zero] suffices (e.symm i : ℕ) = 0 by simp [this] rw [← Nat.lt_one_iff] convert (e.symm i).2 diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index 0c98cefad6ecd4..43de39c1d54b0b 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -66,10 +66,14 @@ theorem filter_dvd_eq_properDivisors (h : n ≠ 0) : theorem properDivisors.not_self_mem : ¬n ∈ properDivisors n := by simp [properDivisors] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp] theorem mem_properDivisors {m : ℕ} : n ∈ properDivisors m ↔ n ∣ m ∧ n < m := by rcases eq_or_ne m 0 with (rfl | hm); · simp [properDivisors] - simp only [and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] + simp only [_root_.and_comm, ← filter_dvd_eq_properDivisors hm, mem_filter, mem_range] theorem insert_self_properDivisors (h : n ≠ 0) : insert n (properDivisors n) = divisors n := by rw [divisors, properDivisors, Ico_succ_right_eq_insert_Ico (one_le_iff_ne_zero.2 h), @@ -96,11 +100,15 @@ theorem dvd_of_mem_divisors {m : ℕ} (h : n ∈ divisors m) : n ∣ m := by · apply dvd_zero · simp [mem_divisors.1 h] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ @[simp] theorem mem_divisorsAntidiagonal {x : ℕ × ℕ} : x ∈ divisorsAntidiagonal n ↔ x.fst * x.snd = n ∧ n ≠ 0 := by simp only [divisorsAntidiagonal, Finset.mem_Ico, Ne, Finset.mem_filter, Finset.mem_product] - rw [and_comm] + rw [_root_.and_comm] apply and_congr_right rintro rfl constructor <;> intro h @@ -421,10 +429,14 @@ theorem sum_properDivisors_eq_one_iff_prime : ∑ x ∈ n.properDivisors, x = 1 (one_mem_properDivisors_iff_one_lt.2 (succ_lt_succ (Nat.succ_pos _)))) ((sum_singleton _ _).trans h.symm) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem mem_properDivisors_prime_pow {p : ℕ} (pp : p.Prime) (k : ℕ) {x : ℕ} : x ∈ properDivisors (p ^ k) ↔ ∃ (j : ℕ) (_ : j < k), x = p ^ j := by rw [mem_properDivisors, Nat.dvd_prime_pow pp, ← exists_and_right] - simp only [exists_prop, and_assoc] + simp only [exists_prop, _root_.and_assoc] apply exists_congr intro a constructor <;> intro h @@ -469,13 +481,17 @@ theorem prod_divisorsAntidiagonal' {M : Type*} [CommMonoid M] (f : ℕ → ℕ rw [← map_swap_divisorsAntidiagonal, Finset.prod_map] exact prod_divisorsAntidiagonal fun i j => f j i +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- The factors of `n` are the prime divisors -/ theorem primeFactors_eq_to_filter_divisors_prime (n : ℕ) : n.primeFactors = (divisors n).filter Prime := by rcases n.eq_zero_or_pos with (rfl | hn) · simp · ext q - simpa [hn, hn.ne', mem_primeFactorsList] using and_comm + simpa [hn, hn.ne', mem_primeFactorsList] using _root_.and_comm @[deprecated (since := "2024-07-17")] alias prime_divisors_eq_to_filter_divisors_prime := primeFactors_eq_to_filter_divisors_prime diff --git a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean index 392736e26084dd..b2cb30daa0d3ef 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/GaussEisensteinLemmas.lean @@ -180,6 +180,10 @@ private theorem sum_Ico_eq_card_lt {p q : ℕ} : (by simp (config := { contextual := true }) only [mem_filter, mem_sigma, and_self_iff, forall_true_iff, mem_product]) (fun _ _ ↦ rfl) (fun _ _ ↦ rfl) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Each of the sums in this lemma is the cardinality of the set of integer points in each of the two triangles formed by the diagonal of the rectangle `(0, p/2) × (0, q/2)`. Adding them gives the number of points in the rectangle. -/ @@ -193,7 +197,7 @@ theorem sum_mul_div_add_sum_mul_div_eq_mul (p q : ℕ) [hp : Fact p.Prime] (hq0 card_equiv (Equiv.prodComm _ _) (fun ⟨_, _⟩ => by simp (config := { contextual := true }) only [mem_filter, and_self_iff, Prod.swap_prod_mk, - forall_true_iff, mem_product, Equiv.prodComm_apply, and_assoc, and_left_comm]) + forall_true_iff, mem_product, Equiv.prodComm_apply, _root_.and_assoc, and_left_comm]) have hdisj : Disjoint ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.2 * p ≤ x.1 * q) ((Ico 1 (p / 2).succ ×ˢ Ico 1 (q / 2).succ).filter fun x : ℕ × ℕ => x.1 * q ≤ x.2 * p) := by diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index 74dce799b0ee14..c4b313daf7c59d 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -156,6 +156,10 @@ theorem mul_left (a₁ a₂ : ℤ) (b : ℕ) : J(a₁ * a₂ | b) = J(a₁ | b) (f := fun x ↦ @legendreSym x {out := prime_of_mem_primeFactorsList x.2} a₁) (g := fun x ↦ @legendreSym x {out := prime_of_mem_primeFactorsList x.2} a₂) +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefixes below. +-/ /-- The symbol `J(a | b)` vanishes iff `a` and `b` are not coprime (assuming `b ≠ 0`). -/ theorem eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [NeZero b] : J(a | b) = 0 ↔ a.gcd b ≠ 1 := List.prod_eq_zero_iff.trans @@ -165,7 +169,7 @@ theorem eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [NeZero b] : J(a | b) = 0 -- been deprecated so we replace them with `and_assoc` and `and_comm` simp_rw [legendreSym.eq_zero_iff _ _, intCast_zmod_eq_zero_iff_dvd, mem_primeFactorsList (NeZero.ne b), ← Int.natCast_dvd, Int.natCast_dvd_natCast, exists_prop, - and_assoc, and_comm]) + _root_.and_assoc, _root_.and_comm]) /-- The symbol `J(a | b)` is nonzero when `a` and `b` are coprime. -/ protected theorem ne_zero {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ≠ 0 := by @@ -214,7 +218,7 @@ theorem sq_one' {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a ^ 2 | b) = 1 := by r /-- The symbol `J(a | b)` depends only on `a` mod `b`. -/ theorem mod_left (a : ℤ) (b : ℕ) : J(a | b) = J(a % b | b) := congr_arg List.prod <| - List.pmap_congr _ + List.pmap_congr_left _ (by -- Porting note: Lean does not synthesize the instance [Fact (Nat.Prime p)] automatically -- (it is needed for `legendreSym.mod` on line 227). Thus, we name the hypothesis @@ -309,7 +313,7 @@ theorem value_at (a : ℤ) {R : Type*} [CommSemiring R] (χ : R →* ℤ) conv_rhs => rw [← prod_primeFactorsList hb.pos.ne', cast_list_prod, map_list_prod χ] rw [jacobiSym, List.map_map, ← List.pmap_eq_map Nat.Prime _ _ fun _ => prime_of_mem_primeFactorsList] - congr 1; apply List.pmap_congr + congr 1; apply List.pmap_congr_left exact fun p h pp _ => hp p pp (hb.ne_two_of_dvd_nat <| dvd_of_mem_primeFactorsList h) /-- If `b` is odd, then `J(-1 | b)` is given by `χ₄ b`. -/ diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 699d1221574f70..12b15e1ddd7d77 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -217,7 +217,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : fun w => ⟨(w x) / 2, div_nonneg (AbsoluteValue.nonneg _ _) (by norm_num)⟩ suffices ∀ w, w ≠ w₁ → f w ≠ 0 by obtain ⟨g, h_geqf, h_gprod⟩ := adjust_f K B this - obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ringOfIntegers_lt (f := g) + obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ringOfIntegers_lt K (f := g) (by rw [convexBodyLT_volume]; convert hB; exact congr_arg ((↑) : NNReal → ENNReal) h_gprod) refine ⟨y, h_ynz, fun w hw => (h_geqf w hw ▸ h_yle w).trans ?_, ?_⟩ · rw [← Rat.cast_le (K := ℝ), Rat.cast_natCast] diff --git a/Mathlib/NumberTheory/PythagoreanTriples.lean b/Mathlib/NumberTheory/PythagoreanTriples.lean index 8522d8eb5a00d0..57ca7ff1090821 100644 --- a/Mathlib/NumberTheory/PythagoreanTriples.lean +++ b/Mathlib/NumberTheory/PythagoreanTriples.lean @@ -179,7 +179,7 @@ theorem normalize : PythagoreanTriple (x / Int.gcd x y) (y / Int.gcd x y) (z / I have hz : z = 0 := by simpa only [PythagoreanTriple, hx, hy, add_zero, zero_eq_mul, mul_zero, or_self_iff] using h - simp only [hx, hy, hz, Int.zero_div] + simp only [hx, hy, hz] exact zero rcases h.gcd_dvd with ⟨z0, rfl⟩ obtain ⟨k, x0, y0, k0, h2, rfl, rfl⟩ : diff --git a/Mathlib/NumberTheory/SmoothNumbers.lean b/Mathlib/NumberTheory/SmoothNumbers.lean index 18f25f99002024..a130280fc27313 100644 --- a/Mathlib/NumberTheory/SmoothNumbers.lean +++ b/Mathlib/NumberTheory/SmoothNumbers.lean @@ -233,7 +233,7 @@ def equivProdNatFactoredNumbers {s : Finset ℕ} {p : ℕ} (hp : p.Prime) (hs : refine prod_eq <| (filter _ <| perm_primeFactorsList_mul (pow_ne_zero e hp.ne_zero) hm₀).trans ?_ rw [filter_append, hp.primeFactorsList_pow, - filter_eq_nil.mpr fun q hq ↦ by rw [mem_replicate] at hq; simp [hq.2, hs], + filter_eq_nil_iff.mpr fun q hq ↦ by rw [mem_replicate] at hq; simp [hq.2, hs], nil_append, filter_eq_self.mpr fun q hq ↦ by simp only [hm q hq, decide_True]] right_inv := by rintro ⟨m, hm₀, hm⟩ diff --git a/Mathlib/NumberTheory/SumFourSquares.lean b/Mathlib/NumberTheory/SumFourSquares.lean index 5d7a765adee868..a08b8c479c8ee0 100644 --- a/Mathlib/NumberTheory/SumFourSquares.lean +++ b/Mathlib/NumberTheory/SumFourSquares.lean @@ -121,6 +121,10 @@ private theorem sum_four_squares_of_two_mul_sum_four_squares {m a b c d : ℤ} have : (∑ x, f (σ x) ^ 2) = ∑ x, f x ^ 2 := Equiv.sum_comp σ (f · ^ 2) simpa only [← hx, ← hy, Fin.sum_univ_four, add_assoc] using this +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- Lagrange's **four squares theorem** for a prime number. Use `Nat.sum_four_squares` instead. -/ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : ∃ a b c d : ℕ, a ^ 2 + b ^ 2 + c ^ 2 + d ^ 2 = p := by @@ -170,9 +174,10 @@ protected theorem Prime.sum_four_squares {p : ℕ} (hp : p.Prime) : -- The quotient `r` is not zero, because otherwise `f a = f b = f c = f d = 0`, hence -- `m` divides each `a`, `b`, `c`, `d`, thus `m ∣ p` which is impossible. rcases (zero_le r).eq_or_gt with rfl | hr₀ - · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [and_assoc] using hr + · replace hr : f a = 0 ∧ f b = 0 ∧ f c = 0 ∧ f d = 0 := by simpa [_root_.and_assoc] using hr obtain ⟨⟨a, rfl⟩, ⟨b, rfl⟩, ⟨c, rfl⟩, ⟨d, rfl⟩⟩ : m ∣ a ∧ m ∣ b ∧ m ∣ c ∧ m ∣ d := by - simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, and_self] + simp only [← ZMod.natCast_zmod_eq_zero_iff_dvd, ← hf_mod, hr, Int.cast_zero, + _root_.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/Order/Directed.lean b/Mathlib/Order/Directed.lean index 39d27f9b220b12..29285fca18124f 100644 --- a/Mathlib/Order/Directed.lean +++ b/Mathlib/Order/Directed.lean @@ -167,7 +167,7 @@ instance OrderDual.isDirected_le [LE α] [IsDirected α (· ≥ ·)] : IsDirecte /-- A monotone function on an upwards-directed type is directed. -/ theorem directed_of_isDirected_le [LE α] [IsDirected α (· ≤ ·)] {f : α → β} {r : β → β → Prop} (H : ∀ ⦃i j⦄, i ≤ j → r (f i) (f j)) : Directed r f := - directed_id.mono_comp H + directed_id.mono_comp _ H theorem Monotone.directed_le [Preorder α] [IsDirected α (· ≤ ·)] [Preorder β] {f : α → β} : Monotone f → Directed (· ≤ ·) f := diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index dfe2683487ad83..66d78953d65388 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -668,7 +668,7 @@ theorem HasBasis.eq_iInf (h : l.HasBasis (fun _ => True) s) : l = ⨅ i, 𝓟 (s theorem hasBasis_iInf_principal {s : ι → Set α} (h : Directed (· ≥ ·) s) [Nonempty ι] : (⨅ i, 𝓟 (s i)).HasBasis (fun _ => True) s := ⟨fun t => by - simpa only [true_and] using mem_iInf_of_directed (h.mono_comp monotone_principal.dual) t⟩ + simpa only [true_and] using mem_iInf_of_directed (h.mono_comp _ monotone_principal.dual) t⟩ /-- If `s : ι → Set α` is an indexed family of sets, then finite intersections of `s i` form a basis of `⨅ i, 𝓟 (s i)`. -/ @@ -683,7 +683,7 @@ theorem hasBasis_biInf_principal {s : β → Set α} {S : Set β} (h : DirectedO ⟨fun t => by refine mem_biInf_of_directed ?_ ne rw [directedOn_iff_directed, ← directed_comp] at h ⊢ - refine h.mono_comp ?_ + refine h.mono_comp _ ?_ exact fun _ _ => principal_mono.2⟩ theorem hasBasis_biInf_principal' {ι : Type*} {p : ι → Prop} {s : ι → Set α} diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 5f4bd78be7ae19..f77c690521b891 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -111,7 +111,8 @@ abbrev OrderHomClass (F : Type*) (α β : outParam Type*) [LE α] [LE β] [FunLi /-- `OrderIsoClass F α β` states that `F` is a type of order isomorphisms. You should extend this class when you extend `OrderIso`. -/ -class OrderIsoClass (F α β : Type*) [LE α] [LE β] [EquivLike F α β] : Prop where +class OrderIsoClass (F : Type*) (α β : outParam Type*) [LE α] [LE β] [EquivLike F α β] : + Prop where /-- An order isomorphism respects `≤`. -/ map_le_map_iff (f : F) {a b : α} : f a ≤ f b ↔ a ≤ b diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 971df11a1b4ff7..92fa0698730b3c 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.lean @@ -59,14 +59,16 @@ section /-- `TopHomClass F α β` states that `F` is a type of `⊤`-preserving morphisms. You should extend this class when you extend `TopHom`. -/ -class TopHomClass (F α β : Type*) [Top α] [Top β] [FunLike F α β] : Prop where +class TopHomClass (F : Type*) (α β : outParam Type*) [Top α] [Top β] [FunLike F α β] : + Prop where /-- A `TopHomClass` morphism preserves the top element. -/ map_top (f : F) : f ⊤ = ⊤ /-- `BotHomClass F α β` states that `F` is a type of `⊥`-preserving morphisms. You should extend this class when you extend `BotHom`. -/ -class BotHomClass (F α β : Type*) [Bot α] [Bot β] [FunLike F α β] : Prop where +class BotHomClass (F : Type*) (α β : outParam Type*) [Bot α] [Bot β] [FunLike F α β] : + Prop where /-- A `BotHomClass` morphism preserves the bottom element. -/ map_bot (f : F) : f ⊥ = ⊥ diff --git a/Mathlib/Order/Interval/Finset/Nat.lean b/Mathlib/Order/Interval/Finset/Nat.lean index 133d9233ac4661..d59c44eaa21395 100644 --- a/Mathlib/Order/Interval/Finset/Nat.lean +++ b/Mathlib/Order/Interval/Finset/Nat.lean @@ -185,10 +185,14 @@ theorem Ico_image_const_sub_eq_Ico (hac : a ≤ c) : rintro ⟨x, hx, rfl⟩ omega +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem Ico_succ_left_eq_erase_Ico : Ico a.succ b = erase (Ico a b) a := by ext x - rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← and_assoc, ne_comm, @and_comm (a ≠ x), - lt_iff_le_and_ne] + rw [Ico_succ_left, mem_erase, mem_Ico, mem_Ioo, ← _root_.and_assoc, ne_comm, + _root_.and_comm (a := a ≠ x), lt_iff_le_and_ne] theorem mod_injOn_Ico (n a : ℕ) : Set.InjOn (· % a) (Finset.Ico n (n + a)) := by induction' n with n ih diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index 982bf1b0e61765..e801c4d599937c 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -58,7 +58,7 @@ satisfy `r a b → s (f a) (f b)`. The relations `r` and `s` are `outParam`s since figuring them out from a goal is a higher-order matching problem that Lean usually can't do unaided. -/ -class RelHomClass (F : Type*) {α β : Type*} (r : outParam <| α → α → Prop) +class RelHomClass (F : Type*) {α β : outParam Type*} (r : outParam <| α → α → Prop) (s : outParam <| β → β → Prop) [FunLike F α β] : Prop where /-- A `RelHomClass` sends related elements to related elements -/ map_rel : ∀ (f : F) {a b}, r a b → s (f a) (f b) diff --git a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean index b2e2b120664d71..f721b4e7769795 100644 --- a/Mathlib/Order/SuccPred/LinearLocallyFinite.lean +++ b/Mathlib/Order/SuccPred/LinearLocallyFinite.lean @@ -185,7 +185,7 @@ end LinearLocallyFiniteOrder section toZ -- Requiring either of `IsSuccArchimedean` or `IsPredArchimedean` is equivalent. -variable [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} +variable [SuccOrder ι] [IsSuccArchimedean ι] [P : PredOrder ι] {i0 i : ι} -- For "to_Z" @@ -194,12 +194,13 @@ variable [SuccOrder ι] [IsSuccArchimedean ι] [PredOrder ι] {i0 i : ι} the range of `toZ`. -/ def toZ (i0 i : ι) : ℤ := dite (i0 ≤ i) (fun hi ↦ Nat.find (exists_succ_iterate_of_le hi)) fun hi ↦ - -Nat.find (exists_pred_iterate_of_le (not_le.mp hi).le) + -Nat.find (exists_pred_iterate_of_le (α := ι) (not_le.mp hi).le) theorem toZ_of_ge (hi : i0 ≤ i) : toZ i0 i = Nat.find (exists_succ_iterate_of_le hi) := dif_pos hi -theorem toZ_of_lt (hi : i < i0) : toZ i0 i = -Nat.find (exists_pred_iterate_of_le hi.le) := +theorem toZ_of_lt (hi : i < i0) : + toZ i0 i = -Nat.find (@exists_pred_iterate_of_le _ _ P _ _ _ hi.le) := dif_neg (not_le.mpr hi) @[simp] @@ -310,7 +311,7 @@ theorem toZ_mono {i j : ι} (h_le : i ≤ j) : toZ i0 i ≤ toZ i0 j := by · exact le_of_not_le h · exact absurd h_le (not_le.mpr (hj.trans_le hi)) · exact (toZ_neg hi).le.trans (toZ_nonneg hj) - · let m := Nat.find (exists_pred_iterate_of_le h_le) + · let m := Nat.find (@exists_pred_iterate_of_le _ _ P _ _ _ h_le) have hm : pred^[m] j = i := Nat.find_spec (exists_pred_iterate_of_le h_le) have hj_eq : i = pred^[(-toZ i0 j).toNat + m] i0 := by rw [← hm, add_comm] diff --git a/Mathlib/Order/SymmDiff.lean b/Mathlib/Order/SymmDiff.lean index 8805a9c3184ad8..fa8d23d3aa0818 100644 --- a/Mathlib/Order/SymmDiff.lean +++ b/Mathlib/Order/SymmDiff.lean @@ -82,7 +82,7 @@ theorem symmDiff_eq_Xor' (p q : Prop) : p ∆ q = Xor' p q := @[simp] theorem bihimp_iff_iff {p q : Prop} : p ⇔ q ↔ (p ↔ q) := - (iff_iff_implies_and_implies _ _).symm.trans Iff.comm + iff_iff_implies_and_implies.symm.trans Iff.comm @[simp] theorem Bool.symmDiff_eq_xor : ∀ p q : Bool, p ∆ q = xor p q := by decide diff --git a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean index ec56f812e2f47c..877ecc41e7f9db 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Basic.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Basic.lean @@ -176,7 +176,8 @@ theorem toOuterMeasure_apply_eq_one_iff : p.toOuterMeasure s = 1 ↔ p.support (fun x => Set.indicator_apply_le fun _ => le_rfl) hsa · suffices ∀ (x) (_ : x ∉ s), p x = 0 from _root_.trans (tsum_congr - fun a => (Set.indicator_apply s p a).trans (ite_eq_left_iff.2 <| symm ∘ this a)) p.tsum_coe + fun a => (Set.indicator_apply s p a).trans + (ite_eq_left_iff.2 <| symm ∘ this a)) p.tsum_coe exact fun a ha => (p.apply_eq_zero_iff a).2 <| Set.not_mem_subset h ha @[simp] diff --git a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean index d137f3c66d02f5..31a9cbfa903d10 100644 --- a/Mathlib/Probability/ProbabilityMassFunction/Monad.lean +++ b/Mathlib/Probability/ProbabilityMassFunction/Monad.lean @@ -67,9 +67,11 @@ theorem toOuterMeasure_pure_apply : (pure a).toOuterMeasure s = if a ∈ s then refine (toOuterMeasure_apply (pure a) s).trans ?_ split_ifs with ha · refine (tsum_congr fun b => ?_).trans (tsum_ite_eq a 1) - exact ite_eq_left_iff.2 fun hb => symm (ite_eq_right_iff.2 fun h => (hb <| h.symm ▸ ha).elim) + exact ite_eq_left_iff.2 fun hb => + symm (ite_eq_right_iff.2 fun h => (hb <| h.symm ▸ ha).elim) · refine (tsum_congr fun b => ?_).trans tsum_zero - exact ite_eq_right_iff.2 fun hb => ite_eq_right_iff.2 fun h => (ha <| h ▸ hb).elim + exact ite_eq_right_iff.2 fun hb => + ite_eq_right_iff.2 fun h => (ha <| h ▸ hb).elim variable [MeasurableSpace α] diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index fbe67ec9e34384..2b83e2348e8b18 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -146,7 +146,7 @@ theorem dZero_comp_eq : dZero A ∘ₗ (zeroCochainsLequiv A) = oneCochainsLequiv A ∘ₗ (inhomogeneousCochains A).d 0 1 := by ext x y show A.ρ y (x default) - x default = _ + ({0} : Finset _).sum _ - simp_rw [Fin.coe_fin_one, zero_add, pow_one, neg_smul, one_smul, + simp_rw [Fin.val_eq_zero, zero_add, pow_one, neg_smul, one_smul, Finset.sum_singleton, sub_eq_add_neg] rcongr i <;> exact Fin.elim0 i diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index 7532a1f66cbdf1..bba7f1e3e74579 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -353,11 +353,11 @@ noncomputable instance instField [Fact (Irreducible f)] : Field (AdjoinRoot f) w ratCast_def q := by rw [← map_natCast (of f), ← map_intCast (of f), ← map_div₀, ← Rat.cast_def]; rfl nnqsmul_def q x := - AdjoinRoot.induction_on (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by + AdjoinRoot.induction_on f (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by simp only [smul_mk, of, RingHom.comp_apply, ← (mk f).map_mul, Polynomial.nnqsmul_eq_C_mul] qsmul_def q x := -- Porting note: I gave the explicit motive and changed `rw` to `simp`. - AdjoinRoot.induction_on (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by + AdjoinRoot.induction_on f (C := fun y ↦ q • y = (of f) q * y) x fun p ↦ by simp only [smul_mk, of, RingHom.comp_apply, ← (mk f).map_mul, Polynomial.qsmul_eq_C_mul] theorem coe_injective (h : degree f ≠ 0) : Function.Injective ((↑) : K → AdjoinRoot f) := diff --git a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean index 0a9301168ce1c8..146fc3d8f12a98 100644 --- a/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean +++ b/Mathlib/RingTheory/IntegralClosure/IsIntegralClosure/Basic.lean @@ -395,14 +395,14 @@ variable [Algebra R A] [Algebra R A'] [IsScalarTower R A B] [IsScalarTower R A' /-- Integral closures are all isomorphic to each other. -/ noncomputable def equiv : A ≃ₐ[R] A' := AlgEquiv.ofAlgHom - (lift _ B (isIntegral := isIntegral_algebra R B)) - (lift _ B (isIntegral := isIntegral_algebra R B)) + (lift R A' B (isIntegral := isIntegral_algebra R B)) + (lift R A B (isIntegral := isIntegral_algebra R B)) (by ext x; apply algebraMap_injective A' R B; simp) (by ext x; apply algebraMap_injective A R B; simp) @[simp] theorem algebraMap_equiv (x : A) : algebraMap A' B (equiv R A B A' x) = algebraMap A B x := - algebraMap_lift A' B (isIntegral := isIntegral_algebra R B) x + algebraMap_lift R A' B (isIntegral := isIntegral_algebra R B) x end Equiv diff --git a/Mathlib/RingTheory/LocalRing/Module.lean b/Mathlib/RingTheory/LocalRing/Module.lean index 838c66b8f38638..332bccee7c8f59 100644 --- a/Mathlib/RingTheory/LocalRing/Module.lean +++ b/Mathlib/RingTheory/LocalRing/Module.lean @@ -178,19 +178,21 @@ theorem free_of_maximalIdeal_rTensor_injective [Module.FinitePresentation R M] Module.finrank_finsupp_self, one_mul, add_zero] rw [Module.finrank_eq_card_chooseBasisIndex] -- On the other hand, `m ⊗ M → M` injective => `Tor₁(k, M) = 0` => `k ⊗ ker(i) → kᴵ` injective. - have := @lTensor_injective_of_exact_of_exact_of_rTensor_injective - (N₁ := LinearMap.ker i) (N₂ := I →₀ R) (N₃ := M) - (f₁ := (𝔪).subtype) (f₂ := Submodule.mkQ 𝔪) inferInstance inferInstance inferInstance - inferInstance inferInstance inferInstance intro x - apply @this (LinearMap.ker i).subtype i (LinearMap.exact_subtype_mkQ 𝔪) - (Submodule.mkQ_surjective _) (LinearMap.exact_subtype_ker_map i) hi H - (Module.Flat.lTensor_preserves_injective_linearMap _ Subtype.val_injective) - apply hi'.injective - rw [LinearMap.baseChange_eq_ltensor] - erw [← LinearMap.comp_apply (i.lTensor k), ← LinearMap.lTensor_comp] - rw [(LinearMap.exact_subtype_ker_map i).linearMap_comp_eq_zero] - simp only [LinearMap.lTensor_zero, LinearMap.zero_apply, map_zero] + refine lTensor_injective_of_exact_of_exact_of_rTensor_injective + (N₁ := LinearMap.ker i) (N₂ := I →₀ R) (N₃ := M) + (f₁ := (𝔪).subtype) (f₂ := Submodule.mkQ 𝔪) + (g₁ := (LinearMap.ker i).subtype) (g₂ := i) (LinearMap.exact_subtype_mkQ 𝔪) + (Submodule.mkQ_surjective _) (LinearMap.exact_subtype_ker_map i) hi H ?_ ?_ + · apply Module.Flat.lTensor_preserves_injective_linearMap + (N := LinearMap.ker i) (N' := I →₀ R) + (L := (LinearMap.ker i).subtype) + exact Subtype.val_injective + · apply hi'.injective + rw [LinearMap.baseChange_eq_ltensor] + erw [← LinearMap.comp_apply (i.lTensor k), ← LinearMap.lTensor_comp] + rw [(LinearMap.exact_subtype_ker_map i).linearMap_comp_eq_zero] + simp only [LinearMap.lTensor_zero, LinearMap.zero_apply, map_zero] -- TODO: Generalise this to finite free modules. theorem free_of_flat_of_localRing [Module.FinitePresentation R P] [Module.Flat R P] : diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index 1960bbcd7e01d1..f4b231a725789d 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -132,7 +132,7 @@ noncomputable abbrev toField : Field K where lemma surjective_iff_isField [IsDomain R] : Function.Surjective (algebraMap R K) ↔ IsField R where mp h := (RingEquiv.ofBijective (algebraMap R K) - ⟨IsFractionRing.injective R K, h⟩).toMulEquiv.isField (IsFractionRing.toField R).toIsField + ⟨IsFractionRing.injective R K, h⟩).toMulEquiv.isField _ (IsFractionRing.toField R).toIsField mpr h := letI := h.toField (IsLocalization.atUnits R _ (S := K) diff --git a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean index 3aa09ead7b4abe..ddadb79adcf83b 100644 --- a/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean +++ b/Mathlib/RingTheory/MvPolynomial/Symmetric/NewtonIdentities.lean @@ -250,13 +250,17 @@ theorem sum_antidiagonal_card_esymm_psum_eq_zero : simp [← sum_filter_add_sum_filter_not (antidiagonal k) (fun a ↦ a.fst < k), ← mul_esymm_eq_sum, mul_add, ← mul_assoc, ← pow_add, mul_comm ↑k (esymm σ R k)] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ /-- A version of Newton's identities which may be more useful in the case that we know the values of the elementary symmetric polynomials and would like to calculate the values of the power sums. -/ theorem psum_eq_mul_esymm_sub_sum (k : ℕ) (h : 0 < k) : psum σ R k = (-1) ^ (k + 1) * k * esymm σ R k - ∑ a ∈ (antidiagonal k).filter (fun a ↦ a.fst ∈ Set.Ioo 0 k), (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd := by - simp only [Set.Ioo, Set.mem_setOf_eq, and_comm] + simp only [Set.Ioo, Set.mem_setOf_eq, _root_.and_comm] have hesymm := mul_esymm_eq_sum σ R k rw [← (sum_filter_add_sum_filter_not ((antidiagonal k).filter (fun a ↦ a.fst < k)) (fun a ↦ 0 < a.fst) (fun a ↦ (-1) ^ a.fst * esymm σ R a.fst * psum σ R a.snd))] at hesymm diff --git a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean index 55bb7d2fd0d446..891e01b70140c6 100644 --- a/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean +++ b/Mathlib/RingTheory/NonUnitalSubsemiring/Basic.lean @@ -29,7 +29,7 @@ variable {R : Type u} {S : Type v} {T : Type w} [NonUnitalNonAssocSemiring R] (M /-- `NonUnitalSubsemiringClass S R` states that `S` is a type of subsets `s ⊆ R` that are both an additive submonoid and also a multiplicative subsemigroup. -/ -class NonUnitalSubsemiringClass (S : Type*) (R : Type u) [NonUnitalNonAssocSemiring R] +class NonUnitalSubsemiringClass (S : Type*) (R : outParam (Type u)) [NonUnitalNonAssocSemiring R] [SetLike S R] extends AddSubmonoidClass S R : Prop where mul_mem : ∀ {s : S} {a b : R}, a ∈ s → b ∈ s → a * b ∈ s diff --git a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean index 4a1ad15664ee85..43e050ec7489de 100644 --- a/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean +++ b/Mathlib/RingTheory/Polynomial/Cyclotomic/Eval.lean @@ -104,12 +104,17 @@ theorem cyclotomic_pos {n : ℕ} (hn : 2 < n) {R} [LinearOrderedCommRing R] (x : exact hn'.ne' hi.2.2.1 · simpa only [eval_X, eval_one, cyclotomic_two, eval_add] using h.right.le +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefix below. +-/ theorem cyclotomic_pos_and_nonneg (n : ℕ) {R} [LinearOrderedCommRing R] (x : R) : (1 < x → 0 < eval x (cyclotomic n R)) ∧ (1 ≤ x → 0 ≤ eval x (cyclotomic n R)) := by rcases n with (_ | _ | _ | n) - · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, and_self] + · simp only [cyclotomic_zero, eval_one, zero_lt_one, implies_true, zero_le_one, + _root_.and_self] · simp only [zero_add, cyclotomic_one, eval_sub, eval_X, eval_one, sub_pos, imp_self, sub_nonneg, - and_self] + _root_.and_self] · simp only [zero_add, reduceAdd, cyclotomic_two, eval_add, eval_X, eval_one] constructor <;> intro <;> linarith · constructor <;> intro <;> [skip; apply le_of_lt] <;> apply cyclotomic_pos (by omega) diff --git a/Mathlib/RingTheory/RingHomProperties.lean b/Mathlib/RingTheory/RingHomProperties.lean index 1b12d138c40924..7008cd0ef9a746 100644 --- a/Mathlib/RingTheory/RingHomProperties.lean +++ b/Mathlib/RingTheory/RingHomProperties.lean @@ -185,9 +185,11 @@ lemma toMorphismProperty_respectsIso_iff : · intro X Y Z e f hf exact h.left f e.commRingCatIsoToRingEquiv hf · intro X Y Z _ _ _ f e hf - exact h.postcomp e.toCommRingCatIso.hom (CommRingCat.ofHom f) hf + exact MorphismProperty.RespectsIso.postcomp (toMorphismProperty P) + e.toCommRingCatIso.hom (CommRingCat.ofHom f) hf · intro X Y Z _ _ _ f e - exact h.precomp e.toCommRingCatIso.hom (CommRingCat.ofHom f) + exact MorphismProperty.RespectsIso.precomp (toMorphismProperty P) + e.toCommRingCatIso.hom (CommRingCat.ofHom f) end ToMorphismProperty diff --git a/Mathlib/RingTheory/SimpleRing/Basic.lean b/Mathlib/RingTheory/SimpleRing/Basic.lean index c9f5a875cc0b81..3f9711cafff025 100644 --- a/Mathlib/RingTheory/SimpleRing/Basic.lean +++ b/Mathlib/RingTheory/SimpleRing/Basic.lean @@ -82,6 +82,6 @@ lemma isField_center (A : Type*) [Ring A] [IsSimpleRing A] : IsField (Subring.ce end IsSimpleRing lemma isSimpleRing_iff_isField (A : Type*) [CommRing A] : IsSimpleRing A ↔ IsField A := - ⟨fun _ ↦ Subring.topEquiv.symm.toMulEquiv.isField <| by + ⟨fun _ ↦ Subring.topEquiv.symm.toMulEquiv.isField _ <| by rw [← Subring.center_eq_top A]; exact IsSimpleRing.isField_center A, fun h ↦ letI := h.toField; inferInstance⟩ diff --git a/Mathlib/RingTheory/Smooth/StandardSmooth.lean b/Mathlib/RingTheory/Smooth/StandardSmooth.lean index f9ed2c10906263..988a61b9c879e4 100644 --- a/Mathlib/RingTheory/Smooth/StandardSmooth.lean +++ b/Mathlib/RingTheory/Smooth/StandardSmooth.lean @@ -544,14 +544,14 @@ instance IsStandardSmooth.baseChange [IsStandardSmooth.{t, w} R S] : IsStandardSmooth.{t, w} T (T ⊗[R] S) where out := by obtain ⟨⟨P⟩⟩ := ‹IsStandardSmooth R S› - exact ⟨P.baseChange T⟩ + exact ⟨P.baseChange R S T⟩ instance IsStandardSmoothOfRelativeDimension.baseChange [IsStandardSmoothOfRelativeDimension.{t, w} n R S] : IsStandardSmoothOfRelativeDimension.{t, w} n T (T ⊗[R] S) where out := by obtain ⟨P, hP⟩ := ‹IsStandardSmoothOfRelativeDimension n R S› - exact ⟨P.baseChange T, hP⟩ + exact ⟨P.baseChange R S T, hP⟩ end BaseChange diff --git a/Mathlib/SetTheory/Cardinal/Finite.lean b/Mathlib/SetTheory/Cardinal/Finite.lean index 1cd7eb9c7c18a7..08b850cedaa5a2 100644 --- a/Mathlib/SetTheory/Cardinal/Finite.lean +++ b/Mathlib/SetTheory/Cardinal/Finite.lean @@ -95,15 +95,19 @@ protected theorem bijective_iff_injective_and_card [Finite β] (f : α → β) : rw [← and_congr_right_iff, ← Bijective, card_eq_fintype_card, card_eq_fintype_card, Fintype.bijective_iff_injective_and_card] +#adaptation_note +/-- +After nightly-2024-09-06 we can remove the `_root_` prefixes below. +-/ protected theorem bijective_iff_surjective_and_card [Finite α] (f : α → β) : Bijective f ↔ Surjective f ∧ Nat.card α = Nat.card β := by classical - rw [and_comm, Bijective, and_congr_left_iff] + rw [_root_.and_comm, Bijective, and_congr_left_iff] intro h have := Fintype.ofFinite α have := Fintype.ofSurjective f h revert h - rw [← and_congr_left_iff, ← Bijective, ← and_comm, + rw [← and_congr_left_iff, ← Bijective, ← _root_.and_comm, card_eq_fintype_card, card_eq_fintype_card, Fintype.bijective_iff_surjective_and_card] theorem _root_.Function.Injective.bijective_of_nat_card_le [Finite β] {f : α → β} diff --git a/Mathlib/Tactic/CC/Addition.lean b/Mathlib/Tactic/CC/Addition.lean index 39157d0e58cab3..51dcb46a44da08 100644 --- a/Mathlib/Tactic/CC/Addition.lean +++ b/Mathlib/Tactic/CC/Addition.lean @@ -842,7 +842,7 @@ def dbgTraceACState : CCM Unit := do def mkACProof (e₁ e₂ : Expr) : MetaM Expr := do let eq ← mkEq e₁ e₂ let .mvar m ← mkFreshExprSyntheticOpaqueMVar eq | failure - AC.rewriteUnnormalized m + AC.rewriteUnnormalizedRefl m let pr ← instantiateMVars (.mvar m) mkExpectedTypeHint pr eq @@ -1470,7 +1470,8 @@ partial def propagateEqUp (e : Expr) : CCM Unit := do if ← isInterpretedValue ra <&&> isInterpretedValue rb <&&> pure (ra.int?.isNone || ra.int? != rb.int?) then raNeRb := some - (Expr.app (.proj ``Iff 0 (← mkAppM ``bne_iff_ne #[ra, rb])) (← mkEqRefl (.const ``true []))) + (Expr.app (.proj ``Iff 0 (← mkAppOptM ``bne_iff_ne #[none, none, none, ra, rb])) + (← mkEqRefl (.const ``true []))) else if let some c₁ ← isConstructorApp? ra then if let some c₂ ← isConstructorApp? rb then @@ -1808,7 +1809,8 @@ def propagateValueInconsistency (e₁ e₂ : Expr) : CCM Unit := do let some eqProof ← getEqProof e₁ e₂ | failure let trueEqFalse ← mkEq (.const ``True []) (.const ``False []) let neProof := - Expr.app (.proj ``Iff 0 (← mkAppM ``bne_iff_ne #[e₁, e₂])) (← mkEqRefl (.const ``true [])) + Expr.app (.proj ``Iff 0 (← mkAppOptM ``bne_iff_ne #[none, none, none, e₁, e₂])) + (← mkEqRefl (.const ``true [])) let H ← mkAbsurd trueEqFalse eqProof neProof pushEq (.const ``True []) (.const ``False []) H diff --git a/Mathlib/Tactic/IntervalCases.lean b/Mathlib/Tactic/IntervalCases.lean index 52169b6406e4cf..37e90dfc742f4b 100644 --- a/Mathlib/Tactic/IntervalCases.lean +++ b/Mathlib/Tactic/IntervalCases.lean @@ -5,6 +5,7 @@ Authors: Kim Morrison, Mario Carneiro -/ import Mathlib.Tactic.NormNum import Mathlib.Tactic.FinCases +import Mathlib.Control.Basic /-! # Case bash on variables in finite intervals diff --git a/Mathlib/Tactic/NormNum/DivMod.lean b/Mathlib/Tactic/NormNum/DivMod.lean index 3037be0ca7da78..821532050c5495 100644 --- a/Mathlib/Tactic/NormNum/DivMod.lean +++ b/Mathlib/Tactic/NormNum/DivMod.lean @@ -147,8 +147,8 @@ theorem isInt_dvd_true : {a b : ℤ} → {a' b' c : ℤ} → | _, _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, rfl => ⟨_, rfl⟩ theorem isInt_dvd_false : {a b : ℤ} → {a' b' : ℤ} → - IsInt a a' → IsInt b b' → Int.mod b' a' != 0 → ¬a ∣ b - | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, e => mt Int.mod_eq_zero_of_dvd (by simpa using e) + IsInt a a' → IsInt b b' → Int.emod b' a' != 0 → ¬a ∣ b + | _, _, _, _, ⟨rfl⟩, ⟨rfl⟩, e => mt Int.emod_eq_zero_of_dvd (by simpa using e) /-- The `norm_num` extension which identifies expressions of the form `(a : ℤ) ∣ b`, such that `norm_num` successfully recognises both `a` and `b`. -/ @@ -167,7 +167,7 @@ such that `norm_num` successfully recognises both `a` and `b`. -/ haveI' : Int.mul $na $c =Q $nb := ⟨⟩ return .isTrue q(isInt_dvd_true $pa $pb (.refl $nb)) else - have : Q(Int.mod $nb $na != 0) := (q(Eq.refl true) : Expr) + have : Q(Int.emod $nb $na != 0) := (q(Eq.refl true) : Expr) return .isFalse q(isInt_dvd_false $pa $pb $this) end Mathlib.Meta.NormNum diff --git a/Mathlib/Tactic/Ring/Basic.lean b/Mathlib/Tactic/Ring/Basic.lean index 498f20124180d4..dcd4f4ddf5f6f7 100644 --- a/Mathlib/Tactic/Ring/Basic.lean +++ b/Mathlib/Tactic/Ring/Basic.lean @@ -1014,7 +1014,7 @@ def ExSum.evalInv {a : Q($α)} (czα : Option Q(CharZero $α)) (va : ExSum sα a match va with | ExSum.zero => pure ⟨_, .zero, (q(inv_zero (R := $α)) : Expr)⟩ | ExSum.add va ExSum.zero => do - let ⟨_, vb, pb⟩ ← va.evalInv dα czα + let ⟨_, vb, pb⟩ ← va.evalInv sα dα czα pure ⟨_, vb.toSum, (q(inv_single $pb) : Expr)⟩ | va => do let ⟨_, vb, pb⟩ ← evalInvAtom sα dα a diff --git a/Mathlib/Tactic/WLOG.lean b/Mathlib/Tactic/WLOG.lean index c996e5552b6dfd..461949d04ab7b6 100644 --- a/Mathlib/Tactic/WLOG.lean +++ b/Mathlib/Tactic/WLOG.lean @@ -86,7 +86,8 @@ def _root_.Lean.MVarId.wlog (goal : MVarId) (h : Option Name) (P : Expr) let hGoal := HExpr.mvarId! /- Begin the "reduction goal" which will contain hypotheses `H` and `¬h`. For now, it only contains `H`. Keep track of that hypothesis' FVarId. -/ - let (HFVarId, reductionGoal) ← goal.assertHypotheses #[⟨H, HType, HExpr⟩] + let (HFVarId, reductionGoal) ← + goal.assertHypotheses #[{ userName := H, type := HType, value := HExpr }] let HFVarId := HFVarId[0]! /- Clear the reverted fvars from the branch that will contain `h` as a hypothesis. -/ let hGoal ← hGoal.tryClearMany revertedFVars diff --git a/Mathlib/Testing/SlimCheck/Functions.lean b/Mathlib/Testing/SlimCheck/Functions.lean index 12325734b9e555..55ef7dc3e20fae 100644 --- a/Mathlib/Testing/SlimCheck/Functions.lean +++ b/Mathlib/Testing/SlimCheck/Functions.lean @@ -363,7 +363,7 @@ theorem applyId_injective [DecidableEq α] {xs ys : List α} (h₀ : List.Nodup · symm; rw [h] rw [← List.applyId_zip_eq] <;> assumption · rw [← h₁.length_eq] - rw [List.getElem?_eq_some] at hx + rw [List.getElem?_eq_some_iff] at hx cases' hx with hx hx' exact hx · rw [← applyId_mem_iff h₀ h₁] at hx hy diff --git a/Mathlib/Topology/Algebra/Valued/NormedValued.lean b/Mathlib/Topology/Algebra/Valued/NormedValued.lean index dbe91b1484df49..2c9ce7c08f2c8d 100644 --- a/Mathlib/Topology/Algebra/Valued/NormedValued.lean +++ b/Mathlib/Topology/Algebra/Valued/NormedValued.lean @@ -102,7 +102,7 @@ def toNormedField : NormedField L := · set δ : ℝ≥0 := hv.hom ε with hδ have hδ_pos : 0 < δ := by rw [hδ, ← _root_.map_zero hv.hom] - exact hv.strictMono (Units.zero_lt ε) + exact hv.strictMono _ (Units.zero_lt ε) use δ, hδ_pos apply subset_trans _ hε intro x hx diff --git a/Mathlib/Topology/Category/LightProfinite/Extend.lean b/Mathlib/Topology/Category/LightProfinite/Extend.lean index bed0658153636f..4f7e5eb315b90b 100644 --- a/Mathlib/Topology/Category/LightProfinite/Extend.lean +++ b/Mathlib/Topology/Category/LightProfinite/Extend.lean @@ -118,7 +118,7 @@ then `cone G c.pt` is a limit cone. -/ noncomputable def isLimitCone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsLimit <| G.mapCone c) : - IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ hc' + IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ _ hc' end Limit @@ -158,7 +158,7 @@ noncomputable def isColimitCocone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsColimit <| G.mapCocone c.op) : IsColimit (cocone G c.pt) := haveI := functorOp_final c hc - (Functor.final_comp (opOpEquivalence ℕ).functor (functorOp c)).isColimitWhiskerEquiv _ hc' + (Functor.final_comp (opOpEquivalence ℕ).functor (functorOp c)).isColimitWhiskerEquiv _ _ hc' end Colimit diff --git a/Mathlib/Topology/Category/Profinite/Extend.lean b/Mathlib/Topology/Category/Profinite/Extend.lean index 8db0cc4dd84e9a..2c00a5e5808e02 100644 --- a/Mathlib/Topology/Category/Profinite/Extend.lean +++ b/Mathlib/Topology/Category/Profinite/Extend.lean @@ -141,7 +141,7 @@ then `cone G c.pt` is a limit cone. -/ noncomputable def isLimitCone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsLimit <| G.mapCone c) : - IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ hc' + IsLimit (cone G c.pt) := (functor_initial c hc).isLimitWhiskerEquiv _ _ hc' end Limit @@ -177,7 +177,7 @@ are epimorphic, then `cocone G c.pt` is a colimit cone. -/ noncomputable def isColimitCocone (hc : IsLimit c) [∀ i, Epi (c.π.app i)] (hc' : IsColimit <| G.mapCocone c.op) : - IsColimit (cocone G c.pt) := (functorOp_final c hc).isColimitWhiskerEquiv _ hc' + IsColimit (cocone G c.pt) := (functorOp_final c hc).isColimitWhiskerEquiv _ _ hc' end Colimit diff --git a/Mathlib/Topology/ContinuousMap/Basic.lean b/Mathlib/Topology/ContinuousMap/Basic.lean index e77f01dabfc8ef..34583c8f810477 100644 --- a/Mathlib/Topology/ContinuousMap/Basic.lean +++ b/Mathlib/Topology/ContinuousMap/Basic.lean @@ -39,8 +39,8 @@ section /-- `ContinuousMapClass F α β` states that `F` is a type of continuous maps. You should extend this class when you extend `ContinuousMap`. -/ -class ContinuousMapClass (F α β : Type*) [TopologicalSpace α] [TopologicalSpace β] - [FunLike F α β] : Prop where +class ContinuousMapClass (F : Type*) (α β : outParam Type*) + [TopologicalSpace α] [TopologicalSpace β] [FunLike F α β] : Prop where /-- Continuity -/ map_continuous (f : F) : Continuous f diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index b94867562a459d..0741095653d324 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -66,7 +66,7 @@ infixl:25 " →ᵈ " => Dilation /-- `DilationClass F α β r` states that `F` is a type of `r`-dilations. You should extend this typeclass when you extend `Dilation`. -/ -class DilationClass (F α β : Type*) [PseudoEMetricSpace α] [PseudoEMetricSpace β] +class DilationClass (F : Type*) (α β : outParam Type*) [PseudoEMetricSpace α] [PseudoEMetricSpace β] [FunLike F α β] : Prop where edist_eq' : ∀ f : F, ∃ r : ℝ≥0, r ≠ 0 ∧ ∀ x y : α, edist (f x) (f y) = r * edist x y diff --git a/Mathlib/Topology/Metrizable/Uniformity.lean b/Mathlib/Topology/Metrizable/Uniformity.lean index efdf16b745cecc..3dc4176d75c215 100644 --- a/Mathlib/Topology/Metrizable/Uniformity.lean +++ b/Mathlib/Topology/Metrizable/Uniformity.lean @@ -158,8 +158,8 @@ theorem le_two_mul_dist_ofPreNNDist (d : X → X → ℝ≥0) (dist_self : ∀ x ← Option.coe_def, Option.toList_some, take_append_of_le_length hMl.le, getElem_cons_succ] · exact single_le_sum (fun x _ => zero_le x) _ (mem_iff_get.2 ⟨⟨M, hM_lt⟩, getElem_zipWith⟩) · rcases hMl.eq_or_lt with (rfl | hMl) - · simp only [getElem_append_right' le_rfl, sub_self, getElem_singleton, dist_self, zero_le] - rw [getElem_append _ hMl] + · simp only [getElem_append_right le_rfl, sub_self, getElem_singleton, dist_self, zero_le] + rw [getElem_append_left hMl] have hlen : length (drop (M + 1) l) = length l - (M + 1) := length_drop _ _ have hlen_lt : length l - (M + 1) < length l := Nat.sub_lt_of_pos_le M.succ_pos hMl refine (ihn _ hlen_lt _ y _ hlen).trans ?_ diff --git a/Mathlib/Topology/Order/LawsonTopology.lean b/Mathlib/Topology/Order/LawsonTopology.lean index 4c730963736e1a..61c08e2a6fe21e 100644 --- a/Mathlib/Topology/Order/LawsonTopology.lean +++ b/Mathlib/Topology/Order/LawsonTopology.lean @@ -144,7 +144,7 @@ instance instIsLawson : IsLawson (WithLawson α) := ⟨rfl⟩ /-- If `α` is equipped with the Lawson topology, then it is homeomorphic to `WithLawson α`. -/ def homeomorph [TopologicalSpace α] [IsLawson α] : WithLawson α ≃ₜ α := - ofLawson.toHomeomorphOfInducing ⟨by erw [@IsLawson.topology_eq_lawson α _ _, induced_id]; rfl⟩ + ofLawson.toHomeomorphOfInducing ⟨by erw [IsLawson.topology_eq_lawson (α := α), induced_id]; rfl⟩ theorem isOpen_preimage_ofLawson {S : Set α} : IsOpen (ofLawson ⁻¹' S) ↔ (lawson α).IsOpen S := Iff.rfl diff --git a/Mathlib/Topology/UniformSpace/Cauchy.lean b/Mathlib/Topology/UniformSpace/Cauchy.lean index a1a7590a0efa45..b2be7ed715fd70 100644 --- a/Mathlib/Topology/UniformSpace/Cauchy.lean +++ b/Mathlib/Topology/UniformSpace/Cauchy.lean @@ -563,7 +563,7 @@ theorem TotallyBounded.image [UniformSpace β] {f : α → β} {s : Set α} (hs simp only [mem_image, iUnion_exists, biUnion_and', iUnion_iUnion_eq_right, image_subset_iff, preimage_iUnion, preimage_setOf_eq] simp? [subset_def] at hct says - simp only [mem_setOf_eq, subset_def, mem_iUnion, exists_prop', nonempty_prop] at hct + simp only [mem_setOf_eq, subset_def, mem_iUnion, exists_prop] at hct intro x hx simpa using hct x hx⟩ diff --git a/Mathlib/Util/CountHeartbeats.lean b/Mathlib/Util/CountHeartbeats.lean index 3b30d85e1aa291..bbc00430ba743b 100644 --- a/Mathlib/Util/CountHeartbeats.lean +++ b/Mathlib/Util/CountHeartbeats.lean @@ -43,8 +43,8 @@ def runTacForHeartbeats (tac : TSyntax `Lean.Parser.Tactic.tacticSeq) (revert : Given a `List Nat`, return the minimum, maximum, and standard deviation. -/ def variation (counts : List Nat) : List Nat := - let min := counts.minimum?.getD 0 - let max := counts.maximum?.getD 0 + let min := counts.min?.getD 0 + let max := counts.max?.getD 0 let toFloat (n : Nat) := n.toUInt64.toFloat let toNat (f : Float) := f.toUInt64.toNat let counts' := counts.map toFloat diff --git a/lake-manifest.json b/lake-manifest.json index 027d26b4ddaa2d..e34a0283a8064c 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f274aed7ae8d1addd3e70adaf3183ccc6e1ed43d", + "rev": "63c1c38b123b0741b7b7fd56fb8510f95bfd0e55", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", @@ -25,7 +25,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "ff420521a0c098891f4f44ecda9dd7ff57b50bad", + "rev": "b20a88676fd00affb90cbc9f1ff004ae588103b3", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "master", @@ -35,10 +35,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "eb08eee94098fe530ccd6d8751a86fe405473d4c", + "rev": "cd20dae87c48495f0220663014dff11671597fcf", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.42", + "inputRev": "v0.0.43-pre", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover/lean4-cli", @@ -55,7 +55,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "e285a7ade149c551c17a4b24f127e1ef782e4bb1", + "rev": "6b6ad220389444229d6b29c386b039e18345a003", "name": "importGraph", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lakefile.lean b/lakefile.lean index 87bbbfd1f21f02..a4a09667f6c2b7 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -10,7 +10,7 @@ open Lake DSL require "leanprover-community" / "batteries" @ git "main" require "leanprover-community" / "Qq" @ git "master" require "leanprover-community" / "aesop" @ git "master" -require "leanprover-community" / "proofwidgets" @ git "v0.0.42" +require "leanprover-community" / "proofwidgets" @ git "v0.0.43-pre" require "leanprover-community" / "importGraph" @ git "main" require "leanprover-community" / "LeanSearchClient" @ git "main" from git "https://github.com/leanprover-community/LeanSearchClient" @ "main" @@ -35,7 +35,8 @@ abbrev mathlibOnlyLinters : Array LeanOption := #[ ⟨`linter.style.longLine, true⟩, ⟨`linter.style.longFile, .ofNat 1500⟩, ⟨`linter.style.missingEnd, true⟩, - ⟨`linter.style.setOption, true⟩ + ⟨`linter.style.setOption, true⟩, + ⟨`aesop.warn.applyIff, false⟩ -- This became a problem after https://github.com/leanprover-community/aesop/commit/29cf094e84ae9852f0011b47b6ddc684ffe4be5f ] /-- These options are passed as `leanOptions` to building mathlib, as well as the diff --git a/lean-toolchain b/lean-toolchain index 89985206aca4e4..a00797801f701f 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.12.0 +leanprover/lean4:v4.13.0-rc1 diff --git a/scripts/noshake.json b/scripts/noshake.json index 735d7aacd17cfb..648722b3701598 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -360,7 +360,8 @@ ["Batteries.Data.Nat.Lemmas", "Mathlib.Data.List.Basic"], "Mathlib.Data.List.Lemmas": ["Mathlib.Data.List.InsertNth"], "Mathlib.Data.List.Defs": ["Batteries.Data.RBMap.Basic"], - "Mathlib.Data.List.Basic": ["Mathlib.Data.Option.Basic"], + "Mathlib.Data.List.Basic": + ["Mathlib.Control.Basic", "Mathlib.Data.Option.Basic"], "Mathlib.Data.LazyList.Basic": ["Mathlib.Lean.Thunk"], "Mathlib.Data.Int.Order.Basic": ["Mathlib.Data.Int.Notation"], "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], diff --git a/test/ValuedCSP.lean b/test/ValuedCSP.lean index f60e3aed3e4212..48b89612eacdc4 100644 --- a/test/ValuedCSP.lean +++ b/test/ValuedCSP.lean @@ -40,6 +40,8 @@ private def exampleFiniteValuedInstance : exampleFiniteValuedCSP.Instance (Fin 2 example : exampleFiniteValuedInstance.IsOptimumSolution ![(0 : ℚ), (0 : ℚ)] := by intro s convert_to 0 ≤ exampleFiniteValuedInstance.evalSolution s + · simp [exampleFiniteValuedInstance, ValuedCSP.Instance.evalSolution] + exact Rat.zero_add 0 rw [ValuedCSP.Instance.evalSolution, exampleFiniteValuedInstance] convert_to 0 ≤ |s 0| + |s 1| · simp [ValuedCSP.unaryTerm, ValuedCSP.Term.evalSolution, Function.OfArity.uncurry] diff --git a/test/aesop_cat.lean b/test/aesop_cat.lean index 7790b9c1b983c3..549cd69b92adb7 100644 --- a/test/aesop_cat.lean +++ b/test/aesop_cat.lean @@ -10,6 +10,8 @@ example : Foo where x := sorry /-- +error: could not synthesize default value for field 'w' of 'Foo' using tactics +--- error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. Initial goal: ⊢ 35 = 37 diff --git a/test/matrix.lean b/test/matrix.lean index 962e193b1ef976..522acbe923c7d2 100644 --- a/test/matrix.lean +++ b/test/matrix.lean @@ -154,7 +154,7 @@ example {α : Type _} [CommRing α] {a b c d : α} : Fin.isValue, of_apply, cons_val', empty_val', cons_val_fin_one, cons_val_zero, det_unique, Fin.default_eq_zero, submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, Fin.zero_succAbove, head_cons, - Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, cons_val_succ, neg_mul, + Finset.univ_unique, Fin.val_succ, Fin.val_eq_zero, zero_add, pow_one, cons_val_succ, neg_mul, Fin.succ_succAbove_zero, Finset.sum_const, Finset.card_singleton, smul_neg, one_smul] ring @@ -167,7 +167,7 @@ example {α : Type _} [CommRing α] {a b c d e f g h i : α} : submatrix_apply, Fin.succ_zero_eq_one, cons_val_one, head_cons, submatrix_submatrix, det_unique, Fin.default_eq_zero, Function.comp_apply, Fin.succ_one_eq_two, cons_val_two, tail_cons, head_fin_const, Fin.sum_univ_succ, Fin.val_zero, pow_zero, one_mul, - Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, Fin.coe_fin_one, zero_add, pow_one, + Fin.zero_succAbove, Finset.univ_unique, Fin.val_succ, Fin.val_eq_zero, 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_two, Even.neg_pow, one_pow, Finset.sum_const, Finset.card_singleton, one_smul] diff --git a/test/says.lean b/test/says.lean index e2f9a3f8346624..4315e51cb870d0 100644 --- a/test/says.lean +++ b/test/says.lean @@ -101,7 +101,8 @@ def very_long_lemma_name_aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa : Q → P := fun _ @[simp] def very_long_lemma_name_bbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbb : Q := trivial /-- -info: Try this: aesop? says simp_all only [very_long_lemma_name_bbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbb, +info: Try this: aesop? says + simp_all only [very_long_lemma_name_bbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbbb, very_long_lemma_name_aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa] -/ #guard_msgs in