Skip to content

Commit e2b29fb

Browse files
committed
chore: adaptations for lean4#5542 (#17564)
In [lean4#5542](leanprover/lean4#5542) we are deprecating `inductive ... :=`, `structure ... :=`, and `class ... :=` for their `... where` counterparts.
1 parent 77a0e20 commit e2b29fb

File tree

44 files changed

+99
-97
lines changed

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

44 files changed

+99
-97
lines changed

Diff for: Mathlib/Algebra/Category/BialgebraCat/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -59,7 +59,7 @@ lemma of_counit {X : Type v} [Ring X] [Bialgebra R X] :
5959
/-- A type alias for `BialgHom` to avoid confusion between the categorical and
6060
algebraic spellings of composition. -/
6161
@[ext]
62-
structure Hom (V W : BialgebraCat.{v} R) :=
62+
structure Hom (V W : BialgebraCat.{v} R) where
6363
/-- The underlying `BialgHom` -/
6464
toBialgHom : V →ₐc[R] W
6565

Diff for: Mathlib/Algebra/Category/CoalgebraCat/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -62,7 +62,7 @@ lemma of_counit {X : Type v} [AddCommGroup X] [Module R X] [Coalgebra R X] :
6262
/-- A type alias for `CoalgHom` to avoid confusion between the categorical and
6363
algebraic spellings of composition. -/
6464
@[ext]
65-
structure Hom (V W : CoalgebraCat.{v} R) :=
65+
structure Hom (V W : CoalgebraCat.{v} R) where
6666
/-- The underlying `CoalgHom` -/
6767
toCoalgHom : V →ₗc[R] W
6868

Diff for: Mathlib/Algebra/Category/HopfAlgebraCat/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -58,7 +58,7 @@ lemma of_counit {X : Type v} [Ring X] [HopfAlgebra R X] :
5858
/-- A type alias for `BialgHom` to avoid confusion between the categorical and
5959
algebraic spellings of composition. -/
6060
@[ext]
61-
structure Hom (V W : HopfAlgebraCat.{v} R) :=
61+
structure Hom (V W : HopfAlgebraCat.{v} R) where
6262
/-- The underlying `BialgHom`. -/
6363
toBialgHom : V →ₐc[R] W
6464

Diff for: Mathlib/Algebra/Homology/ShortComplex/SnakeLemma.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -381,7 +381,7 @@ variable (S₁ S₂ S₃ : SnakeInput C)
381381
/-- A morphism of snake inputs involve four morphisms of short complexes
382382
which make the obvious diagram commute. -/
383383
@[ext]
384-
structure Hom :=
384+
structure Hom where
385385
/-- a morphism between the zeroth lines -/
386386
f₀ : S₁.L₀ ⟶ S₂.L₀
387387
/-- a morphism between the first lines -/

Diff for: Mathlib/Algebra/Lie/Killing.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -46,7 +46,7 @@ namespace LieAlgebra
4646
4747
NB: This is not standard terminology (the literature does not seem to name Lie algebras with this
4848
property). -/
49-
class IsKilling : Prop :=
49+
class IsKilling : Prop where
5050
/-- We say a Lie algebra is Killing if its Killing form is non-singular. -/
5151
killingCompl_top_eq_bot : LieIdeal.killingCompl R L ⊤ = ⊥
5252

Diff for: Mathlib/Algebra/Lie/Weights/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -724,7 +724,7 @@ noncomputable instance Weight.instFintype [NoZeroSMulDivisors R M] [IsNoetherian
724724

725725
/-- A Lie module `M` of a Lie algebra `L` is triangularizable if the endomorhpism of `M` defined by
726726
any `x : L` is triangularizable. -/
727-
class IsTriangularizable : Prop :=
727+
class IsTriangularizable : Prop where
728728
iSup_eq_top : ∀ x, ⨆ φ, ⨆ k, (toEnd R L M x).genEigenspace φ k = ⊤
729729

730730
instance (L' : LieSubalgebra R L) [IsTriangularizable R L M] : IsTriangularizable R L' M where

Diff for: Mathlib/Algebra/Lie/Weights/Linear.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -48,7 +48,7 @@ namespace LieModule
4848

4949
/-- A typeclass encoding the fact that a given Lie module has linear weights, vanishing on the
5050
derived ideal. -/
51-
class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop :=
51+
class LinearWeights [LieAlgebra.IsNilpotent R L] : Prop where
5252
map_add : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y, χ (x + y) = χ x + χ y
5353
map_smul : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ (t : R) x, χ (t • x) = t • χ x
5454
map_lie : ∀ χ : L → R, genWeightSpace M χ ≠ ⊥ → ∀ x y : L, χ ⁅x, y⁆ = 0

Diff for: Mathlib/Algebra/Quandle.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -100,9 +100,9 @@ class Shelf (α : Type u) where
100100
A *unital shelf* is a shelf equipped with an element `1` such that, for all elements `x`,
101101
we have both `x ◃ 1` and `1 ◃ x` equal `x`.
102102
-/
103-
class UnitalShelf (α : Type u) extends Shelf α, One α :=
104-
(one_act : ∀ a : α, act 1 a = a)
105-
(act_one : ∀ a : α, act a 1 = a)
103+
class UnitalShelf (α : Type u) extends Shelf α, One α where
104+
one_act : ∀ a : α, act 1 a = a
105+
act_one : ∀ a : α, act a 1 = a
106106

107107
/-- The type of homomorphisms between shelves.
108108
This is also the notion of rack and quandle homomorphisms.

Diff for: Mathlib/Analysis/Distribution/SchwartzSpace.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -579,7 +579,7 @@ open MeasureTheory Module
579579

580580
/-- A measure `μ` has temperate growth if there is an `n : ℕ` such that `(1 + ‖x‖) ^ (- n)` is
581581
`μ`-integrable. -/
582-
class _root_.MeasureTheory.Measure.HasTemperateGrowth (μ : Measure D) : Prop :=
582+
class _root_.MeasureTheory.Measure.HasTemperateGrowth (μ : Measure D) : Prop where
583583
exists_integrable : ∃ (n : ℕ), Integrable (fun x ↦ (1 + ‖x‖) ^ (- (n : ℝ))) μ
584584

585585
open Classical in

Diff for: Mathlib/Analysis/NormedSpace/OperatorNorm/Mul.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -121,7 +121,7 @@ examples. Any algebra with an approximate identity (e.g., $$L^1$$) is also regul
121121
122122
This is a useful class because it gives rise to a nice norm on the unitization; in particular it is
123123
a C⋆-norm when the norm on `A` is a C⋆-norm. -/
124-
class _root_.RegularNormedAlgebra : Prop :=
124+
class _root_.RegularNormedAlgebra : Prop where
125125
/-- The left regular representation of the algebra on itself is an isometry. -/
126126
isometry_mul' : Isometry (mul 𝕜 𝕜')
127127

Diff for: Mathlib/Analysis/RCLike/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -1086,7 +1086,7 @@ section
10861086
/-- A mixin over a normed field, saying that the norm field structure is the same as `ℝ` or `ℂ`.
10871087
To endow such a field with a compatible `RCLike` structure in a proof, use
10881088
`letI := IsRCLikeNormedField.rclike 𝕜`.-/
1089-
class IsRCLikeNormedField (𝕜 : Type*) [hk : NormedField 𝕜] : Prop :=
1089+
class IsRCLikeNormedField (𝕜 : Type*) [hk : NormedField 𝕜] : Prop where
10901090
out : ∃ h : RCLike 𝕜, hk = h.toNormedField
10911091

10921092
instance (priority := 100) (𝕜 : Type*) [h : RCLike 𝕜] : IsRCLikeNormedField 𝕜 := ⟨⟨h, rfl⟩⟩

Diff for: Mathlib/Analysis/SpecificLimits/Normed.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -252,7 +252,7 @@ alias tendsto_pow_atTop_nhds_0_of_abs_lt_1 := tendsto_pow_atTop_nhds_zero_of_abs
252252
/-- A normed ring has summable geometric series if, for all `ξ` of norm `< 1`, the geometric series
253253
`∑ ξ ^ n` converges. This holds both in complete normed rings and in normed fields, providing a
254254
convenient abstraction of these two classes to avoid repeating the same proofs. -/
255-
class HasSummableGeomSeries (K : Type*) [NormedRing K] : Prop :=
255+
class HasSummableGeomSeries (K : Type*) [NormedRing K] : Prop where
256256
summable_geometric_of_norm_lt_one : ∀ (ξ : K), ‖ξ‖ < 1 → Summable (fun n ↦ ξ ^ n)
257257

258258
lemma summable_geometric_of_norm_lt_one {K : Type*} [NormedRing K] [HasSummableGeomSeries K]

Diff for: Mathlib/CategoryTheory/GradedObject/Trifunctor.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -259,7 +259,7 @@ variable (F₁₂ : C₁ ⥤ C₂ ⥤ C₁₂) (G : C₁₂ ⥤ C₃ ⥤ C₄)
259259
/-- Given a map `r : I₁ × I₂ × I₃ → J`, a `BifunctorComp₁₂IndexData r` consists of the data
260260
of a type `I₁₂`, maps `p : I₁ × I₂ → I₁₂` and `q : I₁₂ × I₃ → J`, such that `r` is obtained
261261
by composition of `p` and `q`. -/
262-
structure BifunctorComp₁₂IndexData :=
262+
structure BifunctorComp₁₂IndexData where
263263
/-- an auxiliary type -/
264264
I₁₂ : Type*
265265
/-- a map `I₁ × I₂ → I₁₂` -/
@@ -439,7 +439,7 @@ variable (F : C₁ ⥤ C₂₃ ⥤ C₄) (G₂₃ : C₂ ⥤ C₃ ⥤ C₂₃)
439439
/-- Given a map `r : I₁ × I₂ × I₃ → J`, a `BifunctorComp₂₃IndexData r` consists of the data
440440
of a type `I₂₃`, maps `p : I₂ × I₃ → I₂₃` and `q : I₁ × I₂₃ → J`, such that `r` is obtained
441441
by composition of `p` and `q`. -/
442-
structure BifunctorComp₂₃IndexData :=
442+
structure BifunctorComp₂₃IndexData where
443443
/-- an auxiliary type -/
444444
I₂₃ : Type*
445445
/-- a map `I₂ × I₃ → I₂₃` -/

Diff for: Mathlib/CategoryTheory/Localization/LocalizerMorphism.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -128,7 +128,7 @@ lemma isEquivalence_iff : G.IsEquivalence ↔ G'.IsEquivalence :=
128128
end
129129

130130
/-- Condition that a `LocalizerMorphism` induces an equivalence on the localized categories -/
131-
class IsLocalizedEquivalence : Prop :=
131+
class IsLocalizedEquivalence : Prop where
132132
/-- the induced functor on the constructed localized categories is an equivalence -/
133133
isEquivalence : (Φ.localizedFunctor W₁.Q W₂.Q).IsEquivalence
134134

Diff for: Mathlib/CategoryTheory/MorphismProperty/Composition.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -24,7 +24,7 @@ namespace MorphismProperty
2424
variable {C : Type u} [Category.{v} C] {D : Type u'} [Category.{v'} D]
2525

2626
/-- Typeclass expressing that a morphism property contain identities. -/
27-
class ContainsIdentities (W : MorphismProperty C) : Prop :=
27+
class ContainsIdentities (W : MorphismProperty C) : Prop where
2828
/-- for all `X : C`, the identity of `X` satisfies the morphism property -/
2929
id_mem : ∀ (X : C), W (𝟙 X)
3030

@@ -63,7 +63,7 @@ instance Pi.containsIdentities {J : Type w} {C : J → Type u}
6363

6464
/-- A morphism property satisfies `IsStableUnderComposition` if the composition of
6565
two such morphisms still falls in the class. -/
66-
class IsStableUnderComposition (P : MorphismProperty C) : Prop :=
66+
class IsStableUnderComposition (P : MorphismProperty C) : Prop where
6767
comp_mem {X Y Z} (f : X ⟶ Y) (g : Y ⟶ Z) : P f → P g → P (f ≫ g)
6868

6969
lemma comp_mem (W : MorphismProperty C) [W.IsStableUnderComposition]
@@ -129,7 +129,7 @@ end naturalityProperty
129129
/-- A morphism property is multiplicative if it contains identities and is stable by
130130
composition. -/
131131
class IsMultiplicative (W : MorphismProperty C)
132-
extends W.ContainsIdentities, W.IsStableUnderComposition : Prop :=
132+
extends W.ContainsIdentities, W.IsStableUnderComposition : Prop
133133

134134
namespace IsMultiplicative
135135

Diff for: Mathlib/CategoryTheory/MorphismProperty/Limits.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -203,7 +203,7 @@ lemma IsStableUnderProductsOfShape.mk (J : Type*)
203203
simp
204204

205205
/-- The condition that a property of morphisms is stable by finite products. -/
206-
class IsStableUnderFiniteProducts : Prop :=
206+
class IsStableUnderFiniteProducts : Prop where
207207
isStableUnderProductsOfShape (J : Type) [Finite J] : W.IsStableUnderProductsOfShape J
208208

209209
lemma isStableUnderProductsOfShape_of_isStableUnderFiniteProducts

Diff for: Mathlib/CategoryTheory/Shift/CommShift.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -246,7 +246,7 @@ variable {C D E J : Type*} [Category C] [Category D] [Category E] [Category J]
246246
/-- If `τ : F₁ ⟶ F₂` is a natural transformation between two functors
247247
which commute with a shift by an additive monoid `A`, this typeclass
248248
asserts a compatibility of `τ` with these shifts. -/
249-
class CommShift : Prop :=
249+
class CommShift : Prop where
250250
comm' (a : A) : (F₁.commShiftIso a).hom ≫ whiskerRight τ _ =
251251
whiskerLeft _ τ ≫ (F₂.commShiftIso a).hom
252252

Diff for: Mathlib/CategoryTheory/Shift/Localization.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -32,7 +32,7 @@ namespace MorphismProperty
3232
/-- A morphism property `W` on a category `C` is compatible with the shift by a
3333
monoid `A` when for all `a : A`, a morphism `f` belongs to `W`
3434
if and only if `f⟦a⟧'` does. -/
35-
class IsCompatibleWithShift : Prop :=
35+
class IsCompatibleWithShift : Prop where
3636
/-- the condition that for all `a : A`, the morphism property `W` is not changed when
3737
we take its inverse image by the shift functor by `a` -/
3838
condition : ∀ (a : A), W.inverseImage (shiftFunctor C a) = W

Diff for: Mathlib/CategoryTheory/Shift/Quotient.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -32,7 +32,7 @@ namespace HomRel
3232

3333
/-- A relation on morphisms is compatible with the shift by a monoid `A` when the
3434
relation if preserved by the shift. -/
35-
class IsCompatibleWithShift : Prop :=
35+
class IsCompatibleWithShift : Prop where
3636
/-- the condition that the relation is preserved by the shift -/
3737
condition : ∀ (a : A) ⦃X Y : C⦄ (f g : X ⟶ Y), r f g → r (f⟦a⟧') (g⟦a⟧')
3838

Diff for: Mathlib/Combinatorics/Colex.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -65,7 +65,7 @@ namespace Finset
6565
/-- Type synonym of `Finset α` equipped with the colexicographic order rather than the inclusion
6666
order. -/
6767
@[ext]
68-
structure Colex (α) :=
68+
structure Colex (α) where
6969
/-- `toColex` is the "identity" function between `Finset α` and `Finset.Colex α`. -/
7070
toColex ::
7171
/-- `ofColex` is the "identity" function between `Finset.Colex α` and `Finset α`. -/

Diff for: Mathlib/Combinatorics/SimpleGraph/Hamiltonian.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -66,7 +66,7 @@ lemma IsHamiltonian.length_eq (hp : p.IsHamiltonian) : p.length = Fintype.card
6666
end
6767

6868
/-- A hamiltonian cycle is a cycle that visits every vertex once. -/
69-
structure IsHamiltonianCycle (p : G.Walk a a) extends p.IsCycle : Prop :=
69+
structure IsHamiltonianCycle (p : G.Walk a a) extends p.IsCycle : Prop where
7070
isHamiltonian_tail : p.tail.IsHamiltonian
7171

7272
variable {p : G.Walk a a}

Diff for: Mathlib/Data/FunLike/Basic.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -17,7 +17,7 @@ There is the "D"ependent version `DFunLike` and the non-dependent version `FunLi
1717
1818
A typical type of morphisms should be declared as:
1919
```
20-
structure MyHom (A B : Type*) [MyClass A] [MyClass B] :=
20+
structure MyHom (A B : Type*) [MyClass A] [MyClass B] where
2121
(toFun : A → B)
2222
(map_op' : ∀ (x y : A), toFun (MyClass.op x y) = MyClass.op (toFun x) (toFun y))
2323
@@ -79,7 +79,7 @@ The second step is to add instances of your new `MyHomClass` for all types exten
7979
Typically, you can just declare a new class analogous to `MyHomClass`:
8080
8181
```
82-
structure CoolerHom (A B : Type*) [CoolClass A] [CoolClass B] extends MyHom A B :=
82+
structure CoolerHom (A B : Type*) [CoolClass A] [CoolClass B] extends MyHom A B where
8383
(map_cool' : toFun CoolClass.cool = CoolClass.cool)
8484
8585
class CoolerHomClass (F : Type*) (A B : outParam Type*) [CoolClass A] [CoolClass B]

Diff for: Mathlib/Data/FunLike/Embedding.lean

+5-5
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ This typeclass is primarily for use by embeddings such as `RelEmbedding`.
1414
1515
A typical type of embeddings should be declared as:
1616
```
17-
structure MyEmbedding (A B : Type*) [MyClass A] [MyClass B] :=
17+
structure MyEmbedding (A B : Type*) [MyClass A] [MyClass B] where
1818
(toFun : A → B)
1919
(injective' : Function.Injective toFun)
2020
(map_op' : ∀ (x y : A), toFun (MyClass.op x y) = MyClass.op (toFun x) (toFun y))
@@ -58,8 +58,8 @@ Continuing the example above:
5858
You should extend this class when you extend `MyEmbedding`. -/
5959
class MyEmbeddingClass (F : Type*) (A B : outParam Type*) [MyClass A] [MyClass B]
6060
[FunLike F A B]
61-
extends EmbeddingLike F A B :=
62-
(map_op : ∀ (f : F) (x y : A), f (MyClass.op x y) = MyClass.op (f x) (f y))
61+
extends EmbeddingLike F A B where
62+
map_op : ∀ (f : F) (x y : A), f (MyClass.op x y) = MyClass.op (f x) (f y)
6363
6464
@[simp]
6565
lemma map_op {F A B : Type*} [MyClass A] [MyClass B] [FunLike F A B] [MyEmbeddingClass F A B]
@@ -84,12 +84,12 @@ The second step is to add instances of your new `MyEmbeddingClass` for all types
8484
Typically, you can just declare a new class analogous to `MyEmbeddingClass`:
8585
8686
```
87-
structure CoolerEmbedding (A B : Type*) [CoolClass A] [CoolClass B] extends MyEmbedding A B :=
87+
structure CoolerEmbedding (A B : Type*) [CoolClass A] [CoolClass B] extends MyEmbedding A B where
8888
(map_cool' : toFun CoolClass.cool = CoolClass.cool)
8989
9090
class CoolerEmbeddingClass (F : Type*) (A B : outParam Type*) [CoolClass A] [CoolClass B]
9191
[FunLike F A B]
92-
extends MyEmbeddingClass F A B :=
92+
extends MyEmbeddingClass F A B where
9393
(map_cool : ∀ (f : F), f CoolClass.cool = CoolClass.cool)
9494
9595
@[simp]

Diff for: Mathlib/Data/FunLike/Equiv.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ This typeclass is primarily for use by isomorphisms like `MonoidEquiv` and `Line
1414
1515
A typical type of isomorphisms should be declared as:
1616
```
17-
structure MyIso (A B : Type*) [MyClass A] [MyClass B] extends Equiv A B :=
17+
structure MyIso (A B : Type*) [MyClass A] [MyClass B] extends Equiv A B where
1818
(map_op' : ∀ (x y : A), toFun (MyClass.op x y) = MyClass.op (toFun x) (toFun y))
1919
2020
namespace MyIso
@@ -77,12 +77,12 @@ The second step is to add instances of your new `MyIsoClass` for all types exten
7777
Typically, you can just declare a new class analogous to `MyIsoClass`:
7878
7979
```
80-
structure CoolerIso (A B : Type*) [CoolClass A] [CoolClass B] extends MyIso A B :=
80+
structure CoolerIso (A B : Type*) [CoolClass A] [CoolClass B] extends MyIso A B where
8181
(map_cool' : toFun CoolClass.cool = CoolClass.cool)
8282
8383
class CoolerIsoClass (F : Type*) (A B : outParam Type*) [CoolClass A] [CoolClass B]
8484
[EquivLike F A B]
85-
extends MyIsoClass F A B :=
85+
extends MyIsoClass F A B where
8686
(map_cool : ∀ (f : F), f CoolClass.cool = CoolClass.cool)
8787
8888
@[simp] lemma map_cool {F A B : Type*} [CoolClass A] [CoolClass B]

Diff for: Mathlib/Data/Opposite.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -30,7 +30,7 @@ variable (α : Sort u)
3030
both `unop (op X) = X` and `op (unop X) = X` are definitional equalities.
3131
3232
-/
33-
structure Opposite :=
33+
structure Opposite where
3434
/-- The canonical map `α → αᵒᵖ`. -/
3535
op ::
3636
/-- The canonical map `αᵒᵖ → α`. -/

Diff for: Mathlib/Data/SetLike/Basic.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -28,7 +28,7 @@ boilerplate for every `SetLike`: a `coe_sort`, a `coe` to set, a
2828
2929
A typical subobject should be declared as:
3030
```
31-
structure MySubobject (X : Type*) [ObjectTypeclass X] :=
31+
structure MySubobject (X : Type*) [ObjectTypeclass X] where
3232
(carrier : Set X)
3333
(op_mem' : ∀ {x : X}, x ∈ carrier → sorry ∈ carrier)
3434
@@ -60,7 +60,7 @@ end MySubobject
6060
6161
An alternative to `SetLike` could have been an extensional `Membership` typeclass:
6262
```
63-
class ExtMembership (α : out_param <| Type u) (β : Type v) extends Membership α β :=
63+
class ExtMembership (α : out_param <| Type u) (β : Type v) extends Membership α β where
6464
(ext_iff : ∀ {s t : β}, s = t ↔ ∀ (x : α), x ∈ s ↔ x ∈ t)
6565
```
6666
While this is equivalent, `SetLike` conveniently uses a carrier set projection directly.

Diff for: Mathlib/GroupTheory/HNNExtension.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -174,7 +174,7 @@ namespace NormalWord
174174
variable (G A B)
175175
/-- To put word in the HNN Extension into a normal form, we must choose an element of each right
176176
coset of both `A` and `B`, such that the chosen element of the subgroup itself is `1`. -/
177-
structure TransversalPair : Type _ :=
177+
structure TransversalPair : Type _ where
178178
/-- The transversal of each subgroup -/
179179
set : ℤˣ → Set G
180180
/-- We have exactly one element of each coset of the subgroup -/
@@ -187,7 +187,7 @@ instance TransversalPair.nonempty : Nonempty (TransversalPair G A B) := by
187187
/-- A reduced word is a `head`, which is an element of `G`, followed by the product list of pairs.
188188
There should also be no sequences of the form `t^u * g * t^-u`, where `g` is in
189189
`toSubgroup A B u` This is a less strict condition than required for `NormalWord`. -/
190-
structure ReducedWord : Type _ :=
190+
structure ReducedWord : Type _ where
191191
/-- Every `ReducedWord` is the product of an element of the group and a word made up
192192
of letters each of which is in the transversal. `head` is that element of the base group. -/
193193
head : G
@@ -215,7 +215,7 @@ The normal form is a `head`, which is an element of `G`, followed by the product
215215
`toSubgroup A B u`. There should also be no sequences of the form `t^u * g * t^-u`
216216
where `g ∈ toSubgroup A B u` -/
217217
structure _root_.HNNExtension.NormalWord (d : TransversalPair G A B)
218-
extends ReducedWord G A B : Type _ :=
218+
extends ReducedWord G A B : Type _ where
219219
/-- Every element `g : G` in the list is the chosen element of its coset -/
220220
mem_set : ∀ (u : ℤˣ) (g : G), (u, g) ∈ toList → g ∈ d.set u
221221

Diff for: Mathlib/LinearAlgebra/PerfectPairing.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,7 @@ open Function Module
3434
variable (R M N : Type*) [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N]
3535

3636
/-- A perfect pairing of two (left) modules over a commutative ring. -/
37-
structure PerfectPairing :=
37+
structure PerfectPairing where
3838
toLin : M →ₗ[R] N →ₗ[R] R
3939
bijectiveLeft : Bijective toLin
4040
bijectiveRight : Bijective toLin.flip

Diff for: Mathlib/LinearAlgebra/QuadraticForm/QuadraticModuleCat.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -43,7 +43,7 @@ def of {X : Type v} [AddCommGroup X] [Module R X] (Q : QuadraticForm R X) :
4343
/-- A type alias for `QuadraticForm.LinearIsometry` to avoid confusion between the categorical and
4444
algebraic spellings of composition. -/
4545
@[ext]
46-
structure Hom (V W : QuadraticModuleCat.{v} R) :=
46+
structure Hom (V W : QuadraticModuleCat.{v} R) where
4747
/-- The underlying isometry -/
4848
toIsometry : V.form →qᵢ W.form
4949

Diff for: Mathlib/LinearAlgebra/RootSystem/Defs.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -86,7 +86,7 @@ evaluates to `2`, and the permutation attached to each element of `ι` is compat
8686
reflections on the corresponding roots and coroots.
8787
8888
It exists to allow for a convenient unification of the theories of root systems and root data. -/
89-
structure RootPairing extends PerfectPairing R M N :=
89+
structure RootPairing extends PerfectPairing R M N where
9090
/-- A parametrized family of vectors, called roots. -/
9191
root : ι ↪ M
9292
/-- A parametrized family of dual vectors, called coroots. -/
@@ -110,7 +110,7 @@ abbrev RootDatum (X₁ X₂ : Type*) [AddCommGroup X₁] [AddCommGroup X₂] :=
110110
111111
Note that this is slightly more general than the usual definition in the sense that `N` is not
112112
required to be the dual of `M`. -/
113-
structure RootSystem extends RootPairing ι R M N :=
113+
structure RootSystem extends RootPairing ι R M N where
114114
span_eq_top : span R (range root) = ⊤
115115

116116
attribute [simp] RootSystem.span_eq_top

Diff for: Mathlib/LinearAlgebra/RootSystem/OfBilinear.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -33,7 +33,7 @@ namespace LinearMap
3333
injective, and for any vector `y`, the norm of `x` divides twice the inner product of `x` and `y`.
3434
These conditions are what we need when describing reflection as a map taking `y` to
3535
`y - 2 • (B x y) / (B x x) • x`. -/
36-
structure IsReflective (B : M →ₗ[R] M →ₗ[R] R) (x : M) : Prop :=
36+
structure IsReflective (B : M →ₗ[R] M →ₗ[R] R) (x : M) : Prop where
3737
regular : IsRegular (B x x)
3838
dvd_two_mul : ∀ y, B x x ∣ 2 * B x y
3939

0 commit comments

Comments
 (0)