From 50eef2713885d1c5ba97242e9b4823201b36c2fd Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 30 Apr 2024 06:07:46 +0000 Subject: [PATCH] chore: adapt to multiple goal linter 1 (#12338) A PR accompanying #12339. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/linting.20against.20.22multi-goal.20proofs.22) --- Archive/Imo/Imo1972Q5.lean | 4 +- Archive/Imo/Imo2008Q2.lean | 6 +- Archive/Wiedijk100Theorems/AbelRuffini.lean | 31 +++++----- Archive/Wiedijk100Theorems/BallotProblem.lean | 4 +- Archive/Wiedijk100Theorems/CubingACube.lean | 60 ++++++++++++------- Archive/Wiedijk100Theorems/Partition.lean | 6 +- Counterexamples/Cyclotomic105.lean | 14 ++--- Counterexamples/DirectSumIsInternal.lean | 4 +- Mathlib/Algebra/BigOperators/Finprod.lean | 6 +- Mathlib/Algebra/BigOperators/List/Basic.lean | 3 +- .../Algebra/Category/Ring/Constructions.lean | 14 ++--- .../Computation/Approximations.lean | 4 +- Mathlib/Algebra/GCDMonoid/Finset.lean | 2 +- Mathlib/Algebra/GeomSum.lean | 4 +- Mathlib/Algebra/GroupPower/Order.lean | 6 +- Mathlib/Algebra/Lie/DirectSum.lean | 12 ++-- Mathlib/Algebra/Lie/IdealOperations.lean | 3 +- Mathlib/Algebra/Module/LinearMap/End.lean | 3 +- .../Order/BigOperators/Group/List.lean | 4 +- .../Algebra/Order/CauSeq/BigOperators.lean | 4 +- Mathlib/Algebra/Order/Field/Basic.lean | 4 +- Mathlib/Algebra/Order/Group/Int.lean | 4 +- Mathlib/Algebra/Order/Group/MinMax.lean | 4 +- Mathlib/Algebra/Order/Interval/Basic.lean | 4 +- Mathlib/Algebra/Order/Pointwise.lean | 32 +++++----- Mathlib/Algebra/Order/Sub/Defs.lean | 6 +- Mathlib/Algebra/Quandle.lean | 11 ++-- Mathlib/AlgebraicGeometry/Gluing.lean | 11 ++-- .../AlgebraicTopology/SimplexCategory.lean | 4 +- Mathlib/Analysis/Calculus/FDeriv/Basic.lean | 2 +- .../Analysis/Calculus/ParametricIntegral.lean | 2 +- .../Analysis/Calculus/UniformLimitsDeriv.lean | 4 +- Mathlib/Analysis/Complex/Polynomial.lean | 9 ++- .../Instances.lean | 2 +- .../SpecialFunctions/Trigonometric/Angle.lean | 8 +-- .../SpecialFunctions/Trigonometric/Basic.lean | 6 +- Mathlib/CategoryTheory/Category/Quiv.lean | 12 ++-- Mathlib/CategoryTheory/GlueData.lean | 2 +- Mathlib/CategoryTheory/IsConnected.lean | 14 ++--- Mathlib/CategoryTheory/Limits/Connected.lean | 24 ++++---- .../Limits/Constructions/BinaryProducts.lean | 2 +- .../FiniteProductsOfBinaryProducts.lean | 4 +- Mathlib/CategoryTheory/Limits/Final.lean | 6 +- Mathlib/CategoryTheory/Limits/Fubini.lean | 4 +- .../CategoryTheory/Limits/KanExtension.lean | 4 +- Mathlib/CategoryTheory/Limits/Lattice.lean | 12 ++-- .../Limits/Preserves/Shapes/Products.lean | 4 +- Mathlib/CategoryTheory/Limits/Presheaf.lean | 8 +-- .../Limits/Shapes/Biproducts.lean | 4 +- .../Limits/Shapes/Equalizers.lean | 5 +- .../Limits/Shapes/Pullbacks.lean | 6 +- .../Limits/Shapes/SingleObj.lean | 4 +- .../Limits/Shapes/WidePullbacks.lean | 16 ++--- .../Limits/Shapes/ZeroMorphisms.lean | 8 +-- .../OfChosenFiniteProducts/Basic.lean | 18 +++--- Mathlib/CategoryTheory/Sites/Canonical.lean | 4 +- Mathlib/CategoryTheory/Sites/Coverage.lean | 6 +- .../CategoryTheory/Subobject/MonoOver.lean | 30 +++++----- .../Enumerative/Composition.lean | 6 +- .../Combinatorics/Quiver/Arborescence.lean | 4 +- .../SetFamily/FourFunctions.lean | 4 +- Mathlib/Combinatorics/SetFamily/LYM.lean | 4 +- .../Combinatorics/SimpleGraph/Acyclic.lean | 6 +- .../SimpleGraph/Connectivity.lean | 12 +++- .../Combinatorics/SimpleGraph/Density.lean | 4 +- Mathlib/Combinatorics/SimpleGraph/Maps.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Prod.lean | 8 +-- Mathlib/Combinatorics/Young/YoungDiagram.lean | 8 +-- Mathlib/Computability/RegularExpressions.lean | 4 +- Mathlib/Computability/TMToPartrec.lean | 34 +++++++---- Mathlib/Computability/TuringMachine.lean | 15 +++-- Mathlib/Control/Fix.lean | 6 +- Mathlib/Control/LawfulFix.lean | 9 ++- Mathlib/Data/Analysis/Filter.lean | 4 +- Mathlib/Data/DFinsupp/Basic.lean | 10 ++-- Mathlib/Data/Fin/Basic.lean | 4 +- Mathlib/Data/FinEnum.lean | 2 +- Mathlib/Data/Finset/Lattice.lean | 2 +- Mathlib/Data/Finset/NoncommProd.lean | 12 ++-- Mathlib/Data/Finset/Sym.lean | 5 +- Mathlib/Data/Finsupp/Basic.lean | 10 ++-- Mathlib/Data/Finsupp/Defs.lean | 4 +- Mathlib/Data/Fintype/Quotient.lean | 4 +- Mathlib/Data/Int/Bitwise.lean | 9 +-- Mathlib/Data/List/Basic.lean | 18 ++++-- Mathlib/Data/List/Chain.lean | 4 +- Mathlib/Data/List/Cycle.lean | 28 ++++----- Mathlib/Data/List/Dedup.lean | 6 +- Mathlib/Data/List/Indexes.lean | 40 +++++++++---- Mathlib/Data/List/MinMax.lean | 6 +- Mathlib/Data/List/Perm.lean | 6 +- Mathlib/Data/List/Permutation.lean | 7 ++- Mathlib/Data/List/Sigma.lean | 9 +-- Mathlib/Data/List/Sublists.lean | 7 ++- Mathlib/Data/List/Zip.lean | 4 +- Mathlib/Data/Multiset/Antidiagonal.lean | 5 +- Mathlib/Data/Multiset/Pi.lean | 2 +- Mathlib/Data/Multiset/Sum.lean | 4 +- Mathlib/Data/NNRat/Defs.lean | 4 +- Mathlib/Data/Nat/Bits.lean | 4 +- Mathlib/Data/Nat/Cast/Field.lean | 10 ++-- Mathlib/Data/Nat/Defs.lean | 9 +-- Mathlib/Data/Nat/EvenOddRec.lean | 10 +++- Mathlib/Data/Nat/Fib/Zeckendorf.lean | 2 +- Mathlib/Data/Nat/Lattice.lean | 5 +- Mathlib/Data/Nat/ModEq.lean | 24 +++++--- Mathlib/Data/Nat/PSub.lean | 4 +- Mathlib/Data/Nat/Pairing.lean | 6 +- Mathlib/Data/Nat/WithBot.lean | 18 +++--- Mathlib/Data/Num/Lemmas.lean | 6 +- Mathlib/Data/Ordmap/Ordset.lean | 53 ++++++++-------- Mathlib/Data/PFunctor/Multivariate/M.lean | 2 +- Mathlib/Data/PFunctor/Univariate/M.lean | 6 +- Mathlib/Data/PNat/Factors.lean | 16 ++--- Mathlib/Data/PNat/Prime.lean | 11 ++-- Mathlib/Data/PNat/Xgcd.lean | 12 ++-- Mathlib/Data/Pi/Lex.lean | 6 +- Mathlib/Data/Prod/TProd.lean | 3 +- .../QPF/Multivariate/Constructions/Fix.lean | 5 +- Mathlib/Data/QPF/Univariate/Basic.lean | 4 +- Mathlib/Data/Rat/Cast/Defs.lean | 12 ++-- Mathlib/Data/Rat/Defs.lean | 2 +- Mathlib/Data/Rat/Lemmas.lean | 4 +- Mathlib/Data/Seq/Parallel.lean | 16 ++--- Mathlib/Data/Seq/Seq.lean | 12 ++-- Mathlib/Data/Seq/WSeq.lean | 6 +- Mathlib/Data/Set/Image.lean | 12 ++-- Mathlib/Data/Stream/Init.lean | 8 ++- Mathlib/Data/Sum/Interval.lean | 4 +- Mathlib/Data/Sum/Order.lean | 5 +- Mathlib/Data/TypeVec.lean | 12 ++-- Mathlib/Data/UInt.lean | 4 +- Mathlib/Data/Vector/Basic.lean | 4 +- Mathlib/Data/Vector3.lean | 7 ++- .../Dynamics/Ergodic/MeasurePreserving.lean | 2 +- .../Manifold/VectorBundle/SmoothSection.lean | 8 +-- .../Manifold/VectorBundle/Tangent.lean | 8 +-- Mathlib/GroupTheory/Congruence.lean | 2 +- Mathlib/GroupTheory/CoprodI.lean | 28 ++++----- Mathlib/GroupTheory/MonoidLocalization.lean | 9 +-- Mathlib/GroupTheory/Perm/ClosureSwap.lean | 9 ++- Mathlib/GroupTheory/Perm/Option.lean | 2 +- Mathlib/GroupTheory/Perm/Support.lean | 14 ++--- Mathlib/GroupTheory/Subgroup/Basic.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 3 +- Mathlib/LinearAlgebra/FiniteDimensional.lean | 12 ++-- Mathlib/LinearAlgebra/LinearIndependent.lean | 6 +- Mathlib/LinearAlgebra/Matrix/Block.lean | 6 +- Mathlib/Logic/Relation.lean | 12 ++-- .../MeasureTheory/Decomposition/Jordan.lean | 13 ++-- .../Decomposition/SignedHahn.lean | 2 +- .../Decomposition/SignedLebesgue.lean | 10 ++-- .../Decomposition/UnsignedHahn.lean | 10 ++-- .../ConditionalExpectation/CondexpL2.lean | 8 +-- Mathlib/MeasureTheory/Group/Convolution.lean | 4 +- .../Group/FundamentalDomain.lean | 4 +- Mathlib/MeasureTheory/Integral/Average.lean | 2 +- .../Integral/DominatedConvergence.lean | 3 +- .../Integral/IntervalIntegral.lean | 2 +- Mathlib/MeasureTheory/Integral/Lebesgue.lean | 6 +- .../MeasureTheory/MeasurableSpace/Basic.lean | 8 +-- Mathlib/MeasureTheory/Measure/Content.lean | 4 +- .../Measure/LevyProkhorovMetric.lean | 20 +++---- Mathlib/MeasureTheory/Measure/Sub.lean | 4 +- .../MeasureTheory/Measure/VectorMeasure.lean | 16 ++--- Mathlib/NumberTheory/ADEInequality.lean | 24 ++++---- Mathlib/NumberTheory/Cyclotomic/Rat.lean | 6 +- Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean | 8 +-- .../LSeries/MellinEqDirichlet.lean | 2 +- .../LegendreSymbol/JacobiSymbol.lean | 15 ++--- .../ModularForms/JacobiTheta/Bounds.lean | 4 +- .../NumberField/CanonicalEmbedding/Basic.lean | 4 +- .../CanonicalEmbedding/ConvexBody.lean | 34 +++++------ Mathlib/NumberTheory/NumberField/Units.lean | 12 ++-- Mathlib/Order/CompactlyGenerated/Basic.lean | 2 +- .../ConditionallyCompleteLattice/Basic.lean | 2 +- Mathlib/Order/Extension/Linear.lean | 10 ++-- Mathlib/Order/Filter/AtTopBot.lean | 6 +- Mathlib/Order/Filter/Basic.lean | 5 +- Mathlib/Order/Filter/Prod.lean | 4 +- Mathlib/Order/Interval/Basic.lean | 2 +- Mathlib/Order/Iterate.lean | 4 +- Mathlib/Order/Minimal.lean | 8 ++- Mathlib/Order/OmegaCompletePartialOrder.lean | 16 ++--- Mathlib/Order/RelSeries.lean | 15 +++-- Mathlib/Order/WellFounded.lean | 10 ++-- Mathlib/Order/WellFoundedSet.lean | 15 ++--- .../Kernel/MeasurableIntegral.lean | 12 ++-- Mathlib/Probability/StrongLaw.lean | 3 +- Mathlib/RingTheory/HahnSeries/Basic.lean | 2 +- Mathlib/RingTheory/Ideal/Cotangent.lean | 4 +- Mathlib/RingTheory/Multiplicity.lean | 4 +- Mathlib/SetTheory/Game/Basic.lean | 6 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 8 ++- Mathlib/Tactic/CancelDenoms/Core.lean | 6 +- Mathlib/Testing/SlimCheck/Gen.lean | 6 +- Mathlib/Topology/Basic.lean | 4 +- .../Topology/ContinuousFunction/Bounded.lean | 4 +- Mathlib/Topology/TietzeExtension.lean | 4 +- 199 files changed, 912 insertions(+), 768 deletions(-) diff --git a/Archive/Imo/Imo1972Q5.lean b/Archive/Imo/Imo1972Q5.lean index 403c7b71135369..f31b9757b199f1 100644 --- a/Archive/Imo/Imo1972Q5.lean +++ b/Archive/Imo/Imo1972Q5.lean @@ -53,8 +53,8 @@ theorem imo1972_q5 (f g : ℝ → ℝ) (hf1 : ∀ x, ∀ y, f (x + y) + f (x - y 0 < ‖f x‖ := norm_pos_iff.mpr hx _ ≤ k := hk₁ x rw [div_lt_iff] - apply lt_mul_of_one_lt_right h₁ hneg - exact zero_lt_one.trans hneg + · apply lt_mul_of_one_lt_right h₁ hneg + · exact zero_lt_one.trans hneg -- Demonstrate that `k ≤ k'` using `hk₂`. have H₂ : k ≤ k' := by have h₁ : ∃ x : ℝ, x ∈ S := by use ‖f 0‖; exact Set.mem_range_self 0 diff --git a/Archive/Imo/Imo2008Q2.lean b/Archive/Imo/Imo2008Q2.lean index 39e299bf497659..2c89203c6a0632 100644 --- a/Archive/Imo/Imo2008Q2.lean +++ b/Archive/Imo/Imo2008Q2.lean @@ -103,9 +103,9 @@ theorem imo2008_q2b : Set.Infinite rationalSolutions := by set z : ℚ := -t * (t + 1) with hz_def simp only [t, W, K, g, Set.mem_image, Prod.exists] use x, y, z; constructor - simp only [Set.mem_setOf_eq] - · use x, y, z; constructor - rfl + · simp only [Set.mem_setOf_eq] + use x, y, z; constructor + · rfl · use t; constructor · simp only [t, gt_iff_lt, lt_max_iff]; right; trivial exact ⟨rfl, rfl, rfl⟩ diff --git a/Archive/Wiedijk100Theorems/AbelRuffini.lean b/Archive/Wiedijk100Theorems/AbelRuffini.lean index c77a6ab08b354b..28d708c441ab48 100644 --- a/Archive/Wiedijk100Theorems/AbelRuffini.lean +++ b/Archive/Wiedijk100Theorems/AbelRuffini.lean @@ -85,21 +85,22 @@ theorem monic_Phi : (Φ R a b).Monic := theorem irreducible_Phi (p : ℕ) (hp : p.Prime) (hpa : p ∣ a) (hpb : p ∣ b) (hp2b : ¬p ^ 2 ∣ b) : Irreducible (Φ ℚ a b) := by rw [← map_Phi a b (Int.castRingHom ℚ), ← IsPrimitive.Int.irreducible_iff_irreducible_map_cast] - apply irreducible_of_eisenstein_criterion - · rwa [span_singleton_prime (Int.natCast_ne_zero.mpr hp.ne_zero), Int.prime_iff_natAbs_prime] - · rw [leadingCoeff_Phi, mem_span_singleton] - exact mod_cast mt Nat.dvd_one.mp hp.ne_one - · intro n hn - rw [mem_span_singleton] - rw [degree_Phi] at hn; norm_cast at hn - interval_cases hn : n <;> - simp (config := {decide := true}) only [Φ, coeff_X_pow, coeff_C, Int.natCast_dvd_natCast.mpr, - hpb, if_true, coeff_C_mul, if_false, coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, - coeff_sub, add_zero, zero_sub, dvd_neg, neg_zero, dvd_mul_of_dvd_left] - · simp only [degree_Phi, ← WithBot.coe_zero, WithBot.coe_lt_coe, Nat.succ_pos'] - decide - · rw [coeff_zero_Phi, span_singleton_pow, mem_span_singleton] - exact mt Int.natCast_dvd_natCast.mp hp2b + on_goal 1 => + apply irreducible_of_eisenstein_criterion + · rwa [span_singleton_prime (Int.natCast_ne_zero.mpr hp.ne_zero), Int.prime_iff_natAbs_prime] + · rw [leadingCoeff_Phi, mem_span_singleton] + exact mod_cast mt Nat.dvd_one.mp hp.ne_one + · intro n hn + rw [mem_span_singleton] + rw [degree_Phi] at hn; norm_cast at hn + interval_cases hn : n <;> + simp (config := {decide := true}) only [Φ, coeff_X_pow, coeff_C, Int.natCast_dvd_natCast.mpr, + hpb, if_true, coeff_C_mul, if_false, coeff_X_zero, hpa, coeff_add, zero_add, mul_zero, + coeff_sub, add_zero, zero_sub, dvd_neg, neg_zero, dvd_mul_of_dvd_left] + · simp only [degree_Phi, ← WithBot.coe_zero, WithBot.coe_lt_coe, Nat.succ_pos'] + decide + · rw [coeff_zero_Phi, span_singleton_pow, mem_span_singleton] + exact mt Int.natCast_dvd_natCast.mp hp2b all_goals exact Monic.isPrimitive (monic_Phi a b) #align abel_ruffini.irreducible_Phi AbelRuffini.irreducible_Phi diff --git a/Archive/Wiedijk100Theorems/BallotProblem.lean b/Archive/Wiedijk100Theorems/BallotProblem.lean index 46c4eb20ffc673..7692712aab2ee3 100644 --- a/Archive/Wiedijk100Theorems/BallotProblem.lean +++ b/Archive/Wiedijk100Theorems/BallotProblem.lean @@ -387,8 +387,8 @@ theorem ballot_problem' : ring all_goals refine' (ENNReal.mul_lt_top _ _).ne - exact (measure_lt_top _ _).ne - simp [Ne, ENNReal.div_eq_top] + · exact (measure_lt_top _ _).ne + · simp [Ne, ENNReal.div_eq_top] #align ballot.ballot_problem' Ballot.ballot_problem' /-- The ballot problem. -/ diff --git a/Archive/Wiedijk100Theorems/CubingACube.lean b/Archive/Wiedijk100Theorems/CubingACube.lean index 756490af096c1b..ee5da4dc7aafae 100644 --- a/Archive/Wiedijk100Theorems/CubingACube.lean +++ b/Archive/Wiedijk100Theorems/CubingACube.lean @@ -182,8 +182,8 @@ theorem zero_le_b {i j} : 0 ≤ (cs i).b j := theorem b_add_w_le_one {j} : (cs i).b j + (cs i).w ≤ 1 := by have : side (cs i) j ⊆ Ico 0 1 := side_subset h rw [side, Ico_subset_Ico_iff] at this - convert this.2 - simp [hw] + · convert this.2 + · simp [hw] #align theorems_100.«82».correct.b_add_w_le_one Theorems100.«82».Correct.b_add_w_le_one theorem nontrivial_fin : Nontrivial (Fin n) := @@ -199,9 +199,12 @@ theorem w_ne_one [Nontrivial ι] (i : ι) : (cs i).w ≠ 1 := by have h2p : p ∈ (cs i).toSet := by intro j; constructor trans (0 : ℝ) - · rw [← add_le_add_iff_right (1 : ℝ)]; convert b_add_w_le_one h; rw [hi]; rw [zero_add] - apply zero_le_b h; apply lt_of_lt_of_le (side_subset h <| (cs i').b_mem_side j).2 - simp [hi, zero_le_b h] + · rw [← add_le_add_iff_right (1 : ℝ)]; convert b_add_w_le_one h + · rw [hi] + · rw [zero_add] + · apply zero_le_b h + · apply lt_of_lt_of_le (side_subset h <| (cs i').b_mem_side j).2 + simp [hi, zero_le_b h] exact (h.PairwiseDisjoint hi').le_bot ⟨hp, h2p⟩ #align theorems_100.«82».correct.w_ne_one Theorems100.«82».Correct.w_ne_one @@ -213,7 +216,9 @@ theorem shiftUp_bottom_subset_bottoms (hc : (cs i).xm ≠ 1) : have : p ∈ (unitCube : Cube (n + 1)).toSet := by simp only [toSet, forall_fin_succ, hp0, side_unitCube, mem_setOf_eq, mem_Ico, head_shiftUp] refine' ⟨⟨_, _⟩, _⟩ - · rw [← zero_add (0 : ℝ)]; apply add_le_add; apply zero_le_b h; apply (cs i).hw' + · rw [← zero_add (0 : ℝ)]; apply add_le_add + · apply zero_le_b h + · apply (cs i).hw' · exact lt_of_le_of_ne (b_add_w_le_one h) hc intro j; exact side_subset h (hps j) rw [← h.2, mem_iUnion] at this; rcases this with ⟨i', hi'⟩ @@ -222,9 +227,9 @@ theorem shiftUp_bottom_subset_bottoms (hc : (cs i).xm ≠ 1) : have := h.1 this rw [onFun, comp_apply, comp_apply, toSet_disjoint, exists_fin_succ] at this rcases this with (h0 | ⟨j, hj⟩) - rw [hp0]; symm; apply eq_of_Ico_disjoint h0 (by simp [hw]) _ - convert hi' 0; rw [hp0]; rfl - exfalso; apply not_disjoint_iff.mpr ⟨tail p j, hps j, hi' j.succ⟩ hj + · rw [hp0]; symm; apply eq_of_Ico_disjoint h0 (by simp [hw]) _ + convert hi' 0; rw [hp0]; rfl + · exfalso; apply not_disjoint_iff.mpr ⟨tail p j, hps j, hi' j.succ⟩ hj #align theorems_100.«82».correct.shift_up_bottom_subset_bottoms Theorems100.«82».Correct.shiftUp_bottom_subset_bottoms end Correct @@ -255,11 +260,15 @@ theorem valley_unitCube [Nontrivial ι] (h : Correct cs) : Valley cs unitCube := intro h0 hv have : v ∈ (unitCube : Cube (n + 1)).toSet := by dsimp only [toSet, unitCube, mem_setOf_eq] - rw [forall_fin_succ, h0]; constructor; norm_num [side, unitCube]; exact hv + rw [forall_fin_succ, h0]; constructor + · norm_num [side, unitCube] + · exact hv rw [← h.2, mem_iUnion] at this; rcases this with ⟨i, hi⟩ use i constructor - · apply le_antisymm; rw [h0]; exact h.zero_le_b; exact (hi 0).1 + · apply le_antisymm + · rw [h0]; exact h.zero_le_b + · exact (hi 0).1 intro j; exact hi _ · intro i _ _; rw [toSet_subset]; intro j; convert h.side_subset using 1; simp [side_tail] · intro i _; exact h.w_ne_one i @@ -290,7 +299,8 @@ theorem b_le_b (hi : i ∈ bcubes cs c) (j : Fin n) : c.b j.succ ≤ (cs i).b j. theorem t_le_t (hi : i ∈ bcubes cs c) (j : Fin n) : (cs i).b j.succ + (cs i).w ≤ c.b j.succ + c.w := by have h' := tail_sub hi j; dsimp only [side] at h'; rw [Ico_subset_Ico_iff] at h' - exact h'.2; simp [hw] + · exact h'.2 + · simp [hw] #align theorems_100.«82».t_le_t Theorems100.«82».t_le_t /-- Every cube in the valley must be smaller than it -/ @@ -362,9 +372,11 @@ theorem mi_strict_minimal (hii' : mi h v ≠ i) (hi : i ∈ bcubes cs c) : /-- The top of `mi` cannot be 1, since there is a larger cube in the valley -/ theorem mi_xm_ne_one : (cs <| mi h v).xm ≠ 1 := by apply ne_of_lt; rcases (nontrivial_bcubes h v).exists_ne (mi h v) with ⟨i, hi, h2i⟩ - apply lt_of_lt_of_le _ h.b_add_w_le_one; exact i; exact 0 - rw [xm, mi_mem_bcubes.1, hi.1, _root_.add_lt_add_iff_left] - exact mi_strict_minimal h2i.symm hi + · apply lt_of_lt_of_le _ h.b_add_w_le_one + · exact i + · exact 0 + rw [xm, mi_mem_bcubes.1, hi.1, _root_.add_lt_add_iff_left] + exact mi_strict_minimal h2i.symm hi #align theorems_100.«82».mi_xm_ne_one Theorems100.«82».mi_xm_ne_one /-- If `mi` lies on the boundary of the valley in dimension j, then this lemma expresses that all @@ -382,8 +394,8 @@ theorem smallest_onBoundary {j} (bi : OnBoundary (mi_mem_bcubes : mi h v ∈ _) · simp [side, bi, hw', w_lt_w h v hi] · intro h'; simpa [lt_irrefl] using h'.2 intro i' hi' i'_i h2i'; constructor - apply le_trans h2i'.1 - · simp [hw'] + · apply le_trans h2i'.1 + simp [hw'] apply lt_of_lt_of_le (add_lt_add_left (mi_strict_minimal i'_i.symm hi') _) simp [bi.symm, b_le_b hi'] let s := bcubes cs c \ {i} @@ -401,7 +413,7 @@ theorem smallest_onBoundary {j} (bi : OnBoundary (mi_mem_bcubes : mi h v ∈ _) rw [add_assoc, le_add_iff_nonneg_right, ← sub_eq_add_neg, sub_nonneg] apply le_of_lt (w_lt_w h v hi') · simp only [side, not_and_or, not_lt, not_le, mem_Ico]; left; exact hx - intro i'' hi'' h2i'' h3i''; constructor; swap; apply lt_trans hx h3i''.2 + intro i'' hi'' h2i'' h3i''; constructor; swap; · apply lt_trans hx h3i''.2 rw [le_sub_iff_add_le] refine' le_trans _ (t_le_t hi'' j); rw [add_le_add_iff_left]; apply h3i' i'' ⟨hi'', _⟩ simp [mem_singleton, h2i''] @@ -429,14 +441,16 @@ theorem mi_not_onBoundary (j : Fin n) : ¬OnBoundary (mi_mem_bcubes : mi h v ∈ have h2i' : i' ∈ bcubes cs c := ⟨hi'.1.symm, v.2.1 i' hi'.1.symm ⟨tail p, hi'.2, hp.2⟩⟩ have i_i' : i ≠ i' := by rintro rfl; simpa [p, side_tail, h2x] using hi'.2 j have : Nonempty (↥((cs i').tail.side j' \ (cs i).tail.side j')) := by - apply nonempty_Ico_sdiff; apply mi_strict_minimal i_i' h2i'; apply hw + apply nonempty_Ico_sdiff + · apply mi_strict_minimal i_i' h2i' + · apply hw rcases this with ⟨⟨x', hx'⟩⟩ let p' : Fin (n + 1) → ℝ := cons (c.b 0) fun j₂ => if j₂ = j' then x' else (cs i).b j₂.succ have hp' : p' ∈ c.bottom := by suffices ∀ j : Fin n, ite (j = j') x' ((cs i).b j.succ) ∈ c.side j.succ by simpa [p', bottom, toSet, tail, side_tail] intro j₂ - by_cases hj₂ : j₂ = j'; simp [hj₂]; apply tail_sub h2i'; apply hx'.1 + by_cases hj₂ : j₂ = j'; · simp [hj₂]; apply tail_sub h2i'; apply hx'.1 simp only [if_congr, if_false, hj₂]; apply tail_sub hi; apply b_mem_side rcases v.1 hp' with ⟨_, ⟨i'', rfl⟩, hi''⟩ have h2i'' : i'' ∈ bcubes cs c := ⟨hi''.1.symm, v.2.1 i'' hi''.1.symm ⟨tail p', hi''.2, hp'.2⟩⟩ @@ -474,7 +488,7 @@ theorem mi_not_onBoundary' (j : Fin n) : have := mi_not_onBoundary h v j simp only [OnBoundary, not_or] at this; cases' this with h1 h2 constructor - apply lt_of_le_of_ne (b_le_b mi_mem_bcubes _) h1 + · apply lt_of_le_of_ne (b_le_b mi_mem_bcubes _) h1 apply lt_of_le_of_ne _ h2 apply ((Ico_subset_Ico_iff _).mp (tail_sub mi_mem_bcubes j)).2 simp [hw] @@ -512,7 +526,9 @@ theorem valley_mi : Valley cs (cs (mi h v)).shiftUp := by rcases v.1 hp with ⟨_, ⟨i'', rfl⟩, hi''⟩ have h2i'' : i'' ∈ bcubes cs c := by use hi''.1.symm; apply v.2.1 i'' hi''.1.symm - use tail p; constructor; exact hi''.2; rw [tail_cons]; exact h3p3 + use tail p; constructor + · exact hi''.2 + · rw [tail_cons]; exact h3p3 have h3i'' : (cs i).w < (cs i'').w := by apply mi_strict_minimal _ h2i''; rintro rfl; apply h2p3; convert hi''.2 let p' := @cons n (fun _ => ℝ) (cs i).xm p3 diff --git a/Archive/Wiedijk100Theorems/Partition.lean b/Archive/Wiedijk100Theorems/Partition.lean index 22e19791ddd52f..014a4c99c19bfe 100644 --- a/Archive/Wiedijk100Theorems/Partition.lean +++ b/Archive/Wiedijk100Theorems/Partition.lean @@ -215,7 +215,7 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) exact fun hi _ => ha.2 i hi · conv_rhs => simp [← a.parts_sum] rw [sum_multiset_count_of_subset _ s] - simp only [smul_eq_mul] + · simp only [smul_eq_mul] · intro i simp only [Multiset.mem_toFinset, not_not, mem_filter] apply ha.2 @@ -229,8 +229,8 @@ theorem partialGF_prop (α : Type*) [CommSemiring α] (n : ℕ) (s : Finset ℕ) by_cases hi : i = 0 · rw [hi] rw [Multiset.count_eq_zero_of_not_mem] - rw [Multiset.count_eq_zero_of_not_mem] - intro a; exact Nat.lt_irrefl 0 (hs 0 (hp₂.2 0 a)) + · rw [Multiset.count_eq_zero_of_not_mem] + intro a; exact Nat.lt_irrefl 0 (hs 0 (hp₂.2 0 a)) intro a; exact Nat.lt_irrefl 0 (hs 0 (hp₁.2 0 a)) · rw [← mul_left_inj' hi] rw [Function.funext_iff] at h diff --git a/Counterexamples/Cyclotomic105.lean b/Counterexamples/Cyclotomic105.lean index f41dd881bc38b9..06d08ac62f418d 100644 --- a/Counterexamples/Cyclotomic105.lean +++ b/Counterexamples/Cyclotomic105.lean @@ -70,7 +70,7 @@ theorem cyclotomic_15 : cyclotomic 15 ℤ = 1 - X + X ^ 3 - X ^ 4 + X ^ 5 - X ^ refine' ((eq_cyclotomic_iff (by norm_num) _).2 _).symm rw [properDivisors_15, Finset.prod_insert _, Finset.prod_insert _, Finset.prod_singleton, cyclotomic_one, cyclotomic_3, cyclotomic_5] - ring + · ring repeat' norm_num #align counterexample.cyclotomic_15 Counterexample.cyclotomic_15 @@ -79,7 +79,7 @@ theorem cyclotomic_21 : refine' ((eq_cyclotomic_iff (by norm_num) _).2 _).symm rw [properDivisors_21, Finset.prod_insert _, Finset.prod_insert _, Finset.prod_singleton, cyclotomic_one, cyclotomic_3, cyclotomic_7] - ring + · ring repeat' norm_num #align counterexample.cyclotomic_21 Counterexample.cyclotomic_21 @@ -90,7 +90,7 @@ theorem cyclotomic_35 : refine' ((eq_cyclotomic_iff (by norm_num) _).2 _).symm rw [properDivisors_35, Finset.prod_insert _, Finset.prod_insert _, Finset.prod_singleton, cyclotomic_one, cyclotomic_5, cyclotomic_7] - ring + · ring repeat' norm_num #align counterexample.cyclotomic_35 Counterexample.cyclotomic_35 @@ -102,10 +102,10 @@ theorem cyclotomic_105 : X ^ 46 + X ^ 47 + X ^ 48 := by refine' ((eq_cyclotomic_iff (by norm_num) _).2 _).symm rw [properDivisors_105] - repeat' rw [Finset.prod_insert (α := ℕ) (β := ℤ[X])] - rw [Finset.prod_singleton, cyclotomic_one, cyclotomic_3, cyclotomic_5, cyclotomic_7, - cyclotomic_15, cyclotomic_21, cyclotomic_35] - ring + repeat rw [Finset.prod_insert (α := ℕ) (β := ℤ[X])] + · rw [Finset.prod_singleton, cyclotomic_one, cyclotomic_3, cyclotomic_5, cyclotomic_7, + cyclotomic_15, cyclotomic_21, cyclotomic_35] + ring repeat' norm_num #align counterexample.cyclotomic_105 Counterexample.cyclotomic_105 diff --git a/Counterexamples/DirectSumIsInternal.lean b/Counterexamples/DirectSumIsInternal.lean index e5af485bd6915b..884b5273321366 100644 --- a/Counterexamples/DirectSumIsInternal.lean +++ b/Counterexamples/DirectSumIsInternal.lean @@ -60,8 +60,8 @@ theorem withSign.isCompl : IsCompl ℤ≥0 ℤ≤0 := by · rw [codisjoint_iff_le_sup] intro x _hx obtain hp | hn := (le_refl (0 : ℤ)).le_or_le x - exact Submodule.mem_sup_left (mem_withSign_one.mpr hp) - exact Submodule.mem_sup_right (mem_withSign_neg_one.mpr hn) + · exact Submodule.mem_sup_left (mem_withSign_one.mpr hp) + · exact Submodule.mem_sup_right (mem_withSign_neg_one.mpr hn) #align counterexample.with_sign.is_compl Counterexample.withSign.isCompl def withSign.independent : CompleteLattice.Independent withSign := by diff --git a/Mathlib/Algebra/BigOperators/Finprod.lean b/Mathlib/Algebra/BigOperators/Finprod.lean index ed0d74a5c634a8..ff63590fa088ec 100644 --- a/Mathlib/Algebra/BigOperators/Finprod.lean +++ b/Mathlib/Algebra/BigOperators/Finprod.lean @@ -1033,9 +1033,9 @@ theorem finprod_subtype_eq_finprod_cond (p : α → Prop) : theorem finprod_mem_inter_mul_diff' (t : Set α) (h : (s ∩ mulSupport f).Finite) : ((∏ᶠ i ∈ s ∩ t, f i) * ∏ᶠ i ∈ s \ t, f i) = ∏ᶠ i ∈ s, f i := by rw [← finprod_mem_union', inter_union_diff] - rw [disjoint_iff_inf_le] - exacts [fun x hx => hx.2.2 hx.1.2, h.subset fun x hx => ⟨hx.1.1, hx.2⟩, - h.subset fun x hx => ⟨hx.1.1, hx.2⟩] + · rw [disjoint_iff_inf_le] + exact fun x hx => hx.2.2 hx.1.2 + exacts [h.subset fun x hx => ⟨hx.1.1, hx.2⟩, h.subset fun x hx => ⟨hx.1.1, hx.2⟩] #align finprod_mem_inter_mul_diff' finprod_mem_inter_mul_diff' #align finsum_mem_inter_add_diff' finsum_mem_inter_add_diff' diff --git a/Mathlib/Algebra/BigOperators/List/Basic.lean b/Mathlib/Algebra/BigOperators/List/Basic.lean index ca246fffc623e8..68e5c7958861d2 100644 --- a/Mathlib/Algebra/BigOperators/List/Basic.lean +++ b/Mathlib/Algebra/BigOperators/List/Basic.lean @@ -106,7 +106,8 @@ theorem prod_cons : (a :: l).prod = a * l.prod := lemma prod_induction (p : M → Prop) (hom : ∀ a b, p a → p b → p (a * b)) (unit : p 1) (base : ∀ x ∈ l, p x) : p l.prod := by - induction' l with a l ih; simpa + induction' l with a l ih + · simpa rw [List.prod_cons] simp only [Bool.not_eq_true, List.mem_cons, forall_eq_or_imp] at base exact hom _ _ (base.1) (ih base.2) diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index be1310396166c4..a14a3e93867ecf 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -42,13 +42,13 @@ def pushoutCocone : Limits.PushoutCocone f g := by letI := RingHom.toAlgebra f letI := RingHom.toAlgebra g fapply Limits.PushoutCocone.mk - show CommRingCat; exact CommRingCat.of (A ⊗[R] B) - show A ⟶ _; exact Algebra.TensorProduct.includeLeftRingHom - show B ⟶ _; exact Algebra.TensorProduct.includeRight.toRingHom - ext r - trans algebraMap R (A ⊗[R] B) r - · exact Algebra.TensorProduct.includeLeft.commutes (R := R) r - · exact (Algebra.TensorProduct.includeRight.commutes (R := R) r).symm + · show CommRingCat; exact CommRingCat.of (A ⊗[R] B) + · show A ⟶ _; exact Algebra.TensorProduct.includeLeftRingHom + · show B ⟶ _; exact Algebra.TensorProduct.includeRight.toRingHom + · ext r + trans algebraMap R (A ⊗[R] B) r + · exact Algebra.TensorProduct.includeLeft.commutes (R := R) r + · exact (Algebra.TensorProduct.includeRight.commutes (R := R) r).symm set_option linter.uppercaseLean3 false in #align CommRing.pushout_cocone CommRingCat.pushoutCocone diff --git a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean index 000ee7afc9b20b..e4a132e16b1c75 100644 --- a/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean +++ b/Mathlib/Algebra/ContinuedFractions/Computation/Approximations.lean @@ -97,7 +97,7 @@ theorem one_le_succ_nth_stream_b {ifp_succ_n : IntFractPair K} obtain ⟨ifp_n, nth_stream_eq, stream_nth_fr_ne_zero, ⟨-⟩⟩ : ∃ ifp_n, IntFractPair.stream v n = some ifp_n ∧ ifp_n.fr ≠ 0 ∧ IntFractPair.of ifp_n.fr⁻¹ = ifp_succ_n - exact succ_nth_stream_eq_some_iff.1 succ_nth_stream_eq + · exact succ_nth_stream_eq_some_iff.1 succ_nth_stream_eq suffices 1 ≤ ifp_n.fr⁻¹ by rwa [IntFractPair.of, le_floor, cast_one] suffices ifp_n.fr ≤ 1 by have h : 0 < ifp_n.fr := @@ -474,7 +474,7 @@ theorem abs_sub_convergents_le (not_terminated_at_n : ¬(of v).TerminatedAt n) : obtain ⟨ifp_n, stream_nth_eq, stream_nth_fr_ne_zero, if_of_eq_ifp_succ_n⟩ : ∃ ifp_n, IntFractPair.stream v n = some ifp_n ∧ ifp_n.fr ≠ 0 ∧ IntFractPair.of ifp_n.fr⁻¹ = ifp_succ_n - exact IntFractPair.succ_nth_stream_eq_some_iff.1 succ_nth_stream_eq + · exact IntFractPair.succ_nth_stream_eq_some_iff.1 succ_nth_stream_eq let denom' := conts.b * (pred_conts.b + ifp_n.fr⁻¹ * conts.b) -- now we can use `sub_convergents_eq` to simplify our goal suffices |(-1) ^ n / denom'| ≤ 1 / denom by diff --git a/Mathlib/Algebra/GCDMonoid/Finset.lean b/Mathlib/Algebra/GCDMonoid/Finset.lean index a82fd399d7ce74..670996b1c53912 100644 --- a/Mathlib/Algebra/GCDMonoid/Finset.lean +++ b/Mathlib/Algebra/GCDMonoid/Finset.lean @@ -275,7 +275,7 @@ theorem extract_gcd (f : β → α) (hs : s.Nonempty) : push_neg at h refine' ⟨fun b ↦ if hb : b ∈ s then g' hb else 0, fun b hb ↦ _, extract_gcd' f _ h fun b hb ↦ _⟩ - simp only [hb, hg, dite_true] + · simp only [hb, hg, dite_true] rw [dif_pos hb, hg hb] #align finset.extract_gcd Finset.extract_gcd diff --git a/Mathlib/Algebra/GeomSum.lean b/Mathlib/Algebra/GeomSum.lean index 5c8e5a7884de92..dd3b2eba74bd6d 100644 --- a/Mathlib/Algebra/GeomSum.lean +++ b/Mathlib/Algebra/GeomSum.lean @@ -507,8 +507,8 @@ theorem geom_sum_alternating_of_lt_neg_one [StrictOrderedRing α] (hx : x + 1 < by_cases hn' : Even n · rw [if_pos hn'] at ihn rw [if_neg, lt_add_iff_pos_left] - exact mul_pos_of_neg_of_neg hx0 ihn - exact not_not_intro hn' + · exact mul_pos_of_neg_of_neg hx0 ihn + · exact not_not_intro hn' · rw [if_neg hn'] at ihn rw [if_pos] swap diff --git a/Mathlib/Algebra/GroupPower/Order.lean b/Mathlib/Algebra/GroupPower/Order.lean index 1cfc9cd3e7a22c..a9390bb31067a2 100644 --- a/Mathlib/Algebra/GroupPower/Order.lean +++ b/Mathlib/Algebra/GroupPower/Order.lean @@ -436,7 +436,11 @@ protected lemma Even.add_pow_le (hn : ∃ k, 2 * k = n) : rw [Commute.mul_pow]; simp [Commute, SemiconjBy, two_mul, mul_two] _ ≤ 2 ^ n * (2 ^ (n - 1) * ((a ^ 2) ^ n + (b ^ 2) ^ n)) := mul_le_mul_of_nonneg_left (add_pow_le (sq_nonneg _) (sq_nonneg _) _) $ pow_nonneg (zero_le_two (α := R)) _ - _ = _ := by simp only [← mul_assoc, ← pow_add, ← pow_mul]; cases n; rfl; simp [Nat.two_mul] + _ = _ := by + simp only [← mul_assoc, ← pow_add, ← pow_mul] + cases n + · rfl + · simp [Nat.two_mul] end LinearOrderedSemiring diff --git a/Mathlib/Algebra/Lie/DirectSum.lean b/Mathlib/Algebra/Lie/DirectSum.lean index a836614e875578..73a2aeb32b1f52 100644 --- a/Mathlib/Algebra/Lie/DirectSum.lean +++ b/Mathlib/Algebra/Lie/DirectSum.lean @@ -166,16 +166,16 @@ def lieAlgebraOf [DecidableEq ι] (j : ι) : L j →ₗ⁅R⁆ ⨁ i, L i := -- with `simp [of, singleAddHom]` simp only [of, singleAddHom, bracket_apply] erw [AddHom.coe_mk, single_apply, single_apply] - simp? [h] says simp only [h, ↓reduceDite, single_apply] - intros - erw [single_add] + · simp? [h] says simp only [h, ↓reduceDite, single_apply] + · intros + erw [single_add] · -- This used to be the end of the proof before leanprover/lean4#2644 -- with `simp [of, singleAddHom]` simp only [of, singleAddHom, bracket_apply] erw [AddHom.coe_mk, single_apply, single_apply] - simp only [h, dite_false, single_apply, lie_self] - intros - erw [single_add] } + · simp only [h, dite_false, single_apply, lie_self] + · intros + erw [single_add] } #align direct_sum.lie_algebra_of DirectSum.lieAlgebraOf /-- The projection map onto one component, as a morphism of Lie algebras. -/ diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index ebc254164bfff0..37d82c6a6346d7 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -329,7 +329,8 @@ theorem comap_bracket_incl {I₁ I₂ : LieIdeal R L} : next => skip rw [← I.incl_idealRange] rw [comap_bracket_eq] - simp only [ker_incl, sup_bot_eq]; exact I.incl_isIdealMorphism + · simp only [ker_incl, sup_bot_eq] + · exact I.incl_isIdealMorphism #align lie_ideal.comap_bracket_incl LieIdeal.comap_bracket_incl /-- This is a very useful result; it allows us to use the fact that inclusion distributes over the diff --git a/Mathlib/Algebra/Module/LinearMap/End.lean b/Mathlib/Algebra/Module/LinearMap/End.lean index d71f6d6988a66f..3d0cb3907bf063 100644 --- a/Mathlib/Algebra/Module/LinearMap/End.lean +++ b/Mathlib/Algebra/Module/LinearMap/End.lean @@ -373,7 +373,8 @@ lemma zero_smulRight (x : M) : (0 : M₁ →ₗ[R] S).smulRight x = 0 := by ext; @[simp] lemma smulRight_apply_eq_zero_iff {f : M₁ →ₗ[R] S} {x : M} [NoZeroSMulDivisors S M] : f.smulRight x = 0 ↔ f = 0 ∨ x = 0 := by - rcases eq_or_ne x 0 with rfl | hx; simp + rcases eq_or_ne x 0 with rfl | hx + · simp refine ⟨fun h ↦ Or.inl ?_, fun h ↦ by simp [h.resolve_right hx]⟩ ext v replace h : f v • x = 0 := by simpa only [LinearMap.zero_apply] using LinearMap.congr_fun h v diff --git a/Mathlib/Algebra/Order/BigOperators/Group/List.lean b/Mathlib/Algebra/Order/BigOperators/Group/List.lean index 668f1a19565312..348ef2d73b143c 100644 --- a/Mathlib/Algebra/Order/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/Order/BigOperators/Group/List.lean @@ -191,7 +191,9 @@ 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]; exact (length l); rfl⟩ + rw [List.eq_replicate.2 ⟨_, h⟩, prod_replicate, one_pow] + · exact (length l) + · rfl⟩ #align list.prod_eq_one_iff List.prod_eq_one_iff #align list.sum_eq_zero_iff List.sum_eq_zero_iff diff --git a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean index 27db9427cb5cdb..a088383a53f4c2 100644 --- a/Mathlib/Algebra/Order/CauSeq/BigOperators.lean +++ b/Mathlib/Algebra/Order/CauSeq/BigOperators.lean @@ -192,8 +192,8 @@ lemma geo_series [Nontrivial β] (x : β) (hx1 : abv x < 1) : refine' @of_mono_bounded _ _ _ _ ((1 : α) / (1 - abv x)) 0 _ _ · intro n _ rw [abs_of_nonneg] - gcongr - · exact sub_le_self _ (abv_pow abv x n ▸ abv_nonneg _ _) + · gcongr + exact sub_le_self _ (abv_pow abv x n ▸ abv_nonneg _ _) refine' div_nonneg (sub_nonneg.2 _) (sub_nonneg.2 <| le_of_lt hx1) exact pow_le_one _ (by positivity) hx1.le · intro n _ diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 6d95439da95c50..40ff2dda9869b4 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -145,8 +145,8 @@ theorem inv_pos_lt_iff_one_lt_mul' (ha : 0 < a) : a⁻¹ < b ↔ 1 < a * b := by /-- One direction of `div_le_iff` where `b` is allowed to be `0` (but `c` must be nonnegative) -/ theorem div_le_of_nonneg_of_le_mul (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c * b) : a / b ≤ c := by rcases eq_or_lt_of_le hb with (rfl | hb') - simp only [div_zero, hc] - rwa [div_le_iff hb'] + · simp only [div_zero, hc] + · rwa [div_le_iff hb'] #align div_le_of_nonneg_of_le_mul div_le_of_nonneg_of_le_mul /-- One direction of `div_le_iff` where `c` is allowed to be `0` (but `b` must be nonnegative) -/ diff --git a/Mathlib/Algebra/Order/Group/Int.lean b/Mathlib/Algebra/Order/Group/Int.lean index 687482049d8a3a..af9caeedccdaec 100644 --- a/Mathlib/Algebra/Order/Group/Int.lean +++ b/Mathlib/Algebra/Order/Group/Int.lean @@ -411,8 +411,8 @@ theorem toNat_le_toNat {a b : ℤ} (h : a ≤ b) : toNat a ≤ toNat b := by #align int.to_nat_le_to_nat Int.toNat_le_toNat theorem toNat_lt_toNat {a b : ℤ} (hb : 0 < b) : toNat a < toNat b ↔ a < b := - ⟨fun h => by cases a; exact lt_toNat.1 h; exact lt_trans (neg_of_sign_eq_neg_one rfl) hb, - fun h => by rw [lt_toNat]; cases a; exact h; exact hb⟩ + ⟨fun h => by cases a; exacts [lt_toNat.1 h, lt_trans (neg_of_sign_eq_neg_one rfl) hb], + fun h => by rw [lt_toNat]; cases a; exacts [h, hb]⟩ #align int.to_nat_lt_to_nat Int.toNat_lt_toNat theorem lt_of_toNat_lt {a b : ℤ} (h : toNat a < toNat b) : a < b := diff --git a/Mathlib/Algebra/Order/Group/MinMax.lean b/Mathlib/Algebra/Order/Group/MinMax.lean index d0f9505ef58441..cec6178fa29833 100644 --- a/Mathlib/Algebra/Order/Group/MinMax.lean +++ b/Mathlib/Algebra/Order/Group/MinMax.lean @@ -85,10 +85,10 @@ variable {α : Type*} [LinearOrderedAddCommGroup α] {a b c : α} theorem max_sub_max_le_max (a b c d : α) : max a b - max c d ≤ max (a - c) (b - d) := by simp only [sub_le_iff_le_add, max_le_iff]; constructor - calc + · calc a = a - c + c := (sub_add_cancel a c).symm _ ≤ max (a - c) (b - d) + max c d := add_le_add (le_max_left _ _) (le_max_left _ _) - calc + · calc b = b - d + d := (sub_add_cancel b d).symm _ ≤ max (a - c) (b - d) + max c d := add_le_add (le_max_right _ _) (le_max_right _ _) #align max_sub_max_le_max max_sub_max_le_max diff --git a/Mathlib/Algebra/Order/Interval/Basic.lean b/Mathlib/Algebra/Order/Interval/Basic.lean index ee4fe91be37ca1..e2ed8ffb17da5f 100644 --- a/Mathlib/Algebra/Order/Interval/Basic.lean +++ b/Mathlib/Algebra/Order/Interval/Basic.lean @@ -711,8 +711,8 @@ theorem length_sum_le (f : ι → Interval α) (s : Finset ι) : (∑ i in s, f i).length ≤ ∑ i in s, (f i).length := by -- Porting note: Old proof was `:= Finset.le_sum_of_subadditive _ length_zero length_add_le _ _` apply Finset.le_sum_of_subadditive - exact length_zero - exact length_add_le + · exact length_zero + · exact length_add_le #align interval.length_sum_le Interval.length_sum_le end Interval diff --git a/Mathlib/Algebra/Order/Pointwise.lean b/Mathlib/Algebra/Order/Pointwise.lean index e840b18d06cf62..53be4c751b29d5 100644 --- a/Mathlib/Algebra/Order/Pointwise.lean +++ b/Mathlib/Algebra/Order/Pointwise.lean @@ -186,8 +186,8 @@ theorem smul_Ioo : r • Ioo a b = Ioo (r • a) (r • b) := by constructor · rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩ constructor - exact (mul_lt_mul_left hr).mpr a_h_left_left - exact (mul_lt_mul_left hr).mpr a_h_left_right + · exact (mul_lt_mul_left hr).mpr a_h_left_left + · exact (mul_lt_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r refine' ⟨⟨(lt_div_iff' hr).mpr a_left, (div_lt_iff' hr).mpr a_right⟩, _⟩ @@ -200,8 +200,8 @@ theorem smul_Icc : r • Icc a b = Icc (r • a) (r • b) := by constructor · rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩ constructor - exact (mul_le_mul_left hr).mpr a_h_left_left - exact (mul_le_mul_left hr).mpr a_h_left_right + · exact (mul_le_mul_left hr).mpr a_h_left_left + · exact (mul_le_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r refine' ⟨⟨(le_div_iff' hr).mpr a_left, (div_le_iff' hr).mpr a_right⟩, _⟩ @@ -214,8 +214,8 @@ theorem smul_Ico : r • Ico a b = Ico (r • a) (r • b) := by constructor · rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩ constructor - exact (mul_le_mul_left hr).mpr a_h_left_left - exact (mul_lt_mul_left hr).mpr a_h_left_right + · exact (mul_le_mul_left hr).mpr a_h_left_left + · exact (mul_lt_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r refine' ⟨⟨(le_div_iff' hr).mpr a_left, (div_lt_iff' hr).mpr a_right⟩, _⟩ @@ -228,8 +228,8 @@ theorem smul_Ioc : r • Ioc a b = Ioc (r • a) (r • b) := by constructor · rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩ constructor - exact (mul_lt_mul_left hr).mpr a_h_left_left - exact (mul_le_mul_left hr).mpr a_h_left_right + · exact (mul_lt_mul_left hr).mpr a_h_left_left + · exact (mul_le_mul_left hr).mpr a_h_left_right · rintro ⟨a_left, a_right⟩ use x / r refine' ⟨⟨(lt_div_iff' hr).mpr a_left, (div_le_iff' hr).mpr a_right⟩, _⟩ @@ -245,8 +245,8 @@ theorem smul_Ioi : r • Ioi a = Ioi (r • a) := by · rintro h use x / r constructor - exact (lt_div_iff' hr).mpr h - exact mul_div_cancel₀ _ (ne_of_gt hr) + · exact (lt_div_iff' hr).mpr h + · exact mul_div_cancel₀ _ (ne_of_gt hr) #align linear_ordered_field.smul_Ioi LinearOrderedField.smul_Ioi theorem smul_Iio : r • Iio a = Iio (r • a) := by @@ -258,8 +258,8 @@ theorem smul_Iio : r • Iio a = Iio (r • a) := by · rintro h use x / r constructor - exact (div_lt_iff' hr).mpr h - exact mul_div_cancel₀ _ (ne_of_gt hr) + · exact (div_lt_iff' hr).mpr h + · exact mul_div_cancel₀ _ (ne_of_gt hr) #align linear_ordered_field.smul_Iio LinearOrderedField.smul_Iio theorem smul_Ici : r • Ici a = Ici (r • a) := by @@ -271,8 +271,8 @@ theorem smul_Ici : r • Ici a = Ici (r • a) := by · rintro h use x / r constructor - exact (le_div_iff' hr).mpr h - exact mul_div_cancel₀ _ (ne_of_gt hr) + · exact (le_div_iff' hr).mpr h + · exact mul_div_cancel₀ _ (ne_of_gt hr) #align linear_ordered_field.smul_Ici LinearOrderedField.smul_Ici theorem smul_Iic : r • Iic a = Iic (r • a) := by @@ -284,8 +284,8 @@ theorem smul_Iic : r • Iic a = Iic (r • a) := by · rintro h use x / r constructor - exact (div_le_iff' hr).mpr h - exact mul_div_cancel₀ _ (ne_of_gt hr) + · exact (div_le_iff' hr).mpr h + · exact mul_div_cancel₀ _ (ne_of_gt hr) #align linear_ordered_field.smul_Iic LinearOrderedField.smul_Iic end LinearOrderedField diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index cbe68060983a56..cbe464251189d3 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -389,9 +389,9 @@ variable [CovariantClass α α (· + ·) (· ≤ ·)] [ContravariantClass α α theorem add_tsub_add_eq_tsub_right (a c b : α) : a + c - (b + c) = a - b := by refine' add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 <| le_of_add_le_add_right _) - exact c - rw [add_assoc] - exact le_tsub_add + · exact c + · rw [add_assoc] + exact le_tsub_add #align add_tsub_add_eq_tsub_right add_tsub_add_eq_tsub_right theorem add_tsub_add_eq_tsub_left (a b c : α) : a + b - (a + c) = b - c := by diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index c189e75cc0a267..761f43fef45841 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -224,14 +224,14 @@ theorem act_invAct_eq (x y : R) : x ◃ x ◃⁻¹ y = y := theorem left_cancel (x : R) {y y' : R} : x ◃ y = x ◃ y' ↔ y = y' := by constructor - apply (act' x).injective + · apply (act' x).injective rintro rfl rfl #align rack.left_cancel Rack.left_cancel theorem left_cancel_inv (x : R) {y y' : R} : x ◃⁻¹ y = x ◃⁻¹ y' ↔ y = y' := by constructor - apply (act' x).symm.injective + · apply (act' x).symm.injective rintro rfl rfl #align rack.left_cancel_inv Rack.left_cancel_inv @@ -305,11 +305,12 @@ theorem self_invAct_act_eq {x y : R} : (x ◃⁻¹ x) ◃ y = x ◃ y := by #align rack.self_inv_act_act_eq Rack.self_invAct_act_eq theorem self_act_eq_iff_eq {x y : R} : x ◃ x = y ◃ y ↔ x = y := by - constructor; swap; rintro rfl; rfl + constructor; swap + · rintro rfl; rfl intro h trans (x ◃ x) ◃⁻¹ x ◃ x - rw [← left_cancel (x ◃ x), right_inv, self_act_act_eq] - rw [h, ← left_cancel (y ◃ y), right_inv, self_act_act_eq] + · rw [← left_cancel (x ◃ x), right_inv, self_act_act_eq] + · rw [h, ← left_cancel (y ◃ y), right_inv, self_act_act_eq] #align rack.self_act_eq_iff_eq Rack.self_act_eq_iff_eq theorem self_invAct_eq_iff_eq {x y : R} : x ◃⁻¹ x = y ◃⁻¹ y ↔ x = y := by diff --git a/Mathlib/AlgebraicGeometry/Gluing.lean b/Mathlib/AlgebraicGeometry/Gluing.lean index 2017fb9c162d10..4a54b470048643 100644 --- a/Mathlib/AlgebraicGeometry/Gluing.lean +++ b/Mathlib/AlgebraicGeometry/Gluing.lean @@ -129,7 +129,8 @@ def gluedScheme : Scheme := by intro x obtain ⟨i, y, rfl⟩ := D.toLocallyRingedSpaceGlueData.ι_jointly_surjective x refine' ⟨_, _ ≫ D.toLocallyRingedSpaceGlueData.toGlueData.ι i, _⟩ - swap; exact (D.U i).affineCover.map y + swap + · exact (D.U i).affineCover.map y constructor · dsimp [-Set.mem_range] rw [coe_comp, Set.range_comp] @@ -329,8 +330,8 @@ theorem glued_cover_cocycle_snd (x y z : 𝒰.J) : theorem glued_cover_cocycle (x y z : 𝒰.J) : gluedCoverT' 𝒰 x y z ≫ gluedCoverT' 𝒰 y z x ≫ gluedCoverT' 𝒰 z x y = 𝟙 _ := by apply pullback.hom_ext <;> simp_rw [Category.id_comp, Category.assoc] - apply glued_cover_cocycle_fst - apply glued_cover_cocycle_snd + · apply glued_cover_cocycle_fst + · apply glued_cover_cocycle_snd #align algebraic_geometry.Scheme.open_cover.glued_cover_cocycle AlgebraicGeometry.Scheme.OpenCover.glued_cover_cocycle /-- The glue data associated with an open cover. @@ -355,7 +356,7 @@ def gluedCover : Scheme.GlueData.{u} where This is an isomorphism, as witnessed by an `IsIso` instance. -/ def fromGlued : 𝒰.gluedCover.glued ⟶ X := by fapply Multicoequalizer.desc - exact fun x => 𝒰.map x + · exact fun x => 𝒰.map x rintro ⟨x, y⟩ change pullback.fst ≫ _ = ((pullbackSymmetry _ _).hom ≫ pullback.fst) ≫ _ simpa using pullback.condition @@ -453,7 +454,7 @@ def glueMorphisms {Y : Scheme} (f : ∀ x, 𝒰.obj x ⟶ Y) X ⟶ Y := by refine' inv 𝒰.fromGlued ≫ _ fapply Multicoequalizer.desc - exact f + · exact f rintro ⟨i, j⟩ change pullback.fst ≫ f i = (_ ≫ _) ≫ f j erw [pullbackSymmetry_hom_comp_fst] diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index a31ec2e53bd5ce..b2264830ae0146 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -761,8 +761,8 @@ theorem eq_comp_δ_of_not_surjective' {n : ℕ} {Δ : SimplexCategory} (θ : Δ dsimp [σ, δ] erw [Fin.predAbove_of_castSucc_lt _ _ (by rwa [Fin.castSucc_castPred])] rw [Fin.succAbove_of_le_castSucc i _] - erw [Fin.succ_pred] - exact Nat.le_sub_one_of_lt (Fin.lt_iff_val_lt_val.mp h') + · erw [Fin.succ_pred] + · exact Nat.le_sub_one_of_lt (Fin.lt_iff_val_lt_val.mp h') · obtain rfl := le_antisymm (Fin.le_last i) (not_lt.mp h) use θ ≫ σ (Fin.last _) ext x : 3 diff --git a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean index b65fa838e97ddc..d5bbed4a6d21a5 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Basic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Basic.lean @@ -336,7 +336,7 @@ only assumes that `‖f x - f x₀‖ ≤ C * ‖x - x₀‖` in a neighborhood theorem HasFDerivAt.le_of_lip' {f : E → F} {f' : E →L[𝕜] F} {x₀ : E} (hf : HasFDerivAt f f' x₀) {C : ℝ} (hC₀ : 0 ≤ C) (hlip : ∀ᶠ x in 𝓝 x₀, ‖f x - f x₀‖ ≤ C * ‖x - x₀‖) : ‖f'‖ ≤ C := by refine' le_of_forall_pos_le_add fun ε ε0 => opNorm_le_of_nhds_zero _ _ - exact add_nonneg hC₀ ε0.le + · exact add_nonneg hC₀ ε0.le rw [← map_add_left_nhds_zero x₀, eventually_map] at hlip filter_upwards [isLittleO_iff.1 (hasFDerivAt_iff_isLittleO_nhds_zero.1 hf) ε0, hlip] with y hy hyC rw [add_sub_cancel_left] at hyC diff --git a/Mathlib/Analysis/Calculus/ParametricIntegral.lean b/Mathlib/Analysis/Calculus/ParametricIntegral.lean index 88448428a2a12b..5504db4f54ef6e 100644 --- a/Mathlib/Analysis/Calculus/ParametricIntegral.lean +++ b/Mathlib/Analysis/Calculus/ParametricIntegral.lean @@ -129,7 +129,7 @@ theorem hasFDerivAt_integral_of_dominated_loc_of_lip' {F' : α → H →L[𝕜] exact ((hF_meas _ x_in).sub (hF_meas _ x₀_in)).sub (hF'_meas.apply_continuousLinearMap _) · refine mem_of_superset h_ball fun x hx ↦ ?_ apply (h_diff.and h_lipsch).mono - rintro a ⟨-, ha_bound⟩ + on_goal 1 => rintro a ⟨-, ha_bound⟩ show ‖‖x - x₀‖⁻¹ • (F x a - F x₀ a - F' a (x - x₀))‖ ≤ b a + ‖F' a‖ replace ha_bound : ‖F x a - F x₀ a‖ ≤ b a * ‖x - x₀‖ := ha_bound x hx calc diff --git a/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean b/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean index 9ffd1ee20a7389..d1b4f14745ca4f 100644 --- a/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean +++ b/Mathlib/Analysis/Calculus/UniformLimitsDeriv.lean @@ -350,8 +350,8 @@ theorem hasFDerivAt_of_tendstoUniformlyOnFilter [NeBot l] have : 𝓝 (0 : G) = 𝓝 (0 + 0 + 0) := by simp only [add_zero] rw [this] refine' Tendsto.add (Tendsto.add _ _) _ - simp only - · have := difference_quotients_converge_uniformly hf' hf hfg + · simp only + have := difference_quotients_converge_uniformly hf' hf hfg rw [Metric.tendstoUniformlyOnFilter_iff] at this rw [Metric.tendsto_nhds] intro ε hε diff --git a/Mathlib/Analysis/Complex/Polynomial.lean b/Mathlib/Analysis/Complex/Polynomial.lean index 1fc1dcdb358368..4eeb5fb4d7b0c1 100644 --- a/Mathlib/Analysis/Complex/Polynomial.lean +++ b/Mathlib/Analysis/Complex/Polynomial.lean @@ -77,8 +77,8 @@ theorem card_complex_roots_eq_card_real_add_card_not_gal_inv (p : ℚ[X]) : rw [← Finset.card_image_of_injective _ Subtype.coe_injective, ← Finset.card_image_of_injective _ inj] let a : Finset ℂ := ?_ - let b : Finset ℂ := ?_ - let c : Finset ℂ := ?_ + on_goal 1 => let b : Finset ℂ := ?_ + on_goal 1 => let c : Finset ℂ := ?_ -- Porting note: was -- change a.card = b.card + c.card suffices a.card = b.card + c.card by exact this @@ -91,7 +91,10 @@ theorem card_complex_roots_eq_card_real_add_card_not_gal_inv (p : ℚ[X]) : · rintro ⟨w, hw, rfl⟩ exact ⟨by rw [aeval_algHom_apply, hw, AlgHom.map_zero], rfl⟩ · rintro ⟨hz1, hz2⟩ - have key : IsScalarTower.toAlgHom ℚ ℝ ℂ z.re = z := by ext; rfl; rw [hz2]; rfl + have key : IsScalarTower.toAlgHom ℚ ℝ ℂ z.re = z := by + ext + · rfl + · rw [hz2]; rfl exact ⟨z.re, inj (by rwa [← aeval_algHom_apply, key, AlgHom.map_zero]), key⟩ have hc0 : ∀ w : p.rootSet ℂ, galActionHom p ℂ (restrict p ℂ (Complex.conjAe.restrictScalars ℚ)) w = w ↔ diff --git a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean b/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean index 58a537cb9bd3cc..dc30c3024a2847 100644 --- a/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean +++ b/Mathlib/Analysis/NormedSpace/Star/ContinuousFunctionalCalculus/Instances.lean @@ -177,7 +177,7 @@ lemma SpectrumRestricts.nnreal_iff_nnnorm {a : A} {t : ℝ≥0} (ha : IsSelfAdjo have : IsSelfAdjoint (algebraMap ℝ A t - a) := IsSelfAdjoint.algebraMap A (.all (t : ℝ)) |>.sub ha rw [← ENNReal.coe_le_coe, ← IsSelfAdjoint.spectralRadius_eq_nnnorm, ← SpectrumRestricts.spectralRadius_eq (f := Complex.reCLM)] at ht ⊢ - exact SpectrumRestricts.nnreal_iff_spectralRadius_le ht + · exact SpectrumRestricts.nnreal_iff_spectralRadius_le ht all_goals try apply IsSelfAdjoint.spectrumRestricts assumption diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean index 03df556cd7b89b..888520717eabc7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Angle.lean @@ -254,8 +254,8 @@ theorem cos_eq_iff_coe_eq_or_eq_neg {θ ψ : ℝ} : zero_add] · rw [angle_eq_iff_two_pi_dvd_sub, ← coe_neg, angle_eq_iff_two_pi_dvd_sub] rintro (⟨k, H⟩ | ⟨k, H⟩) - rw [← sub_eq_zero, cos_sub_cos, H, mul_assoc 2 π k, mul_div_cancel_left₀ _ (two_ne_zero' ℝ), - mul_comm π _, sin_int_mul_pi, mul_zero] + · rw [← sub_eq_zero, cos_sub_cos, H, mul_assoc 2 π k, mul_div_cancel_left₀ _ (two_ne_zero' ℝ), + mul_comm π _, sin_int_mul_pi, mul_zero] rw [← sub_eq_zero, cos_sub_cos, ← sub_neg_eq_add, H, mul_assoc 2 π k, mul_div_cancel_left₀ _ (two_ne_zero' ℝ), mul_comm π _, sin_int_mul_pi, mul_zero, zero_mul] @@ -276,8 +276,8 @@ theorem sin_eq_iff_coe_eq_or_add_eq_pi {θ ψ : ℝ} : exact h.symm · rw [angle_eq_iff_two_pi_dvd_sub, ← eq_sub_iff_add_eq, ← coe_sub, angle_eq_iff_two_pi_dvd_sub] rintro (⟨k, H⟩ | ⟨k, H⟩) - rw [← sub_eq_zero, sin_sub_sin, H, mul_assoc 2 π k, mul_div_cancel_left₀ _ (two_ne_zero' ℝ), - mul_comm π _, sin_int_mul_pi, mul_zero, zero_mul] + · rw [← sub_eq_zero, sin_sub_sin, H, mul_assoc 2 π k, mul_div_cancel_left₀ _ (two_ne_zero' ℝ), + mul_comm π _, sin_int_mul_pi, mul_zero, zero_mul] have H' : θ + ψ = 2 * k * π + π := by rwa [← sub_add, sub_add_eq_add_sub, sub_eq_iff_eq_add, mul_assoc, mul_comm π _, ← mul_assoc] at H diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean index 1de38626f96b39..826f9e2504f2bd 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/Basic.lean @@ -792,9 +792,9 @@ theorem sin_sq_pi_over_two_pow (n : ℕ) : theorem sin_sq_pi_over_two_pow_succ (n : ℕ) : sin (π / 2 ^ (n + 2)) ^ 2 = 1 / 2 - sqrtTwoAddSeries 0 n / 4 := by rw [sin_sq_pi_over_two_pow, sqrtTwoAddSeries, div_pow, sq_sqrt, add_div, ← sub_sub] - congr - · norm_num - · norm_num + · congr + · norm_num + · norm_num · exact add_nonneg two_pos.le (sqrtTwoAddSeries_zero_nonneg _) #align real.sin_sq_pi_over_two_pow_succ Real.sin_sq_pi_over_two_pow_succ diff --git a/Mathlib/CategoryTheory/Category/Quiv.lean b/Mathlib/CategoryTheory/Category/Quiv.lean index 19e5bdbf1e5ad2..0d1c9fab62fc7a 100644 --- a/Mathlib/CategoryTheory/Category/Quiv.lean +++ b/Mathlib/CategoryTheory/Category/Quiv.lean @@ -81,13 +81,13 @@ def free : Quiv.{v, u} ⥤ Cat.{max u v, u} map_id V := by change (show Paths V ⥤ _ from _) = _ ext; swap - apply eq_conj_eqToHom - rfl + · apply eq_conj_eqToHom + · rfl map_comp {U _ _} F G := by change (show Paths U ⥤ _ from _) = _ ext; swap - apply eq_conj_eqToHom - rfl + · apply eq_conj_eqToHom + · rfl set_option linter.uppercaseLean3 false in #align category_theory.Cat.free CategoryTheory.Cat.free @@ -125,8 +125,8 @@ def adj : Cat.free ⊣ Quiv.forget := homEquiv_naturality_left_symm := fun {V _ _} f g => by change (show Paths V ⥤ _ from _) = _ ext; swap - apply eq_conj_eqToHom - rfl } + · apply eq_conj_eqToHom + · rfl } set_option linter.uppercaseLean3 false in #align category_theory.Quiv.adj CategoryTheory.Quiv.adj diff --git a/Mathlib/CategoryTheory/GlueData.lean b/Mathlib/CategoryTheory/GlueData.lean index 9553ddbc8d95ac..ce8b53420c2bb9 100644 --- a/Mathlib/CategoryTheory/GlueData.lean +++ b/Mathlib/CategoryTheory/GlueData.lean @@ -381,7 +381,7 @@ def vPullbackConeIsLimitOfMap (i j : D.J) [ReflectsLimit (cospan (D.ι i) (D.ι apply hc.ofIsoLimit refine Cones.ext (Iso.refl _) ?_ rintro (_ | _ | _) - change _ = _ ≫ (_ ≫ _) ≫ _ + on_goal 1 => change _ = _ ≫ (_ ≫ _) ≫ _ all_goals change _ = 𝟙 _ ≫ _ ≫ _; aesop_cat set_option linter.uppercaseLean3 false in #align category_theory.glue_data.V_pullback_cone_is_limit_of_map CategoryTheory.GlueData.vPullbackConeIsLimitOfMap diff --git a/Mathlib/CategoryTheory/IsConnected.lean b/Mathlib/CategoryTheory/IsConnected.lean index 000f173d440d43..fb6095d345f6d6 100644 --- a/Mathlib/CategoryTheory/IsConnected.lean +++ b/Mathlib/CategoryTheory/IsConnected.lean @@ -210,13 +210,13 @@ theorem IsConnected.of_induct [Nonempty J] {j₀ : J} instance [hc : IsConnected J] : IsConnected (ULiftHom.{v₂} (ULift.{u₂} J)) := by have : Nonempty (ULiftHom.{v₂} (ULift.{u₂} J)) := by simp [ULiftHom, hc.is_nonempty] apply IsConnected.of_induct - rintro p hj₀ h ⟨j⟩ - let p' : Set J := {j : J | p ⟨j⟩} - have hj₀' : Classical.choice hc.is_nonempty ∈ p' := by - simp only [p', (eq_self p')] - exact hj₀ - apply induct_on_objects p' hj₀' @fun _ _ f => - h ((ULiftHomULiftCategory.equiv J).functor.map f) + · rintro p hj₀ h ⟨j⟩ + let p' : Set J := {j : J | p ⟨j⟩} + have hj₀' : Classical.choice hc.is_nonempty ∈ p' := by + simp only [p', (eq_self p')] + exact hj₀ + apply induct_on_objects p' hj₀' @fun _ _ f => + h ((ULiftHomULiftCategory.equiv J).functor.map f) /-- Another induction principle for `IsPreconnected J`: given a type family `Z : J → Sort*` and diff --git a/Mathlib/CategoryTheory/Limits/Connected.lean b/Mathlib/CategoryTheory/Limits/Connected.lean index e59291da841813..36233f9d18705a 100644 --- a/Mathlib/CategoryTheory/Limits/Connected.lean +++ b/Mathlib/CategoryTheory/Limits/Connected.lean @@ -34,18 +34,18 @@ section Examples instance widePullbackShape_connected (J : Type v₁) : IsConnected (WidePullbackShape J) := by apply IsConnected.of_induct - introv hp t - cases j - · exact hp - · rwa [t (WidePullbackShape.Hom.term _)] + · introv hp t + cases j + · exact hp + · rwa [t (WidePullbackShape.Hom.term _)] #align category_theory.wide_pullback_shape_connected CategoryTheory.widePullbackShape_connected instance widePushoutShape_connected (J : Type v₁) : IsConnected (WidePushoutShape J) := by apply IsConnected.of_induct - introv hp t - cases j - · exact hp - · rwa [← t (WidePushoutShape.Hom.init _)] + · introv hp t + cases j + · exact hp + · rwa [← t (WidePushoutShape.Hom.init _)] #align category_theory.wide_pushout_shape_connected CategoryTheory.widePushoutShape_connected instance parallelPairInhabited : Inhabited WalkingParallelPair := @@ -54,10 +54,10 @@ instance parallelPairInhabited : Inhabited WalkingParallelPair := instance parallel_pair_connected : IsConnected WalkingParallelPair := by apply IsConnected.of_induct - introv _ t - cases j - · rwa [t WalkingParallelPairHom.left] - · assumption + · introv _ t + cases j + · rwa [t WalkingParallelPairHom.left] + · assumption #align category_theory.parallel_pair_connected CategoryTheory.parallel_pair_connected end Examples diff --git a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean index ba346924faf0b3..2e9793a848c707 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/BinaryProducts.lean @@ -143,7 +143,7 @@ def isBinaryCoproductOfIsInitialIsPushout (F : Discrete WalkingPair ⥤ C) (c : apply hc.hom_ext rintro (_ | (_ | _)) <;> simp only [PushoutCocone.mk_ι_app_zero, PushoutCocone.mk_ι_app, Category.assoc] - congr 1 + on_goal 1 => congr 1 exacts [(hc.fac c' WalkingSpan.left).symm, (hc.fac c' WalkingSpan.left).symm, (hc.fac c' WalkingSpan.right).symm] #align is_binary_coproduct_of_is_initial_is_pushout isBinaryCoproductOfIsInitialIsPushout diff --git a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean index 3bad4fa7c6f3d8..e89ff8cc7d416c 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/FiniteProductsOfBinaryProducts.lean @@ -140,7 +140,7 @@ noncomputable def preservesFinOfPreservesBinaryAndTerminal : (isLimitOfHasBinaryProductOfPreservesLimit F _ _) refine' IsLimit.ofIsoLimit this _ apply Cones.ext _ _ - apply Iso.refl _ + · apply Iso.refl _ rintro ⟨j⟩ refine Fin.inductionOn j ?_ ?_ · apply (Category.id_comp _).symm @@ -280,7 +280,7 @@ noncomputable def preservesFinOfPreservesBinaryAndInitial : (isColimitOfHasBinaryCoproductOfPreservesColimit F _ _) refine' IsColimit.ofIsoColimit this _ apply Cocones.ext _ _ - apply Iso.refl _ + · apply Iso.refl _ rintro ⟨j⟩ refine Fin.inductionOn j ?_ ?_ · apply Category.comp_id diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index dfdad42583d7ad..c160f2d85c965d 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -394,7 +394,8 @@ theorem zigzag_of_eqvGen_quot_rel {F : C ⥤ D} {d : D} {f₁ f₂ : ΣX, d ⟶ | rel x y r => obtain ⟨f, w⟩ := r fconstructor - swap; fconstructor + swap + · fconstructor left; fconstructor exact StructuredArrow.homMk f | refl => fconstructor @@ -403,7 +404,8 @@ theorem zigzag_of_eqvGen_quot_rel {F : C ⥤ D} {d : D} {f₁ f₂ : ΣX, d ⟶ exact ih | trans x y z _ _ ih₁ ih₂ => apply Relation.ReflTransGen.trans - exact ih₁; exact ih₂ + · exact ih₁ + · exact ih₂ #align category_theory.functor.final.zigzag_of_eqv_gen_quot_rel CategoryTheory.Functor.Final.zigzag_of_eqvGen_quot_rel end Final diff --git a/Mathlib/CategoryTheory/Limits/Fubini.lean b/Mathlib/CategoryTheory/Limits/Fubini.lean index 8231d8055ca0c4..a94aa117f59d82 100644 --- a/Mathlib/CategoryTheory/Limits/Fubini.lean +++ b/Mathlib/CategoryTheory/Limits/Fubini.lean @@ -482,8 +482,8 @@ noncomputable def colimitIsoColimitCurryCompColim : colimit G ≅ colimit (curry have i : G ≅ uncurry.obj ((@curry J _ K _ C _).obj G) := currying.symm.unitIso.app G haveI : Limits.HasColimit (uncurry.obj ((@curry J _ K _ C _).obj G)) := hasColimitOfIso i.symm trans colimit (uncurry.obj ((@curry J _ K _ C _).obj G)) - apply HasColimit.isoOfNatIso i - exact colimitUncurryIsoColimitCompColim ((@curry J _ K _ C _).obj G) + · apply HasColimit.isoOfNatIso i + · exact colimitUncurryIsoColimitCompColim ((@curry J _ K _ C _).obj G) @[simp, reassoc] theorem colimitIsoColimitCurryCompColim_ι_ι_inv {j} {k} : diff --git a/Mathlib/CategoryTheory/Limits/KanExtension.lean b/Mathlib/CategoryTheory/Limits/KanExtension.lean index fd1ea39c49c993..dd9683f3784ef5 100644 --- a/Mathlib/CategoryTheory/Limits/KanExtension.lean +++ b/Mathlib/CategoryTheory/Limits/KanExtension.lean @@ -118,8 +118,8 @@ def equiv (F : S ⥤ D) [h : ∀ x, HasLimit (diagram ι F x)] (G : L ⥤ D) : let t : StructuredArrow.mk (𝟙 (ι.obj x)) ⟶ (StructuredArrow.map (ι.map ff)).obj (StructuredArrow.mk (𝟙 (ι.obj y))) := StructuredArrow.homMk ff ?_ - convert (limit.w (diagram ι F (ι.obj x)) t).symm using 1 - simp } + · convert (limit.w (diagram ι F (ι.obj x)) t).symm using 1 + · simp } invFun f := { app := fun x => limit.lift (diagram ι F x) (cone _ f) naturality := by diff --git a/Mathlib/CategoryTheory/Limits/Lattice.lean b/Mathlib/CategoryTheory/Limits/Lattice.lean index 9a8eeeed781eae..5ac300b39c5be9 100644 --- a/Mathlib/CategoryTheory/Limits/Lattice.lean +++ b/Mathlib/CategoryTheory/Limits/Lattice.lean @@ -85,9 +85,9 @@ A finite product in the category of a `SemilatticeInf` with `OrderTop` is the sa theorem finite_product_eq_finset_inf [SemilatticeInf α] [OrderTop α] {ι : Type u} [Fintype ι] (f : ι → α) : ∏ f = Fintype.elems.inf f := by trans - exact - (IsLimit.conePointUniqueUpToIso (limit.isLimit _) - (finiteLimitCone (Discrete.functor f)).isLimit).to_eq + · exact + (IsLimit.conePointUniqueUpToIso (limit.isLimit _) + (finiteLimitCone (Discrete.functor f)).isLimit).to_eq change Finset.univ.inf (f ∘ discreteEquiv.toEmbedding) = Fintype.elems.inf f simp only [← Finset.inf_map, Finset.univ_map_equiv_to_embedding] rfl @@ -99,9 +99,9 @@ supremum. theorem finite_coproduct_eq_finset_sup [SemilatticeSup α] [OrderBot α] {ι : Type u} [Fintype ι] (f : ι → α) : ∐ f = Fintype.elems.sup f := by trans - exact - (IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) - (finiteColimitCocone (Discrete.functor f)).isColimit).to_eq + · exact + (IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) + (finiteColimitCocone (Discrete.functor f)).isColimit).to_eq change Finset.univ.sup (f ∘ discreteEquiv.toEmbedding) = Fintype.elems.sup f simp only [← Finset.sup_map, Finset.univ_map_equiv_to_embedding] rfl diff --git a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean index b396695cf3193c..f5beb101d49b16 100644 --- a/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Preserves/Shapes/Products.lean @@ -40,7 +40,7 @@ def isLimitMapConeFanMkEquiv {P : C} (g : ∀ j, P ⟶ f j) : IsLimit (Functor.mapCone G (Fan.mk P g)) ≃ IsLimit (Fan.mk _ fun j => G.map (g j) : Fan fun j => G.obj (f j)) := by refine' (IsLimit.postcomposeHomEquiv _ _).symm.trans (IsLimit.equivIsoLimit _) - refine' Discrete.natIso fun j => Iso.refl (G.obj (f j.as)) + · refine' Discrete.natIso fun j => Iso.refl (G.obj (f j.as)) refine' Cones.ext (Iso.refl _) fun j => by dsimp; cases j; simp #align category_theory.limits.is_limit_map_cone_fan_mk_equiv CategoryTheory.Limits.isLimitMapConeFanMkEquiv @@ -112,7 +112,7 @@ def isColimitMapCoconeCofanMkEquiv {P : C} (g : ∀ j, f j ⟶ P) : IsColimit (Functor.mapCocone G (Cofan.mk P g)) ≃ IsColimit (Cofan.mk _ fun j => G.map (g j) : Cofan fun j => G.obj (f j)) := by refine' (IsColimit.precomposeHomEquiv _ _).symm.trans (IsColimit.equivIsoColimit _) - refine' Discrete.natIso fun j => Iso.refl (G.obj (f j.as)) + · refine' Discrete.natIso fun j => Iso.refl (G.obj (f j.as)) refine' Cocones.ext (Iso.refl _) fun j => by dsimp; cases j; simp #align category_theory.limits.is_colimit_map_cocone_cofan_mk_equiv CategoryTheory.Limits.isColimitMapCoconeCofanMkEquiv diff --git a/Mathlib/CategoryTheory/Limits/Presheaf.lean b/Mathlib/CategoryTheory/Limits/Presheaf.lean index 3ee8bc2ae1970a..9bd6095e13729f 100644 --- a/Mathlib/CategoryTheory/Limits/Presheaf.lean +++ b/Mathlib/CategoryTheory/Limits/Presheaf.lean @@ -245,8 +245,8 @@ noncomputable def extendAlongYonedaIsoKanApp (X) : hom_inv_id := by erw [colimit.pre_pre ((CategoryOfElements.π X).leftOp ⋙ A) eq.inverse] trans colimit.pre ((CategoryOfElements.π X).leftOp ⋙ A) (𝟭 _) - congr - · exact congr_arg Functor.op (CategoryOfElements.from_toCostructuredArrow_eq X) + · congr + exact congr_arg Functor.op (CategoryOfElements.from_toCostructuredArrow_eq X) · ext simp only [colimit.ι_pre] erw [Category.comp_id] @@ -254,8 +254,8 @@ noncomputable def extendAlongYonedaIsoKanApp (X) : inv_hom_id := by erw [colimit.pre_pre (Lan.diagram (yoneda : C ⥤ _ ⥤ Type u₁) A X) eq.functor] trans colimit.pre (Lan.diagram (yoneda : C ⥤ _ ⥤ Type u₁) A X) (𝟭 _) - congr - · exact CategoryOfElements.to_fromCostructuredArrow_eq X + · congr + exact CategoryOfElements.to_fromCostructuredArrow_eq X · ext simp only [colimit.ι_pre] erw [Category.comp_id] diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean index d3c59bf379e7e3..9bdbf3e940aead 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Biproducts.lean @@ -968,8 +968,8 @@ def cokernelCoforkBiproductFromSubtype (p : Set K) : simp only [Category.assoc, Pi.compl_apply, biproduct.ι_fromSubtype_assoc, biproduct.ι_toSubtype_assoc, comp_zero, zero_comp] rw [dif_neg] - simp only [zero_comp] - exact not_not.mpr k.2) + · simp only [zero_comp] + · exact not_not.mpr k.2) isColimit := CokernelCofork.IsColimit.ofπ _ _ (fun {W} g _ => biproduct.fromSubtype f pᶜ ≫ g) (by diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean index 6804da1fa1b157..930cc129ae891a 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Equalizers.lean @@ -360,7 +360,10 @@ theorem Cofork.app_zero_eq_comp_π_right (s : Cofork f g) : s.ι.app zero = g def Fork.ofι {P : C} (ι : P ⟶ X) (w : ι ≫ f = ι ≫ g) : Fork f g where pt := P π := - { app := fun X => by cases X; exact ι; exact ι ≫ f + { app := fun X => by + cases X + · exact ι + · exact ι ≫ f naturality := fun {X} {Y} f => by cases X <;> cases Y <;> cases f <;> dsimp <;> simp; assumption } #align category_theory.limits.fork.of_ι CategoryTheory.Limits.Fork.ofι diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean index 06457123edeb04..be6f44ef447a9f 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Pullbacks.lean @@ -1962,7 +1962,9 @@ theorem fst_eq_snd_of_mono_eq [Mono f] : (pullback.fst : pullback f f ⟶ _) = p @[simp] theorem pullbackSymmetry_hom_of_mono_eq [Mono f] : (pullbackSymmetry f f).hom = 𝟙 _ := by - ext; simp [fst_eq_snd_of_mono_eq]; simp [fst_eq_snd_of_mono_eq] + ext + · simp [fst_eq_snd_of_mono_eq] + · simp [fst_eq_snd_of_mono_eq] #align category_theory.limits.pullback_symmetry_hom_of_mono_eq CategoryTheory.Limits.pullbackSymmetry_hom_of_mono_eq instance fst_iso_of_mono_eq [Mono f] : IsIso (pullback.fst : pullback f f ⟶ _) := by @@ -2346,7 +2348,7 @@ local notation "l₂'" => (pullback.snd : W' ⟶ Z₂) def pullbackPullbackLeftIsPullback [HasPullback (g₂ ≫ f₃) f₄] : IsLimit (PullbackCone.mk l₁ l₂ (show l₁ ≫ g₂ = l₂ ≫ g₃ from (pullback.lift_fst _ _ _).symm)) := by apply leftSquareIsPullback - exact pullbackIsPullback f₃ f₄ + · exact pullbackIsPullback f₃ f₄ convert pullbackIsPullback (g₂ ≫ f₃) f₄ rw [pullback.lift_snd] #align category_theory.limits.pullback_pullback_left_is_pullback CategoryTheory.Limits.pullbackPullbackLeftIsPullback diff --git a/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean b/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean index 90b889e3388314..5168baa185fa41 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/SingleObj.lean @@ -41,8 +41,8 @@ instance (J : SingleObj M ⥤ Type u) : MulAction M (J.obj (SingleObj.star M)) w mul_smul g h x := by show J.map (g * h) x = (J.map h ≫ J.map g) x rw [← SingleObj.comp_as_mul] - simp only [FunctorToTypes.map_comp_apply, types_comp_apply] - rfl + · simp only [FunctorToTypes.map_comp_apply, types_comp_apply] + rfl section Limits diff --git a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean index ed7266bade0e82..d994d30869cefe 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/WidePullbacks.lean @@ -69,7 +69,7 @@ instance struct : CategoryStruct (WidePullbackShape J) where id j := Hom.id j comp f g := by cases f - exact g + · exact g cases g apply Hom.term _ #align category_theory.limits.wide_pullback_shape.struct CategoryTheory.Limits.WidePullbackShape.struct @@ -93,7 +93,9 @@ instance subsingleton_hom : Quiver.IsThin (WidePullbackShape J) := fun _ _ => by constructor intro a b casesm* WidePullbackShape _, (_: WidePullbackShape _) ⟶ (_ : WidePullbackShape _) - rfl; rfl; rfl + · rfl + · rfl + · rfl #align category_theory.limits.wide_pullback_shape.subsingleton_hom CategoryTheory.Limits.WidePullbackShape.subsingleton_hom instance category : SmallCategory (WidePullbackShape J) := @@ -186,7 +188,7 @@ instance struct : CategoryStruct (WidePushoutShape J) where id j := Hom.id j comp f g := by cases f - exact g + · exact g cases g apply Hom.init _ #align category_theory.limits.wide_pushout_shape.struct CategoryTheory.Limits.WidePushoutShape.struct @@ -378,8 +380,8 @@ theorem eq_lift_of_comp_eq (g : X ⟶ widePullback _ _ arrows) : theorem hom_eq_lift (g : X ⟶ widePullback _ _ arrows) : g = lift (g ≫ base arrows) (fun j => g ≫ π arrows j) (by aesop_cat) := by apply eq_lift_of_comp_eq - aesop_cat - rfl -- Porting note: quite a few missing refl's in aesop_cat now + · aesop_cat + · rfl -- Porting note: quite a few missing refl's in aesop_cat now #align category_theory.limits.wide_pullback.hom_eq_lift CategoryTheory.Limits.WidePullback.hom_eq_lift @[ext 1100] @@ -457,8 +459,8 @@ theorem hom_eq_desc (g : widePushout _ _ arrows ⟶ X) : rw [← Category.assoc] simp := by apply eq_desc_of_comp_eq - aesop_cat - rfl -- Porting note: another missing rfl + · aesop_cat + · rfl -- Porting note: another missing rfl #align category_theory.limits.wide_pushout.hom_eq_desc CategoryTheory.Limits.WidePushout.hom_eq_desc @[ext 1100] diff --git a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean index be95547cb2178d..cd97d90a6e045d 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/ZeroMorphisms.lean @@ -517,12 +517,12 @@ def isIsoZeroEquivIsoZero (X Y : C) : IsIso (0 : X ⟶ Y) ≃ (X ≅ 0) × (Y fconstructor · rintro ⟨eX, eY⟩ fconstructor - exact (idZeroEquivIsoZero X).symm eX - exact (idZeroEquivIsoZero Y).symm eY + · exact (idZeroEquivIsoZero X).symm eX + · exact (idZeroEquivIsoZero Y).symm eY · rintro ⟨hX, hY⟩ fconstructor - exact (idZeroEquivIsoZero X) hX - exact (idZeroEquivIsoZero Y) hY + · exact (idZeroEquivIsoZero X) hX + · exact (idZeroEquivIsoZero Y) hY · aesop_cat · aesop_cat #align category_theory.limits.is_iso_zero_equiv_iso_zero CategoryTheory.Limits.isIsoZeroEquivIsoZero diff --git a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean index 883754fdd9f4d8..ccd1cfaf9961ec 100644 --- a/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/OfChosenFiniteProducts/Basic.lean @@ -147,15 +147,15 @@ def IsLimit.assoc {X Y Z : C} {sXY : BinaryFan X Y} (P : IsLimit sXY) {sYZ : Bin have h := R.uniq (BinaryFan.assocInv P t) m rw [h] rintro ⟨⟨⟩⟩ <;> simp - apply P.hom_ext - rintro ⟨⟨⟩⟩ <;> simp - · exact w ⟨WalkingPair.left⟩ - · specialize w ⟨WalkingPair.right⟩ - simp? at w says - simp only [pair_obj_right, BinaryFan.π_app_right, BinaryFan.assoc_snd, - Functor.const_obj_obj, pair_obj_left] at w - rw [← w] - simp + · apply P.hom_ext + rintro ⟨⟨⟩⟩ <;> simp + · exact w ⟨WalkingPair.left⟩ + · specialize w ⟨WalkingPair.right⟩ + simp? at w says + simp only [pair_obj_right, BinaryFan.π_app_right, BinaryFan.assoc_snd, + Functor.const_obj_obj, pair_obj_left] at w + rw [← w] + simp · specialize w ⟨WalkingPair.right⟩ simp? at w says simp only [pair_obj_right, BinaryFan.π_app_right, BinaryFan.assoc_snd, diff --git a/Mathlib/CategoryTheory/Sites/Canonical.lean b/Mathlib/CategoryTheory/Sites/Canonical.lean index a99b2107ed1e71..77a8fb9dd291b3 100644 --- a/Mathlib/CategoryTheory/Sites/Canonical.lean +++ b/Mathlib/CategoryTheory/Sites/Canonical.lean @@ -131,8 +131,8 @@ theorem isSheafFor_trans (P : Cᵒᵖ ⥤ Type v) (R S : Sieve X) rintro Z f ⟨W, f, g, hg, hf : S _, rfl⟩ apply hf apply Presieve.isSheafFor_subsieve_aux P this - apply isSheafFor_bind _ _ _ hR hS - · intro Y f hf Z g + · apply isSheafFor_bind _ _ _ hR hS + intro Y f hf Z g rw [← pullback_comp] apply (hS (R.downward_closed hf _)).isSeparatedFor · intro Y f hf diff --git a/Mathlib/CategoryTheory/Sites/Coverage.lean b/Mathlib/CategoryTheory/Sites/Coverage.lean index 7a950f5bcaba86..95e33ec997222b 100644 --- a/Mathlib/CategoryTheory/Sites/Coverage.lean +++ b/Mathlib/CategoryTheory/Sites/Coverage.lean @@ -232,7 +232,7 @@ def toGrothendieck (K : Coverage C) : GrothendieckTopology C where | top X => apply saturate.top | transitive X R S _ hS H1 _ => apply saturate.transitive - apply H1 f + · apply H1 f intro Z g hg rw [← Sieve.pullback_comp] exact hS hg @@ -376,10 +376,10 @@ theorem isSheaf_coverage (K : Coverage C) (P : Cᵒᵖ ⥤ Type*) : refine ⟨t, fun Z g hg => ?_⟩ refine (H1 (g ≫ f)).ext (fun ZZ gg hgg => ?_) rw [← types_comp_apply _ (P.map gg.op), ← P.map_comp, ← op_comp, ht] - swap; simpa using hgg + on_goal 2 => simpa using hgg refine (H2 hgg (𝟙 _)).ext (fun ZZZ ggg hggg => ?_) rw [← types_comp_apply _ (P.map ggg.op), ← P.map_comp, ← op_comp, hz] - swap; simpa using hggg + on_goal 2 => simpa using hggg refine (H2 hgg ggg).ext (fun ZZZZ gggg _ => ?_) rw [← types_comp_apply _ (P.map gggg.op), ← P.map_comp, ← op_comp] apply hx diff --git a/Mathlib/CategoryTheory/Subobject/MonoOver.lean b/Mathlib/CategoryTheory/Subobject/MonoOver.lean index ed9a7e7165ff29..eebc0b14c2ca73 100644 --- a/Mathlib/CategoryTheory/Subobject/MonoOver.lean +++ b/Mathlib/CategoryTheory/Subobject/MonoOver.lean @@ -367,12 +367,12 @@ def image : Over X ⥤ MonoOver X where map {f g} k := by apply (forget X).preimage _ apply Over.homMk _ _ - exact - image.lift - { I := Limits.image _ - m := image.ι g.hom - e := k.left ≫ factorThruImage g.hom } - apply image.lift_fac + · exact + image.lift + { I := Limits.image _ + m := image.ι g.hom + e := k.left ≫ factorThruImage g.hom } + · apply image.lift_fac #align category_theory.mono_over.image CategoryTheory.MonoOver.image /-- `MonoOver.image : Over X ⥤ MonoOver X` is left adjoint to @@ -388,13 +388,13 @@ def imageForgetAdj : image ⊣ forget X := apply image.fac invFun := fun k => by refine' Over.homMk _ _ - exact - image.lift - { I := g.obj.left - m := g.arrow - e := k.left - fac := Over.w k } - apply image.lift_fac + · exact + image.lift + { I := g.obj.left + m := g.arrow + e := k.left + fac := Over.w k } + · apply image.lift_fac left_inv := fun k => Subsingleton.elim _ _ right_inv := fun k => by ext1 @@ -441,8 +441,8 @@ def existsIsoMap (f : X ⟶ Y) [Mono f] : «exists» f ≅ map f := suffices (forget _).obj ((«exists» f).obj Z) ≅ (forget _).obj ((map f).obj Z) by apply (forget _).preimageIso this apply Over.isoMk _ _ - apply imageMonoIsoSource (Z.arrow ≫ f) - apply imageMonoIsoSource_hom_self) + · apply imageMonoIsoSource (Z.arrow ≫ f) + · apply imageMonoIsoSource_hom_self) #align category_theory.mono_over.exists_iso_map CategoryTheory.MonoOver.existsIsoMap /-- `exists` is adjoint to `pullback` when images exist -/ diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index bd351d85d6c4e7..1374af66714568 100644 --- a/Mathlib/Combinatorics/Enumerative/Composition.lean +++ b/Mathlib/Combinatorics/Enumerative/Composition.lean @@ -388,7 +388,7 @@ theorem disjoint_range {i₁ i₂ : Fin c.length} (h : i₁ ≠ i₂) : Disjoint (Set.range (c.embedding i₁)) (Set.range (c.embedding i₂)) := by classical wlog h' : i₁ < i₂ - exact (this c h.symm (h.lt_or_lt.resolve_left h')).symm + · exact (this c h.symm (h.lt_or_lt.resolve_left h')).symm by_contra d obtain ⟨x, hx₁, hx₂⟩ : ∃ x : Fin n, x ∈ Set.range (c.embedding i₁) ∧ x ∈ Set.range (c.embedding i₂) := @@ -1031,8 +1031,8 @@ theorem CompositionAsSet.toComposition_boundaries (c : CompositionAsSet n) : constructor · rintro ⟨i, _, hi⟩ refine' ⟨i.1, _, _⟩ - simpa [c.card_boundaries_eq_succ_length] using i.2 - simp [Composition.boundary, Composition.sizeUpTo, ← hi] + · simpa [c.card_boundaries_eq_succ_length] using i.2 + · simp [Composition.boundary, Composition.sizeUpTo, ← hi] · rintro ⟨i, i_lt, hi⟩ refine' ⟨i, by simp, _⟩ rw [c.card_boundaries_eq_succ_length] at i_lt diff --git a/Mathlib/Combinatorics/Quiver/Arborescence.lean b/Mathlib/Combinatorics/Quiver/Arborescence.lean index 4487233d713f6b..c55d68f4bd58c9 100644 --- a/Mathlib/Combinatorics/Quiver/Arborescence.lean +++ b/Mathlib/Combinatorics/Quiver/Arborescence.lean @@ -75,8 +75,8 @@ noncomputable def arborescenceMk {V : Type u} [Quiver V] (r : V) (height : V → have height_le : ∀ {a b}, Path a b → height a ≤ height b := by intro a b p induction' p with b c _ e ih - rfl - exact le_of_lt (lt_of_le_of_lt ih (height_lt e)) + · rfl + · exact le_of_lt (lt_of_le_of_lt ih (height_lt e)) suffices ∀ p q : Path r b, p = q by intro p apply this diff --git a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean index 253d4c7950cd28..67c9691c423829 100644 --- a/Mathlib/Combinatorics/SetFamily/FourFunctions.lean +++ b/Mathlib/Combinatorics/SetFamily/FourFunctions.lean @@ -286,8 +286,8 @@ lemma four_functions_theorem [DecidableEq α] (h₁ : 0 ≤ f₁) (h₂ : 0 ≤ (extend g (f₃ ∘ (↑)) 0) (extend g (f₄ ∘ (↑)) 0) (extend_nonneg (fun _ ↦ h₁ _) le_rfl) (extend_nonneg (fun _ ↦ h₂ _) le_rfl) (extend_nonneg (fun _ ↦ h₃ _) le_rfl) (extend_nonneg (fun _ ↦ h₄ _) le_rfl) ?_ (s'.map ⟨g, hg⟩) (t'.map ⟨g, hg⟩) - simpa only [← hs', ← ht', ← map_sups, ← map_infs, sum_map, Embedding.coeFn_mk, hg.extend_apply] - using this + · simpa only [← hs', ← ht', ← map_sups, ← map_infs, sum_map, Embedding.coeFn_mk, hg.extend_apply] + using this rintro s t classical obtain ⟨a, rfl⟩ | hs := em (∃ a, g a = s) diff --git a/Mathlib/Combinatorics/SetFamily/LYM.lean b/Mathlib/Combinatorics/SetFamily/LYM.lean index 35c88ec1818838..d201b5cdb40d56 100644 --- a/Mathlib/Combinatorics/SetFamily/LYM.lean +++ b/Mathlib/Combinatorics/SetFamily/LYM.lean @@ -234,8 +234,8 @@ theorem IsAntichain.sperner [Fintype α] {𝒜 : Finset (Finset α)} suffices (∑ r in Iic (Fintype.card α), ((𝒜 # r).card : ℚ) / (Fintype.card α).choose (Fintype.card α / 2)) ≤ 1 by rw [← sum_div, ← Nat.cast_sum, div_le_one] at this - simp only [cast_le] at this - rwa [sum_card_slice] at this + · simp only [cast_le] at this + rwa [sum_card_slice] at this simp only [cast_pos] exact choose_pos (Nat.div_le_self _ _) rw [Iic_eq_Icc, ← Ico_succ_right, bot_eq_zero, Ico_zero_eq_range] diff --git a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean index 5dfdac08c9ba9f..cb51b99272064a 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Acyclic.lean @@ -181,8 +181,8 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : · have h1 : ((f a).firstDart <| not_nil_of_ne (by simpa using ha)).snd = b := congrArg (·.snd) h have h3 := congrArg length (hf' _ (((f _).tail _).copy h1 rfl) ?_) - rw [length_copy, ← add_left_inj 1, length_tail_add_one] at h3 - · omega + · rw [length_copy, ← add_left_inj 1, length_tail_add_one] at h3 + omega · simp only [ne_eq, eq_mp_eq_cast, id_eq, isPath_copy] exact (hf _).tail _ case surj => @@ -198,7 +198,7 @@ lemma IsTree.card_edgeFinset [Fintype V] [Fintype G.edgeSet] (hG : G.IsTree) : length_cons, length_nil] at h' simp [Nat.le_zero, Nat.one_ne_zero] at h' rw [← hf' _ (.cons h.symm (f x)) ((cons_isPath_iff _ _).2 ⟨hf _, fun hy => ?contra⟩)] - rfl + · rfl case contra => suffices (f x).takeUntil y hy = .cons h .nil by rw [← take_spec _ hy] at h' diff --git a/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean b/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean index 83096f9a013306..8c85659b32099c 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Connectivity.lean @@ -1150,7 +1150,10 @@ def takeUntil {v w : V} : ∀ (p : G.Walk v w) (u : V), u ∈ p.support → G.Wa if hx : v = u then by subst u; exact Walk.nil else - cons r (takeUntil p u <| by cases h; exact (hx rfl).elim; assumption) + cons r (takeUntil p u <| by + cases h + · exact (hx rfl).elim + · assumption) #align simple_graph.walk.take_until SimpleGraph.Walk.takeUntil /-- Given a vertex in the support of a path, give the path from (and including) that vertex to @@ -1162,7 +1165,10 @@ def dropUntil {v w : V} : ∀ (p : G.Walk v w) (u : V), u ∈ p.support → G.Wa if hx : v = u then by subst u exact cons r p - else dropUntil p u <| by cases h; exact (hx rfl).elim; assumption + else dropUntil p u <| by + cases h + · exact (hx rfl).elim + · assumption #align simple_graph.walk.drop_until SimpleGraph.Walk.dropUntil /-- The `takeUntil` and `dropUntil` functions split a walk into two pieces. @@ -1491,7 +1497,7 @@ theorem length_bypass_le {u v : V} (p : G.Walk u v) : p.bypass.length ≤ p.leng simp only [bypass] split_ifs · trans - apply length_dropUntil_le + · apply length_dropUntil_le rw [length_cons] exact le_add_right ih · rw [length_cons, length_cons] diff --git a/Mathlib/Combinatorics/SimpleGraph/Density.lean b/Mathlib/Combinatorics/SimpleGraph/Density.lean index cd638cb4791716..c7de40464863b7 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Density.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Density.lean @@ -259,8 +259,8 @@ theorem abs_edgeDensity_sub_edgeDensity_le_two_mul (hs : s₂ ⊆ s₁) (ht : t rw [two_mul] refine' (abs_sub _ _).trans (add_le_add (le_trans _ h) (le_trans _ h)) <;> · rw [abs_of_nonneg] - exact mod_cast edgeDensity_le_one r _ _ - exact mod_cast edgeDensity_nonneg r _ _ + · exact mod_cast edgeDensity_le_one r _ _ + · exact mod_cast edgeDensity_nonneg r _ _ #align rel.abs_edge_density_sub_edge_density_le_two_mul Rel.abs_edgeDensity_sub_edgeDensity_le_two_mul end Asymmetric diff --git a/Mathlib/Combinatorics/SimpleGraph/Maps.lean b/Mathlib/Combinatorics/SimpleGraph/Maps.lean index 43aed9ca4d4477..3a473c72d2f411 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Maps.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Maps.lean @@ -321,7 +321,7 @@ def mapSpanningSubgraphs {G G' : SimpleGraph V} (h : G ≤ G') : G →g G' where theorem mapEdgeSet.injective (hinj : Function.Injective f) : Function.Injective f.mapEdgeSet := by rintro ⟨e₁, h₁⟩ ⟨e₂, h₂⟩ dsimp [Hom.mapEdgeSet] - repeat' rw [Subtype.mk_eq_mk] + repeat rw [Subtype.mk_eq_mk] apply Sym2.map.injective hinj #align simple_graph.hom.map_edge_set.injective SimpleGraph.Hom.mapEdgeSet.injective diff --git a/Mathlib/Combinatorics/SimpleGraph/Prod.lean b/Mathlib/Combinatorics/SimpleGraph/Prod.lean index 118f0c67193f13..af740f704f220b 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Prod.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Prod.lean @@ -151,8 +151,8 @@ theorem ofBoxProdLeft_boxProdLeft [DecidableEq β] [DecidableRel G.Adj] {a₁ a | nil => rfl | cons' x y z h w => by rw [Walk.boxProdLeft, map_cons, ofBoxProdLeft, Or.by_cases, dif_pos, ← Walk.boxProdLeft] - simp [ofBoxProdLeft_boxProdLeft] - exact ⟨h, rfl⟩ + · simp [ofBoxProdLeft_boxProdLeft] + · exact ⟨h, rfl⟩ #align simple_graph.walk.of_box_prod_left_box_prod_left SimpleGraph.Walk.ofBoxProdLeft_boxProdLeft set_option autoImplicit true in @@ -163,8 +163,8 @@ theorem ofBoxProdLeft_boxProdRight [DecidableEq α] [DecidableRel G.Adj] {b₁ b | cons' x y z h w => by rw [Walk.boxProdRight, map_cons, ofBoxProdRight, Or.by_cases, dif_pos, ← Walk.boxProdRight] - simp [ofBoxProdLeft_boxProdRight] - exact ⟨h, rfl⟩ + · simp [ofBoxProdLeft_boxProdRight] + · exact ⟨h, rfl⟩ #align simple_graph.walk.of_box_prod_left_box_prod_right SimpleGraph.Walk.ofBoxProdLeft_boxProdRight end Walk diff --git a/Mathlib/Combinatorics/Young/YoungDiagram.lean b/Mathlib/Combinatorics/Young/YoungDiagram.lean index 84e4ca3e2fd41f..cbf3f88e60e632 100644 --- a/Mathlib/Combinatorics/Young/YoungDiagram.lean +++ b/Mathlib/Combinatorics/Young/YoungDiagram.lean @@ -174,10 +174,10 @@ theorem cells_bot : (⊥ : YoungDiagram).cells = ∅ := @[norm_cast] theorem coe_bot : (⊥ : YoungDiagram).cells = (∅ : Set (ℕ × ℕ)) := by refine' Set.eq_of_subset_of_subset _ _ - intros x h - simp? [mem_mk, Finset.coe_empty, Set.mem_empty_iff_false] at h says - simp only [cells_bot, Finset.coe_empty, Set.mem_empty_iff_false] at h - simp only [cells_bot, Finset.coe_empty, Set.empty_subset] + · intros x h + simp? [mem_mk, Finset.coe_empty, Set.mem_empty_iff_false] at h says + simp only [cells_bot, Finset.coe_empty, Set.mem_empty_iff_false] at h + · simp only [cells_bot, Finset.coe_empty, Set.empty_subset] #align young_diagram.coe_bot YoungDiagram.coe_bot @[simp] diff --git a/Mathlib/Computability/RegularExpressions.lean b/Mathlib/Computability/RegularExpressions.lean index 7a3f747767d2a3..9d7537b190c755 100644 --- a/Mathlib/Computability/RegularExpressions.lean +++ b/Mathlib/Computability/RegularExpressions.lean @@ -246,7 +246,7 @@ theorem add_rmatch_iff (P Q : RegularExpression α) (x : List α) : (P + Q).rmatch x ↔ P.rmatch x ∨ Q.rmatch x := by induction' x with _ _ ih generalizing P Q · simp only [rmatch, matchEpsilon, Bool.coe_or_iff] - · repeat' rw [rmatch] + · repeat rw [rmatch] rw [deriv_add] exact ih _ _ #align regular_expression.add_rmatch_iff RegularExpression.add_rmatch_iff @@ -264,7 +264,7 @@ theorem mul_rmatch_iff (P Q : RegularExpression α) (x : List α) : cases' List.append_eq_nil.1 h₁.symm with ht hu subst ht subst hu - repeat' rw [rmatch] at h₂ + repeat rw [rmatch] at h₂ simp [h₂] · rw [rmatch]; simp [deriv] split_ifs with hepsilon diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 1e11a15e8211f7..26716f56a3632c 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -649,7 +649,7 @@ theorem cont_eval_fix {f k v} (fok : Code.Ok f) : obtain ⟨v'', h₁, h₂⟩ := this rw [reaches_eval] at h₂ swap - exact ReflTransGen.single rfl + · exact ReflTransGen.single rfl cases Part.mem_unique h₂ (mem_eval.2 ⟨ReflTransGen.refl, rfl⟩) refine' ⟨v', h₁, _⟩ rw [stepRet] at h @@ -657,7 +657,7 @@ theorem cont_eval_fix {f k v} (fok : Code.Ok f) : by_cases he : v'.headI = 0 <;> simp only [exists_prop, if_pos, if_false, he] <;> intro h · refine' ⟨_, Part.mem_some _, _⟩ rw [reaches_eval] - exact h + · exact h exact ReflTransGen.single rfl · obtain ⟨k₀, v₀, e₀⟩ := stepNormal.is_ret f Cont.halt v'.tail have e₁ := stepNormal_then f Cont.halt (Cont.fix f k) v'.tail @@ -674,7 +674,7 @@ theorem cont_eval_fix {f k v} (fok : Code.Ok f) : · rintro ⟨v', he, hr⟩ rw [reaches_eval] at hr swap - exact ReflTransGen.single rfl + · exact ReflTransGen.single rfl refine' PFun.fixInduction he fun v (he : v' ∈ f.fix.eval v) IH => _ rw [fok, Part.bind_eq_bind, Part.mem_bind_iff] obtain he | ⟨v'', he₁', _⟩ := PFun.mem_fix_iff.1 he @@ -683,7 +683,7 @@ theorem cont_eval_fix {f k v} (fok : Code.Ok f) : refine' ⟨_, he₁, _⟩ rw [reaches_eval] swap - exact ReflTransGen.single rfl + · exact ReflTransGen.single rfl rwa [stepRet, if_pos h] · obtain ⟨v₁, he₁, he₂⟩ := (Part.mem_map_iff _).1 he₁' split_ifs at he₂ with h; cases he₂ @@ -691,7 +691,7 @@ theorem cont_eval_fix {f k v} (fok : Code.Ok f) : refine' ⟨_, he₁, _⟩ rw [reaches_eval] swap - exact ReflTransGen.single rfl + · exact ReflTransGen.single rfl rw [stepRet, if_neg h] exact IH v₁.tail ((Part.mem_map_iff _).2 ⟨_, he₁, if_neg h⟩) #align turing.to_partrec.cont_eval_fix Turing.ToPartrec.cont_eval_fix @@ -701,13 +701,15 @@ theorem code_is_ok (c) : Code.Ok c := by | cons f fs IHf IHfs => rw [Code.eval, IHf] simp only [bind_assoc, Cont.eval, pure_bind]; congr; funext v - rw [reaches_eval]; swap; exact ReflTransGen.single rfl + rw [reaches_eval]; swap + · exact ReflTransGen.single rfl rw [stepRet, IHfs]; congr; funext v' refine' Eq.trans _ (Eq.symm _) <;> try exact reaches_eval (ReflTransGen.single rfl) | comp f g IHf IHg => rw [Code.eval, IHg] simp only [bind_assoc, Cont.eval, pure_bind]; congr; funext v - rw [reaches_eval]; swap; exact ReflTransGen.single rfl + rw [reaches_eval]; swap + · exact ReflTransGen.single rfl rw [stepRet, IHf] | case f g IHf IHg => simp only [Code.eval] @@ -728,13 +730,15 @@ theorem stepRet_eval {k v} : eval step (stepRet k v) = Cfg.halt <$> k.eval v := | cons₁ fs as k IH => rw [Cont.eval, stepRet, code_is_ok] simp only [← bind_pure_comp, bind_assoc]; congr; funext v' - rw [reaches_eval]; swap; exact ReflTransGen.single rfl + rw [reaches_eval]; swap + · exact ReflTransGen.single rfl rw [stepRet, IH, bind_pure_comp] | cons₂ ns k IH => rw [Cont.eval, stepRet]; exact IH | comp f k IH => rw [Cont.eval, stepRet, code_is_ok] simp only [← bind_pure_comp, bind_assoc]; congr; funext v' - rw [reaches_eval]; swap; exact ReflTransGen.single rfl + rw [reaches_eval]; swap + · exact ReflTransGen.single rfl rw [IH, bind_pure_comp] | fix f k IH => rw [Cont.eval, stepRet]; simp only [bind_pure_comp] @@ -1596,7 +1600,9 @@ theorem trNormal_respects (c k v s) : | succ => refine' ⟨_, ⟨none, rfl⟩, head_main_ok.trans succ_ok⟩ | tail => let o : Option Γ' := List.casesOn v none fun _ _ => some Γ'.cons - refine' ⟨_, ⟨o, rfl⟩, _⟩; convert clear_ok _ using 2; simp; rfl; swap + refine' ⟨_, ⟨o, rfl⟩, _⟩; convert clear_ok _ using 2 + · simp; rfl + swap refine' splitAtPred_eq _ _ (trNat v.headI) _ _ (trNat_natEnd _) _ cases v <;> simp [o] | cons f fs IHf _ => @@ -1632,8 +1638,10 @@ theorem tr_ret_respects (k v s) : ∃ b₂, refine' (move₂_ok (by decide) _ (splitAtPred_false _)).trans _; · rfl simp only [TM2.step, Option.mem_def, Option.elim, id_eq, elim_update_main, elim_main, elim_aux, List.append_nil, elim_update_aux] - refine' (move₂_ok (by decide) _ _).trans _; pick_goal 4; · rfl - pick_goal 4; + refine' (move₂_ok (by decide) _ _).trans _ + pick_goal 4 + · rfl + pick_goal 4 · exact splitAtPred_eq _ _ _ (some Γ'.consₗ) _ (fun x h => Bool.decide_false (trList_ne_consₗ _ _ h)) ⟨rfl, rfl⟩ @@ -1726,7 +1734,7 @@ theorem trStmts₁_trans {q q'} : q' ∈ trStmts₁ q → trStmts₁ q' ⊆ trSt induction' q with _ _ _ q q_ih _ _ q q_ih q q_ih _ _ q q_ih q q_ih q q_ih q₁ q₂ q₁_ih q₂_ih _ <;> simp (config := { contextual := true }) only [trStmts₁, Finset.mem_insert, Finset.mem_union, or_imp, Finset.mem_singleton, Finset.Subset.refl, imp_true_iff, true_and_iff] - iterate 4 exact fun h => Finset.Subset.trans (q_ih h) (Finset.subset_insert _ _) + repeat exact fun h => Finset.Subset.trans (q_ih h) (Finset.subset_insert _ _) · simp intro s h x h' simp only [Finset.mem_biUnion, Finset.mem_univ, true_and, Finset.mem_insert] diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 584ee9b788cf01..6c1a6f53e5c610 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -148,8 +148,8 @@ def BlankRel.above {Γ} [Inhabited Γ] {l₁ l₂ : List Γ} (h : BlankRel l₁ refine' if hl : l₁.length ≤ l₂.length then ⟨l₂, Or.elim h id fun h' ↦ _, BlankExtends.refl _⟩ else ⟨l₁, BlankExtends.refl _, Or.elim h (fun h' ↦ _) id⟩ - exact (BlankExtends.refl _).above_of_le h' hl - exact (BlankExtends.refl _).above_of_le h' (le_of_not_ge hl) + · exact (BlankExtends.refl _).above_of_le h' hl + · exact (BlankExtends.refl _).above_of_le h' (le_of_not_ge hl) #align turing.blank_rel.above Turing.BlankRel.above /-- Given two `BlankRel` lists, there exists (constructively) a common meet. -/ @@ -158,8 +158,8 @@ def BlankRel.below {Γ} [Inhabited Γ] {l₁ l₂ : List Γ} (h : BlankRel l₁ refine' if hl : l₁.length ≤ l₂.length then ⟨l₁, BlankExtends.refl _, Or.elim h id fun h' ↦ _⟩ else ⟨l₂, Or.elim h (fun h' ↦ _) id, BlankExtends.refl _⟩ - exact (BlankExtends.refl _).above_of_le h' hl - exact (BlankExtends.refl _).above_of_le h' (le_of_not_ge hl) + · exact (BlankExtends.refl _).above_of_le h' hl + · exact (BlankExtends.refl _).above_of_le h' (le_of_not_ge hl) #align turing.blank_rel.below Turing.BlankRel.below theorem BlankRel.equivalence (Γ) [Inhabited Γ] : Equivalence (@BlankRel Γ _) := @@ -471,7 +471,8 @@ def ListBlank.bind {Γ Γ'} [Inhabited Γ] [Inhabited Γ'] (l : ListBlank Γ) (f apply l.liftOn (fun l ↦ ListBlank.mk (List.bind l f)) rintro l _ ⟨i, rfl⟩; cases' hf with n e; refine' Quotient.sound' (Or.inl ⟨i * n, _⟩) rw [List.append_bind, mul_comm]; congr - induction' i with i IH; rfl + induction' i with i IH + · rfl simp only [IH, e, List.replicate_add, Nat.mul_succ, add_comm, List.replicate_succ, List.cons_bind] #align turing.list_blank.bind Turing.ListBlank.bind @@ -2650,7 +2651,9 @@ theorem tr_respects_aux₁ {k} (o q v) {S : List (Γ k)} {L : ListBlank (∀ k, rw [iterate_succ_apply']; simp only [TM1.step, TM1.stepAux, tr, Tape.mk'_nth_nat, Tape.move_right_n_head, addBottom_nth_snd, Option.mem_def] - rw [stk_nth_val _ hL, List.get?_eq_get]; rfl; rwa [List.length_reverse] + rw [stk_nth_val _ hL, List.get?_eq_get] + · rfl + · rwa [List.length_reverse] #align turing.TM2to1.tr_respects_aux₁ Turing.TM2to1.tr_respects_aux₁ theorem tr_respects_aux₃ {q v} {L : ListBlank (∀ k, Option (Γ k))} (n) : Reaches₀ (TM1.step (tr M)) diff --git a/Mathlib/Control/Fix.lean b/Mathlib/Control/Fix.lean index 5967a6d5a4944c..6ef10f27f6960c 100644 --- a/Mathlib/Control/Fix.lean +++ b/Mathlib/Control/Fix.lean @@ -91,9 +91,9 @@ protected theorem fix_def {x : α} (h' : ∃ i, (Fix.approx f i x).Dom) : congr ext x: 1 rw [assert_neg] - rfl - rw [Nat.zero_add] at _this - simpa only [not_not, Coe] + · rfl + · rw [Nat.zero_add] at _this + simpa only [not_not, Coe] | succ n n_ih => intro x' rw [Fix.approx, WellFounded.fix_eq, fixAux] diff --git a/Mathlib/Control/LawfulFix.lean b/Mathlib/Control/LawfulFix.lean index 1cba5e051aafaa..a5e6dcd587c70a 100644 --- a/Mathlib/Control/LawfulFix.lean +++ b/Mathlib/Control/LawfulFix.lean @@ -148,8 +148,8 @@ theorem fix_eq_ωSup : Part.fix f = ωSup (approxChain f) := by cases' exists_fix_le_approx f x with i hx trans approx f i.succ x · trans - apply hx - apply approx_mono' f + · apply hx + · apply approx_mono' f apply le_ωSup_of_le i.succ dsimp [approx] rfl @@ -166,7 +166,10 @@ theorem fix_le {X : (a : _) → Part <| β a} (hX : f X ≤ X) : Part.fix f ≤ intro i induction i with | zero => dsimp [Fix.approx]; apply bot_le - | succ _ i_ih => trans f X; apply f.monotone i_ih; apply hX + | succ _ i_ih => + trans f X + · apply f.monotone i_ih + · apply hX #align part.fix_le Part.fix_le variable {f} (hc : Continuous f) diff --git a/Mathlib/Data/Analysis/Filter.lean b/Mathlib/Data/Analysis/Filter.lean index a4ea300eb4a54a..f0b5e629a7abbf 100644 --- a/Mathlib/Data/Analysis/Filter.lean +++ b/Mathlib/Data/Analysis/Filter.lean @@ -268,8 +268,8 @@ protected def inf {f g : Filter α} (F : f.Realizer) (G : g.Realizer) : (f ⊓ g constructor · rintro ⟨s, t, h⟩ apply mem_inf_of_inter _ _ h - use s - use t + · use s + · use t · rintro ⟨_, ⟨a, ha⟩, _, ⟨b, hb⟩, rfl⟩ exact ⟨a, b, inter_subset_inter ha hb⟩⟩ #align filter.realizer.inf Filter.Realizer.inf diff --git a/Mathlib/Data/DFinsupp/Basic.lean b/Mathlib/Data/DFinsupp/Basic.lean index 2d03ed52677b34..47b0746030c447 100644 --- a/Mathlib/Data/DFinsupp/Basic.lean +++ b/Mathlib/Data/DFinsupp/Basic.lean @@ -1009,7 +1009,7 @@ theorem add_closure_iUnion_range_single : AddSubmonoid.closure (⋃ i : ι, Set.range (single i : β i → Π₀ i, β i)) = ⊤ := top_unique fun x _ => by apply DFinsupp.induction x - exact AddSubmonoid.zero_mem _ + · exact AddSubmonoid.zero_mem _ exact fun a b f _ _ hf => AddSubmonoid.add_mem _ (AddSubmonoid.subset_closure <| Set.mem_iUnion.2 ⟨a, Set.mem_range_self _⟩) hf @@ -1247,8 +1247,8 @@ theorem erase_def (i : ι) (f : Π₀ i, β i) : f.erase i = mk (f.support.erase theorem support_erase (i : ι) (f : Π₀ i, β i) : (f.erase i).support = f.support.erase i := by ext j by_cases h1 : j = i - simp only [h1, mem_support_toFun, erase_apply, ite_true, ne_eq, not_true, not_not, - Finset.mem_erase, false_and] + · simp only [h1, mem_support_toFun, erase_apply, ite_true, ne_eq, not_true, not_not, + Finset.mem_erase, false_and] by_cases h2 : f j ≠ 0 <;> simp at h2 <;> simp [h1, h2] #align dfinsupp.support_erase DFinsupp.support_erase @@ -1955,8 +1955,8 @@ theorem sumAddHom_apply [∀ i, AddZeroClass (β i)] [∀ (i) (x : β i), Decida intro i _ dsimp only [coe_mk', Subtype.coe_mk] at * split_ifs with h - rfl - rw [not_not.mp h, AddMonoidHom.map_zero] + · rfl + · rw [not_not.mp h, AddMonoidHom.map_zero] #align dfinsupp.sum_add_hom_apply DFinsupp.sumAddHom_apply theorem _root_.dfinsupp_sumAddHom_mem [∀ i, AddZeroClass (β i)] [AddCommMonoid γ] {S : Type*} diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 6eec2447f94b11..54399bcd119166 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1592,8 +1592,8 @@ theorem coe_sub_one {n} (a : Fin (n + 1)) : ↑(a - 1) = if a = 0 then n else a @[simp] theorem lt_sub_one_iff {n : ℕ} {k : Fin (n + 2)} : k < k - 1 ↔ k = 0 := by rcases k with ⟨_ | k, hk⟩ - simp only [zero_eq, zero_eta, zero_sub, lt_iff_val_lt_val, val_zero, coe_neg_one, add_pos_iff, - _root_.zero_lt_one, or_true] + · simp only [zero_eq, zero_eta, zero_sub, lt_iff_val_lt_val, val_zero, coe_neg_one, add_pos_iff, + _root_.zero_lt_one, or_true] have : (k + 1 + (n + 1)) % (n + 2) = k % (n + 2) := by rw [add_right_comm, add_assoc, add_mod_right] simp [lt_iff_val_lt_val, ext_iff, Fin.coe_sub, succ_eq_add_one, this, diff --git a/Mathlib/Data/FinEnum.lean b/Mathlib/Data/FinEnum.lean index 26cc839e3db9fe..39f2a2d9636b75 100644 --- a/Mathlib/Data/FinEnum.lean +++ b/Mathlib/Data/FinEnum.lean @@ -232,7 +232,7 @@ theorem mem_pi {β : α → Type _} [FinEnum α] [∀ a, FinEnum (β a)] (xs : L · ext a ⟨⟩ · exists Pi.cons xs_hd xs_tl (f _ (List.mem_cons_self _ _)) constructor - exact ⟨_, rfl⟩ + · exact ⟨_, rfl⟩ exists Pi.tail f constructor · apply xs_ih diff --git a/Mathlib/Data/Finset/Lattice.lean b/Mathlib/Data/Finset/Lattice.lean index 3bd74496c9c341..2489f2658be2ad 100644 --- a/Mathlib/Data/Finset/Lattice.lean +++ b/Mathlib/Data/Finset/Lattice.lean @@ -608,7 +608,7 @@ theorem inf_sup {κ : ι → Type*} (s : Finset ι) (t : ∀ i, Finset (κ i)) ( -- Porting note: `simpa` doesn't support placeholders in proof terms have := h (fun j hj => if hji : j = i then cast (congr_arg κ hji.symm) a else g _ <| mem_of_mem_insert_of_ne hj hji) (fun j hj => ?_) - simpa only [cast_eq, dif_pos, Function.comp, Subtype.coe_mk, dif_neg, aux] using this + · simpa only [cast_eq, dif_pos, Function.comp, Subtype.coe_mk, dif_neg, aux] using this rw [mem_insert] at hj obtain (rfl | hj) := hj · simpa diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 578167b61d8a16..075033b0467eaa 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -235,7 +235,8 @@ theorem mul_noncommProd_erase [DecidableEq α] (s : Multiset α) {a : α} (h : a simp only [quot_mk_to_coe, mem_coe, coe_erase, noncommProd_coe] at comm h ⊢ suffices ∀ x ∈ l, ∀ y ∈ l, x * y = y * x by rw [List.prod_erase_of_comm h this] intro x hx y hy - rcases eq_or_ne x y with rfl | hxy; rfl + rcases eq_or_ne x y with rfl | hxy + · rfl exact comm hx hy hxy theorem noncommProd_erase_mul [DecidableEq α] (s : Multiset α) {a : α} (h : a ∈ s) (comm) @@ -244,7 +245,8 @@ theorem noncommProd_erase_mul [DecidableEq α] (s : Multiset α) {a : α} (h : a suffices ∀ b ∈ erase s a, Commute a b by rw [← (noncommProd_commute (s.erase a) comm' a this).eq, mul_noncommProd_erase s h comm comm'] intro b hb - rcases eq_or_ne a b with rfl | hab; rfl + rcases eq_or_ne a b with rfl | hab + · rfl exact comm h (mem_of_mem_erase hb) hab end Multiset @@ -365,8 +367,8 @@ theorem noncommProd_map [MonoidHomClass F β γ] (s : Finset α) (f : α → β) theorem noncommProd_eq_pow_card (s : Finset α) (f : α → β) (comm) (m : β) (h : ∀ x ∈ s, f x = m) : s.noncommProd f comm = m ^ s.card := by rw [noncommProd, Multiset.noncommProd_eq_pow_card _ _ m] - simp only [Finset.card_def, Multiset.card_map] - simpa using h + · simp only [Finset.card_def, Multiset.card_map] + · simpa using h #align finset.noncomm_prod_eq_pow_card Finset.noncommProd_eq_pow_card #align finset.noncomm_sum_eq_card_nsmul Finset.noncommSum_eq_card_nsmul @@ -473,7 +475,7 @@ theorem noncommProd_mul_single [Fintype ι] [DecidableEq ι] (x : ∀ i, M i) : · intro j _; dsimp · rw [noncommProd_insert_of_not_mem _ _ _ _ (not_mem_erase _ _), noncommProd_eq_pow_card (univ.erase i), one_pow, mul_one] - simp only [MonoidHom.mulSingle_apply, ne_eq, Pi.mulSingle_eq_same] + · simp only [MonoidHom.mulSingle_apply, ne_eq, Pi.mulSingle_eq_same] · intro j hj simp? at hj says simp only [mem_erase, ne_eq, mem_univ, and_true] at hj simp only [MonoidHom.mulSingle_apply, Pi.mulSingle, Function.update, Eq.ndrec, Pi.one_apply, diff --git a/Mathlib/Data/Finset/Sym.lean b/Mathlib/Data/Finset/Sym.lean index f5bae68c769491..72b2c2c099c390 100644 --- a/Mathlib/Data/Finset/Sym.lean +++ b/Mathlib/Data/Finset/Sym.lean @@ -289,8 +289,9 @@ def symInsertEquiv (h : a ∉ s) : (insert a s).sym n ≃ Σi : Fin (n + 1), s.s left_inv m := Subtype.ext <| m.1.fill_filterNe a right_inv := fun ⟨i, m, hm⟩ ↦ by refine' Function.Injective.sigma_map (Function.injective_id) (fun i ↦ _) _ - exact fun i ↦ Sym α (n - i) - swap; exact Subtype.coe_injective + · exact fun i ↦ Sym α (n - i) + swap + · exact Subtype.coe_injective refine Eq.trans ?_ (Sym.filter_ne_fill a _ ?_) exacts [rfl, h ∘ mem_sym_iff.1 hm a] #align finset.sym_insert_equiv Finset.symInsertEquiv diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 733942d6311ee5..1190dd8cce3edb 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -851,8 +851,8 @@ theorem prod_option_index [AddCommMonoid M] [CommMonoid N] (f : Option α →₀ · simp [some_zero, h_zero] · intro f₁ f₂ h₁ h₂ rw [Finsupp.prod_add_index, h₁, h₂, some_add, Finsupp.prod_add_index] - simp only [h_add, Pi.add_apply, Finsupp.coe_add] - rw [mul_mul_mul_comm] + · simp only [h_add, Pi.add_apply, Finsupp.coe_add] + rw [mul_mul_mul_comm] all_goals simp [h_zero, h_add] · rintro (_ | a) m <;> simp [h_zero, h_add] #align finsupp.prod_option_index Finsupp.prod_option_index @@ -1599,9 +1599,9 @@ theorem mapRange_smul {_ : Monoid R} [AddMonoid M] [DistribMulAction R M] [AddMo [DistribMulAction R N] {f : M → N} {hf : f 0 = 0} (c : R) (v : α →₀ M) (hsmul : ∀ x, f (c • x) = c • f x) : mapRange f hf (c • v) = c • mapRange f hf v := by erw [← mapRange_comp] - have : f ∘ (c • ·) = (c • ·) ∘ f := funext hsmul - simp_rw [this] - apply mapRange_comp + · have : f ∘ (c • ·) = (c • ·) ∘ f := funext hsmul + simp_rw [this] + apply mapRange_comp simp only [Function.comp_apply, smul_zero, hf] #align finsupp.map_range_smul Finsupp.mapRange_smul diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 54309f8e59be48..c2d34d8361684a 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -633,8 +633,8 @@ def erase (a : α) (f : α →₀ M) : α →₀ M where classical rw [mem_erase, mem_support_iff]; dsimp split_ifs with h - exact ⟨fun H _ => H.1 h, fun H => (H rfl).elim⟩ - exact and_iff_right h + · exact ⟨fun H _ => H.1 h, fun H => (H rfl).elim⟩ + · exact and_iff_right h #align finsupp.erase Finsupp.erase @[simp] diff --git a/Mathlib/Data/Fintype/Quotient.lean b/Mathlib/Data/Fintype/Quotient.lean index 24a49d0b77d364..5b5be6c3898674 100644 --- a/Mathlib/Data/Fintype/Quotient.lean +++ b/Mathlib/Data/Fintype/Quotient.lean @@ -32,8 +32,8 @@ def Quotient.finChoiceAux {ι : Type*} [DecidableEq ι] {α : ι → Type*} [S : refine' Quotient.liftOn₂ (f i (List.mem_cons_self _ _)) (Quotient.finChoiceAux l fun j h => f j (List.mem_cons_of_mem _ h)) _ _ - exact fun a l => ⟦fun j h => - if e : j = i then by rw [e]; exact a else l _ ((List.mem_cons.1 h).resolve_left e)⟧ + · exact fun a l => ⟦fun j h => + if e : j = i then by rw [e]; exact a else l _ ((List.mem_cons.1 h).resolve_left e)⟧ refine' fun a₁ l₁ a₂ l₂ h₁ h₂ => Quotient.sound fun j h => _ by_cases e : j = i <;> simp [e] · subst j diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index 6b6a959d81893b..a081698f1de3e1 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -227,8 +227,8 @@ theorem bit1_val (n : ℤ) : bit1 n = 2 * n + 1 := theorem bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by cases b - apply (bit0_val n).trans (add_zero _).symm - apply bit1_val + · apply (bit0_val n).trans (add_zero _).symm + · apply bit1_val #align int.bit_val Int.bit_val theorem bit_decomp (n : ℤ) : bit (bodd n) (div2 n) = n := @@ -539,8 +539,9 @@ theorem shiftLeft_eq_mul_pow : ∀ (m : ℤ) (n : ℕ), m <<< (n : ℤ) = m * (2 theorem shiftRight_eq_div_pow : ∀ (m : ℤ) (n : ℕ), m >>> (n : ℤ) = m / (2 ^ n : ℕ) | (m : ℕ), n => by rw [shiftRight_coe_nat, Nat.shiftRight_eq_div_pow _ _]; simp | -[m+1], n => by - rw [shiftRight_negSucc, negSucc_ediv, Nat.shiftRight_eq_div_pow]; rfl - exact ofNat_lt_ofNat_of_lt (Nat.pow_pos (by decide)) + rw [shiftRight_negSucc, negSucc_ediv, Nat.shiftRight_eq_div_pow] + · rfl + · exact ofNat_lt_ofNat_of_lt (Nat.pow_pos (by decide)) #align int.shiftr_eq_div_pow Int.shiftRight_eq_div_pow theorem one_shiftLeft (n : ℕ) : 1 <<< (n : ℤ) = (2 ^ n : ℕ) := diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 633f18240722fa..c4b01f44a804a8 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -618,8 +618,8 @@ theorem isEmpty_iff_eq_nil {l : List α} : l.isEmpty ↔ l = [] := by cases l <; theorem getLast_cons {a : α} {l : List α} : ∀ h : l ≠ nil, getLast (a :: l) (cons_ne_nil a l) = getLast l h := by induction l <;> intros - contradiction - rfl + · contradiction + · rfl #align list.last_cons List.getLast_cons theorem getLast_append_singleton {a : α} (l : List α) : @@ -802,11 +802,16 @@ theorem mem_of_mem_head? {x : α} {l : List α} (h : x ∈ l.head?) : x ∈ l := @[simp] theorem head!_append [Inhabited α] (t : List α) {s : List α} (h : s ≠ []) : - head! (s ++ t) = head! s := by induction s; contradiction; rfl + head! (s ++ t) = head! s := by + induction s + · contradiction + · rfl #align list.head_append List.head!_append theorem head?_append {s t : List α} {x : α} (h : x ∈ s.head?) : x ∈ (s ++ t).head? := by - cases s; contradiction; exact h + cases s + · contradiction + · exact h #align list.head'_append List.head?_append theorem head?_append_of_ne_nil : @@ -816,7 +821,9 @@ theorem head?_append_of_ne_nil : theorem tail_append_singleton_of_ne_nil {a : α} {l : List α} (h : l ≠ nil) : tail (l ++ [a]) = tail l ++ [a] := by - induction l; contradiction; rw [tail, cons_append, tail] + induction l + · contradiction + · rw [tail, cons_append, tail] #align list.tail_append_singleton_of_ne_nil List.tail_append_singleton_of_ne_nil theorem cons_head?_tail : ∀ {l : List α} {a : α}, a ∈ head? l → a :: tail l = l @@ -991,7 +998,6 @@ theorem bidirectionalRec_cons_append {motive : List α → Sort*} | nil => rfl | cons x xs => simp only [List.cons_append] - congr dsimp only [← List.cons_append] suffices ∀ (ys init : List α) (hinit : init = ys) (last : α) (hlast : last = b), (cons_append a init last diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 9723c484225648..8450979dc9d23c 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -420,8 +420,8 @@ theorem Chain.induction (p : α → Prop) (l : List α) (h : Chain r a l) · rw [chain_cons] at h simp only [mem_cons] rintro _ (rfl | H) - apply carries h.1 (l_ih h.2 hb _ (mem_cons.2 (Or.inl rfl))) - apply l_ih h.2 hb _ (mem_cons.2 H) + · apply carries h.1 (l_ih h.2 hb _ (mem_cons.2 (Or.inl rfl))) + · apply l_ih h.2 hb _ (mem_cons.2 H) #align list.chain.induction List.Chain.induction /-- Given a chain from `a` to `b`, and a predicate true at `b`, if `r x y → p y → p x` then diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 9a6c71dd3e472d..d9868f44735de3 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -187,7 +187,7 @@ theorem next_getLast_cons (h : x ∈ l) (y : α) (h : x ∈ y :: l) (hy : x ≠ · rw [← Option.some_inj] at hk' rw [← get?_eq_get, dropLast_eq_take, get?_take, get?_zero, head?_cons, Option.some_inj] at hk' - exact hy (Eq.symm hk') + · exact hy (Eq.symm hk') rw [Nat.zero_eq, length_cons, Nat.pred_succ] exact length_pos_of_mem (by assumption) suffices k.succ = l.length by simp [this] at hk @@ -199,9 +199,9 @@ theorem next_getLast_cons (h : x ∈ l) (y : α) (h : x ∈ y :: l) (hy : x ≠ ⟨tl.length, by simp⟩ ?_ rw [← Option.some_inj] at hk' rw [← get?_eq_get, dropLast_eq_take, get?_take, get?, get?_eq_get, Option.some_inj] at hk' - rw [hk'] - simp only [getLast_eq_get, length_cons, ge_iff_le, Nat.succ_sub_succ_eq_sub, - nonpos_iff_eq_zero, add_eq_zero_iff, and_false, tsub_zero, get_cons_succ] + · rw [hk'] + simp only [getLast_eq_get, length_cons, ge_iff_le, Nat.succ_sub_succ_eq_sub, + nonpos_iff_eq_zero, add_eq_zero_iff, and_false, tsub_zero, get_cons_succ] simpa using hk #align list.next_last_cons List.next_getLast_cons @@ -289,14 +289,14 @@ theorem next_get : ∀ (l : List α) (_h : Nodup l) (i : Fin l.length), · simp [getLast_eq_get] · exact hn.of_cons · rw [next_ne_head_ne_getLast _ _ _ _ _ hx'] - simp only [get_cons_succ] - rw [next_get (y::l), ← get_cons_succ (a := x)] - congr - dsimp - rw [Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 hi'), - Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 (Nat.succ_lt_succ_iff.2 hi'))] - · simp [Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 hi'), Nat.succ_eq_add_one, hi'] - · exact hn.of_cons + · simp only [get_cons_succ] + rw [next_get (y::l), ← get_cons_succ (a := x)] + · congr + dsimp + rw [Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 hi'), + Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 (Nat.succ_lt_succ_iff.2 hi'))] + · simp [Nat.mod_eq_of_lt (Nat.succ_lt_succ_iff.2 hi'), Nat.succ_eq_add_one, hi'] + · exact hn.of_cons · rw [getLast_eq_get] intro h have := nodup_iff_injective_get.1 hn h @@ -956,7 +956,7 @@ theorem chain_map {β : Type*} {r : α → α → Prop} (f : β → α) {s : Cyc Chain r (s.map f) ↔ Chain (fun a b => r (f a) (f b)) s := Quotient.inductionOn' s fun l => by cases' l with a l - rfl + · rfl dsimp only [Chain, ← mk''_eq_coe, Quotient.liftOn'_mk'', Cycle.map, Quotient.map', Quot.map, Quotient.mk'', Quotient.liftOn', Quotient.liftOn, Quot.liftOn_mk, List.map] rw [← concat_eq_append, ← List.map_concat, List.chain_map f] @@ -985,7 +985,7 @@ theorem chain_mono : Monotone (Chain : (α → α → Prop) → Cycle α → Pro theorem chain_of_pairwise : (∀ a ∈ s, ∀ b ∈ s, r a b) → Chain r s := by induction' s using Cycle.induction_on with a l _ - exact fun _ => Cycle.Chain.nil r + · exact fun _ => Cycle.Chain.nil r intro hs have Ha : a ∈ (a :: l : Cycle α) := by simp have Hl : ∀ {b} (_hb : b ∈ l), b ∈ (a :: l : Cycle α) := @fun b hb => by simp [hb] diff --git a/Mathlib/Data/List/Dedup.lean b/Mathlib/Data/List/Dedup.lean index 26af9e520f3720..94f57df9fed6c0 100644 --- a/Mathlib/Data/List/Dedup.lean +++ b/Mathlib/Data/List/Dedup.lean @@ -43,9 +43,9 @@ theorem dedup_cons_of_not_mem' {a : α} {l : List α} (h : a ∉ dedup l) : @[simp] theorem mem_dedup {a : α} {l : List α} : a ∈ dedup l ↔ a ∈ l := by have := not_congr (@forall_mem_pwFilter α (· ≠ ·) _ ?_ a l) - simpa only [dedup, forall_mem_ne, not_not] using this - intros x y z xz - exact not_and_or.1 <| mt (fun h ↦ h.1.trans h.2) xz + · simpa only [dedup, forall_mem_ne, not_not] using this + · intros x y z xz + exact not_and_or.1 <| mt (fun h ↦ h.1.trans h.2) xz #align list.mem_dedup List.mem_dedup @[simp] diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 458080b3df6b9e..b3f51526c6ac68 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -77,8 +77,14 @@ protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) ( generalize e : (l₁ ++ l₂).length = len revert n l₁ l₂ induction' len with len ih <;> intros n l₁ l₂ h - · have l₁_nil : l₁ = [] := by cases l₁; rfl; contradiction - have l₂_nil : l₂ = [] := by cases l₂; rfl; rw [List.length_append] at h; contradiction + · have l₁_nil : l₁ = [] := by + cases l₁ + · rfl + · contradiction + have l₂_nil : l₂ = [] := by + cases l₂ + · rfl + · rw [List.length_append] at h; contradiction simp only [l₁_nil, l₂_nil]; rfl · cases' l₁ with head tail · rfl @@ -105,8 +111,14 @@ theorem mapIdxGo_append : ∀ (f : ℕ → α → β) (l₁ l₂ : List α) (arr generalize e : (l₁ ++ l₂).length = len revert l₁ l₂ arr induction' len with len ih <;> intros l₁ l₂ arr h - · have l₁_nil : l₁ = [] := by cases l₁; rfl; contradiction - have l₂_nil : l₂ = [] := by cases l₂; rfl; rw [List.length_append] at h; contradiction + · have l₁_nil : l₁ = [] := by + cases l₁ + · rfl + · contradiction + have l₂_nil : l₂ = [] := by + 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] · cases' l₁ with head tail <;> simp only [mapIdx.go] · simp only [nil_append, Array.toList_eq, Array.toArray_data] @@ -149,17 +161,20 @@ theorem map_enumFrom_eq_zipWith : ∀ (l : List α) (n : ℕ) (f : ℕ → α generalize e : l.length = len revert l induction' len with len ih <;> intros l e n f - · have : l = [] := by cases l; rfl; contradiction + · have : l = [] := by + cases l + · rfl + · contradiction rw [this]; rfl · cases' l with head tail · contradiction · simp only [map, uncurry_apply_pair, range_succ_eq_map, zipWith, zero_add, zipWith_map_left] rw [ih] - suffices (fun i ↦ f (i + (n + 1))) = ((fun i ↦ f (i + n)) ∘ Nat.succ) by - rw [this] - rfl - funext n' a - simp only [comp, Nat.add_assoc, Nat.add_comm, Nat.add_succ] + · suffices (fun i ↦ f (i + (n + 1))) = ((fun i ↦ f (i + n)) ∘ Nat.succ) by + rw [this] + rfl + funext n' a + simp only [comp, Nat.add_assoc, Nat.add_comm, Nat.add_succ] simp only [length_cons, Nat.succ.injEq] at e; exact e theorem mapIdx_eq_enum_map (l : List α) (f : ℕ → α → β) : @@ -386,7 +401,10 @@ theorem mapIdxMGo_eq_mapIdxMAuxSpec {α β} (f : ℕ → α → m β) (arr : Arr generalize e : as.length = len revert as arr induction' len with len ih <;> intro arr as h - · have : as = [] := by cases as; rfl; contradiction + · have : as = [] := by + cases as + · rfl + · contradiction simp only [this, mapIdxM.go, mapIdxMAuxSpec, List.traverse, map_pure, append_nil] · match as with | nil => contradiction diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 832c6fc083a218..148266e7e2adaf 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -218,7 +218,7 @@ theorem index_of_argmax : · cases not_le_of_lt ‹_› ‹_› · rw [if_pos rfl] · rw [if_neg, if_neg] - exact Nat.succ_le_succ (index_of_argmax h (by assumption) ham) + · exact Nat.succ_le_succ (index_of_argmax h (by assumption) ham) · exact ne_of_apply_ne f (lt_of_lt_of_le ‹_› ‹_›).ne · exact ne_of_apply_ne _ ‹f hd < f _›.ne · rw [if_pos rfl] @@ -509,8 +509,8 @@ theorem le_max_of_le {l : List α} {a x : α} (hx : x ∈ l) (h : a ≤ x) : a induction' l with y l IH · exact absurd hx (not_mem_nil _) · obtain hl | hl := hx - simp only [foldr, foldr_cons] - · exact le_max_of_le_left h + · simp only [foldr, foldr_cons] + exact le_max_of_le_left h · exact le_max_of_le_right (IH (by assumption)) #align list.le_max_of_le List.le_max_of_le diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 6aae15cdf759a3..a6c3a4480e6362 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -853,9 +853,9 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' · obtain ⟨m, rfl⟩ := Nat.exists_eq_add_of_lt H' erw [length_insertNth _ _ hk.le, Nat.succ_lt_succ_iff, Nat.succ_add] at hn rw [nthLe_insertNth_add_succ] - convert nthLe_insertNth_add_succ s x k m.succ (by simpa using hn) using 2 - · simp [Nat.add_succ, Nat.succ_add] - · simp [Nat.add_left_comm, Nat.add_comm] + · convert nthLe_insertNth_add_succ s x k m.succ (by simpa using hn) using 2 + · simp [Nat.add_succ, Nat.succ_add] + · simp [Nat.add_left_comm, Nat.add_comm] · simpa [Nat.succ_add] using hn #align list.nodup_permutations'_aux_iff List.nodup_permutations'Aux_iff diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index e221621e81437b..31ee71adeee0c0 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -95,15 +95,16 @@ theorem map_permutationsAux2' {α β α' β'} (g : α → α') (g' : β → β') · simp · simp only [map, permutationsAux2_snd_cons, cons_append, cons.injEq] rw [ys_ih, permutationsAux2_fst] - refine' ⟨_, rfl⟩ - · simp only [← map_cons, ← map_append]; apply H + · refine' ⟨_, rfl⟩ + simp only [← map_cons, ← map_append]; apply H · intro a; apply H #align list.map_permutations_aux2' List.map_permutationsAux2' /-- The `f` argument to `permutationsAux2` when `r = []` can be eliminated. -/ theorem map_permutationsAux2 (t : α) (ts : List α) (ys : List α) (f : List α → β) : (permutationsAux2 t ts [] ys id).2.map f = (permutationsAux2 t ts [] ys f).2 := by - rw [map_permutationsAux2' id, map_id, map_id]; rfl + rw [map_permutationsAux2' id, map_id, map_id] + · rfl simp #align list.map_permutations_aux2 List.map_permutationsAux2 diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index be04d046c7395d..1201b804906455 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -434,7 +434,7 @@ theorem exists_of_kerase {a : α} {l : List (Sigma β)} (h : a ∈ l.keys) : exact ⟨hd.2, [], tl, by simp, by cases hd; rfl, by simp⟩ · simp only [keys_cons, mem_cons] at h cases' h with h h - exact absurd h e + · exact absurd h e rcases ih h with ⟨b, tl₁, tl₂, h₁, h₂, h₃⟩ exact ⟨b, hd :: tl₁, tl₂, not_mem_cons_of_ne_of_not_mem e h₁, by (rw [h₂]; rfl), by simp [e, h₃]⟩ @@ -655,7 +655,8 @@ theorem nodupKeys_dedupKeys (l : List (Sigma β)) : NodupKeys (dedupKeys l) := b #align list.nodupkeys_dedupkeys List.nodupKeys_dedupKeys theorem dlookup_dedupKeys (a : α) (l : List (Sigma β)) : dlookup a (dedupKeys l) = dlookup a l := by - induction' l with l_hd _ l_ih; rfl + induction' l with l_hd _ l_ih + · rfl cases' l_hd with a' b by_cases h : a = a' · subst a' @@ -671,8 +672,8 @@ theorem sizeOf_dedupKeys {α} {β : α → Type*} [DecidableEq α] [SizeOf (Sigm · simp [dedupKeys] · simp only [dedupKeys_cons, kinsert_def, Nat.add_le_add_iff_left, Sigma.eta] trans - apply sizeOf_kerase - assumption + · apply sizeOf_kerase + · assumption #align list.sizeof_dedupkeys List.sizeOf_dedupKeys /-! ### `kunion` -/ diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 86a4f755c0d68b..a36c9153d02f66 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -84,9 +84,10 @@ theorem mem_sublists' {s t : List α} : s ∈ sublists' t ↔ s <+ t := by · simp only [sublists'_nil, mem_singleton] exact ⟨fun h => by rw [h], eq_nil_of_sublist_nil⟩ simp only [sublists'_cons, mem_append, IH, mem_map] - constructor <;> intro h; rcases h with (h | ⟨s, h, rfl⟩) - · exact sublist_cons_of_sublist _ h - · exact h.cons_cons _ + constructor <;> intro h + · rcases h with (h | ⟨s, h, rfl⟩) + · exact sublist_cons_of_sublist _ h + · exact h.cons_cons _ · cases' h with _ _ _ h s _ _ h · exact Or.inl h · exact Or.inr ⟨s, h, rfl⟩ diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index 25453065634fc5..a35f11808fb935 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -264,8 +264,8 @@ theorem get?_zip_with (f : α → β → γ) (l₁ : List α) (l₂ : List β) ( induction' l₁ with head tail generalizing l₂ i · rw [zipWith] <;> simp · cases l₂ - simp only [zipWith, Seq.seq, Functor.map, get?, Option.map_none'] - · cases (head :: tail).get? i <;> rfl + · simp only [zipWith, Seq.seq, Functor.map, get?, Option.map_none'] + cases (head :: tail).get? i <;> rfl · cases i <;> simp only [Option.map_some', get?, Option.some_bind', *] #align list.nth_zip_with List.get?_zip_with diff --git a/Mathlib/Data/Multiset/Antidiagonal.lean b/Mathlib/Data/Multiset/Antidiagonal.lean index fe6da4ec11371c..c3aba2716c0b2a 100644 --- a/Mathlib/Data/Multiset/Antidiagonal.lean +++ b/Mathlib/Data/Multiset/Antidiagonal.lean @@ -78,7 +78,10 @@ theorem antidiagonal_cons (a : α) (s) : simp only [revzip, reverse_append, quot_mk_to_coe, coe_eq_coe, powersetAux'_cons, cons_coe, map_coe, antidiagonal_coe', coe_add] rw [← zip_map, ← zip_map, zip_append, (_ : _ ++ _ = _)] - · congr; simp only [List.map_id]; rw [map_reverse]; simp + · congr + · simp only [List.map_id] + · rw [map_reverse] + · simp · simp #align multiset.antidiagonal_cons Multiset.antidiagonal_cons diff --git a/Mathlib/Data/Multiset/Pi.lean b/Mathlib/Data/Multiset/Pi.lean index 4cffc055bfc999..62d6a5a5f4418d 100644 --- a/Mathlib/Data/Multiset/Pi.lean +++ b/Mathlib/Data/Multiset/Pi.lean @@ -54,7 +54,7 @@ theorem Pi.cons_swap {a a' : α} {b : δ a} {b' : δ a'} {m : Multiset α} {f : rintro a'' _ rfl refine' hfunext (by rw [Multiset.cons_swap]) fun ha₁ ha₂ _ => _ rcases ne_or_eq a'' a with (h₁ | rfl) - rcases eq_or_ne a'' a' with (rfl | h₂) + on_goal 1 => rcases eq_or_ne a'' a' with (rfl | h₂) all_goals simp [*, Pi.cons_same, Pi.cons_ne] #align multiset.pi.cons_swap Multiset.Pi.cons_swap diff --git a/Mathlib/Data/Multiset/Sum.lean b/Mathlib/Data/Multiset/Sum.lean index 60e290627adda7..5f711c575147ff 100644 --- a/Mathlib/Data/Multiset/Sum.lean +++ b/Mathlib/Data/Multiset/Sum.lean @@ -55,7 +55,7 @@ theorem mem_disjSum : x ∈ s.disjSum t ↔ (∃ a, a ∈ s ∧ inl a = x) ∨ theorem inl_mem_disjSum : inl a ∈ s.disjSum t ↔ a ∈ s := by rw [mem_disjSum, or_iff_left] -- Porting note: Previous code for L62 was: simp only [exists_eq_right] - simp only [inl.injEq, exists_eq_right] + · simp only [inl.injEq, exists_eq_right] rintro ⟨b, _, hb⟩ exact inr_ne_inl hb #align multiset.inl_mem_disj_sum Multiset.inl_mem_disjSum @@ -64,7 +64,7 @@ theorem inl_mem_disjSum : inl a ∈ s.disjSum t ↔ a ∈ s := by theorem inr_mem_disjSum : inr b ∈ s.disjSum t ↔ b ∈ t := by rw [mem_disjSum, or_iff_right] -- Porting note: Previous code for L72 was: simp only [exists_eq_right] - simp only [inr.injEq, exists_eq_right] + · simp only [inr.injEq, exists_eq_right] rintro ⟨a, _, ha⟩ exact inl_ne_inr ha #align multiset.inr_mem_disj_sum Multiset.inr_mem_disjSum diff --git a/Mathlib/Data/NNRat/Defs.lean b/Mathlib/Data/NNRat/Defs.lean index 6bd37d010afaf9..81cb6619e1b4af 100644 --- a/Mathlib/Data/NNRat/Defs.lean +++ b/Mathlib/Data/NNRat/Defs.lean @@ -370,8 +370,8 @@ lemma coprime_num_den (q : ℚ≥0) : q.num.Coprime q.den := by simpa [num, den] theorem ext_num_den (hn : p.num = q.num) (hd : p.den = q.den) : p = q := by refine ext <| Rat.ext ?_ ?_ · apply (Int.natAbs_inj_of_nonneg_of_nonneg _ _).1 hn - exact Rat.num_nonneg.2 p.2 - exact Rat.num_nonneg.2 q.2 + · exact Rat.num_nonneg.2 p.2 + · exact Rat.num_nonneg.2 q.2 · exact hd #align nnrat.ext_num_denom NNRat.ext_num_den diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 1394bcc40075e7..a55ea42627f024 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -149,8 +149,8 @@ lemma bit1_val (n : Nat) : bit1 n = 2 * n + 1 := congr_arg succ (bit0_val _) lemma bit_val (b n) : bit b n = 2 * n + cond b 1 0 := by cases b - apply bit0_val - apply bit1_val + · apply bit0_val + · apply bit1_val #align nat.bit_val Nat.bit_val lemma bit_decomp (n : Nat) : bit (bodd n) (div2 n) = n := diff --git a/Mathlib/Data/Nat/Cast/Field.lean b/Mathlib/Data/Nat/Cast/Field.lean index d9212da4cc2a98..53d09a2ee21caf 100644 --- a/Mathlib/Data/Nat/Cast/Field.lean +++ b/Mathlib/Data/Nat/Cast/Field.lean @@ -55,7 +55,7 @@ theorem cast_div_le {m n : ℕ} : ((m / n : ℕ) : α) ≤ m / n := by cases n · rw [cast_zero, div_zero, Nat.div_zero, cast_zero] rw [le_div_iff, ← Nat.cast_mul, @Nat.cast_le] - exact Nat.div_mul_le_self m _ + · exact Nat.div_mul_le_self m _ · exact Nat.cast_pos.2 (Nat.succ_pos _) #align nat.cast_div_le Nat.cast_div_le @@ -70,14 +70,14 @@ theorem one_div_pos_of_nat {n : ℕ} : 0 < 1 / ((n : α) + 1) := by theorem one_div_le_one_div {n m : ℕ} (h : n ≤ m) : 1 / ((m : α) + 1) ≤ 1 / ((n : α) + 1) := by refine' one_div_le_one_div_of_le _ _ - exact Nat.cast_add_one_pos _ - simpa + · exact Nat.cast_add_one_pos _ + · simpa #align nat.one_div_le_one_div Nat.one_div_le_one_div theorem one_div_lt_one_div {n m : ℕ} (h : n < m) : 1 / ((m : α) + 1) < 1 / ((n : α) + 1) := by refine' one_div_lt_one_div_of_lt _ _ - exact Nat.cast_add_one_pos _ - simpa + · exact Nat.cast_add_one_pos _ + · simpa #align nat.one_div_lt_one_div Nat.one_div_lt_one_div end LinearOrderedSemifield diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index fc9ed2a3454a0a..2735e4d8bc7223 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -1197,7 +1197,8 @@ lemma dvd_sub_mod (k : ℕ) : n ∣ k - k % n := lemma add_mod_eq_ite : (m + n) % k = if k ≤ m % k + n % k then m % k + n % k - k else m % k + n % k := by - cases k; simp only [zero_eq, mod_zero, zero_le, ↓reduceIte, Nat.sub_zero] + cases k + · simp rw [Nat.add_mod] split_ifs with h · rw [Nat.mod_eq_sub_mod h, Nat.mod_eq_of_lt] @@ -1292,9 +1293,9 @@ set_option linter.deprecated false in protected theorem not_two_dvd_bit1 (n : ℕ) : ¬2 ∣ bit1 n := by rw [bit1, Nat.dvd_add_right, Nat.dvd_one] -- Porting note: was `cc` - decide - rw [bit0, ← Nat.two_mul] - exact Nat.dvd_mul_right _ _ + · decide + · rw [bit0, ← Nat.two_mul] + exact Nat.dvd_mul_right _ _ #align nat.not_two_dvd_bit1 Nat.not_two_dvd_bit1 /-- A natural number `m` divides the sum `m + n` if and only if `m` divides `n`.-/ diff --git a/Mathlib/Data/Nat/EvenOddRec.lean b/Mathlib/Data/Nat/EvenOddRec.lean index 38a477d54a9a30..d74179ad109d26 100644 --- a/Mathlib/Data/Nat/EvenOddRec.lean +++ b/Mathlib/Data/Nat/EvenOddRec.lean @@ -38,7 +38,10 @@ theorem evenOddRec_even (n : ℕ) (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i (2 * n).evenOddRec h0 h_even h_odd = h_even n (evenOddRec h0 h_even h_odd n) := have : ∀ a, bit false n = a → HEq (@evenOddRec _ h0 h_even h_odd a) (h_even n (evenOddRec h0 h_even h_odd n)) - | _, rfl => by rw [evenOddRec, binaryRec_eq]; apply eq_rec_heq; exact H + | _, rfl => by + rw [evenOddRec, binaryRec_eq] + · apply eq_rec_heq + · exact H eq_of_heq (this _ (bit0_val _)) #align nat.even_odd_rec_even Nat.evenOddRec_even @@ -48,7 +51,10 @@ theorem evenOddRec_odd (n : ℕ) (P : ℕ → Sort*) (h0 : P 0) (h_even : ∀ i, (2 * n + 1).evenOddRec h0 h_even h_odd = h_odd n (evenOddRec h0 h_even h_odd n) := have : ∀ a, bit true n = a → HEq (@evenOddRec _ h0 h_even h_odd a) (h_odd n (evenOddRec h0 h_even h_odd n)) - | _, rfl => by rw [evenOddRec, binaryRec_eq]; apply eq_rec_heq; exact H + | _, rfl => by + rw [evenOddRec, binaryRec_eq] + · apply eq_rec_heq + · exact H eq_of_heq (this _ (bit1_val _)) #align nat.even_odd_rec_odd Nat.evenOddRec_odd diff --git a/Mathlib/Data/Nat/Fib/Zeckendorf.lean b/Mathlib/Data/Nat/Fib/Zeckendorf.lean index cc927c5f49ec2a..3210f4cdd72670 100644 --- a/Mathlib/Data/Nat/Fib/Zeckendorf.lean +++ b/Mathlib/Data/Nat/Fib/Zeckendorf.lean @@ -102,7 +102,7 @@ lemma greatestFib_sub_fib_greatestFib_le_greatestFib (hn : n ≠ 0) : greatestFib (n - fib (greatestFib n)) ≤ greatestFib n - 2 := by rw [← Nat.lt_succ_iff, greatestFib_lt, tsub_lt_iff_right n.fib_greatestFib_le, Nat.sub_succ, succ_pred, ← fib_add_one] - exact n.lt_fib_greatestFib_add_one + · exact n.lt_fib_greatestFib_add_one · simpa · simpa [← succ_le_iff, tsub_eq_zero_iff_le] using hn.bot_lt diff --git a/Mathlib/Data/Nat/Lattice.lean b/Mathlib/Data/Nat/Lattice.lean index 7cdcfb161a1c4a..1ba3d8ee00b467 100644 --- a/Mathlib/Data/Nat/Lattice.lean +++ b/Mathlib/Data/Nat/Lattice.lean @@ -111,8 +111,9 @@ theorem sInf_upward_closed_eq_succ_iff {s : Set ℕ} (hs : ∀ k₁ k₂ : ℕ, constructor · intro H rw [eq_Ici_of_nonempty_of_upward_closed (nonempty_of_sInf_eq_succ _) hs, H, mem_Ici, mem_Ici] - exact ⟨le_rfl, k.not_succ_le_self⟩; - exact k; assumption + · exact ⟨le_rfl, k.not_succ_le_self⟩; + · exact k + · assumption · rintro ⟨H, H'⟩ rw [sInf_def (⟨_, H⟩ : s.Nonempty), find_eq_iff] exact ⟨H, fun n hnk hns ↦ H' <| hs n k (Nat.lt_succ_iff.mp hnk) hns⟩ diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index e66b40dedb68be..a9c866e48ae8b0 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -286,13 +286,13 @@ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b have hcd := gcd_dvd_right m c rw [modEq_iff_dvd] refine' @Int.dvd_of_dvd_mul_right_of_gcd_one (m / d) (c / d) (b - a) _ _ - show (m / d : ℤ) ∣ c / d * (b - a) - · rw [mul_comm, ← Int.mul_ediv_assoc (b - a) (Int.natCast_dvd_natCast.mpr hcd), mul_comm] + · show (m / d : ℤ) ∣ c / d * (b - a) + rw [mul_comm, ← Int.mul_ediv_assoc (b - a) (Int.natCast_dvd_natCast.mpr hcd), mul_comm] apply Int.ediv_dvd_ediv (Int.natCast_dvd_natCast.mpr hmd) rw [mul_sub] exact modEq_iff_dvd.mp h - show Int.gcd (m / d) (c / d) = 1 - · simp only [← Int.natCast_div, Int.coe_nat_gcd (m / d) (c / d), gcd_div hmd hcd, + · show Int.gcd (m / d) (c / d) = 1 + simp only [← Int.natCast_div, Int.coe_nat_gcd (m / d) (c / d), gcd_div hmd hcd, Nat.div_self (gcd_pos_of_pos_left c hm)] #align nat.modeq.cancel_left_div_gcd Nat.ModEq.cancel_left_div_gcd @@ -331,9 +331,15 @@ end ModEq /-- The natural number less than `lcm n m` congruent to `a` mod `n` and `b` mod `m` -/ def chineseRemainder' (h : a ≡ b [MOD gcd n m]) : { k // k ≡ a [MOD n] ∧ k ≡ b [MOD m] } := - if hn : n = 0 then ⟨a, by rw [hn, gcd_zero_left] at h; constructor; rfl; exact h⟩ + if hn : n = 0 then ⟨a, by + rw [hn, gcd_zero_left] at h; constructor + · rfl + · exact h⟩ else - if hm : m = 0 then ⟨b, by rw [hm, gcd_zero_right] at h; constructor; exact h.symm; rfl⟩ + if hm : m = 0 then ⟨b, by + rw [hm, gcd_zero_right] at h; constructor + · exact h.symm + · rfl⟩ else ⟨let (c, d) := xgcd n m; Int.toNat ((n * c * b + m * d * a) / gcd n m % lcm n m), by rw [xgcd_val] @@ -353,7 +359,7 @@ def chineseRemainder' (h : a ≡ b [MOD gcd n m]) : { k // k ≡ a [MOD n] ∧ k rw [← this, sub_mul, ← add_sub_assoc, add_comm, add_sub_assoc, ← mul_sub, Int.add_ediv_of_dvd_left, Int.mul_ediv_cancel_left _ hnonzero, Int.mul_ediv_assoc _ h.dvd, ← sub_sub, sub_self, zero_sub, dvd_neg, mul_assoc] - exact dvd_mul_right _ _ + · exact dvd_mul_right _ _ norm_cast exact dvd_mul_right _ _ · exact dvd_lcm_left n m @@ -361,8 +367,8 @@ def chineseRemainder' (h : a ≡ b [MOD gcd n m]) : { k // k ≡ a [MOD n] ∧ k rw [← this, sub_mul, sub_add, ← mul_sub, Int.sub_ediv_of_dvd, Int.mul_ediv_cancel_left _ hnonzero, Int.mul_ediv_assoc _ h.dvd, ← sub_add, sub_self, zero_add, mul_assoc] - exact dvd_mul_right _ _ - exact hcoedvd _ + · exact dvd_mul_right _ _ + · exact hcoedvd _ · exact dvd_lcm_right n m⟩ #align nat.chinese_remainder' Nat.chineseRemainder' diff --git a/Mathlib/Data/Nat/PSub.lean b/Mathlib/Data/Nat/PSub.lean index 6eec16f9eee34f..1fb4b1243edaff 100644 --- a/Mathlib/Data/Nat/PSub.lean +++ b/Mathlib/Data/Nat/PSub.lean @@ -118,8 +118,8 @@ def psub' (m n : ℕ) : Option ℕ := theorem psub'_eq_psub (m n) : psub' m n = psub m n := by rw [psub'] split_ifs with h - exact (psub_eq_sub h).symm - exact (psub_eq_none.2 (not_le.1 h)).symm + · exact (psub_eq_sub h).symm + · exact (psub_eq_none.2 (not_le.1 h)).symm #align nat.psub'_eq_psub Nat.psub'_eq_psub end Nat diff --git a/Mathlib/Data/Nat/Pairing.lean b/Mathlib/Data/Nat/Pairing.lean index aeef9c860f86b4..205bb2516bbac8 100644 --- a/Mathlib/Data/Nat/Pairing.lean +++ b/Mathlib/Data/Nat/Pairing.lean @@ -127,8 +127,8 @@ theorem pair_lt_pair_left {a₁ a₂} (b) (h : a₁ < a₂) : pair a₁ b < pair · by_cases h₂ : a₂ < b <;> simp [pair, h₂, h] simp? at h₂ says simp only [not_lt] at h₂ apply Nat.add_lt_add_of_le_of_lt - exact Nat.mul_self_le_mul_self h₂ - exact Nat.lt_add_right _ h + · exact Nat.mul_self_le_mul_self h₂ + · exact Nat.lt_add_right _ h · simp at h₁ simp only [not_lt_of_gt (lt_of_le_of_lt h₁ h), ite_false] apply add_lt_add @@ -155,7 +155,7 @@ theorem pair_lt_max_add_one_sq (m n : ℕ) : pair m n < (max m n + 1) ^ 2 := by theorem max_sq_add_min_le_pair (m n : ℕ) : max m n ^ 2 + min m n ≤ pair m n := by rw [pair] cases' lt_or_le m n with h h - rw [if_pos h, max_eq_right h.le, min_eq_left h.le, Nat.pow_two] + · rw [if_pos h, max_eq_right h.le, min_eq_left h.le, Nat.pow_two] rw [if_neg h.not_lt, max_eq_left h, min_eq_right h, Nat.pow_two, Nat.add_assoc, Nat.add_le_add_iff_left] exact Nat.le_add_left _ _ diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index f7ac76ee1d9a0e..4698995f25cbbf 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -26,35 +26,35 @@ instance : WellFoundedRelation (WithBot ℕ) where theorem add_eq_zero_iff {n m : WithBot ℕ} : n + m = 0 ↔ n = 0 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals (exact ⟨fun h => Option.noConfusion h, fun h => Option.noConfusion h.1⟩) - exact ⟨fun h => Option.noConfusion h, fun h => Option.noConfusion h.2⟩ - repeat' erw [WithBot.coe_eq_coe] + repeat (· exact ⟨fun h => Option.noConfusion h, fun h => Option.noConfusion h.1⟩) + · exact ⟨fun h => Option.noConfusion h, fun h => Option.noConfusion h.2⟩ + repeat erw [WithBot.coe_eq_coe] exact add_eq_zero_iff' (zero_le _) (zero_le _) #align nat.with_bot.add_eq_zero_iff Nat.WithBot.add_eq_zero_iff theorem add_eq_one_iff {n m : WithBot ℕ} : n + m = 1 ↔ n = 0 ∧ m = 1 ∨ n = 1 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + repeat refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop (simp_config := { decide := true }) - repeat' erw [WithBot.coe_eq_coe] + repeat erw [WithBot.coe_eq_coe] exact Nat.add_eq_one_iff #align nat.with_bot.add_eq_one_iff Nat.WithBot.add_eq_one_iff theorem add_eq_two_iff {n m : WithBot ℕ} : n + m = 2 ↔ n = 0 ∧ m = 2 ∨ n = 1 ∧ m = 1 ∨ n = 2 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + repeat refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop (simp_config := { decide := true }) - repeat' erw [WithBot.coe_eq_coe] + repeat erw [WithBot.coe_eq_coe] exact Nat.add_eq_two_iff #align nat.with_bot.add_eq_two_iff Nat.WithBot.add_eq_two_iff theorem add_eq_three_iff {n m : WithBot ℕ} : n + m = 3 ↔ n = 0 ∧ m = 3 ∨ n = 1 ∧ m = 2 ∨ n = 2 ∧ m = 1 ∨ n = 3 ∧ m = 0 := by rcases n, m with ⟨_ | _, _ | _⟩ - any_goals refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; + repeat refine' ⟨fun h => Option.noConfusion h, fun h => _⟩; aesop (simp_config := { decide := true }) - repeat' erw [WithBot.coe_eq_coe] + repeat erw [WithBot.coe_eq_coe] exact Nat.add_eq_three_iff #align nat.with_bot.add_eq_three_iff Nat.WithBot.add_eq_three_iff diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index afce2d3aa82d82..f934c992b2a620 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -894,7 +894,7 @@ theorem castNum_eq_bitwise {f : Num → Num → Num} {g : Bool → Bool → Bool show PosNum.bit0 = PosNum.bit false from rfl, show PosNum.bit1 = PosNum.bit true from rfl, show ((1 : Num) : ℕ) = Nat.bit true 0 from rfl] all_goals - repeat' + repeat rw [show ∀ b n, (pos (PosNum.bit b n) : ℕ) = Nat.bit b ↑n by intros b _; cases b <;> rfl] rw [Nat.bitwise_bit gff] @@ -956,13 +956,13 @@ theorem castNum_shiftRight (m : Num) (n : Nat) : ↑(m >>> n) = (m : ℕ) >>> (n apply Nat.div_eq_of_lt simp · trans - apply IH + · apply IH change Nat.shiftRight m n = Nat.shiftRight (_root_.bit1 m) (n + 1) rw [add_comm n 1, @Nat.shiftRight_eq _ (1 + n), Nat.shiftRight_add] apply congr_arg fun x => Nat.shiftRight x n simp [Nat.shiftRight_succ, Nat.shiftRight_zero, ← Nat.div2_val] · trans - apply IH + · apply IH change Nat.shiftRight m n = Nat.shiftRight (_root_.bit0 m) (n + 1) rw [add_comm n 1, @Nat.shiftRight_eq _ (1 + n), Nat.shiftRight_add] apply congr_arg fun x => Nat.shiftRight x n diff --git a/Mathlib/Data/Ordmap/Ordset.lean b/Mathlib/Data/Ordmap/Ordset.lean index 2563156b6220dc..00405ab65d5abd 100644 --- a/Mathlib/Data/Ordmap/Ordset.lean +++ b/Mathlib/Data/Ordmap/Ordset.lean @@ -784,9 +784,10 @@ def Raised (n m : ℕ) : Prop := #align ordnode.raised Ordnode.Raised theorem raised_iff {n m} : Raised n m ↔ n ≤ m ∧ m ≤ n + 1 := by - constructor; rintro (rfl | rfl) - · exact ⟨le_rfl, Nat.le_succ _⟩ - · exact ⟨Nat.le_succ _, le_rfl⟩ + constructor + · rintro (rfl | rfl) + · exact ⟨le_rfl, Nat.le_succ _⟩ + · exact ⟨Nat.le_succ _, le_rfl⟩ · rintro ⟨h₁, h₂⟩ rcases eq_or_lt_of_le h₁ with (rfl | h₁) · exact Or.inl rfl @@ -1427,30 +1428,30 @@ theorem Valid'.glue_aux {l r o₁ o₂} (hl : Valid' o₁ l o₂) (hr : Valid' o cases' r with rs rl rx rr; · exact ⟨hl, rfl⟩ dsimp [glue]; split_ifs · rw [splitMax_eq] - cases' Valid'.eraseMax_aux hl with v e - suffices H : _ by - refine' ⟨Valid'.balanceR v (hr.of_gt _ _) H, _⟩ - · refine' findMax'_all (P := fun a : α => Bounded nil (a : WithTop α) o₂) - lx lr hl.1.2.to_nil (sep.2.2.imp _) - exact fun x h => hr.1.2.to_nil.mono_left (le_of_lt h.2.1) - · exact @findMax'_all _ (fun a => All (· > a) (.node rs rl rx rr)) lx lr sep.2.1 sep.2.2 - · rw [size_balanceR v.3 hr.3 v.2 hr.2 H, add_right_comm, ← e, hl.2.1]; rfl - refine' Or.inl ⟨_, Or.inr e, _⟩ - rwa [hl.2.eq_node'] at bal + · cases' Valid'.eraseMax_aux hl with v e + suffices H : _ by + refine' ⟨Valid'.balanceR v (hr.of_gt _ _) H, _⟩ + · refine' findMax'_all (P := fun a : α => Bounded nil (a : WithTop α) o₂) + lx lr hl.1.2.to_nil (sep.2.2.imp _) + exact fun x h => hr.1.2.to_nil.mono_left (le_of_lt h.2.1) + · exact @findMax'_all _ (fun a => All (· > a) (.node rs rl rx rr)) lx lr sep.2.1 sep.2.2 + · rw [size_balanceR v.3 hr.3 v.2 hr.2 H, add_right_comm, ← e, hl.2.1]; rfl + refine' Or.inl ⟨_, Or.inr e, _⟩ + rwa [hl.2.eq_node'] at bal · rw [splitMin_eq] - cases' Valid'.eraseMin_aux hr with v e - 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 - 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 - (all_iff_forall.2 fun x hx => sep.imp fun y hy => all_iff_forall.1 hy.1 _ hx) - (sep.imp fun y hy => hy.2.1) - · rw [size_balanceL hl.3 v.3 hl.2 v.2 H, add_assoc, ← e, hr.2.1]; rfl - refine' Or.inr ⟨_, Or.inr e, _⟩ - rwa [hr.2.eq_node'] at bal + · cases' Valid'.eraseMin_aux hr with v e + 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 + 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 + (all_iff_forall.2 fun x hx => sep.imp fun y hy => all_iff_forall.1 hy.1 _ hx) + (sep.imp fun y hy => hy.2.1) + · rw [size_balanceL hl.3 v.3 hl.2 v.2 H, add_assoc, ← e, hr.2.1]; rfl + refine' Or.inr ⟨_, Or.inr e, _⟩ + rwa [hr.2.eq_node'] at bal #align ordnode.valid'.glue_aux Ordnode.Valid'.glue_aux theorem Valid'.glue {l} {x : α} {r o₁ o₂} (hl : Valid' o₁ l x) (hr : Valid' x r o₂) : diff --git a/Mathlib/Data/PFunctor/Multivariate/M.lean b/Mathlib/Data/PFunctor/Multivariate/M.lean index 7f73b764925aa7..9fc2d934fb5ff0 100644 --- a/Mathlib/Data/PFunctor/Multivariate/M.lean +++ b/Mathlib/Data/PFunctor/Multivariate/M.lean @@ -213,7 +213,7 @@ theorem M.dest_corec' {α : TypeVec.{u} n} {β : Type u} (g₀ : β → P.A) theorem M.dest_corec {α : TypeVec n} {β : Type u} (g : β → P (α.append1 β)) (x : β) : M.dest P (M.corec P g x) = appendFun id (M.corec P g) <$$> g x := by trans - apply M.dest_corec' + · apply M.dest_corec' cases' g x with a f; dsimp rw [MvPFunctor.map_eq]; congr conv => diff --git a/Mathlib/Data/PFunctor/Univariate/M.lean b/Mathlib/Data/PFunctor/Univariate/M.lean index 3c3d96a904f032..742b20c90db18c 100644 --- a/Mathlib/Data/PFunctor/Univariate/M.lean +++ b/Mathlib/Data/PFunctor/Univariate/M.lean @@ -249,7 +249,7 @@ def children (x : M F) (i : F.B (head x)) : M F := have P' := x.2 (succ n) apply agree_children _ _ _ P' trans i - apply cast_heq + · apply cast_heq symm apply cast_heq } set_option linter.uppercaseLean3 false in @@ -516,8 +516,8 @@ theorem iselect_eq_default [DecidableEq F.A] [Inhabited (M F)] (ps : Path F) (x induction' x using PFunctor.M.casesOn' with x_a x_f simp only [iselect, isubtree] at ps_ih ⊢ by_cases h'' : a = x_a - subst x_a - · simp only [dif_pos, eq_self_iff_true, casesOn_mk'] + · subst x_a + simp only [dif_pos, eq_self_iff_true, casesOn_mk'] rw [ps_ih] intro h' apply h diff --git a/Mathlib/Data/PNat/Factors.lean b/Mathlib/Data/PNat/Factors.lean index 0896d803bb7bba..61e4d4c36db961 100644 --- a/Mathlib/Data/PNat/Factors.lean +++ b/Mathlib/Data/PNat/Factors.lean @@ -364,12 +364,12 @@ theorem factorMultiset_gcd (m n : ℕ+) : factorMultiset (gcd m n) = factorMultiset m ⊓ factorMultiset n := by apply le_antisymm · apply le_inf_iff.mpr; constructor <;> apply factorMultiset_le_iff.mpr - exact gcd_dvd_left m n - exact gcd_dvd_right m n + · exact gcd_dvd_left m n + · exact gcd_dvd_right m n · rw [← PrimeMultiset.prod_dvd_iff, prod_factorMultiset] apply dvd_gcd <;> rw [PrimeMultiset.prod_dvd_iff'] - exact inf_le_left - exact inf_le_right + · exact inf_le_left + · exact inf_le_right #align pnat.factor_multiset_gcd PNat.factorMultiset_gcd theorem factorMultiset_lcm (m n : ℕ+) : @@ -377,11 +377,11 @@ theorem factorMultiset_lcm (m n : ℕ+) : apply le_antisymm · rw [← PrimeMultiset.prod_dvd_iff, prod_factorMultiset] apply lcm_dvd <;> rw [← factorMultiset_le_iff'] - exact le_sup_left - exact le_sup_right + · exact le_sup_left + · exact le_sup_right · apply sup_le_iff.mpr; constructor <;> apply factorMultiset_le_iff.mpr - exact dvd_lcm_left m n - exact dvd_lcm_right m n + · exact dvd_lcm_left m n + · exact dvd_lcm_right m n #align pnat.factor_multiset_lcm PNat.factorMultiset_lcm /-- The number of occurrences of p in the factor multiset of m diff --git a/Mathlib/Data/PNat/Prime.lean b/Mathlib/Data/PNat/Prime.lean index 8ffcabad7faa06..f49467f0726a34 100644 --- a/Mathlib/Data/PNat/Prime.lean +++ b/Mathlib/Data/PNat/Prime.lean @@ -101,8 +101,9 @@ theorem gcd_mul_lcm (n m : ℕ+) : gcd n m * lcm n m = n * m := #align pnat.gcd_mul_lcm PNat.gcd_mul_lcm theorem eq_one_of_lt_two {n : ℕ+} : n < 2 → n = 1 := by - intro h; apply le_antisymm; swap; apply PNat.one_le - exact PNat.lt_add_one_iff.1 h + intro h; apply le_antisymm; swap + · apply PNat.one_le + · exact PNat.lt_add_one_iff.1 h #align pnat.eq_one_of_lt_two PNat.eq_one_of_lt_two section Prime @@ -189,13 +190,13 @@ theorem coprime_coe {m n : ℕ+} : Nat.Coprime ↑m ↑n ↔ m.Coprime n := by #align pnat.coprime_coe PNat.coprime_coe theorem Coprime.mul {k m n : ℕ+} : m.Coprime k → n.Coprime k → (m * n).Coprime k := by - repeat' rw [← coprime_coe] + repeat rw [← coprime_coe] rw [mul_coe] apply Nat.Coprime.mul #align pnat.coprime.mul PNat.Coprime.mul theorem Coprime.mul_right {k m n : ℕ+} : k.Coprime m → k.Coprime n → k.Coprime (m * n) := by - repeat' rw [← coprime_coe] + repeat rw [← coprime_coe] rw [mul_coe] apply Nat.Coprime.mul_right #align pnat.coprime.mul_right PNat.Coprime.mul_right @@ -271,7 +272,7 @@ theorem coprime_one {n : ℕ+} : n.Coprime 1 := theorem Coprime.coprime_dvd_left {m k n : ℕ+} : m ∣ k → k.Coprime n → m.Coprime n := by rw [dvd_iff] - repeat' rw [← coprime_coe] + repeat rw [← coprime_coe] apply Nat.Coprime.coprime_dvd_left #align pnat.coprime.coprime_dvd_left PNat.Coprime.coprime_dvd_left diff --git a/Mathlib/Data/PNat/Xgcd.lean b/Mathlib/Data/PNat/Xgcd.lean index c43183d9b8b21c..5dc6c9b0dc7918 100644 --- a/Mathlib/Data/PNat/Xgcd.lean +++ b/Mathlib/Data/PNat/Xgcd.lean @@ -473,7 +473,7 @@ theorem gcd_props : have hb' : (b' : ℕ) = y + z := gcdB'_coe a b have hdet : w * z = succPNat (x * y) := u.reduce_isSpecial' rfl constructor - exact hdet + · exact hdet have hdet' : (w * z : ℕ) = x * y + 1 := by rw [← mul_coe, hdet, succPNat_coe] have _ : u.v = ⟨a, b⟩ := XgcdType.start_v a b let hv : Prod.mk (w * d + x * ur.b : ℕ) (y * d + z * ur.b : ℕ) = ⟨a, b⟩ := @@ -482,9 +482,9 @@ theorem gcd_props : have ha'' : (a : ℕ) = a' * d := (congr_arg Prod.fst hv).symm have hb'' : (b : ℕ) = b' * d := (congr_arg Prod.snd hv).symm constructor - exact eq ha'' + · exact eq ha'' constructor - exact eq hb'' + · exact eq hb'' have hza' : (z * a' : ℕ) = x * b' + 1 := by rw [ha', hb', mul_add, mul_add, mul_comm (z : ℕ), hdet'] ring @@ -498,7 +498,7 @@ theorem gcd_props : · apply eq rw [succPNat_coe, Nat.succ_eq_add_one, mul_coe, hwb'] rw [ha'', hb''] - repeat' rw [← @mul_assoc] + repeat rw [← @mul_assoc] rw [hza', hwb'] constructor <;> ring #align pnat.gcd_props PNat.gcd_props @@ -507,8 +507,8 @@ theorem gcd_eq : gcdD a b = gcd a b := by rcases gcd_props a b with ⟨_, h₁, h₂, _, _, h₅, _⟩ apply dvd_antisymm · apply dvd_gcd - exact Dvd.intro (gcdA' a b) (h₁.trans (mul_comm _ _)).symm - exact Dvd.intro (gcdB' a b) (h₂.trans (mul_comm _ _)).symm + · exact Dvd.intro (gcdA' a b) (h₁.trans (mul_comm _ _)).symm + · exact Dvd.intro (gcdB' a b) (h₂.trans (mul_comm _ _)).symm · have h₇ : (gcd a b : ℕ) ∣ gcdZ a b * a := (Nat.gcd_dvd_left a b).trans (dvd_mul_left _ _) have h₈ : (gcd a b : ℕ) ∣ gcdX a b * b := (Nat.gcd_dvd_right a b).trans (dvd_mul_left _ _) rw [h₅] at h₇ diff --git a/Mathlib/Data/Pi/Lex.lean b/Mathlib/Data/Pi/Lex.lean index 6f7239630b3d8c..0912e500880553 100644 --- a/Mathlib/Data/Pi/Lex.lean +++ b/Mathlib/Data/Pi/Lex.lean @@ -191,9 +191,9 @@ instance [Preorder ι] [∀ i, LT (β i)] [∀ i, DenselyOrdered (β i)] : obtain ⟨a, ha₁, ha₂⟩ := exists_between hi classical refine' ⟨Function.update a₂ _ a, ⟨i, fun j hj => _, _⟩, i, fun j hj => _, _⟩ - rw [h j hj] - dsimp only at hj - · rw [Function.update_noteq hj.ne a] + · rw [h j hj] + dsimp only at hj + rw [Function.update_noteq hj.ne a] · rwa [Function.update_same i a] · rw [Function.update_noteq hj.ne a] · rwa [Function.update_same i a]⟩ diff --git a/Mathlib/Data/Prod/TProd.lean b/Mathlib/Data/Prod/TProd.lean index 6609b922cc63df..be9e3a1946981f 100644 --- a/Mathlib/Data/Prod/TProd.lean +++ b/Mathlib/Data/Prod/TProd.lean @@ -117,7 +117,8 @@ theorem ext : (_ : ∀ (i) (hi : i ∈ l), v.elim hi = w.elim hi), v = w | [], _, v, w, _ => PUnit.ext v w | i :: is, hl, v, w, hvw => by - apply Prod.ext; rw [← elim_self v, hvw, elim_self] + apply Prod.ext + · rw [← elim_self v, hvw, elim_self] refine' ext (nodup_cons.mp hl).2 fun j hj => _ rw [← elim_of_mem hl, hvw, elim_of_mem hl] #align list.tprod.ext List.TProd.ext diff --git a/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean b/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean index 7ab940a84543f3..b95c0872c8b2f5 100644 --- a/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean +++ b/Mathlib/Data/QPF/Multivariate/Constructions/Fix.lean @@ -164,7 +164,10 @@ theorem wEquiv_map {α β : TypeVec n} (g : α ⟹ β) (x y : q.P.W α) : abs (q.P.objAppend1 a₀ (g ⊚ f'₀) fun x => q.P.wMap g (f₀ x)) = abs (q.P.objAppend1 a₁ (g ⊚ f'₁) fun x => q.P.wMap g (f₁ x)) rw [← q.P.map_objAppend1, ← q.P.map_objAppend1, abs_map, abs_map, h] - | trans x y z _ _ ih₁ ih₂ => apply MvQPF.WEquiv.trans; apply ih₁; apply ih₂ + | trans x y z _ _ ih₁ ih₂ => + apply MvQPF.WEquiv.trans + · apply ih₁ + · apply ih₂ set_option linter.uppercaseLean3 false in #align mvqpf.Wequiv_map MvQPF.wEquiv_map diff --git a/Mathlib/Data/QPF/Univariate/Basic.lean b/Mathlib/Data/QPF/Univariate/Basic.lean index 58d7fdb49d9e5f..c58fb00437cb0e 100644 --- a/Mathlib/Data/QPF/Univariate/Basic.lean +++ b/Mathlib/Data/QPF/Univariate/Basic.lean @@ -553,8 +553,8 @@ def comp : QPF (Functor.Comp F₂ F₁) where cases' a with b h; dsimp symm trans - symm - apply abs_map + · symm + apply abs_map congr rw [PFunctor.map_eq] dsimp [Function.comp_def] diff --git a/Mathlib/Data/Rat/Cast/Defs.lean b/Mathlib/Data/Rat/Cast/Defs.lean index 523041b554924c..ba11da9369e766 100644 --- a/Mathlib/Data/Rat/Cast/Defs.lean +++ b/Mathlib/Data/Rat/Cast/Defs.lean @@ -179,8 +179,8 @@ lemma cast_add_of_ne_zero {q r : ℚ} (hq : (q.den : α) ≠ 0) (hr : (r.den : (q + r : ℚ) = (q + r : α) := by rw [add_def', cast_mkRat_of_ne_zero, cast_def, cast_def, mul_comm r.num, (Nat.cast_commute _ _).div_add_div (Nat.commute_cast _ _) hq hr] - push_cast - rfl + · push_cast + rfl · push_cast exact mul_ne_zero hq hr #align rat.cast_add_of_ne_zero Rat.cast_add_of_ne_zero @@ -196,8 +196,8 @@ lemma cast_add_of_ne_zero {q r : ℚ} (hq : (q.den : α) ≠ 0) (hr : (r.den : ↑(p * q) = (p * q : α) := by rw [mul_eq_mkRat, cast_mkRat_of_ne_zero, cast_def, cast_def, (Nat.commute_cast _ _).div_mul_div_comm (Int.commute_cast _ _)] - push_cast - rfl + · push_cast + rfl · push_cast exact mul_ne_zero hp hq #align rat.cast_mul_of_ne_zero Rat.cast_mul_of_ne_zero @@ -211,8 +211,8 @@ lemma cast_inv_of_ne_zero (hq : (q.num : α) ≠ 0) : ↑(q⁻¹) = (q⁻¹ : α ↑(p / q) = (p / q : α) := by rw [div_def', cast_divInt_of_ne_zero, cast_def, cast_def, div_eq_mul_inv (_ / _), inv_div, (Int.commute_cast _ _).div_mul_div_comm (Nat.commute_cast _ _)] - push_cast - rfl + · push_cast + rfl · push_cast exact mul_ne_zero hp hq #align rat.cast_div_of_ne_zero Rat.cast_div_of_ne_zero diff --git a/Mathlib/Data/Rat/Defs.lean b/Mathlib/Data/Rat/Defs.lean index 8fd1f774ceb314..d0a24b60e3078a 100644 --- a/Mathlib/Data/Rat/Defs.lean +++ b/Mathlib/Data/Rat/Defs.lean @@ -499,7 +499,7 @@ theorem divInt_div_divInt_cancel_right {x : ℤ} (hx : x ≠ 0) (n d : ℤ) : #align rat.mk_div_mk_cancel_right Rat.divInt_div_divInt_cancel_right theorem intCast_div_eq_divInt {n d : ℤ} : (n : ℚ) / (d) = n /. d := by - repeat' rw [intCast_eq_divInt] + repeat rw [intCast_eq_divInt] exact divInt_div_divInt_cancel_left one_ne_zero n d #align rat.coe_int_div_eq_mk Rat.intCast_div_eq_divInt diff --git a/Mathlib/Data/Rat/Lemmas.lean b/Mathlib/Data/Rat/Lemmas.lean index 2e0fd0f06d692f..7c7c960aac6d1d 100644 --- a/Mathlib/Data/Rat/Lemmas.lean +++ b/Mathlib/Data/Rat/Lemmas.lean @@ -130,7 +130,7 @@ theorem mul_num_den' (q r : ℚ) : exists_eq_mul_div_num_and_eq_mul_div_den (q.num * r.num) hs rw [c_mul_num, mul_assoc, mul_comm] nth_rw 1 [c_mul_den] - repeat' rw [Int.mul_assoc] + repeat rw [Int.mul_assoc] apply mul_eq_mul_left_iff.2 rw [or_iff_not_imp_right] intro @@ -147,7 +147,7 @@ theorem add_num_den' (q r : ℚ) : exists_eq_mul_div_num_and_eq_mul_div_den (q.num * r.den + r.num * q.den) hs rw [c_mul_num, mul_assoc, mul_comm] nth_rw 1 [c_mul_den] - repeat' rw [Int.mul_assoc] + repeat rw [Int.mul_assoc] apply mul_eq_mul_left_iff.2 rw [or_iff_not_imp_right] intro diff --git a/Mathlib/Data/Seq/Parallel.lean b/Mathlib/Data/Seq/Parallel.lean index 049c7d689dd8ee..31a8a619cbc641 100644 --- a/Mathlib/Data/Seq/Parallel.lean +++ b/Mathlib/Data/Seq/Parallel.lean @@ -143,8 +143,8 @@ theorem terminates_parallel {S : WSeq (Computation α)} {c} (h : c ∈ S) [T : T infer_instance · have C : corec parallel.aux1 (l, S) = _ := by apply destruct_eq_think - simp only [corec_eq, rmap, parallel.aux1._eq_1] - rw [h, H] + · simp only [corec_eq, rmap, parallel.aux1._eq_1] + rw [h, H] rw [C] refine @Computation.think_terminates _ _ ?_ apply terminates_parallel.aux _ T @@ -162,8 +162,8 @@ theorem terminates_parallel {S : WSeq (Computation α)} {c} (h : c ∈ S) [T : T infer_instance · have C : corec parallel.aux1 (l, S) = _ := by apply destruct_eq_think - simp only [corec_eq, rmap, parallel.aux1._eq_1] - rw [h] + · simp only [corec_eq, rmap, parallel.aux1._eq_1] + rw [h] rw [C] refine @Computation.think_terminates _ _ ?_ have TT : ∀ l', Terminates (corec parallel.aux1 (l', S.tail)) := by @@ -199,8 +199,8 @@ theorem exists_of_mem_parallel {S : WSeq (Computation α)} {a} (h : a ∈ parall let F : List (Computation α) → Sum α (List (Computation α)) → Prop := by intro l a cases' a with a l' - exact ∃ c ∈ l, a ∈ c - exact ∀ a', (∃ c ∈ l', a' ∈ c) → ∃ c ∈ l, a' ∈ c + · exact ∃ c ∈ l, a ∈ c + · exact ∀ a', (∃ c ∈ l', a' ∈ c) → ∃ c ∈ l, a' ∈ c have lem1 : ∀ l : List (Computation α), F l (parallel.aux2 l) := by intro l induction' l with c l IH <;> simp only [parallel.aux2, List.foldr] @@ -333,8 +333,8 @@ def parallelRec {S : WSeq (Computation α)} (C : α → Sort v) (H : ∀ s ∈ S rcases exists_of_mem_map ad' with ⟨a', ac', e'⟩ injection e' with i1 i2 constructor - rwa [i1, i2] at ac' - rwa [i2] at cs' + · rwa [i1, i2] at ac' + · rwa [i2] at cs' cases' this with ac cs apply H _ cs _ ac #align computation.parallel_rec Computation.parallelRec diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index abbb963691f974..4c5f1139d4b0ce 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -377,9 +377,9 @@ def IsBisimulation := theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s₁ = s₂ := by apply Subtype.eq apply Stream'.eq_of_bisim fun x y => ∃ s s' : Seq α, s.1 = x ∧ s'.1 = y ∧ R s s' - dsimp [Stream'.IsBisimulation] - intro t₁ t₂ e - exact + · dsimp [Stream'.IsBisimulation] + intro t₁ t₂ e + exact match t₁, t₂, e with | _, _, ⟨s, s', rfl, rfl, r⟩ => by suffices head s = head s' ∧ R (tail s) (tail s') from @@ -401,9 +401,9 @@ theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s rw [head_cons, head_cons, tail_cons, tail_cons] cases' this with h1 h2 constructor - rw [h1] - exact h2 - exact ⟨s₁, s₂, rfl, rfl, r⟩ + · rw [h1] + · exact h2 + · exact ⟨s₁, s₂, rfl, rfl, r⟩ #align stream.seq.eq_of_bisim Stream'.Seq.eq_of_bisim end Bisim diff --git a/Mathlib/Data/Seq/WSeq.lean b/Mathlib/Data/Seq/WSeq.lean index 55d0e46c3f4ad6..96f8d57d8efe4d 100644 --- a/Mathlib/Data/Seq/WSeq.lean +++ b/Mathlib/Data/Seq/WSeq.lean @@ -509,8 +509,8 @@ theorem liftRel_destruct_iff {R : α → β → Prop} {s : WSeq α} {t : WSeq β Or.inr h, fun {s t} h => by have h : Computation.LiftRel (LiftRelO R (LiftRel R)) (destruct s) (destruct t) := by cases' h with h h - exact liftRel_destruct h - assumption + · exact liftRel_destruct h + · assumption apply Computation.LiftRel.imp _ _ _ h intro a b apply LiftRelO.imp_right @@ -1041,7 +1041,7 @@ theorem liftRel_dropn_destruct {R : α → β → Prop} {s t} (H : LiftRel R s t | n + 1 => by simp only [LiftRelO, drop, Nat.add_eq, Nat.add_zero, destruct_tail, tail.aux] apply liftRel_bind - apply liftRel_dropn_destruct H n + · apply liftRel_dropn_destruct H n exact fun {a b} o => match a, b, o with | none, none, _ => by diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 88d66e149a0bf5..1f87ce0b9ccf08 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -1055,10 +1055,10 @@ theorem surjective_onto_range : Surjective (rangeFactorization f) := fun ⟨_, theorem image_eq_range (f : α → β) (s : Set α) : f '' s = range fun x : s => f x := by ext constructor - rintro ⟨x, h1, h2⟩ - exact ⟨⟨x, h1⟩, h2⟩ - rintro ⟨⟨x, h1⟩, h2⟩ - exact ⟨x, h1, h2⟩ + · rintro ⟨x, h1, h2⟩ + exact ⟨⟨x, h1⟩, h2⟩ + · rintro ⟨⟨x, h1⟩, h2⟩ + exact ⟨x, h1, h2⟩ #align set.image_eq_range Set.image_eq_range theorem _root_.Sum.range_eq (f : Sum α β → γ) : @@ -1083,8 +1083,8 @@ theorem range_ite_subset' {p : Prop} [Decidable p] {f g : α → β} : theorem range_ite_subset {p : α → Prop} [DecidablePred p] {f g : α → β} : (range fun x => if p x then f x else g x) ⊆ range f ∪ range g := by rw [range_subset_iff]; intro x; by_cases h : p x - simp only [if_pos h, mem_union, mem_range, exists_apply_eq_apply, true_or] - simp [if_neg h, mem_union, mem_range_self] + · simp only [if_pos h, mem_union, mem_range, exists_apply_eq_apply, true_or] + · simp [if_neg h, mem_union, mem_range_self] #align set.range_ite_subset Set.range_ite_subset @[simp] diff --git a/Mathlib/Data/Stream/Init.lean b/Mathlib/Data/Stream/Init.lean index 59766ad360edff..306f0d13d3e2ab 100644 --- a/Mathlib/Data/Stream/Init.lean +++ b/Mathlib/Data/Stream/Init.lean @@ -326,7 +326,9 @@ theorem bisim_simple (s₁ s₂ : Stream' α) : head s₁ = head s₂ → s₁ = tail s₁ → s₂ = tail s₂ → s₁ = s₂ := fun hh ht₁ ht₂ => eq_of_bisim (fun s₁ s₂ => head s₁ = head s₂ ∧ s₁ = tail s₁ ∧ s₂ = tail s₂) (fun s₁ s₂ ⟨h₁, h₂, h₃⟩ => by - constructor; exact h₁; rw [← h₂, ← h₃] + constructor + · exact h₁ + rw [← h₂, ← h₃] (repeat' constructor) <;> assumption) (And.intro hh (And.intro ht₁ ht₂)) #align stream.bisim_simple Stream'.bisim_simple @@ -761,7 +763,9 @@ theorem get_nats (n : Nat) : get nats n = n := theorem nats_eq : nats = cons 0 (map succ nats) := by apply Stream'.ext; intro n - cases n; rfl; rw [get_succ]; rfl + cases n + · rfl + rw [get_succ]; rfl #align stream.nats_eq Stream'.nats_eq end Stream' diff --git a/Mathlib/Data/Sum/Interval.lean b/Mathlib/Data/Sum/Interval.lean index 512777a3909420..40e405677ce159 100644 --- a/Mathlib/Data/Sum/Interval.lean +++ b/Mathlib/Data/Sum/Interval.lean @@ -60,7 +60,7 @@ theorem mem_sumLift₂ : theorem inl_mem_sumLift₂ {c₁ : γ₁} : inl c₁ ∈ sumLift₂ f g a b ↔ ∃ a₁ b₁, a = inl a₁ ∧ b = inl b₁ ∧ c₁ ∈ f a₁ b₁ := by rw [mem_sumLift₂, or_iff_left] - simp only [inl.injEq, exists_and_left, exists_eq_left'] + · simp only [inl.injEq, exists_and_left, exists_eq_left'] rintro ⟨_, _, c₂, _, _, h, _⟩ exact inl_ne_inr h #align finset.inl_mem_sum_lift₂ Finset.inl_mem_sumLift₂ @@ -68,7 +68,7 @@ theorem inl_mem_sumLift₂ {c₁ : γ₁} : theorem inr_mem_sumLift₂ {c₂ : γ₂} : inr c₂ ∈ sumLift₂ f g a b ↔ ∃ a₂ b₂, a = inr a₂ ∧ b = inr b₂ ∧ c₂ ∈ g a₂ b₂ := by rw [mem_sumLift₂, or_iff_right] - simp only [inr.injEq, exists_and_left, exists_eq_left'] + · simp only [inr.injEq, exists_and_left, exists_eq_left'] rintro ⟨_, _, c₂, _, _, h, _⟩ exact inr_ne_inl h #align finset.inr_mem_sum_lift₂ Finset.inr_mem_sumLift₂ diff --git a/Mathlib/Data/Sum/Order.lean b/Mathlib/Data/Sum/Order.lean index f7682e56ebac9b..35d019debe755a 100644 --- a/Mathlib/Data/Sum/Order.lean +++ b/Mathlib/Data/Sum/Order.lean @@ -688,8 +688,9 @@ theorem sumLexAssoc_symm_apply_inr_inr : (sumLexAssoc α β γ).symm (inr (inr c def sumLexDualAntidistrib (α β : Type*) [LE α] [LE β] : (α ⊕ₗ β)ᵒᵈ ≃o βᵒᵈ ⊕ₗ αᵒᵈ := { Equiv.sumComm α β with map_rel_iff' := @fun a b => by - rcases a with (a | a) <;> rcases b with (b | b); simp - · change + rcases a with (a | a) <;> rcases b with (b | b); + · simp + change toLex (inr <| toDual a) ≤ toLex (inr <| toDual b) ↔ toDual (toLex <| inl a) ≤ toDual (toLex <| inl b) simp [toDual_le_toDual, Lex.inl_le_inl_iff, Lex.inr_le_inr_iff] diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index 65b301a2c1984e..9f1e37587cea73 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -530,7 +530,7 @@ end theorem prod_fst_mk {α β : TypeVec n} (i : Fin2 n) (a : α i) (b : β i) : TypeVec.prod.fst i (prod.mk i a b) = a := by induction' i with _ _ _ i_ih - simp_all only [prod.fst, prod.mk] + · simp_all only [prod.fst, prod.mk] apply i_ih #align typevec.prod_fst_mk TypeVec.prod_fst_mk @@ -538,7 +538,7 @@ theorem prod_fst_mk {α β : TypeVec n} (i : Fin2 n) (a : α i) (b : β i) : theorem prod_snd_mk {α β : TypeVec n} (i : Fin2 n) (a : α i) (b : β i) : TypeVec.prod.snd i (prod.mk i a b) = b := by induction' i with _ _ _ i_ih - simp_all [prod.snd, prod.mk] + · simp_all [prod.snd, prod.mk] apply i_ih #align typevec.prod_snd_mk TypeVec.prod_snd_mk @@ -582,7 +582,7 @@ theorem snd_diag {α : TypeVec n} : TypeVec.prod.snd ⊚ (prod.diag : α ⟹ _) theorem repeatEq_iff_eq {α : TypeVec n} {i x y} : ofRepeat (repeatEq α i (prod.mk _ x y)) ↔ x = y := by induction' i with _ _ _ i_ih - rfl + · rfl erw [repeatEq, i_ih] #align typevec.repeat_eq_iff_eq TypeVec.repeatEq_iff_eq @@ -645,7 +645,7 @@ theorem subtypeVal_nil {α : TypeVec.{u} 0} (ps : α ⟹ «repeat» 0 Prop) : theorem diag_sub_val {n} {α : TypeVec.{u} n} : subtypeVal (repeatEq α) ⊚ diagSub = prod.diag := by ext i x induction' i with _ _ _ i_ih - simp only [comp, subtypeVal, repeatEq._eq_2, diagSub, prod.diag] + · simp only [comp, subtypeVal, repeatEq._eq_2, diagSub, prod.diag] apply @i_ih (drop α) #align typevec.diag_sub_val TypeVec.diag_sub_val @@ -762,8 +762,8 @@ theorem prod_map_id {α β : TypeVec n} : (@TypeVec.id _ α ⊗' @TypeVec.id _ ext i x : 2 induction i <;> simp only [TypeVec.prod.map, *, dropFun_id] cases x - rfl - rfl + · rfl + · rfl #align typevec.prod_map_id TypeVec.prod_map_id @[simp] diff --git a/Mathlib/Data/UInt.lean b/Mathlib/Data/UInt.lean index f17bc2830191cb..5df46004ac8483 100644 --- a/Mathlib/Data/UInt.lean +++ b/Mathlib/Data/UInt.lean @@ -130,8 +130,8 @@ def isAlphanum (c : UInt8) : Bool := theorem toChar_aux (n : Nat) (h : n < size) : Nat.isValidChar (UInt32.ofNat n).1 := by rw [UInt32.val_eq_of_lt] - exact Or.inl <| Nat.lt_trans h <| by decide - exact Nat.lt_trans h <| by decide + · exact Or.inl <| Nat.lt_trans h <| by decide + · exact Nat.lt_trans h <| by decide /-- The numbers from 0 to 256 are all valid UTF-8 characters, so we can embed one in the other. -/ def toChar (n : UInt8) : Char := ⟨n.toUInt32, toChar_aux n.1 n.1.2⟩ diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 6ab33270b43ae9..a2e329b439e682 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -702,8 +702,8 @@ protected theorem comp_traverse (f : β → F γ) (g : α → G β) (x : Vector Vector.traverse (Comp.mk ∘ Functor.map f ∘ g) x = Comp.mk (Vector.traverse f <$> Vector.traverse g x) := by induction' x using Vector.inductionOn with n x xs ih - simp! [cast, *, functor_norm] - · rfl + · simp! [cast, *, functor_norm] + rfl · rw [Vector.traverse_def, ih] simp [functor_norm, (· ∘ ·)] #align vector.comp_traverse Vector.comp_traverse diff --git a/Mathlib/Data/Vector3.lean b/Mathlib/Data/Vector3.lean index 656875438d0fff..0548ded40f0fd6 100644 --- a/Mathlib/Data/Vector3.lean +++ b/Mathlib/Data/Vector3.lean @@ -45,8 +45,8 @@ def nil : Vector3 α 0 := @[match_pattern] def cons (a : α) (v : Vector3 α n) : Vector3 α (succ n) := fun i => by refine' i.cases' _ _ - exact a - exact v + · exact a + · exact v #align vector3.cons Vector3.cons section @@ -203,7 +203,8 @@ theorem insert_fs (a : α) (b : α) (v : Vector3 α n) (i : Fin2 (succ n)) : theorem append_insert (a : α) (t : Vector3 α m) (v : Vector3 α n) (i : Fin2 (succ n)) (e : succ n + m = succ (n + m)) : insert a (t +-+ v) (Eq.recOn e (i.add m)) = Eq.recOn e (t +-+ insert a v i) := by - refine' Vector3.recOn t (fun e => _) (@fun k b t IH _ => _) e; rfl + refine' Vector3.recOn t (fun e => _) (@fun k b t IH _ => _) e + · rfl have e' := succ_add n k change insert a (b :: t +-+ v) (Eq.recOn (congr_arg succ e') (fs (add i k))) = diff --git a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean index 026d392d4b6fb6..71e4addbad1bd4 100644 --- a/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean +++ b/Mathlib/Dynamics/Ergodic/MeasurePreserving.lean @@ -153,7 +153,7 @@ open scoped symmDiff in lemma measure_symmDiff_preimage_iterate_le (hf : MeasurePreserving f μ μ) (hs : MeasurableSet s) (n : ℕ) : μ (s ∆ (f^[n] ⁻¹' s)) ≤ n • μ (s ∆ (f ⁻¹' s)) := by - induction' n with n ih; simp + induction' n with n ih; · simp simp only [add_smul, one_smul, ← n.add_one] refine' le_trans (measure_symmDiff_le s (f^[n] ⁻¹' s) (f^[n+1] ⁻¹' s)) (add_le_add ih _) replace hs : MeasurableSet (s ∆ (f ⁻¹' s)) := hs.symmDiff <| hf.measurable hs diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 8c2226162aa23e..7860a0b9ded69a 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -211,10 +211,10 @@ instance instZSMul : SMul ℤ Cₛ^n⟮I; F, V⟯ := @[simp] theorem coe_zsmul (s : Cₛ^n⟮I; F, V⟯) (z : ℤ) : ⇑(z • s : Cₛ^n⟮I; F, V⟯) = z • ⇑s := by cases' z with n n - refine' (coe_nsmul s n).trans _ - simp only [Int.ofNat_eq_coe, natCast_zsmul] - refine' (congr_arg Neg.neg (coe_nsmul s (n + 1))).trans _ - simp only [negSucc_zsmul, neg_inj] + · refine' (coe_nsmul s n).trans _ + simp only [Int.ofNat_eq_coe, natCast_zsmul] + · refine' (congr_arg Neg.neg (coe_nsmul s (n + 1))).trans _ + simp only [negSucc_zsmul, neg_inj] #align cont_mdiff_section.coe_zsmul ContMDiffSection.coe_zsmul instance instAddCommGroup : AddCommGroup Cₛ^n⟮I; F, V⟯ := diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index 96c5907f222f81..b321c06542f41d 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -352,10 +352,10 @@ instance tangentBundleCore.isSmooth : (tangentBundleCore I M).IsSmooth I := by refine' ⟨fun i j => _⟩ rw [SmoothOn, contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas I i.2), contMDiffOn_iff_contDiffOn] - refine' ((contDiffOn_fderiv_coord_change I i j).congr fun x hx => _).mono _ - · rw [PartialEquiv.trans_source'] at hx - simp_rw [Function.comp_apply, tangentBundleCore_coordChange, (i.1.extend I).right_inv hx.1] - · exact (i.1.extend_image_source_inter j.1 I).subset + · refine' ((contDiffOn_fderiv_coord_change I i j).congr fun x hx => _).mono _ + · rw [PartialEquiv.trans_source'] at hx + simp_rw [Function.comp_apply, tangentBundleCore_coordChange, (i.1.extend I).right_inv hx.1] + · exact (i.1.extend_image_source_inter j.1 I).subset · apply inter_subset_left #align tangent_bundle_core.is_smooth tangentBundleCore.isSmooth diff --git a/Mathlib/GroupTheory/Congruence.lean b/Mathlib/GroupTheory/Congruence.lean index dacacb474b43c9..7a47622946048a 100644 --- a/Mathlib/GroupTheory/Congruence.lean +++ b/Mathlib/GroupTheory/Congruence.lean @@ -1338,7 +1338,7 @@ def liftOnUnits (u : Units c.Quotient) (f : ∀ x y : M, c (x * y) 1 → c (y * f x y (c.eq.1 hxy) (c.eq.1 hyx)) (fun x y x' y' hx hy => _) u.3 u.4 refine' Function.hfunext _ _ - rw [c.eq.2 hx, c.eq.2 hy] + · rw [c.eq.2 hx, c.eq.2 hy] · rintro Hxy Hxy' - refine' Function.hfunext _ _ · rw [c.eq.2 hx, c.eq.2 hy] diff --git a/Mathlib/GroupTheory/CoprodI.lean b/Mathlib/GroupTheory/CoprodI.lean index c2cd1f0eb7a651..839f581902ea41 100644 --- a/Mathlib/GroupTheory/CoprodI.lean +++ b/Mathlib/GroupTheory/CoprodI.lean @@ -350,8 +350,8 @@ def cons {i} (m : M i) (w : Word M) (hmw : w.fstIdx ≠ some i) (h1 : m ≠ 1) : ne_one := by simp only [List.mem_cons] rintro l (rfl | hl) - exact h1 - exact w.ne_one l hl + · exact h1 + · exact w.ne_one l hl chain_ne := w.chain_ne.cons' (fstIdx_ne_iff.mp hmw) } /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ @@ -809,8 +809,8 @@ def replaceHead : ∀ {i j : ι} (x : M i) (_hnotone : x ≠ 1) (_w : NeWord M i theorem replaceHead_head {i j : ι} (x : M i) (hnotone : x ≠ 1) (w : NeWord M i j) : (replaceHead x hnotone w).head = x := by induction w - rfl - simp [*] + · rfl + · simp [*] #align free_product.neword.replace_head_head Monoid.CoprodI.NeWord.replaceHead_head /-- One can multiply an element from the left to a non-empty reduced word if it does not cancel @@ -823,8 +823,8 @@ def mulHead {i j : ι} (w : NeWord M i j) (x : M i) (hnotone : x * w.head ≠ 1) theorem mulHead_head {i j : ι} (w : NeWord M i j) (x : M i) (hnotone : x * w.head ≠ 1) : (mulHead w x hnotone).head = x * w.head := by induction w - rfl - simp [*] + · rfl + · simp [*] #align free_product.neword.mul_head_head Monoid.CoprodI.NeWord.mulHead_head @[simp] @@ -1018,8 +1018,8 @@ instance {ι : Type*} (G : ι → Type*) [∀ i, Group (G i)] [∀ i, IsFreeGrou def _root_.freeGroupEquivCoprodI {ι : Type u_1} : FreeGroup ι ≃* CoprodI fun _ : ι => FreeGroup Unit := by refine' MonoidHom.toMulEquiv _ _ _ _ - exact FreeGroup.lift fun i => @CoprodI.of ι _ _ i (FreeGroup.of Unit.unit) - exact CoprodI.lift fun i => FreeGroup.lift fun _ => FreeGroup.of i + · exact FreeGroup.lift fun i => @CoprodI.of ι _ _ i (FreeGroup.of Unit.unit) + · exact CoprodI.lift fun i => FreeGroup.lift fun _ => FreeGroup.of i · ext; simp · ext i a; cases a; simp #align free_group_equiv_free_product freeGroupEquivCoprodI @@ -1068,18 +1068,18 @@ theorem _root_.FreeGroup.injective_lift_of_ping_pong : Function.Injective (FreeG let f : ∀ i, H i →* G := fun i => FreeGroup.lift fun _ => a i let X' : ι → Set α := fun i => X i ∪ Y i apply lift_injective_of_ping_pong f _ X' - show ∀ i, (X' i).Nonempty - · exact fun i => Set.Nonempty.inl (hXnonempty i) - show Pairwise fun i j => Disjoint (X' i) (X' j) - · intro i j hij + · show ∀ i, (X' i).Nonempty + exact fun i => Set.Nonempty.inl (hXnonempty i) + · show Pairwise fun i j => Disjoint (X' i) (X' j) + intro i j hij simp only [X'] apply Disjoint.union_left <;> apply Disjoint.union_right · exact hXdisj hij · exact hXYdisj i j · exact (hXYdisj j i).symm · exact hYdisj hij - show Pairwise fun i j => ∀ h : H i, h ≠ 1 → f i h • X' j ⊆ X' i - · rintro i j hij + · show Pairwise fun i j => ∀ h : H i, h ≠ 1 → f i h • X' j ⊆ X' i + rintro i j hij -- use free_group unit ≃ ℤ refine' FreeGroup.freeGroupUnitEquivInt.forall_congr_left'.mpr _ intro n hne1 diff --git a/Mathlib/GroupTheory/MonoidLocalization.lean b/Mathlib/GroupTheory/MonoidLocalization.lean index daadf3d788d079..f9990e1c26446a 100644 --- a/Mathlib/GroupTheory/MonoidLocalization.lean +++ b/Mathlib/GroupTheory/MonoidLocalization.lean @@ -961,7 +961,7 @@ noncomputable def lift : N →* P where dsimp only rw [mul_inv_left hg, ← mul_assoc, ← mul_assoc, mul_inv_right hg, mul_comm _ (g (f.sec y).1), ← mul_assoc, ← mul_assoc, mul_inv_right hg] - repeat' rw [← g.map_mul] + repeat rw [← g.map_mul] exact f.eq_of_eq hg (by simp_rw [f.toMap.map_mul, sec_spec']; ac_rfl) #align submonoid.localization_map.lift Submonoid.LocalizationMap.lift #align add_submonoid.localization_map.lift AddSubmonoid.LocalizationMap.lift @@ -2099,9 +2099,10 @@ instance partialOrder : PartialOrder (Localization s) where rwa [mul_left_comm, mul_left_comm (b.2 : α), mul_le_mul_iff_left] le_antisymm a b := by induction' a using Localization.rec with a₁ a₂ - induction' b using Localization.rec with b₁ b₂ - simp_rw [mk_le_mk, mk_eq_mk_iff, r_iff_exists] - exact fun hab hba => ⟨1, by rw [hab.antisymm hba]⟩ + on_goal 1 => + induction' b using Localization.rec with b₁ b₂ + · simp_rw [mk_le_mk, mk_eq_mk_iff, r_iff_exists] + exact fun hab hba => ⟨1, by rw [hab.antisymm hba]⟩ all_goals rfl lt_iff_le_not_le a b := Localization.induction_on₂ a b fun a b => lt_iff_le_not_le diff --git a/Mathlib/GroupTheory/Perm/ClosureSwap.lean b/Mathlib/GroupTheory/Perm/ClosureSwap.lean index daf9ba2d9c37d2..f1e2739089fd38 100644 --- a/Mathlib/GroupTheory/Perm/ClosureSwap.lean +++ b/Mathlib/GroupTheory/Perm/ClosureSwap.lean @@ -47,8 +47,10 @@ theorem Equiv.Perm.IsSwap.finite_compl_fixedBy {σ : Perm α} (h : σ.IsSwap) : theorem SubmonoidClass.swap_mem_trans {a b c : α} {C} [SetLike C (Perm α)] [SubmonoidClass C (Perm α)] (M : C) (hab : swap a b ∈ M) (hbc : swap b c ∈ M) : swap a c ∈ M := by - obtain rfl | hab' := eq_or_ne a b; exact hbc - obtain rfl | hac := eq_or_ne a c; exact swap_self a ▸ one_mem M + obtain rfl | hab' := eq_or_ne a b + · exact hbc + obtain rfl | hac := eq_or_ne a c + · exact swap_self a ▸ one_mem M rw [swap_comm, ← swap_mul_swap_mul_swap hab' hac] exact mul_mem (mul_mem hbc hab) hbc @@ -107,7 +109,8 @@ theorem mem_closure_isSwap {S : Set (Perm α)} (hS : ∀ f ∈ S, f.IsSwap) {f : · contrapose! hb simp_rw [not_mem_compl_iff, mem_fixedBy, Perm.smul_def, Perm.mul_apply, swap_apply_def, apply_eq_iff_eq] - by_cases hb' : f b = b; rw [hb']; split_ifs with h <;> simp only [h] + by_cases hb' : f b = b + · rw [hb']; split_ifs with h <;> simp only [h] simp [show b = a by simpa [hb] using supp_subset hb'] /-- A permutation is a product of transpositions if and only if it has finite support. -/ diff --git a/Mathlib/GroupTheory/Perm/Option.lean b/Mathlib/GroupTheory/Perm/Option.lean index d4d2e3ca94bc4a..239791126233a5 100644 --- a/Mathlib/GroupTheory/Perm/Option.lean +++ b/Mathlib/GroupTheory/Perm/Option.lean @@ -29,7 +29,7 @@ theorem Equiv.optionCongr_swap {α : Type*} [DecidableEq α] (x y : α) : ext (_ | i) · simp [swap_apply_of_ne_of_ne] · by_cases hx : i = x - simp only [hx, optionCongr_apply, Option.map_some', swap_apply_left, Option.mem_def, + · simp only [hx, optionCongr_apply, Option.map_some', swap_apply_left, Option.mem_def, Option.some.injEq] by_cases hy : i = y <;> simp [hx, hy, swap_apply_of_ne_of_ne] #align equiv.option_congr_swap Equiv.optionCongr_swap diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index 5337213e45e8c4..dd8bcd9a78420f 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -226,10 +226,10 @@ theorem ofSubtype_swap_eq {p : α → Prop} [DecidablePred p] (x y : Subtype p) · rw [swap_apply_of_ne_of_ne] <;> simp [Subtype.ext_iff, *] · rw [ofSubtype_apply_of_not_mem _ hz, swap_apply_of_ne_of_ne] - intro h - apply hz - rw [h] - exact Subtype.prop x + · intro h + apply hz + rw [h] + exact Subtype.prop x intro h apply hz rw [h] @@ -343,7 +343,7 @@ theorem exists_mem_support_of_mem_support_prod {l : List (Perm α)} {x : α} induction' l with f l ih · rfl · rw [List.prod_cons, mul_apply, ih, hx] - simp only [List.find?, List.mem_cons, true_or] + · simp only [List.find?, List.mem_cons, true_or] intros f' hf' refine' hx f' _ simp only [List.find?, List.mem_cons] @@ -506,8 +506,8 @@ theorem mem_support_swap_mul_imp_mem_support_ne {x y : α} (hy : y ∈ support ( · constructor <;> intro <;> simp_all only [if_true, eq_self_iff_true, not_true, Ne] · split_ifs at hy with hf heq <;> simp_all only [not_true] - exact ⟨h, hy⟩ - exact ⟨hy, heq⟩ + · exact ⟨h, hy⟩ + · exact ⟨hy, heq⟩ #align equiv.perm.mem_support_swap_mul_imp_mem_support_ne Equiv.Perm.mem_support_swap_mul_imp_mem_support_ne theorem Disjoint.mem_imp (h : Disjoint f g) {x : α} (hx : x ∈ f.support) : x ∉ g.support := diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 3b48cd739c4a52..51d36fa2d3b05c 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -3123,7 +3123,7 @@ theorem comap_normalizer_eq_of_injective_of_le_range {N : Type*} [Group N] (H : theorem subgroupOf_normalizer_eq {H N : Subgroup G} (h : H.normalizer ≤ N) : H.normalizer.subgroupOf N = (H.subgroupOf N).normalizer := by apply comap_normalizer_eq_of_injective_of_le_range - exact Subtype.coe_injective + · exact Subtype.coe_injective simpa #align subgroup.subgroup_of_normalizer_eq Subgroup.subgroupOf_normalizer_eq #align add_subgroup.add_subgroup_of_normalizer_eq AddSubgroup.addSubgroupOf_normalizer_eq diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 44d04cc222267b..00fe12210ff16f 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -369,7 +369,8 @@ theorem independent_generalizedEigenspace [NoZeroSMulDivisors R M] (f : End R M) Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s - induction' s using Finset.induction_on with μ₂ s _ ih; simp + induction' s using Finset.induction_on with μ₂ s _ ih + · simp intro hμ₁₂ obtain ⟨hμ₁₂ : μ₁ ≠ μ₂, hμ₁ : μ₁ ∉ s⟩ := by rwa [Finset.mem_insert, not_or] at hμ₁₂ specialize ih hμ₁ diff --git a/Mathlib/LinearAlgebra/FiniteDimensional.lean b/Mathlib/LinearAlgebra/FiniteDimensional.lean index 9987a6bb723f91..11164142ef1315 100644 --- a/Mathlib/LinearAlgebra/FiniteDimensional.lean +++ b/Mathlib/LinearAlgebra/FiniteDimensional.lean @@ -157,11 +157,11 @@ theorem of_finite_basis {ι : Type w} {s : Set ι} (h : Basis s K V) (hs : Set.F instance finiteDimensional_submodule [FiniteDimensional K V] (S : Submodule K V) : FiniteDimensional K S := by letI : IsNoetherian K V := iff_fg.2 ?_ - exact - iff_fg.1 - (IsNoetherian.iff_rank_lt_aleph0.2 - (lt_of_le_of_lt (rank_submodule_le _) (_root_.rank_lt_aleph0 K V))) - infer_instance + · exact + iff_fg.1 + (IsNoetherian.iff_rank_lt_aleph0.2 + (lt_of_le_of_lt (rank_submodule_le _) (_root_.rank_lt_aleph0 K V))) + · infer_instance #align finite_dimensional.finite_dimensional_submodule FiniteDimensional.finiteDimensional_submodule /-- A quotient of a finite-dimensional space is also finite-dimensional. -/ @@ -487,7 +487,7 @@ theorem finrank_sup_add_finrank_inf_eq (s t : Submodule K V) [FiniteDimensional finrank K ↑(s ⊔ t) + finrank K ↑(s ⊓ t) = finrank K ↑s + finrank K ↑t := by have key : Module.rank K ↑(s ⊔ t) + Module.rank K ↑(s ⊓ t) = Module.rank K s + Module.rank K t := rank_sup_add_rank_inf_eq s t - repeat' rw [← finrank_eq_rank] at key + repeat rw [← finrank_eq_rank] at key norm_cast at key #align submodule.finrank_sup_add_finrank_inf_eq Submodule.finrank_sup_add_finrank_inf_eq diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 2115abc5c0272d..47436f55e6a076 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -835,8 +835,8 @@ theorem linearIndependent_iUnion_finite {η : Type*} {ιs : η → Type*} {f : apply LinearIndependent.of_subtype_range · rintro ⟨x₁, x₂⟩ ⟨y₁, y₂⟩ hxy by_cases h_cases : x₁ = y₁ - subst h_cases - · apply Sigma.eq + · subst h_cases + apply Sigma.eq rw [LinearIndependent.injective (hindep _) hxy] rfl · have h0 : f x₁ x₂ = 0 := by @@ -866,7 +866,7 @@ def LinearIndependent.totalEquiv (hv : LinearIndependent R v) : apply LinearEquiv.ofBijective (LinearMap.codRestrict (span R (range v)) (Finsupp.total ι M R v) _) constructor · rw [← LinearMap.ker_eq_bot, LinearMap.ker_codRestrict] - apply hv + · apply hv · intro l rw [← Finsupp.range_total] rw [LinearMap.mem_range] diff --git a/Mathlib/LinearAlgebra/Matrix/Block.lean b/Mathlib/LinearAlgebra/Matrix/Block.lean index 70cca037a819eb..52ef160bfb4f9e 100644 --- a/Mathlib/LinearAlgebra/Matrix/Block.lean +++ b/Mathlib/LinearAlgebra/Matrix/Block.lean @@ -211,9 +211,9 @@ theorem twoBlockTriangular_det' (M : Matrix m m R) (p : m → Prop) [DecidablePr (h : ∀ i, p i → ∀ j, ¬p j → M i j = 0) : M.det = (toSquareBlockProp M p).det * (toSquareBlockProp M fun i => ¬p i).det := by rw [M.twoBlockTriangular_det fun i => ¬p i, mul_comm] - congr 1 - exact equiv_block_det _ fun _ => not_not.symm - simpa only [Classical.not_not] using h + · congr 1 + exact equiv_block_det _ fun _ => not_not.symm + · simpa only [Classical.not_not] using h #align matrix.two_block_triangular_det' Matrix.twoBlockTriangular_det' protected theorem BlockTriangular.det [DecidableEq α] [LinearOrder α] (hM : BlockTriangular M b) : diff --git a/Mathlib/Logic/Relation.lean b/Mathlib/Logic/Relation.lean index 3e1ea55bd04c03..8e438c31631fc2 100644 --- a/Mathlib/Logic/Relation.lean +++ b/Mathlib/Logic/Relation.lean @@ -160,16 +160,16 @@ theorem comp_assoc : (r ∘r p) ∘r q = r ∘r p ∘r q := by funext a d apply propext constructor - exact fun ⟨c, ⟨b, hab, hbc⟩, hcd⟩ ↦ ⟨b, hab, c, hbc, hcd⟩ - exact fun ⟨b, hab, c, hbc, hcd⟩ ↦ ⟨c, ⟨b, hab, hbc⟩, hcd⟩ + · exact fun ⟨c, ⟨b, hab, hbc⟩, hcd⟩ ↦ ⟨b, hab, c, hbc, hcd⟩ + · exact fun ⟨b, hab, c, hbc, hcd⟩ ↦ ⟨c, ⟨b, hab, hbc⟩, hcd⟩ #align relation.comp_assoc Relation.comp_assoc theorem flip_comp : flip (r ∘r p) = flip p ∘r flip r := by funext c a apply propext constructor - exact fun ⟨b, hab, hbc⟩ ↦ ⟨b, hbc, hab⟩ - exact fun ⟨b, hbc, hab⟩ ↦ ⟨b, hab, hbc⟩ + · exact fun ⟨b, hab, hbc⟩ ↦ ⟨b, hbc, hab⟩ + · exact fun ⟨b, hbc, hab⟩ ↦ ⟨b, hab, hbc⟩ #align relation.flip_comp Relation.flip_comp end Comp @@ -375,8 +375,8 @@ namespace TransGen theorem to_reflTransGen {a b} (h : TransGen r a b) : ReflTransGen r a b := by induction' h with b h b c _ bc ab - exact ReflTransGen.single h - exact ReflTransGen.tail ab bc + · exact ReflTransGen.single h + · exact ReflTransGen.tail ab bc -- Porting note: in Lean 3 this function was called `to_refl` which seems wrong. #align relation.trans_gen.to_refl Relation.TransGen.to_reflTransGen diff --git a/Mathlib/MeasureTheory/Decomposition/Jordan.lean b/Mathlib/MeasureTheory/Decomposition/Jordan.lean index d4159206fdfd1f..6ecaf800b86fef 100644 --- a/Mathlib/MeasureTheory/Decomposition/Jordan.lean +++ b/Mathlib/MeasureTheory/Decomposition/Jordan.lean @@ -306,12 +306,13 @@ between the two sets. -/ theorem of_diff_eq_zero_of_symmDiff_eq_zero_positive (hu : MeasurableSet u) (hv : MeasurableSet v) (hsu : 0 ≤[u] s) (hsv : 0 ≤[v] s) (hs : s (u ∆ v) = 0) : s (u \ v) = 0 ∧ s (v \ u) = 0 := by rw [restrict_le_restrict_iff] at hsu hsv - have a := hsu (hu.diff hv) (u.diff_subset v) - have b := hsv (hv.diff hu) (v.diff_subset u) - erw [of_union (Set.disjoint_of_subset_left (u.diff_subset v) disjoint_sdiff_self_right) - (hu.diff hv) (hv.diff hu)] at hs - rw [zero_apply] at a b - constructor + on_goal 1 => + have a := hsu (hu.diff hv) (u.diff_subset v) + have b := hsv (hv.diff hu) (v.diff_subset u) + erw [of_union (Set.disjoint_of_subset_left (u.diff_subset v) disjoint_sdiff_self_right) + (hu.diff hv) (hv.diff hu)] at hs + rw [zero_apply] at a b + constructor all_goals first | linarith | assumption #align measure_theory.signed_measure.of_diff_eq_zero_of_symm_diff_eq_zero_positive MeasureTheory.SignedMeasure.of_diff_eq_zero_of_symmDiff_eq_zero_positive diff --git a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean index 45ae6e520a71f5..9aee41bc18d8dd 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedHahn.lean @@ -279,7 +279,7 @@ theorem exists_subset_restrict_nonpos (hi : s i < 0) : exact fun _ => Nat.lt_succ_iff.symm have h₁ : s i = s A + ∑' l, s (restrictNonposSeq s i l) := by rw [hA, ← s.of_disjoint_iUnion_nat, add_comm, of_add_of_diff] - exact MeasurableSet.iUnion fun _ => restrictNonposSeq_measurableSet _ + · exact MeasurableSet.iUnion fun _ => restrictNonposSeq_measurableSet _ exacts [hi₁, Set.iUnion_subset fun _ => restrictNonposSeq_subset _, fun _ => restrictNonposSeq_measurableSet _, restrictNonposSeq_disjoint] have h₂ : s A ≤ s i := by diff --git a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean index cd49a571e94ae6..ce40d15e6bb82e 100644 --- a/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean +++ b/Mathlib/MeasureTheory/Decomposition/SignedLebesgue.lean @@ -218,11 +218,11 @@ theorem singularPart_add_withDensity_rnDeriv_eq [s.HaveLebesgueDecomposition μ] add_assoc (-(s.toJordanDecomposition.negPart.singularPart μ).toSignedMeasure), ← toSignedMeasure_add, add_comm, ← add_assoc, ← neg_add, ← toSignedMeasure_add, add_comm, ← sub_eq_add_neg] - convert rfl - -- `convert rfl` much faster than `congr` - · exact s.toJordanDecomposition.posPart.haveLebesgueDecomposition_add μ - · rw [add_comm] - exact s.toJordanDecomposition.negPart.haveLebesgueDecomposition_add μ + · convert rfl + -- `convert rfl` much faster than `congr` + · exact s.toJordanDecomposition.posPart.haveLebesgueDecomposition_add μ + · rw [add_comm] + exact s.toJordanDecomposition.negPart.haveLebesgueDecomposition_add μ all_goals first | exact (lintegral_rnDeriv_lt_top _ _).ne diff --git a/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean b/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean index 2c3a4fcc7f24d6..93735b351e556e 100644 --- a/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean +++ b/Mathlib/MeasureTheory/Decomposition/UnsignedHahn.lean @@ -58,8 +58,8 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : intro s hm refine' Tendsto.sub _ _ <;> refine' NNReal.tendsto_coe.2 <| (ENNReal.tendsto_toNNReal _).comp <| tendsto_measure_iUnion hm - exact hμ _ - exact hν _ + · exact hμ _ + · exact hν _ have d_Inter : ∀ s : ℕ → Set α, (∀ n, MeasurableSet (s n)) → @@ -122,9 +122,9 @@ theorem hahn_decomposition [IsFiniteMeasure μ] [IsFiniteMeasure ν] : rw [f_succ _ _ hmn, d_split (f m n) (e (n + 1)) (hf _ _) (he₁ _), add_assoc] _ = d (e (n + 1) ∪ f m n) + d (f m (n + 1)) := by rw [d_split (e (n + 1) ∪ f m n) (e (n + 1)), union_diff_left, union_inter_cancel_left] - abel - exact (he₁ _).union (hf _ _) - exact he₁ _ + · abel + · exact (he₁ _).union (hf _ _) + · exact he₁ _ _ ≤ γ + d (f m (n + 1)) := add_le_add_right (d_le_γ _ <| (he₁ _).union (hf _ _)) _ exact (add_le_add_iff_left γ).1 this diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index 91b949e17d646e..17eaf85ef4f15e 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -186,10 +186,10 @@ theorem condexpL2_ae_eq_zero_of_ae_eq_zero (hs : MeasurableSet[m] s) (hμs : μ (hf : f =ᵐ[μ.restrict s] 0) : condexpL2 ℝ ℝ hm f =ᵐ[μ.restrict s] (0 : α → ℝ) := by suffices h_nnnorm_eq_zero : ∫⁻ x in s, ‖(condexpL2 ℝ ℝ hm f : α → ℝ) x‖₊ ∂μ = 0 by rw [lintegral_eq_zero_iff] at h_nnnorm_eq_zero - refine' h_nnnorm_eq_zero.mono fun x hx => _ - dsimp only at hx - rw [Pi.zero_apply] at hx ⊢ - · rwa [ENNReal.coe_eq_zero, nnnorm_eq_zero] at hx + · refine' h_nnnorm_eq_zero.mono fun x hx => _ + dsimp only at hx + rw [Pi.zero_apply] at hx ⊢ + · rwa [ENNReal.coe_eq_zero, nnnorm_eq_zero] at hx · refine' Measurable.coe_nnreal_ennreal (Measurable.nnnorm _) rw [lpMeas_coe] exact (Lp.stronglyMeasurable _).measurable diff --git a/Mathlib/MeasureTheory/Group/Convolution.lean b/Mathlib/MeasureTheory/Group/Convolution.lean index be4e768d4c0336..f1f7c0180bb451 100644 --- a/Mathlib/MeasureTheory/Group/Convolution.lean +++ b/Mathlib/MeasureTheory/Group/Convolution.lean @@ -42,7 +42,7 @@ theorem dirac_one_mconv [MeasurableMul₂ M] (μ : Measure M) [SFinite μ] : (Measure.dirac 1) ∗ μ = μ := by unfold mconv rw [MeasureTheory.Measure.dirac_prod, map_map] - simp only [Function.comp_def, one_mul, map_id'] + · simp only [Function.comp_def, one_mul, map_id'] all_goals { measurability } /-- Convolution of a measure μ with the dirac measure at 1 returns μ. -/ @@ -51,7 +51,7 @@ theorem mconv_dirac_one [MeasurableMul₂ M] (μ : Measure M) [SFinite μ] : μ ∗ (Measure.dirac 1) = μ := by unfold mconv rw [MeasureTheory.Measure.prod_dirac, map_map] - simp only [Function.comp_def, mul_one, map_id'] + · simp only [Function.comp_def, mul_one, map_id'] all_goals { measurability } /-- Convolution of the zero measure with a measure μ returns the zero measure. -/ diff --git a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean index cd78e685b51211..8c1567590ea69c 100644 --- a/Mathlib/MeasureTheory/Group/FundamentalDomain.lean +++ b/Mathlib/MeasureTheory/Group/FundamentalDomain.lean @@ -669,8 +669,8 @@ theorem pairwise_disjoint_fundamentalInterior : rintro _ ⟨x, hx, rfl⟩ ⟨y, hy, hxy⟩ rw [mem_fundamentalInterior] at hx hy refine' hx.2 (a⁻¹ * b) _ _ - rwa [Ne, inv_mul_eq_iff_eq_mul, mul_one, eq_comm] - simpa [mul_smul, ← hxy, mem_inv_smul_set_iff] using hy.1 + · rwa [Ne, inv_mul_eq_iff_eq_mul, mul_one, eq_comm] + · simpa [mul_smul, ← hxy, mem_inv_smul_set_iff] using hy.1 #align measure_theory.pairwise_disjoint_fundamental_interior MeasureTheory.pairwise_disjoint_fundamentalInterior #align measure_theory.pairwise_disjoint_add_fundamental_interior MeasureTheory.pairwise_disjoint_addFundamentalInterior diff --git a/Mathlib/MeasureTheory/Integral/Average.lean b/Mathlib/MeasureTheory/Integral/Average.lean index 7aa5eba86062e6..07f50023235120 100644 --- a/Mathlib/MeasureTheory/Integral/Average.lean +++ b/Mathlib/MeasureTheory/Integral/Average.lean @@ -695,7 +695,7 @@ theorem measure_setLaverage_le_pos (hμ : μ s ≠ 0) (hs : NullMeasurableSet s rintro x ⟨hfx, hx⟩ dsimp at hfx rw [← toReal_laverage hg.aemeasurable, toReal_le_toReal (setLaverage_lt_top hint).ne hx] at hfx - exact hfx.trans (hgf _) + · exact hfx.trans (hgf _) · simp_rw [ae_iff, not_ne_iff] exact measure_eq_top_of_lintegral_ne_top hg.aemeasurable hint #align measure_theory.measure_set_laverage_le_pos MeasureTheory.measure_setLaverage_le_pos diff --git a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean index f9dc65877e8bd9..e1a2b2b045e407 100644 --- a/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean +++ b/Mathlib/MeasureTheory/Integral/DominatedConvergence.lean @@ -245,7 +245,8 @@ theorem hasSum_intervalIntegral_of_summable_norm [Countable ι] {f : ι → C( · exact ae_of_all _ fun x _ => hf_sum · exact intervalIntegrable_const · refine ae_of_all _ fun x hx => Summable.hasSum ?_ - let x : (⟨uIcc a b, isCompact_uIcc⟩ : Compacts ℝ) := ⟨x, ?_⟩; swap; exact ⟨hx.1.le, hx.2⟩ + let x : (⟨uIcc a b, isCompact_uIcc⟩ : Compacts ℝ) := ⟨x, ?_⟩; swap + · exact ⟨hx.1.le, hx.2⟩ have := hf_sum.of_norm simpa only [Compacts.coe_mk, ContinuousMap.restrict_apply] using ContinuousMap.summable_apply this x diff --git a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean index 70ab94bcfb855f..14fe072171f99a 100644 --- a/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/IntervalIntegral.lean @@ -1043,7 +1043,7 @@ nonrec theorem integral_indicator {a₁ a₂ a₃ : ℝ} (h : a₂ ∈ Icc a₁ have : {x | x ≤ a₂} ∩ Ioc a₁ a₃ = Ioc a₁ a₂ := Iic_inter_Ioc_of_le h.2 rw [integral_of_le h.1, integral_of_le (h.1.trans h.2), integral_indicator, Measure.restrict_restrict, this] - exact measurableSet_Iic + · exact measurableSet_Iic all_goals apply measurableSet_Iic #align interval_integral.integral_indicator intervalIntegral.integral_indicator diff --git a/Mathlib/MeasureTheory/Integral/Lebesgue.lean b/Mathlib/MeasureTheory/Integral/Lebesgue.lean index 90157241aaed91..7829250014d8dc 100644 --- a/Mathlib/MeasureTheory/Integral/Lebesgue.lean +++ b/Mathlib/MeasureTheory/Integral/Lebesgue.lean @@ -805,8 +805,8 @@ theorem set_lintegral_eq_const {f : α → ℝ≥0∞} (hf : Measurable f) (r : ∫⁻ x in { x | f x = r }, f x ∂μ = r * μ { x | f x = r } := by have : ∀ᵐ x ∂μ, x ∈ { x | f x = r } → f x = r := ae_of_all μ fun _ hx => hx rw [set_lintegral_congr_fun _ this] - rw [lintegral_const, Measure.restrict_apply MeasurableSet.univ, Set.univ_inter] - exact hf (measurableSet_singleton r) + · rw [lintegral_const, Measure.restrict_apply MeasurableSet.univ, Set.univ_inter] + · exact hf (measurableSet_singleton r) #align measure_theory.set_lintegral_eq_const MeasureTheory.set_lintegral_eq_const theorem lintegral_indicator_one_le (s : Set α) : ∫⁻ a, s.indicator 1 a ∂μ ≤ μ s := @@ -1264,7 +1264,7 @@ theorem lintegral_iSup_directed [Countable β] {f : β → α → ℝ≥0∞} (h refine' ⟨z, _, _⟩ <;> · intro x by_cases hx : x ∈ aeSeqSet hf p - · repeat' rw [aeSeq.aeSeq_eq_fun_of_mem_aeSeqSet hf hx] + · repeat rw [aeSeq.aeSeq_eq_fun_of_mem_aeSeqSet hf hx] apply_rules [hz₁, hz₂] · simp only [aeSeq, hx, if_false] exact le_rfl diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index b1fd49e69ad5f3..e613a69c5f606e 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -1128,8 +1128,8 @@ theorem measurable_tProd_elim' [DecidableEq δ] {l : List δ} (h : ∀ i, i ∈ theorem MeasurableSet.tProd (l : List δ) {s : ∀ i, Set (π i)} (hs : ∀ i, MeasurableSet (s i)) : MeasurableSet (Set.tprod l s) := by induction' l with i l ih - exact MeasurableSet.univ - exact (hs i).prod ih + · exact MeasurableSet.univ + · exact (hs i).prod ih #align measurable_set.tprod MeasurableSet.tProd end TProd @@ -1835,8 +1835,8 @@ def sumPiEquivProdPi (α : δ ⊕ δ' → Type*) [∀ i, MeasurableSpace (α i)] · refine Measurable.prod ?_ ?_ <;> rw [measurable_pi_iff] <;> rintro i <;> apply measurable_pi_apply · rw [measurable_pi_iff]; rintro (i|i) - exact measurable_pi_iff.1 measurable_fst _ - exact measurable_pi_iff.1 measurable_snd _ + · exact measurable_pi_iff.1 measurable_fst _ + · exact measurable_pi_iff.1 measurable_snd _ theorem coe_sumPiEquivProdPi (α : δ ⊕ δ' → Type*) [∀ i, MeasurableSpace (α i)] : ⇑(MeasurableEquiv.sumPiEquivProdPi α) = Equiv.sumPiEquivProdPi α := by rfl diff --git a/Mathlib/MeasureTheory/Measure/Content.lean b/Mathlib/MeasureTheory/Measure/Content.lean index 0a6363d7d7373e..20c95e4f2515e4 100644 --- a/Mathlib/MeasureTheory/Measure/Content.lean +++ b/Mathlib/MeasureTheory/Measure/Content.lean @@ -333,8 +333,8 @@ theorem outerMeasure_caratheodory (A : Set G) : ∀ U : Opens G, μ.outerMeasure (U ∩ A) + μ.outerMeasure (U \ A) ≤ μ.outerMeasure U := by rw [Opens.forall] apply inducedOuterMeasure_caratheodory - apply innerContent_iUnion_nat - apply innerContent_mono' + · apply innerContent_iUnion_nat + · apply innerContent_mono' #align measure_theory.content.outer_measure_caratheodory MeasureTheory.Content.outerMeasure_caratheodory @[to_additive] diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index ee37be31132e8c..4a681d0aae2068 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -298,13 +298,13 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)).mpr <| measure_mono (subset_univ _) apply le_trans (setIntegral_mono (s := Ioc 0 ‖f‖) ?_ ?_ key) - rw [integral_add] - · apply add_le_add_left - simp only [integral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter, - Real.volume_Ioc, sub_zero, norm_nonneg, toReal_ofReal, smul_eq_mul, - (mul_comm _ ε).le] - · exact intble₂ - · exact integrable_const ε + · rw [integral_add] + · apply add_le_add_left + simp only [integral_const, MeasurableSet.univ, Measure.restrict_apply, univ_inter, + Real.volume_Ioc, sub_zero, norm_nonneg, toReal_ofReal, smul_eq_mul, + (mul_comm _ ε).le] + · exact intble₂ + · exact integrable_const ε · exact intble₁ · exact intble₂.add <| integrable_const ε @@ -330,9 +330,9 @@ lemma tendsto_integral_meas_thickening_le (f : Ω →ᵇ ℝ) · apply eventually_of_forall (fun t ↦ ?_) simp only [NNReal.tendsto_coe] apply (ENNReal.tendsto_toNNReal _).comp - apply tendsto_measure_thickening_of_isClosed ?_ ?_ - · exact ⟨1, ⟨Real.zero_lt_one, measure_ne_top _ _⟩⟩ - · exact isClosed_le continuous_const f.continuous + · apply tendsto_measure_thickening_of_isClosed ?_ ?_ + · exact ⟨1, ⟨Real.zero_lt_one, measure_ne_top _ _⟩⟩ + · exact isClosed_le continuous_const f.continuous · exact measure_ne_top _ _ /-- The coercion `LevyProkhorov (ProbabilityMeasure Ω) → ProbabilityMeasure Ω` is continuous. -/ diff --git a/Mathlib/MeasureTheory/Measure/Sub.lean b/Mathlib/MeasureTheory/Measure/Sub.lean index 6bf2729451376d..c333a51c454cf9 100644 --- a/Mathlib/MeasureTheory/Measure/Sub.lean +++ b/Mathlib/MeasureTheory/Measure/Sub.lean @@ -104,7 +104,7 @@ theorem sub_add_cancel_of_le [IsFiniteMeasure ν] (h₁ : ν ≤ μ) : μ - ν + theorem restrict_sub_eq_restrict_sub_restrict (h_meas_s : MeasurableSet s) : (μ - ν).restrict s = μ.restrict s - ν.restrict s := by - repeat' rw [sub_def] + repeat rw [sub_def] have h_nonempty : { d | μ ≤ d + ν }.Nonempty := ⟨μ, Measure.le_add_right le_rfl⟩ rw [restrict_sInf_eq_sInf_restrict h_nonempty h_meas_s] apply le_antisymm @@ -115,7 +115,7 @@ theorem restrict_sub_eq_restrict_sub_restrict (h_meas_s : MeasurableSet s) : refine' ⟨ν' + (⊤ : Measure α).restrict sᶜ, _, _⟩ · rw [mem_setOf_eq, add_right_comm, Measure.le_iff] intro t h_meas_t - repeat' rw [← measure_inter_add_diff t h_meas_s] + repeat rw [← measure_inter_add_diff t h_meas_s] refine' add_le_add _ _ · rw [add_apply, add_apply] apply le_add_right _ diff --git a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean index a99afa0c3cfcb9..36b59ba8e713bb 100644 --- a/Mathlib/MeasureTheory/Measure/VectorMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/VectorMeasure.lean @@ -920,15 +920,15 @@ theorem restrict_le_restrict_iUnion {f : ℕ → Set α} (hf₁ : ∀ n, Measura have ha₄ : Pairwise (Disjoint on fun n => a ∩ disjointed f n) := (disjoint_disjointed _).mono fun i j => Disjoint.mono inf_le_right inf_le_right rw [← ha₃, v.of_disjoint_iUnion_nat _ ha₄, w.of_disjoint_iUnion_nat _ ha₄] - refine' tsum_le_tsum (fun n => (restrict_le_restrict_iff v w (hf₁ n)).1 (hf₂ n) _ _) _ _ - · exact ha₁.inter (MeasurableSet.disjointed hf₁ n) - · exact Set.Subset.trans (Set.inter_subset_right _ _) (disjointed_subset _ _) - · refine' (v.m_iUnion (fun n => _) _).summable + · refine' tsum_le_tsum (fun n => (restrict_le_restrict_iff v w (hf₁ n)).1 (hf₂ n) _ _) _ _ · exact ha₁.inter (MeasurableSet.disjointed hf₁ n) - · exact (disjoint_disjointed _).mono fun i j => Disjoint.mono inf_le_right inf_le_right - · refine' (w.m_iUnion (fun n => _) _).summable - · exact ha₁.inter (MeasurableSet.disjointed hf₁ n) - · exact (disjoint_disjointed _).mono fun i j => Disjoint.mono inf_le_right inf_le_right + · exact Set.Subset.trans (Set.inter_subset_right _ _) (disjointed_subset _ _) + · refine' (v.m_iUnion (fun n => _) _).summable + · exact ha₁.inter (MeasurableSet.disjointed hf₁ n) + · exact (disjoint_disjointed _).mono fun i j => Disjoint.mono inf_le_right inf_le_right + · refine' (w.m_iUnion (fun n => _) _).summable + · exact ha₁.inter (MeasurableSet.disjointed hf₁ n) + · exact (disjoint_disjointed _).mono fun i j => Disjoint.mono inf_le_right inf_le_right · intro n exact ha₁.inter (MeasurableSet.disjointed hf₁ n) · exact fun n => ha₁.inter (MeasurableSet.disjointed hf₁ n) diff --git a/Mathlib/NumberTheory/ADEInequality.lean b/Mathlib/NumberTheory/ADEInequality.lean index 738f65c09ccb4e..a37cfcd8bba92e 100644 --- a/Mathlib/NumberTheory/ADEInequality.lean +++ b/Mathlib/NumberTheory/ADEInequality.lean @@ -179,16 +179,16 @@ theorem lt_three {p q r : ℕ+} (hpq : p ≤ q) (hqr : q ≤ r) (H : 1 < sumInv have h3r := h3q.trans hqr have hp: (p : ℚ)⁻¹ ≤ 3⁻¹ := by rw [inv_le_inv _ h3] - assumption_mod_cast - norm_num + · assumption_mod_cast + · norm_num have hq: (q : ℚ)⁻¹ ≤ 3⁻¹ := by rw [inv_le_inv _ h3] - assumption_mod_cast - norm_num + · assumption_mod_cast + · norm_num have hr: (r : ℚ)⁻¹ ≤ 3⁻¹ := by rw [inv_le_inv _ h3] - assumption_mod_cast - norm_num + · assumption_mod_cast + · norm_num calc (p : ℚ)⁻¹ + (q : ℚ)⁻¹ + (r : ℚ)⁻¹ ≤ 3⁻¹ + 3⁻¹ + 3⁻¹ := add_le_add (add_le_add hp hq) hr _ = 1 := by norm_num @@ -201,12 +201,12 @@ theorem lt_four {q r : ℕ+} (hqr : q ≤ r) (H : 1 < sumInv {2, q, r}) : q < 4 have h4r := H.trans hqr have hq: (q : ℚ)⁻¹ ≤ 4⁻¹ := by rw [inv_le_inv _ h4] - assumption_mod_cast - norm_num + · assumption_mod_cast + · norm_num have hr: (r : ℚ)⁻¹ ≤ 4⁻¹ := by rw [inv_le_inv _ h4] - assumption_mod_cast - norm_num + · assumption_mod_cast + · norm_num calc (2⁻¹ + (q : ℚ)⁻¹ + (r : ℚ)⁻¹) ≤ 2⁻¹ + 4⁻¹ + 4⁻¹ := add_le_add (add_le_add le_rfl hq) hr _ = 1 := by norm_num @@ -218,8 +218,8 @@ theorem lt_six {r : ℕ+} (H : 1 < sumInv {2, 3, r}) : r < 6 := by rw [sumInv_pqr] have hr: (r : ℚ)⁻¹ ≤ 6⁻¹ := by rw [inv_le_inv _ h6] - assumption_mod_cast - norm_num + · assumption_mod_cast + · norm_num calc (2⁻¹ + 3⁻¹ + (r : ℚ)⁻¹ : ℚ) ≤ 2⁻¹ + 3⁻¹ + 6⁻¹ := add_le_add (add_le_add le_rfl le_rfl) hr _ = 1 := by norm_num diff --git a/Mathlib/NumberTheory/Cyclotomic/Rat.lean b/Mathlib/NumberTheory/Cyclotomic/Rat.lean index 718407b0cfbb18..8f337bb75ba856 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Rat.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Rat.lean @@ -250,9 +250,9 @@ noncomputable def subOneIntegralPowerBasis [IsCyclotomicExtension {p ^ k} ℚ K] convert Subalgebra.add_mem _ (self_mem_adjoin_singleton ℤ (⟨ζ - 1, _⟩ : 𝓞 K)) (Subalgebra.one_mem _) -- Porting note: `simp` was able to finish the proof. - simp only [Subsemiring.coe_add, Subalgebra.coe_toSubsemiring, - OneMemClass.coe_one, sub_add_cancel] - exact Subalgebra.sub_mem _ (hζ.isIntegral (by simp)) (Subalgebra.one_mem _)) + · simp only [Subsemiring.coe_add, Subalgebra.coe_toSubsemiring, + OneMemClass.coe_one, sub_add_cancel] + · exact Subalgebra.sub_mem _ (hζ.isIntegral (by simp)) (Subalgebra.one_mem _)) #align is_primitive_root.sub_one_integral_power_basis IsPrimitiveRoot.subOneIntegralPowerBasis @[simp] diff --git a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean index 91f55a9d303dae..efed081687958c 100644 --- a/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean +++ b/Mathlib/NumberTheory/Harmonic/ZetaAsymp.lean @@ -217,7 +217,7 @@ lemma term_tsum_of_lt {s : ℝ} (hs : 1 < s) : linarith · apply le_of_eq rw [rpow_sub_one, ← div_mul, div_one, mul_comm, one_div, inv_rpow, ← div_eq_mul_inv] - norm_cast + · norm_cast all_goals positivity /-- Reformulation of `ZetaAsymptotics.term_tsum_of_lt` which is useful for some computations @@ -271,9 +271,9 @@ lemma continuousOn_term_tsum : ContinuousOn term_tsum (Ici 1) := by · exact (term_welldef n.succ_pos (zero_lt_one.trans_le hs)).1 · exact (term_welldef n.succ_pos zero_lt_one).1 · rw [div_le_div_left] -- leave side-goals to end and kill them all together - apply rpow_le_rpow_of_exponent_le - · exact (lt_of_le_of_lt (by simp) hx.1).le - · linarith [mem_Ici.mp hs] + · apply rpow_le_rpow_of_exponent_le + · exact (lt_of_le_of_lt (by simp) hx.1).le + · linarith [mem_Ici.mp hs] · linarith [hx.1] all_goals apply rpow_pos_of_pos ((Nat.cast_nonneg _).trans_lt hx.1) · rw [intervalIntegral.integral_of_le (by linarith)] diff --git a/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean b/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean index 28d79700dbad63..4dbb0d550803b0 100644 --- a/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean +++ b/Mathlib/NumberTheory/LSeries/MellinEqDirichlet.lean @@ -119,7 +119,7 @@ lemma hasSum_mellin_pi_mul_sq {a : ι → ℂ} {r : ι → ℝ} {F : ℝ → ℂ convert hasSum_mellin_pi_mul₀ (fun i ↦ sq_nonneg (r i)) hs' hF ?_ using 3 with i · rw [← neg_div, Gammaℝ_def] · rw [← _root_.sq_abs, ofReal_pow, ← cpow_nat_mul'] - ring_nf + · ring_nf all_goals rw [arg_ofReal_of_nonneg (abs_nonneg _)]; linarith [pi_pos] · convert h_sum using 3 with i rw [← _root_.sq_abs, ← rpow_natCast_mul (abs_nonneg _), div_ofNat_re, Nat.cast_ofNat, diff --git a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean index 138a2ea492d965..d621c4c130f500 100644 --- a/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean +++ b/Mathlib/NumberTheory/LegendreSymbol/JacobiSymbol.lean @@ -441,7 +441,7 @@ theorem quadratic_reciprocity' {a b : ℕ} (ha : Odd a) (hb : Odd b) : -- define the right hand side for fixed `a` as a `ℕ →* ℤ` let rhs : ℕ → ℕ →* ℤ := fun a => { toFun := fun x => qrSign x a * J(x | a) - map_one' := by convert ← mul_one (M := ℤ) _; symm; all_goals apply one_left + map_one' := by convert ← mul_one (M := ℤ) _; (on_goal 1 => symm); all_goals apply one_left map_mul' := fun x y => by -- Porting note: `simp_rw` on line 423 replaces `rw` to allow the rewrite rules to be -- applied under the binder `fun ↦ ...` @@ -506,12 +506,13 @@ theorem mod_right' (a : ℕ) {b : ℕ} (hb : Odd b) : J(a | b) = J(a | b % (4 * rw [Nat.cast_mul, mul_left, mul_left, quadratic_reciprocity' ha₁ hb, quadratic_reciprocity' ha₁ hb', Nat.cast_pow, pow_left, pow_left, Nat.cast_two, at_two hb, at_two hb'] - congr 1; swap; congr 1 - · simp_rw [qrSign] - rw [χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * a)), mod_mod_of_dvd b (dvd_mul_right 4 a)] - · rw [mod_left ↑(b % _), mod_left b, Int.natCast_mod, Int.emod_emod_of_dvd b] - simp only [ha₂, Nat.cast_mul, ← mul_assoc] - apply dvd_mul_left + congr 1; swap; + · congr 1 + · simp_rw [qrSign] + rw [χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * a)), mod_mod_of_dvd b (dvd_mul_right 4 a)] + · rw [mod_left ↑(b % _), mod_left b, Int.natCast_mod, Int.emod_emod_of_dvd b] + simp only [ha₂, Nat.cast_mul, ← mul_assoc] + apply dvd_mul_left -- Porting note: In mathlib3, it was written `cases' e`. In Lean 4, this resulted in the choice -- of a name other than e (for the case distinction of line 482) so we indicate the name -- to use explicitly. diff --git a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean index 112c1013b80781..cdd96609f933d4 100644 --- a/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean +++ b/Mathlib/NumberTheory/ModularForms/JacobiTheta/Bounds.lean @@ -250,8 +250,8 @@ lemma isBigO_atTop_F_int_zero_sub (a : UnitAddCircle) : ∃ p, 0 < p ∧ obtain ⟨q, hq, hq'⟩ := isBigO_atTop_F_nat_zero_sub (sub_nonneg.mpr ha.2.le) have ha' : (a : UnitAddCircle) = 0 ↔ a = 0 := by rw [← AddCircle.coe_eq_coe_iff_of_mem_Ico (hp := ⟨zero_lt_one' ℝ⟩), QuotientAddGroup.mk_zero] - rw [zero_add]; exact ha - simp + · rw [zero_add]; exact ha + · simp simp_rw [ha'] simp_rw [eq_false_intro (by linarith [ha.2] : 1 - a ≠ 0), if_false, sub_zero] at hq' refine ⟨_, lt_min hp hq, ?_⟩ diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index c8857ee4035302..9ad583ee170061 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -78,8 +78,8 @@ theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : obtain hr | hr := lt_or_le r 0 · obtain ⟨φ⟩ := (inferInstance : Nonempty (K →+* ℂ)) refine iff_of_false ?_ ?_ - exact (hr.trans_le (norm_nonneg _)).not_le - exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) + · exact (hr.trans_le (norm_nonneg _)).not_le + · exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) · lift r to NNReal using hr simp_rw [← coe_nnnorm, nnnorm_eq, NNReal.coe_le_coe, Finset.sup_le_iff, Finset.mem_univ, forall_true_left] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 54016369c69b8b..547d55908859bf 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -88,7 +88,7 @@ instance : IsAddHaarMeasure (volume : Measure (E K)) := prod.instIsAddHaarMeasur instance : NoAtoms (volume : Measure (E K)) := by obtain ⟨w⟩ := (inferInstance : Nonempty (InfinitePlace K)) by_cases hw : IsReal w - exact @prod.instNoAtoms_fst _ _ _ _ volume volume _ (pi_noAtoms ⟨w, hw⟩) + · exact @prod.instNoAtoms_fst _ _ _ _ volume volume _ (pi_noAtoms ⟨w, hw⟩) · exact @prod.instNoAtoms_snd _ _ _ _ volume volume _ (pi_noAtoms ⟨w, not_isReal_iff_isComplex.mp hw⟩) @@ -223,16 +223,16 @@ theorem convexBodyLT'_volume : have vol_box : ∀ B : ℝ≥0, volume {x : ℂ | |x.re| < 1 ∧ |x.im| < B^2} = 4*B^2 := by intro B rw [← (Complex.volume_preserving_equiv_real_prod.symm).measure_preimage] - simp_rw [Set.preimage_setOf_eq, Complex.measurableEquivRealProd_symm_apply] - rw [show {a : ℝ × ℝ | |a.1| < 1 ∧ |a.2| < B ^ 2} = - Set.Ioo (-1:ℝ) (1:ℝ) ×ˢ Set.Ioo (- (B:ℝ) ^ 2) ((B:ℝ) ^ 2) by - ext; simp_rw [Set.mem_setOf_eq, Set.mem_prod, Set.mem_Ioo, abs_lt]] - simp_rw [volume_eq_prod, prod_prod, Real.volume_Ioo, sub_neg_eq_add, one_add_one_eq_two, - ← two_mul, ofReal_mul zero_le_two, ofReal_pow (coe_nonneg B), ofReal_ofNat, - ofReal_coe_nnreal, ← mul_assoc, show (2:ℝ≥0∞) * 2 = 4 by norm_num] - refine MeasurableSet.inter ?_ ?_ - · exact measurableSet_lt (measurable_norm.comp Complex.measurable_re) measurable_const - · exact measurableSet_lt (measurable_norm.comp Complex.measurable_im) measurable_const + · simp_rw [Set.preimage_setOf_eq, Complex.measurableEquivRealProd_symm_apply] + rw [show {a : ℝ × ℝ | |a.1| < 1 ∧ |a.2| < B ^ 2} = + Set.Ioo (-1:ℝ) (1:ℝ) ×ˢ Set.Ioo (- (B:ℝ) ^ 2) ((B:ℝ) ^ 2) by + ext; simp_rw [Set.mem_setOf_eq, Set.mem_prod, Set.mem_Ioo, abs_lt]] + simp_rw [volume_eq_prod, prod_prod, Real.volume_Ioo, sub_neg_eq_add, one_add_one_eq_two, + ← two_mul, ofReal_mul zero_le_two, ofReal_pow (coe_nonneg B), ofReal_ofNat, + ofReal_coe_nnreal, ← mul_assoc, show (2:ℝ≥0∞) * 2 = 4 by norm_num] + · refine MeasurableSet.inter ?_ ?_ + · exact measurableSet_lt (measurable_norm.comp Complex.measurable_re) measurable_const + · exact measurableSet_lt (measurable_norm.comp Complex.measurable_im) measurable_const calc _ = (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (2 * (f x.val))) * ((∏ x in Finset.univ.erase w₀, ENNReal.ofReal (f x.val) ^ 2 * pi) * @@ -384,8 +384,8 @@ theorem convexBodySum_convex : Convex ℝ (convexBodySum K B) := by theorem convexBodySum_isBounded : Bornology.IsBounded (convexBodySum K B) := by refine Metric.isBounded_iff.mpr ⟨B + B, fun x hx y hy => ?_⟩ refine le_trans (norm_sub_le x y) (add_le_add ?_ ?_) - exact le_trans (norm_le_convexBodySumFun x) hx - exact le_trans (norm_le_convexBodySumFun y) hy + · exact le_trans (norm_le_convexBodySumFun x) hx + · exact le_trans (norm_le_convexBodySumFun y) hy theorem convexBodySum_compact : IsCompact (convexBodySum K B) := by rw [Metric.isCompact_iff_isClosed_bounded] @@ -486,10 +486,10 @@ theorem volume_fundamentalDomain_fractionalIdealLatticeBasis : fractionalIdeal_rank] rw [← fundamentalDomain_reindex (fractionalIdealLatticeBasis K I) e, measure_fundamentalDomain ((fractionalIdealLatticeBasis K I).reindex e)] - rw [show (fractionalIdealLatticeBasis K I).reindex e = (mixedEmbedding K) ∘ - (basisOfFractionalIdeal K I) ∘ e.symm by - ext1; simp only [Basis.coe_reindex, Function.comp_apply, fractionalIdealLatticeBasis_apply]] - rw [mixedEmbedding.det_basisOfFractionalIdeal_eq_norm] + · rw [show (fractionalIdealLatticeBasis K I).reindex e = (mixedEmbedding K) ∘ + (basisOfFractionalIdeal K I) ∘ e.symm by + ext1; simp only [Basis.coe_reindex, Function.comp_apply, fractionalIdealLatticeBasis_apply]] + rw [mixedEmbedding.det_basisOfFractionalIdeal_eq_norm] theorem minkowskiBound_lt_top : minkowskiBound K I < ⊤ := by refine ENNReal.mul_lt_top ?_ ?_ diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 7f791fcfbc43da..f93cac74b30936 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -442,12 +442,12 @@ theorem unitLattice_span_eq_top : simp_rw [Real.norm_eq_abs, B, Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply] rw [← sub_pos, sum_congr rfl (fun x hx => abs_of_neg ?_), sum_neg_distrib, sub_neg_eq_add, sum_erase_eq_sub (mem_univ _), ← add_comm_sub] - refine add_pos_of_nonneg_of_pos ?_ ?_ - · rw [sub_nonneg] - exact le_abs_self _ - · rw [sum_logEmbedding_component (exists_unit K w).choose] - refine mul_pos_of_neg_of_neg ?_ ((exists_unit K w).choose_spec _ w.prop.symm) - rw [mult]; split_ifs <;> norm_num + · refine add_pos_of_nonneg_of_pos ?_ ?_ + · rw [sub_nonneg] + exact le_abs_self _ + · rw [sum_logEmbedding_component (exists_unit K w).choose] + refine mul_pos_of_neg_of_neg ?_ ((exists_unit K w).choose_spec _ w.prop.symm) + rw [mult]; split_ifs <;> norm_num · refine mul_neg_of_pos_of_neg ?_ ((exists_unit K w).choose_spec x ?_) · rw [mult]; split_ifs <;> norm_num · exact Subtype.ext_iff_val.not.mp (ne_of_mem_erase hx) diff --git a/Mathlib/Order/CompactlyGenerated/Basic.lean b/Mathlib/Order/CompactlyGenerated/Basic.lean index d733e05ab21d58..d82b297d917146 100644 --- a/Mathlib/Order/CompactlyGenerated/Basic.lean +++ b/Mathlib/Order/CompactlyGenerated/Basic.lean @@ -593,13 +593,13 @@ theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } ⟨CompleteLattice.independent_sUnion_of_directed hc2.directedOn fun s hs => (hc1 hs).1, ?_, fun a ⟨s, sc, as⟩ => (hc1 sc).2.2 a as⟩, fun _ => Set.subset_sUnion_of_mem⟩ - obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := this swap · rw [sSup_sUnion, ← sSup_image, DirectedOn.disjoint_sSup_right] · rintro _ ⟨s, hs, rfl⟩ exact (hc1 hs).2.1 · rw [directedOn_image] exact hc2.directedOn.mono @fun s t => sSup_le_sSup + obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := this refine' ⟨s, s_ind, ⟨b_inf_Sup_s, _⟩, s_atoms⟩ rw [codisjoint_iff_le_sup, ← h, sSup_le_iff] intro a ha diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index 44ed77dc0cbb24..b3406ad525ffcb 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -1670,7 +1670,7 @@ noncomputable instance WithTop.WithBot.completeLattice {α : Type*} show ite _ _ _ ≤ a by split_ifs with h₁ · cases' a with a - exact le_rfl + · exact le_rfl cases h₁ haS · cases a · exact le_top diff --git a/Mathlib/Order/Extension/Linear.lean b/Mathlib/Order/Extension/Linear.lean index 310e29a5881b87..0550b9623c2ef5 100644 --- a/Mathlib/Order/Extension/Linear.lean +++ b/Mathlib/Order/Extension/Linear.lean @@ -61,11 +61,11 @@ theorem extend_partialOrder {α : Type u} (r : α → α → Prop) [IsPartialOrd { refl := fun x ↦ Or.inl (refl _) trans := _ antisymm := _ } - rintro a b c (ab | ⟨ax : s a x, yb : s y b⟩) (bc | ⟨bx : s b x, yc : s y c⟩) - · exact Or.inl (_root_.trans ab bc) - · exact Or.inr ⟨_root_.trans ab bx, yc⟩ - · exact Or.inr ⟨ax, _root_.trans yb bc⟩ - · exact Or.inr ⟨ax, yc⟩ + · rintro a b c (ab | ⟨ax : s a x, yb : s y b⟩) (bc | ⟨bx : s b x, yc : s y c⟩) + · exact Or.inl (_root_.trans ab bc) + · exact Or.inr ⟨_root_.trans ab bx, yc⟩ + · exact Or.inr ⟨ax, _root_.trans yb bc⟩ + · exact Or.inr ⟨ax, yc⟩ rintro a b (ab | ⟨ax : s a x, yb : s y b⟩) (ba | ⟨bx : s b x, ya : s y a⟩) · exact antisymm ab ba · exact (h.2 (_root_.trans ya (_root_.trans ab bx))).elim diff --git a/Mathlib/Order/Filter/AtTopBot.lean b/Mathlib/Order/Filter/AtTopBot.lean index 2ba1a9c43b324a..713496e5698676 100644 --- a/Mathlib/Order/Filter/AtTopBot.lean +++ b/Mathlib/Order/Filter/AtTopBot.lean @@ -1492,8 +1492,10 @@ theorem tendsto_finset_preimage_atTop_atTop {f : α → β} (hf : Function.Injec -- Porting note: generalized from `SemilatticeSup` to `Preorder` theorem prod_atTop_atTop_eq [Preorder α] [Preorder β] : (atTop : Filter α) ×ˢ (atTop : Filter β) = (atTop : Filter (α × β)) := by - cases isEmpty_or_nonempty α; exact Subsingleton.elim _ _ - cases isEmpty_or_nonempty β; exact Subsingleton.elim _ _ + cases isEmpty_or_nonempty α + · exact Subsingleton.elim _ _ + cases isEmpty_or_nonempty β + · exact Subsingleton.elim _ _ simpa [atTop, prod_iInf_left, prod_iInf_right, iInf_prod] using iInf_comm #align filter.prod_at_top_at_top_eq Filter.prod_atTop_atTop_eq diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index e1153828db6a8a..3f5210874f88c2 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -786,8 +786,9 @@ theorem NeBot.nonempty (f : Filter α) [hf : f.NeBot] : Nonempty α := equal. -/ theorem eq_top_of_neBot [Subsingleton α] (l : Filter α) [NeBot l] : l = ⊤ := by refine' top_unique fun s hs => _ - obtain rfl : s = univ; exact Subsingleton.eq_univ_of_nonempty (nonempty_of_mem hs) - exact univ_mem + obtain rfl : s = univ + · exact Subsingleton.eq_univ_of_nonempty (nonempty_of_mem hs) + · exact univ_mem #align filter.eq_top_of_ne_bot Filter.eq_top_of_neBot theorem forall_mem_nonempty_iff_neBot {f : Filter α} : diff --git a/Mathlib/Order/Filter/Prod.lean b/Mathlib/Order/Filter/Prod.lean index f5a0942bc403d9..8f93bfbb2c1c6b 100644 --- a/Mathlib/Order/Filter/Prod.lean +++ b/Mathlib/Order/Filter/Prod.lean @@ -392,8 +392,8 @@ protected theorem map_prod (m : α × β → γ) (f : Filter α) (g : Filter β) simp only [Filter.ext_iff, mem_map, mem_prod_iff, mem_map_seq_iff, exists_and_left] intro s constructor - exact fun ⟨t, ht, s, hs, h⟩ => ⟨s, hs, t, ht, fun x hx y hy => @h ⟨x, y⟩ ⟨hx, hy⟩⟩ - exact fun ⟨s, hs, t, ht, h⟩ => ⟨t, ht, s, hs, fun ⟨x, y⟩ ⟨hx, hy⟩ => h x hx y hy⟩ + · exact fun ⟨t, ht, s, hs, h⟩ => ⟨s, hs, t, ht, fun x hx y hy => @h ⟨x, y⟩ ⟨hx, hy⟩⟩ + · exact fun ⟨s, hs, t, ht, h⟩ => ⟨t, ht, s, hs, fun ⟨x, y⟩ ⟨hx, hy⟩ => h x hx y hy⟩ #align filter.map_prod Filter.map_prod theorem prod_eq : f ×ˢ g = (f.map Prod.mk).seq g := f.map_prod id g diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 3e6587e1e4bb79..7a9ac1a395ff87 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -557,7 +557,7 @@ instance lattice : Lattice (Interval α) := change _ ≤ dite _ _ _ simp only [WithBot.some_eq_coe, WithBot.coe_le_coe] at hb hc ⊢ rw [dif_pos, WithBot.coe_le_coe] - exact ⟨sup_le hb.1 hc.1, le_inf hb.2 hc.2⟩ + · exact ⟨sup_le hb.1 hc.1, le_inf hb.2 hc.2⟩ -- Porting note: had to add the next 6 lines including the changes because -- it seems that lean cannot automatically turn `NonemptyInterval.toDualProd s` -- into `s.toProd` anymore. diff --git a/Mathlib/Order/Iterate.lean b/Mathlib/Order/Iterate.lean index a59942b391de34..6946df96ed01fe 100644 --- a/Mathlib/Order/Iterate.lean +++ b/Mathlib/Order/Iterate.lean @@ -44,8 +44,8 @@ theorem seq_le_seq (hf : Monotone f) (n : ℕ) (h₀ : x 0 ≤ y 0) (hx : ∀ k induction' n with n ihn · exact h₀ · refine' (hx _ n.lt_succ_self).trans ((hf <| ihn _ _).trans (hy _ n.lt_succ_self)) - exact fun k hk => hx _ (hk.trans n.lt_succ_self) - exact fun k hk => hy _ (hk.trans n.lt_succ_self) + · exact fun k hk => hx _ (hk.trans n.lt_succ_self) + · exact fun k hk => hy _ (hk.trans n.lt_succ_self) #align monotone.seq_le_seq Monotone.seq_le_seq theorem seq_pos_lt_seq_of_lt_of_le (hf : Monotone f) {n : ℕ} (hn : 0 < n) (h₀ : x 0 ≤ y 0) diff --git a/Mathlib/Order/Minimal.lean b/Mathlib/Order/Minimal.lean index 9b55e6eedbf3a0..6dc8388906dcfe 100644 --- a/Mathlib/Order/Minimal.lean +++ b/Mathlib/Order/Minimal.lean @@ -423,7 +423,9 @@ theorem RelEmbedding.inter_preimage_minimals_eq_of_subset (f : r ↪r s) (h : y theorem RelEmbedding.minimals_preimage_eq (f : r ↪r s) (y : Set β) : minimals r (f ⁻¹' y) = f ⁻¹' minimals s (y ∩ range f) := by - convert (f.inter_preimage_minimals_eq univ y).symm; simp [univ_inter]; simp [inter_comm] + convert (f.inter_preimage_minimals_eq univ y).symm + · simp [univ_inter] + · simp [inter_comm] theorem inter_maximals_preimage_inter_eq_of_rel_iff_rel_on (hf : ∀ ⦃a a'⦄, a ∈ x → a' ∈ x → (r a a' ↔ s (f a) (f a'))) (y : Set β) : @@ -447,7 +449,9 @@ theorem RelEmbedding.inter_preimage_maximals_eq_of_subset (f : r ↪r s) (h : y theorem RelEmbedding.maximals_preimage_eq (f : r ↪r s) (y : Set β) : maximals r (f ⁻¹' y) = f ⁻¹' maximals s (y ∩ range f) := by - convert (f.inter_preimage_maximals_eq univ y).symm; simp [univ_inter]; simp [inter_comm] + convert (f.inter_preimage_maximals_eq univ y).symm + · simp [univ_inter] + · simp [inter_comm] end Image diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index f5bd1dc1eb1263..9243444ca85ae0 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -232,8 +232,8 @@ theorem ωSup_le_ωSup_of_le {c₀ c₁ : Chain α} (h : c₀ ≤ c₁) : ωSup theorem ωSup_le_iff (c : Chain α) (x : α) : ωSup c ≤ x ↔ ∀ i, c i ≤ x := by constructor <;> intros · trans ωSup c - exact le_ωSup _ _ - assumption + · exact le_ωSup _ _ + · assumption exact ωSup_le _ _ ‹_› #align omega_complete_partial_order.ωSup_le_iff OmegaCompletePartialOrder.ωSup_le_iff @@ -295,9 +295,9 @@ lemma isLUB_of_scottContinuous {c : Chain α} {f : α → β} (hf : ScottContinu lemma ScottContinuous.continuous' {f : α → β} (hf : ScottContinuous f) : Continuous' f := by constructor - intro c - rw [← (ωSup_eq_of_isLUB (isLUB_of_scottContinuous hf))] - simp only [OrderHom.coe_mk] + · intro c + rw [← (ωSup_eq_of_isLUB (isLUB_of_scottContinuous hf))] + simp only [OrderHom.coe_mk] theorem Continuous'.to_monotone {f : α → β} (hf : Continuous' f) : Monotone f := hf.fst @@ -417,7 +417,7 @@ theorem mem_ωSup (x : α) (c : Chain (Part α)) : x ∈ ωSup c ↔ some x ∈ constructor · split_ifs with h swap - rintro ⟨⟨⟩⟩ + · rintro ⟨⟨⟩⟩ intro h' have hh := Classical.choose_spec h simp only [mem_some_iff] at h' @@ -857,8 +857,8 @@ def apply : (α →𝒄 β) × α →𝒄 β where intro j apply le_ωSup_of_le (max i j) apply apply_mono - exact monotone_fst (OrderHom.mono _ (le_max_left _ _)) - exact monotone_snd (OrderHom.mono _ (le_max_right _ _)) + · exact monotone_fst (OrderHom.mono _ (le_max_left _ _)) + · exact monotone_snd (OrderHom.mono _ (le_max_right _ _)) · apply ωSup_le intro i apply le_ωSup_of_le i diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index 894debf0532985..4664cb41be95c6 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -277,15 +277,18 @@ def insertNth (p : RelSeries r) (i : Fin p.length) (a : α) · convert p.step ⟨m, hm.trans i.2⟩ · show Fin.insertNth _ _ _ _ = _ rw [Fin.insertNth_apply_below] - pick_goal 2; exact hm.trans (lt_add_one _) + pick_goal 2 + · exact hm.trans (lt_add_one _) simp · show Fin.insertNth _ _ _ _ = _ rw [Fin.insertNth_apply_below] - pick_goal 2; change m.1 + 1 < i.1 + 1; rwa [add_lt_add_iff_right] + pick_goal 2 + · change m.1 + 1 < i.1 + 1; rwa [add_lt_add_iff_right] simp; rfl · rw [show x = p m from show Fin.insertNth _ _ _ _ = _ by rw [Fin.insertNth_apply_below] - pick_goal 2; show m.1 < i.1 + 1; exact hm ▸ lt_add_one _ + pick_goal 2 + · show m.1 < i.1 + 1; exact hm ▸ lt_add_one _ simp] convert prev_connect · ext; exact hm @@ -300,7 +303,8 @@ def insertNth (p : RelSeries r) (i : Fin p.length) (a : α) aesop · change Fin.insertNth _ _ _ _ = _ rw [Fin.insertNth_apply_above] - swap; exact hm.trans (lt_add_one _) + swap + · exact hm.trans (lt_add_one _) simp only [Fin.val_succ, Nat.zero_eq, Fin.pred_succ, eq_rec_constant, ge_iff_le, Fin.succ_mk] congr @@ -310,7 +314,8 @@ def insertNth (p : RelSeries r) (i : Fin p.length) (a : α) rw [show m.castSucc = i.succ.castSucc from Fin.ext hm.symm, Fin.insertNth_apply_same] · change Fin.insertNth _ _ _ _ = _ rw [Fin.insertNth_apply_above] - swap; change i.1 + 1 < m.1 + 1; rw [hm]; exact lt_add_one _ + swap + · change i.1 + 1 < m.1 + 1; rw [hm]; exact lt_add_one _ simp only [Fin.pred_succ, eq_rec_constant] congr; ext; exact hm.symm diff --git a/Mathlib/Order/WellFounded.lean b/Mathlib/Order/WellFounded.lean index 7520b187f45642..f2b4b3329b5b98 100644 --- a/Mathlib/Order/WellFounded.lean +++ b/Mathlib/Order/WellFounded.lean @@ -129,13 +129,13 @@ protected theorem lt_succ_iff {r : α → α → Prop} [wo : IsWellOrder α r] { rw [WellFounded.succ, dif_pos] at h' exact wo.wf.not_lt_min _ h hy h' rcases trichotomous_of r x y with (hy | hy | hy) - exfalso - exact this hy - right - exact hy.symm + · exfalso + exact this hy + · right + exact hy.symm left exact hy - rintro (hy | rfl); exact _root_.trans hy (wo.wf.lt_succ h); exact wo.wf.lt_succ h + rintro (hy | rfl); (· exact _root_.trans hy (wo.wf.lt_succ h)); exact wo.wf.lt_succ h #align well_founded.lt_succ_iff WellFounded.lt_succ_iff section LinearOrder diff --git a/Mathlib/Order/WellFoundedSet.lean b/Mathlib/Order/WellFoundedSet.lean index 999bf02f246a2c..a81f42033c4a6e 100644 --- a/Mathlib/Order/WellFoundedSet.lean +++ b/Mathlib/Order/WellFoundedSet.lean @@ -303,7 +303,8 @@ theorem PartiallyWellOrderedOn.image_of_monotone_on (hs : s.PartiallyWellOrdered (hf : ∀ a₁ ∈ s, ∀ a₂ ∈ s, r a₁ a₂ → r' (f a₁) (f a₂)) : (f '' s).PartiallyWellOrderedOn r' := by intro g' hg' choose g hgs heq using hg' - obtain rfl : f ∘ g = g'; exact funext heq + obtain rfl : f ∘ g = g' + · exact funext heq obtain ⟨m, n, hlt, hmn⟩ := hs g hgs exact ⟨m, n, hlt, hf _ (hgs m) _ (hgs n) hmn⟩ #align set.partially_well_ordered_on.image_of_monotone_on Set.PartiallyWellOrderedOn.image_of_monotone_on @@ -870,7 +871,7 @@ theorem WellFounded.prod_lex_of_wellFoundedOn_fiber (hα : WellFounded (rα on f (hβ : ∀ a, (f ⁻¹' {a}).WellFoundedOn (rβ on g)) : WellFounded (Prod.Lex rα rβ on fun c => (f c, g c)) := by refine' (PSigma.lex_wf (wellFoundedOn_range.2 hα) fun a => hβ a).onFun.mono fun c c' h => _ - exact fun c => ⟨⟨_, c, rfl⟩, c, rfl⟩ + · exact fun c => ⟨⟨_, c, rfl⟩, c, rfl⟩ obtain h' | h' := Prod.lex_iff.1 h · exact PSigma.Lex.left _ _ h' · dsimp only [InvImage, (· on ·)] at h' ⊢ @@ -883,8 +884,8 @@ theorem Set.WellFoundedOn.prod_lex_of_wellFoundedOn_fiber (hα : s.WellFoundedOn s.WellFoundedOn (Prod.Lex rα rβ on fun c => (f c, g c)) := by refine' WellFounded.prod_lex_of_wellFoundedOn_fiber hα fun a ↦ (hβ a).onFun.mono (fun b c h ↦ _) swap - exact fun _ x => ⟨x, x.1.2, x.2⟩ - assumption + · exact fun _ x => ⟨x, x.1.2, x.2⟩ + · assumption #align set.well_founded_on.prod_lex_of_well_founded_on_fiber Set.WellFoundedOn.prod_lex_of_wellFoundedOn_fiber end ProdLex @@ -899,7 +900,7 @@ theorem WellFounded.sigma_lex_of_wellFoundedOn_fiber (hι : WellFounded (rι on (hπ : ∀ i, (f ⁻¹' {i}).WellFoundedOn (rπ i on g i)) : WellFounded (Sigma.Lex rι rπ on fun c => ⟨f c, g (f c) c⟩) := by refine' (PSigma.lex_wf (wellFoundedOn_range.2 hι) fun a => hπ a).onFun.mono fun c c' h => _ - exact fun c => ⟨⟨_, c, rfl⟩, c, rfl⟩ + · exact fun c => ⟨⟨_, c, rfl⟩, c, rfl⟩ obtain h' | ⟨h', h''⟩ := Sigma.lex_iff.1 h · exact PSigma.Lex.left _ _ h' · dsimp only [InvImage, (· on ·)] at h' ⊢ @@ -921,8 +922,8 @@ theorem Set.WellFoundedOn.sigma_lex_of_wellFoundedOn_fiber (hι : s.WellFoundedO @WellFounded.sigma_lex_of_wellFoundedOn_fiber _ s _ _ rπ (fun c => f c) (fun i c => g _ c) hι fun i => (hπ i).onFun.mono (fun b c h => _) swap - exact fun _ x => ⟨x, x.1.2, x.2⟩ - assumption + · exact fun _ x => ⟨x, x.1.2, x.2⟩ + · assumption #align set.well_founded_on.sigma_lex_of_well_founded_on_fiber Set.WellFoundedOn.sigma_lex_of_wellFoundedOn_fiber end SigmaLex diff --git a/Mathlib/Probability/Kernel/MeasurableIntegral.lean b/Mathlib/Probability/Kernel/MeasurableIntegral.lean index 96da308fc211f1..52796afbec74f8 100644 --- a/Mathlib/Probability/Kernel/MeasurableIntegral.lean +++ b/Mathlib/Probability/Kernel/MeasurableIntegral.lean @@ -321,9 +321,9 @@ theorem StronglyMeasurable.integral_kernel_prod_right'' {f : β × γ → E} -- Porting note: was (`Function.comp` reducibility) -- refine' MeasureTheory.StronglyMeasurable.integral_kernel_prod_right' _ -- exact hf.comp_measurable (measurable_fst.snd.prod_mk measurable_snd) - have := MeasureTheory.StronglyMeasurable.integral_kernel_prod_right' (κ := η) - (hf.comp_measurable (measurable_fst.snd.prod_mk measurable_snd)) - simpa using this + · have := MeasureTheory.StronglyMeasurable.integral_kernel_prod_right' (κ := η) + (hf.comp_measurable (measurable_fst.snd.prod_mk measurable_snd)) + simpa using this #align measure_theory.strongly_measurable.integral_kernel_prod_right'' MeasureTheory.StronglyMeasurable.integral_kernel_prod_right'' theorem StronglyMeasurable.integral_kernel_prod_left ⦃f : β → α → E⦄ @@ -345,9 +345,9 @@ theorem StronglyMeasurable.integral_kernel_prod_left'' {f : γ × β → E} (hf -- Porting note: was (`Function.comp` reducibility) -- refine' MeasureTheory.StronglyMeasurable.integral_kernel_prod_left' _ -- exact hf.comp_measurable (measurable_fst.prod_mk measurable_snd.snd) - have := MeasureTheory.StronglyMeasurable.integral_kernel_prod_left' (κ := η) - (hf.comp_measurable (measurable_fst.prod_mk measurable_snd.snd)) - simpa using this + · have := MeasureTheory.StronglyMeasurable.integral_kernel_prod_left' (κ := η) + (hf.comp_measurable (measurable_fst.prod_mk measurable_snd.snd)) + simpa using this #align measure_theory.strongly_measurable.integral_kernel_prod_left'' MeasureTheory.StronglyMeasurable.integral_kernel_prod_left'' end MeasureTheory diff --git a/Mathlib/Probability/StrongLaw.lean b/Mathlib/Probability/StrongLaw.lean index 7a6ee4f4172c81..043c7cf0e2744c 100644 --- a/Mathlib/Probability/StrongLaw.lean +++ b/Mathlib/Probability/StrongLaw.lean @@ -367,7 +367,8 @@ theorem sum_variance_truncation_le {X : Ω → ℝ} (hint : Integrable X) (hnonn ↑2 / (↑k + ↑1) * x ^ 2 = x / (k + 1) * (2 * x) := by ring _ ≤ 1 * (2 * x) := (mul_le_mul_of_nonneg_right (by - convert (div_le_one _).2 hx.2; norm_cast + convert (div_le_one _).2 hx.2 + · norm_cast simp only [Nat.cast_add, Nat.cast_one] linarith only [show (0 : ℝ) ≤ k from Nat.cast_nonneg k]) (mul_nonneg zero_le_two ((Nat.cast_nonneg k).trans hx.1.le))) diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index 3d8d915ef2e728..0e04bae1700c58 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -383,7 +383,7 @@ theorem order_ofForallLtEqZero [Zero Γ] (f : Γ → R) (hf : f ≠ 0) (n : Γ) dsimp only [order] by_cases h : ofSuppBddBelow f (forallLTEqZero_supp_BddBelow f n hn) = 0 cases h - exact (hf rfl).elim + · exact (hf rfl).elim simp_all only [dite_false] rw [Set.IsWF.le_min_iff] intro m hm diff --git a/Mathlib/RingTheory/Ideal/Cotangent.lean b/Mathlib/RingTheory/Ideal/Cotangent.lean index b1498c87a2f813..de2affe43b205b 100644 --- a/Mathlib/RingTheory/Ideal/Cotangent.lean +++ b/Mathlib/RingTheory/Ideal/Cotangent.lean @@ -236,8 +236,8 @@ lemma CotangentSpace.span_image_eq_top_iff [IsNoetherianRing R] {s : Set (maxima Submodule.span R s = ⊤ := by rw [← map_eq_top_iff, ← (Submodule.restrictScalars_injective R ..).eq_iff, Submodule.restrictScalars_span] - simp only [Ideal.toCotangent_apply, Submodule.restrictScalars_top, Submodule.map_span] - exact Ideal.Quotient.mk_surjective + · simp only [Ideal.toCotangent_apply, Submodule.restrictScalars_top, Submodule.map_span] + · exact Ideal.Quotient.mk_surjective open FiniteDimensional diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index db83e381836013..eb03fed035c4b0 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -448,8 +448,8 @@ theorem multiplicity_add_of_gt {p a b : α} (h : multiplicity p b < multiplicity cases' PartENat.ne_top_iff.mp (PartENat.ne_top_of_lt h) with k hk rw [hk] rw_mod_cast [multiplicity_lt_iff_not_dvd, dvd_add_right] - intro h_dvd - · apply multiplicity.is_greatest _ h_dvd + · intro h_dvd + apply multiplicity.is_greatest _ h_dvd rw [hk, ← Nat.succ_eq_add_one] norm_cast apply Nat.lt_succ_self k diff --git a/Mathlib/SetTheory/Game/Basic.lean b/Mathlib/SetTheory/Game/Basic.lean index ec92a0d3dca5ef..fac70589fe2c06 100644 --- a/Mathlib/SetTheory/Game/Basic.lean +++ b/Mathlib/SetTheory/Game/Basic.lean @@ -494,8 +494,8 @@ def negMulRelabelling (x y : PGame.{u}) : -x * y ≡r -(x * y) := apply ((negAddRelabelling _ _).trans _).symm apply ((negAddRelabelling _ _).trans (Relabelling.addCongr _ _)).subCongr -- Porting note: we used to just do `<;> exact (negMulRelabelling _ _).symm` from here. - exact (negMulRelabelling _ _).symm - exact (negMulRelabelling _ _).symm + · exact (negMulRelabelling _ _).symm + · exact (negMulRelabelling _ _).symm -- Porting note: not sure what has gone wrong here. -- The goal is hideous here, and the `exact` doesn't work, -- but if we just `change` it to look like the mathlib3 goal then we're fine!? @@ -866,7 +866,7 @@ instance uniqueInvTy (l r : Type u) [IsEmpty l] [IsEmpty r] : Unique (InvTy l r { InvTy.instInhabited l r with uniq := by rintro (a | a | a) - rfl + · rfl all_goals exact isEmptyElim a } #align pgame.unique_inv_ty SetTheory.PGame.uniqueInvTy diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index bf0e87367c8be8..0e2e7592fef046 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -619,8 +619,8 @@ theorem one_add_omega : 1 + ω = ω := by rw [omega, ← lift_one.{_, 0}, ← lift_add, lift_le, ← type_unit, ← type_sum_lex] refine' ⟨RelEmbedding.collapse (RelEmbedding.ofMonotone _ _)⟩ · apply Sum.rec - exact fun _ => 0 - exact Nat.succ + · exact fun _ => 0 + · exact Nat.succ · intro a b cases a <;> cases b <;> intro H <;> cases' H with _ _ H _ _ H <;> [exact H.elim; exact Nat.succ_pos _; exact Nat.succ_lt_succ H] @@ -983,7 +983,9 @@ theorem isLimit_add_iff {a b} : IsLimit (a + b) ↔ IsLimit b ∨ b = 0 ∧ IsLi apply sub_isLimit h suffices a + 0 < a + b by simpa only [add_zero] using this rwa [add_lt_add_iff_left, Ordinal.pos_iff_ne_zero] - rcases h with (h | ⟨rfl, h⟩); exact add_isLimit a h; simpa only [add_zero] + rcases h with (h | ⟨rfl, h⟩) + · exact add_isLimit a h + · simpa only [add_zero] #align ordinal.is_limit_add_iff Ordinal.isLimit_add_iff theorem dvd_add_iff : ∀ {a b c : Ordinal}, a ∣ b → (a ∣ b + c ↔ a ∣ c) diff --git a/Mathlib/Tactic/CancelDenoms/Core.lean b/Mathlib/Tactic/CancelDenoms/Core.lean index c2d3242a5e9b1c..8cbf9b57136133 100644 --- a/Mathlib/Tactic/CancelDenoms/Core.lean +++ b/Mathlib/Tactic/CancelDenoms/Core.lean @@ -96,9 +96,9 @@ theorem cancel_factors_eq {α} [Field α] {a b ad bd a' b' gcd : α} (ha : ad * · intro h simp only [← mul_assoc] at h refine' mul_left_cancel₀ (mul_ne_zero _ _) h - apply mul_ne_zero - apply div_ne_zero - exact one_ne_zero + on_goal 1 => apply mul_ne_zero + on_goal 1 => apply div_ne_zero + · exact one_ne_zero all_goals assumption #align cancel_factors.cancel_factors_eq CancelDenoms.cancel_factors_eq diff --git a/Mathlib/Testing/SlimCheck/Gen.lean b/Mathlib/Testing/SlimCheck/Gen.lean index be98d44acb90e7..bf8c7ebfa511c2 100644 --- a/Mathlib/Testing/SlimCheck/Gen.lean +++ b/Mathlib/Testing/SlimCheck/Gen.lean @@ -53,9 +53,9 @@ lemma chooseNatLt_aux {lo hi : Nat} (a : Nat) (h : Nat.succ lo ≤ a ∧ a ≤ h lo ≤ Nat.pred a ∧ Nat.pred a < hi := And.intro (Nat.le_sub_one_of_lt (Nat.lt_of_succ_le h.left)) <| show a.pred.succ ≤ hi by - rw [Nat.succ_pred_eq_of_pos] - exact h.right - exact lt_of_le_of_lt (Nat.zero_le lo) h.left + rw [Nat.succ_pred_eq_of_pos] + · exact h.right + · exact lt_of_le_of_lt (Nat.zero_le lo) h.left /-- Generate a `Nat` example between `x` and `y` (exclusively). -/ def chooseNatLt (lo hi : Nat) (h : lo < hi) : Gen {a // lo ≤ a ∧ a < hi} := diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index ebe1fb565ca4eb..881c2acd3c08c0 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -648,8 +648,8 @@ theorem dense_compl_singleton_iff_not_open : exact (hd.inter_open_nonempty _ ho (singleton_nonempty _)).ne_empty (inter_compl_self _) · refine' fun ho => dense_iff_inter_open.2 fun U hU hne => inter_compl_nonempty_iff.2 fun hUx => _ obtain rfl : U = {x} - exact eq_singleton_iff_nonempty_unique_mem.2 ⟨hne, hUx⟩ - exact ho hU + · exact eq_singleton_iff_nonempty_unique_mem.2 ⟨hne, hUx⟩ + · exact ho hU #align dense_compl_singleton_iff_not_open dense_compl_singleton_iff_not_open /-! diff --git a/Mathlib/Topology/ContinuousFunction/Bounded.lean b/Mathlib/Topology/ContinuousFunction/Bounded.lean index 2218d7444bb9f3..61b2918f77224e 100644 --- a/Mathlib/Topology/ContinuousFunction/Bounded.lean +++ b/Mathlib/Topology/ContinuousFunction/Bounded.lean @@ -668,8 +668,8 @@ instance instAdd : Add (α →ᵇ β) where rw [Prod.dist_eq] refine' mul_le_mul_of_nonneg_left _ (LipschitzAdd.C β).coe_nonneg apply max_le_max - exact Classical.choose_spec f.bounded x y - exact Classical.choose_spec g.bounded x y) + · exact Classical.choose_spec f.bounded x y + · exact Classical.choose_spec g.bounded x y) @[simp] theorem coe_add : ⇑(f + g) = f + g := rfl diff --git a/Mathlib/Topology/TietzeExtension.lean b/Mathlib/Topology/TietzeExtension.lean index 9af762108638f6..9d92cfefea8a1b 100644 --- a/Mathlib/Topology/TietzeExtension.lean +++ b/Mathlib/Topology/TietzeExtension.lean @@ -317,9 +317,9 @@ theorem exists_extension_forall_exists_le_ge_of_closedEmbedding [Nonempty X] (f inhabit X -- Put `a = ⨅ x, f x` and `b = ⨆ x, f x` obtain ⟨a, ha⟩ : ∃ a, IsGLB (range f) a - exact ⟨_, isGLB_ciInf (Real.isBounded_iff_bddBelow_bddAbove.1 f.isBounded_range).1⟩ + · exact ⟨_, isGLB_ciInf (Real.isBounded_iff_bddBelow_bddAbove.1 f.isBounded_range).1⟩ obtain ⟨b, hb⟩ : ∃ b, IsLUB (range f) b - exact ⟨_, isLUB_ciSup (Real.isBounded_iff_bddBelow_bddAbove.1 f.isBounded_range).2⟩ + · exact ⟨_, isLUB_ciSup (Real.isBounded_iff_bddBelow_bddAbove.1 f.isBounded_range).2⟩ -- Then `f x ∈ [a, b]` for all `x` have hmem : ∀ x, f x ∈ Icc a b := fun x => ⟨ha.1 ⟨x, rfl⟩, hb.1 ⟨x, rfl⟩⟩ -- Rule out the trivial case `a = b`