Skip to content

Commit 8214f46

Browse files
Ruben-VandeVeldeParcly-TaxelChrisHughes24
committed
chore: tidy various files (#4304)
Co-authored-by: Jeremy Tan Jie Rui <reddeloostw@gmail.com> Co-authored-by: Chris Hughes <chrishughes24@gmail.com>
1 parent 518697a commit 8214f46

File tree

24 files changed

+123
-199
lines changed

24 files changed

+123
-199
lines changed

Mathlib/Algebra/Category/GroupCat/Limits.lean

Lines changed: 26 additions & 32 deletions
Original file line numberDiff line numberDiff line change
@@ -24,9 +24,7 @@ the underlying types are just the limits in the category of types.
2424
-/
2525

2626

27-
open CategoryTheory
28-
29-
open CategoryTheory.Limits
27+
open CategoryTheory CategoryTheory.Limits
3028

3129
universe v u
3230

@@ -52,15 +50,12 @@ set_option linter.uppercaseLean3 false in
5250
set_option linter.uppercaseLean3 false in
5351
#align AddGroup.add_group_obj AddGroupCat.addGroupObj
5452

55-
/-- The flat sections of a functor into `Group` form a subgroup of all sections.
53+
/-- The flat sections of a functor into `GroupCat` form a subgroup of all sections.
5654
-/
5755
@[to_additive
58-
"The flat sections of a functor into `AddGroup` form an additive subgroup of all sections."]
56+
"The flat sections of a functor into `AddGroupCat` form an additive subgroup of all sections."]
5957
def sectionsSubgroup (F : J ⥤ GroupCat) : Subgroup (∀ j, F.obj j) :=
60-
{
61-
MonCat.sectionsSubmonoid
62-
(F ⋙ forget₂ GroupCat
63-
MonCat) with
58+
{ MonCat.sectionsSubmonoid (F ⋙ forget₂ GroupCat MonCat) with
6459
carrier := (F ⋙ forget GroupCat).sections
6560
inv_mem' := fun {a} ah j j' f => by
6661
simp only [Functor.comp_map, Pi.inv_apply, MonoidHom.map_inv, inv_inj]
@@ -81,17 +76,17 @@ set_option linter.uppercaseLean3 false in
8176
set_option linter.uppercaseLean3 false in
8277
#align AddGroup.limit_add_group AddGroupCat.limitAddGroup
8378

84-
/-- We show that the forgetful functor `GroupMon` creates limits.
79+
/-- We show that the forgetful functor `GroupCatMonCat` creates limits.
8580
86-
All we need to do is notice that the limit point has a `group` instance available, and then reuse
81+
All we need to do is notice that the limit point has a `Group` instance available, and then reuse
8782
the existing limit. -/
88-
@[to_additive "We show that the forgetful functor `AddGroupAddMon` creates limits.
83+
@[to_additive "We show that the forgetful functor `AddGroupCatAddMonCat` creates limits.
8984
90-
All we need to do is notice that the limit point has an `add_group` instance available, and then
85+
All we need to do is notice that the limit point has an `AddGroup` instance available, and then
9186
reuse the existing limit."]
9287
noncomputable instance Forget₂.createsLimit (F : J ⥤ GroupCatMax.{v, u}) :
9388
CreatesLimit F (forget₂ GroupCatMax.{v, u} MonCatMax.{v, u}) :=
94-
-- Porting note: need to add `forget₂ Grp Mon` reflects isomorphism
89+
-- Porting note: need to add `forget₂ GrpCat MonCat` reflects isomorphism
9590
letI : ReflectsIsomorphisms (forget₂ GroupCatMax.{v, u} MonCatMax.{v, u}) :=
9691
CategoryTheory.reflectsIsomorphisms_forget₂ _ _
9792
createsLimitOfReflectsIso (K := F) (F := (forget₂ GroupCat.{max v u} MonCat.{max v u}))
@@ -112,10 +107,10 @@ set_option linter.uppercaseLean3 false in
112107
set_option linter.uppercaseLean3 false in
113108
#align AddGroup.forget₂.creates_limit AddGroupCat.Forget₂.createsLimit
114109

115-
/-- A choice of limit cone for a functor into `Group`.
110+
/-- A choice of limit cone for a functor into `GroupCat`.
116111
(Generally, you'll just want to use `limit F`.)
117112
-/
118-
@[to_additive "A choice of limit cone for a functor into `Group`.
113+
@[to_additive "A choice of limit cone for a functor into `GroupCat`.
119114
(Generally, you'll just want to use `limit F`.)"]
120115
noncomputable def limitCone (F : J ⥤ GroupCatMax.{v, u}) : Cone F :=
121116
-- Porting note: add this instance to help Lean unify universe levels
@@ -172,8 +167,8 @@ This means the underlying monoid of a limit can be computed as a limit in the ca
172167
This means the underlying additive monoid of a limit can be computed as a limit in the category of
173168
additive monoids."]
174169
noncomputable instance forget₂MonPreservesLimitsOfSize :
175-
PreservesLimitsOfSize.{v, v} (forget₂ GroupCatMax.{v, u} MonCat.{max v u})
176-
where preservesLimitsOfShape {J _} := { preservesLimit := fun {F} =>
170+
PreservesLimitsOfSize.{v, v} (forget₂ GroupCatMax.{v, u} MonCat.{max v u}) where
171+
preservesLimitsOfShape {J _} := { preservesLimit := fun {F} =>
177172
-- Porting note: add this instance to help Lean unify universe levels
178173
letI : HasLimit (F ⋙ forget₂ GroupCatMax.{v, u} MonCat.{max v u}) :=
179174
(MonCat.hasLimitsOfSize.{v, u}.1 J).1 _
@@ -255,14 +250,14 @@ set_option linter.uppercaseLean3 false in
255250
set_option linter.uppercaseLean3 false in
256251
#align AddCommGroup.limit_add_comm_group AddCommGroupCat.limitAddCommGroup
257252

258-
/-- We show that the forgetful functor `CommGroupGroup` creates limits.
253+
/-- We show that the forgetful functor `CommGroupCatGroupCat` creates limits.
259254
260-
All we need to do is notice that the limit point has a `comm_group` instance available,
255+
All we need to do is notice that the limit point has a `CommGroup` instance available,
261256
and then reuse the existing limit.
262257
-/
263-
@[to_additive "We show that the forgetful functor `AddCommGroupAddGroup` creates limits.
258+
@[to_additive "We show that the forgetful functor `AddCommGroupCatAddGroupCat` creates limits.
264259
265-
All we need to do is notice that the limit point has an `add_comm_group` instance available,
260+
All we need to do is notice that the limit point has an `AddCommGroup` instance available,
266261
and then reuse the existing limit."]
267262
noncomputable instance Forget₂.createsLimit (F : J ⥤ CommGroupCatMax.{v, u}) :
268263
CreatesLimit F (forget₂ CommGroupCat GroupCatMax.{v, u}) :=
@@ -284,11 +279,11 @@ set_option linter.uppercaseLean3 false in
284279
set_option linter.uppercaseLean3 false in
285280
#align AddCommGroup.forget₂.creates_limit AddCommGroupCat.Forget₂.createsLimit
286281

287-
/-- A choice of limit cone for a functor into `CommGroup`.
282+
/-- A choice of limit cone for a functor into `CommGroupCat`.
288283
(Generally, you'll just want to use `limit F`.)
289284
-/
290285
@[to_additive
291-
"A choice of limit cone for a functor into `CommGroup`.
286+
"A choice of limit cone for a functor into `AddCommGroupCat`.
292287
(Generally, you'll just want to use `limit F`.)"]
293288
noncomputable def limitCone (F : J ⥤ CommGroupCat.{max v u}) : Cone F :=
294289
liftLimit (limit.isLimit (F ⋙ forget₂ CommGroupCatMax.{v, u} GroupCatMax.{v, u}))
@@ -302,7 +297,7 @@ set_option linter.uppercaseLean3 false in
302297
-/
303298
@[to_additive
304299
"The chosen cone is a limit cone.
305-
(Generally, you'll just wantto use `limit.cone F`.)"]
300+
(Generally, you'll just want to use `limit.cone F`.)"]
306301
noncomputable def limitConeIsLimit (F : J ⥤ CommGroupCatMax.{v, u}) :
307302
IsLimit (limitCone.{v, u} F) :=
308303
liftedLimitIsLimit _
@@ -376,10 +371,9 @@ in the category of commutative monoids.)
376371
preserves all limits. (That is, the underlying additive commutative monoids could have been
377372
computed instead as limits in the category of additive commutative monoids.)"]
378373
noncomputable instance forget₂CommMonPreservesLimitsOfSize :
379-
PreservesLimitsOfSize.{v, v} (forget₂ CommGroupCat CommMonCat.{max v u})
380-
where preservesLimitsOfShape :=
381-
{
382-
preservesLimit := fun {F} =>
374+
PreservesLimitsOfSize.{v, v} (forget₂ CommGroupCat CommMonCat.{max v u}) where
375+
preservesLimitsOfShape :=
376+
{ preservesLimit := fun {F} =>
383377
preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F)
384378
(forget₂CommMonPreservesLimitsAux.{v, u} F) }
385379
set_option linter.uppercaseLean3 false in
@@ -395,8 +389,8 @@ underlying types could have been computed instead as limits in the category of t
395389
(That is, the underlying types could have been computed instead as limits in the category of
396390
types.)"]
397391
noncomputable instance forgetPreservesLimitsOfSize :
398-
PreservesLimitsOfSize.{v, v} (forget CommGroupCatMax.{v, u})
399-
where preservesLimitsOfShape {J _} :=
392+
PreservesLimitsOfSize.{v, v} (forget CommGroupCatMax.{v, u}) where
393+
preservesLimitsOfShape {J _} :=
400394
{ preservesLimit := fun {F} =>
401395
-- Porting note : add these instance to help Lean unify universe levels
402396
letI : HasLimit (F ⋙ forget₂ CommGroupCatMax.{v, u} GroupCat.{max v u}) :=
@@ -416,7 +410,7 @@ end CommGroupCat
416410

417411
namespace AddCommGroupCat
418412

419-
/-- The categorical kernel of a morphism in `AddCommGroup`
413+
/-- The categorical kernel of a morphism in `AddCommGroupCat`
420414
agrees with the usual group-theoretical kernel.
421415
-/
422416
def kernelIsoKer {G H : AddCommGroupCat.{u}} (f : G ⟶ H) :

Mathlib/Algebra/Category/Ring/Instances.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -24,11 +24,11 @@ instance localization_unit_isIso (R : CommRingCat) :
2424
IsIso.of_iso (IsLocalization.atOne R (Localization.Away (1 : R))).toRingEquiv.toCommRingCatIso
2525
#align localization_unit_is_iso localization_unit_isIso
2626

27-
instance localization_unit_is_iso' (R : CommRingCat) :
27+
instance localization_unit_isIso' (R : CommRingCat) :
2828
@IsIso CommRingCat _ R _ (CommRingCat.ofHom <| algebraMap R (Localization.Away (1 : R))) := by
2929
cases R
3030
exact localization_unit_isIso _
31-
#align localization_unit_is_iso' localization_unit_is_iso'
31+
#align localization_unit_is_iso' localization_unit_isIso'
3232

3333
theorem IsLocalization.epi {R : Type _} [CommRing R] (M : Submonoid R) (S : Type _) [CommRing S]
3434
[Algebra R S] [IsLocalization M S] : Epi (CommRingCat.ofHom <| algebraMap R S) :=

Mathlib/Algebra/Module/Bimodule.lean

Lines changed: 8 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -39,22 +39,22 @@ This file is a place to collect results which are specific to bimodules.
3939
4040
## Main definitions
4141
42-
* `subbimodule.mk`
43-
* `subbimodule.smul_mem`
44-
* `subbimodule.smul_mem'`
45-
* `subbimodule.to_submodule`
46-
* `subbimodule.to_submodule'`
42+
* `Subbimodule.mk`
43+
* `Subbimodule.smul_mem`
44+
* `Subbimodule.smul_mem'`
45+
* `Subbimodule.toSubmodule`
46+
* `Subbimodule.toSubmodule'`
4747
4848
## Implementation details
4949
5050
For many definitions and lemmas it is preferable to set things up without opposites, i.e., as:
51-
`[module S M] [smul_comm_class R S M]` rather than `[module Sᵐᵒᵖ M] [smul_comm_class R Sᵐᵒᵖ M]`.
51+
`[Module S M] [SMulCommClass R S M]` rather than `[Module Sᵐᵒᵖ M] [SMulCommClass R Sᵐᵒᵖ M]`.
5252
The corresponding results for opposites then follow automatically and do not require taking
5353
advantage of the fact that `(Sᵐᵒᵖ)ᵐᵒᵖ` is defeq to `S`.
5454
5555
## TODO
5656
57-
Develop the theory of two-sided ideals, which have type `submodule (R ⊗[ℕ] Rᵐᵒᵖ) R`.
57+
Develop the theory of two-sided ideals, which have type `Submodule (R ⊗[ℕ] Rᵐᵒᵖ) R`.
5858
5959
-/
6060

@@ -83,7 +83,7 @@ variable [SMulCommClass A B M]
8383
individually, rather than jointly via their tensor product.
8484
8585
Note that `R` plays no role but it is convenient to make this generalisation to support the cases
86-
`R = ℕ` and `R = ℤ` which both show up naturally. See also `base_change`. -/
86+
`R = ℕ` and `R = ℤ` which both show up naturally. See also `Subbimodule.baseChange`. -/
8787
@[simps]
8888
def mk (p : AddSubmonoid M) (hA : ∀ (a : A) {m : M}, m ∈ p → a • m ∈ p)
8989
(hB : ∀ (b : B) {m : M}, m ∈ p → b • m ∈ p) : Submodule (A ⊗[R] B) M :=
@@ -155,4 +155,3 @@ def toSubbimoduleNat (p : Submodule (R ⊗[ℤ] S) M) : Submodule (R ⊗[ℕ] S)
155155
end Ring
156156

157157
end Subbimodule
158-

Mathlib/Analysis/Complex/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -292,12 +292,12 @@ theorem imClm_apply (z : ℂ) : (imClm : ℂ → ℝ) z = z.im :=
292292
rfl
293293
#align complex.im_clm_apply Complex.imClm_apply
294294

295-
theorem restrictScalars_one_smul_right' (x : E) :
295+
theorem restrictScalars_one_smulRight' (x : E) :
296296
ContinuousLinearMap.restrictScalars ℝ ((1 : ℂ →L[ℂ] ℂ).smulRight x : ℂ →L[ℂ] E) =
297297
reClm.smulRight x + I • imClm.smulRight x := by
298298
ext ⟨a, b⟩
299299
simp [mk_eq_add_mul_I, mul_smul, smul_comm I b x]
300-
#align complex.restrict_scalars_one_smul_right' Complex.restrictScalars_one_smul_right'
300+
#align complex.restrict_scalars_one_smul_right' Complex.restrictScalars_one_smulRight'
301301

302302
theorem restrictScalars_one_smulRight (x : ℂ) :
303303
ContinuousLinearMap.restrictScalars ℝ ((1 : ℂ →L[ℂ] ℂ).smulRight x : ℂ →L[ℂ] ℂ) =

Mathlib/CategoryTheory/Category/Preorder.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -108,9 +108,9 @@ theorem homOfLE_leOfHom {x y : X} (h : x ⟶ y) : h.le.hom = h :=
108108
#align category_theory.hom_of_le_le_of_hom CategoryTheory.homOfLE_leOfHom
109109

110110
/-- Construct a morphism in the opposite of a preorder category from an inequality. -/
111-
def opHomOfLe {x y : Xᵒᵖ} (h : unop x ≤ unop y) : y ⟶ x :=
111+
def opHomOfLE {x y : Xᵒᵖ} (h : unop x ≤ unop y) : y ⟶ x :=
112112
(homOfLE h).op
113-
#align category_theory.op_hom_of_le CategoryTheory.opHomOfLe
113+
#align category_theory.op_hom_of_le CategoryTheory.opHomOfLE
114114

115115
theorem le_of_op_hom {x y : Xᵒᵖ} (h : x ⟶ y) : unop y ≤ unop x :=
116116
h.unop.le

Mathlib/CategoryTheory/Limits/Constructions/Over/Products.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -116,8 +116,8 @@ def conesEquivCounitIso (B : C) (F : Discrete J ⥤ Over B) :
116116
(by aesop_cat)
117117
#align category_theory.over.construct_products.cones_equiv_counit_iso CategoryTheory.Over.ConstructProducts.conesEquivCounitIso
118118

119-
-- TODO: Can we add `. obviously` to the second arguments of `nat_iso.of_components` and
120-
-- `cones.ext`?
119+
-- TODO: Can we add `:= by aesop` to the second arguments of `NatIso.ofComponents` and
120+
-- `Cones.ext`?
121121
/-- (Impl) Establish an equivalence between the category of cones for `F` and for the "grown" `F`.
122122
-/
123123
@[simps]

Mathlib/Combinatorics/SimpleGraph/Ends/Defs.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -282,7 +282,7 @@ theorem infinite_iff_in_eventualRange {K : (Finset V)ᵒᵖ} (C : G.componentCom
282282
Set.mem_range, componentComplFunctor_map]
283283
exact
284284
fun h Lop KL => h Lop.unop (le_of_op_hom KL), fun h L KL =>
285-
h (Opposite.op L) (opHomOfLe KL)⟩
285+
h (Opposite.op L) (opHomOfLE KL)⟩
286286
#align simple_graph.infinite_iff_in_eventual_range SimpleGraph.infinite_iff_in_eventualRange
287287

288288
end Ends

Mathlib/Computability/Partrec.lean

Lines changed: 18 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -178,22 +178,25 @@ theorem of_eq_tot {f : ℕ →. ℕ} {g : ℕ → ℕ} (hf : Partrec f) (H : ∀
178178
#align nat.partrec.of_eq_tot Nat.Partrec.of_eq_tot
179179

180180
theorem of_primrec {f : ℕ → ℕ} (hf : Nat.Primrec f) : Partrec f := by
181-
induction hf
182-
case zero => exact zero
183-
case succ => exact succ
184-
case left => exact left
185-
case right => exact right
186-
case pair f g _ _ pf pg =>
181+
induction hf with
182+
| zero => exact zero
183+
| succ => exact succ
184+
| left => exact left
185+
| right => exact right
186+
| pair _ _ pf pg =>
187187
refine' (pf.pair pg).of_eq_tot fun n => _
188188
simp [Seq.seq]
189-
case comp f g _ _ pf pg =>
189+
| comp _ _ pf pg =>
190190
refine' (pf.comp pg).of_eq_tot fun n => _
191191
simp
192-
case prec f g _ _ pf pg =>
192+
| prec _ _ pf pg =>
193193
refine' (pf.prec pg).of_eq_tot fun n => _
194-
simp
195-
induction' n.unpair.2 with m IH; · simp
196-
simp; exact ⟨_, IH, rfl⟩
194+
simp only [unpaired, PFun.coe_val, bind_eq_bind]
195+
induction n.unpair.2 with
196+
| zero => simp
197+
| succ m IH =>
198+
simp only [mem_bind_iff, mem_some_iff]
199+
exact ⟨_, IH, rfl⟩
197200
#align nat.partrec.of_primrec Nat.Partrec.of_primrec
198201

199202
protected theorem some : Partrec some :=
@@ -764,14 +767,10 @@ open Computable
764767

765768
theorem option_some_iff {f : α →. σ} : (Partrec fun a => (f a).map Option.some) ↔ Partrec f :=
766769
fun h => (Nat.Partrec.ppred.comp h).of_eq fun n => by
767-
simp [Part.bind_assoc]
768-
-- Porting note: `simp` can't match `Part.some ∘ f` with `fun x => Part.some (f x)`,
769-
-- so `conv` & `erw` are needed.
770-
conv_lhs =>
771-
congr
772-
· skip
773-
· ext x
774-
erw [bind_some_eq_map],
770+
-- Porting note: needed to help with applying bind_some_eq_map because `Function.comp` got
771+
-- less reducible.
772+
simp [Part.bind_assoc, ← Function.comp_apply (f := Part.some) (g := encode), bind_some_eq_map,
773+
-Function.comp_apply],
775774
fun hf => hf.map (option_some.comp snd).to₂⟩
776775
#align partrec.option_some_iff Partrec.option_some_iff
777776

Mathlib/Data/Fin/VecNotation.lean

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -338,10 +338,11 @@ theorem vecAlt0_vecAppend (v : Fin n → α) : vecAlt0 rfl (vecAppend rfl v v) =
338338
theorem vecAlt1_vecAppend (v : Fin (n + 1) → α) : vecAlt1 rfl (vecAppend rfl v v) = v ∘ bit1 := by
339339
ext i
340340
simp_rw [Function.comp, vecAlt1, vecAppend_eq_ite]
341-
cases n
342-
· cases' i with i hi
341+
cases n with
342+
| zero =>
343+
cases' i with i hi
343344
simp only [Nat.zero_eq, zero_add, Nat.lt_one_iff] at hi; subst i; rfl
344-
case succ n =>
345+
| succ n =>
345346
split_ifs with h <;> simp_rw [bit1, bit0] <;> congr
346347
· simp only [Fin.ext_iff, Fin.val_add, Fin.val_mk]
347348
rw [Fin.val_mk] at h

Mathlib/GroupTheory/Congruence.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1275,7 +1275,7 @@ instance zpowinst : Pow c.Quotient ℤ :=
12751275

12761276
/-- The quotient of a group by a congruence relation is a group. -/
12771277
@[to_additive "The quotient of an `AddGroup` by an additive congruence relation is
1278-
an `add_group`."]
1278+
an `AddGroup`."]
12791279
instance group : Group c.Quotient :=
12801280
Function.Surjective.group Quotient.mk''
12811281
Quotient.surjective_Quotient_mk'' rfl (fun _ _ => rfl) (fun _ => rfl)

0 commit comments

Comments
 (0)