diff --git a/Mathlib/Algebra/BigOperators/Basic.lean b/Mathlib/Algebra/BigOperators/Basic.lean index 452c55e29a58a..1bd95a0eef0ab 100644 --- a/Mathlib/Algebra/BigOperators/Basic.lean +++ b/Mathlib/Algebra/BigOperators/Basic.lean @@ -2111,8 +2111,7 @@ theorem finset_sum_eq_sup_iff_disjoint {β : Type _} {i : Finset β} {f : β → i.sum f = i.sup f ↔ ∀ (x) (_ : x ∈ i) (y) (_ : y ∈ i), x ≠ y → Multiset.Disjoint (f x) (f y) := by induction' i using Finset.cons_induction_on with z i hz hr - · - simp only [Finset.not_mem_empty, IsEmpty.forall_iff, imp_true_iff, Finset.sum_empty, + · simp only [Finset.not_mem_empty, IsEmpty.forall_iff, imp_true_iff, Finset.sum_empty, Finset.sup_empty, bot_eq_zero, eq_self_iff_true] · simp_rw [Finset.sum_cons hz, Finset.sup_cons, Finset.mem_cons, Multiset.sup_eq_union, forall_eq_or_imp, Ne.def, IsEmpty.forall_iff, true_and_iff, diff --git a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean index 711926cff0cbf..b913ec0eb3fe8 100644 --- a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean @@ -62,8 +62,8 @@ def adjoinOneAdj : adjoinOne ⊣ forget₂ MonCat.{u} SemigroupCat.{u} := simp only [Equiv.symm_symm, adjoinOne_map, coe_comp] simp_rw [WithOne.map] cases x - . rfl - . simp + · rfl + · simp rfl } #align adjoin_one_adj adjoinOneAdj #align adjoin_zero_adj adjoinZeroAdj diff --git a/Mathlib/Algebra/Category/MonCat/Colimits.lean b/Mathlib/Algebra/Category/MonCat/Colimits.lean index 3cd8eef8d7d62..f5669427b8414 100644 --- a/Mathlib/Algebra/Category/MonCat/Colimits.lean +++ b/Mathlib/Algebra/Category/MonCat/Colimits.lean @@ -227,17 +227,17 @@ def descFun (s : Cocone F) : ColimitType F → s.pt := by · intro x y r induction' r with _ _ _ _ h _ _ _ _ _ h₁ h₂ _ _ f x _ _ _ _ _ _ _ _ h _ _ _ _ h <;> try simp -- symm - . exact h.symm + · exact h.symm -- trans - . exact h₁.trans h₂ + · exact h₁.trans h₂ -- map - . exact s.w_apply f x + · exact s.w_apply f x -- mul_1 - . rw [h] + · rw [h] -- mul_2 - . rw [h] + · rw [h] -- mul_assoc - . rw [mul_assoc] + · rw [mul_assoc] set_option linter.uppercaseLean3 false in #align Mon.colimits.desc_fun MonCat.Colimits.descFun @@ -261,12 +261,12 @@ def colimitIsColimit : IsColimit (colimitCocone F) where ext x induction' x using Quot.inductionOn with x induction' x with j x x y hx hy - . change _ = s.ι.app j _ + · change _ = s.ι.app j _ rw [← w j] rfl - . rw [quot_one, map_one] + · rw [quot_one, map_one] rfl - . rw [quot_mul, map_mul, hx, hy] + · rw [quot_mul, map_mul, hx, hy] dsimp [descMorphism, FunLike.coe, descFun] simp only [← quot_mul, descFunLift] set_option linter.uppercaseLean3 false in diff --git a/Mathlib/Algebra/DirectLimit.lean b/Mathlib/Algebra/DirectLimit.lean index 37c3a788bb219..1638a8b1b65c9 100644 --- a/Mathlib/Algebra/DirectLimit.lean +++ b/Mathlib/Algebra/DirectLimit.lean @@ -489,8 +489,7 @@ theorem of.zero_exact_aux2 {x : FreeCommRing (Σi, G i)} {s t} (hxs : IsSupporte f' j k hjk (lift (fun ix : s => f' ix.1.1 j (hj ix ix.2) ix.1.2) (restriction s x)) = lift (fun ix : t => f' ix.1.1 k (hk ix ix.2) ix.1.2) (restriction t x) := by refine' Subring.InClosure.recOn hxs _ _ _ _ - · - rw [(restriction _).map_one, (FreeCommRing.lift _).map_one, (f' j k hjk).map_one, + · rw [(restriction _).map_one, (FreeCommRing.lift _).map_one, (f' j k hjk).map_one, (restriction _).map_one, (FreeCommRing.lift _).map_one] · -- porting note: Lean 3 had `(FreeCommRing.lift _).map_neg` but I needed to replace it with -- `RingHom.map_neg` to get the rewrite to compile diff --git a/Mathlib/Algebra/GCDMonoid/Finset.lean b/Mathlib/Algebra/GCDMonoid/Finset.lean index e40b8ba2270ec..54c67e13a2f43 100644 --- a/Mathlib/Algebra/GCDMonoid/Finset.lean +++ b/Mathlib/Algebra/GCDMonoid/Finset.lean @@ -168,8 +168,7 @@ theorem dvd_gcd {a : α} : (∀ b ∈ s, a ∣ f b) → a ∣ s.gcd f := theorem gcd_insert [DecidableEq β] {b : β} : (insert b s : Finset β).gcd f = GCDMonoid.gcd (f b) (s.gcd f) := by by_cases h : b ∈ s - · - rw [insert_eq_of_mem h, + · rw [insert_eq_of_mem h, (gcd_eq_right_iff (f b) (s.gcd f) (Multiset.normalize_gcd (s.1.map f))).2 (gcd_dvd h)] apply fold_insert h #align finset.gcd_insert Finset.gcd_insert diff --git a/Mathlib/Algebra/GroupPower/Order.lean b/Mathlib/Algebra/GroupPower/Order.lean index 98219ab4e00c3..6a146e413cc15 100644 --- a/Mathlib/Algebra/GroupPower/Order.lean +++ b/Mathlib/Algebra/GroupPower/Order.lean @@ -403,10 +403,10 @@ theorem zero_pow_le_one : ∀ n : ℕ, (0 : R) ^ n ≤ 1 theorem pow_add_pow_le (hx : 0 ≤ x) (hy : 0 ≤ y) (hn : n ≠ 0) : x ^ n + y ^ n ≤ (x + y) ^ n := by rcases Nat.exists_eq_succ_of_ne_zero hn with ⟨k, rfl⟩ induction' k with k ih; - . have eqn : Nat.succ Nat.zero = 1 := rfl + · have eqn : Nat.succ Nat.zero = 1 := rfl rw [eqn] simp only [pow_one, le_refl] - . let n := k.succ + · let n := k.succ have h1 := add_nonneg (mul_nonneg hx (pow_nonneg hy n)) (mul_nonneg hy (pow_nonneg hx n)) have h2 := add_nonneg hx hy calc diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 07f2b6b4a0dee..67fcf47de0be1 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -359,12 +359,12 @@ def single₀MapHomologicalComplex (F : V ⥤ W) [F.Additive] : | _ + 1 => F.mapZeroObject.inv } hom_inv_id := by ext (_|_) - . simp - . exact IsZero.eq_of_src (IsZero.of_iso (isZero_zero _) F.mapZeroObject) _ _ + · simp + · exact IsZero.eq_of_src (IsZero.of_iso (isZero_zero _) F.mapZeroObject) _ _ inv_hom_id := by ext (_|_) - . simp - . exact IsZero.eq_of_src (isZero_zero _) _ _ }) + · simp + · exact IsZero.eq_of_src (isZero_zero _) _ _ }) fun f => by ext (_|_) <;> simp #align chain_complex.single₀_map_homological_complex ChainComplex.single₀MapHomologicalComplex @@ -415,12 +415,12 @@ def single₀MapHomologicalComplex (F : V ⥤ W) [F.Additive] : | _ + 1 => F.mapZeroObject.inv } hom_inv_id := by ext (_|_) - . simp - . exact IsZero.eq_of_src (IsZero.of_iso (isZero_zero _) F.mapZeroObject) _ _ + · simp + · exact IsZero.eq_of_src (IsZero.of_iso (isZero_zero _) F.mapZeroObject) _ _ inv_hom_id := by ext (_|_) - . simp - . exact IsZero.eq_of_src (isZero_zero _) _ _ }) + · simp + · exact IsZero.eq_of_src (isZero_zero _) _ _ }) fun f => by ext (_|_) <;> simp #align cochain_complex.single₀_map_homological_complex CochainComplex.single₀MapHomologicalComplex diff --git a/Mathlib/Algebra/Homology/HomologicalComplex.lean b/Mathlib/Algebra/Homology/HomologicalComplex.lean index fa308422e79df..9e4c03777eb1a 100644 --- a/Mathlib/Algebra/Homology/HomologicalComplex.lean +++ b/Mathlib/Algebra/Homology/HomologicalComplex.lean @@ -75,7 +75,7 @@ theorem d_comp_d (C : HomologicalComplex V c) (i j k : ι) : C.d i j ≫ C.d j k by_cases hij : c.Rel i j · by_cases hjk : c.Rel j k · exact C.d_comp_d' i j k hij hjk - . rw [C.shape j k hjk, comp_zero] + · rw [C.shape j k hjk, comp_zero] · rw [C.shape i j hij, zero_comp] #align homological_complex.d_comp_d HomologicalComplex.d_comp_d @@ -91,8 +91,8 @@ theorem ext {C₁ C₂ : HomologicalComplex V c} (h_X : C₁.X = C₂.X) simp only [mk.injEq, heq_eq_eq, true_and] ext i j by_cases hij: c.Rel i j - . simpa only [comp_id, id_comp, eqToHom_refl] using h_d i j hij - . rw [s₁ i j hij, s₂ i j hij] + · simpa only [comp_id, id_comp, eqToHom_refl] using h_d i j hij + · rw [s₁ i j hij, s₂ i j hij] #align homological_complex.ext HomologicalComplex.ext end HomologicalComplex @@ -188,7 +188,7 @@ theorem Hom.comm {A B : HomologicalComplex V c} (f : A.Hom B) (i j : ι) : f.f i ≫ B.d i j = A.d i j ≫ f.f j := by by_cases hij : c.Rel i j · exact f.comm' i j hij - . rw [A.shape i j hij, B.shape i j hij, comp_zero, zero_comp] + · rw [A.shape i j hij, B.shape i j hij, comp_zero, zero_comp] #align homological_complex.hom.comm HomologicalComplex.Hom.comm instance (A B : HomologicalComplex V c) : Inhabited (Hom A B) := diff --git a/Mathlib/Algebra/Homology/Homotopy.lean b/Mathlib/Algebra/Homology/Homotopy.lean index cb15060031e18..7f26205603302 100644 --- a/Mathlib/Algebra/Homology/Homotopy.lean +++ b/Mathlib/Algebra/Homology/Homotopy.lean @@ -122,7 +122,7 @@ theorem prevD_nat (C D : CochainComplex V ℕ) (i : ℕ) (f : ∀ i j, C.X i ⟶ cases i · simp only [shape, CochainComplex.prev_nat_zero, ComplexShape.up_Rel, Nat.one_ne_zero, not_false_iff, comp_zero] - . congr <;> simp + · congr <;> simp #align prev_d_nat prevD_nat -- porting note: removed @[has_nonempty_instance] @@ -555,16 +555,16 @@ def mkInductive : Homotopy e 0 where simp only [add_zero] refine' (mkInductiveAux₂ e zero comm_zero one comm_one succ i).2.2.trans _ congr - . cases i - . dsimp [fromNext, mkInductiveAux₂] + · cases i + · dsimp [fromNext, mkInductiveAux₂] rw [dif_neg] simp only - . dsimp [fromNext] + · dsimp [fromNext] simp only [ChainComplex.next_nat_succ, dite_true] rw [mkInductiveAux₃ e zero comm_zero one comm_one succ] dsimp [xNextIso] rw [Category.id_comp] - . dsimp [toPrev] + · dsimp [toPrev] erw [dif_pos, Category.comp_id] simp only [ChainComplex.prev] #align homotopy.mk_inductive Homotopy.mkInductive @@ -680,16 +680,16 @@ def mkCoinductive : Homotopy e 0 where rw [add_comm] refine' (mkCoinductiveAux₂ e zero comm_zero one comm_one succ i).2.2.trans _ congr - . cases i - . dsimp [toPrev, mkCoinductiveAux₂] + · cases i + · dsimp [toPrev, mkCoinductiveAux₂] rw [dif_neg] simp only - . dsimp [toPrev] + · dsimp [toPrev] simp only [CochainComplex.prev_nat_succ, dite_true] rw [mkCoinductiveAux₃ e zero comm_zero one comm_one succ] dsimp [xPrevIso] rw [Category.comp_id] - . dsimp [fromNext] + · dsimp [fromNext] erw [dif_pos, Category.id_comp] simp only [CochainComplex.next] #align homotopy.mk_coinductive Homotopy.mkCoinductive diff --git a/Mathlib/Algebra/Homology/Single.lean b/Mathlib/Algebra/Homology/Single.lean index db67869d17ac9..e91d23221f774 100644 --- a/Mathlib/Algebra/Homology/Single.lean +++ b/Mathlib/Algebra/Homology/Single.lean @@ -136,12 +136,12 @@ def single₀ : V ⥤ ChainComplex V ℕ where | n + 1 => 0 } map_id X := by ext (_|_) - . rfl - . simp + · rfl + · simp map_comp f g := by ext (_|_) - . rfl - . simp + · rfl + · simp #align chain_complex.single₀ ChainComplex.single₀ @[simp] @@ -239,10 +239,10 @@ def toSingle₀Equiv (C : ChainComplex V ℕ) (X : V) : | n + 1 => 0 comm' := fun i j h => by rcases i with (_|_|i) <;> cases j <;> simp only [single₀_obj_X_d, comp_zero] - . rw [C.shape, zero_comp] + · rw [C.shape, zero_comp] simp - . exact f.2.symm - . rw [C.shape, zero_comp] + · exact f.2.symm + · rw [C.shape, zero_comp] exact i.succ_succ_ne_one.symm } left_inv f := by ext i @@ -297,9 +297,9 @@ def single₀IsoSingle : single₀ V ≅ single V _ 0 := hom_inv_id := to_single₀_ext _ _ (by simp) inv_hom_id := by ext (_|_) - . dsimp + · dsimp simp - . dsimp + · dsimp rw [Category.comp_id] }) fun f => by ext (_|_) <;> aesop_cat #align chain_complex.single₀_iso_single ChainComplex.single₀IsoSingle @@ -333,12 +333,12 @@ def single₀ : V ⥤ CochainComplex V ℕ where | n + 1 => 0 } map_id X := by ext (_|_) - . rfl - . simp + · rfl + · simp map_comp f g := by ext (_|_) - . rfl - . simp + · rfl + · simp #align cochain_complex.single₀ CochainComplex.single₀ @[simp] @@ -436,10 +436,10 @@ def fromSingle₀Equiv (C : CochainComplex V ℕ) (X : V) : comm' := fun i j h => by rcases f with ⟨f, hf⟩ rcases j with (_|_|j) <;> cases i <;> simp only [single₀_obj_X_d, zero_comp] - . rw [C.shape, comp_zero] + · rw [C.shape, comp_zero] simp - . exact hf - . rw [C.shape, comp_zero] + · exact hf + · rw [C.shape, comp_zero] simp exact j.succ_succ_ne_one.symm } left_inv f := by @@ -470,9 +470,9 @@ def single₀IsoSingle : single₀ V ≅ single V _ 0 := hom_inv_id := from_single₀_ext _ _ (by simp) inv_hom_id := by ext (_|_) - . dsimp + · dsimp simp - . dsimp + · dsimp rw [Category.id_comp] rfl } #align cochain_complex.single₀_iso_single CochainComplex.single₀IsoSingle diff --git a/Mathlib/Algebra/IndicatorFunction.lean b/Mathlib/Algebra/IndicatorFunction.lean index 741628e7e1e98..42f5647fa786b 100644 --- a/Mathlib/Algebra/IndicatorFunction.lean +++ b/Mathlib/Algebra/IndicatorFunction.lean @@ -464,11 +464,9 @@ theorem mulIndicator_mul_compl_eq_piecewise [DecidablePred (· ∈ s)] (f g : α s.mulIndicator f * sᶜ.mulIndicator g = s.piecewise f g := by ext x by_cases h : x ∈ s - · - rw [piecewise_eq_of_mem _ _ _ h, Pi.mul_apply, Set.mulIndicator_of_mem h, + · rw [piecewise_eq_of_mem _ _ _ h, Pi.mul_apply, Set.mulIndicator_of_mem h, Set.mulIndicator_of_not_mem (Set.not_mem_compl_iff.2 h), mul_one] - · - rw [piecewise_eq_of_not_mem _ _ _ h, Pi.mul_apply, Set.mulIndicator_of_not_mem h, + · rw [piecewise_eq_of_not_mem _ _ _ h, Pi.mul_apply, Set.mulIndicator_of_not_mem h, Set.mulIndicator_of_mem (Set.mem_compl h), one_mul] #align set.mul_indicator_mul_compl_eq_piecewise Set.mulIndicator_mul_compl_eq_piecewise #align set.indicator_add_compl_eq_piecewise Set.indicator_add_compl_eq_piecewise diff --git a/Mathlib/Algebra/Module/LinearMap.lean b/Mathlib/Algebra/Module/LinearMap.lean index 3e05f3933bf1a..d968392ee17e9 100644 --- a/Mathlib/Algebra/Module/LinearMap.lean +++ b/Mathlib/Algebra/Module/LinearMap.lean @@ -373,8 +373,7 @@ theorem _root_.preimage_smul_setₛₗ [SemilinearMapClass F σ M M₃] {c : R} apply Set.Subset.antisymm · rintro x ⟨y, ys, hy⟩ refine' ⟨(hc.unit.inv : R) • x, _, _⟩ - · - simp only [← hy, smul_smul, Set.mem_preimage, Units.inv_eq_val_inv, map_smulₛₗ h, ← map_mul, + · simp only [← hy, smul_smul, Set.mem_preimage, Units.inv_eq_val_inv, map_smulₛₗ h, ← map_mul, IsUnit.val_inv_mul, one_smul, map_one, ys] · simp only [smul_smul, IsUnit.mul_val_inv, one_smul, Units.inv_eq_val_inv] · rintro x ⟨y, hy, rfl⟩ diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 43627a128a539..bf975009fc558 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -237,8 +237,8 @@ theorem div_le_of_nonneg_of_le_mul (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ c * /-- One direction of `div_le_iff` where `c` is allowed to be `0` (but `b` must be nonnegative) -/ lemma mul_le_of_nonneg_of_le_div (hb : 0 ≤ b) (hc : 0 ≤ c) (h : a ≤ b / c) : a * c ≤ b := by obtain rfl | hc := hc.eq_or_lt - . simpa using hb - . rwa [le_div_iff hc] at h + · simpa using hb + · rwa [le_div_iff hc] at h #align mul_le_of_nonneg_of_le_div mul_le_of_nonneg_of_le_div theorem div_le_one_of_le (h : a ≤ b) (hb : 0 ≤ b) : a / b ≤ 1 := diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index a9abf46e8ff60..5a8b6b4954355 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -273,7 +273,7 @@ theorem isAffineOpen_iff_of_isOpenImmersion {X Y : Scheme} (f : X ⟶ Y) [H : Is dsimp [Opens.inclusion] rw [ContinuousMap.coe_mk, ContinuousMap.coe_mk, Subtype.range_coe, Subtype.range_coe] rfl - . infer_instance + · infer_instance #align algebraic_geometry.is_affine_open_iff_of_is_open_immersion AlgebraicGeometry.isAffineOpen_iff_of_isOpenImmersion instance Scheme.quasi_compact_of_affine (X : Scheme) [IsAffine X] : CompactSpace X := @@ -349,7 +349,7 @@ theorem IsAffineOpen.basicOpenIsAffine {X : Scheme} {U : Opens X} (hU : IsAffine Scheme.Spec.map (CommRingCat.ofHom (algebraMap ((X.presheaf.obj <| op U)) (Localization.Away f))).op ≫ hU.fromSpec - . exact PresheafedSpace.IsOpenImmersion.comp (hf := inferInstance) (hg := inferInstance) + · exact PresheafedSpace.IsOpenImmersion.comp (hf := inferInstance) (hg := inferInstance) convert rangeIsAffineOpenOfOpenImmersion (Scheme.Spec.map @@ -406,7 +406,7 @@ theorem Scheme.map_PrimeSpectrum_basicOpen_of_affine (X : Scheme) [IsAffine X] ((Scheme.Spec.obj (op (Scheme.Γ.obj (op X)))).basicOpen ((inv (X.isoSpec.hom.1.c.app (op ((Opens.map (inv X.isoSpec.hom).val.base).obj ⊤)))) ((X.presheaf.map (eqToHom <| by congr)) f))) - . congr + · congr · rw [← IsIso.inv_eq_inv, IsIso.inv_inv, IsIso.Iso.inv_inv, NatIso.app_hom] -- Porting note : added this `change` to prevent timeout change SpecΓIdentity.hom.app (X.presheaf.obj <| op ⊤) = _ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 2329c2d84f1df..d9137b19f0659 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -106,7 +106,7 @@ instance (priority := 100) isIso_of_isEmpty {X Y : Scheme} (f : X ⟶ Y) [IsEmpt IsIso f := by haveI : IsEmpty X.carrier := ⟨fun x => isEmptyElim (show Y.carrier from f.1.base x)⟩ have : Epi f.1.base - . rw [TopCat.epi_iff_surjective]; rintro (x : Y.carrier) + · rw [TopCat.epi_iff_surjective]; rintro (x : Y.carrier) exact isEmptyElim x apply IsOpenImmersion.to_iso #align algebraic_geometry.is_iso_of_is_empty AlgebraicGeometry.isIso_of_isEmpty diff --git a/Mathlib/AlgebraicGeometry/LocallyRingedSpace/HasColimits.lean b/Mathlib/AlgebraicGeometry/LocallyRingedSpace/HasColimits.lean index 6065e86cd9721..f7f820c4c5d91 100644 --- a/Mathlib/AlgebraicGeometry/LocallyRingedSpace/HasColimits.lean +++ b/Mathlib/AlgebraicGeometry/LocallyRingedSpace/HasColimits.lean @@ -299,7 +299,7 @@ noncomputable def coequalizerCoforkIsColimit : IsColimit (coequalizerCofork f g) set h := _ change IsLocalRingHom h suffices : IsLocalRingHom ((PresheafedSpace.stalkMap (coequalizerCofork f g).π.1 _).comp h) - . apply isLocalRingHom_of_comp _ (PresheafedSpace.stalkMap (coequalizerCofork f g).π.1 _) + · apply isLocalRingHom_of_comp _ (PresheafedSpace.stalkMap (coequalizerCofork f g).π.1 _) change IsLocalRingHom (_ ≫ PresheafedSpace.stalkMap (coequalizerCofork f g).π.val y) erw [← PresheafedSpace.stalkMap.comp] apply isLocalRingHom_stalkMap_congr _ _ (coequalizer.π_desc s.π.1 e).symm y @@ -331,7 +331,7 @@ noncomputable instance preservesCoequalizer : -- of colimit is provided later suffices : PreservesColimit (parallelPair (F.map WalkingParallelPairHom.left) (F.map WalkingParallelPairHom.right)) forgetToSheafedSpace - . apply preservesColimitOfIsoDiagram _ (diagramIsoParallelPair F).symm + · apply preservesColimitOfIsoDiagram _ (diagramIsoParallelPair F).symm apply preservesColimitOfPreservesColimitCocone (coequalizerCoforkIsColimit _ _) apply (isColimitMapCoconeCoforkEquiv _ _).symm _ dsimp only [forgetToSheafedSpace] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index 3b5d952873043..9ac40cbcc776e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -324,8 +324,8 @@ theorem AffineTargetMorphismProperty.IsLocal.affine_openCover_iff {P : AffineTar targetAffineLocally P f ↔ ∀ i, @P _ _ (pullback.snd : pullback f (𝒰.map i) ⟶ _) (h𝒰 i) := by refine' ⟨fun H => let h := ((hP.affine_openCover_TFAE f).out 0 2).mp H; _, fun H => let h := ((hP.affine_openCover_TFAE f).out 1 0).mp; _⟩ - . exact fun i => h 𝒰 i - . exact h ⟨𝒰, inferInstance, H⟩ + · exact fun i => h 𝒰 i + · exact h ⟨𝒰, inferInstance, H⟩ #align algebraic_geometry.affine_target_morphism_property.is_local.affine_open_cover_iff AlgebraicGeometry.AffineTargetMorphismProperty.IsLocal.affine_openCover_iff theorem AffineTargetMorphismProperty.IsLocal.affine_target_iff {P : AffineTargetMorphismProperty} @@ -440,8 +440,8 @@ theorem PropertyIsLocalAtTarget.openCover_iff {P : MorphismProperty Scheme} -- Porting note : couldn't get the term mode proof work refine ⟨fun H => let h := ((hP.openCover_TFAE f).out 0 2).mp H; fun i => ?_, fun H => let h := ((hP.openCover_TFAE f).out 1 0).mp; ?_⟩ - . exact h 𝒰 i - . exact h ⟨𝒰, H⟩ + · exact h 𝒰 i + · exact h ⟨𝒰, H⟩ #align algebraic_geometry.property_is_local_at_target.open_cover_iff AlgebraicGeometry.PropertyIsLocalAtTarget.openCover_iff namespace AffineTargetMorphismProperty @@ -563,8 +563,8 @@ theorem AffineTargetMorphismProperty.diagonalOfTargetAffineLocally -- Porting note : added this instance haveI hg₁ : IsOpenImmersion g₁ := by apply (config := { allowSynthFailures := true }) Scheme.pullback_map_isOpenImmersion - . exact PresheafedSpace.IsOpenImmersion.comp (hf := hf₁) _ - . exact PresheafedSpace.IsOpenImmersion.comp (hf := hf₂) _ + · exact PresheafedSpace.IsOpenImmersion.comp (hf := hf₁) _ + · exact PresheafedSpace.IsOpenImmersion.comp (hf := hf₂) _ specialize H g₁ rw [← affine_cancel_left_isIso hP.1 (pullbackDiagonalMapIso f _ f₁ f₂).hom] convert H diff --git a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean index 42a4cca290829..5395baf925ff8 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/QuasiCompact.lean @@ -326,7 +326,7 @@ theorem exists_pow_mul_eq_zero_of_res_basicOpen_eq_zero_of_isCompact (X : Scheme convert congr_arg (X.presheaf.map (homOfLE _).op) H · simp only [← comp_apply, ← Functor.map_comp] rfl - . rw [map_zero] + · rw [map_zero] simp only [Scheme.basicOpen_res, ge_iff_le, inf_le_right] choose n hn using H' haveI := hs.to_subtype diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion/Basic.lean b/Mathlib/AlgebraicGeometry/OpenImmersion/Basic.lean index 0efbc34a31429..c6c624cf32ba4 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion/Basic.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion/Basic.lean @@ -117,7 +117,7 @@ noncomputable def isoRestrict : X ≅ Y.restrict H.base_open := PresheafedSpace.isoOfComponents (Iso.refl _) <| by symm fapply NatIso.ofComponents - . intro U + · intro U refine' asIso (f.c.app (op (H.openFunctor.obj (unop U)))) ≪≫ X.presheaf.mapIso (eqToIso _) · induction U using Opposite.rec' with | h U => ?_ cases U @@ -299,7 +299,7 @@ instance ofRestrict {X : TopCat} (Y : PresheafedSpace C) {f : X ⟶ Y.carrier} apply Subsingleton.helim rw [this] rfl - . infer_instance + · infer_instance #align algebraic_geometry.PresheafedSpace.is_open_immersion.of_restrict AlgebraicGeometry.PresheafedSpace.IsOpenImmersion.ofRestrict @[elementwise, simp] @@ -317,7 +317,7 @@ theorem to_iso (f : X ⟶ Y) [h : IsOpenImmersion f] [h' : Epi f.base] : IsIso f -- Porting Note : was `apply (config := { instances := False }) ...` -- See https://github.com/leanprover/lean4/issues/2273 have : ∀ (U : (Opens Y)ᵒᵖ), IsIso (f.c.app U) - . intro U + · intro U have : U = op (h.openFunctor.obj ((Opens.map f.base).obj (unop U))) := by induction U using Opposite.rec' with | h U => ?_ cases U @@ -531,7 +531,7 @@ instance forgetPreservesLimitsOfLeft : PreservesLimit (cospan f g) (forget C) := apply (IsLimit.postcomposeHomEquiv (diagramIsoCospan _) _).toFun refine' (IsLimit.equivIsoLimit _).toFun (limit.isLimit (cospan f.base g.base)) fapply Cones.ext - . exact Iso.refl _ + · exact Iso.refl _ change ∀ j, _ = 𝟙 _ ≫ _ ≫ _ simp_rw [Category.id_comp] rintro (_ | _ | _) <;> symm @@ -728,7 +728,7 @@ instance forgetMapIsOpenImmersion : PresheafedSpace.IsOpenImmersion ((forget).ma instance hasLimit_cospan_forget_of_left : HasLimit (cospan f g ⋙ forget) := by have : HasLimit (cospan ((cospan f g ⋙ forget).map Hom.inl) ((cospan f g ⋙ forget).map Hom.inr)) - . change HasLimit (cospan ((forget).map f) ((forget).map g)) + · change HasLimit (cospan ((forget).map f) ((forget).map g)) infer_instance apply hasLimitOfIso (diagramIsoCospan _).symm #align algebraic_geometry.SheafedSpace.is_open_immersion.has_limit_cospan_forget_of_left AlgebraicGeometry.SheafedSpace.IsOpenImmersion.hasLimit_cospan_forget_of_left @@ -740,7 +740,7 @@ instance hasLimit_cospan_forget_of_left' : instance hasLimit_cospan_forget_of_right : HasLimit (cospan g f ⋙ forget) := by have : HasLimit (cospan ((cospan g f ⋙ forget).map Hom.inl) ((cospan g f ⋙ forget).map Hom.inr)) - . change HasLimit (cospan ((forget).map g) ((forget).map f)) + · change HasLimit (cospan ((forget).map g) ((forget).map f)) infer_instance apply hasLimitOfIso (diagramIsoCospan _).symm #align algebraic_geometry.SheafedSpace.is_open_immersion.has_limit_cospan_forget_of_right AlgebraicGeometry.SheafedSpace.IsOpenImmersion.hasLimit_cospan_forget_of_right @@ -772,7 +772,7 @@ instance sheafedSpaceForgetPreservesOfLeft : PreservesLimit (cospan f g) (Sheafe have : PreservesLimit (cospan ((cospan f g ⋙ forget).map Hom.inl) ((cospan f g ⋙ forget).map Hom.inr)) (PresheafedSpace.forget C) - . dsimp + · dsimp infer_instance apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm #align algebraic_geometry.SheafedSpace.is_open_immersion.SheafedSpace_forget_preserves_of_left AlgebraicGeometry.SheafedSpace.IsOpenImmersion.sheafedSpaceForgetPreservesOfLeft @@ -1013,7 +1013,7 @@ def pullbackConeOfLeftIsLimit : IsLimit (pullbackConeOfLeft f g) := rw [PresheafedSpace.stalkMap.comp, ← IsIso.eq_inv_comp] at this rw [this] infer_instance - . intro m _ h₂ + · intro m _ h₂ rw [← cancel_mono (pullbackConeOfLeft f g).snd] exact h₂.trans <| LocallyRingedSpace.Hom.ext _ _ (PresheafedSpace.IsOpenImmersion.pullbackConeOfLeftLift_snd f.1 g.1 <| @@ -1084,10 +1084,10 @@ instance forgetToTopPreservesPullbackOfLeft : WalkingCospan.Hom.inl) ((cospan f g ⋙ forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace).map WalkingCospan.Hom.inr)) (PresheafedSpace.forget CommRingCat) - . dsimp; infer_instance + · dsimp; infer_instance have : PreservesLimit (cospan f g ⋙ forgetToSheafedSpace ⋙ SheafedSpace.forgetToPresheafedSpace) (PresheafedSpace.forget CommRingCat) - . apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm + · apply preservesLimitOfIsoDiagram _ (diagramIsoCospan _).symm apply Limits.compPreservesLimit #align algebraic_geometry.LocallyRingedSpace.is_open_immersion.forget_to_Top_preserves_pullback_of_left AlgebraicGeometry.LocallyRingedSpace.IsOpenImmersion.forgetToTopPreservesPullbackOfLeft diff --git a/Mathlib/AlgebraicGeometry/OpenImmersion/Scheme.lean b/Mathlib/AlgebraicGeometry/OpenImmersion/Scheme.lean index 1e50056d38cb1..b920cac10acf1 100644 --- a/Mathlib/AlgebraicGeometry/OpenImmersion/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/OpenImmersion/Scheme.lean @@ -58,7 +58,7 @@ protected def scheme (X : LocallyRingedSpace) skip apply PresheafedSpace.IsOpenImmersion.isoOfRangeEq (PresheafedSpace.ofRestrict _ _) f.1 · exact Subtype.range_coe_subtype - . exact Opens.openEmbedding _ -- Porting note : was `infer_instance` + · exact Opens.openEmbedding _ -- Porting note : was `infer_instance` #align algebraic_geometry.LocallyRingedSpace.IsOpenImmersion.Scheme AlgebraicGeometry.LocallyRingedSpace.IsOpenImmersion.scheme end LocallyRingedSpace.IsOpenImmersion @@ -215,8 +215,8 @@ instance basic_open_isOpenImmersion {R : CommRingCat} (f : R) : AlgebraicGeometry.IsOpenImmersion (Scheme.Spec.map (CommRingCat.ofHom (algebraMap R (Localization.Away f))).op) := by apply SheafedSpace.IsOpenImmersion.of_stalk_iso (H := ?_) - . exact (PrimeSpectrum.localization_away_openEmbedding (Localization.Away f) f : _) - . intro x + · exact (PrimeSpectrum.localization_away_openEmbedding (Localization.Away f) f : _) + · intro x exact Spec_map_localization_isIso R (Submonoid.powers f) x #align algebraic_geometry.Scheme.basic_open_IsOpenImmersion AlgebraicGeometry.Scheme.basic_open_isOpenImmersion @@ -538,7 +538,7 @@ instance forgetToTopPreservesOfLeft : PreservesLimit (cospan f g) Scheme.forgetT delta Scheme.forgetToTop apply @Limits.compPreservesLimit (K := cospan f g) (F := forget) (G := LocallyRingedSpace.forgetToTop) ?_ ?_ - . infer_instance + · infer_instance apply @preservesLimitOfIsoDiagram (F := _) _ _ _ _ _ _ (diagramIsoCospan.{u} _).symm ?_ dsimp [LocallyRingedSpace.forgetToTop] infer_instance @@ -1030,7 +1030,7 @@ theorem morphismRestrict_c_app {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : O ext1; exact Set.preimage_image_eq _ Subtype.coe_injective have : _ ≫ X.presheaf.map _ = _ := (((f ∣_ U).1.c.naturality (eqToHom e).op).symm.trans ?_).trans this - . rw [← IsIso.eq_comp_inv, ← Functor.map_inv, Category.assoc] at this + · rw [← IsIso.eq_comp_inv, ← Functor.map_inv, Category.assoc] at this rw [this] congr 1 erw [← X.presheaf.map_comp, ← X.presheaf.map_comp] @@ -1092,10 +1092,10 @@ def morphismRestrictRestrict {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (V : Ope ((Category.comp_id _).trans (Category.id_comp _).symm) (by aesop_cat) ≫ (pullbackRightPullbackFstIso _ _ _).hom ≫ (pullbackSymmetry _ _).hom have i3 : IsIso h - . repeat + · repeat apply (config := { allowSynthFailures := true }) IsIso.comp_isIso have : (f ∣_ U ∣_ V) ≫ (Iso.refl _).hom = (asIso h).hom ≫ pullback.snd (f := f) (g := g) - . simp only [Category.comp_id, pullbackRightPullbackFstIso_hom_fst, Iso.refl_hom, + · simp only [Category.comp_id, pullbackRightPullbackFstIso_hom_fst, Iso.refl_hom, Category.assoc, pullbackSymmetry_hom_comp_snd, asIso_hom, pullback.lift_fst, pullbackSymmetry_hom_comp_fst] rfl @@ -1152,7 +1152,7 @@ def morphismRestrictStalkMap {X Y : Scheme} (f : X ⟶ Y) (U : Opens Y) (x) : erw [PresheafedSpace.stalkMap_germ _ (U.openEmbedding.isOpenMap.functor.obj V) ⟨x.1, ⟨⟨f.1.base x.1, x.2⟩, _, rfl⟩⟩] swap - . rw [morphismRestrict_val_base] at hxV + · rw [morphismRestrict_val_base] at hxV exact hxV erw [PresheafedSpace.restrictStalkIso_hom_eq_germ] rw [morphismRestrict_c_app, Category.assoc, TopCat.Presheaf.germ_res] diff --git a/Mathlib/AlgebraicGeometry/PresheafedSpace.lean b/Mathlib/AlgebraicGeometry/PresheafedSpace.lean index fa2f0039bedd6..e3a3f797873c5 100644 --- a/Mathlib/AlgebraicGeometry/PresheafedSpace.lean +++ b/Mathlib/AlgebraicGeometry/PresheafedSpace.lean @@ -178,7 +178,7 @@ instance categoryOfPresheafedSpaces : Category (PresheafedSpace C) where id_comp _ := by dsimp ext - . dsimp + · dsimp simp only [map_id, whiskerRight_id', assoc] erw [comp_id, comp_id] · dsimp @@ -392,7 +392,7 @@ instance ofRestrict_mono {U : TopCat} (X : PresheafedSpace C) (f : U ⟶ X.1) (h simp only [PresheafedSpace.comp_base, PresheafedSpace.ofRestrict_base] at this rw [cancel_mono] at this exact this - . ext V + · ext V have hV : (Opens.map (X.ofRestrict hf).base).obj (hf.isOpenMap.functor.obj V) = V := by ext1 exact Set.preimage_image_eq _ hf.inj diff --git a/Mathlib/AlgebraicGeometry/Pullbacks.lean b/Mathlib/AlgebraicGeometry/Pullbacks.lean index 07be759c433d2..65e16d38c87e3 100644 --- a/Mathlib/AlgebraicGeometry/Pullbacks.lean +++ b/Mathlib/AlgebraicGeometry/Pullbacks.lean @@ -675,7 +675,7 @@ def openCoverOfBase' (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCove rfl · simp only [Category.comp_id, Category.id_comp] -- Porting note : this `IsIso` instance was `inferInstance` - . apply IsIso.comp_isIso + · apply IsIso.comp_isIso #align algebraic_geometry.Scheme.pullback.open_cover_of_base' AlgebraicGeometry.Scheme.Pullback.openCoverOfBase' /-- Given an open cover `{ Zᵢ }` of `Z`, then `X ×[Z] Y` is covered by `Xᵢ ×[Zᵢ] Yᵢ`, where @@ -695,7 +695,7 @@ def openCoverOfBase (𝒰 : OpenCover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : OpenCover -- Porting note : deviated from original proof a bit so that it won't timeout. rw [Iso.refl_hom, Category.id_comp, openCoverOfBase'_map] apply pullback.hom_ext <;> dsimp <;> - . simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.assoc, + · simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.assoc, limit.lift_π_assoc, cospan_left, Category.comp_id, limit.isoLimitCone_inv_π, limit.isoLimitCone_inv_π_assoc, pullbackSymmetry_hom_comp_fst_assoc, pullbackSymmetry_hom_comp_snd_assoc] diff --git a/Mathlib/AlgebraicGeometry/RingedSpace.lean b/Mathlib/AlgebraicGeometry/RingedSpace.lean index 79d204ecca2ce..7fb952e86cce0 100644 --- a/Mathlib/AlgebraicGeometry/RingedSpace.lean +++ b/Mathlib/AlgebraicGeometry/RingedSpace.lean @@ -69,7 +69,7 @@ theorem isUnit_res_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) (x : -- desired form replace heq : (X.presheaf.germ ⟨x.val, hxW⟩) ((X.presheaf.map (U.infLELeft V).op) f * (X.presheaf.map (U.infLERight V).op) g) = (X.presheaf.germ ⟨x.val, hxW⟩) 1 - . dsimp [germ] + · dsimp [germ] erw [map_mul, map_one, show X.presheaf.germ ⟨x, hxW⟩ ((X.presheaf.map (U.infLELeft V).op) f) = X.presheaf.germ x f from X.presheaf.germ_res_apply (Opens.infLELeft U V) ⟨x.1, hxW⟩ f, show X.presheaf.germ ⟨x, hxW⟩ (X.presheaf.map (U.infLERight V).op g) = @@ -89,7 +89,7 @@ theorem isUnit_of_isUnit_germ (U : Opens X) (f : X.presheaf.obj (op U)) -- We pick a cover of `U` by open sets `V x`, such that `f` is a unit on each `V x`. choose V iVU m h_unit using fun x : U => X.isUnit_res_of_isUnit_germ U f x (h x) have hcover : U ≤ iSup V - . intro x hxU + · intro x hxU -- Porting note : in Lean3 `rw` is sufficient erw [Opens.mem_iSup] exact ⟨⟨x, hxU⟩, m ⟨x, hxU⟩⟩ diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index fbf9064cafeb0..5604a12d905a7 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -302,9 +302,9 @@ theorem mem_basicOpen (x : U) : ↑x ∈ X.basicOpen f ↔ IsUnit (X.presheaf.ge theorem mem_basicOpen_top' {U : Opens X} (f : X.presheaf.obj (op U)) (x : X.carrier) : x ∈ X.basicOpen f ↔ ∃ (m : x ∈ U), IsUnit (X.presheaf.germ (⟨x, m⟩ : U) f) := by fconstructor - . rintro ⟨y, hy1, rfl⟩ + · rintro ⟨y, hy1, rfl⟩ exact ⟨y.2, hy1⟩ - . rintro ⟨m, hm⟩ + · rintro ⟨m, hm⟩ exact ⟨⟨x, m⟩, hm, rfl⟩ @[simp] diff --git a/Mathlib/AlgebraicGeometry/Spec.lean b/Mathlib/AlgebraicGeometry/Spec.lean index 3cd4e6c573db2..ecd551980707c 100644 --- a/Mathlib/AlgebraicGeometry/Spec.lean +++ b/Mathlib/AlgebraicGeometry/Spec.lean @@ -420,7 +420,7 @@ theorem isLocalizedModule_toPushforwardStalkAlgHom_aux (y) : ⟨p, hpr⟩ _).trans e set s' := (Spec.topMap (algebraMap R S) _* (structureSheaf S).1).map (homOfLE hrU).op s with h replace e : ((Spec.topMap (algebraMap R S) _* (structureSheaf S).val).germ ⟨p, hpr⟩) s' = y - . rw [h]; exact e + · rw [h]; exact e clear_value s'; clear! U obtain ⟨⟨s, ⟨_, n, rfl⟩⟩, hsn⟩ := @IsLocalization.surj _ _ _ _ _ _ diff --git a/Mathlib/AlgebraicGeometry/StructureSheaf.lean b/Mathlib/AlgebraicGeometry/StructureSheaf.lean index 2868b3eb814e2..8048924b3ca59 100644 --- a/Mathlib/AlgebraicGeometry/StructureSheaf.lean +++ b/Mathlib/AlgebraicGeometry/StructureSheaf.lean @@ -854,7 +854,7 @@ theorem toBasicOpen_surjective (f : R) : Function.Surjective (toBasicOpen R f) : rw [← SetLike.coe_subset_coe, Opens.coe_iSup] at ht_cover replace ht_cover : (PrimeSpectrum.basicOpen f : Set <| PrimeSpectrum R) ⊆ ⋃ (i : ι) (x : i ∈ t), (PrimeSpectrum.basicOpen (h i) : Set _) - . convert ht_cover using 2 + · convert ht_cover using 2 exact funext fun j => by rw [Opens.coe_iSup] -- Next we show that some power of `f` is a linear combination of the `h i` obtain ⟨n, hn⟩ : f ∈ (Ideal.span (h '' ↑t)).radical := by @@ -862,9 +862,9 @@ theorem toBasicOpen_surjective (f : R) : Function.Surjective (toBasicOpen R f) : -- Porting note : simp with `PrimeSpectrum.basicOpen_eq_zeroLocus_compl` does not work replace ht_cover : (PrimeSpectrum.zeroLocus {f})ᶜ ⊆ ⋃ (i : ι) (x : i ∈ t), (PrimeSpectrum.zeroLocus {h i})ᶜ - . convert ht_cover - . rw [PrimeSpectrum.basicOpen_eq_zeroLocus_compl] - . simp only [Opens.iSup_mk, Opens.carrier_eq_coe, PrimeSpectrum.basicOpen_eq_zeroLocus_compl] + · convert ht_cover + · rw [PrimeSpectrum.basicOpen_eq_zeroLocus_compl] + · simp only [Opens.iSup_mk, Opens.carrier_eq_coe, PrimeSpectrum.basicOpen_eq_zeroLocus_compl] rw [Set.compl_subset_comm] at ht_cover -- Why doesn't `simp_rw` do this? simp_rw [Set.compl_iUnion, compl_compl, ← PrimeSpectrum.zeroLocus_iUnion, diff --git a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean index 45b06f4ba9a3b..0e16638251b37 100644 --- a/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean +++ b/Mathlib/AlgebraicTopology/AlternatingFaceMapComplex.lean @@ -98,9 +98,9 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by dsimp simp only [zsmul_comp, comp_zsmul, smul_smul, ← neg_smul] congr 1 - . simp only [Fin.val_succ, pow_add, pow_one, mul_neg, neg_neg, mul_one] + · simp only [Fin.val_succ, pow_add, pow_one, mul_neg, neg_neg, mul_one] apply mul_comm - . rw [CategoryTheory.SimplicialObject.δ_comp_δ''] + · rw [CategoryTheory.SimplicialObject.δ_comp_δ''] simpa using hij · -- φ : S → Sᶜ is injective rintro ⟨i, j⟩ ⟨i', j'⟩ hij hij' h @@ -112,11 +112,11 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by simp only [Finset.mem_univ, forall_true_left, Prod.forall, ge_iff_le, Finset.compl_filter, not_le, Finset.mem_filter, true_and] at hij' refine' ⟨(j'.pred _, Fin.castSucc i'), _, _⟩ - . rintro rfl + · rintro rfl simp only [Fin.val_zero, not_lt_zero'] at hij' - . simpa only [Finset.mem_univ, forall_true_left, Prod.forall, ge_iff_le, Finset.mem_filter, + · simpa only [Finset.mem_univ, forall_true_left, Prod.forall, ge_iff_le, Finset.mem_filter, Fin.coe_castSucc, Fin.coe_pred, true_and] using Nat.le_pred_of_lt hij' - . simp only [Fin.castLT_castSucc, Fin.succ_pred] + · simp only [Fin.castLT_castSucc, Fin.succ_pred] #align algebraic_topology.alternating_face_map_complex.d_squared AlgebraicTopology.AlternatingFaceMapComplex.d_squared /-! @@ -264,7 +264,7 @@ def inclusionOfMooreComplexMap (X : SimplicialObject A) : simp only [AlternatingFaceMapComplex.objD, comp_sum] rw [Fin.sum_univ_succ, Fintype.sum_eq_zero] swap - . intro j + · intro j rw [NormalizedMooreComplex.objX, comp_zsmul, ← factorThru_arrow _ _ (finset_inf_arrow_factors Finset.univ _ _ (Finset.mem_univ j)), Category.assoc, kernelSubobject_arrow_comp, comp_zero, smul_zero] diff --git a/Mathlib/AlgebraicTopology/CechNerve.lean b/Mathlib/AlgebraicTopology/CechNerve.lean index 525db5815b113..b2975f9fe50ce 100644 --- a/Mathlib/AlgebraicTopology/CechNerve.lean +++ b/Mathlib/AlgebraicTopology/CechNerve.lean @@ -138,10 +138,10 @@ def equivalenceLeftToRight (X : SimplicialObject.Augmented C) (F : Arrow C) intro x y f dsimp ext - . dsimp + · dsimp simp only [WidePullback.lift_π, Category.assoc, ← X.left.map_comp_assoc] rfl - . dsimp + · dsimp simp } right := G.right #align category_theory.simplicial_object.equivalence_left_to_right CategoryTheory.SimplicialObject.equivalenceLeftToRight @@ -173,11 +173,11 @@ def cechNerveEquiv (X : SimplicialObject.Augmented C) (F : Arrow C) : dsimp ext x : 2 · refine' WidePullback.hom_ext _ _ _ (fun j => _) _ - . dsimp + · dsimp simp rfl - . simpa using congr_app A.w.symm x - . rfl + · simpa using congr_app A.w.symm x + · rfl #align category_theory.simplicial_object.cech_nerve_equiv CategoryTheory.SimplicialObject.cechNerveEquiv /-- The augmented Čech nerve construction is right adjoint to the `toArrow` functor. -/ @@ -314,14 +314,14 @@ def cechConerveEquiv (F : Arrow C) (X : CosimplicialObject.Augmented C) : intro A dsimp ext x : 2 - . rfl - . refine' WidePushout.hom_ext _ _ _ (fun j => _) _ - . dsimp + · rfl + · refine' WidePushout.hom_ext _ _ _ (fun j => _) _ + · dsimp simp only [Category.assoc, ← NatTrans.naturality A.right, Arrow.augmentedCechConerve_right, SimplexCategory.len_mk, Arrow.cechConerve_map, colimit.ι_desc, WidePushoutShape.mkCocone_ι_app, colimit.ι_desc_assoc] rfl - . dsimp + · dsimp rw [colimit.ι_desc] exact congr_app A.w x right_inv := by diff --git a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean index c5e8499498cae..e42a52cac3e78 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Decomposition.lean @@ -75,7 +75,7 @@ theorem decomposition_Q (n q : ℕ) : simp only [Fin.val_mk, (HigherFacesVanish.of_P q n).comp_Hσ_eq hnaq', q'.rev_eq hnaq', neg_neg] rfl - . ext ⟨i, hi⟩ + · ext ⟨i, hi⟩ simp only [Nat.succ_eq_add_one, Nat.lt_succ_iff_lt_or_eq, Finset.mem_univ, forall_true_left, Finset.mem_filter, lt_self_iff_false, or_true, and_self, not_true, Finset.mem_erase, ne_eq, Fin.mk.injEq, true_and] @@ -121,7 +121,7 @@ theorem id_φ : (id X n).φ = 𝟙 _ := by simp only [← P_add_Q_f (n + 1) (n + 1), φ] congr 1 · simp only [id, PInfty_f, P_f_idem] - . exact Eq.trans (by congr ; simp) (decomposition_Q n (n + 1)).symm + · exact Eq.trans (by congr ; simp) (decomposition_Q n (n + 1)).symm #align algebraic_topology.dold_kan.morph_components.id_φ AlgebraicTopology.DoldKan.MorphComponents.id_φ variable {X n} diff --git a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean index 1e3e038b23136..8d3b81c70a451 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Faces.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Faces.lean @@ -94,21 +94,21 @@ theorem comp_Hσ_eq {Y : C} {n a q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFac swap · rintro ⟨k, hk⟩ rw [assoc, X.δ_comp_σ_of_gt', v.comp_δ_eq_zero_assoc, zero_comp, zsmul_zero] - . simp only [Fin.lt_iff_val_lt_val] + · simp only [Fin.lt_iff_val_lt_val] dsimp [Fin.natAdd, Fin.cast] linarith - . intro h + · intro h rw [Fin.pred_eq_iff_eq_succ, Fin.ext_iff] at h dsimp [Fin.cast] at h linarith - . dsimp [Fin.cast, Fin.pred] + · dsimp [Fin.cast, Fin.pred] rw [Nat.pred_eq_sub_one, Nat.succ_add_sub_one] linarith simp only [assoc] conv_lhs => congr - . rw [Fin.sum_univ_castSucc] - . rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc] + · rw [Fin.sum_univ_castSucc] + · rw [Fin.sum_univ_castSucc, Fin.sum_univ_castSucc] dsimp [Fin.cast, Fin.castLE, Fin.castLT] /- the purpose of the following `simplif` is to create three subgoals in order to finish the proof -/ @@ -159,13 +159,13 @@ theorem comp_Hσ_eq_zero {Y : C} {n q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : Higher · intro j dsimp [Fin.cast, Fin.castLE, Fin.castLT] rw [comp_zsmul, comp_zsmul, δ_comp_σ_of_gt', v.comp_δ_eq_zero_assoc, zero_comp, zsmul_zero] - . simp only [Fin.lt_iff_val_lt_val] + · simp only [Fin.lt_iff_val_lt_val] dsimp [Fin.succ] linarith - . intro h + · intro h simp only [Fin.pred, Fin.ext_iff, Nat.pred_eq_sub_one, Nat.succ_add_sub_one, Fin.val_zero, add_eq_zero, false_and] at h - . simp only [Fin.pred, Nat.pred_eq_sub_one, Nat.succ_add_sub_one] + · simp only [Fin.pred, Nat.pred_eq_sub_one, Nat.succ_add_sub_one] linarith set_option linter.uppercaseLean3 false in #align algebraic_topology.dold_kan.higher_faces_vanish.comp_Hσ_eq_zero AlgebraicTopology.DoldKan.HigherFacesVanish.comp_Hσ_eq_zero diff --git a/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean b/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean index 6b7f764f33d1a..b36555aa7d399 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/HomotopyEquivalence.lean @@ -71,7 +71,7 @@ def homotopyPInftyToId : Homotopy (PInfty : K[X] ⟶ _) (𝟙 _) where zero i j hij := Homotopy.zero _ i j hij comm n := by rcases n with _|n - . simpa only [Homotopy.dNext_zero_chainComplex, Homotopy.prevD_chainComplex, + · simpa only [Homotopy.dNext_zero_chainComplex, Homotopy.prevD_chainComplex, PInfty_f, Nat.zero_eq, P_f_0_eq, zero_add] using (homotopyPToId X 2).comm 0 · -- Porting note: this branch had been: -- simpa only [Homotopy.dNext_succ_chainComplex, Homotopy.prevD_chainComplex, diff --git a/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean b/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean index d750ba10006ed..65c41cad5ede4 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NCompGamma.lean @@ -67,7 +67,7 @@ theorem PInfty_comp_map_mono_eq_zero (X : SimplicialObject C) {n : ℕ} {Δ' : S dsimp at h' linarith by_cases hj₁ : j₁ = 0 - . subst hj₁ + · subst hj₁ rw [assoc, ← SimplexCategory.δ_comp_δ'' (Fin.zero_le _)] simp only [op_comp, X.map_comp, assoc, PInfty_f] erw [(HigherFacesVanish.of_P _ _).comp_δ_eq_zero_assoc _ j₂.succ_ne_zero, zero_comp] diff --git a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean index ba9be12b5dee3..b9c69b49e2369 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/NReflectsIso.lean @@ -76,7 +76,7 @@ theorem compatibility_N₂_N₁_karoubi : · ext n · dsimp simp only [karoubi_PInfty_f, comp_id, PInfty_f_naturality, id_comp, eqToHom_refl] - . rfl + · rfl · rintro _ n (rfl : n + 1 = _) ext have h := (AlternatingFaceMapComplex.map P.p).comm (n + 1) n diff --git a/Mathlib/AlgebraicTopology/DoldKan/Projections.lean b/Mathlib/AlgebraicTopology/DoldKan/Projections.lean index 2cbe83468235f..d63e2fdcbbc9b 100644 --- a/Mathlib/AlgebraicTopology/DoldKan/Projections.lean +++ b/Mathlib/AlgebraicTopology/DoldKan/Projections.lean @@ -129,7 +129,7 @@ theorem comp_P_eq_self {Y : C} {n q : ℕ} {φ : Y ⟶ X _[n + 1]} (v : HigherFa comp_id, ← assoc, hq v.of_succ, add_right_eq_self] by_cases hqn : n < q · exact v.of_succ.comp_Hσ_eq_zero hqn - . obtain ⟨a, ha⟩ := Nat.le.dest (not_lt.mp hqn) + · obtain ⟨a, ha⟩ := Nat.le.dest (not_lt.mp hqn) have hnaq : n = a + q := by linarith simp only [v.of_succ.comp_Hσ_eq hnaq, neg_eq_zero, ← assoc] have eq := v ⟨a, by linarith⟩ (by diff --git a/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean b/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean index c068e12a9ea02..63b5a8becf3da 100644 --- a/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean +++ b/Mathlib/AlgebraicTopology/ExtraDegeneracy.lean @@ -218,9 +218,9 @@ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) : ext j : 2 dsimp [SimplicialObject.δ, SimplexCategory.δ, SSet.standardSimplex] by_cases j = 0 - . subst h + · subst h simp only [Fin.succ_succAbove_zero, shiftFun_0] - . obtain ⟨_, rfl⟩ := Fin.eq_succ_of_ne_zero h + · obtain ⟨_, rfl⟩ := Fin.eq_succ_of_ne_zero h simp only [Fin.succ_succAbove_succ, shiftFun_succ, Function.comp_apply] s_comp_σ n i := by ext1 φ @@ -399,7 +399,7 @@ noncomputable def homotopyEquiv {C : Type _} [Category C] [Preadditive C] [HasZe · simp only [eq_self_iff_true] comm := fun i => by rcases i with _|i - . rw [Homotopy.prevD_chainComplex, Homotopy.dNext_zero_chainComplex, zero_add] + · rw [Homotopy.prevD_chainComplex, Homotopy.dNext_zero_chainComplex, zero_add] dsimp [ChainComplex.fromSingle₀Equiv, ChainComplex.toSingle₀Equiv] simp only [comp_id, ite_true, zero_add, ComplexShape.down_Rel, not_true, AlternatingFaceMapComplex.obj_d_eq, Preadditive.neg_comp] diff --git a/Mathlib/AlgebraicTopology/MooreComplex.lean b/Mathlib/AlgebraicTopology/MooreComplex.lean index dc568a1d77e92..c7d6f016d1137 100644 --- a/Mathlib/AlgebraicTopology/MooreComplex.lean +++ b/Mathlib/AlgebraicTopology/MooreComplex.lean @@ -95,11 +95,11 @@ theorem d_squared (n : ℕ) : objD X (n + 1) ≫ objD X n = 0 := by -- It's a pity we need to do a case split here; -- after the first erw the proofs are almost identical rcases n with _ | n <;> dsimp [objD] - . erw [Subobject.factorThru_arrow_assoc, Category.assoc, + · erw [Subobject.factorThru_arrow_assoc, Category.assoc, ← X.δ_comp_δ_assoc (Fin.zero_le (0 : Fin 2)), ← factorThru_arrow _ _ (finset_inf_arrow_factors Finset.univ _ (0 : Fin 2) (by simp)), Category.assoc, kernelSubobject_arrow_comp_assoc, zero_comp, comp_zero] - . erw [factorThru_right, factorThru_eq_zero, factorThru_arrow_assoc, Category.assoc, + · erw [factorThru_right, factorThru_eq_zero, factorThru_arrow_assoc, Category.assoc, ← X.δ_comp_δ (Fin.zero_le (0 : Fin (n + 3))), ← factorThru_arrow _ _ (finset_inf_arrow_factors Finset.univ _ (0 : Fin (n + 3)) (by simp)), Category.assoc, kernelSubobject_arrow_comp_assoc, zero_comp, comp_zero] diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 66cdc8c7bdea8..f027fe1200ba9 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -408,9 +408,9 @@ instance : EssSurj skeletalFunctor.{v} where rw [← hf.le_iff_le] show f (f.symm i) ≤ f (f.symm j) simpa only [OrderIso.apply_symm_apply] - . ext1 ⟨i⟩ + · ext1 ⟨i⟩ exact congr_arg ULift.up (f.symm_apply_apply i) - . ext1 i + · ext1 i exact f.apply_symm_apply i⟩⟩ noncomputable instance isEquivalence : IsEquivalence skeletalFunctor.{v} := @@ -525,7 +525,7 @@ instance {n : ℕ} {i : Fin (n + 1)} : Epi (σ i) := by intro b simp only [σ, mkHom, Hom.toOrderHom_mk, OrderHom.coe_fun_mk] by_cases b ≤ i - . use b + · use b rw [Fin.predAbove_below i b (by simpa only [Fin.coe_eq_castSucc] using h)] simp only [len_mk, Fin.coe_eq_castSucc, Fin.castPred_castSucc] · use b.succ diff --git a/Mathlib/AlgebraicTopology/TopologicalSimplex.lean b/Mathlib/AlgebraicTopology/TopologicalSimplex.lean index 1f775f3b88841..acb9a82a3ec08 100644 --- a/Mathlib/AlgebraicTopology/TopologicalSimplex.lean +++ b/Mathlib/AlgebraicTopology/TopologicalSimplex.lean @@ -60,7 +60,7 @@ def toTopMap {x y : SimplexCategory} (f : x ⟶ y) : x.toTopObj → y.toTopObj : dsimp [toTopObj] at hg convert hg · simp [Finset.eq_univ_iff_forall] - . intro i _ j _ h + · intro i _ j _ h rw [Function.onFun, disjoint_iff_inf_le] intro e he simp only [Finset.bot_eq_empty, Finset.not_mem_empty] @@ -107,10 +107,10 @@ def toTop : SimplexCategory ⥤ TopCat where rw [CategoryTheory.comp_apply, ContinuousMap.coe_mk, ContinuousMap.coe_mk, ContinuousMap.coe_mk] simp only [coe_toTopMap] erw [← Finset.sum_biUnion] - . apply Finset.sum_congr - . exact Finset.ext (fun j => ⟨fun hj => by simpa using hj, fun hj => by simpa using hj⟩) - . tauto - . intro j _ k _ h + · apply Finset.sum_congr + · exact Finset.ext (fun j => ⟨fun hj => by simpa using hj, fun hj => by simpa using hj⟩) + · tauto + · intro j _ k _ h rw [Function.onFun, disjoint_iff_inf_le] intro e he simp only [Finset.bot_eq_empty, Finset.not_mem_empty] diff --git a/Mathlib/Analysis/Convex/Segment.lean b/Mathlib/Analysis/Convex/Segment.lean index f39c97a087895..72290e5fdfe7f 100644 --- a/Mathlib/Analysis/Convex/Segment.lean +++ b/Mathlib/Analysis/Convex/Segment.lean @@ -495,8 +495,7 @@ theorem Icc_subset_segment : Icc x y ⊆ [x -[𝕜] y] := by rw [← sub_pos] at h refine' ⟨(y - z) / (y - x), (z - x) / (y - x), div_nonneg hyz h.le, div_nonneg hxz h.le, _, _⟩ · rw [← add_div, sub_add_sub_cancel, div_self h.ne'] - · - rw [smul_eq_mul, smul_eq_mul, ← mul_div_right_comm, ← mul_div_right_comm, ← add_div, + · rw [smul_eq_mul, smul_eq_mul, ← mul_div_right_comm, ← mul_div_right_comm, ← add_div, div_eq_iff h.ne', add_comm, sub_mul, sub_mul, mul_comm x, sub_add_sub_cancel, mul_sub] #align Icc_subset_segment Icc_subset_segment diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index c6bbcd0dad29c..25c40d9221305 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -65,7 +65,7 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s · exact mul_nonneg hau hav have hw : ∑ i, w i = az * av + bz * au := by trans az * av * bu + (bz * au * bv + au * av) - . simp [Fin.sum_univ_succ, Fin.sum_univ_zero] + · simp [Fin.sum_univ_succ, Fin.sum_univ_zero] rw [← one_mul (au * av), ← habz, add_mul, ← add_assoc, add_add_add_comm, mul_assoc, ← mul_add, mul_assoc, ← mul_add, mul_comm av, ← add_mul, ← mul_add, add_comm bu, add_comm bv, habu, habv, one_mul, mul_one] diff --git a/Mathlib/Analysis/LocallyConvex/Basic.lean b/Mathlib/Analysis/LocallyConvex/Basic.lean index 6298fec0749a8..a23d470d8f6df 100644 --- a/Mathlib/Analysis/LocallyConvex/Basic.lean +++ b/Mathlib/Analysis/LocallyConvex/Basic.lean @@ -100,8 +100,7 @@ theorem absorbs_iUnion_finset {ι : Type _} {t : Finset ι} {f : ι → Set E} : Absorbs 𝕜 s (⋃ i ∈ t, f i) ↔ ∀ i ∈ t, Absorbs 𝕜 s (f i) := by classical induction' t using Finset.induction_on with i t _ht hi - · - simp only [Finset.not_mem_empty, Set.iUnion_false, Set.iUnion_empty, absorbs_empty, + · simp only [Finset.not_mem_empty, Set.iUnion_false, Set.iUnion_empty, absorbs_empty, IsEmpty.forall_iff, imp_true_iff] rw [Finset.set_biUnion_insert, absorbs_union, hi] constructor <;> intro h diff --git a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean index 35e3a22942899..dbf3b9d70df72 100644 --- a/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean +++ b/Mathlib/Analysis/LocallyConvex/WithSeminorms.lean @@ -692,9 +692,9 @@ lemma bound_of_continuous_normedSpace (q : Seminorm 𝕜 F) have : 0 < ‖c‖ / ε := by positivity refine ⟨‖c‖ / ε, this, fun x ↦ ?_⟩ by_cases hx : ‖x‖ = 0 - . rw [hx, mul_zero] + · rw [hx, mul_zero] exact le_of_eq (map_eq_zero_of_norm_zero q hq hx) - . refine (normSeminorm 𝕜 F).bound_of_shell q ε_pos hc (fun x hle hlt ↦ ?_) hx + · refine (normSeminorm 𝕜 F).bound_of_shell q ε_pos hc (fun x hle hlt ↦ ?_) hx refine (le_of_lt <| show q x < _ from hε hlt).trans ?_ rwa [← div_le_iff' this, one_div_div] diff --git a/Mathlib/Analysis/Normed/Ring/Seminorm.lean b/Mathlib/Analysis/Normed/Ring/Seminorm.lean index c2ebae282008a..08c0fb1116121 100644 --- a/Mathlib/Analysis/Normed/Ring/Seminorm.lean +++ b/Mathlib/Analysis/Normed/Ring/Seminorm.lean @@ -129,8 +129,7 @@ instance [DecidableEq R] : One (RingSeminorm R) := ⟨{ (1 : AddGroupSeminorm R) with mul_le' := fun x y => by by_cases h : x * y = 0 - · - refine' (if_pos h).trans_le (mul_nonneg _ _) <;> + · refine' (if_pos h).trans_le (mul_nonneg _ _) <;> · change _ ≤ ite _ _ _ split_ifs exacts [le_rfl, zero_le_one] diff --git a/Mathlib/Analysis/NormedSpace/lpSpace.lean b/Mathlib/Analysis/NormedSpace/lpSpace.lean index fe3361d3ccf97..a775b21e424d4 100644 --- a/Mathlib/Analysis/NormedSpace/lpSpace.lean +++ b/Mathlib/Analysis/NormedSpace/lpSpace.lean @@ -1236,8 +1236,8 @@ lemma LipschitzWith.uniformly_bounded [PseudoMetricSpace α] (g : α → ι → _ ≤ |g a i - g a₀ i| + |g a₀ i| := abs_add _ _ _ ≤ ↑K * dist a a₀ + M := by gcongr - . exact lipschitzWith_iff_dist_le_mul.1 (hg i) a a₀ - . exact hM ⟨i, rfl⟩ + · exact lipschitzWith_iff_dist_le_mul.1 (hg i) a a₀ + · exact hM ⟨i, rfl⟩ theorem LipschitzOnWith.coordinate [PseudoMetricSpace α] (f : α → ℓ^∞(ι)) (s : Set α) (K : ℝ≥0) : LipschitzOnWith K f s ↔ ∀ i : ι, LipschitzOnWith K (fun a : α ↦ f a i) s := by diff --git a/Mathlib/Analysis/Seminorm.lean b/Mathlib/Analysis/Seminorm.lean index a2b3edfb3f9dd..69149ddc26f4b 100644 --- a/Mathlib/Analysis/Seminorm.lean +++ b/Mathlib/Analysis/Seminorm.lean @@ -408,8 +408,8 @@ theorem exists_apply_eq_finset_sup (p : ι → Seminorm 𝕜 E) {s : Finset ι} theorem zero_or_exists_apply_eq_finset_sup (p : ι → Seminorm 𝕜 E) (s : Finset ι) (x : E) : s.sup p x = 0 ∨ ∃ i ∈ s, s.sup p x = p i x := by rcases Finset.eq_empty_or_nonempty s with (rfl|hs) - . left; rfl - . right; exact exists_apply_eq_finset_sup p hs x + · left; rfl + · right; exact exists_apply_eq_finset_sup p hs x theorem finset_sup_smul (p : ι → Seminorm 𝕜 E) (s : Finset ι) (C : ℝ≥0) : s.sup (C • p) = C • s.sup p := by @@ -1226,17 +1226,17 @@ lemma rescale_to_shell_zpow (p : Seminorm 𝕜 E) {c : 𝕜} (hc : 1 < ‖c‖) have cpos : 0 < ‖c‖ := lt_trans (zero_lt_one : (0 :ℝ) < 1) hc have cnpos : 0 < ‖c^(n+1)‖ := by rw [norm_zpow]; exact lt_trans xεpos hn.2 refine ⟨-(n+1), ?_, ?_, ?_, ?_⟩ - . show c ^ (-(n + 1)) ≠ 0; exact zpow_ne_zero _ (norm_pos_iff.1 cpos) - . show p ((c ^ (-(n + 1))) • x) < ε + · show c ^ (-(n + 1)) ≠ 0; exact zpow_ne_zero _ (norm_pos_iff.1 cpos) + · show p ((c ^ (-(n + 1))) • x) < ε rw [map_smul_eq_mul, zpow_neg, norm_inv, ← div_eq_inv_mul, div_lt_iff cnpos, mul_comm, norm_zpow] exact (div_lt_iff εpos).1 (hn.2) - . show ε / ‖c‖ ≤ p (c ^ (-(n + 1)) • x) + · show ε / ‖c‖ ≤ p (c ^ (-(n + 1)) • x) rw [zpow_neg, div_le_iff cpos, map_smul_eq_mul, norm_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), zpow_one, mul_inv_rev, mul_comm, ← mul_assoc, ← mul_assoc, mul_inv_cancel (ne_of_gt cpos), one_mul, ← div_eq_inv_mul, le_div_iff (zpow_pos_of_pos cpos _), mul_comm] exact (le_div_iff εpos).1 hn.1 - . show ‖(c ^ (-(n + 1)))‖⁻¹ ≤ ε⁻¹ * ‖c‖ * p x + · show ‖(c ^ (-(n + 1)))‖⁻¹ ≤ ε⁻¹ * ‖c‖ * p x have : ε⁻¹ * ‖c‖ * p x = ε⁻¹ * p x * ‖c‖ := by ring rw [zpow_neg, norm_inv, inv_inv, norm_zpow, zpow_add₀ (ne_of_gt cpos), zpow_one, this, ← div_eq_inv_mul] diff --git a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean index 1814b8a618f4d..fb33c0fd4293b 100644 --- a/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean +++ b/Mathlib/Analysis/SpecialFunctions/Complex/Arg.lean @@ -73,8 +73,7 @@ set_option linter.uppercaseLean3 false in theorem abs_eq_one_iff (z : ℂ) : abs z = 1 ↔ ∃ θ : ℝ, exp (θ * I) = z := by refine' ⟨fun hz => ⟨arg z, _⟩, _⟩ - · - calc + · calc exp (arg z * I) = abs z * exp (arg z * I) := by rw [hz, ofReal_one, one_mul] _ = z := abs_mul_exp_arg_mul_I z @@ -385,8 +384,7 @@ theorem arg_neg_eq_arg_add_pi_of_im_neg {x : ℂ} (hi : x.im < 0) : arg (-x) = a theorem arg_neg_eq_arg_sub_pi_iff {x : ℂ} : arg (-x) = arg x - π ↔ 0 < x.im ∨ x.im = 0 ∧ x.re < 0 := by rcases lt_trichotomy x.im 0 with (hi | hi | hi) - · - simp [hi, hi.ne, hi.not_lt, arg_neg_eq_arg_add_pi_of_im_neg, sub_eq_add_neg, ← + · simp [hi, hi.ne, hi.not_lt, arg_neg_eq_arg_add_pi_of_im_neg, sub_eq_add_neg, ← add_eq_zero_iff_eq_neg, Real.pi_ne_zero] · rw [(ext rfl hi : x = x.re)] rcases lt_trichotomy x.re 0 with (hr | hr | hr) @@ -409,8 +407,7 @@ theorem arg_neg_eq_arg_add_pi_iff {x : ℂ} : · simp [hr, hi, Real.pi_ne_zero.symm] · rw [arg_ofReal_of_nonneg hr.le, ← ofReal_neg, arg_ofReal_of_neg (Left.neg_neg_iff.2 hr)] simp [hr] - · - simp [hi, hi.ne.symm, hi.not_lt, arg_neg_eq_arg_sub_pi_of_im_pos, sub_eq_add_neg, ← + · simp [hi, hi.ne.symm, hi.not_lt, arg_neg_eq_arg_sub_pi_of_im_pos, sub_eq_add_neg, ← add_eq_zero_iff_neg_eq, Real.pi_ne_zero] #align complex.arg_neg_eq_arg_add_pi_iff Complex.arg_neg_eq_arg_add_pi_iff diff --git a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean index cb091e3d8c29d..d3ca012fd82c8 100644 --- a/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Adjunction/FullyFaithful.lean @@ -216,8 +216,8 @@ def Adjunction.restrictFullyFaithful (iC : C ⥤ C') (iD : D ⥤ D') {L' : C' homEquiv_naturality_right := fun {X Y' Y} f g => by apply iC.map_injective suffices : R'.map (iD.map g) ≫ comm2.hom.app Y = comm2.hom.app Y' ≫ iC.map (R.map g) - . simp [Trans.trans, this] - . apply comm2.hom.naturality g } + · simp [Trans.trans, this] + · apply comm2.hom.naturality g } #align category_theory.adjunction.restrict_fully_faithful CategoryTheory.Adjunction.restrictFullyFaithful end CategoryTheory diff --git a/Mathlib/CategoryTheory/Category/Cat/Limit.lean b/Mathlib/CategoryTheory/Category/Cat/Limit.lean index 81340e40df1bf..8cccb2361ce88 100644 --- a/Mathlib/CategoryTheory/Category/Cat/Limit.lean +++ b/Mathlib/CategoryTheory/Category/Cat/Limit.lean @@ -145,11 +145,11 @@ def limitConeIsLimit (F : J ⥤ Cat.{v, v}) : IsLimit (limitCone F) where uniq s m w := by symm refine' CategoryTheory.Functor.ext _ _ - . intro X + · intro X apply Types.limit_ext.{v, v} intro j simp [Types.Limit.lift_π_apply', ← w j] - . intro X Y f + · intro X Y f dsimp simp [fun j => Functor.congr_hom (w j).symm f] set_option linter.uppercaseLean3 false in diff --git a/Mathlib/CategoryTheory/Category/PartialFun.lean b/Mathlib/CategoryTheory/Category/PartialFun.lean index 1dad8bc69c50b..004bec6ea7e57 100644 --- a/Mathlib/CategoryTheory/Category/PartialFun.lean +++ b/Mathlib/CategoryTheory/Category/PartialFun.lean @@ -108,9 +108,9 @@ def pointedToPartialFun : Pointed.{u} ⥤ PartialFun where apply PFun.ext _ rintro ⟨a, ha⟩ ⟨c, hc⟩ constructor - . rintro ⟨h₁, h₂⟩ + · rintro ⟨h₁, h₂⟩ exact ⟨⟨fun h₀ => h₁ ((congr_arg g.toFun h₀).trans g.map_point), h₁⟩, h₂⟩ - . rintro ⟨_, _, _⟩ + · rintro ⟨_, _, _⟩ exact ⟨_, rfl⟩ #align Pointed_to_PartialFun pointedToPartialFun @@ -152,12 +152,12 @@ noncomputable def partialFunEquivPointed : PartialFun.{u} ≌ Pointed := refine' (Part.mem_bind_iff.trans _).trans PFun.mem_toSubtype_iff.symm obtain ⟨b | b, hb⟩ := b · exact (hb rfl).elim - . dsimp [Part.toOption] + · dsimp [Part.toOption] simp_rw [Part.mem_some_iff, Subtype.mk_eq_mk] constructor - . rintro ⟨_, ⟨h₁, h₂⟩, h₃⟩ + · rintro ⟨_, ⟨h₁, h₂⟩, h₃⟩ rw [h₃, ← h₂, dif_pos h₁] - . intro h + · intro h split_ifs at h with ha rw [some_inj] at h refine' ⟨b, ⟨ha, h.symm⟩, rfl⟩) @@ -173,9 +173,9 @@ noncomputable def partialFunEquivPointed : PartialFun.{u} ≌ Pointed := right_inv := fun a => by dsimp split_ifs with h - . rw [h] + · rw [h] rfl - . rfl} rfl) + · rfl} rfl) fun {X Y} f => Pointed.Hom.ext _ _ <| funext fun a => @@ -187,8 +187,8 @@ noncomputable def partialFunEquivPointed : PartialFun.{u} ≌ Pointed := -- conflicting `Decidable` instances rw [Option.elim'_eq_elim, @Part.elim_toOption _ _ _ (Classical.propDecidable _)] split_ifs with h - . rfl - . exact Eq.symm (of_not_not h))) + · rfl + · exact Eq.symm (of_not_not h))) #align PartialFun_equiv_Pointed partialFunEquivPointed /-- Forgetting that maps are total and making them total again by adding a point is the same as just diff --git a/Mathlib/CategoryTheory/Category/Pointed.lean b/Mathlib/CategoryTheory/Category/Pointed.lean index ce72c2f7be2ed..005ad1cf7109f 100644 --- a/Mathlib/CategoryTheory/Category/Pointed.lean +++ b/Mathlib/CategoryTheory/Category/Pointed.lean @@ -144,8 +144,8 @@ def typeToPointedForgetAdjunction : typeToPointed ⊣ forget Pointed := apply Pointed.Hom.ext funext x cases x - . exact f.map_point.symm - . rfl + · exact f.map_point.symm + · rfl right_inv := fun f => funext fun _ => rfl } homEquiv_naturality_left_symm := fun f g => by apply Pointed.Hom.ext diff --git a/Mathlib/CategoryTheory/Extensive.lean b/Mathlib/CategoryTheory/Extensive.lean index 571041ccc90eb..f57052ff5b31c 100644 --- a/Mathlib/CategoryTheory/Extensive.lean +++ b/Mathlib/CategoryTheory/Extensive.lean @@ -442,8 +442,8 @@ noncomputable def finitaryExtensiveTopCatAux (Z : TopCat.{u}) ext ⟨⟨x, ⟨⟩⟩, (hx : f x = Sum.inl PUnit.unit)⟩ change dite _ _ _ = _ split_ifs with h - . rfl - . cases (h hx) -- Porting note : in Lean3 it is `rfl` + · rfl + · cases (h hx) -- Porting note : in Lean3 it is `rfl` · intro s ext ⟨⟨x, ⟨⟩⟩, hx⟩ change dite _ _ _ = _ diff --git a/Mathlib/CategoryTheory/Filtered.lean b/Mathlib/CategoryTheory/Filtered.lean index 02e51be43296a..035724588ed9d 100644 --- a/Mathlib/CategoryTheory/Filtered.lean +++ b/Mathlib/CategoryTheory/Filtered.lean @@ -240,8 +240,8 @@ variable [IsFiltered C] theorem sup_objs_exists (O : Finset C) : ∃ S : C, ∀ {X}, X ∈ O → _root_.Nonempty (X ⟶ S) := by classical induction' O using Finset.induction with X O' nm h - . exact ⟨Classical.choice IsFiltered.Nonempty, by intro; simp⟩ - . obtain ⟨S', w'⟩ := h + · exact ⟨Classical.choice IsFiltered.Nonempty, by intro; simp⟩ + · obtain ⟨S', w'⟩ := h use max X S' rintro Y mY obtain rfl | h := eq_or_ne Y X @@ -278,9 +278,9 @@ theorem sup_exists : · rw [@w' _ _ mX mY f'] simp only [Finset.mem_insert, PSigma.mk.injEq, heq_eq_eq, true_and] at mf' rcases mf' with mf' | mf' - . exfalso + · exfalso exact hf mf'.symm - . exact mf' + · exact mf' · rw [@w' _ _ mX' mY' f' _] apply Finset.mem_of_mem_insert_of_ne mf' contrapose! h @@ -684,8 +684,8 @@ variable [IsCofiltered C] theorem inf_objs_exists (O : Finset C) : ∃ S : C, ∀ {X}, X ∈ O → _root_.Nonempty (S ⟶ X) := by classical induction' O using Finset.induction with X O' nm h - . exact ⟨Classical.choice IsCofiltered.Nonempty, by intro; simp⟩ - . obtain ⟨S', w'⟩ := h + · exact ⟨Classical.choice IsCofiltered.Nonempty, by intro; simp⟩ + · obtain ⟨S', w'⟩ := h use min X S' rintro Y mY obtain rfl | h := eq_or_ne Y X @@ -722,9 +722,9 @@ theorem inf_exists : · rw [@w' _ _ mX mY f'] simp only [Finset.mem_insert, PSigma.mk.injEq, heq_eq_eq, true_and] at mf' rcases mf' with mf' | mf' - . exfalso + · exfalso exact hf mf'.symm - . exact mf' + · exact mf' · rw [@w' _ _ mX' mY' f' _] apply Finset.mem_of_mem_insert_of_ne mf' contrapose! h diff --git a/Mathlib/CategoryTheory/Functor/LeftDerived.lean b/Mathlib/CategoryTheory/Functor/LeftDerived.lean index 36bc6ffd516d8..0ee9051b0c217 100644 --- a/Mathlib/CategoryTheory/Functor/LeftDerived.lean +++ b/Mathlib/CategoryTheory/Functor/LeftDerived.lean @@ -123,8 +123,8 @@ theorem Functor.leftDerived_map_eq (F : C ⥤ D) [F.Additive] (n : ℕ) {X Y : C apply HomotopyCategory.eq_of_homotopy apply Functor.mapHomotopy apply ProjectiveResolution.liftHomotopy f - . simp - . simp [w] + · simp + · simp [w] #align category_theory.functor.left_derived_map_eq CategoryTheory.Functor.leftDerived_map_eq /-- The natural transformation between left-derived functors induced by a natural transformation. -/ diff --git a/Mathlib/CategoryTheory/Idempotents/Biproducts.lean b/Mathlib/CategoryTheory/Idempotents/Biproducts.lean index f32dfe4853f5a..f5b43a7263222 100644 --- a/Mathlib/CategoryTheory/Idempotents/Biproducts.lean +++ b/Mathlib/CategoryTheory/Idempotents/Biproducts.lean @@ -64,10 +64,10 @@ def bicone [HasFiniteBiproducts C] {J : Type} [Finite J] (F : J → Karoubi C) : comm := by simp only [biproduct.ι_map, assoc, idem_assoc] } ι_π j j' := by split_ifs with h - . subst h + · subst h simp only [biproduct.ι_map, biproduct.bicone_π, biproduct.map_π, eqToHom_refl, id_eq, hom_ext_iff, comp_f, assoc, bicone_ι_π_self_assoc, idem] - . dsimp + · dsimp simp only [hom_ext_iff, biproduct.ι_map, biproduct.map_π, comp_f, assoc, ne_eq, biproduct.ι_π_ne_assoc _ h, comp_zero, zero_comp] #align category_theory.idempotents.karoubi.biproducts.bicone CategoryTheory.Idempotents.Karoubi.Biproducts.bicone @@ -84,10 +84,10 @@ theorem karoubi_hasFiniteBiproducts [HasFiniteBiproducts C] : HasFiniteBiproduct biproduct.bicone_π, biproduct.map_π, Biproducts.bicone_ι_f, biproduct.ι_map, assoc, idem_assoc, id_eq, Biproducts.bicone_pt_p, comp_sum] rw [Finset.sum_eq_single j] - . simp only [bicone_ι_π_self_assoc] - . intro b _ hb + · simp only [bicone_ι_π_self_assoc] + · intro b _ hb simp only [biproduct.ι_π_ne_assoc _ hb.symm, zero_comp] - . intro hj + · intro hj simp only [Finset.mem_univ, not_true] at hj } } #align category_theory.idempotents.karoubi.karoubi_has_finite_biproducts CategoryTheory.Idempotents.Karoubi.karoubi_hasFiniteBiproducts @@ -132,13 +132,13 @@ def decomposition (P : Karoubi C) : P ⊞ P.complement ≅ (toKaroubi _).obj P.X inv := biprod.lift P.decompId_p P.complement.decompId_p hom_inv_id := by apply biprod.hom_ext' - . rw [biprod.inl_desc_assoc, comp_id, biprod.lift_eq, comp_add, ← decompId_assoc, + · rw [biprod.inl_desc_assoc, comp_id, biprod.lift_eq, comp_add, ← decompId_assoc, add_right_eq_self, ← assoc] refine' (_ =≫ _).trans zero_comp ext simp only [comp_f, toKaroubi_obj_X, decompId_i_f, decompId_p_f, complement_p, comp_sub, comp_id, idem, sub_self, instAddCommGroupHom_zero] - . rw [biprod.inr_desc_assoc, comp_id, biprod.lift_eq, comp_add, ← decompId_assoc, + · rw [biprod.inr_desc_assoc, comp_id, biprod.lift_eq, comp_add, ← decompId_assoc, add_left_eq_self, ← assoc] refine' (_ =≫ _).trans zero_comp ext diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean index 64a7aa516ed38..02abed2353210 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorCategories.lean @@ -82,11 +82,11 @@ instance functor_category_isIdempotentComplete [IsIdempotentComplete C] : naturality := fun j j' φ => equalizer.hom_ext (by simp) } use Y, i, e constructor - . ext j + · ext j apply equalizer.hom_ext dsimp rw [assoc, equalizer.lift_ι, ← equalizer.condition, id_comp, comp_id] - . ext j + · ext j simp namespace KaroubiFunctorCategoryEmbedding diff --git a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean index fdf61ecbc1233..d455feb607472 100644 --- a/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean +++ b/Mathlib/CategoryTheory/Idempotents/FunctorExtension.lean @@ -121,8 +121,8 @@ theorem functorExtension₁_comp_whiskeringLeft_toKaroubi : refine' Functor.ext _ _ · intro X ext - . simp - . simp + · simp + · simp · aesop_cat · intro F G φ aesop_cat @@ -194,8 +194,8 @@ theorem functorExtension₁_comp (F : C ⥤ Karoubi D) (G : D ⥤ Karoubi E) : (functorExtension₁ C E).obj (F ⋙ (functorExtension₁ D E).obj G) = (functorExtension₁ C D).obj F ⋙ (functorExtension₁ D E).obj G := by refine' Functor.ext _ _ - . aesop_cat - . intro X Y f + · aesop_cat + · intro X Y f ext simp only [eqToHom_refl, id_comp, comp_id] rfl diff --git a/Mathlib/CategoryTheory/Limits/Bicones.lean b/Mathlib/CategoryTheory/Limits/Bicones.lean index 9dcac4ffe1068..21650e207bbfd 100644 --- a/Mathlib/CategoryTheory/Limits/Bicones.lean +++ b/Mathlib/CategoryTheory/Limits/Bicones.lean @@ -84,13 +84,13 @@ instance biconeCategoryStruct : CategoryStruct (Bicone J) id j := Bicone.casesOn j BiconeHom.left_id BiconeHom.right_id fun k => BiconeHom.diagram (𝟙 k) comp f g := by rcases f with (_ | _ | _ | _ | f) - . exact g - . exact g - . cases g + · exact g + · exact g + · cases g apply BiconeHom.left - . cases g + · cases g apply BiconeHom.right - . rcases g with (_|_|_|_|g) + · rcases g with (_|_|_|_|g) exact BiconeHom.diagram (f ≫ g) #align category_theory.bicone_category_struct CategoryTheory.biconeCategoryStruct @@ -115,51 +115,51 @@ def biconeMk {C : Type u₁} [Category.{v₁} C] {F : J ⥤ C} (c₁ c₂ : Cone obj X := Bicone.casesOn X c₁.pt c₂.pt fun j => F.obj j map f := by rcases f with (_|_|_|_|f) - . exact 𝟙 _ - . exact 𝟙 _ - . exact c₁.π.app _ - . exact c₂.π.app _ - . exact F.map f + · exact 𝟙 _ + · exact 𝟙 _ + · exact c₁.π.app _ + · exact c₂.π.app _ + · exact F.map f map_id X := by cases X <;> simp map_comp f g := by rcases f with (_|_|_|_|_) - . exact (Category.id_comp _).symm - . exact (Category.id_comp _).symm - . cases g + · exact (Category.id_comp _).symm + · exact (Category.id_comp _).symm + · cases g exact (Category.id_comp _).symm.trans (c₁.π.naturality _) - . cases g + · cases g exact (Category.id_comp _).symm.trans (c₂.π.naturality _) - . cases g + · cases g apply F.map_comp #align category_theory.bicone_mk CategoryTheory.biconeMk instance finBiconeHom [FinCategory J] (j k : Bicone J) : Fintype (j ⟶ k) := by cases j <;> cases k - . exact + · exact { elems := {BiconeHom.left_id} complete := fun f => by cases f; simp } - . exact + · exact { elems := ∅ complete := fun f => by cases f } - . exact + · exact { elems := {BiconeHom.left _} complete := fun f => by cases f; simp } - . exact + · exact { elems := ∅ complete := fun f => by cases f } - . exact + · exact { elems := {BiconeHom.right_id} complete := fun f => by cases f; simp } - . exact + · exact { elems := {BiconeHom.right _} complete := fun f => by cases f; simp } - . exact + · exact { elems := ∅ complete := fun f => by cases f } - . exact + · exact { elems := ∅ complete := fun f => by cases f } - . exact + · exact { elems := Finset.image BiconeHom.diagram Fintype.elems complete := fun f => by rcases f with (_|_|_|_|f) diff --git a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean index 95b5222e8b8b1..dc9ff416ffff3 100644 --- a/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean +++ b/Mathlib/CategoryTheory/Limits/Constructions/Filtered.lean @@ -70,9 +70,9 @@ def liftToFinsetColimitCocone [HasFiniteCoproducts C] [HasFilteredColimitsOfSize apply colimit.hom_ext rintro ⟨⟨j, hj⟩⟩ convert h j using 1 - . simp [← colimit.w (liftToFinset F) ⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩] + · simp [← colimit.w (liftToFinset F) ⟨⟨Finset.singleton_subset_iff.2 hj⟩⟩] rfl - . aesop_cat } + · aesop_cat } #align category_theory.limits.coproducts_from_finite_filtered.lift_to_finset_colimit_cocone CategoryTheory.Limits.CoproductsFromFiniteFiltered.liftToFinsetColimitCocone end CoproductsFromFiniteFiltered diff --git a/Mathlib/CategoryTheory/Limits/Opposites.lean b/Mathlib/CategoryTheory/Limits/Opposites.lean index df22af11c7248..3c899b2925ba9 100644 --- a/Mathlib/CategoryTheory/Limits/Opposites.lean +++ b/Mathlib/CategoryTheory/Limits/Opposites.lean @@ -575,10 +575,10 @@ in the opposite category is a limit cone. -/ def isColimitEquivIsLimitOp {X Y Z : C} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : IsColimit c ≃ IsLimit c.op := by apply equivOfSubsingletonOfSubsingleton - . intro h + · intro h exact (IsLimit.postcomposeHomEquiv _ _).invFun ((IsLimit.whiskerEquivalenceEquiv walkingSpanOpEquiv.symm).toFun (isLimitCoconeOp _ h)) - . intro h + · intro h exact (IsColimit.equivIsoColimit c.opUnop).toFun (isColimitConeUnop _ ((IsLimit.postcomposeHomEquiv _ _).invFun ((IsLimit.whiskerEquivalenceEquiv _).toFun h))) @@ -589,10 +589,10 @@ in `C` is a limit cone. -/ def isColimitEquivIsLimitUnop {X Y Z : Cᵒᵖ} {f : X ⟶ Y} {g : X ⟶ Z} (c : PushoutCocone f g) : IsColimit c ≃ IsLimit c.unop := by apply equivOfSubsingletonOfSubsingleton - . intro h + · intro h exact isLimitCoconeUnop _ ((IsColimit.precomposeHomEquiv _ _).invFun ((IsColimit.whiskerEquivalenceEquiv _).toFun h)) - . intro h + · intro h exact (IsColimit.equivIsoColimit c.unopOp).toFun ((IsColimit.precomposeHomEquiv _ _).invFun ((IsColimit.whiskerEquivalenceEquiv walkingCospanOpEquiv.symm).toFun (isColimitConeOp _ h))) diff --git a/Mathlib/CategoryTheory/Limits/Over.lean b/Mathlib/CategoryTheory/Limits/Over.lean index 6bc01bc7ee9d5..3b6ccfee388b8 100644 --- a/Mathlib/CategoryTheory/Limits/Over.lean +++ b/Mathlib/CategoryTheory/Limits/Over.lean @@ -107,9 +107,9 @@ def mapPullbackAdj {A B : C} (f : A ⟶ B) : Over.map f ⊣ pullback f := -- TODO: It would be nice to replace the next two lines with just `ext`. apply OverMorphism.ext apply pullback.hom_ext - . dsimp + · dsimp simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app] - . dsimp + · dsimp simp only [limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, ← Over.w Y ] rfl } } #align category_theory.over.map_pullback_adj CategoryTheory.Over.mapPullbackAdj diff --git a/Mathlib/CategoryTheory/Limits/Shapes/CommSq.lean b/Mathlib/CategoryTheory/Limits/Shapes/CommSq.lean index fc382fc00acd6..8f92067da33e8 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/CommSq.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/CommSq.lean @@ -1005,7 +1005,7 @@ theorem Functor.map_isPullback [PreservesLimit (cospan h i) F] (s : IsPullback f (isLimitOfPreserves F s.isLimit)) · rfl · simp - . simp + · simp #align category_theory.functor.map_is_pullback CategoryTheory.Functor.map_isPullback theorem Functor.map_isPushout [PreservesColimit (span f g) F] (s : IsPushout f g h i) : diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean b/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean index 3eaf2748de519..54cdc99fb1e3b 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Diagonal.lean @@ -137,16 +137,16 @@ def pullbackDiagonalMapIso : hom := pullback.lift (pullback.snd ≫ pullback.fst) (pullback.snd ≫ pullback.snd) (by ext - . simp [Category.assoc, pullback_diagonal_map_snd_fst_fst, pullback_diagonal_map_snd_snd_fst] - . simp [Category.assoc, pullback.condition, pullback.condition_assoc]) + · simp [Category.assoc, pullback_diagonal_map_snd_fst_fst, pullback_diagonal_map_snd_snd_fst] + · simp [Category.assoc, pullback.condition, pullback.condition_assoc]) inv := pullback.lift (pullback.fst ≫ i₁ ≫ pullback.fst) (pullback.map _ _ _ _ (𝟙 _) (𝟙 _) pullback.snd (Category.id_comp _).symm (Category.id_comp _).symm) (by ext - . simp only [Category.assoc, diagonal_fst, Category.comp_id, limit.lift_π, + · simp only [Category.assoc, diagonal_fst, Category.comp_id, limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, limit.lift_π_assoc, cospan_left] - . simp only [condition_assoc, Category.assoc, diagonal_snd, Category.comp_id, + · simp only [condition_assoc, Category.assoc, diagonal_snd, Category.comp_id, limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, limit.lift_π_assoc, cospan_right]) #align category_theory.limits.pullback_diagonal_map_iso CategoryTheory.Limits.pullbackDiagonalMapIso @@ -227,16 +227,16 @@ def pullbackDiagonalMapIdIso : pullback f g := by refine' _ ≪≫ pullbackDiagonalMapIso i (𝟙 _) (f ≫ inv pullback.fst) (g ≫ inv pullback.fst) ≪≫ _ - . refine' @asIso _ _ _ _ (pullback.map _ _ _ _ (𝟙 T) ((pullback.congrHom _ _).hom) (𝟙 _) _ _) ?_ - . rw [← Category.comp_id pullback.snd, ← condition, Category.assoc, IsIso.inv_hom_id_assoc] - . rw [← Category.comp_id pullback.snd, ← condition, Category.assoc, IsIso.inv_hom_id_assoc] + · refine' @asIso _ _ _ _ (pullback.map _ _ _ _ (𝟙 T) ((pullback.congrHom _ _).hom) (𝟙 _) _ _) ?_ + · rw [← Category.comp_id pullback.snd, ← condition, Category.assoc, IsIso.inv_hom_id_assoc] + · rw [← Category.comp_id pullback.snd, ← condition, Category.assoc, IsIso.inv_hom_id_assoc] · rw [Category.comp_id, Category.id_comp] - . ext <;> simp - . infer_instance - . refine' @asIso _ _ _ _ (pullback.map _ _ _ _ (𝟙 _) (𝟙 _) pullback.fst _ _) ?_ - . rw [Category.assoc, IsIso.inv_hom_id, Category.comp_id, Category.id_comp] - . rw [Category.assoc, IsIso.inv_hom_id, Category.comp_id, Category.id_comp] - . infer_instance + · ext <;> simp + · infer_instance + · refine' @asIso _ _ _ _ (pullback.map _ _ _ _ (𝟙 _) (𝟙 _) pullback.fst _ _) ?_ + · rw [Category.assoc, IsIso.inv_hom_id, Category.comp_id, Category.id_comp] + · rw [Category.assoc, IsIso.inv_hom_id, Category.comp_id, Category.id_comp] + · infer_instance #align category_theory.limits.pullback_diagonal_map_id_iso CategoryTheory.Limits.pullbackDiagonalMapIdIso @[reassoc (attr := simp)] @@ -407,24 +407,24 @@ def pullbackFstFstIso {X Y S X' Y' S' : C} (f : X ⟶ S) (g : Y ⟶ S) (f' : X' -- We could use `ext` here to immediately descend to the leaf goals, -- but it only obscures the structure. apply pullback.hom_ext - . apply pullback.hom_ext - . apply pullback.hom_ext - . simp only [Category.assoc, lift_fst, lift_fst_assoc, Category.id_comp] + · apply pullback.hom_ext + · apply pullback.hom_ext + · simp only [Category.assoc, lift_fst, lift_fst_assoc, Category.id_comp] rw [condition] - . simp [Category.assoc, lift_snd] + · simp [Category.assoc, lift_snd] rw [condition_assoc, condition] - . simp only [Category.assoc, lift_fst_assoc, lift_snd, lift_fst, Category.id_comp] - . apply pullback.hom_ext - . apply pullback.hom_ext - . simp only [Category.assoc, lift_snd_assoc, lift_fst_assoc, lift_fst, Category.id_comp] + · simp only [Category.assoc, lift_fst_assoc, lift_snd, lift_fst, Category.id_comp] + · apply pullback.hom_ext + · apply pullback.hom_ext + · simp only [Category.assoc, lift_snd_assoc, lift_fst_assoc, lift_fst, Category.id_comp] rw [← condition_assoc, condition] - . simp only [Category.assoc, lift_snd, lift_fst_assoc, lift_snd_assoc, Category.id_comp] + · simp only [Category.assoc, lift_snd, lift_fst_assoc, lift_snd_assoc, Category.id_comp] rw [condition] - . simp only [Category.assoc, lift_snd_assoc, lift_snd, Category.id_comp] + · simp only [Category.assoc, lift_snd_assoc, lift_snd, Category.id_comp] inv_hom_id := by apply pullback.hom_ext - . simp only [Category.assoc, lift_fst, lift_fst_assoc, lift_snd, Category.id_comp] - . simp only [Category.assoc, lift_snd, lift_snd_assoc, Category.id_comp] + · simp only [Category.assoc, lift_fst, lift_fst_assoc, lift_snd, Category.id_comp] + · simp only [Category.assoc, lift_snd, lift_snd_assoc, Category.id_comp] #align category_theory.limits.pullback_fst_fst_iso CategoryTheory.Limits.pullbackFstFstIso theorem pullback_map_eq_pullbackFstFstIso_inv {X Y S X' Y' S' : C} (f : X ⟶ S) (g : Y ⟶ S) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/KernelPair.lean b/Mathlib/CategoryTheory/Limits/Shapes/KernelPair.lean index bf441f4868eac..2b0998311f908 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/KernelPair.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/KernelPair.lean @@ -122,9 +122,9 @@ theorem cancel_right {f₁ : X ⟶ Y} {f₂ : Y ⟶ Z} (comm : a ≫ f₁ = b big_k.isLimit.fac _ WalkingCospan.right, fun m₁ m₂ => _⟩ apply big_k.isLimit.hom_ext refine' (PullbackCone.mk a b _ : PullbackCone (f₁ ≫ f₂) _).equalizer_ext _ _ - . apply reassoc_of% comm - . apply m₁.trans (big_k.isLimit.fac s' WalkingCospan.left).symm - . apply m₂.trans (big_k.isLimit.fac s' WalkingCospan.right).symm⟩ } + · apply reassoc_of% comm + · apply m₁.trans (big_k.isLimit.fac s' WalkingCospan.left).symm + · apply m₂.trans (big_k.isLimit.fac s' WalkingCospan.right).symm⟩ } #align category_theory.is_kernel_pair.cancel_right CategoryTheory.IsKernelPair.cancel_right /-- If `(a,b)` is a kernel pair for `f₁ ≫ f₂` and `f₂` is mono, then `(a,b)` is a kernel pair for @@ -147,11 +147,11 @@ theorem comp_of_mono {f₁ : X ⟶ Y} {f₂ : Y ⟶ Z} [Mono f₂] (small_k : Is refine' PullbackCone.isLimitAux _ (fun s => small_k.lift s.fst s.snd (by rw [← cancel_mono f₂, assoc, s.condition, assoc])) (by simp) (by simp) _ - . intro s m hm + · intro s m hm apply small_k.isLimit.hom_ext apply PullbackCone.equalizer_ext small_k.cone _ _ - . exact (hm WalkingCospan.left).trans (by simp) - . exact (hm WalkingCospan.right).trans (by simp)⟩ } + · exact (hm WalkingCospan.left).trans (by simp) + · exact (hm WalkingCospan.right).trans (by simp)⟩ } #align category_theory.is_kernel_pair.comp_of_mono CategoryTheory.IsKernelPair.comp_of_mono /-- @@ -166,8 +166,8 @@ def toCoequalizer (k : IsKernelPair f a b) [r : RegularEpi f] : IsColimit (Cofor (fun s => Cofork.IsColimit.desc r.isColimit s.π (by rw [← ht, assoc, s.condition, reassoc_of% kt])) (fun s => _) (fun s m w => _) - . apply Cofork.IsColimit.π_desc' r.isColimit - . apply Cofork.IsColimit.hom_ext r.isColimit + · apply Cofork.IsColimit.π_desc' r.isColimit + · apply Cofork.IsColimit.hom_ext r.isColimit exact w.trans (Cofork.IsColimit.π_desc' r.isColimit _ _).symm #align category_theory.is_kernel_pair.to_coequalizer CategoryTheory.IsKernelPair.toCoequalizer @@ -182,17 +182,17 @@ protected theorem pullback {X Y Z A : C} {g : Y ⟶ Z} {a₁ a₂ : A ⟶ Y} (h (fun s => pullback.lift (s.fst ≫ pullback.fst) (h.lift (s.fst ≫ pullback.snd) (s.snd ≫ pullback.snd) _ ) _) (fun s => _) (fun s => _) (fun s m hm => _)⟩⟩ - . simp_rw [Category.assoc, ← pullback.condition, ← Category.assoc, s.condition] - . simp only [assoc, lift_fst_assoc, pullback.condition] - . ext <;> simp - . ext - . simp [s.condition] - . simp - . apply pullback.hom_ext - . simpa using hm WalkingCospan.left =≫ pullback.fst - . apply PullbackCone.IsLimit.hom_ext h.isLimit - . simpa using hm WalkingCospan.left =≫ pullback.snd - . simpa using hm WalkingCospan.right =≫ pullback.snd + · simp_rw [Category.assoc, ← pullback.condition, ← Category.assoc, s.condition] + · simp only [assoc, lift_fst_assoc, pullback.condition] + · ext <;> simp + · ext + · simp [s.condition] + · simp + · apply pullback.hom_ext + · simpa using hm WalkingCospan.left =≫ pullback.fst + · apply PullbackCone.IsLimit.hom_ext h.isLimit + · simpa using hm WalkingCospan.left =≫ pullback.snd + · simpa using hm WalkingCospan.right =≫ pullback.snd #align category_theory.is_kernel_pair.pullback CategoryTheory.IsKernelPair.pullback theorem mono_of_isIso_fst (h : IsKernelPair f a b) [IsIso a] : Mono f := by diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean index 1c37a43ab374c..faf16c4a3eb23 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Types.lean @@ -221,8 +221,8 @@ noncomputable def binaryProductIsoProd : binaryProductFunctor ≅ (prod.functor refine' NatIso.ofComponents (fun X => _) (fun _ => _) · refine' NatIso.ofComponents (fun Y => _) (fun _ => _) · exact ((limit.isLimit _).conePointUniqueUpToIso (binaryProductLimit X Y)).symm - . apply Limits.prod.hom_ext <;> simp <;> rfl - . ext : 2 + · apply Limits.prod.hom_ext <;> simp <;> rfl + · ext : 2 apply Limits.prod.hom_ext <;> simp <;> rfl #align category_theory.limits.types.binary_product_iso_prod CategoryTheory.Limits.Types.binaryProductIsoProd @@ -416,9 +416,9 @@ noncomputable def typeEqualizerOfUnique (t : ∀ y : Y, g y = h y → ∃! x : X refine' ⟨fun i => _, _, _⟩ · apply Classical.choose (t (s.ι i) _) apply congr_fun s.condition i - . funext i + · funext i exact (Classical.choose_spec (t (s.ι i) (congr_fun s.condition i))).1 - . intro m hm + · intro m hm funext i exact (Classical.choose_spec (t (s.ι i) (congr_fun s.condition i))).2 _ (congr_fun hm i) #align category_theory.limits.types.type_equalizer_of_unique CategoryTheory.Limits.Types.typeEqualizerOfUnique diff --git a/Mathlib/CategoryTheory/Limits/Types.lean b/Mathlib/CategoryTheory/Limits/Types.lean index 6586665df3b95..2ae9171b2ecf5 100644 --- a/Mathlib/CategoryTheory/Limits/Types.lean +++ b/Mathlib/CategoryTheory/Limits/Types.lean @@ -427,8 +427,8 @@ theorem rel_of_quot_rel (x y : Σ j, F.obj j) : theorem eqvGen_quot_rel_of_rel (x y : Σ j, F.obj j) : FilteredColimit.Rel.{v, u} F x y → EqvGen (Quot.Rel.{v, u} F) x y := fun ⟨k, f, g, h⟩ => by refine' EqvGen.trans _ ⟨k, F.map f x.2⟩ _ _ _ - . exact (EqvGen.rel _ _ ⟨f, rfl⟩) - . exact (EqvGen.symm _ _ (EqvGen.rel _ _ ⟨g, h⟩)) + · exact (EqvGen.rel _ _ ⟨f, rfl⟩) + · exact (EqvGen.symm _ _ (EqvGen.rel _ _ ⟨g, h⟩)) #align category_theory.limits.types.filtered_colimit.eqv_gen_quot_rel_of_rel CategoryTheory.Limits.Types.FilteredColimit.eqvGen_quot_rel_of_rel --attribute [local elab_without_expected_type] nat_trans.app @@ -457,12 +457,12 @@ noncomputable def isColimitOf (t : Cocone F) (hsurj : ∀ x : t.pt, ∃ i xi, x rw [← colimit.w F f, ← colimit.w F g] change colimit.ι F k (F.map f xi) = colimit.ι F k (F.map g xj) rw [h'] - . show Function.Surjective _ + · show Function.Surjective _ intro x rcases hsurj x with ⟨i, xi, rfl⟩ use colimit.ι F i xi apply Colimit.ι_desc_apply.{v, u} - . intro j + · intro j apply colimit.ι_desc #align category_theory.limits.types.filtered_colimit.is_colimit_of CategoryTheory.Limits.Types.FilteredColimit.isColimitOf diff --git a/Mathlib/CategoryTheory/Localization/Construction.lean b/Mathlib/CategoryTheory/Localization/Construction.lean index 15b75de5cc6e5..f3aee6a582af2 100644 --- a/Mathlib/CategoryTheory/Localization/Construction.lean +++ b/Mathlib/CategoryTheory/Localization/Construction.lean @@ -158,8 +158,8 @@ def liftToPathCategory : Paths (LocQuiver W) ⥤ D := map := by intros X Y rintro (f | ⟨g, hg⟩) - . exact G.map f - . haveI := hG g hg + · exact G.map f + · haveI := hG g hg exact inv (G.map g) } #align category_theory.localization.construction.lift_to_path_category CategoryTheory.Localization.Construction.liftToPathCategory @@ -171,8 +171,8 @@ def lift : W.Localization ⥤ D := rintro ⟨X⟩ ⟨Y⟩ f₁ f₂ r --Porting note: rest of proof was `rcases r with ⟨⟩; tidy` rcases r with (_|_|⟨f,hf⟩|⟨f,hf⟩) - . aesop_cat - . aesop_cat + · aesop_cat + · aesop_cat all_goals dsimp haveI := hG f hf @@ -237,7 +237,7 @@ theorem morphismProperty_is_top (P : MorphismProperty W.Localization) funext X Y f ext constructor - . intro + · intro apply MorphismProperty.top_apply · intro let G : _ ⥤ W.Localization := Quotient.functor _ @@ -249,12 +249,12 @@ theorem morphismProperty_is_top (P : MorphismProperty W.Localization) intros X₁ X₂ p induction' p with X₂ X₃ p g hp · simpa only [Functor.map_id] using hP₁ (𝟙 X₁.obj) - . let p' : X₁ ⟶X₂ := p + · let p' : X₁ ⟶X₂ := p rw [show p'.cons g = p' ≫ Quiver.Hom.toPath g by rfl, G.map_comp] refine' hP₃ _ _ hp _ rcases g with (g | ⟨g, hg⟩) - . apply hP₁ - . apply hP₂ + · apply hP₁ + · apply hP₂ #align category_theory.localization.construction.morphism_property_is_top CategoryTheory.Localization.Construction.morphismProperty_is_top /-- A `MorphismProperty` in `W.Localization` is satisfied by all diff --git a/Mathlib/CategoryTheory/Localization/Predicate.lean b/Mathlib/CategoryTheory/Localization/Predicate.lean index db062c6663b88..63147fad1e7f8 100644 --- a/Mathlib/CategoryTheory/Localization/Predicate.lean +++ b/Mathlib/CategoryTheory/Localization/Predicate.lean @@ -236,7 +236,7 @@ instance : IsEquivalence (whiskeringLeftFunctor L W E) := by dsimp [Construction.whiskeringLeftEquivalence, equivalenceFromModel, whiskerLeft] erw [NatTrans.comp_app, NatTrans.comp_app, eqToHom_app, eqToHom_app, eqToHom_refl, eqToHom_refl, comp_id, id_comp] - . rfl + · rfl all_goals change (W.Q ⋙ Localization.Construction.lift L (inverts L W)) ⋙ _ = L ⋙ _ rw [Construction.fac] diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean index 29c5186eb87f4..12c765c4085d8 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Basic.lean @@ -324,12 +324,12 @@ def project : MonoidalFunctor (F C) D where μ X Y := 𝟙 _ μ_natural := @fun _ _ _ _ f g => by induction' f using Quotient.recOn - . induction' g using Quotient.recOn - . dsimp + · induction' g using Quotient.recOn + · dsimp simp rfl - . rfl - . rfl + · rfl + · rfl #align category_theory.free_monoidal_category.project CategoryTheory.FreeMonoidalCategory.project end Functor diff --git a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean index b0d158add5de1..492a3e635b3aa 100644 --- a/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean +++ b/Mathlib/CategoryTheory/Monoidal/Free/Coherence.lean @@ -245,8 +245,8 @@ def normalizeIso : tensorFunc C ≅ normalize' C := rintro X Y f induction' f using Quotient.recOn with f ; swap ; rfl induction' f with _ X₁ X₂ X₃ _ _ _ _ _ _ _ _ _ _ _ _ h₁ h₂ X₁ X₂ Y₁ Y₂ f g h₁ h₂ - . simp only [mk_id, Functor.map_id, Category.comp_id, Category.id_comp] - . ext n + · simp only [mk_id, Functor.map_id, Category.comp_id, Category.id_comp] + · ext n dsimp rw [mk_α_hom, NatTrans.comp_app, NatTrans.comp_app] dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] @@ -258,41 +258,41 @@ def normalizeIso : tensorFunc C ≅ normalize' C := pentagon_inv_assoc (inclusionObj n.as) X₁ X₂ X₃, tensor_inv_hom_id_assoc, tensor_id, Category.id_comp, Iso.inv_hom_id, Category.comp_id] - . ext n + · ext n dsimp rw [mk_α_inv, NatTrans.comp_app, NatTrans.comp_app] dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] simp only [Category.assoc, comp_tensor_id, tensor_id, Category.comp_id, pentagon_inv_assoc, ← associator_inv_naturality_assoc] rfl - . ext n + · ext n dsimp [Functor.comp] rw [mk_l_hom, NatTrans.comp_app, NatTrans.comp_app] dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] simp only [triangle_assoc_comp_right_assoc, Category.assoc, Category.comp_id] rfl - . ext n + · ext n dsimp [Functor.comp] rw [mk_l_inv, NatTrans.comp_app, NatTrans.comp_app] dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] simp only [triangle_assoc_comp_left_inv_assoc, inv_hom_id_tensor_assoc, tensor_id, Category.id_comp, Category.comp_id] rfl - . ext n + · ext n dsimp rw [mk_ρ_hom, NatTrans.comp_app, NatTrans.comp_app] dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] simp only [← (Iso.inv_comp_eq _).2 (rightUnitor_tensor _ _), Category.assoc, ← rightUnitor_naturality, Category.comp_id]; rfl - . ext n + · ext n dsimp rw [mk_ρ_inv, NatTrans.comp_app, NatTrans.comp_app] dsimp [NatIso.ofComponents, normalizeMapAux, whiskeringRight, whiskerRight, Functor.comp] simp only [← (Iso.eq_comp_inv _).1 (rightUnitor_tensor_inv _ _), rightUnitor_conjugation, Category.assoc, Iso.hom_inv_id_assoc, Iso.inv_hom_id_assoc, Iso.inv_hom_id, Discrete.functor, Category.comp_id, Function.comp] - . rw [mk_comp, Functor.map_comp, Functor.map_comp, Category.assoc, h₂, reassoc_of% h₁] - . ext ⟨n⟩ + · rw [mk_comp, Functor.map_comp, Functor.map_comp, Category.assoc, h₂, reassoc_of% h₁] + · ext ⟨n⟩ replace h₁ := NatTrans.congr_app h₁ ⟨n⟩ replace h₂ := NatTrans.congr_app h₂ ((Discrete.functor (normalizeObj X₁)).obj ⟨n⟩) have h₃ := (normalizeIsoAux _ Y₂).hom.naturality ((normalizeMapAux f).app ⟨n⟩) diff --git a/Mathlib/CategoryTheory/Monoidal/Transport.lean b/Mathlib/CategoryTheory/Monoidal/Transport.lean index d661f304c2847..58f7064c31978 100644 --- a/Mathlib/CategoryTheory/Monoidal/Transport.lean +++ b/Mathlib/CategoryTheory/Monoidal/Transport.lean @@ -135,8 +135,8 @@ def transport (e : C ≌ D) : MonoidalCategory.{v₂} D where slice_lhs 1 2 => rw [tensor_id_comp_id_tensor] conv_rhs => congr - . skip - . rw [← id_tensor_comp_tensor_id, id_tensor_comp] + · skip + · rw [← id_tensor_comp_tensor_id, id_tensor_comp] simp only [Category.assoc] slice_rhs 1 2 => rw [← id_tensor_comp, Iso.hom_inv_id_app] diff --git a/Mathlib/CategoryTheory/Preadditive/Injective.lean b/Mathlib/CategoryTheory/Preadditive/Injective.lean index 14032b3f61a5e..9856287861e66 100644 --- a/Mathlib/CategoryTheory/Preadditive/Injective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Injective.lean @@ -365,12 +365,10 @@ def injectivePresentationOfMapInjectivePresentation (X : C) theorem enoughInjectives_iff (F : C ≌ D) : EnoughInjectives C ↔ EnoughInjectives D := by constructor all_goals intro H; constructor; intro X; constructor - · - exact + · exact F.symm.injectivePresentationOfMapInjectivePresentation _ (Nonempty.some (H.presentation (F.inverse.obj X))) - · - exact + · exact F.injectivePresentationOfMapInjectivePresentation X (Nonempty.some (H.presentation (F.functor.obj X))) #align category_theory.equivalence.enough_injectives_iff CategoryTheory.Equivalence.enoughInjectives_iff diff --git a/Mathlib/CategoryTheory/Preadditive/Mat.lean b/Mathlib/CategoryTheory/Preadditive/Mat.lean index 6a510e61cd5a8..2874bec3079df 100644 --- a/Mathlib/CategoryTheory/Preadditive/Mat.lean +++ b/Mathlib/CategoryTheory/Preadditive/Mat.lean @@ -221,9 +221,9 @@ instance hasFiniteBiproducts : HasFiniteBiproducts (Mat_ C) split_ifs with h h' · substs h h' simp only [CategoryTheory.eqToHom_refl, CategoryTheory.Mat_.id_apply_self] - . subst h + · subst h rw [eqToHom_refl, id_apply_of_ne _ _ _ h'] - . rfl } + · rfl } (by dsimp ext1 ⟨i, j⟩ @@ -231,30 +231,30 @@ instance hasFiniteBiproducts : HasFiniteBiproducts (Mat_ C) rw [Finset.sum_apply, Finset.sum_apply] dsimp rw [Finset.sum_eq_single i] ; rotate_left - . intro b _ hb + · intro b _ hb apply Finset.sum_eq_zero intro x _ rw [dif_neg hb.symm, zero_comp] - . intro hi + · intro hi simp at hi rw [Finset.sum_eq_single j] ; rotate_left - . intro b _ hb + · intro b _ hb rw [dif_pos rfl, dif_neg, zero_comp] simp only tauto - . intro hj + · intro hj simp at hj simp only [eqToHom_refl, dite_eq_ite, ite_true, Category.id_comp, ne_eq, Sigma.mk.inj_iff, not_and, id_def] by_cases i' = i - . subst h + · subst h rw [dif_pos rfl] simp only [heq_eq_eq, true_and] by_cases j' = j - . subst h + · subst h simp - . rw [dif_neg h, dif_neg (Ne.symm h)] - . rw [dif_neg h, dif_neg] + · rw [dif_neg h, dif_neg (Ne.symm h)] + · rw [dif_neg h, dif_neg] tauto ) } set_option linter.uppercaseLean3 false in #align category_theory.Mat_.has_finite_biproducts CategoryTheory.Mat_.hasFiniteBiproducts @@ -343,11 +343,11 @@ def isoBiproductEmbedding (M : Mat_ C) : M ≅ ⨁ fun i => (embedding C).obj (M funext i j dsimp [id_def] rw [Finset.sum_apply, Finset.sum_apply, Finset.sum_eq_single i] ; rotate_left - . intro b _ hb + · intro b _ hb dsimp simp only [Finset.sum_const, Finset.card_singleton, one_smul] rw [dif_neg hb.symm, zero_comp] - . intro h + · intro h simp at h simp inv_hom_id := by @@ -361,9 +361,9 @@ def isoBiproductEmbedding (M : Mat_ C) : M ≅ ⨁ fun i => (embedding C).obj (M simp only [embedding, comp_apply, comp_dite, dite_comp, comp_zero, zero_comp, Finset.sum_dite_eq', Finset.mem_univ, ite_true, eqToHom_refl, Category.comp_id] split_ifs with h - . subst h + · subst h simp - . rfl + · rfl set_option linter.uppercaseLean3 false in #align category_theory.Mat_.iso_biproduct_embedding CategoryTheory.Mat_.isoBiproductEmbedding @@ -442,8 +442,8 @@ def lift (F : C ⥤ D) [Functor.Additive F] : Mat_ C ⥤ D where dsimp ext i j by_cases h : j = i - . subst h; simp - . simp [h] + · subst h; simp + · simp [h] set_option linter.uppercaseLean3 false in #align category_theory.Mat_.lift CategoryTheory.Mat_.lift diff --git a/Mathlib/CategoryTheory/Preadditive/Schur.lean b/Mathlib/CategoryTheory/Preadditive/Schur.lean index f263165613f6e..43af2b3517189 100644 --- a/Mathlib/CategoryTheory/Preadditive/Schur.lean +++ b/Mathlib/CategoryTheory/Preadditive/Schur.lean @@ -206,7 +206,7 @@ theorem finrank_hom_simple_simple_eq_zero_iff (X Y : C) [FiniteDimensional 𝕜 refine' ⟨fun h => by rw [h]; simp, fun h => _⟩ have := finrank_hom_simple_simple_le_one 𝕜 X Y interval_cases finrank 𝕜 (X ⟶ Y) - . rfl + · rfl · exact False.elim (h rfl) #align category_theory.finrank_hom_simple_simple_eq_zero_iff CategoryTheory.finrank_hom_simple_simple_eq_zero_iff diff --git a/Mathlib/CategoryTheory/Sites/LeftExact.lean b/Mathlib/CategoryTheory/Sites/LeftExact.lean index c369a03d12180..c5205ff5dfd0d 100644 --- a/Mathlib/CategoryTheory/Sites/LeftExact.lean +++ b/Mathlib/CategoryTheory/Sites/LeftExact.lean @@ -188,7 +188,7 @@ instance preservesLimitsOfShape_plusFunctor refine' ⟨fun S => liftToPlusObjLimitObj.{w, v, u} F X.unop S, _, _⟩ · intro S k apply liftToPlusObjLimitObj_fac - . intro S m hm + · intro S m hm dsimp [liftToPlusObjLimitObj] simp_rw [← Category.assoc, Iso.eq_comp_inv, ← Iso.comp_inv_eq] refine' limit.hom_ext (fun k => _) diff --git a/Mathlib/CategoryTheory/Sites/Plus.lean b/Mathlib/CategoryTheory/Sites/Plus.lean index 952af4f5ec66d..0139e868070f0 100644 --- a/Mathlib/CategoryTheory/Sites/Plus.lean +++ b/Mathlib/CategoryTheory/Sites/Plus.lean @@ -265,12 +265,12 @@ theorem plusMap_toPlus : J.plusMap (J.toPlus P) = J.toPlus (J.plusObj P) := by convert (Multiequalizer.condition (S.unop.index P) ⟨_, _, _, II.f, 𝟙 _, I.f, II.f ≫ I.f, I.hf, Sieve.downward_closed _ I.hf _, by simp⟩) using 1 - . dsimp [diagram] + · dsimp [diagram] cases I simp only [Category.assoc, limit.lift_π, Multifork.ofι_pt, Multifork.ofι_π_app, Cover.Arrow.map_Y, Cover.Arrow.map_f] rfl - . erw [Multiequalizer.lift_ι] + · erw [Multiequalizer.lift_ι] dsimp [Cover.index] simp only [Functor.map_id, Category.comp_id] rfl diff --git a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean index 789fdc93bbf73..f3bc5dd3cfc18 100644 --- a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean +++ b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean @@ -975,11 +975,11 @@ theorem compatible_iff (x : FirstObj P S) : ((firstObjEqFamily P S).hom x).Compatible ↔ firstMap P S x = secondMap P S x := by rw [Presieve.compatible_iff_sieveCompatible] constructor - . intro t + · intro t apply SecondObj.ext intros Y Z g f hf simpa [firstMap, secondMap] using t _ g hf - . intro t Y Z f g hf + · intro t Y Z f g hf rw [Types.limit_ext_iff'] at t simpa [firstMap, secondMap] using t ⟨⟨Y, Z, g, f, hf⟩⟩ #align category_theory.equalizer.sieve.compatible_iff CategoryTheory.Equalizer.Sieve.compatible_iff @@ -998,7 +998,7 @@ theorem equalizer_sheaf_condition : rw [← Iso.toEquiv_symm_fun] rw [Equiv.eq_symm_apply] constructor - . intro q + · intro q funext Y f hf simpa [firstObjEqFamily, forkMap] using q _ _ · intro q Y f hf @@ -1056,7 +1056,7 @@ theorem compatible_iff (x : FirstObj P R) : ((firstObjEqFamily P R).hom x).Compatible ↔ firstMap P R x = secondMap P R x := by rw [Presieve.pullbackCompatible_iff] constructor - . intro t + · intro t apply Limits.Types.limit_ext.{max u₁ v₁, u₁} rintro ⟨⟨Y, f, hf⟩, Z, g, hg⟩ simpa [firstMap, secondMap] using t hf hg diff --git a/Mathlib/CategoryTheory/Sites/Sieves.lean b/Mathlib/CategoryTheory/Sites/Sieves.lean index db5c779db6bd9..80de35eea001f 100644 --- a/Mathlib/CategoryTheory/Sites/Sieves.lean +++ b/Mathlib/CategoryTheory/Sites/Sieves.lean @@ -339,8 +339,8 @@ instance : CompleteLattice (Sieve X) le_sup_right _ _ _ _ := Or.inr sup_le _ _ _ h₁ h₂ _ f := by--ℰ S hS Y f := by rintro (hf | hf) - . exact h₁ _ hf - . exact h₂ _ hf + · exact h₁ _ hf + · exact h₂ _ hf inf_le_left _ _ _ _ := And.left inf_le_right _ _ _ _ := And.right le_inf _ _ _ p q _ _ z := ⟨p _ z, q _ z⟩ diff --git a/Mathlib/CategoryTheory/Subobject/Basic.lean b/Mathlib/CategoryTheory/Subobject/Basic.lean index 0369aa8204f1f..e63f0e8fb0827 100644 --- a/Mathlib/CategoryTheory/Subobject/Basic.lean +++ b/Mathlib/CategoryTheory/Subobject/Basic.lean @@ -610,14 +610,14 @@ def mapIsoToOrderIso (e : X ≅ Y) : Subobject X ≃o Subobject Y where map_rel_iff' {A B} := by dsimp constructor - . intro h + · intro h apply_fun (map e.inv).obj at h - . simpa only [← map_comp, e.hom_inv_id, map_id] using h - . apply Functor.monotone - . intro h + · simpa only [← map_comp, e.hom_inv_id, map_id] using h + · apply Functor.monotone + · intro h apply_fun (map e.hom).obj at h - . exact h - . apply Functor.monotone + · exact h + · apply Functor.monotone #align category_theory.subobject.map_iso_to_order_iso CategoryTheory.Subobject.mapIsoToOrderIso @[simp] diff --git a/Mathlib/CategoryTheory/Subobject/Comma.lean b/Mathlib/CategoryTheory/Subobject/Comma.lean index a5b0a031dceec..f12668118f9e4 100644 --- a/Mathlib/CategoryTheory/Subobject/Comma.lean +++ b/Mathlib/CategoryTheory/Subobject/Comma.lean @@ -112,8 +112,8 @@ def subobjectEquiv [HasLimits C] [PreservesLimits T] (A : StructuredArrow S T) : refine' ⟨fun h => Subobject.mk_le_mk_of_comm _ _, fun h => _⟩ · exact homMk (Subobject.ofMkLEMk _ _ h) ((cancel_mono (T.map g.right)).1 (by simp [← T.map_comp])) - . aesop_cat - . refine' Subobject.mk_le_mk_of_comm (Subobject.ofMkLEMk _ _ h).right _ + · aesop_cat + · refine' Subobject.mk_le_mk_of_comm (Subobject.ofMkLEMk _ _ h).right _ exact congr_arg CommaMorphism.right (Subobject.ofMkLEMk_comp h) #align category_theory.structured_arrow.subobject_equiv CategoryTheory.StructuredArrow.subobjectEquiv @@ -224,7 +224,7 @@ def quotientEquiv [HasColimits C] [PreservesColimits S] (A : CostructuredArrow S dsimp simp only [← S.map_comp_assoc, unop_left_comp_ofMkLEMk_unop, unop_op, CommaMorphism.w, Functor.const_obj_obj, right_eq_id, Functor.const_obj_map, Category.comp_id] - . apply Quiver.Hom.unop_inj + · apply Quiver.Hom.unop_inj ext exact unop_left_comp_ofMkLEMk_unop _ · refine' Subobject.mk_le_mk_of_comm (Subobject.ofMkLEMk _ _ h).unop.left.op _ diff --git a/Mathlib/CategoryTheory/Subobject/Lattice.lean b/Mathlib/CategoryTheory/Subobject/Lattice.lean index a2a639e219fc9..25d6b377d1e68 100644 --- a/Mathlib/CategoryTheory/Subobject/Lattice.lean +++ b/Mathlib/CategoryTheory/Subobject/Lattice.lean @@ -205,7 +205,7 @@ def supLe {A : C} (f g h : MonoOver A) : (f ⟶ h) → (g ⟶ h) → ((sup.obj f intro k₁ k₂ refine' homMk _ _ apply image.lift ⟨_, h.arrow, coprod.desc k₁.left k₂.left, _⟩ - . ext + · ext · simp [w k₁] · simp [w k₂] · apply image.lift_fac @@ -438,8 +438,8 @@ theorem finset_inf_factors {I : Type _} {A B : C} {s : Finset I} {P : I → Subo (s.inf P).Factors f ↔ ∀ i ∈ s, (P i).Factors f := by classical induction' s using Finset.induction_on with _ _ _ ih - . simp [top_factors] - . simp [ih] + · simp [top_factors] + · simp [ih] #align category_theory.subobject.finset_inf_factors CategoryTheory.Subobject.finset_inf_factors -- `i` is explicit here because often we'd like to defer a proof of `m` @@ -448,8 +448,8 @@ theorem finset_inf_arrow_factors {I : Type _} {B : C} (s : Finset I) (P : I → classical revert i m induction' s using Finset.induction_on with _ _ _ ih - . rintro _ ⟨⟩ - . intro _ m + · rintro _ ⟨⟩ + · intro _ m rw [Finset.inf_insert] simp only [Finset.mem_insert] at m rcases m with (rfl | m) @@ -475,8 +475,8 @@ theorem inf_eq_map_pullback {A : C} (f₁ : MonoOver A) (f₂ : Subobject A) : theorem prod_eq_inf {A : C} {f₁ f₂ : Subobject A} [HasBinaryProduct f₁ f₂] : (f₁ ⨯ f₂) = f₁ ⊓ f₂ := by apply le_antisymm - . refine' le_inf _ _ _ (Limits.prod.fst.le) (Limits.prod.snd.le) - . apply leOfHom + · refine' le_inf _ _ _ (Limits.prod.fst.le) (Limits.prod.snd.le) + · apply leOfHom exact prod.lift (inf_le_left _ _).hom (inf_le_right _ _).hom #align category_theory.subobject.prod_eq_inf CategoryTheory.Subobject.prod_eq_inf @@ -639,7 +639,7 @@ def sInf {A : C} (s : Set (Subobject A)) : Subobject A := theorem sInf_le {A : C} (s : Set (Subobject A)) (f) (hf : f ∈ s) : sInf s ≤ f := by fapply le_of_comm - . exact (underlyingIso _).hom ≫ + · exact (underlyingIso _).hom ≫ Limits.limit.π (wideCospan s) (some ⟨equivShrink (Subobject A) f, Set.mem_image_of_mem (equivShrink (Subobject A)) hf⟩) ≫ @@ -703,7 +703,7 @@ theorem symm_apply_mem_iff_mem_image {α β : Type _} (e : α ≃ β) (s : Set theorem sSup_le {A : C} (s : Set (Subobject A)) (f : Subobject A) (k : ∀ g ∈ s, g ≤ f) : sSup s ≤ f := by fapply le_of_comm - . refine'(underlyingIso _).hom ≫ image.lift ⟨_, f.arrow, _, _⟩ + · refine'(underlyingIso _).hom ≫ image.lift ⟨_, f.arrow, _, _⟩ · refine' Sigma.desc _ rintro ⟨g, m⟩ refine' underlying.map (homOfLE (k _ _)) @@ -711,7 +711,7 @@ theorem sSup_le {A : C} (s : Set (Subobject A)) (f : Subobject A) (k : ∀ g ∈ · ext dsimp [smallCoproductDesc] simp - . dsimp [sSup] + · dsimp [sSup] rw [assoc, image.lift_fac, underlyingIso_hom_comp_eq_mk] #align category_theory.subobject.Sup_le CategoryTheory.Subobject.sSup_le @@ -763,10 +763,10 @@ def subobjectOrderIso {X : C} (Y : Subobject X) : Subobject (Y : C) ≃o Set.Iic map_rel_iff' {W Z} := by dsimp constructor - . intro h + · intro h exact le_of_comm (((underlyingIso _).inv ≫ ofLE _ _ (Subtype.mk_le_mk.mp h) ≫ (underlyingIso _).hom)) (by aesop_cat) - . intro h + · intro h exact Subtype.mk_le_mk.mpr (le_of_comm ((underlyingIso _).hom ≫ ofLE _ _ h ≫ (underlyingIso _).inv) (by simp)) #align category_theory.subobject.subobject_order_iso CategoryTheory.Subobject.subobjectOrderIso diff --git a/Mathlib/CategoryTheory/WithTerminal.lean b/Mathlib/CategoryTheory/WithTerminal.lean index 6a3a7964ba394..8a3ee6e524f18 100644 --- a/Mathlib/CategoryTheory/WithTerminal.lean +++ b/Mathlib/CategoryTheory/WithTerminal.lean @@ -102,9 +102,9 @@ instance : Category.{v} (WithTerminal C) where -- so the `false_of_from_star` destruct rule below can be used here. -- That works, but causes mysterious failures of `aesop_cat` in `map`. cases a <;> cases b <;> cases c <;> cases d <;> try aesop_cat - . exact (h : PEmpty).elim - . exact (g : PEmpty).elim - . exact (h : PEmpty).elim + · exact (h : PEmpty).elim + · exact (g : PEmpty).elim + · exact (h : PEmpty).elim /-- Helper function for typechecking. -/ def down {X Y : C} (f : of X ⟶ of Y) : X ⟶ Y := f @@ -289,9 +289,9 @@ instance : Category.{v} (WithInitial C) where -- Porting note: it would be nice to automate this away as well. -- See the note on `Category (WithTerminal C)` cases a <;> cases b <;> cases c <;> cases d <;> try aesop_cat - . exact (g : PEmpty).elim - . exact (f : PEmpty).elim - . exact (f : PEmpty).elim + · exact (g : PEmpty).elim + · exact (f : PEmpty).elim + · exact (f : PEmpty).elim /-- Helper function for typechecking. -/ def down {X Y : C} (f : of X ⟶ of Y) : X ⟶ Y := f diff --git a/Mathlib/Combinatorics/Composition.lean b/Mathlib/Combinatorics/Composition.lean index 02854d5eadbb2..37358f35c1ea9 100644 --- a/Mathlib/Combinatorics/Composition.lean +++ b/Mathlib/Combinatorics/Composition.lean @@ -677,8 +677,8 @@ theorem splitWrtCompositionAux_cons (l : List α) (n ns) : theorem length_splitWrtCompositionAux (l : List α) (ns) : length (l.splitWrtCompositionAux ns) = ns.length := by induction ns generalizing l - . simp [splitWrtCompositionAux, *] - . simp [*] + · simp [splitWrtCompositionAux, *] + · simp [*] #align list.length_split_wrt_composition_aux List.length_splitWrtCompositionAux /-- When one splits a list along a composition `c`, the number of sublists thus created is @@ -693,12 +693,12 @@ theorem length_splitWrtComposition (l : List α) (c : Composition n) : theorem map_length_splitWrtCompositionAux {ns : List ℕ} : ∀ {l : List α}, ns.sum ≤ l.length → map length (l.splitWrtCompositionAux ns) = ns := by induction' ns with n ns IH <;> intro l h <;> simp at h - . simp + · simp have := le_trans (Nat.le_add_right _ _) h simp only [splitWrtCompositionAux_cons, this] ; dsimp rw [length_take, IH] <;> simp [length_drop] - . assumption - . exact le_tsub_of_add_le_left h + · assumption + · exact le_tsub_of_add_le_left h #align list.map_length_split_wrt_composition_aux List.map_length_splitWrtCompositionAux /-- When one splits a list along a composition `c`, the lengths of the sublists thus created are @@ -730,9 +730,9 @@ theorem nthLe_splitWrtCompositionAux (l : List α) (ns : List ℕ) {i : ℕ} (hi induction' ns with n ns IH generalizing l i · cases hi cases' i with i - . rw [Nat.add_zero, List.take_zero, sum_nil, nthLe_zero]; dsimp + · rw [Nat.add_zero, List.take_zero, sum_nil, nthLe_zero]; dsimp simp only [splitWrtCompositionAux_cons, head!, sum, foldl, zero_add] - . simp only [splitWrtCompositionAux_cons, take, sum_cons, + · simp only [splitWrtCompositionAux_cons, take, sum_cons, Nat.add_eq, add_zero, gt_iff_lt, nthLe_cons, IH]; dsimp rw [Nat.succ_sub_succ_eq_sub, ←Nat.succ_eq_add_one,tsub_zero] simp only [← drop_take, drop_drop] @@ -763,7 +763,7 @@ theorem join_splitWrtCompositionAux {ns : List ℕ} : simp only [splitWrtCompositionAux_cons] ; dsimp rw [IH] · simp - . rw [length_drop, ← h, add_tsub_cancel_left] + · rw [length_drop, ← h, add_tsub_cancel_left] #align list.join_split_wrt_composition_aux List.join_splitWrtCompositionAux /-- If one splits a list along a composition, and then joins the sublists, one gets back the @@ -855,9 +855,9 @@ def compositionAsSetEquiv (n : ℕ) : CompositionAsSet n ≃ Finset (Fin (n - 1) constructor · intro h cases' h with n h - . rw [add_comm] at this + · rw [add_comm] at this contradiction - . cases' h with w h; cases' h with h₁ h₂ + · cases' h with w h; cases' h with h₁ h₂ rw [←Fin.ext_iff] at h₂ rwa [h₂] · intro h diff --git a/Mathlib/Combinatorics/Derangements/Basic.lean b/Mathlib/Combinatorics/Derangements/Basic.lean index 28f94562fc4fc..b8d41fcc9c952 100644 --- a/Mathlib/Combinatorics/Derangements/Basic.lean +++ b/Mathlib/Combinatorics/Derangements/Basic.lean @@ -91,10 +91,10 @@ def atMostOneFixedPointEquivSum_derangements [DecidableEq α] (a : α) : { f : Perm α // fixedPoints f ⊆ {a} ∧ a ∉ fixedPoints f } := by -- porting note: `subtypeSubtypeEquivSubtypeInter` no longer works with placeholder `_`s. refine' Equiv.sumCongr _ _ - . exact subtypeSubtypeEquivSubtypeInter + · exact subtypeSubtypeEquivSubtypeInter (fun x : Perm α => fixedPoints x ⊆ {a}) (a ∈ fixedPoints ·) - . exact subtypeSubtypeEquivSubtypeInter + · exact subtypeSubtypeEquivSubtypeInter (fun x : Perm α => fixedPoints x ⊆ {a}) (¬a ∈ fixedPoints ·) _ ≃ Sum { f : Perm α // fixedPoints f = {a} } { f : Perm α // fixedPoints f = ∅ } := by diff --git a/Mathlib/Combinatorics/Young/SemistandardTableau.lean b/Mathlib/Combinatorics/Young/SemistandardTableau.lean index 291830f8984f7..0e0bf9a1b0429 100644 --- a/Mathlib/Combinatorics/Young/SemistandardTableau.lean +++ b/Mathlib/Combinatorics/Young/SemistandardTableau.lean @@ -128,15 +128,15 @@ theorem zeros {μ : YoungDiagram} (T : Ssyt μ) {i j : ℕ} (not_cell : (i, j) theorem row_weak_of_le {μ : YoungDiagram} (T : Ssyt μ) {i j1 j2 : ℕ} (hj : j1 ≤ j2) (cell : (i, j2) ∈ μ) : T i j1 ≤ T i j2 := by cases' eq_or_lt_of_le hj with h h - . rw [h] - . exact T.row_weak h cell + · rw [h] + · exact T.row_weak h cell #align ssyt.row_weak_of_le Ssyt.row_weak_of_le theorem col_weak {μ : YoungDiagram} (T : Ssyt μ) {i1 i2 j : ℕ} (hi : i1 ≤ i2) (cell : (i2, j) ∈ μ) : T i1 j ≤ T i2 j := by cases' eq_or_lt_of_le hi with h h - . rw [h] - . exact le_of_lt (T.col_strict h cell) + · rw [h] + · exact le_of_lt (T.col_strict h cell) #align ssyt.col_weak Ssyt.col_weak /-- The "highest weight" SSYT of a given shape has all i's in row i, for each i. -/ diff --git a/Mathlib/Control/EquivFunctor/Instances.lean b/Mathlib/Control/EquivFunctor/Instances.lean index 6089fcb8c453f..9223a742c4ca6 100644 --- a/Mathlib/Control/EquivFunctor/Instances.lean +++ b/Mathlib/Control/EquivFunctor/Instances.lean @@ -39,9 +39,9 @@ instance EquivFunctorFinset : EquivFunctor Finset where map_refl' α := by ext; simp map_trans' k h := by ext _ a; simp; constructor <;> intro h' - . let ⟨a, ha₁, ha₂⟩ := h' + · let ⟨a, ha₁, ha₂⟩ := h' rw [← ha₂]; simp; apply ha₁ - . exists (Equiv.symm k) ((Equiv.symm h) a) + · exists (Equiv.symm k) ((Equiv.symm h) a) simp [h'] #align equiv_functor_finset EquivFunctorFinset diff --git a/Mathlib/Control/Fold.lean b/Mathlib/Control/Fold.lean index ea5994ceb241c..a95516f3dc50e 100644 --- a/Mathlib/Control/Fold.lean +++ b/Mathlib/Control/Fold.lean @@ -398,8 +398,8 @@ theorem length_toList {xs : t α} : length xs = List.length (toList xs) := by rw [← Nat.add_zero ys.length] generalize 0 = n induction' ys with _ _ ih generalizing n - . simp - . simp_arith [ih] + · simp + · simp_arith [ih] #align traversable.length_to_list Traversable.length_toList variable {m : Type u → Type u} [Monad m] [LawfulMonad m] diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 9f45645695ddd..3666cecbfb67d 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -556,8 +556,8 @@ theorem exp_sub : exp (x - y) = exp x / exp y := by theorem exp_int_mul (z : ℂ) (n : ℤ) : Complex.exp (n * z) = Complex.exp z ^ n := by cases n - . simp [exp_nat_mul] - . simp [exp_add, add_mul, pow_add, exp_neg, exp_nat_mul] + · simp [exp_nat_mul] + · simp [exp_add, add_mul, pow_add, exp_neg, exp_nat_mul] #align complex.exp_int_mul Complex.exp_int_mul @[simp] diff --git a/Mathlib/Data/ENat/Basic.lean b/Mathlib/Data/ENat/Basic.lean index 1cfb472ae5db4..e8b8cc914316e 100644 --- a/Mathlib/Data/ENat/Basic.lean +++ b/Mathlib/Data/ENat/Basic.lean @@ -210,8 +210,8 @@ theorem nat_induction {P : ℕ∞ → Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ (htop : (∀ n : ℕ, P n) → P ⊤) : P a := by have A : ∀ n : ℕ, P n := fun n => Nat.recOn n h0 hsuc cases a - . exact htop A - . exact A _ + · exact htop A + · exact A _ #align enat.nat_induction ENat.nat_induction end ENat diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index a2b92498ec4a8..1d34958229382 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -255,14 +255,14 @@ theorem forall_iff {p : Fin n → Prop} : (∀ i, p i) ↔ ∀ i h, p ⟨i, h⟩ lemma ite_val {n : ℕ} {c : Prop} [Decidable c] {x : c → Fin n} (y : ¬c → Fin n) : (if h : c then x h else y h).val = if h : c then (x h).val else (y h).val := by by_cases c - . simp only [dif_pos h] - . simp only [dif_neg h] + · simp only [dif_pos h] + · simp only [dif_neg h] lemma dite_val {n : ℕ} {c : Prop} [Decidable c] {x y : Fin n} : (if c then x else y).val = if c then x.val else y.val := by by_cases c - . simp only [if_pos h] - . simp only [if_neg h] + · simp only [if_pos h] + · simp only [if_neg h] end coe diff --git a/Mathlib/Data/Finset/NoncommProd.lean b/Mathlib/Data/Finset/NoncommProd.lean index 99b53e76123dc..952549cc6e706 100644 --- a/Mathlib/Data/Finset/NoncommProd.lean +++ b/Mathlib/Data/Finset/NoncommProd.lean @@ -422,15 +422,15 @@ theorem noncommProd_mul_single [Fintype ι] [DecidableEq ι] (x : ∀ i, M i) : ext i apply (univ.noncommProd_map (fun i => MonoidHom.single M i (x i)) _ (Pi.evalMonoidHom M i)).trans refine' (noncommProd_congr (insert_erase (mem_univ i)).symm _ _).trans _ - . intro i _ j _ _ + · intro i _ j _ _ exact Pi.mulSingle_apply_commute x i j - . intro j + · intro j exact Pi.mulSingle j (x j) i - . intro j _; dsimp - . rw [noncommProd_insert_of_not_mem _ _ _ _ (not_mem_erase _ _), + · intro j _; dsimp + · rw [noncommProd_insert_of_not_mem _ _ _ _ (not_mem_erase _ _), noncommProd_eq_pow_card (univ.erase i), one_pow, mul_one] simp - . intro j hj + · intro j hj simp at hj simp [Pi.mulSingle, Function.update] intro h diff --git a/Mathlib/Data/Finset/Powerset.lean b/Mathlib/Data/Finset/Powerset.lean index d66c51edd418f..82e5953bf9f18 100644 --- a/Mathlib/Data/Finset/Powerset.lean +++ b/Mathlib/Data/Finset/Powerset.lean @@ -336,7 +336,7 @@ theorem powersetLen_map {β : Type _} (f : α ↪ β) (n : ℕ) (s : Finset α) ext <| fun t => by simp only [card_map, mem_powersetLen, le_eq_subset, gt_iff_lt, mem_map, mapEmbedding_apply] constructor - . classical + · classical intro h have : map f (filter (fun x => (f x ∈ t)) s) = t := by ext x @@ -345,7 +345,7 @@ theorem powersetLen_map {β : Type _} (f : α ↪ β) (n : ℕ) (s : Finset α) fun hx => let ⟨y, hy⟩ := mem_map.1 (h.1 hx); ⟨y, ⟨hy.1, hy.2 ▸ hx⟩, hy.2⟩⟩ refine' ⟨_, _, this⟩ rw [← card_map f, this, h.2]; simp - . rintro ⟨a, ⟨has, rfl⟩, rfl⟩ + · rintro ⟨a, ⟨has, rfl⟩, rfl⟩ dsimp [RelEmbedding.coe_toEmbedding] --Porting note: Why is `rw` required here and not `simp`? rw [mapEmbedding_apply] diff --git a/Mathlib/Data/Finset/Prod.lean b/Mathlib/Data/Finset/Prod.lean index 8d78923bbb1c7..d29d7607d213f 100644 --- a/Mathlib/Data/Finset/Prod.lean +++ b/Mathlib/Data/Finset/Prod.lean @@ -170,8 +170,8 @@ theorem filter_product_card (s : Finset α) (t : Finset β) (p : α → Prop) (q ext ⟨a, b⟩ simp only [filter_union_right, mem_filter, mem_product] constructor <;> intro h <;> use h.1 - . simp only [h.2, Function.comp_apply, Decidable.em, and_self] - . revert h + · simp only [h.2, Function.comp_apply, Decidable.em, and_self] + · revert h simp only [Function.comp_apply, and_imp] rintro _ _ (_|_) <;> simp [*] · apply Finset.disjoint_filter_filter' diff --git a/Mathlib/Data/Finset/Sigma.lean b/Mathlib/Data/Finset/Sigma.lean index 310eaa4c3b1ca..323a0c6db09ed 100644 --- a/Mathlib/Data/Finset/Sigma.lean +++ b/Mathlib/Data/Finset/Sigma.lean @@ -175,8 +175,8 @@ theorem sigmaLift_nonempty : theorem sigmaLift_eq_empty : sigmaLift f a b = ∅ ↔ ∀ h : a.1 = b.1, f (h ▸ a.2) b.2 = ∅ := by simp_rw [sigmaLift] split_ifs with h - . simp [h, forall_prop_of_true h] - . simp [h, forall_prop_of_false h] + · simp [h, forall_prop_of_true h] + · simp [h, forall_prop_of_false h] #align finset.sigma_lift_eq_empty Finset.sigmaLift_eq_empty theorem sigmaLift_mono (h : ∀ ⦃i⦄ ⦃a : α i⦄ ⦃b : β i⦄, f a b ⊆ g a b) (a : Σi, α i) (b : Σi, β i) : diff --git a/Mathlib/Data/Finsupp/AList.lean b/Mathlib/Data/Finsupp/AList.lean index 13dd756859d9e..58c8dfc8c31d5 100644 --- a/Mathlib/Data/Finsupp/AList.lean +++ b/Mathlib/Data/Finsupp/AList.lean @@ -88,8 +88,8 @@ theorem lookupFinsupp_support [DecidableEq α] [DecidableEq M] (l : AList fun _x l.lookupFinsupp.support = (l.1.filter fun x => Sigma.snd x ≠ 0).keys.toFinset := by -- porting note: was `convert rfl` simp only [lookupFinsupp, ne_eq, Finsupp.coe_mk]; congr - . apply Subsingleton.elim - . funext ; congr + · apply Subsingleton.elim + · funext ; congr #align alist.lookup_finsupp_support AList.lookupFinsupp_support theorem lookupFinsupp_eq_iff_of_ne_zero [DecidableEq α] {l : AList fun _x : α => M} {a : α} {x : M} diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 78d4c82147995..99dbd049eff7e 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -662,8 +662,7 @@ theorem mapDomain_injOn (S : Set α) {f : α → β} (hf : Set.InjOn f S) : ext a classical by_cases h : a ∈ v₁.support ∪ v₂.support - · - rw [← mapDomain_apply' S _ hv₁ hf _, ← mapDomain_apply' S _ hv₂ hf _, eq] <;> + · rw [← mapDomain_apply' S _ hv₁ hf _, ← mapDomain_apply' S _ hv₂ hf _, eq] <;> · apply Set.union_subset hv₁ hv₂ exact_mod_cast h · simp only [not_or, mem_union, not_not, mem_support_iff] at h @@ -1224,8 +1223,7 @@ theorem sum_curry_index (f : α × β →₀ M) (g : α → β → M → N) (hg (f.curry.sum fun a f => f.sum (g a)) = f.sum fun p c => g p.1 p.2 c := by rw [Finsupp.curry] trans - · - exact + · exact sum_sum_index (fun a => sum_zero_index) fun a b₀ b₁ => sum_add_index' (fun a => hg₀ _ _) fun c d₀ d₁ => hg₁ _ _ _ _ congr ; funext p c diff --git a/Mathlib/Data/Int/Bitwise.lean b/Mathlib/Data/Int/Bitwise.lean index 798bbf8a1a1fe..2bc3495029519 100644 --- a/Mathlib/Data/Int/Bitwise.lean +++ b/Mathlib/Data/Int/Bitwise.lean @@ -225,15 +225,15 @@ theorem bitwise_or : bitwise or = lor := by cases' m with m m <;> cases' n with n n <;> try {rfl} <;> simp only [bitwise, natBitwise, Bool.not_false, Bool.or_true, cond_true, lor, Nat.ldiff', negSucc.injEq, Bool.true_or, Nat.land'] - . rw [Nat.bitwise'_swap, Function.swap] + · rw [Nat.bitwise'_swap, Function.swap] congr funext x y cases x <;> cases y <;> rfl rfl - . congr + · congr funext x y cases x <;> cases y <;> rfl - . congr + · congr funext x y cases x <;> cases y <;> rfl #align int.bitwise_or Int.bitwise_or @@ -245,12 +245,12 @@ theorem bitwise_and : bitwise and = land := by <;> simp only [bitwise, natBitwise, Bool.not_false, Bool.or_true, cond_false, cond_true, lor, Nat.ldiff', Bool.and_true, negSucc.injEq, Bool.and_false, Nat.land'] - . rw [Nat.bitwise'_swap, Function.swap] + · rw [Nat.bitwise'_swap, Function.swap] congr funext x y cases x <;> cases y <;> rfl rfl - . congr + · congr funext x y cases x <;> cases y <;> rfl #align int.bitwise_and Int.bitwise_and @@ -262,13 +262,13 @@ theorem bitwise_diff : (bitwise fun a b => a && not b) = ldiff' := by <;> simp only [bitwise, natBitwise, Bool.not_false, Bool.or_true, cond_false, cond_true, lor, Nat.ldiff', Bool.and_true, negSucc.injEq, Bool.and_false, Nat.land', Bool.not_true, ldiff', Nat.lor'] - . congr + · congr funext x y cases x <;> cases y <;> rfl - . congr + · congr funext x y cases x <;> cases y <;> rfl - . rw [Nat.bitwise'_swap, Function.swap] + · rw [Nat.bitwise'_swap, Function.swap] congr funext x y cases x <;> cases y <;> rfl @@ -282,13 +282,13 @@ theorem bitwise_xor : bitwise xor = lxor' := by <;> simp only [bitwise, natBitwise, Bool.not_false, Bool.or_true, cond_false, cond_true, lor, Nat.ldiff', Bool.and_true, negSucc.injEq, Bool.false_xor, Bool.true_xor, Bool.and_false, Nat.land', Bool.not_true, ldiff', Nat.lor', lxor', Nat.lxor'] - . congr + · congr funext x y cases x <;> cases y <;> rfl - . congr + · congr funext x y cases x <;> cases y <;> rfl - . congr + · congr funext x y cases x <;> cases y <;> rfl #align int.bitwise_xor Int.bitwise_xor @@ -299,10 +299,10 @@ theorem bitwise_bit (f : Bool → Bool → Bool) (a m b n) : cases' m with m m <;> cases' n with n n <;> simp only [bitwise, ofNat_eq_coe, bit_coe_nat, natBitwise, Bool.not_false, Bool.not_eq_false', bit_negSucc] - . by_cases h : f false false <;> simp [h] - . by_cases h : f false true <;> simp [h] - . by_cases h : f true false <;> simp [h] - . by_cases h : f true true <;> simp [h] + · by_cases h : f false false <;> simp [h] + · by_cases h : f false true <;> simp [h] + · by_cases h : f true false <;> simp [h] + · by_cases h : f true true <;> simp [h] #align int.bitwise_bit Int.bitwise_bit @[simp] @@ -335,10 +335,10 @@ theorem lnot_bit (b) : ∀ n, lnot (bit b n) = bit (not b) (lnot n) theorem testBit_bitwise (f : Bool → Bool → Bool) (m n k) : testBit (bitwise f m n) k = f (testBit m k) (testBit n k) := by cases m <;> cases n <;> simp only [testBit, bitwise, natBitwise] - . by_cases h : f false false <;> simp [h] - . by_cases h : f false true <;> simp [h] - . by_cases h : f true false <;> simp [h] - . by_cases h : f true true <;> simp [h] + · by_cases h : f false false <;> simp [h] + · by_cases h : f false true <;> simp [h] + · by_cases h : f true false <;> simp [h] + · by_cases h : f true true <;> simp [h] #align int.test_bit_bitwise Int.testBit_bitwise @[simp] diff --git a/Mathlib/Data/Int/Range.lean b/Mathlib/Data/Int/Range.lean index b005fd5f886ff..3749379154638 100644 --- a/Mathlib/Data/Int/Range.lean +++ b/Mathlib/Data/Int/Range.lean @@ -51,9 +51,9 @@ instance decidableLELE (P : Int → Prop) [DecidablePred P] (m n : ℤ) : -- `Decidable (∀ (r : ℤ), r ∈ range m (n + 1) → P r)` apply decidable_of_iff (∀ r ∈ range m (n + 1), P r) apply Iff.intro <;> intros h _ _ - . intro _; apply h + · intro _; apply h simp_all only [mem_range_iff, and_imp, lt_add_one_iff] - . simp_all only [mem_range_iff, and_imp, lt_add_one_iff] + · simp_all only [mem_range_iff, and_imp, lt_add_one_iff] #align int.decidable_le_le Int.decidableLELE instance decidableLTLT (P : Int → Prop) [DecidablePred P] (m n : ℤ) : diff --git a/Mathlib/Data/List/AList.lean b/Mathlib/Data/List/AList.lean index 4fe605ff511d7..8976437d96dc1 100644 --- a/Mathlib/Data/List/AList.lean +++ b/Mathlib/Data/List/AList.lean @@ -506,7 +506,7 @@ theorem union_comm_of_disjoint {s₁ s₂ : AList β} (h : Disjoint s₁ s₂) : (by intros ; simp constructor <;> intro h' - . cases' h' with h' h' + · cases' h' with h' h' · right refine' ⟨_, h'⟩ apply h @@ -514,7 +514,7 @@ theorem union_comm_of_disjoint {s₁ s₂ : AList β} (h : Disjoint s₁ s₂) : exact rfl · left rw [h'.2] - . cases' h' with h' h' + · cases' h' with h' h' · right refine' ⟨_, h'⟩ intro h'' diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 3679a3382698e..574edb4d7f7c1 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -3181,8 +3181,8 @@ theorem get?_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : induction' l with hd tl hl generalizing n · simp · cases' n with n - . simp - . simp [hl] + · simp + · simp [hl] #align list.nth_pmap List.get?_pmap theorem get_pmap {p : α → Prop} (f : ∀ a, p a → β) {l : List α} (h : ∀ a ∈ l, p a) {n : ℕ} @@ -3543,8 +3543,8 @@ theorem mem_of_mem_filter {a : α} {l} (h : a ∈ filter p l) : a ∈ l := theorem mem_filter_of_mem {a : α} : ∀ {l}, a ∈ l → p a → a ∈ filter p l | x :: l, h, h1 => by rcases mem_cons.1 h with rfl | h - . simp [filter, h1] - . rw [filter] + · simp [filter, h1] + · rw [filter] cases p x <;> simp [mem_filter_of_mem h h1] #align list.mem_filter_of_mem List.mem_filter_of_mem diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 89eeb1bf4fd65..2b771ff9e4546 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -155,7 +155,7 @@ theorem chain_iff_get {R} : ∀ {a : α} {l : List α}, Chain R a l ↔ intro i w cases' i with i · apply h0 - . exact h i (lt_pred_iff.2 <| by simpa using w) + · exact h i (lt_pred_iff.2 <| by simpa using w) rintro ⟨h0, h⟩; constructor · apply h0 simp diff --git a/Mathlib/Data/List/Cycle.lean b/Mathlib/Data/List/Cycle.lean index 8982cc877279f..03a7e8647ad5f 100644 --- a/Mathlib/Data/List/Cycle.lean +++ b/Mathlib/Data/List/Cycle.lean @@ -73,8 +73,8 @@ theorem nextOr_eq_nextOr_of_mem_of_ne (xs : List α) (x d d' : α) (x_mem : x by_cases h : x = y · rw [h, nextOr_self_cons_cons, nextOr_self_cons_cons] · rw [nextOr, nextOr, IH] - . simpa [h] using x_mem - . simpa using x_ne + · simpa [h] using x_mem + · simpa using x_ne #align list.next_or_eq_next_or_of_mem_of_ne List.nextOr_eq_nextOr_of_mem_of_ne theorem mem_of_nextOr_ne {xs : List α} {x d : α} (h : nextOr xs x d ≠ d) : x ∈ xs := by @@ -285,24 +285,24 @@ theorem next_get : ∀ (l : List α) (_h : Nodup l) (i : Fin l.length), · subst hi' rw [next_getLast_cons] · simp [hi', get] - . rw [get_cons_succ]; exact get_mem _ _ _ - . exact hx' - . simp [getLast_eq_get] - . exact hn.of_cons - . rw [next_ne_head_ne_getLast _ _ _ _ _ hx'] + · rw [get_cons_succ]; exact get_mem _ _ _ + · exact hx' + · 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 - . rw [getLast_eq_get] + · 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 simp at this; simp [this] at hi' - . rw [get_cons_succ]; exact get_mem _ _ _ + · rw [get_cons_succ]; exact get_mem _ _ _ set_option linter.deprecated false in @[deprecated next_get] @@ -414,7 +414,7 @@ theorem prev_reverse_eq_next (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) rw [← nthLe_reverse] · simp [tsub_tsub_cancel_of_le (Nat.le_pred_of_lt hk)] · simpa using (Nat.sub_le _ _).trans_lt (tsub_lt_self lpos Nat.succ_pos') - . simpa + · simpa #align list.prev_reverse_eq_next List.prev_reverse_eq_next theorem next_reverse_eq_prev (l : List α) (h : Nodup l) (x : α) (hx : x ∈ l) : diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index 0f5984fbc2316..ffc709bb81fb9 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -202,8 +202,8 @@ theorem forall₂_iff_zip {l₁ l₂} : · cases length_eq_zero.1 h₁.symm constructor · cases' l₂ with b l₂ - . simp at h₁ - . simp only [length_cons, succ.injEq] at h₁ + · simp at h₁ + · simp only [length_cons, succ.injEq] at h₁ exact Forall₂.cons (h₂ <| by simp [zip]) (IH h₁ <| fun h => h₂ <| by simp only [zip, zipWith, find?, mem_cons, Prod.mk.injEq]; right diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index b79eeeb8e521c..01a0e76761e92 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -217,11 +217,11 @@ theorem index_of_argmax : obtain ha | ha := ha <;> split_ifs at hm <;> injection hm with hm <;> subst hm · cases not_le_of_lt ‹_› ‹_› · rw [if_pos rfl] - . rw [if_neg, if_neg] + · rw [if_neg, if_neg] 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] + · rw [if_pos rfl] exact Nat.zero_le _ #align list.index_of_argmax List.index_of_argmax diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index 9f079a586a4a0..9626437f5d5fa 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -99,9 +99,9 @@ theorem nodup_iff_injective_get {l : List α} : ⟨fun h i j hg => by cases' i with i hi; cases' j with j hj rcases lt_trichotomy i j with (hij | rfl | hji) - . exact (h ⟨i, hi⟩ ⟨j, hj⟩ hij hg).elim - . rfl - . exact (h ⟨j, hj⟩ ⟨i, hi⟩ hji hg.symm).elim, + · exact (h ⟨i, hi⟩ ⟨j, hj⟩ hij hg).elim + · rfl + · exact (h ⟨j, hj⟩ ⟨i, hi⟩ hji hg.symm).elim, fun hinj i j hij h => Nat.ne_of_lt hij (Fin.veq_of_eq (hinj h))⟩ set_option linter.deprecated false in @@ -128,10 +128,10 @@ theorem nodup_iff_get?_ne_get? {l : List α} : l.Nodup ↔ ∀ i j : ℕ, i < j → j < l.length → l.get? i ≠ l.get? j := by rw [Nodup, pairwise_iff_get] constructor - . intro h i j hij hj + · intro h i j hij hj rw [get?_eq_get (lt_trans hij hj), get?_eq_get hj, Ne.def, Option.some_inj] exact h _ _ hij - . intro h i j hij + · intro h i j hij rw [Ne.def, ← Option.some_inj, ← get?_eq_get, ← get?_eq_get] exact h i j hij j.2 #align list.nodup_iff_nth_ne_nth List.nodup_iff_get?_ne_get? diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index 4f9e8937161bd..3721ec96d8ef0 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -63,7 +63,7 @@ theorem get?_ofFn {n} (f : Fin n → α) (i) : get? (ofFn f) i = ofFnNthVal f i if h : i < (ofFn f).length then by rw [get?_eq_get h, get_ofFn] - . simp at h; simp [ofFnNthVal, h] + · simp at h; simp [ofFnNthVal, h] else by rw [ofFnNthVal, dif_neg] <;> simpa using h @@ -117,8 +117,8 @@ theorem ofFn_zero (f : Fin 0 → α) : ofFn f = [] := theorem ofFn_succ {n} (f : Fin (succ n) → α) : ofFn f = f 0 :: ofFn fun i => f i.succ := ext_get (by simp) (fun i hi₁ hi₂ => by cases i - . simp - . simp) + · simp + · simp) #align list.of_fn_succ List.ofFn_succ theorem ofFn_succ' {n} (f : Fin (succ n) → α) : diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index 733e5d16411fd..4f22a04850068 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -143,9 +143,9 @@ theorem Pairwise.forall_of_forall (H : Symmetric R) (H₁ : ∀ x ∈ l, R x x) theorem Pairwise.forall (hR : Symmetric R) (hl : l.Pairwise R) : ∀ ⦃a⦄, a ∈ l → ∀ ⦃b⦄, b ∈ l → a ≠ b → R a b := by apply Pairwise.forall_of_forall - . exact fun a b h hne => hR (h hne.symm) - . exact fun _ _ hx => (hx rfl).elim - . exact hl.imp (@fun a b h _ => by exact h) + · exact fun a b h hne => hR (h hne.symm) + · exact fun _ _ hx => (hx rfl).elim + · exact hl.imp (@fun a b h _ => by exact h) #align list.pairwise.forall List.Pairwise.forall theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x ∈ l }.Pairwise R := @@ -329,7 +329,7 @@ theorem pairwise_iff_get : ∀ {l : List α}, Pairwise R l ↔ have := H ⟨0, show 0 < (a::l).length from Nat.succ_pos _⟩ ⟨n.succ, Nat.succ_lt_succ n.2⟩ (Nat.succ_pos n) simpa - . simpa using H i.succ j.succ (show i.1.succ < j.1.succ from Nat.succ_lt_succ hij) + · simpa using H i.succ j.succ (show i.1.succ < j.1.succ from Nat.succ_lt_succ hij) set_option linter.deprecated false in @[deprecated pairwise_iff_get] diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index d39613c28293d..03e9ef36f25b1 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -567,10 +567,10 @@ theorem Perm.prod_eq' [M : Monoid α] {l₁ l₂ : List α} (h : l₁ ~ l₂) (h l₁.prod = l₂.prod := by refine h.foldl_eq' ?_ _ apply Pairwise.forall_of_forall - . intro x y h z + · intro x y h z exact (h z).symm - . intros; rfl - . apply hc.imp + · intros; rfl + · apply hc.imp intro a b h z rw [mul_assoc z, mul_assoc z, h] #align list.perm.prod_eq' List.Perm.prod_eq' @@ -1190,8 +1190,8 @@ theorem Perm.dropSlice_inter {α} [DecidableEq α] {xs ys : List α} (n m : ℕ) have : n ≤ n + m := Nat.le_add_right _ _ have h₂ := h.nodup_iff.2 h' apply Perm.trans _ (Perm.inter_append _).symm - . exact Perm.append (Perm.take_inter _ h h') (Perm.drop_inter _ h h') - . exact disjoint_take_drop h₂ this + · exact Perm.append (Perm.take_inter _ h h') (Perm.drop_inter _ h h') + · exact disjoint_take_drop h₂ this #align list.perm.slice_inter List.Perm.dropSlice_inter -- enumerating permutations @@ -1426,11 +1426,9 @@ theorem nodup_permutations'Aux_iff {s : List α} {x : α} : Nodup (permutations' rw [length_insertNth _ _ hk.le, length_insertNth _ _ (Nat.succ_le_of_lt hk)] refine' ext_nthLe hl fun n hn hn' => _ rcases lt_trichotomy n k with (H | rfl | H) - · - rw [nthLe_insertNth_of_lt _ _ _ _ H (H.trans hk), + · rw [nthLe_insertNth_of_lt _ _ _ _ H (H.trans hk), nthLe_insertNth_of_lt _ _ _ _ (H.trans (Nat.lt_succ_self _))] - · - rw [nthLe_insertNth_self _ _ _ hk.le, nthLe_insertNth_of_lt _ _ _ _ (Nat.lt_succ_self _) hk, + · rw [nthLe_insertNth_self _ _ _ hk.le, nthLe_insertNth_of_lt _ _ _ _ (Nat.lt_succ_self _) hk, hk'] · rcases(Nat.succ_le_of_lt H).eq_or_lt with (rfl | H') · rw [nthLe_insertNth_self _ _ _ (Nat.succ_le_of_lt hk)] diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 3cf0366061dcf..5e136e61e4afd 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -93,12 +93,12 @@ theorem map_permutationsAux2' {α β α' β'} (g : α → α') (g' : β → β') map g' (permutationsAux2 t ts r ys f).2 = (permutationsAux2 (g t) (map g ts) (map g' r) (map g ys) f').2 := by induction' ys with ys_hd _ ys_ih generalizing f f' - . simp - . simp only [map, permutationsAux2_snd_cons, cons_append, cons.injEq] + · 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 - . intro a; apply H + · 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. -/ diff --git a/Mathlib/Data/List/Sigma.lean b/Mathlib/Data/List/Sigma.lean index 0c9ac517b3e27..5c893bd68ce8e 100644 --- a/Mathlib/Data/List/Sigma.lean +++ b/Mathlib/Data/List/Sigma.lean @@ -285,8 +285,8 @@ theorem head?_lookupAll (a : α) : ∀ l : List (Sigma β), head? (lookupAll a l | [] => by simp | ⟨a', b⟩ :: l => by by_cases h : a = a' - . subst h; simp - . rw [lookupAll_cons_ne, dlookup_cons_ne, head?_lookupAll a l] <;> assumption + · subst h; simp + · rw [lookupAll_cons_ne, dlookup_cons_ne, head?_lookupAll a l] <;> assumption #align list.head_lookup_all List.head?_lookupAll theorem mem_lookupAll {a : α} {b : β a} : @@ -296,7 +296,7 @@ theorem mem_lookupAll {a : α} {b : β a} : by_cases h : a = a' · subst h simp [*, mem_lookupAll] - . simp [*, mem_lookupAll] + · simp [*, mem_lookupAll] #align list.mem_lookup_all List.mem_lookupAll theorem lookupAll_sublist (a : α) : ∀ l : List (Sigma β), (lookupAll a l).map (Sigma.mk a) <+ l @@ -348,7 +348,7 @@ theorem kreplace_of_forall_not (a : α) (b : β a) {l : List (Sigma β)} rintro ⟨a', b'⟩ h; dsimp; split_ifs · subst a' exact H _ h - . rfl + · rfl #align list.kreplace_of_forall_not List.kreplace_of_forall_not theorem kreplace_self {a : α} {b : β a} {l : List (Sigma β)} (nd : NodupKeys l) diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 725523c93acc0..0efea1358b023 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -65,8 +65,8 @@ theorem sublists'_eq_sublists'Aux (l : List α) : simp only [sublists', sublists'Aux_eq_array_foldl] dsimp only rw [← List.foldr_hom Array.toList] - . rfl - . intros _ _; congr <;> simp + · rfl + · intros _ _; congr <;> simp theorem sublists'Aux_eq_map (a : α) (r₁ : List (List α)) : ∀ (r₂ : List (List α)), sublists'Aux a r₁ r₂ = r₂ ++ map (cons a) r₁ := @@ -143,8 +143,8 @@ theorem sublists_eq_sublistsAux (l : List α) : sublists l = l.foldr sublistsAux [[]] := by simp only [sublists, sublistsAux_eq_array_foldl, Array.foldr_eq_foldr_data] rw [← foldr_hom Array.toList] - . rfl - . intros _ _; congr <;> simp + · rfl + · intros _ _; congr <;> simp #noalign list.sublists_aux₁_eq_sublists_aux #noalign list.sublists_aux_cons_eq_sublists_aux₁ @@ -160,8 +160,8 @@ theorem sublists_append (l₁ l₂ : List α) : sublists (l₁ ++ l₂) = (sublists l₂) >>= (fun x => (sublists l₁).map (. ++ x)) := by simp only [sublists_eq_sublistsAux, foldr_append, sublistsAux_eq_bind] induction l₁ - . case nil => simp - . case cons a l₁ ih => + · case nil => simp + · case cons a l₁ ih => rw [foldr_cons, ih] simp [List.bind, join_join, Function.comp] #align list.sublists_append List.sublists_append @@ -430,10 +430,10 @@ theorem sublists_cons_perm_append (a : α) (l : List α) : theorem revzip_sublists (l : List α) : ∀ l₁ l₂, (l₁, l₂) ∈ revzip l.sublists → l₁ ++ l₂ ~ l := by rw [revzip] induction' l using List.reverseRecOn with l' a ih - . intro l₁ l₂ h + · intro l₁ l₂ h simp at h simp [h] - . intro l₁ l₂ h + · intro l₁ l₂ h rw [sublists_concat, reverse_append, zip_append, ← map_reverse, zip_map_right, zip_map_left] at * <;> [skip; simp] simp only [Prod.mk.inj_iff, mem_map, mem_append, Prod.map_mk, Prod.exists] at h diff --git a/Mathlib/Data/MvPolynomial/Rename.lean b/Mathlib/Data/MvPolynomial/Rename.lean index 204d0bdd03bfe..68263de8a2a4d 100644 --- a/Mathlib/Data/MvPolynomial/Rename.lean +++ b/Mathlib/Data/MvPolynomial/Rename.lean @@ -236,8 +236,7 @@ theorem exists_finset_rename (p : MvPolynomial σ R) : exact ⟨∅, C r, by rw [rename_C]⟩ · rintro p q ⟨s, p, rfl⟩ ⟨t, q, rfl⟩ refine' ⟨s ∪ t, ⟨_, _⟩⟩ - · - refine' rename (Subtype.map id _) p + rename (Subtype.map id _) q <;> + · refine' rename (Subtype.map id _) p + rename (Subtype.map id _) q <;> simp (config := { contextual := true }) only [id.def, true_or_iff, or_true_iff, Finset.mem_union, forall_true_iff] · simp only [rename_rename, AlgHom.map_add] diff --git a/Mathlib/Data/Nat/Choose/Basic.lean b/Mathlib/Data/Nat/Choose/Basic.lean index 63c13f390773a..a04380cc012b9 100644 --- a/Mathlib/Data/Nat/Choose/Basic.lean +++ b/Mathlib/Data/Nat/Choose/Basic.lean @@ -100,7 +100,7 @@ theorem triangle_succ (n : ℕ) : (n + 1) * (n + 1 - 1) / 2 = n * (n - 1) / 2 + /-- `choose n 2` is the `n`-th triangle number. -/ theorem choose_two_right (n : ℕ) : choose n 2 = n * (n - 1) / 2 := by induction' n with n ih - . simp + · simp · rw [triangle_succ n, choose, ih] simp [add_comm] #align nat.choose_two_right Nat.choose_two_right @@ -187,7 +187,7 @@ theorem factorial_mul_factorial_dvd_factorial {n k : ℕ} (hk : k ≤ n) : k ! * theorem factorial_mul_factorial_dvd_factorial_add (i j : ℕ) : i ! * j ! ∣ (i + j)! := by suffices : i ! * (i + j - i) ! ∣ (i + j)! - . rwa [add_tsub_cancel_left i j] at this + · rwa [add_tsub_cancel_left i j] at this exact factorial_mul_factorial_dvd_factorial (Nat.le_add_right _ _) #align nat.factorial_mul_factorial_dvd_factorial_add Nat.factorial_mul_factorial_dvd_factorial_add @@ -199,7 +199,7 @@ theorem choose_symm {n k : ℕ} (hk : k ≤ n) : choose n (n - k) = choose n k : theorem choose_symm_of_eq_add {n a b : ℕ} (h : n = a + b) : Nat.choose n a = Nat.choose n b := by suffices : choose n (n - b) = choose n b - . rw [h, add_tsub_cancel_right] at this; rwa [h] + · rw [h, add_tsub_cancel_right] at this; rwa [h] exact choose_symm (h ▸ le_add_left _ _) #align nat.choose_symm_of_eq_add Nat.choose_symm_of_eq_add diff --git a/Mathlib/Data/Nat/Digits.lean b/Mathlib/Data/Nat/Digits.lean index 388f9d501c161..2865b9f78c32b 100644 --- a/Mathlib/Data/Nat/Digits.lean +++ b/Mathlib/Data/Nat/Digits.lean @@ -310,7 +310,7 @@ theorem digits_eq_cons_digits_div {b n : ℕ} (h : 1 < b) (w : n ≠ 0) : · norm_num at h rcases n with (_ | n) · norm_num at w - . simp only [digits_add_two_add_one, ne_eq] + · simp only [digits_add_two_add_one, ne_eq] #align nat.digits_eq_cons_digits_div Nat.digits_eq_cons_digits_div theorem digits_getLast {b : ℕ} (m : ℕ) (h : 1 < b) (p q) : @@ -381,8 +381,8 @@ theorem digits_lt_base' {b m : ℕ} : ∀ {d}, d ∈ digits (b + 2) m → d < b -- Porting note: Previous code (single line) contained linarith. -- . exact IH _ (Nat.div_lt_self (Nat.succ_pos _) (by linarith)) hd · apply IH ((n + 1) / (b + 2)) - . apply Nat.div_lt_self <;> simp - . assumption + · apply Nat.div_lt_self <;> simp + · assumption #align nat.digits_lt_base' Nat.digits_lt_base' /-- The digits in the base b expansion of n are all less than b, if b ≥ 2 -/ diff --git a/Mathlib/Data/Nat/ModEq.lean b/Mathlib/Data/Nat/ModEq.lean index 5e8a057f944d6..137c354d51688 100644 --- a/Mathlib/Data/Nat/ModEq.lean +++ b/Mathlib/Data/Nat/ModEq.lean @@ -293,8 +293,7 @@ lemma cancel_left_div_gcd (hm : 0 < m) (h : c * a ≡ c * b [MOD m]) : a ≡ b rw [mul_sub] exact modEq_iff_dvd.mp h show Int.gcd (m / d) (c / d) = 1 - · - simp only [← Int.coe_nat_div, Int.coe_nat_gcd (m / d) (c / d), gcd_div hmd hcd, + · simp only [← Int.coe_nat_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 diff --git a/Mathlib/Data/Ordmap/Ordnode.lean b/Mathlib/Data/Ordmap/Ordnode.lean index 7a22c17beb925..fd1929c2caefd 100644 --- a/Mathlib/Data/Ordmap/Ordnode.lean +++ b/Mathlib/Data/Ordmap/Ordnode.lean @@ -208,8 +208,7 @@ def balanceL (l : Ordnode α) (x : α) (r : Ordnode α) : Ordnode α := by · exact node 3 (ι lx) lrx ι x · cases' id lr with lrs lrl lrx lrr · exact node 3 ll lx ι x - · - exact + · exact if lrs < ratio * lls then node (ls + 1) ll lx (node (lrs + 1) lr x nil) else node (ls + 1) (node (lls + size lrl + 1) ll lx lrl) lrx @@ -245,8 +244,7 @@ def balanceR (l : Ordnode α) (x : α) (r : Ordnode α) : Ordnode α := by · exact node 3 (ι x) rlx ι rx · cases' id rl with rls rll rlx rlr · exact node 3 (ι x) rx rr - · - exact + · exact if rls < ratio * rrs then node (rs + 1) (node (rls + 1) nil x rl) rx rr else node (rs + 1) (node (size rll + 1) nil x rll) rlx diff --git a/Mathlib/Data/PEquiv.lean b/Mathlib/Data/PEquiv.lean index d4434d980303e..80070ce44ebc7 100644 --- a/Mathlib/Data/PEquiv.lean +++ b/Mathlib/Data/PEquiv.lean @@ -335,12 +335,12 @@ def single (a : α) (b : β) : inv x y := by dsimp only split_ifs with h1 h2 - . simp [*] - . simp only [mem_def, some.injEq, iff_false] at * + · simp [*] + · simp only [mem_def, some.injEq, iff_false] at * exact Ne.symm h2 - . simp only [mem_def, some.injEq, false_iff] at * + · simp only [mem_def, some.injEq, false_iff] at * exact Ne.symm h1 - . simp + · simp #align pequiv.single PEquiv.single theorem mem_single (a : α) (b : β) : b ∈ single a b a := diff --git a/Mathlib/Data/PFunctor/Multivariate/Basic.lean b/Mathlib/Data/PFunctor/Multivariate/Basic.lean index 7e64e04d16d62..43023fc6d8f83 100644 --- a/Mathlib/Data/PFunctor/Multivariate/Basic.lean +++ b/Mathlib/Data/PFunctor/Multivariate/Basic.lean @@ -174,7 +174,7 @@ theorem liftP_iff' {α : TypeVec n} (p : ∀ ⦃i⦄, α i → Prop) (a : P.A) ( simp only [liftP_iff, Sigma.mk.inj_iff]; constructor · rintro ⟨_, _, ⟨⟩, _⟩ assumption - . intro + · intro repeat' first |constructor|assumption #align mvpfunctor.liftp_iff' MvPFunctor.liftP_iff' diff --git a/Mathlib/Data/PNat/Xgcd.lean b/Mathlib/Data/PNat/Xgcd.lean index cf403fb3157cf..d5ac7bacc354d 100644 --- a/Mathlib/Data/PNat/Xgcd.lean +++ b/Mathlib/Data/PNat/Xgcd.lean @@ -155,8 +155,8 @@ theorem isSpecial_iff : u.IsSpecial ↔ u.IsSpecial' := by let ⟨wp, x, y, zp, ap, bp⟩ := u constructor <;> intro h <;> simp [w, z, succPNat] at * <;> simp only [← coe_inj, mul_coe, mk_coe] at * - . simp_all [← h, Nat.mul, Nat.succ_eq_add_one]; ring - . simp [Nat.succ_eq_add_one, Nat.mul_add, Nat.add_mul, ← Nat.add_assoc] at h; rw [← h]; ring + · simp_all [← h, Nat.mul, Nat.succ_eq_add_one]; ring + · simp [Nat.succ_eq_add_one, Nat.mul_add, Nat.add_mul, ← Nat.add_assoc] at h; rw [← h]; ring -- Porting note: Old code has been removed as it was much more longer. #align pnat.xgcd_type.is_special_iff PNat.XgcdType.isSpecial_iff diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 7f2edcdc55c9c..fd64cdf7acf08 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -343,8 +343,8 @@ theorem ofOption_dom {α} : ∀ o : Option α, (ofOption o).Dom ↔ o.isSome theorem ofOption_eq_get {α} (o : Option α) : ofOption o = ⟨_, @Option.get _ o⟩ := Part.ext' (ofOption_dom o) fun h₁ h₂ => by cases o - . simp at h₂ - . rfl + · simp at h₂ + · rfl #align part.of_option_eq_get Part.ofOption_eq_get instance : Coe (Option α) (Part α) := diff --git a/Mathlib/Data/Polynomial/RingDivision.lean b/Mathlib/Data/Polynomial/RingDivision.lean index 015ee21974f67..49bac25287fa8 100644 --- a/Mathlib/Data/Polynomial/RingDivision.lean +++ b/Mathlib/Data/Polynomial/RingDivision.lean @@ -286,8 +286,7 @@ theorem irreducible_of_monic (hp : p.Monic) (hp1 : p ≠ 1) : (isUnit_of_mul_eq_one g _)⟩⟩ · rwa [Monic, leadingCoeff_mul, leadingCoeff_C, ← leadingCoeff_mul, mul_comm, ← hfg, ← Monic] · rwa [Monic, leadingCoeff_mul, leadingCoeff_C, ← leadingCoeff_mul, ← hfg, ← Monic] - · - rw [mul_mul_mul_comm, ← C_mul, ← leadingCoeff_mul, ← hfg, hp.leadingCoeff, C_1, mul_one, + · rw [mul_mul_mul_comm, ← C_mul, ← leadingCoeff_mul, ← hfg, hp.leadingCoeff, C_1, mul_one, mul_comm, ← hfg] #align polynomial.irreducible_of_monic Polynomial.irreducible_of_monic @@ -734,8 +733,7 @@ theorem roots_pow (p : R[X]) (n : ℕ) : (p ^ n).roots = n • p.roots := by · rw [pow_zero, roots_one, Nat.zero_eq, zero_smul, empty_eq_zero] · rcases eq_or_ne p 0 with (rfl | hp) · rw [zero_pow n.succ_pos, roots_zero, smul_zero] - · - rw [pow_succ', roots_mul (mul_ne_zero (pow_ne_zero _ hp) hp), ihn, Nat.succ_eq_add_one, + · rw [pow_succ', roots_mul (mul_ne_zero (pow_ne_zero _ hp) hp), ihn, Nat.succ_eq_add_one, add_smul, one_smul] #align polynomial.roots_pow Polynomial.roots_pow @@ -897,8 +895,7 @@ theorem comp_eq_zero_iff : p.comp q = 0 ↔ p = 0 ∨ p.eval (q.coeff 0) = 0 ∧ exact Or.inl (key.trans h) · rw [key, comp_C, C_eq_zero] at h exact Or.inr ⟨h, key⟩ - · - exact fun h => + · exact fun h => Or.rec (fun h => by rw [h, zero_comp]) (fun h => by rw [h.2, comp_C, h.1, C_0]) h #align polynomial.comp_eq_zero_iff Polynomial.comp_eq_zero_iff diff --git a/Mathlib/Data/Rat/Floor.lean b/Mathlib/Data/Rat/Floor.lean index 04af58e1d2f69..35e1d2d42e70d 100644 --- a/Mathlib/Data/Rat/Floor.lean +++ b/Mathlib/Data/Rat/Floor.lean @@ -36,8 +36,8 @@ variable {α : Type _} [LinearOrderedField α] [FloorRing α] protected theorem floor_def' (a : ℚ) : a.floor = a.num / a.den := by rw [Rat.floor] split - . next h => simp [h] - . next => rfl + · next h => simp [h] + · next => rfl protected theorem le_floor {z : ℤ} : ∀ {r : ℚ}, z ≤ Rat.floor r ↔ (z : ℚ) ≤ r | ⟨n, d, h, c⟩ => by diff --git a/Mathlib/Data/Seq/Computation.lean b/Mathlib/Data/Seq/Computation.lean index c72e51785363e..43bf76562f02e 100644 --- a/Mathlib/Data/Seq/Computation.lean +++ b/Mathlib/Data/Seq/Computation.lean @@ -58,8 +58,8 @@ instance : CoeTC α (Computation α) := def think (c : Computation α) : Computation α := ⟨Stream'.cons none c.1, fun n a h => by cases' n with n - . contradiction - . exact c.2 h⟩ + · contradiction + · exact c.2 h⟩ #align computation.think Computation.think /-- `thinkN c n` is the computation that delays for `n` ticks and then performs @@ -304,8 +304,8 @@ theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s have h := bisim r; revert r h apply recOn s _ _ <;> intro r' <;> apply recOn s' _ _ <;> intro a' r h · constructor <;> dsimp at h - . rw [h] - . rw [h] at r + · rw [h] + · rw [h] at r rw [tail_pure, tail_pure,h] assumption · rw [destruct_pure, destruct_think] at h @@ -653,7 +653,7 @@ def map (f : α → β) : Computation α → Computation β ⟨s.map fun o => Option.casesOn o none (some ∘ f), fun n b => by dsimp [Stream'.map, Stream'.nth] induction' e : s n with a <;> intro h - . contradiction + · contradiction · rw [al e]; exact h⟩ #align computation.map Computation.map @@ -761,13 +761,13 @@ theorem bind_pure (f : α → β) (s) : bind s (pure ∘ f) = map f s := by @[simp] theorem bind_pure' (s : Computation α) : bind s pure = s := by apply eq_of_bisim fun c₁ c₂ => c₁ = c₂ ∨ ∃ s, c₁ = bind s (pure) ∧ c₂ = s - . intro c₁ c₂ h + · intro c₁ c₂ h exact match c₁, c₂, h with | _, c₂, Or.inl (Eq.refl _) => by cases' destruct c₂ with b cb <;> simp | _, _, Or.inr ⟨s, rfl, rfl⟩ => by apply recOn s <;> intro s <;> simp - . exact Or.inr ⟨s, rfl, rfl⟩ + · exact Or.inr ⟨s, rfl, rfl⟩ #align computation.bind_ret' Computation.bind_pure' @[simp] diff --git a/Mathlib/Data/Seq/Seq.lean b/Mathlib/Data/Seq/Seq.lean index 9d81450dd176d..e8c0024852a15 100644 --- a/Mathlib/Data/Seq/Seq.lean +++ b/Mathlib/Data/Seq/Seq.lean @@ -394,8 +394,8 @@ theorem eq_of_bisim (bisim : IsBisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s apply recOn s _ _ <;> apply recOn s' _ _ · intro r _ constructor - . rfl - . assumption + · rfl + · assumption · intro x s _ this rw [destruct_nil, destruct_cons] at this exact False.elim this @@ -528,7 +528,7 @@ def map (f : α → β) : Seq α → Seq β induction' e : s n with e <;> intro · rw [al e] assumption - . contradiction⟩ + · contradiction⟩ #align stream.seq.map Stream'.Seq.map /-- Flatten a sequence of sequences. (It is required that the @@ -862,8 +862,8 @@ theorem exists_of_mem_map {f} {b : β} : ∀ {s : Seq α}, b ∈ map f s → ∃ | ⟨g, al⟩ => let ⟨o, om, oe⟩ := @Stream'.exists_of_mem_map _ _ (Option.map f) (some b) g h cases' o with a - . injection oe - . injection oe with h' ; exact ⟨a, om, h'⟩ + · injection oe + · injection oe with h' ; exact ⟨a, om, h'⟩ #align stream.seq.exists_of_mem_map Stream'.Seq.exists_of_mem_map theorem of_mem_append {s₁ s₂ : Seq α} {a : α} (h : a ∈ append s₁ s₂) : a ∈ s₁ ∨ a ∈ s₂ := by diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 908470e20d926..d8b3d033ec5aa 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -2188,8 +2188,8 @@ theorem powerset_singleton (x : α) : 𝒫({x} : Set α) = {∅, {x}} := by theorem mem_dite (p : Prop) [Decidable p] (s : p → Set α) (t : ¬ p → Set α) (x : α) : (x ∈ if h : p then s h else t h) ↔ (∀ h : p, x ∈ s h) ∧ ∀ h : ¬p, x ∈ t h := by split_ifs with hp - . exact ⟨fun hx => ⟨fun _ => hx, fun hnp => (hnp hp).elim⟩, fun hx => hx.1 hp⟩ - . exact ⟨fun hx => ⟨fun h => (hp h).elim, fun _ => hx⟩, fun hx => hx.2 hp⟩ + · exact ⟨fun hx => ⟨fun _ => hx, fun hnp => (hnp hp).elim⟩, fun hx => hx.1 hp⟩ + · exact ⟨fun hx => ⟨fun h => (hp h).elim, fun _ => hx⟩, fun hx => hx.2 hp⟩ --Porting note: Old proof was `split_ifs; simp [h]` theorem mem_dite_univ_right (p : Prop) [Decidable p] (t : p → Set α) (x : α) : diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index a5e6514abb914..db23b74022a07 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -1014,11 +1014,11 @@ theorem range_subtype_map {p : α → Prop} {q : β → Prop} (f : α → β) (h ext ⟨x, hx⟩ rw [mem_preimage, mem_range, mem_image, Subtype.exists, Subtype.coe_mk] apply Iff.intro - . rintro ⟨a, b, hab⟩ + · rintro ⟨a, b, hab⟩ rw [Subtype.map, Subtype.mk.injEq] at hab use a trivial - . rintro ⟨a, b, hab⟩ + · rintro ⟨a, b, hab⟩ use a use b rw [Subtype.map, Subtype.mk.injEq] @@ -1135,12 +1135,12 @@ theorem range_inclusion (h : s ⊆ t) : range (inclusion h) = { x : t | (x : α) ext ⟨x, hx⟩ -- Porting note: `simp [inclusion]` doesn't solve goal apply Iff.intro - . rw [mem_range] + · rw [mem_range] rintro ⟨a, ha⟩ rw [inclusion, Subtype.mk.injEq] at ha rw [mem_setOf, Subtype.coe_mk, ← ha] exact Subtype.coe_prop _ - . rw [mem_setOf, Subtype.coe_mk, mem_range] + · rw [mem_setOf, Subtype.coe_mk, mem_range] intro hx' use ⟨x, hx'⟩ trivial diff --git a/Mathlib/Data/Set/Opposite.lean b/Mathlib/Data/Set/Opposite.lean index 84619f0bb57f8..6296f90eae227 100644 --- a/Mathlib/Data/Set/Opposite.lean +++ b/Mathlib/Data/Set/Opposite.lean @@ -80,32 +80,32 @@ def opEquiv : Set α ≃ Set αᵒᵖ := theorem singleton_op (x : α) : ({x} : Set α).op = {op x} := by ext constructor - . apply unop_injective - . apply op_injective + · apply unop_injective + · apply op_injective #align set.singleton_op Set.singleton_op @[simp] theorem singleton_unop (x : αᵒᵖ) : ({x} : Set αᵒᵖ).unop = {unop x} := by ext constructor - . apply op_injective - . apply unop_injective + · apply op_injective + · apply unop_injective #align set.singleton_unop Set.singleton_unop @[simp 1100] theorem singleton_op_unop (x : α) : ({op x} : Set αᵒᵖ).unop = {x} := by ext constructor - . apply op_injective - . apply unop_injective + · apply op_injective + · apply unop_injective #align set.singleton_op_unop Set.singleton_op_unop @[simp 1100] theorem singleton_unop_op (x : αᵒᵖ) : ({unop x} : Set α).op = {x} := by ext constructor - . apply unop_injective - . apply op_injective + · apply unop_injective + · apply op_injective #align set.singleton_unop_op Set.singleton_unop_op end Set diff --git a/Mathlib/Data/TypeVec.lean b/Mathlib/Data/TypeVec.lean index 1c67cbe77d0f4..c4f6738d49c48 100644 --- a/Mathlib/Data/TypeVec.lean +++ b/Mathlib/Data/TypeVec.lean @@ -175,8 +175,8 @@ theorem eq_of_drop_last_eq {α β : TypeVec (n + 1)} {f g : α ⟹ β} (h₀ : d -- porting note: FIXME: congr_fun h₀ <;> ext1 ⟨⟩ <;> apply_assumption refine funext (fun x => ?_) cases x - . apply h₁ - . apply congr_fun h₀ + · apply h₁ + · apply congr_fun h₀ #align typevec.eq_of_drop_last_eq TypeVec.eq_of_drop_last_eq @[simp] @@ -675,9 +675,9 @@ theorem append_prod_appendFun {n} {α α' β β' : TypeVec.{u} n} {φ φ' ψ ψ' ((f₀ ⊗' g₀) ::: (_root_.Prod.map f₁ g₁)) = ((f₀ ::: f₁) ⊗' (g₀ ::: g₁)) := by ext i a cases i - . cases a + · cases a rfl - . rfl + · rfl #align typevec.append_prod_append_fun TypeVec.append_prod_appendFun end Liftp' @@ -782,8 +782,8 @@ theorem prod_map_id {α β : TypeVec n} : (@TypeVec.id _ α ⊗' @TypeVec.id _ theorem subtypeVal_diagSub {α : TypeVec n} : subtypeVal (repeatEq α) ⊚ diagSub = prod.diag := by ext i x induction' i with _ _ _ i_ih - . simp [comp, diagSub, subtypeVal, prod.diag] - . simp [prod.diag] + · simp [comp, diagSub, subtypeVal, prod.diag] + · simp [prod.diag] simp [comp, diagSub, subtypeVal, prod.diag] at * apply @i_ih (drop _) #align typevec.subtype_val_diag_sub TypeVec.subtypeVal_diagSub diff --git a/Mathlib/Data/Vector/Basic.lean b/Mathlib/Data/Vector/Basic.lean index 26b1593be0fcf..8806d69a8321e 100644 --- a/Mathlib/Data/Vector/Basic.lean +++ b/Mathlib/Data/Vector/Basic.lean @@ -581,14 +581,14 @@ theorem removeNth_insertNth' {v : Vector α (n + 1)} : · rcases Nat.exists_eq_succ_of_ne_zero (Nat.pos_iff_ne_zero.1 (lt_of_le_of_lt (Nat.zero_le _) hij)) with ⟨j, rfl⟩ rw [← List.insertNth_removeNth_of_ge] - . simp; rfl - . simpa - . simpa [Nat.lt_succ_iff] using hij + · simp; rfl + · simpa + · simpa [Nat.lt_succ_iff] using hij · dsimp rw [← List.insertNth_removeNth_of_le i j _ _ _] - . rfl - . simpa - . simpa [not_lt] using hij + · rfl + · simpa + · simpa [not_lt] using hij #align vector.remove_nth_insert_nth' Vector.removeNth_insertNth' theorem insertNth_comm (a b : α) (i j : Fin (n + 1)) (h : i ≤ j) : @@ -631,8 +631,8 @@ theorem get_set_of_ne {v : Vector α n} {i j : Fin n} (h : i ≠ j) (a : α) : cases v; cases i; cases j simp [Vector.set, Vector.get_eq_get, List.get_set_of_ne (Fin.vne_of_ne h)] rw [List.get_set_of_ne] - . rfl - . simpa using h + · rfl + · simpa using h #align vector.nth_update_nth_of_ne Vector.get_set_of_ne theorem get_set_eq_if {v : Vector α n} {i j : Fin n} (a : α) : @@ -713,8 +713,8 @@ protected theorem comp_traverse (f : β → F γ) (g : α → G β) (x : Vector 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 - . rw [Vector.traverse_def, ih] + · rfl + · rw [Vector.traverse_def, ih] simp [functor_norm, (. ∘ .)] #align vector.comp_traverse Vector.comp_traverse @@ -728,8 +728,8 @@ variable (η : ApplicativeTransformation F G) protected theorem naturality {α β : Type _} (f : α → F β) (x : Vector α n) : η (x.traverse f) = x.traverse (@η _ ∘ f) := by induction' x using Vector.inductionOn with n x xs ih - . simp! [functor_norm, cast, η.preserves_pure] - . rw [Vector.traverse_def, Vector.traverse_def, ← ih, η.preserves_seq, η.preserves_map] + · simp! [functor_norm, cast, η.preserves_pure] + · rw [Vector.traverse_def, Vector.traverse_def, ← ih, η.preserves_seq, η.preserves_map] rfl #align vector.naturality Vector.naturality @@ -807,8 +807,8 @@ theorem get_map₂ (v₁ : Vector α n) (v₂ : Vector β n) (f : α → β → case cons x xs y ys ih => rw [map₂_cons] cases i using Fin.cases - . simp only [get_zero, head_cons] - . simp only [get_cons_succ, ih] + · simp only [get_zero, head_cons] + · simp only [get_cons_succ, ih] diff --git a/Mathlib/Data/Vector/Snoc.lean b/Mathlib/Data/Vector/Snoc.lean index e4af114c0d6be..39cd695bad044 100644 --- a/Mathlib/Data/Vector/Snoc.lean +++ b/Mathlib/Data/Vector/Snoc.lean @@ -141,8 +141,8 @@ theorem mapAccumr_snoc : let r := mapAccumr f xs q.1 (r.1, r.2.snoc q.2) := by induction xs using Vector.inductionOn - . rfl - . simp[*] + · rfl + · simp[*] variable (ys : Vector β n) diff --git a/Mathlib/Data/ZMod/Basic.lean b/Mathlib/Data/ZMod/Basic.lean index e1a564628991b..6eeb68c1dce0a 100644 --- a/Mathlib/Data/ZMod/Basic.lean +++ b/Mathlib/Data/ZMod/Basic.lean @@ -161,8 +161,8 @@ instance (priority := 900) (n : ℕ) : CoeTC (ZMod n) R := theorem cast_zero : ((0 : ZMod n) : R) = 0 := by delta ZMod.cast cases n - . exact Int.cast_zero - . simp + · exact Int.cast_zero + · simp #align zmod.cast_zero ZMod.cast_zero theorem cast_eq_val [NeZero n] (a : ZMod n) : (a : R) = a.val := by @@ -176,15 +176,15 @@ variable {S : Type _} [AddGroupWithOne S] @[simp] theorem _root_.Prod.fst_zmod_cast (a : ZMod n) : (a : R × S).fst = a := by cases n - . rfl - . simp [ZMod.cast] + · rfl + · simp [ZMod.cast] #align prod.fst_zmod_cast Prod.fst_zmod_cast @[simp] theorem _root_.Prod.snd_zmod_cast (a : ZMod n) : (a : R × S).snd = a := by cases n - . rfl - . simp [ZMod.cast] + · rfl + · simp [ZMod.cast] #align prod.snd_zmod_cast Prod.snd_zmod_cast end @@ -750,14 +750,14 @@ def chineseRemainder {m n : ℕ} (h : m.coprime n) : ZMod (m * n) ≃+* ZMod m have inv : Function.LeftInverse inv_fun to_fun ∧ Function.RightInverse inv_fun to_fun := if hmn0 : m * n = 0 then by rcases h.eq_of_mul_eq_zero hmn0 with (⟨rfl, rfl⟩ | ⟨rfl, rfl⟩) - . constructor - . intro x; rfl - . rintro ⟨x, y⟩ + · constructor + · intro x; rfl + · rintro ⟨x, y⟩ fin_cases y simp [castHom, Prod.ext_iff] - . constructor - . intro x; rfl - . rintro ⟨x, y⟩ + · constructor + · intro x; rfl + · rintro ⟨x, y⟩ fin_cases x simp [castHom, Prod.ext_iff] else by @@ -1162,8 +1162,8 @@ instance subsingleton_ringEquiv [Semiring R] : Subsingleton (ZMod n ≃+* R) := @[simp] theorem ringHom_map_cast [Ring R] (f : R →+* ZMod n) (k : ZMod n) : f k = k := by cases n - . dsimp [ZMod, ZMod.cast] at f k ⊢; simp - . dsimp [ZMod, ZMod.cast] at f k ⊢ + · dsimp [ZMod, ZMod.cast] at f k ⊢; simp + · dsimp [ZMod, ZMod.cast] at f k ⊢ erw [map_natCast, Fin.cast_val_eq_self] #align zmod.ring_hom_map_cast ZMod.ringHom_map_cast diff --git a/Mathlib/FieldTheory/Normal.lean b/Mathlib/FieldTheory/Normal.lean index 144aab95b4e51..daf45d3754534 100644 --- a/Mathlib/FieldTheory/Normal.lean +++ b/Mathlib/FieldTheory/Normal.lean @@ -521,7 +521,7 @@ theorem restrictScalars_eq_iSup_adjoin [h : Normal F L] : apply PowerBasis.lift_gen change aeval y (minpoly F (AdjoinSimple.gen F x)) = 0 suffices : minpoly F (AdjoinSimple.gen F x) = minpoly F x - . exact this ▸ aeval_eq_zero_of_mem_rootSet (Multiset.mem_toFinset.mpr hy) + · exact this ▸ aeval_eq_zero_of_mem_rootSet (Multiset.mem_toFinset.mpr hy) exact minpoly_gen ((isIntegral_algebraMap_iff (algebraMap K L).injective).mp (h.isIntegral (algebraMap K L x))) diff --git a/Mathlib/FieldTheory/RatFunc.lean b/Mathlib/FieldTheory/RatFunc.lean index 4bb8995af21ad..6009bdc0dc1d3 100644 --- a/Mathlib/FieldTheory/RatFunc.lean +++ b/Mathlib/FieldTheory/RatFunc.lean @@ -253,8 +253,7 @@ theorem liftOn_mk {P : Sort v} (p q : K[X]) (f : ∀ _p _q : K[X], P) (f0 : ∀ · subst hq simp only [mk_zero, f0, ← Localization.mk_zero 1, Localization.liftOn_mk, liftOn_ofFractionRing_mk, Submonoid.coe_one] - · - simp only [mk_eq_localization_mk _ hq, Localization.liftOn_mk, liftOn_ofFractionRing_mk] + · simp only [mk_eq_localization_mk _ hq, Localization.liftOn_mk, liftOn_ofFractionRing_mk] #align ratfunc.lift_on_mk RatFunc.liftOn_mk /-- Non-dependent recursion principle for `RatFunc K`: if `f p q : P` for all `p q`, diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index e6c5beaf93043..74a1619da6c60 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -422,7 +422,7 @@ theorem Trivialization.contMDiffOn_symm_trans : rintro ⟨b, x⟩ hb refine Prod.ext ?_ rfl · have : (e.toLocalHomeomorph.symm (b, x)).1 ∈ e'.baseSet - . simp_all only [Trivialization.mem_target, mfld_simps] + · simp_all only [Trivialization.mem_target, mfld_simps] exact (e'.coe_fst' this).trans (e.proj_symm_apply hb.1) variable {IB e e'} diff --git a/Mathlib/GroupTheory/Congruence.lean b/Mathlib/GroupTheory/Congruence.lean index 7ca24681e3e00..d8d17c9fdacb1 100644 --- a/Mathlib/GroupTheory/Congruence.lean +++ b/Mathlib/GroupTheory/Congruence.lean @@ -512,11 +512,11 @@ theorem conGen_eq (r : M → M → Prop) : conGen r = sInf { s : Con M | ∀ x y (le_sInf (fun s hs x y (hxy : (conGen r).r x y) => show s.r x y by apply ConGen.Rel.recOn (motive := fun x y _ => s.r x y) hxy - . exact fun x y h => hs x y h - . exact s.refl' - . exact fun _ => s.symm' - . exact fun _ _ => s.trans' - . exact fun _ _ => s.mul)) + · exact fun x y h => hs x y h + · exact s.refl' + · exact fun _ => s.symm' + · exact fun _ _ => s.trans' + · exact fun _ _ => s.mul)) (sInf_le ConGen.Rel.of) #align con.con_gen_eq Con.conGen_eq #align add_con.add_con_gen_eq AddCon.addConGen_eq @@ -716,10 +716,10 @@ def correspondence : { d // c ≤ d } ≃o Con c.Quotient Con.induction_on₂ x y fun w z h => ⟨w, z, rfl, rfl, h⟩⟩ map_rel_iff' := @fun s t => by constructor - . intros h x y hs + · intros h x y hs rcases h ⟨x, y, rfl, rfl, hs⟩ with ⟨a, b, hx, hy, ht⟩ exact t.1.trans (t.1.symm <| t.2 <| eq_rel.1 hx) (t.1.trans ht (t.2 <| eq_rel.1 hy)) - . intros h _ _ hs + · intros h _ _ hs rcases hs with ⟨a, b, hx, hy, Hs⟩ exact ⟨a, b, hx, hy, h Hs⟩ #align con.correspondence Con.correspondence @@ -1305,10 +1305,10 @@ def liftOnUnits (u : Units c.Quotient) (f : ∀ x y : M, c (x * y) 1 → c (y * (fun x y x' y' hx hy => _) u.3 u.4 refine' Function.hfunext _ _ rw [c.eq.2 hx, c.eq.2 hy] - . rintro Hxy Hxy' - + · rintro Hxy Hxy' - refine' Function.hfunext _ _ · rw [c.eq.2 hx, c.eq.2 hy] - . rintro Hyx Hyx' - + · rintro Hyx Hyx' - exact heq_of_eq (Hf _ _ _ _ _ _ _ _ hx hy) #align con.lift_on_units Con.liftOnUnits #align add_con.lift_on_add_units AddCon.liftOnAddUnits diff --git a/Mathlib/GroupTheory/FreeProduct.lean b/Mathlib/GroupTheory/FreeProduct.lean index d143cad46390b..63f282ae051e1 100644 --- a/Mathlib/GroupTheory/FreeProduct.lean +++ b/Mathlib/GroupTheory/FreeProduct.lean @@ -428,8 +428,8 @@ theorem of_smul_def (i) (w : Word M) (m : M i) : theorem cons_eq_smul {i} {m : M i} {ls h1 h2} : Word.mk (⟨i, m⟩::ls) h1 h2 = of m • mkAux ls h1 h2 := by rw [cons_eq_rcons, of_smul_def, equivPair_eq_of_fstIdx_ne _] - . simp - . rw [fstIdx_ne_iff] + · simp + · rw [fstIdx_ne_iff] exact (List.chain'_cons'.1 h2).1 #align free_product.word.cons_eq_smul FreeProduct.Word.cons_eq_smul @@ -950,10 +950,10 @@ theorem _root_.FreeGroup.injective_lift_of_ping_pong : Function.Injective (FreeG _ ⊆ X i := by clear hnne0 hlt refine Int.le_induction (P := fun n => a i ^ n • (Y i)ᶜ ⊆ X i) ?_ ?_ n h1n - . dsimp + · dsimp rw [zpow_one] exact hX i - . dsimp + · dsimp intro n _hle hi calc a i ^ (n + 1) • (Y i)ᶜ = (a i ^ n * a i) • (Y i)ᶜ := by rw [zpow_add, zpow_one] diff --git a/Mathlib/GroupTheory/Nilpotent.lean b/Mathlib/GroupTheory/Nilpotent.lean index d2861cb83ad02..d7affaa034c3f 100644 --- a/Mathlib/GroupTheory/Nilpotent.lean +++ b/Mathlib/GroupTheory/Nilpotent.lean @@ -712,8 +712,7 @@ theorem lowerCentralSeries_prod (n : ℕ) : lowerCentralSeries (G₁ × G₂) n = (lowerCentralSeries G₁ n).prod (lowerCentralSeries G₂ n) := by induction' n with n ih · simp - · - calc + · calc lowerCentralSeries (G₁ × G₂) n.succ = ⁅lowerCentralSeries (G₁ × G₂) n, ⊤⁆ := rfl _ = ⁅(lowerCentralSeries G₁ n).prod (lowerCentralSeries G₂ n), ⊤⁆ := by rw [ih] _ = ⁅(lowerCentralSeries G₁ n).prod (lowerCentralSeries G₂ n), (⊤ : Subgroup G₁).prod ⊤⁆ := @@ -755,8 +754,7 @@ theorem lowerCentralSeries_pi_le (n : ℕ) : let pi := fun f : ∀ i, Subgroup (Gs i) => Subgroup.pi Set.univ f induction' n with n ih · simp [pi_top] - · - calc + · calc lowerCentralSeries (∀ i, Gs i) n.succ = ⁅lowerCentralSeries (∀ i, Gs i) n, ⊤⁆ := rfl _ ≤ ⁅pi fun i => lowerCentralSeries (Gs i) n, ⊤⁆ := (commutator_mono ih (le_refl _)) _ = ⁅pi fun i => lowerCentralSeries (Gs i) n, pi fun i => ⊤⁆ := by simp [pi_top] @@ -790,8 +788,7 @@ theorem lowerCentralSeries_pi_of_finite [Finite η] (n : ℕ) : let pi := fun f : ∀ i, Subgroup (Gs i) => Subgroup.pi Set.univ f induction' n with n ih · simp [pi_top] - · - calc + · calc lowerCentralSeries (∀ i, Gs i) n.succ = ⁅lowerCentralSeries (∀ i, Gs i) n, ⊤⁆ := rfl _ = ⁅pi fun i => lowerCentralSeries (Gs i) n, ⊤⁆ := by rw [ih] _ = ⁅pi fun i => lowerCentralSeries (Gs i) n, pi fun i => ⊤⁆ := by simp [pi_top] diff --git a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean index e8d8b12923b4f..fa378b342fa14 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Basic.lean @@ -1713,8 +1713,7 @@ theorem isConj_of_support_equiv by_cases hx : x ∈ σ.support · rw [Equiv.extendSubtype_apply_of_mem, Equiv.extendSubtype_apply_of_mem] · exact hf x (Finset.mem_coe.2 hx) - · - rwa [Classical.not_not.1 ((not_congr mem_support).1 (Equiv.extendSubtype_not_mem f _ _)), + · rwa [Classical.not_not.1 ((not_congr mem_support).1 (Equiv.extendSubtype_not_mem f _ _)), Classical.not_not.1 ((not_congr mem_support).mp hx)] #align equiv.perm.is_conj_of_support_equiv Equiv.Perm.isConj_of_support_equiv diff --git a/Mathlib/GroupTheory/Perm/List.lean b/Mathlib/GroupTheory/Perm/List.lean index 4f4c89f836373..d41beb9d97d8b 100644 --- a/Mathlib/GroupTheory/Perm/List.lean +++ b/Mathlib/GroupTheory/Perm/List.lean @@ -95,9 +95,9 @@ theorem formPerm_apply_mem_of_mem (x : α) (l : List α) (h : x ∈ l) : formPer · by_cases hx : x ∈ z :: l · rw [formPerm_cons_cons, mul_apply, swap_apply_def] split_ifs - . simp [IH _ _ hx] - . simp - . simpa [*] using IH _ _ hx + · simp [IH _ _ hx] + · simp + · simpa [*] using IH _ _ hx · replace h : x = y := Or.resolve_right (mem_cons.1 h) hx simp [formPerm_apply_of_not_mem _ _ hx, ← h] #align list.form_perm_apply_mem_of_mem List.formPerm_apply_mem_of_mem @@ -116,7 +116,7 @@ theorem mem_of_formPerm_apply_mem (x : α) (l : List α) (h : l.formPerm x ∈ l List.mem_cons, swap_apply_def, ite_eq_left_iff] at h simp only [List.mem_cons] rcases h with h | h | h <;> split_ifs at h with h1 <;> try { aesop } - . simp [h1, imp_false] at h + · simp [h1, imp_false] at h simp [h] #align list.mem_of_form_perm_apply_mem List.mem_of_formPerm_apply_mem @@ -221,7 +221,7 @@ theorem formPerm_apply_lt (xs : List α) (h : Nodup xs) (n : ℕ) (hn : n + 1 < simp only [swap_apply_eq_iff, coe_mul, formPerm_cons_cons, Function.comp] simp only [nthLe, get_cons_succ] at * rw [← IH, swap_apply_of_ne_of_ne] <;> - . intro hx + · intro hx rw [← hx, IH] at h simp [get_mem] at h #align list.form_perm_apply_lt List.formPerm_apply_lt @@ -304,7 +304,7 @@ theorem formPerm_reverse (l : List α) (h : Nodup l) : formPerm l.reverse = (for -- We only have to check for `x ∈ l` that `formPerm l (formPerm l.reverse x)` rw [mul_apply, one_apply] cases' Classical.em (x ∈ l) with hx hx - . obtain ⟨k, hk, rfl⟩ := nthLe_of_mem ((mem_reverse _ _).mpr hx) + · obtain ⟨k, hk, rfl⟩ := nthLe_of_mem ((mem_reverse _ _).mpr hx) have h1 : l.length - 1 - k < l.length := by rw [Nat.sub_sub, add_comm] exact Nat.sub_lt_self (Nat.succ_pos _) (Nat.succ_le_of_lt (by simpa using hk)) @@ -317,10 +317,10 @@ theorem formPerm_reverse (l : List α) (h : Nodup l) : formPerm l.reverse = (for congr rw [length_reverse] at * cases' lt_or_eq_of_le (Nat.succ_le_of_lt hk) with h h - . rw [Nat.mod_eq_of_lt h, ← Nat.sub_add_comm, Nat.succ_sub_succ_eq_sub, + · rw [Nat.mod_eq_of_lt h, ← Nat.sub_add_comm, Nat.succ_sub_succ_eq_sub, Nat.mod_eq_of_lt h1] exact (Nat.le_sub_iff_add_le (length_pos_of_mem hx)).2 (Nat.succ_le_of_lt h) - . rw [← h]; simp + · rw [← h]; simp · rw [formPerm_apply_of_not_mem x l.reverse, formPerm_apply_of_not_mem _ _ hx] simpa using hx #align list.form_perm_reverse List.formPerm_reverse @@ -389,8 +389,7 @@ theorem formPerm_apply_mem_eq_self_iff (hl : Nodup l) (x : α) (hx : x ∈ l) : · rw [hn] at hk cases' (Nat.le_of_lt_succ hk).eq_or_lt with hk' hk' · simp [← hk', Nat.succ_le_succ_iff, eq_comm] - · - simpa [Nat.mod_eq_of_lt (Nat.succ_lt_succ hk'), Nat.succ_lt_succ_iff] using + · simpa [Nat.mod_eq_of_lt (Nat.succ_lt_succ hk'), Nat.succ_lt_succ_iff] using k.zero_le.trans_lt hk' #align list.form_perm_apply_mem_eq_self_iff List.formPerm_apply_mem_eq_self_iff diff --git a/Mathlib/GroupTheory/Perm/Sign.lean b/Mathlib/GroupTheory/Perm/Sign.lean index 2fed0f3b84cb8..bbe12adac0db9 100644 --- a/Mathlib/GroupTheory/Perm/Sign.lean +++ b/Mathlib/GroupTheory/Perm/Sign.lean @@ -50,9 +50,9 @@ def modSwap [DecidableEq α] (i j : α) : Setoid (Perm α) := fun {σ τ υ} hστ hτυ => by cases' hστ with hστ hστ <;> cases' hτυ with hτυ hτυ <;> try rw [hστ, hτυ, swap_mul_self_mul] <;> simp [hστ, hτυ] -- porting note: should close goals, but doesn't - . simp [hστ, hτυ] - . simp [hστ, hτυ] - . simp [hστ, hτυ]⟩ + · simp [hστ, hτυ] + · simp [hστ, hτυ] + · simp [hστ, hτυ]⟩ #align equiv.perm.mod_swap Equiv.Perm.modSwap noncomputable instance {α : Type _} [Fintype α] [DecidableEq α] (i j : α) : @@ -184,8 +184,7 @@ theorem Disjoint.extendDomain {α : Type _} {p : β → Prop} [DecidablePred p] {σ τ : Perm α} (h : Disjoint σ τ) : Disjoint (σ.extendDomain f) (τ.extendDomain f) := by intro b by_cases pb : p b - · - refine' (h (f.symm ⟨b, pb⟩)).imp _ _ <;> + · refine' (h (f.symm ⟨b, pb⟩)).imp _ _ <;> · intro h rw [extendDomain_apply_subtype _ _ pb, h, apply_symm_apply, Subtype.coe_mk] · left @@ -339,8 +338,8 @@ theorem signBijAux_inj {n : ℕ} {f : Perm (Fin n)} : have : ¬b₁ < b₂ := hb.le.not_lt split_ifs at h <;> simp_all [(Equiv.injective f).eq_iff, eq_self_iff_true, and_self_iff, heq_iff_eq] - . exact absurd this (not_le.mpr ha) - . exact absurd this (not_le.mpr ha) + · exact absurd this (not_le.mpr ha) + · exact absurd this (not_le.mpr ha) #align equiv.perm.sign_bij_aux_inj Equiv.Perm.signBijAux_inj theorem signBijAux_surj {n : ℕ} {f : Perm (Fin n)} : @@ -377,18 +376,18 @@ theorem signAux_inv {n : ℕ} (f : Perm (Fin n)) : signAux f⁻¹ = signAux f := simp_all [signBijAux, dif_pos h, if_neg h.not_le, apply_inv_self, apply_inv_self, if_neg (mem_finPairsLT.1 hab).not_le] split_ifs with h₁ - . dsimp [finPairsLT] at hab + · dsimp [finPairsLT] at hab simp at hab exact absurd h₁ (not_le_of_gt hab) - . rfl + · rfl else by simp_all [signBijAux, if_pos (le_of_not_gt h), dif_neg h, apply_inv_self, apply_inv_self, if_pos (mem_finPairsLT.1 hab).le] split_ifs with h₁ h₂ h₃ - . rfl - . exact absurd h (not_le_of_gt h₁) - . rfl - . dsimp at * + · rfl + · exact absurd h (not_le_of_gt h₁) + · rfl + · dsimp at * dsimp [finPairsLT] at hab simp at * exact absurd h₃ (asymm_of LT.lt hab)) @@ -410,8 +409,8 @@ theorem signAux_mul {n : ℕ} (f g : Perm (Fin n)) : signAux (f * g) = signAux f · rw [dif_pos h] simp only [not_le_of_gt hab, mul_one, mul_ite, mul_neg, Perm.inv_apply_self, if_false] split_ifs with h₁ h₂ h₃ <;> dsimp at * - . exact absurd hab (not_lt_of_ge h₂) - . exact absurd hab (not_lt_of_ge h₃) + · exact absurd hab (not_lt_of_ge h₂) + · exact absurd hab (not_lt_of_ge h₃) · rw [dif_neg h, inv_apply_self, inv_apply_self, if_pos hab.le] by_cases h₁ : f (g b) ≤ f (g a) · have : f (g b) ≠ f (g a) := by diff --git a/Mathlib/GroupTheory/Perm/Support.lean b/Mathlib/GroupTheory/Perm/Support.lean index 1231fc5864727..ad497c6727537 100644 --- a/Mathlib/GroupTheory/Perm/Support.lean +++ b/Mathlib/GroupTheory/Perm/Support.lean @@ -518,8 +518,7 @@ theorem eq_on_support_mem_disjoint {l : List (Perm α)} (h : f ∈ l) (hl : l.Pa rw [List.pairwise_cons] at hl rw [List.mem_cons] at h rcases h with (rfl | h) - · - rw [List.prod_cons, mul_apply, + · rw [List.prod_cons, mul_apply, not_mem_support.mp ((disjoint_prod_right tl hl.left).mem_imp hx)] · rw [List.prod_cons, mul_apply, ← IH h hl.right _ hx, eq_comm, ← not_mem_support] refine' (hl.left _ h).symm.mem_imp _ diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 13919064d6099..30f9babbedc0a 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -199,7 +199,7 @@ instance Subgroup.isCyclic {α : Type u} [Group α] [IsCyclic α] (H : Subgroup cases' k with k k · rw [Int.ofNat_eq_coe, Int.natAbs_cast k, ← zpow_ofNat, ←Int.ofNat_eq_coe, hk] exact hx₁ - . rw [Int.natAbs_negSucc, ← Subgroup.inv_mem_iff H]; simp_all⟩ + · rw [Int.natAbs_negSucc, ← Subgroup.inv_mem_iff H]; simp_all⟩ ⟨⟨⟨g ^ Nat.find hex, (Nat.find_spec hex).2⟩, fun ⟨x, hx⟩ => let ⟨k, hk⟩ := hg x have hk : g ^ k = x := hk diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index 4b63e8191af23..49a729a7d3eef 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -251,8 +251,7 @@ theorem primaryComponent.disjoint {p' : ℕ} [hp' : Fact p'.Prime] (hne : p ≠ Submonoid.disjoint_def.mpr <| by rintro g ⟨_ | n, hn⟩ ⟨n', hn'⟩ · rwa [pow_zero, orderOf_eq_one_iff] at hn - · - exact + · exact absurd (eq_of_prime_pow_eq hp.out.prime hp'.out.prime n.succ_pos (hn.symm.trans hn')) hne #align comm_monoid.primary_component.disjoint CommMonoid.primaryComponent.disjoint #align add_comm_monoid.primary_component.disjoint AddCommMonoid.primaryComponent.disjoint diff --git a/Mathlib/Init/Data/List/Instances.lean b/Mathlib/Init/Data/List/Instances.lean index 999d90487ca79..2327600de9786 100644 --- a/Mathlib/Init/Data/List/Instances.lean +++ b/Mathlib/Init/Data/List/Instances.lean @@ -24,8 +24,8 @@ instance decidableBex : ∀ (l : List α), Decidable (∃ x ∈ l, p x) | isFalse h₂ => isFalse <| by intro h; cases' h with y h; cases' h with hm hp; cases' mem_cons.1 hm with h h - . rw [h] at hp; contradiction - . exact absurd ⟨y, h, hp⟩ h₂ + · rw [h] at hp; contradiction + · exact absurd ⟨y, h, hp⟩ h₂ #align list.decidable_bex List.decidableBex instance decidableBall (l : List α) : Decidable (∀ x ∈ l, p x) := diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index c3f4c6df09a15..1fc13b1654007 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -1624,8 +1624,8 @@ isomorphism `R ⊗ R ≃ R`. noncomputable def dualDistribEquivOfBasis (b : Basis ι R M) (c : Basis κ R N) : Dual R M ⊗[R] Dual R N ≃ₗ[R] Dual R (M ⊗[R] N) := by refine' LinearEquiv.ofLinear (dualDistrib R M N) (dualDistribInvOfBasis b c) _ _ - . exact dualDistrib_dualDistribInvOfBasis_left_inverse _ _ - . exact dualDistrib_dualDistribInvOfBasis_right_inverse _ _ + · exact dualDistrib_dualDistribInvOfBasis_left_inverse _ _ + · exact dualDistrib_dualDistribInvOfBasis_right_inverse _ _ #align tensor_product.dual_distrib_equiv_of_basis TensorProduct.dualDistribEquivOfBasis variable (R M N) diff --git a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean index c8c913e67f105..d035d3fd903a1 100644 --- a/Mathlib/LinearAlgebra/Matrix/Adjugate.lean +++ b/Mathlib/LinearAlgebra/Matrix/Adjugate.lean @@ -348,8 +348,7 @@ theorem adjugate_diagonal (v : n → α) : ext i j simp only [adjugate_def, cramer_apply, diagonal_transpose, of_apply] obtain rfl | hij := eq_or_ne i j - · - rw [diagonal_apply_eq, diagonal_updateColumn_single, det_diagonal, + · rw [diagonal_apply_eq, diagonal_updateColumn_single, det_diagonal, prod_update_of_mem (Finset.mem_univ _), sdiff_singleton_eq_erase, one_mul] · rw [diagonal_apply_ne _ hij] refine' det_eq_zero_of_row_eq_zero j fun k => _ diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant.lean b/Mathlib/LinearAlgebra/Matrix/Determinant.lean index 15d69608cd330..7888e3ebc1a48 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant.lean @@ -716,8 +716,7 @@ theorem det_succ_column_zero {n : ℕ} (A : Matrix (Fin n.succ) (Fin n.succ) R) rw [Matrix.det_apply, Finset.univ_perm_fin_succ, ← Finset.univ_product_univ] simp only [Finset.sum_map, Equiv.toEmbedding_apply, Finset.sum_product, Matrix.submatrix] refine' Finset.sum_congr rfl fun i _ => Fin.cases _ (fun i => _) i - · - simp only [Fin.prod_univ_succ, Matrix.det_apply, Finset.mul_sum, + · simp only [Fin.prod_univ_succ, Matrix.det_apply, Finset.mul_sum, Equiv.Perm.decomposeFin_symm_apply_zero, Fin.val_zero, one_mul, Equiv.Perm.decomposeFin.symm_sign, Equiv.swap_self, if_true, id.def, eq_self_iff_true, Equiv.Perm.decomposeFin_symm_apply_succ, Fin.succAbove_zero, Equiv.coe_refl, pow_zero, diff --git a/Mathlib/LinearAlgebra/PiTensorProduct.lean b/Mathlib/LinearAlgebra/PiTensorProduct.lean index d8b5699de80c1..df6ae79d669bd 100644 --- a/Mathlib/LinearAlgebra/PiTensorProduct.lean +++ b/Mathlib/LinearAlgebra/PiTensorProduct.lean @@ -533,7 +533,7 @@ def isEmptyEquiv [IsEmpty ι] : (⨂[R] _ : ι, M) ≃ₗ[R] R where smul_eq_mul, mul_one] congr aesop - . simp only + · simp only intro x y hx hy rw [map_add, add_smul, hx, hy] right_inv t := by simp diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index ec42614a8b39e..7a53f0012af5a 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -870,7 +870,7 @@ protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : protected theorem forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p x ↔ q (f x)) : (∀ x, p x) ↔ (∀ y, q y) := by constructor <;> intro h₂ x - . rw [← f.right_inv x]; apply h.mp; apply h₂ + · rw [← f.right_inv x]; apply h.mp; apply h₂ · apply h.mpr; apply h₂ #align equiv.forall_congr Equiv.forall_congr diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index ac8618976f170..290984b3e4e07 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -668,8 +668,8 @@ theorem preimage_piEquivPiSubtypeProd_symm_pi {α : Type _} {β : α → Type _} -- Porting note: Two lines below were `by_cases hi <;> simp [hi]` -- This regression is https://github.com/leanprover/lean4/issues/1926 by_cases hi : p i - . simp [forall_prop_of_true hi, forall_prop_of_false (not_not.2 hi), hi] - . simp [forall_prop_of_false hi, hi, forall_prop_of_true hi] + · simp [forall_prop_of_true hi, forall_prop_of_false (not_not.2 hi), hi] + · simp [forall_prop_of_false hi, hi, forall_prop_of_true hi] #align equiv.preimage_pi_equiv_pi_subtype_prod_symm_pi Equiv.preimage_piEquivPiSubtypeProd_symm_pi -- See also `Equiv.sigmaFiberEquiv`. diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean index cb65880031bbe..c32aeb3f62670 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Metrizable.lean @@ -111,8 +111,7 @@ theorem aemeasurable_of_tendsto_metrizable_ae {ι} {μ : Measure α} {f : ι → · simp_rw [aeSeq.mk_eq_fun_of_mem_aeSeqSet h'f hx] exact @aeSeq.fun_prop_of_mem_aeSeqSet _ α β _ _ _ _ _ h'f x hx · exact tendsto_const_nhds - · - exact + · exact (ite_ae_eq_of_measure_compl_zero g (fun x => (⟨f (v 0) x⟩ : Nonempty β).some) (aeSeqSet h'f p) (aeSeq.measure_compl_aeSeqSet_eq_zero h'f hp)).symm #align ae_measurable_of_tendsto_metrizable_ae aemeasurable_of_tendsto_metrizable_ae diff --git a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean index 0ee41a10e2927..47f0a91305590 100644 --- a/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean +++ b/Mathlib/MeasureTheory/Function/StronglyMeasurable/Basic.lean @@ -865,8 +865,7 @@ theorem _root_.MeasurableEmbedding.stronglyMeasurable_extend {f : α → β} {g · rcases hx with ⟨y, rfl⟩ simpa only [SimpleFunc.extend_apply, hg.injective, Injective.extend_apply] using hf.tendsto_approx y - · - simpa only [hx, SimpleFunc.extend_apply', not_false_iff, extend_apply'] using + · simpa only [hx, SimpleFunc.extend_apply', not_false_iff, extend_apply'] using hg'.tendsto_approx x #align measurable_embedding.strongly_measurable_extend MeasurableEmbedding.stronglyMeasurable_extend diff --git a/Mathlib/MeasureTheory/Integral/Layercake.lean b/Mathlib/MeasureTheory/Integral/Layercake.lean index 38cc5f41fd196..82a8ff8e063f7 100644 --- a/Mathlib/MeasureTheory/Integral/Layercake.lean +++ b/Mathlib/MeasureTheory/Integral/Layercake.lean @@ -111,14 +111,12 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable (μ : Measure α) (Ici s).indicator (fun _ : ℝ => (1 : ℝ≥0∞)) (f x) := by funext a by_cases s ∈ Ioc (0 : ℝ) (f a) - · - simp only [h, show s ∈ Ioi (0 : ℝ) from h.1, show f a ∈ Ici s from h.2, indicator_of_mem, + · simp only [h, show s ∈ Ioi (0 : ℝ) from h.1, show f a ∈ Ici s from h.2, indicator_of_mem, mul_one] · have h_copy := h simp only [mem_Ioc, not_and, not_le] at h by_cases h' : 0 < s - · - simp only [h_copy, h h', indicator_of_not_mem, not_false_iff, mem_Ici, not_le, + · simp only [h_copy, h h', indicator_of_not_mem, not_false_iff, mem_Ici, not_le, MulZeroClass.mul_zero] · have : s ∉ Ioi (0 : ℝ) := h' simp only [this, h', indicator_of_not_mem, not_false_iff, MulZeroClass.mul_zero, diff --git a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean index 279961053511f..56ffad3c2b059 100644 --- a/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean +++ b/Mathlib/MeasureTheory/Integral/VitaliCaratheodory.lean @@ -252,8 +252,7 @@ theorem exists_lt_lowerSemicontinuous_lintegral_ge_of_aemeasurable [SigmaFinite · have : f x = fmeas.mk f x := by rw [Set.compl_subset_comm] at hs ; exact hs h rw [this] exact (f_lt_g0 x).trans_le le_self_add - · - calc + · calc ∫⁻ x, g0 x + g1 x ∂μ = (∫⁻ x, g0 x ∂μ) + ∫⁻ x, g1 x ∂μ := lintegral_add_left g0_cont.measurable _ _ ≤ (∫⁻ x, f x ∂μ) + ε / 2 + (0 + ε / 2) := by @@ -301,8 +300,7 @@ theorem exists_lt_lowerSemicontinuous_integral_gt_nnreal [SigmaFinite μ] (f : simp only [hasFiniteIntegral_iff_norm, Real.norm_eq_abs, abs_of_nonneg ENNReal.toReal_nonneg] convert gint_ne.lt_top using 1 · rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] - · - calc + · calc ENNReal.toReal (∫⁻ a : α, ENNReal.ofReal (g a).toReal ∂μ) = ENNReal.toReal (∫⁻ a : α, g a ∂μ) := by congr 1 diff --git a/Mathlib/MeasureTheory/MeasurableSpace.lean b/Mathlib/MeasureTheory/MeasurableSpace.lean index 9c60120ca0c8e..7eb6baa3b0d53 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace.lean @@ -1677,7 +1677,7 @@ theorem measurable_injection_nat_bool_of_countablyGenerated [MeasurableSpace α] obtain ⟨e, he⟩ := Set.Countable.exists_eq_range (bct.insert ∅) (insert_nonempty _ _) rw [← generateFrom_insert_empty, he] at hb refine' ⟨fun x n => x ∈ e n, _, _⟩ - . rw [measurable_pi_iff] + · rw [measurable_pi_iff] intro n apply measurable_to_bool simp only [preimage, mem_singleton_iff, Bool.decide_iff] diff --git a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean index 677af520601b8..f209d34353109 100644 --- a/Mathlib/MeasureTheory/Measure/NullMeasurable.lean +++ b/Mathlib/MeasureTheory/Measure/NullMeasurable.lean @@ -314,8 +314,7 @@ theorem measure_inter_add_diff₀ (s : Set α) (ht : NullMeasurableSet t μ) : (@disjoint_inf_sdiff _ s' t _).aedisjoint).symm _ = μ s' := (congr_arg μ (inter_union_diff _ _)) _ = μ s := hs' - · - calc + · calc μ s = μ (s ∩ t ∪ s \ t) := by rw [inter_union_diff] _ ≤ μ (s ∩ t) + μ (s \ t) := measure_union_le _ _ diff --git a/Mathlib/MeasureTheory/Measure/OuterMeasure.lean b/Mathlib/MeasureTheory/Measure/OuterMeasure.lean index 3c9df9e2c1a18..4262c00c670f8 100644 --- a/Mathlib/MeasureTheory/Measure/OuterMeasure.lean +++ b/Mathlib/MeasureTheory/Measure/OuterMeasure.lean @@ -742,8 +742,7 @@ theorem ofFunction_union_of_top_of_nonempty_inter {s t : Set α} refine' le_antisymm (OuterMeasure.union _ _ _) (le_iInf fun f => le_iInf fun hf => _) set μ := OuterMeasure.ofFunction m m_empty rcases Classical.em (∃ i, (s ∩ f i).Nonempty ∧ (t ∩ f i).Nonempty) with (⟨i, hs, ht⟩ | he) - · - calc + · calc μ s + μ t ≤ ∞ := le_top _ = m (f i) := (h (f i) hs ht).symm _ ≤ ∑' i, m (f i) := ENNReal.le_tsum i @@ -1007,8 +1006,8 @@ theorem isCaratheodory_iUnion_nat {s : ℕ → Set α} (h : ∀ i, IsCaratheodor intro t have hp : m (t ∩ ⋃ i, s i) ≤ ⨆ n, m (t ∩ ⋃ i < n, s i) := by convert m.iUnion fun i => t ∩ s i using 1 - . simp [inter_iUnion] - . simp [ENNReal.tsum_eq_iSup_nat, isCaratheodory_sum m h hd] + · simp [inter_iUnion] + · simp [ENNReal.tsum_eq_iSup_nat, isCaratheodory_sum m h hd] refine' le_trans (add_le_add_right hp _) _ rw [ENNReal.iSup_add] refine' @@ -1417,8 +1416,7 @@ theorem extend_iUnion {β} [Countable β] {f : β → Set α} (hd : Pairwise (Di (hm : ∀ i, P (f i)) : extend m (⋃ i, f i) = ∑' i, extend m (f i) := by cases nonempty_encodable β rw [← Encodable.iUnion_decode₂, ← tsum_iUnion_decode₂] - · - exact + · exact extend_iUnion_nat PU (fun n => Encodable.iUnion_decode₂_cases P0 hm) (mU _ (Encodable.iUnion_decode₂_disjoint_on hd)) · exact extend_empty P0 m0 diff --git a/Mathlib/ModelTheory/Semantics.lean b/Mathlib/ModelTheory/Semantics.lean index c9710fd6a1646..e7aeb68672b1c 100644 --- a/Mathlib/ModelTheory/Semantics.lean +++ b/Mathlib/ModelTheory/Semantics.lean @@ -185,8 +185,8 @@ theorem realize_varsToConstants [L[[α]].Structure M] [(lhomWithConstants L α). induction' t with ab n f ts ih · cases' ab with a b --Porting note: both cases were `simp [Language.con]` - . simp [Language.con, realize, constantMap, funMap_eq_coe_constants] - . simp [realize, constantMap] + · simp [Language.con, realize, constantMap, funMap_eq_coe_constants] + · simp [realize, constantMap] · simp only [realize, constantsOn, mk₂_Functions, ih] --Porting note: below lemma does not work with simp for some reason rw [withConstants_funMap_sum_inl] @@ -335,8 +335,7 @@ theorem realize_foldr_sup (l : List (L.BoundedFormula α n)) (v : α → M) (xs (l.foldr (· ⊔ ·) ⊥).Realize v xs ↔ ∃ φ ∈ l, BoundedFormula.Realize φ v xs := by induction' l with φ l ih · simp - · - simp_rw [List.foldr_cons, realize_sup, ih, exists_prop, List.mem_cons, or_and_right, exists_or, + · simp_rw [List.foldr_cons, realize_sup, ih, exists_prop, List.mem_cons, or_and_right, exists_or, exists_eq_left] #align first_order.language.bounded_formula.realize_foldr_sup FirstOrder.Language.BoundedFormula.realize_foldr_sup @@ -557,9 +556,9 @@ theorem realize_toPrenexImp {φ ψ : L.BoundedFormula α n} (hφ : IsPrenex φ) theorem realize_toPrenex (φ : L.BoundedFormula α n) {v : α → M} : ∀ {xs : Fin n → M}, φ.toPrenex.Realize v xs ↔ φ.Realize v xs := by induction' φ with _ _ _ _ _ _ _ _ _ f1 f2 h1 h2 _ _ h - . exact Iff.rfl - . exact Iff.rfl - . exact Iff.rfl + · exact Iff.rfl + · exact Iff.rfl + · exact Iff.rfl · intros rw [toPrenex, realize_toPrenexImp f1.toPrenex_isPrenex f2.toPrenex_isPrenex, realize_imp, realize_imp, h1, h2] @@ -741,8 +740,8 @@ theorem realize_equivSentence_symm_con [L[[α]].Structure M] refine' _root_.trans _ BoundedFormula.realize_constantsVarsEquiv rw [iff_iff_eq] congr with (_ | a) - . simp - . cases a + · simp + · cases a #align first_order.language.formula.realize_equiv_sentence_symm_con FirstOrder.Language.Formula.realize_equivSentence_symm_con @[simp] @@ -919,8 +918,7 @@ theorem realize_toFormula (φ : L.BoundedFormula α n) (v : Sum α (Fin n) → M · rfl · simp [BoundedFormula.Realize] · simp [BoundedFormula.Realize] - · - rw [toFormula, Formula.Realize, realize_imp, ← Formula.Realize, ih1, ← Formula.Realize, ih2, + · rw [toFormula, Formula.Realize, realize_imp, ← Formula.Realize, ih1, ← Formula.Realize, ih2, realize_imp] · rw [toFormula, Formula.Realize, realize_all, realize_all] refine' forall_congr' fun a => _ diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 4713bdafeb660..bf9ae5ef7f34d 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -567,8 +567,8 @@ theorem ppow_succ {f : ArithmeticFunction R} {k : ℕ} : f.ppow (k + 1) = f.pmul ext x simp_rw [ppow_apply (Nat.succ_pos k), pow_succ] -- porting note: was `rw [..., pow_succ]` induction k <;> simp - . exact pow_one _ -- porting note: added - . exact pow_succ'' _ _ -- porting note: added + · exact pow_one _ -- porting note: added + · exact pow_succ'' _ _ -- porting note: added #align nat.arithmetic_function.ppow_succ Nat.ArithmeticFunction.ppow_succ theorem ppow_succ' {f : ArithmeticFunction R} {k : ℕ} {kpos : 0 < k} : @@ -666,22 +666,19 @@ theorem mul [CommSemiring R] {f g : ArithmeticFunction R} (hf : f.IsMultiplicati rw [← hcd.1.1, h.1, Nat.gcd_mul_left, cop.coprime_mul_left.coprime_mul_right_right.gcd_eq_one, mul_one] · trans Nat.gcd (a1 * a2) (a2 * b2) - · - rw [mul_comm, Nat.gcd_mul_left, cop.coprime_mul_right.coprime_mul_left_right.gcd_eq_one, + · rw [mul_comm, Nat.gcd_mul_left, cop.coprime_mul_right.coprime_mul_left_right.gcd_eq_one, mul_one] · rw [← hcd.1.1, ← hcd.2.1] at cop rw [← hcd.1.1, h.2, mul_comm, Nat.gcd_mul_left, cop.coprime_mul_right.coprime_mul_left_right.gcd_eq_one, mul_one] · trans Nat.gcd (b1 * b2) (a1 * b1) - · - rw [mul_comm, Nat.gcd_mul_right, + · rw [mul_comm, Nat.gcd_mul_right, cop.coprime_mul_right.coprime_mul_left_right.symm.gcd_eq_one, one_mul] · rw [← hcd.1.1, ← hcd.2.1] at cop rw [← hcd.2.1, h.1, mul_comm c1 d1, Nat.gcd_mul_left, cop.coprime_mul_right.coprime_mul_left_right.symm.gcd_eq_one, mul_one] · trans Nat.gcd (b1 * b2) (a2 * b2) - · - rw [Nat.gcd_mul_right, cop.coprime_mul_left.coprime_mul_right_right.symm.gcd_eq_one, + · rw [Nat.gcd_mul_right, cop.coprime_mul_left.coprime_mul_right_right.symm.gcd_eq_one, one_mul] · rw [← hcd.1.1, ← hcd.2.1] at cop rw [← hcd.2.1, h.2, Nat.gcd_mul_right, diff --git a/Mathlib/NumberTheory/BernoulliPolynomials.lean b/Mathlib/NumberTheory/BernoulliPolynomials.lean index c968632b151ac..f89c4ae4d91cb 100644 --- a/Mathlib/NumberTheory/BernoulliPolynomials.lean +++ b/Mathlib/NumberTheory/BernoulliPolynomials.lean @@ -131,11 +131,11 @@ nonrec theorem sum_bernoulli (n : ℕ) : simp_rw [smul_monomial, mul_comm (_root_.bernoulli _) _, smul_eq_mul, ← mul_assoc] conv_lhs => apply_congr - . skip - . conv => + · skip + · conv => apply_congr - . skip - . rw [← Nat.cast_mul, choose_mul ((le_tsub_iff_left <| mem_range_le (by assumption)).1 <| + · skip + · rw [← Nat.cast_mul, choose_mul ((le_tsub_iff_left <| mem_range_le (by assumption)).1 <| mem_range_le (by assumption)) (le.intro rfl), Nat.cast_mul, add_tsub_cancel_left, mul_assoc, mul_comm, ← smul_eq_mul, ← smul_monomial] @@ -203,17 +203,17 @@ theorem bernoulli_eval_one_add (n : ℕ) (x : ℚ) : bernoulli_eq_sub_sum, eval_sub, eval_finset_sum] conv_lhs => congr - . skip - . apply_congr - . skip - . rw [eval_smul, hd _ (mem_range.1 (by assumption))] + · skip + · apply_congr + · skip + · rw [eval_smul, hd _ (mem_range.1 (by assumption))] rw [eval_sub, eval_finset_sum] simp_rw [eval_smul, smul_add] rw [sum_add_distrib, sub_add, sub_eq_sub_iff_sub_eq_sub, _root_.add_sub_sub_cancel] conv_rhs => congr - . skip - . congr + · skip + · congr rw [succ_eq_add_one, ← choose_succ_self_right d] rw [Nat.cast_succ, ← smul_eq_mul, ← sum_range_succ _ d, eval_monomial_one_add_sub] simp_rw [smul_eq_mul] diff --git a/Mathlib/NumberTheory/Bertrand.lean b/Mathlib/NumberTheory/Bertrand.lean index 3932002cf1b77..42d801e69053b 100644 --- a/Mathlib/NumberTheory/Bertrand.lean +++ b/Mathlib/NumberTheory/Bertrand.lean @@ -77,7 +77,7 @@ theorem real_main_inequality {x : ℝ} (n_large : (512 : ℝ) ≤ x) : (Set.Ioi_subset_Ioi (by norm_num)) (convex_Ioi 0.5) convert ((strictConcaveOn_sqrt_mul_log_Ioi.concaveOn.comp_linearMap ((2 : ℝ) • LinearMap.id))) using 1 - . ext x + · ext x simp only [Set.mem_Ioi, Set.mem_preimage, LinearMap.smul_apply, LinearMap.id_coe, id_eq, smul_eq_mul] rw [← mul_lt_mul_left (two_pos)] @@ -90,7 +90,7 @@ theorem real_main_inequality {x : ℝ} (n_large : (512 : ℝ) ≤ x) : obtain ⟨x1, x2, h1, h2, h0, h3, h4⟩ := this exact (h.right_le_of_le_left'' h1 ((h1.trans h2).trans_le h0) h2 h0 (h4.trans h3)).trans h4 refine' ⟨18, 512, by norm_num1, by norm_num1, n_large, _, _⟩ - . have : sqrt (2 * 18) = 6 := (sqrt_eq_iff_mul_self_eq_of_pos (by norm_num1)).mpr (by norm_num1) + · have : sqrt (2 * 18) = 6 := (sqrt_eq_iff_mul_self_eq_of_pos (by norm_num1)).mpr (by norm_num1) rw [hf, log_nonneg_iff, this] rw [one_le_div] <;> norm_num1 apply le_trans _ (le_mul_of_one_le_left _ _) <;> norm_num1 @@ -99,17 +99,17 @@ theorem real_main_inequality {x : ℝ} (n_large : (512 : ℝ) ≤ x) : apply rpow_pos_of_pos ; norm_num1 apply hf' 18 ; norm_num1 norm_num1 - . have : sqrt (2 * 512) = 32 := + · have : sqrt (2 * 512) = 32 := (sqrt_eq_iff_mul_self_eq_of_pos (by norm_num1)).mpr (by norm_num1) rw [hf, log_nonpos_iff (hf' _ _), this, div_le_one] <;> norm_num1 have : (512 : ℝ) = 2 ^ (9 : ℕ) - . rw [rpow_nat_cast 2 9] ; norm_num1 + · rw [rpow_nat_cast 2 9] ; norm_num1 conv_lhs => rw [this] have : (1024 : ℝ) = 2 ^ (10 : ℕ) - . rw [rpow_nat_cast 2 10] ; norm_num1 + · rw [rpow_nat_cast 2 10] ; norm_num1 rw [this, ← rpow_mul, ← rpow_add] <;> norm_num1 have : (4 : ℝ) = 2 ^ (2 : ℕ) - . rw [rpow_nat_cast 2 2] ; norm_num1 + · rw [rpow_nat_cast 2 2] ; norm_num1 rw [this, ← rpow_mul] <;> norm_num1 apply rpow_le_rpow_of_exponent_le <;> norm_num1 apply rpow_pos_of_pos four_pos diff --git a/Mathlib/NumberTheory/ClassNumber/Finite.lean b/Mathlib/NumberTheory/ClassNumber/Finite.lean index 9e3e5fd69f37d..44e7cd7ea6d7a 100644 --- a/Mathlib/NumberTheory/ClassNumber/Finite.lean +++ b/Mathlib/NumberTheory/ClassNumber/Finite.lean @@ -102,7 +102,7 @@ theorem norm_le (a : S) {y : ℤ} (hy : ∀ k, abv (bS.repr a k) ≤ y) : -- exact finset.mem_image.mpr ⟨⟨i, j, k⟩, Finset.mem_univ _, rfl⟩ rw [← LinearMap.det_toMatrix bS] convert Matrix.det_sum_smul_le (n := ι) Finset.univ _ hy using 3 - . simp; rfl + · simp; rfl · rw [Finset.card_univ, smul_mul_assoc, mul_comm] · intro i j k apply Finset.le_max' diff --git a/Mathlib/NumberTheory/Cyclotomic/Basic.lean b/Mathlib/NumberTheory/Cyclotomic/Basic.lean index dc298fabae79a..14694e3047211 100644 --- a/Mathlib/NumberTheory/Cyclotomic/Basic.lean +++ b/Mathlib/NumberTheory/Cyclotomic/Basic.lean @@ -555,9 +555,9 @@ instance isCyclotomicExtension [NeZero ((n : ℕ) : K)] : rw [← eval_map, ← IsRoot.def, map_cyclotomic, isRoot_cyclotomic_iff] at hζ -- Porting note: the first `?_` was `forall_eq.2 ⟨ζ, hζ⟩` that now fails. refine ⟨?_, ?_⟩ - . simp only [mem_singleton_iff, forall_eq] + · simp only [mem_singleton_iff, forall_eq] exact ⟨ζ, hζ⟩ - . rw [← Algebra.eq_top_iff, ← SplittingField.adjoin_rootSet, eq_comm] + · rw [← Algebra.eq_top_iff, ← SplittingField.adjoin_rootSet, eq_comm] exact IsCyclotomicExtension.adjoin_roots_cyclotomic_eq_adjoin_nth_roots hζ #align cyclotomic_field.is_cyclotomic_extension CyclotomicField.isCyclotomicExtension diff --git a/Mathlib/NumberTheory/PellMatiyasevic.lean b/Mathlib/NumberTheory/PellMatiyasevic.lean index fe88b1c56b13b..15515c624719c 100644 --- a/Mathlib/NumberTheory/PellMatiyasevic.lean +++ b/Mathlib/NumberTheory/PellMatiyasevic.lean @@ -809,7 +809,7 @@ theorem modEq_of_xn_modEq {i j n} (ipos : 0 < i) (hin : i ≤ n) have jj : j ≡ j' [MOD 4 * n] := by delta ModEq; rw [Nat.mod_eq_of_lt jl] have : ∀ j q, xn a1 (j + 4 * n * q) ≡ xn a1 j [MOD xn a1 n] := by intro j q; induction' q with q IH - . simp [ModEq.refl] + · simp [ModEq.refl] rw [Nat.mul_succ, ← add_assoc, add_comm] exact (xn_modEq_x4n_add _ _ _).trans IH Or.imp (fun ji : j' = i => by rwa [← ji]) diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 8544a76ae28c8..6c460832d5de0 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -586,7 +586,7 @@ protected def booleanAlgebra {α} [DecidableEq α] [Lattice α] [BoundedOrder α inf_compl_le_bot := fun x => by rcases eq_bot_or_eq_top x with (rfl | rfl) · simp - . simp + · simp top_le_sup_compl := fun x => by rcases eq_bot_or_eq_top x with (rfl | rfl) <;> simp } #align is_simple_order.boolean_algebra IsSimpleOrder.booleanAlgebra diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index d77be44fe167f..02f571ce1033b 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -1593,8 +1593,7 @@ theorem isLUB_prod [Preorder α] [Preorder β] {s : Set (α × β)} (p : α × · suffices h : (p.1, a) ∈ upperBounds s from (H.2 h).2 exact fun q hq => ⟨(H.1 hq).1, ha <| mem_image_of_mem _ hq⟩ · exact fun q hq => ⟨H.1.1 <| mem_image_of_mem _ hq, H.2.1 <| mem_image_of_mem _ hq⟩ - · - exact fun q hq => + · exact fun q hq => ⟨H.1.2 <| monotone_fst.mem_upperBounds_image hq, H.2.2 <| monotone_snd.mem_upperBounds_image hq⟩ #align is_lub_prod isLUB_prod diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index 2b88dc9ddb983..4b95baffcd675 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -391,12 +391,12 @@ instance CompleteLinearOrder.toLinearOrder [i : CompleteLinearOrder α] : Linear max := Sup.sup min_def := fun a b => by split_ifs with h - . simp [h] - . simp [(CompleteLinearOrder.le_total a b).resolve_left h] + · simp [h] + · simp [(CompleteLinearOrder.le_total a b).resolve_left h] max_def := fun a b => by split_ifs with h - . simp [h] - . simp [(CompleteLinearOrder.le_total a b).resolve_left h] } + · simp [h] + · simp [(CompleteLinearOrder.le_total a b).resolve_left h] } namespace OrderDual diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index d58b6c225bf8b..2106e150cc114 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -1190,8 +1190,7 @@ theorem isGLB_sInf' {β : Type _} [ConditionallyCompleteLattice β] {s : Set (Wi intro b hb exact Set.mem_singleton_iff.2 (top_le_iff.1 (ha hb)) · refine' some_le_some.2 (le_csInf _ _) - · - classical + · classical contrapose! h rintro (⟨⟩ | a) ha · exact mem_singleton ⊤ diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 723d2d42d517a..1e7dd0cb4861a 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -316,8 +316,8 @@ theorem eq_of_chain {c : Chain (Part α)} {a b : α} (ha : some a ∈ c) (hb : s cases' hb with j hb; replace hb := hb.symm rw [eq_some_iff] at ha hb cases' le_total i j with hij hji - . have := c.monotone hij _ ha; apply mem_unique this hb - . have := c.monotone hji _ hb; apply Eq.symm; apply mem_unique this ha + · have := c.monotone hij _ ha; apply mem_unique this hb + · have := c.monotone hji _ hb; apply Eq.symm; apply mem_unique this ha --Porting note: Old proof -- wlog h : i ≤ j := le_total i j using a b i j, b a j i -- rw [eq_some_iff] at ha hb diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index 37222926569b2..0331d2324539d 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -1021,7 +1021,7 @@ instance : SuccOrder (WithTop α) where · exact le_top change _ ≤ ite _ _ _ split_ifs - . exact le_top + · exact le_top · exact some_le_some.2 (le_succ a) max_of_succ_le {a} ha := by cases a diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 3591e02fad35f..65ee0c3a13f47 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -393,7 +393,7 @@ theorem unbot'_bot_le_iff [LE α] [OrderBot α] {a : WithBot α} {b : α} : theorem unbot'_lt_iff [LT α] {a : WithBot α} {b c : α} (ha : a ≠ ⊥) : a.unbot' b < c ↔ a < c := by cases a · exact (ha rfl).elim - . rw [some_eq_coe, unbot'_coe, coe_lt_coe] + · rw [some_eq_coe, unbot'_coe, coe_lt_coe] #align with_bot.unbot'_lt_iff WithBot.unbot'_lt_iff instance semilatticeSup [SemilatticeSup α] : SemilatticeSup (WithBot α) := @@ -1289,10 +1289,10 @@ instance trichotomous.lt [Preorder α] [IsTrichotomous α (· < ·)] : IsTrichotomous (WithTop α) (· < ·) := ⟨by rintro (a | a) (b | b) - . simp - . simp - . simp - . simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (. < .) _ a b⟩ + · simp + · simp + · simp + · simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (. < .) _ a b⟩ #align with_top.trichotomous.lt WithTop.trichotomous.lt instance IsWellOrder.lt [Preorder α] [h : IsWellOrder α (· < ·)] : @@ -1303,10 +1303,10 @@ instance trichotomous.gt [Preorder α] [IsTrichotomous α (· > ·)] : IsTrichotomous (WithTop α) (· > ·) := ⟨by rintro (a | a) (b | b) - . simp - . simp - . simp - . simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (. > .) _ a b⟩ + · simp + · simp + · simp + · simpa [some_eq_coe, IsTrichotomous, coe_eq_coe] using @trichotomous α (. > .) _ a b⟩ #align with_top.trichotomous.gt WithTop.trichotomous.gt instance IsWellOrder.gt [Preorder α] [h : IsWellOrder α (· > ·)] : diff --git a/Mathlib/Order/Zorn.lean b/Mathlib/Order/Zorn.lean index 073ad165df0a9..72ad43e4e20eb 100644 --- a/Mathlib/Order/Zorn.lean +++ b/Mathlib/Order/Zorn.lean @@ -137,7 +137,7 @@ theorem zorn_nonempty_preorder₀ (s : Set α) -- rcases zorn_preorder₀ ({ y ∈ s | x ≤ y }) fun c hcs hc => ?_ with ⟨m, ⟨hms, hxm⟩, hm⟩ -- · exact ⟨m, hms, hxm, fun z hzs hmz => hm _ ⟨hzs, hxm.trans hmz⟩ hmz⟩ have H := zorn_preorder₀ ({ y ∈ s | x ≤ y }) fun c hcs hc => ?_ - . rcases H with ⟨m, ⟨hms, hxm⟩, hm⟩ + · rcases H with ⟨m, ⟨hms, hxm⟩, hm⟩ exact ⟨m, hms, hxm, fun z hzs hmz => hm _ ⟨hzs, hxm.trans hmz⟩ hmz⟩ · rcases c.eq_empty_or_nonempty with (rfl | ⟨y, hy⟩) · exact ⟨x, ⟨hxs, le_rfl⟩, fun z => False.elim⟩ diff --git a/Mathlib/Probability/Kernel/CondCdf.lean b/Mathlib/Probability/Kernel/CondCdf.lean index d972770e5e570..8957b876587e7 100644 --- a/Mathlib/Probability/Kernel/CondCdf.lean +++ b/Mathlib/Probability/Kernel/CondCdf.lean @@ -94,8 +94,7 @@ theorem Real.iInter_Iic_rat : ⋂ r : ℚ, Iic (r : ℝ) = ∅ := by theorem atBot_le_nhds_bot {α : Type _} [TopologicalSpace α] [LinearOrder α] [OrderBot α] [OrderTopology α] : (atBot : Filter α) ≤ 𝓝 ⊥ := by cases subsingleton_or_nontrivial α - · - simp only [nhds_discrete, le_pure_iff, mem_atBot_sets, mem_singleton_iff, + · simp only [nhds_discrete, le_pure_iff, mem_atBot_sets, mem_singleton_iff, eq_iff_true_of_subsingleton, imp_true_iff, exists_const] have h : atBot.HasBasis (fun _ : α => True) Iic := @atBot_basis α _ _ have h_nhds : (𝓝 ⊥).HasBasis (fun a : α => ⊥ < a) fun a => Iio a := @nhds_bot_basis α _ _ _ _ _ @@ -730,8 +729,7 @@ theorem condCdf'_eq_condCdfRat (ρ : Measure (α × ℝ)) (a : α) (r : ℚ) : condCdf' ρ a r = condCdfRat ρ a r := by rw [← inf_gt_condCdfRat ρ a r, condCdf'] refine' Equiv.iInf_congr _ _ - · - exact + · exact { toFun := fun t => ⟨t.1, by exact_mod_cast t.2⟩ invFun := fun t => ⟨t.1, by exact_mod_cast t.2⟩ left_inv := fun t => by simp only [Subtype.coe_eta] diff --git a/Mathlib/RingTheory/Adjoin/Basic.lean b/Mathlib/RingTheory/Adjoin/Basic.lean index 8380fd0480330..99f375c271eb7 100644 --- a/Mathlib/RingTheory/Adjoin/Basic.lean +++ b/Mathlib/RingTheory/Adjoin/Basic.lean @@ -254,8 +254,7 @@ theorem adjoin_inl_union_inr_eq_prod (s) (t) : adjoin R (LinearMap.inl R A B '' (s ∪ {1}) ∪ LinearMap.inr R A B '' (t ∪ {1})) = (adjoin R s).prod (adjoin R t) := by apply le_antisymm - · - simp only [adjoin_le_iff, Set.insert_subset_iff, Subalgebra.zero_mem, Subalgebra.one_mem, + · simp only [adjoin_le_iff, Set.insert_subset_iff, Subalgebra.zero_mem, Subalgebra.one_mem, subset_adjoin,-- the rest comes from `squeeze_simp` Set.union_subset_iff, LinearMap.coe_inl, Set.mk_preimage_prod_right, Set.image_subset_iff, SetLike.mem_coe, diff --git a/Mathlib/RingTheory/AdjoinRoot.lean b/Mathlib/RingTheory/AdjoinRoot.lean index 830349fac7223..69fc9947b6b47 100644 --- a/Mathlib/RingTheory/AdjoinRoot.lean +++ b/Mathlib/RingTheory/AdjoinRoot.lean @@ -402,8 +402,8 @@ noncomputable instance field [Fact (Irreducible f)] : Field (AdjoinRoot f) := -- porting note: was -- `rw [Rat.cast_mk' (K := ℚ), _root_.map_mul, _root_.map_intCast, map_inv₀, map_natCast]` convert_to ((Rat.mk' a b h1 h2 : K) : AdjoinRoot f) = ((↑a * (↑b)⁻¹ : K) : AdjoinRoot f) - . simp only [_root_.map_mul, map_intCast, map_inv₀, map_natCast] - . simp only [Rat.cast_mk', _root_.map_mul, map_intCast, map_inv₀, map_natCast] + · simp only [_root_.map_mul, map_intCast, map_inv₀, map_natCast] + · simp only [Rat.cast_mk', _root_.map_mul, map_intCast, map_inv₀, map_natCast] qsmul := (· • ·) qsmul_eq_mul' := fun a x => -- porting note: I gave the explicit motive and changed `rw` to `simp`. diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index 16d9647bf92db..bf0a4f853a717 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -344,8 +344,8 @@ theorem isArtinian_of_fg_of_artinian {R M} [Ring R] [AddCommGroup M] [Module R M haveI := Classical.decEq R have : ∀ x ∈ s, x ∈ N := fun x hx => hs ▸ Submodule.subset_span hx refine' @isArtinian_of_surjective _ ((↑s : Set M) →₀ R) N _ _ _ _ _ _ _ isArtinian_finsupp - . exact Finsupp.total (↑s : Set M) N R (fun i => ⟨i, hs ▸ subset_span i.2⟩) - . rw [← LinearMap.range_eq_top, eq_top_iff, + · exact Finsupp.total (↑s : Set M) N R (fun i => ⟨i, hs ▸ subset_span i.2⟩) + · rw [← LinearMap.range_eq_top, eq_top_iff, ← map_le_map_iff_of_injective (show Injective (Submodule.subtype N) from Subtype.val_injective), Submodule.map_top, range_subtype, ← Submodule.map_top, ← Submodule.map_comp, Submodule.map_top] diff --git a/Mathlib/RingTheory/HahnSeries.lean b/Mathlib/RingTheory/HahnSeries.lean index b581e6e642bc1..e0b2a7b063eed 100644 --- a/Mathlib/RingTheory/HahnSeries.lean +++ b/Mathlib/RingTheory/HahnSeries.lean @@ -391,7 +391,7 @@ theorem min_order_le_order_add {Γ} [LinearOrderedCancelAddCommMonoid Γ] {x y : by_cases hy : y = 0; · simp [hy] rw [order_of_ne hx, order_of_ne hy, order_of_ne hxy] refine' le_of_eq_of_le _ (Set.IsWf.min_le_min_of_subset (support_add_subset (x := x) (y := y))) - . exact (Set.IsWf.min_union _ _ _ _).symm + · exact (Set.IsWf.min_union _ _ _ _).symm #align hahn_series.min_order_le_order_add HahnSeries.min_order_le_order_add /-- `single` as an additive monoid/group homomorphism -/ @@ -1226,7 +1226,7 @@ theorem ofPowerSeries_X_pow {R} [CommSemiring R] (n : ℕ) : induction' n with n ih · simp rfl - . rw [pow_succ, pow_succ, ih, ofPowerSeries_X, mul_comm, single_mul_single, one_mul, + · rw [pow_succ, pow_succ, ih, ofPowerSeries_X, mul_comm, single_mul_single, one_mul, Nat.cast_succ, add_comm] #align hahn_series.of_power_series_X_pow HahnSeries.ofPowerSeries_X_pow diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 49542d64b38ec..186131a7ac107 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -780,8 +780,7 @@ division (semi)ring. This result actually holds for all division semirings, but we lack the predicate to state it. -/ theorem isField_iff_isSimpleOrder_ideal : IsField R ↔ IsSimpleOrder (Ideal R) := by cases subsingleton_or_nontrivial R - · - exact + · exact ⟨fun h => (not_isField_of_subsingleton _ h).elim, fun h => (false_of_nontrivial_of_subsingleton <| Ideal R).elim⟩ rw [← not_iff_not, Ring.not_isField_iff_exists_ideal_bot_lt_and_lt_top, ← not_iff_not] diff --git a/Mathlib/RingTheory/Ideal/LocalRing.lean b/Mathlib/RingTheory/Ideal/LocalRing.lean index b0f2d52b31571..0fbef5b27ec98 100644 --- a/Mathlib/RingTheory/Ideal/LocalRing.lean +++ b/Mathlib/RingTheory/Ideal/LocalRing.lean @@ -300,22 +300,22 @@ theorem local_hom_TFAE (f : R →+* S) : (maximalIdeal R).map f ≤ maximalIdeal S, maximalIdeal R ≤ (maximalIdeal S).comap f, (maximalIdeal S).comap f = maximalIdeal R] := by tfae_have 1 → 2 - . rintro _ _ ⟨a, ha, rfl⟩ + · rintro _ _ ⟨a, ha, rfl⟩ exact map_nonunit f a ha tfae_have 2 → 4 - . exact Set.image_subset_iff.1 + · exact Set.image_subset_iff.1 tfae_have 3 ↔ 4 - . exact Ideal.map_le_iff_le_comap + · exact Ideal.map_le_iff_le_comap tfae_have 4 → 1 - . intro h + · intro h constructor exact fun x => not_imp_not.1 (@h x) tfae_have 1 → 5 - . intro + · intro ext exact not_iff_not.2 (isUnit_map_iff f _) tfae_have 5 → 4 - . exact fun h => le_of_eq h.symm + · exact fun h => le_of_eq h.symm tfae_finish #align local_ring.local_hom_tfae LocalRing.local_hom_TFAE diff --git a/Mathlib/RingTheory/Ideal/MinimalPrime.lean b/Mathlib/RingTheory/Ideal/MinimalPrime.lean index 065846b5aa402..ccabc314606e8 100644 --- a/Mathlib/RingTheory/Ideal/MinimalPrime.lean +++ b/Mathlib/RingTheory/Ideal/MinimalPrime.lean @@ -77,11 +77,11 @@ theorem Ideal.radical_minimalPrimes : I.radical.minimalPrimes = I.minimalPrimes congr ext p refine' ⟨_, _⟩ <;> rintro ⟨⟨a, ha⟩, b⟩ - . refine' ⟨⟨a, a.radical_le_iff.1 ha⟩, _⟩ - . simp only [Set.mem_setOf_eq, and_imp] at * + · refine' ⟨⟨a, a.radical_le_iff.1 ha⟩, _⟩ + · simp only [Set.mem_setOf_eq, and_imp] at * exact fun _ h2 h3 h4 => b h2 (h2.radical_le_iff.2 h3) h4 - . refine' ⟨⟨a, a.radical_le_iff.2 ha⟩, _⟩ - . simp only [Set.mem_setOf_eq, and_imp] at * + · refine' ⟨⟨a, a.radical_le_iff.2 ha⟩, _⟩ + · simp only [Set.mem_setOf_eq, and_imp] at * exact fun _ h2 h3 h4 => b h2 (h2.radical_le_iff.1 h3) h4 #align ideal.radical_minimal_primes Ideal.radical_minimalPrimes @@ -138,7 +138,7 @@ theorem Ideal.exists_comap_eq_of_mem_minimalPrimes {I : Ideal S} (f : R →+* S) (RingHom.kerLift_injective f') (p.map <| Ideal.Quotient.mk <| RingHom.ker f') this refine' ⟨p'.comap <| Ideal.Quotient.mk I, Ideal.IsPrime.comap _, _, _⟩ · exact Ideal.mk_ker.symm.trans_le (Ideal.comap_mono bot_le) - . convert congr_arg (Ideal.comap <| Ideal.Quotient.mk <| RingHom.ker f') hp₂ + · convert congr_arg (Ideal.comap <| Ideal.Quotient.mk <| RingHom.ker f') hp₂ rwa [Ideal.comap_map_of_surjective (Ideal.Quotient.mk <| RingHom.ker f') Ideal.Quotient.mk_surjective, eq_comm, sup_eq_left] refine' ⟨⟨_, bot_le⟩, _⟩ @@ -199,8 +199,7 @@ theorem Ideal.minimalPrimes_eq_comap : theorem Ideal.minimalPrimes_eq_subsingleton (hI : I.IsPrimary) : I.minimalPrimes = {I.radical} := by ext J constructor - · - exact fun H => + · exact fun H => let e := H.1.1.radical_le_iff.mpr H.1.2 (H.2 ⟨Ideal.isPrime_radical hI, Ideal.le_radical⟩ e).antisymm e · rintro (rfl : J = I.radical) diff --git a/Mathlib/RingTheory/Ideal/Over.lean b/Mathlib/RingTheory/Ideal/Over.lean index 70eead40c0623..13686c4a60883 100644 --- a/Mathlib/RingTheory/Ideal/Over.lean +++ b/Mathlib/RingTheory/Ideal/Over.lean @@ -69,8 +69,8 @@ theorem exists_coeff_ne_zero_mem_comap_of_non_zero_divisor_root_mem {r : S} rw [eval₂_mul, eval₂_X] at hp obtain ⟨i, hi, mem⟩ := ih p_nonzero (r_non_zero_divisor hp) refine' ⟨i + 1, _, _⟩ - . simp [hi, mem] - . simpa [hi] using mem + · simp [hi, mem] + · simpa [hi] using mem #align ideal.exists_coeff_ne_zero_mem_comap_of_non_zero_divisor_root_mem Ideal.exists_coeff_ne_zero_mem_comap_of_non_zero_divisor_root_mem /-- Let `P` be an ideal in `R[x]`. The map diff --git a/Mathlib/RingTheory/IntegralDomain.lean b/Mathlib/RingTheory/IntegralDomain.lean index 3696b85ed6dd7..edcc2b8199a3e 100644 --- a/Mathlib/RingTheory/IntegralDomain.lean +++ b/Mathlib/RingTheory/IntegralDomain.lean @@ -160,8 +160,8 @@ variable (S : Subgroup Rˣ) [Finite S] instance subgroup_units_cyclic : IsCyclic S := by -- porting note: the original proof used a `coe`, but I was not able to get it to work. apply isCyclic_of_subgroup_isDomain (R := R) (G := S) _ _ - . exact MonoidHom.mk (OneHom.mk (fun s => ↑s.val) rfl) (by simp) - . exact Units.ext.comp Subtype.val_injective + · exact MonoidHom.mk (OneHom.mk (fun s => ↑s.val) rfl) (by simp) + · exact Units.ext.comp Subtype.val_injective #align subgroup_units_cyclic subgroup_units_cyclic end diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 6dbb40d12eebe..aab83fd53462d 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -498,10 +498,10 @@ theorem coeff_isUnit_isNilpotent_of_isUnit {P : Polynomial R} (hunit : IsUnit P) IsUnit (P.coeff 0) ∧ (∀ i, i ≠ 0 → IsNilpotent (P.coeff i)) := by obtain ⟨Q, hQ⟩ := IsUnit.exists_right_inv hunit constructor - . refine' isUnit_of_mul_eq_one _ (Q.coeff 0) _ + · refine' isUnit_of_mul_eq_one _ (Q.coeff 0) _ have h := (mul_coeff_zero P Q).symm rwa [hQ, coeff_one_zero] at h - . intros n hn + · intros n hn rw [nilpotent_iff_mem_prime] intros I hI let f := mapRingHom (Ideal.Quotient.mk I) diff --git a/Mathlib/RingTheory/Polynomial/Content.lean b/Mathlib/RingTheory/Polynomial/Content.lean index ddce458bd0401..12cd5d429a653 100644 --- a/Mathlib/RingTheory/Polynomial/Content.lean +++ b/Mathlib/RingTheory/Polynomial/Content.lean @@ -180,8 +180,8 @@ theorem normalize_content {p : R[X]} : normalize p.content = p.content := @[simp] theorem normUnit_content {p : R[X]} : normUnit (content p) = 1 := by by_cases hp0 : p.content = 0 - . simp [hp0] - . ext + · simp [hp0] + · ext apply mul_left_cancel₀ hp0 erw [← normalize_apply, normalize_content, mul_one] @@ -388,8 +388,7 @@ theorem content_mul {p q : R[X]} : (p * q).content = p.content * q.content := by ← Nat.cast_withBot, ← degree_eq_natDegree q.primPart_ne_zero] at heq rw [p.eq_C_content_mul_primPart, q.eq_C_content_mul_primPart] suffices h : (q.primPart * p.primPart).content = 1 - · - rw [mul_assoc, content_C_mul, content_C_mul, mul_comm p.primPart, mul_assoc, content_C_mul, + · rw [mul_assoc, content_C_mul, content_C_mul, mul_comm p.primPart, mul_assoc, content_C_mul, content_C_mul, h, mul_one, content_primPart, content_primPart, mul_one, mul_one] rw [← normalize_content, normalize_eq_one, isUnit_iff_dvd_one, content_eq_gcd_leadingCoeff_content_eraseLead, leadingCoeff_mul, gcd_comm] diff --git a/Mathlib/RingTheory/TensorProduct.lean b/Mathlib/RingTheory/TensorProduct.lean index 4838164037fe0..06d1c49491571 100644 --- a/Mathlib/RingTheory/TensorProduct.lean +++ b/Mathlib/RingTheory/TensorProduct.lean @@ -731,11 +731,11 @@ algEquivOfLinearEquivTensorProduct f (fun x₁ x₂ c₁ c₂ => by try intros trivial - . intros ab₁ ab₂ h₁ h₂ a b + · intros ab₁ ab₂ h₁ h₂ a b rw [mul_add, add_tmul, map_add, h₁, h₂, map_add, mul_add] - . intros a b ab₁ ab₂ h₁ h₂ + · intros a b ab₁ ab₂ h₁ h₂ rw [add_mul, add_tmul, map_add, h₁, h₂, map_add, add_mul] - . intros ab₁ ab₂ _ _ x y hx hy + · intros ab₁ ab₂ _ _ x y hx hy rw [add_mul, add_tmul, map_add, hx, hy, map_add, map_add, mul_add, mul_add, add_mul, mul_add]) w₂ #align algebra.tensor_product.alg_equiv_of_linear_equiv_triple_tensor_product Algebra.TensorProduct.algEquivOfLinearEquivTripleTensorProduct diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain.lean b/Mathlib/RingTheory/UniqueFactorizationDomain.lean index 353d41a21dada..5191e523d0443 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain.lean @@ -1266,7 +1266,7 @@ theorem prod_mono : ∀ {a b : FactorSet α}, a ≤ b → a.prod ≤ b.prod theorem FactorSet.prod_eq_zero_iff [Nontrivial α] (p : FactorSet α) : p.prod = 0 ↔ p = ⊤ := by induction p using WithTop.recTopCoe · simp only [iff_self_iff, eq_self_iff_true, Associates.prod_top] - . rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, + · rw [prod_coe, Multiset.prod_eq_zero_iff, Multiset.mem_map, eq_false WithTop.coe_ne_top, iff_false_iff, not_exists] exact fun a => not_and_of_not_right _ a.prop.ne_zero #align associates.factor_set.prod_eq_zero_iff Associates.FactorSet.prod_eq_zero_iff diff --git a/Mathlib/RingTheory/WittVector/WittPolynomial.lean b/Mathlib/RingTheory/WittVector/WittPolynomial.lean index d2402f1e5ee23..83400362ee0b3 100644 --- a/Mathlib/RingTheory/WittVector/WittPolynomial.lean +++ b/Mathlib/RingTheory/WittVector/WittPolynomial.lean @@ -273,14 +273,14 @@ theorem xInTermsOfW_vars_aux (n : ℕ) : rcases H with ⟨j, hj, H⟩ rw [vars_C_mul] at H swap - . apply pow_ne_zero + · apply pow_ne_zero exact_mod_cast hp.1.ne_zero rw [mem_range] at hj replace H := (ih j hj).2 (vars_pow _ _ H) rw [mem_range] at H - . rw [mem_range] + · rw [mem_range] linarith - . linarith + · linarith set_option linter.uppercaseLean3 false in #align X_in_terms_of_W_vars_aux xInTermsOfW_vars_aux diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index dc371000d85f3..eb78c82c9c1b7 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1958,7 +1958,7 @@ theorem toPartENat_lift (c : Cardinal.{v}) : toPartENat (lift.{u, v} c) = toPart simp only [toNat_lift] rw [lift_lt_aleph0] exact hc - . rw [toPartENat_apply_of_aleph0_le hc, toPartENat_apply_of_aleph0_le _] + · rw [toPartENat_apply_of_aleph0_le hc, toPartENat_apply_of_aleph0_le _] rw [aleph0_le_lift] exact hc #align cardinal.to_part_enat_lift Cardinal.toPartENat_lift diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 87337162fa9b0..32194fb94773a 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -529,7 +529,7 @@ theorem cof_eq_one_iff_is_succ {o} : cof.{u} o = 1 ↔ ∃ a, o = succ a := simp [Subrel, Order.Preimage, EmptyRelation] exact x.2 · suffices : r x a ∨ ∃ _ : PUnit.{u}, ↑a = x - . convert this + · convert this dsimp [RelEmbedding.ofMonotone]; simp rcases trichotomous_of r x a with (h | h | h) · exact Or.inl h diff --git a/Mathlib/SetTheory/Cardinal/Ordinal.lean b/Mathlib/SetTheory/Cardinal/Ordinal.lean index 4fb76d149b3b0..1e184ccfd73a5 100644 --- a/Mathlib/SetTheory/Cardinal/Ordinal.lean +++ b/Mathlib/SetTheory/Cardinal/Ordinal.lean @@ -528,8 +528,7 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by typein_inj, mem_setOf_eq] at h exact max_le_iff.1 (le_iff_lt_or_eq.2 <| h.imp_right And.left) suffices H : (insert (g p) { x | r x (g p) } : Set α) ≃ Sum { x | r x (g p) } PUnit - · - exact + · exact ⟨(Set.embeddingOfSubset _ _ this).trans ((Equiv.Set.prod _ _).trans (H.prodCongr H)).toEmbedding⟩ refine' (Equiv.Set.insert _).trans ((Equiv.refl _).sumCongr punitEquivPUnit) @@ -1032,8 +1031,7 @@ theorem mk_finset_of_infinite (α : Type u) [Infinite α] : #(Finset α) = #α : theorem mk_finsupp_lift_of_infinite (α : Type u) (β : Type v) [Infinite α] [Zero β] [Nontrivial β] : #(α →₀ β) = max (lift.{v} #α) (lift.{u} #β) := by apply le_antisymm - · - calc + · calc #(α →₀ β) ≤ #(Finset (α × β)) := mk_le_of_injective (Finsupp.graph_injective α β) _ = #(α × β) := mk_finset_of_infinite _ _ = max (lift.{v} #α) (lift.{u} #β) := diff --git a/Mathlib/SetTheory/Lists.lean b/Mathlib/SetTheory/Lists.lean index ec4ce866511b9..f971dffea6286 100644 --- a/Mathlib/SetTheory/Lists.lean +++ b/Mathlib/SetTheory/Lists.lean @@ -301,7 +301,7 @@ theorem Equiv.antisymm_iff {l₁ l₂ : Lists' α true} : of' l₁ ~ of' l₂ refine' ⟨fun h => _, fun ⟨h₁, h₂⟩ => Equiv.antisymm h₁ h₂⟩ cases' h with _ _ _ h₁ h₂ · simp [Lists'.Subset.refl] - . exact ⟨h₁, h₂⟩ + · exact ⟨h₁, h₂⟩ #align lists.equiv.antisymm_iff Lists.Equiv.antisymm_iff attribute [refl] Equiv.refl @@ -347,9 +347,9 @@ theorem Equiv.trans : ∀ {l₁ l₂ l₃ : Lists α}, l₁ ~ l₂ → l₂ ~ l -- Assumption fails. simp only [Lists'.toList, Sigma.eta, List.find?, List.mem_cons, forall_eq_or_imp] constructor - . intros l₂ l₃ h₁ h₂ + · intros l₂ l₃ h₁ h₂ exact IH₁ h₁ h₂ - . intros a h₁ l₂ l₃ h₂ h₃ + · intros a h₁ l₂ l₃ h₂ h₃ exact IH _ h₁ h₂ h₃ #align lists.equiv.trans Lists.Equiv.trans diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 62f8fb71c0a69..b151d561c1383 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -910,8 +910,7 @@ theorem le_div {a b c : Ordinal} (c0 : c ≠ 0) : a ≤ b / c ↔ c * a ≤ b := · simp only [mul_zero, Ordinal.zero_le] · intros rw [succ_le_iff, lt_div c0] - · - simp (config := { contextual := true }) only [mul_le_of_limit, limit_le, iff_self_iff, + · simp (config := { contextual := true }) only [mul_le_of_limit, limit_le, iff_self_iff, forall_true_iff] #align ordinal.le_div Ordinal.le_div diff --git a/Mathlib/SetTheory/Ordinal/Exponential.lean b/Mathlib/SetTheory/Ordinal/Exponential.lean index f5f2455514b1f..fcc3652431aa1 100644 --- a/Mathlib/SetTheory/Ordinal/Exponential.lean +++ b/Mathlib/SetTheory/Ordinal/Exponential.lean @@ -156,8 +156,7 @@ theorem opow_le_opow_left {a b : Ordinal} (c) (ab : a ≤ b) : (a^c) ≤ (b^c) : · simp only [opow_zero, le_refl] · intro c IH simpa only [opow_succ] using mul_le_mul' IH ab - · - exact fun c l IH => + · exact fun c l IH => (opow_le_of_limit a0 l).2 fun b' h => (IH _ h).trans (opow_le_opow_right ((Ordinal.pos_iff_ne_zero.2 a0).trans_le ab) h.le) #align ordinal.opow_le_opow_left Ordinal.opow_le_opow_left diff --git a/Mathlib/SetTheory/Ordinal/Notation.lean b/Mathlib/SetTheory/Ordinal/Notation.lean index bdcc465c7c600..7002c8f958347 100644 --- a/Mathlib/SetTheory/Ordinal/Notation.lean +++ b/Mathlib/SetTheory/Ordinal/Notation.lean @@ -341,21 +341,21 @@ theorem cmp_compares : ∀ (a b : ONote) [NF a] [NF b], (cmp a b).Compares a b rw [cmpUsing, ite_eq_iff, not_lt] at nh case lt => cases' nh with nh nh - . exact oadd_lt_oadd_2 h₁ nh.left - . rw [ite_eq_iff] at nh; cases' nh.right with nh nh <;> cases nh <;> contradiction + · exact oadd_lt_oadd_2 h₁ nh.left + · rw [ite_eq_iff] at nh; cases' nh.right with nh nh <;> cases nh <;> contradiction case gt => cases' nh with nh nh - . cases nh; contradiction - . cases' nh with _ nh + · cases nh; contradiction + · cases' nh with _ nh rw [ite_eq_iff] at nh; cases' nh with nh nh - . exact oadd_lt_oadd_2 h₂ nh.left - . cases nh; contradiction + · exact oadd_lt_oadd_2 h₂ nh.left + · cases nh; contradiction cases' nh with nh nh - . cases nh; contradiction + · cases nh; contradiction cases' nh with nhl nhr rw [ite_eq_iff] at nhr cases' nhr with nhr nhr - . cases nhr; contradiction + · cases nhr; contradiction obtain rfl := Subtype.eq (eq_of_incomp ⟨(not_lt_of_ge nhl), nhr.left⟩) have IHa := @cmp_compares _ _ h₁.snd h₂.snd revert IHa; cases cmp a₁ a₂ <;> intro IHa <;> dsimp at IHa @@ -540,7 +540,7 @@ theorem repr_sub : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ - o₂) = rep have ee := @cmp_compares _ _ h₁.fst h₂.fst cases h : cmp e₁ e₂ <;> simp only [h] at ee · rw [Ordinal.sub_eq_zero_iff_le.2] - . rfl + · rfl exact le_of_lt (oadd_lt_oadd_1 h₁ ee) · change e₁ = e₂ at ee subst e₂ @@ -817,10 +817,10 @@ instance nf_opowAux (e a0 a) [NF e] [NF a0] [NF a] : ∀ k m, NF (opowAux e a0 a intro k m unfold opowAux cases' m with m m - . cases k <;> exact NF.zero + · cases k <;> exact NF.zero cases' k with k k - . exact NF.oadd_zero _ _ - . haveI := nf_opowAux e a0 a k + · exact NF.oadd_zero _ _ + · haveI := nf_opowAux e a0 a k simp only [Nat.succ_ne_zero m] ; infer_instance #align onote.NF_opow_aux ONote.nf_opowAux @@ -839,8 +839,8 @@ instance nf_opow (o₁ o₂) [NF o₁] [NF o₂] : NF (o₁ ^ o₂) := by · simp [(· ^ ·),Pow.pow,pow, opow, opowAux2, e₁, e₂, split_eq_scale_split' e₂] have := na.fst cases' k with k <;> simp <;> skip <;> dsimp - . infer_instance - . cases k <;> cases m <;> infer_instance + · infer_instance + · cases k <;> cases m <;> infer_instance #align onote.NF_opow ONote.nf_opow theorem scale_opowAux (e a0 a : ONote) [NF e] [NF a0] [NF a] : diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index b794d53206755..ac0649adb4513 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -697,8 +697,8 @@ theorem tsum_iSup_decode₂ [CompleteLattice β] (m : β → α) (m0 : m ⊥ = 0 intro n h generalize decode₂ γ n = foo at * cases' foo with b - . refine' (h <| by simp [m0]).elim - . exact rfl + · refine' (h <| by simp [m0]).elim + · exact rfl symm refine' tsum_eq_tsum_of_ne_zero_bij (fun a => Option.get _ (H a.1 a.2)) _ _ _ · dsimp only [] @@ -842,7 +842,7 @@ theorem HasSum.update (hf : HasSum f a₁) (b : β) [DecidableEq β] (a : α) : by_cases h : b' = b · rw [h, update_same] simp [eq_self_iff_true, if_true, sub_add_cancel] - . simp only [h, update_noteq, if_false, Ne.def, zero_add, not_false_iff] + · simp only [h, update_noteq, if_false, Ne.def, zero_add, not_false_iff] #align has_sum.update HasSum.update theorem Summable.update (hf : Summable f) (b : β) [DecidableEq β] (a : α) : diff --git a/Mathlib/Topology/Category/TopCat/Opens.lean b/Mathlib/Topology/Category/TopCat/Opens.lean index 6628717490148..6ac7564f09de1 100644 --- a/Mathlib/Topology/Category/TopCat/Opens.lean +++ b/Mathlib/Topology/Category/TopCat/Opens.lean @@ -355,10 +355,10 @@ theorem adjunction_counit_app_self {X : TopCat} (U : Opens X) : theorem inclusion_top_functor (X : TopCat) : (@Opens.openEmbedding X ⊤).isOpenMap.functor = map (inclusionTopIso X).inv := by refine' CategoryTheory.Functor.ext _ _ - . intro U + · intro U ext x exact ⟨fun ⟨⟨_, _⟩, h, rfl⟩ => h, fun h => ⟨⟨x, trivial⟩, h, rfl⟩⟩ - . intros U V f + · intros U V f apply Subsingleton.elim #align topological_space.opens.inclusion_top_functor TopologicalSpace.Opens.inclusion_top_functor @@ -377,9 +377,9 @@ lemma set_range_forget_map_inclusion {X : TopCat} (U : Opens X) : Set.range ((forget TopCat).map (inclusion U)) = (U : Set X) := by ext x constructor - . rintro ⟨x, rfl⟩ + · rintro ⟨x, rfl⟩ exact x.2 - . intro h + · intro h exact ⟨⟨x, h⟩, rfl⟩ @[simp] diff --git a/Mathlib/Topology/Instances/RatLemmas.lean b/Mathlib/Topology/Instances/RatLemmas.lean index 7a27bf59f5976..2bcc661a48494 100644 --- a/Mathlib/Topology/Instances/RatLemmas.lean +++ b/Mathlib/Topology/Instances/RatLemmas.lean @@ -86,7 +86,7 @@ instance : TotallyDisconnectedSpace ℚ := by refine' ⟨fun s hsu hs x hx y hy => _⟩; clear hsu by_contra' H : x ≠ y wlog hlt : x < y - . refine' this s hs y hy x hx H.symm <| H.lt_or_lt.resolve_left hlt <;> assumption + · refine' this s hs y hy x hx H.symm <| H.lt_or_lt.resolve_left hlt <;> assumption rcases exists_irrational_btwn (Rat.cast_lt.2 hlt) with ⟨z, hz, hxz, hzy⟩ have := hs.image _ continuous_coe_real.continuousOn rw [isPreconnected_iff_ordConnected] at this diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 0eee480494ce1..a584e9bbf929d 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -150,10 +150,10 @@ theorem LipschitzOnWith.extend_lp_infty [PseudoMetricSpace α] {s : Set α} {f : exact LipschitzOnWith.extend_real (hfl i) -- use the nonlinear Hahn-Banach theorem here! choose g hgl hgeq using this rcases s.eq_empty_or_nonempty with rfl | ⟨a₀, ha₀_in_s⟩ - . exact ⟨0, LipschitzWith.const' 0, by simp⟩ + · exact ⟨0, LipschitzWith.const' 0, by simp⟩ · -- Show that the extensions are uniformly bounded have hf_extb : ∀ a : α, Memℓp (swap g a) ∞ - . apply LipschitzWith.uniformly_bounded (swap g) hgl a₀ + · apply LipschitzWith.uniformly_bounded (swap g) hgl a₀ use ‖f a₀‖ rintro - ⟨i, rfl⟩ simp_rw [←hgeq i ha₀_in_s] diff --git a/Mathlib/Topology/Sheaves/Functors.lean b/Mathlib/Topology/Sheaves/Functors.lean index 9c42ad85beadb..f94b7b6065d24 100644 --- a/Mathlib/Topology/Sheaves/Functors.lean +++ b/Mathlib/Topology/Sheaves/Functors.lean @@ -46,7 +46,7 @@ theorem map_diagram : Pairwise.diagram U ⋙ Opens.map f = Pairwise.diagram ((Opens.map f).obj ∘ U) := by have obj_eq : ∀ (j : Pairwise ι), (Pairwise.diagram U ⋙ Opens.map f).obj j = (Pairwise.diagram ((Opens.map f).toPrefunctor.obj ∘ U)).obj j - . rintro ⟨i⟩ <;> rfl + · rintro ⟨i⟩ <;> rfl refine Functor.hext obj_eq ?_ intro i j g; apply Subsingleton.helim rw [obj_eq, obj_eq] diff --git a/Mathlib/Topology/Sheaves/LocalPredicate.lean b/Mathlib/Topology/Sheaves/LocalPredicate.lean index 3a2bf05f874f3..d8c06528e4016 100644 --- a/Mathlib/Topology/Sheaves/LocalPredicate.lean +++ b/Mathlib/Topology/Sheaves/LocalPredicate.lean @@ -299,8 +299,8 @@ theorem stalkToFiber_injective (P : LocalPredicate T) (x : X) -- and put it back together again in the correct order. refine' ⟨op W, fun w => fU (iU w : (unop U).1), P.res _ _ hU, _⟩ rcases W with ⟨W, m⟩ - . exact iU - . exact ⟨colimit_sound iU.op (Subtype.eq rfl), colimit_sound iV.op (Subtype.eq (funext w).symm)⟩ + · exact iU + · exact ⟨colimit_sound iU.op (Subtype.eq rfl), colimit_sound iV.op (Subtype.eq (funext w).symm)⟩ set_option linter.uppercaseLean3 false in #align Top.stalk_to_fiber_injective TopCat.stalkToFiber_injective diff --git a/Mathlib/Topology/Sheaves/Operations.lean b/Mathlib/Topology/Sheaves/Operations.lean index ccde073d32f3a..f816b687c5175 100644 --- a/Mathlib/Topology/Sheaves/Operations.lean +++ b/Mathlib/Topology/Sheaves/Operations.lean @@ -121,7 +121,7 @@ instance (F : X.Sheaf CommRingCat.{w}) : Mono F.presheaf.toTotalQuotientPresheaf -- Porting note : was an `apply (config := { instances := false })` -- See https://github.com/leanprover/lean4/issues/2273 suffices : ∀ (U : (Opens ↑X)ᵒᵖ), Mono (F.presheaf.toTotalQuotientPresheaf.app U) - . apply NatTrans.mono_of_mono_app + · apply NatTrans.mono_of_mono_app intro U apply ConcreteCategory.mono_of_injective dsimp [toTotalQuotientPresheaf, CommRingCat.ofHom] diff --git a/Mathlib/Topology/Sheaves/Presheaf.lean b/Mathlib/Topology/Sheaves/Presheaf.lean index 0074c99d17204..b6a29327ade8d 100644 --- a/Mathlib/Topology/Sheaves/Presheaf.lean +++ b/Mathlib/Topology/Sheaves/Presheaf.lean @@ -327,10 +327,10 @@ def pullbackObjObjOfImageOpen {X Y : TopCat.{v}} (f : X ⟶ Y) (ℱ : Y.Presheaf { lift := fun s ↦ by fapply CostructuredArrow.homMk change op (unop _) ⟶ op (⟨_, H⟩ : Opens _) - . refine' (homOfLE _).op + · refine' (homOfLE _).op apply (Set.image_subset f s.pt.hom.unop.le).trans exact Set.image_preimage.l_u_le (SetLike.coe s.pt.left.unop) - . simp + · simp -- porting note : add `fac`, `uniq` manually fac := fun _ _ => by ext; simp uniq := fun _ _ _ => by ext; simp } @@ -398,7 +398,7 @@ theorem id_pushforward {X : TopCat.{w}} : pushforward C (𝟙 X) = 𝟭 (X.Presh apply CategoryTheory.Functor.ext · intros a b f ext U - . erw [NatTrans.congr f (Opens.op_map_id_obj (op U))] + · erw [NatTrans.congr f (Opens.op_map_id_obj (op U))] simp only [Functor.op_obj, eqToHom_refl, CategoryTheory.Functor.map_id, Category.comp_id, Category.id_comp, Functor.id_obj, Functor.id_map] apply Pushforward.id_eq diff --git a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean index 40a6a3c9cd54f..d09e301f99dcd 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/PairwiseIntersections.lean @@ -378,7 +378,7 @@ def interUnionPullbackConeLift : s.pt ⟶ F.1.obj (op (U ⊔ V)) := by rcases g with ⟨⟩ <;> dsimp [Pairwise.diagram] <;> simp only [Category.id_comp, s.condition, CategoryTheory.Functor.map_id, Category.comp_id] - . rw [← cancel_mono (F.1.map (eqToHom <| inf_comm : U ⊓ V ⟶ _).op), Category.assoc, + · rw [← cancel_mono (F.1.map (eqToHom <| inf_comm : U ⊓ V ⟶ _).op), Category.assoc, Category.assoc, ←F.1.map_comp, ←F.1.map_comp] exact s.condition.symm set_option linter.uppercaseLean3 false in diff --git a/Mathlib/Topology/Sheaves/Skyscraper.lean b/Mathlib/Topology/Sheaves/Skyscraper.lean index f042497d3124b..6637607f45aae 100644 --- a/Mathlib/Topology/Sheaves/Skyscraper.lean +++ b/Mathlib/Topology/Sheaves/Skyscraper.lean @@ -235,8 +235,8 @@ theorem skyscraperPresheaf_isSheaf : (skyscraperPresheaf p₀ A).IsSheaf := by (Presheaf.isSheaf_on_pUnit_of_isTerminal _ (by dsimp [skyscraperPresheaf] rw [if_neg] - . exact terminalIsTerminal - . exact Set.not_mem_empty PUnit.unit))) + · exact terminalIsTerminal + · exact Set.not_mem_empty PUnit.unit))) #align skyscraper_presheaf_is_sheaf skyscraperPresheaf_isSheaf /-- diff --git a/Mathlib/Topology/UniformSpace/Completion.lean b/Mathlib/Topology/UniformSpace/Completion.lean index fad2ddb7316b9..3cb5a07217caa 100644 --- a/Mathlib/Topology/UniformSpace/Completion.lean +++ b/Mathlib/Topology/UniformSpace/Completion.lean @@ -192,10 +192,10 @@ theorem denseRange_pureCauchy : DenseRange (pureCauchy : α → CauchyFilter α) simp only [closure_eq_cluster_pts, ClusterPt, nhds_eq_uniformity, lift'_inf_principal_eq, Set.inter_comm _ (range pureCauchy), mem_setOf_eq] refine (lift'_neBot_iff ?_).mpr (fun s hs => ?_) - . refine monotone_const.inter ?_ + · refine monotone_const.inter ?_ simp_rw [UniformSpace.ball] exact monotone_preimage - . let ⟨y, hy⟩ := h_ex s hs + · let ⟨y, hy⟩ := h_ex s hs have : pureCauchy y ∈ range pureCauchy ∩ { y : CauchyFilter α | (f, y) ∈ s } := ⟨mem_range_self y, hy⟩ exact ⟨_, this⟩ @@ -644,16 +644,16 @@ def completionSeparationQuotientEquiv (α : Type u) [UniformSpace α] : -- porting note: had to insert rewrites to switch between Quot.mk, Quotient.mk, Quotient.mk' rw [← Quotient.mk,extension_coe (SeparationQuotient.uniformContinuous_lift _), SeparationQuotient.lift_mk (uniformContinuous_coe α), map_coe] - . rfl - . exact uniformContinuous_quotient_mk + · rfl + · exact uniformContinuous_quotient_mk · intro a refine' Completion.induction_on a (isClosed_eq (continuous_extension.comp continuous_map) continuous_id) fun a => _ rw [map_coe] -- porting note: add SeparationQuotient.lift_mk' for Quotient.mk' ? - . rw [extension_coe (SeparationQuotient.uniformContinuous_lift _), Quotient.mk', + · rw [extension_coe (SeparationQuotient.uniformContinuous_lift _), Quotient.mk', SeparationQuotient.lift_mk (uniformContinuous_coe α) _] - . exact uniformContinuous_quotient_mk + · exact uniformContinuous_quotient_mk #align uniform_space.completion.completion_separation_quotient_equiv UniformSpace.Completion.completionSeparationQuotientEquiv theorem uniformContinuous_completionSeparationQuotientEquiv : diff --git a/scripts/lint-style.py b/scripts/lint-style.py index f712c20e99e33..083c8819b20b7 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -47,6 +47,7 @@ ERR_TAC = 9 # imported Mathlib.Tactic ERR_UNF = 10 # unfreeze_local_instances ERR_IBY = 11 # isolated by +ERR_DOT = 12 # isolated or low focusing dot exceptions = [] @@ -231,7 +232,7 @@ def banned_import_check(lines, path): errors += [(ERR_TAC, line_nr, path)] return errors -def isolated_by_check(lines, path): +def isolated_by_dot_check(lines, path): errors = [] for line_nr, line in enumerate(lines, 1): if line.strip() == "by": @@ -241,6 +242,8 @@ def isolated_by_check(lines, path): prev_line = lines[line_nr - 2].rstrip() if not prev_line.endswith(",") and not re.search(", fun [^,]* =>$", prev_line): errors += [(ERR_IBY, line_nr, path)] + if line.lstrip().startswith(". ") or line.strip() in (".", "·"): + errors += [(ERR_DOT, line_nr, path)] return errors def output_message(path, line_nr, code, msg): @@ -283,13 +286,15 @@ def format_errors(errors): output_message(path, line_nr, "ERR_TAC", "Files in mathlib cannot import the whole tactic folder") if errno == ERR_IBY: output_message(path, line_nr, "ERR_IBY", "Line is an isolated 'by'") + if errno == ERR_DOT: + output_message(path, line_nr, "ERR_DOT", "Line is an isolated focusing dot or uses . instead of ·") def lint(path): with path.open(encoding="utf-8") as f: lines = f.readlines() errs = long_lines_check(lines, path) format_errors(errs) - errs = isolated_by_check(lines, path) + errs = isolated_by_dot_check(lines, path) format_errors(errs) (b, errs) = import_only_check(lines, path) if b: diff --git a/test/interval_cases.lean b/test/interval_cases.lean index 9095332f3530c..ef7a8d848b8fd 100644 --- a/test/interval_cases.lean +++ b/test/interval_cases.lean @@ -27,59 +27,59 @@ example (n : ℕ) (w₂ : n < 1) : n = 0 := by example (n : ℕ) (w₂ : n < 2) : n = 0 ∨ n = 1 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : 1 ≤ n) (w₂ : n < 3) : n = 1 ∨ n = 2 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : 1 ≤ n) (w₂ : n < 3) : n = 1 ∨ n = 2 := by interval_cases using w₁, w₂ - . left; rfl - . right; rfl + · left; rfl + · right; rfl -- make sure we only pick up bounds on the specified variable: example (n m : ℕ) (w₁ : 1 ≤ n) (w₂ : n < 3) (_ : m < 2) : n = 1 ∨ n = 2 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : 1 < n) (w₂ : n < 4) : n = 2 ∨ n = 3 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : n ≥ 3) (w₂ : n < 5) : n = 3 ∨ n = 4 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₀ : n ≥ 2) (w₁ : n ≥ 3) (w₂ : n < 5) : n = 3 ∨ n = 4 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : n > 2) (w₂ : n < 5) : n = 3 ∨ n = 4 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : n > 2) (w₂ : n ≤ 4) : n = 3 ∨ n = 4 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (w₁ : 2 < n) (w₂ : 4 ≥ n) : n = 3 ∨ n = 4 := by interval_cases n - . left; rfl - . right; rfl + · left; rfl + · right; rfl example (n : ℕ) (h1 : 4 < n) (h2 : n ≤ 6) : n < 20 := by interval_cases n - . guard_target =ₛ 5 < 20; norm_num - . guard_target =ₛ 6 < 20; norm_num + · guard_target =ₛ 5 < 20; norm_num + · guard_target =ₛ 6 < 20; norm_num example (n : ℕ) (w₁ : n % 3 < 1) : n % 3 = 0 := by interval_cases h : n % 3 @@ -124,15 +124,15 @@ example (z : ℤ) (h1 : z ≥ -3) (h2 : z < 2) : z < 20 := by example (z : ℤ) (h1 : z ≥ -3) (h2 : z < 2) : z < 20 := by interval_cases z - . guard_target =ₛ (-3 : ℤ) < 20 + · guard_target =ₛ (-3 : ℤ) < 20 norm_num - . guard_target =ₛ (-2 : ℤ) < 20 + · guard_target =ₛ (-2 : ℤ) < 20 norm_num - . guard_target =ₛ (-1 : ℤ) < 20 + · guard_target =ₛ (-1 : ℤ) < 20 norm_num - . guard_target =ₛ (0 : ℤ) < 20 + · guard_target =ₛ (0 : ℤ) < 20 norm_num - . guard_target =ₛ (1 : ℤ) < 20 + · guard_target =ₛ (1 : ℤ) < 20 norm_num example (n : ℕ) : n % 2 = 0 ∨ n % 2 = 1 := by @@ -140,9 +140,9 @@ example (n : ℕ) : n % 2 = 0 ∨ n % 2 = 1 := by have h2 : r < 2 := by exact Nat.mod_lt _ (by decide) interval_cases hrv : r -- Porting note: old syntax was `interval_cases r with hrv` - . left; exact hrv.symm.trans hrv + · left; exact hrv.symm.trans hrv --^ hover says `hrv : r = 0` and jumps to `hrv :` above - . right; exact hrv.symm.trans hrv + · right; exact hrv.symm.trans hrv --^ hover says `hrv : r = 1` and jumps to `hrv :` above /- https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/interval_cases.20bug -/ @@ -151,8 +151,8 @@ example {x : ℕ} (hx2 : x < 2) (h : False) : False := by -- `interval_cases` deliberately not focussed, -- this is testing that the `interval_cases` only acts on `have` side goal, not on both interval_cases x - . exact zero_le_one - . rfl -- done for free in the mathlib3 version + · exact zero_le_one + · rfl -- done for free in the mathlib3 version exact h /-