Skip to content

Commit 6fff716

Browse files
chore: tidy various files (#2321)
1 parent 2657489 commit 6fff716

File tree

10 files changed

+154
-177
lines changed

10 files changed

+154
-177
lines changed

Mathlib/Algebra/GradedMonoid.lean

Lines changed: 40 additions & 42 deletions
Original file line numberDiff line numberDiff line change
@@ -104,7 +104,7 @@ def GradedMonoid (A : ι → Type _) :=
104104

105105
namespace GradedMonoid
106106

107-
instance {A : ι → Type _} [Inhabited ι] [Inhabited (A default)] : Inhabited (GradedMonoid A) where
107+
instance {A : ι → Type _} [Inhabited ι] [Inhabited (A default)] : Inhabited (GradedMonoid A) where
108108
default := Sigma.instInhabitedSigma.default
109109

110110
/-- Construct an element of a graded monoid. -/
@@ -194,7 +194,7 @@ class GMonoid [AddMonoid ι] extends GMul A, GOne A where
194194
/-- Successor powers behave as expected -/
195195
gnpow_succ' :
196196
∀ (n : ℕ) (a : GradedMonoid A),
197-
(GradedMonoid.mk _ <| gnpow n.succ a.snd) = a * ⟨_, gnpow n a.snd⟩ := by
197+
(GradedMonoid.mk _ <| gnpow n.succ a.snd) = a * ⟨_, gnpow n a.snd⟩ := by
198198
apply_gmonoid_gnpowRec_succ_tac
199199
#align graded_monoid.gmonoid GradedMonoid.GMonoid
200200

@@ -214,11 +214,11 @@ instance GMonoid.toMonoid [AddMonoid ι] [GMonoid A] : Monoid (GradedMonoid A)
214214
theorem mk_pow [AddMonoid ι] [GMonoid A] {i} (a : A i) (n : ℕ) :
215215
mk i a ^ n = mk (n • i) (GMonoid.gnpow _ a) :=
216216
by
217-
match n with
218-
| 0 =>
217+
match n with
218+
| 0 =>
219219
rw [pow_zero]
220220
exact (GMonoid.gnpow_zero' ⟨_, a⟩).symm
221-
| n+1 =>
221+
| n+1 =>
222222
rw [pow_succ, mk_pow a n, mk_mul_mk]
223223
exact (GMonoid.gnpow_succ' n ⟨_, a⟩).symm
224224
#align graded_monoid.mk_pow GradedMonoid.mk_pow
@@ -268,7 +268,7 @@ variable [AddZeroClass ι] [GMul A]
268268
an `Eq.rec` to turn `A (0 + i)` into `A i`.
269269
-/
270270
instance GradeZero.smul (i : ι) : SMul (A 0) (A i)
271-
where smul x y := @Eq.rec ι (0+i) (fun a _ => A a) (GMul.mul x y) i (zero_add i)
271+
where smul x y := @Eq.rec ι (0+i) (fun a _ => A a) (GMul.mul x y) i (zero_add i)
272272
#align graded_monoid.grade_zero.has_smul GradedMonoid.GradeZero.smul
273273

274274
/-- `(*) : A 0 → A 0 → A 0` is the value provided in `GradedMonoid.GMul.mul`, composed with
@@ -281,7 +281,7 @@ variable {A}
281281

282282
@[simp]
283283
theorem mk_zero_smul {i} (a : A 0) (b : A i) : mk _ (a • b) = mk _ a * mk _ b :=
284-
Sigma.ext (zero_add _).symm <| eq_rec_heq _ _
284+
Sigma.ext (zero_add _).symm <| eq_rec_heq _ _
285285
#align graded_monoid.mk_zero_smul GradedMonoid.mk_zero_smul
286286

287287
@[simp]
@@ -295,7 +295,7 @@ section Monoid
295295

296296
variable [AddMonoid ι] [GMonoid A]
297297

298-
instance : Pow (A 0) ℕ where
298+
instance : Pow (A 0) ℕ where
299299
pow x n := @Eq.rec ι (n • (0:ι)) (fun a _ => A a) (GMonoid.gnpow n x) 0 (nsmul_zero n)
300300

301301
variable {A}
@@ -376,9 +376,9 @@ theorem List.dProdIndex_cons (a : α) (l : List α) (fι : α → ι) :
376376

377377
theorem List.dProdIndex_eq_map_sum (l : List α) (fι : α → ι) : l.dProdIndex fι = (l.map fι).sum :=
378378
by
379-
match l with
380-
| [] => simp
381-
| head::tail => simp [List.dProdIndex_eq_map_sum tail fι]
379+
match l with
380+
| [] => simp
381+
| head::tail => simp [List.dProdIndex_eq_map_sum tail fι]
382382
#align list.dprod_index_eq_map_sum List.dProdIndex_eq_map_sum
383383

384384
/-- A dependent product for graded monoids represented by the indexed family of types `A i`.
@@ -407,9 +407,9 @@ theorem List.dProd_cons (fι : α → ι) (fA : ∀ a, A (fι a)) (a : α) (l :
407407
theorem GradedMonoid.mk_list_dProd (l : List α) (fι : α → ι) (fA : ∀ a, A (fι a)) :
408408
GradedMonoid.mk _ (l.dProd fι fA) = (l.map fun a => GradedMonoid.mk (fι a) (fA a)).prod :=
409409
by
410-
match l with
411-
| [] => simp; rfl
412-
| head::tail =>
410+
match l with
411+
| [] => simp; rfl
412+
| head::tail =>
413413
simp[← GradedMonoid.mk_list_dProd tail _ _, GradedMonoid.mk_mul_mk, List.prod_cons]
414414
#align graded_monoid.mk_list_dprod GradedMonoid.mk_list_dProd
415415

@@ -450,7 +450,7 @@ structure. -/
450450
instance Monoid.gMonoid [AddMonoid ι] [Monoid R] : GradedMonoid.GMonoid fun _ : ι => R :=
451451
-- { Mul.gMul ι, One.gOne ι with
452452
{ One.gOne ι with
453-
mul := fun x y => x * y
453+
mul := fun x y => x * y
454454
one_mul := fun _ => Sigma.ext (zero_add _) (heq_of_eq (one_mul _))
455455
mul_one := fun _ => Sigma.ext (add_zero _) (heq_of_eq (mul_one _))
456456
mul_assoc := fun _ _ _ => Sigma.ext (add_assoc _ _ _) (heq_of_eq (mul_assoc _ _ _))
@@ -473,11 +473,11 @@ instance CommMonoid.gCommMonoid [AddCommMonoid ι] [CommMonoid R] :
473473
theorem List.dProd_monoid {α} [AddMonoid ι] [Monoid R] (l : List α) (fι : α → ι) (fA : α → R) :
474474
@List.dProd _ _ (fun _:ι => R) _ _ l fι fA = (l.map fA).prod :=
475475
by
476-
match l with
477-
| [] =>
476+
match l with
477+
| [] =>
478478
rw [List.dProd_nil, List.map_nil, List.prod_nil]
479479
rfl
480-
| head::tail =>
480+
| head::tail =>
481481
rw [List.dProd_cons, List.map_cons, List.prod_cons, List.dProd_monoid tail _ _]
482482
rfl
483483
#align list.dprod_monoid List.dProd_monoid
@@ -530,18 +530,18 @@ instance SetLike.gMul {S : Type _} [SetLike S R] [Mul R] [Add ι] (A : ι → S)
530530
#align set_like.ghas_mul SetLike.gMul
531531

532532
/-
533-
Porting note: simpNF linter returns
533+
Porting note: simpNF linter returns
534534
535535
"Left-hand side does not simplify, when using the simp lemma on itself."
536536
537537
However, simp does indeed solve the following. Possibly related std#71,std#78
538538
539539
example {S : Type _} [SetLike S R] [Mul R] [Add ι] (A : ι → S)
540540
[SetLike.GradedMul A] {i j : ι} (x : A i) (y : A j) :
541-
↑(@GradedMonoid.GMul.mul _ (fun i => A i) _ _ _ _ x y) = (x * y : R) := by simp
541+
↑(@GradedMonoid.GMul.mul _ (fun i => A i) _ _ _ _ x y) = (x * y : R) := by simp
542542
543543
-/
544-
@[simp,nolint simpNF]
544+
@[simp,nolint simpNF]
545545
theorem SetLike.coe_gMul {S : Type _} [SetLike S R] [Mul R] [Add ι] (A : ι → S)
546546
[SetLike.GradedMul A] {i j : ι} (x : A i) (y : A j) :
547547
↑(@GradedMonoid.GMul.mul _ (fun i => A i) _ _ _ _ x y) = (x * y : R) :=
@@ -562,23 +562,23 @@ variable {A : ι → S} [SetLike.GradedMonoid A]
562562

563563
theorem pow_mem_graded (n : ℕ) {r : R} {i : ι} (h : r ∈ A i) : r ^ n ∈ A (n • i) :=
564564
by
565-
match n with
566-
| 0 =>
565+
match n with
566+
| 0 =>
567567
rw [pow_zero, zero_nsmul]
568568
exact one_mem_graded _
569-
| n+1 =>
569+
| n+1 =>
570570
rw [pow_succ', succ_nsmul']
571571
exact mul_mem_graded (pow_mem_graded n h) h
572572
#align set_like.pow_mem_graded SetLike.pow_mem_graded
573573

574574
theorem list_prod_map_mem_graded {ι'} (l : List ι') (i : ι' → ι) (r : ι' → R)
575575
(h : ∀ j ∈ l, r j ∈ A (i j)) : (l.map r).prod ∈ A (l.map i).sum :=
576576
by
577-
match l with
578-
| [] =>
577+
match l with
578+
| [] =>
579579
rw [List.map_nil, List.map_nil, List.prod_nil, List.sum_nil]
580580
exact one_mem_graded _
581-
| head::tail =>
581+
| head::tail =>
582582
rw [List.map_cons, List.map_cons, List.prod_cons, List.sum_cons]
583583
exact
584584
mul_mem_graded (h _ <| List.mem_cons_self _ _)
@@ -598,8 +598,7 @@ end SetLike
598598
instance SetLike.gMonoid {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A : ι → S)
599599
[SetLike.GradedMonoid A] : GradedMonoid.GMonoid fun i => A i :=
600600
{ SetLike.gOne A,
601-
SetLike.gMul
602-
A with
601+
SetLike.gMul A with
603602
one_mul := fun ⟨_, _, _⟩ => Sigma.subtype_ext (zero_add _) (one_mul _)
604603
mul_one := fun ⟨_, _, _⟩ => Sigma.subtype_ext (add_zero _) (mul_one _)
605604
mul_assoc := fun ⟨_, _, _⟩ ⟨_, _, _⟩ ⟨_, _, _⟩ =>
@@ -610,7 +609,7 @@ instance SetLike.gMonoid {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A
610609
#align set_like.gmonoid SetLike.gMonoid
611610

612611
/-
613-
Porting note: simpNF linter returns
612+
Porting note: simpNF linter returns
614613
615614
"Left-hand side does not simplify, when using the simp lemma on itself."
616615
@@ -621,7 +620,7 @@ example {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A : ι → S)
621620
↑(@GradedMonoid.GMonoid.gnpow _ (fun i => A i) _ _ n _ x) = (x:R)^n := by simp
622621
623622
-/
624-
@[simp,nolint simpNF]
623+
@[simp,nolint simpNF]
625624
theorem SetLike.coe_gnpow {S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι] (A : ι → S)
626625
[SetLike.GradedMonoid A] {i : ι} (x : A i) (n : ℕ) :
627626
↑(@GradedMonoid.GMonoid.gnpow _ (fun i => A i) _ _ n _ x) = (x:R)^n :=
@@ -642,30 +641,30 @@ open SetLike SetLike.GradedMonoid
642641
variable {α S : Type _} [SetLike S R] [Monoid R] [AddMonoid ι]
643642

644643
/-
645-
Porting note: simpNF linter returns
644+
Porting note: simpNF linter returns
646645
647646
"Left-hand side does not simplify, when using the simp lemma on itself."
648647
649648
However, simp does indeed solve the following. Possibly related std#71,std#78
650649
651650
example (A : ι → S) [SetLike.GradedMonoid A] (fι : α → ι)
652-
(fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
651+
(fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
653652
= (List.prod (l.map fun a => fA a) : R) := by simp
654653
-/
655654
/-- Coercing a dependent product of subtypes is the same as taking the regular product of the
656655
coercions. -/
657-
@[simp,nolint simpNF]
656+
@[simp,nolint simpNF]
658657
theorem SetLike.coe_list_dProd (A : ι → S) [SetLike.GradedMonoid A] (fι : α → ι)
659-
(fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
658+
(fA : ∀ a, A (fι a)) (l : List α) : ↑(@List.dProd _ _ (fun i => ↥(A i)) _ _ l fι fA)
660659
= (List.prod (l.map fun a => fA a) : R) := by
661-
match l with
662-
| [] =>
660+
match l with
661+
| [] =>
663662
rw [List.dProd_nil, coe_gOne, List.map_nil, List.prod_nil]
664-
| head::tail =>
665-
rw [List.dProd_cons, coe_gMul, List.map_cons, List.prod_cons,
663+
| head::tail =>
664+
rw [List.dProd_cons, coe_gMul, List.map_cons, List.prod_cons,
666665
SetLike.coe_list_dProd _ _ _ tail]
667666
#align set_like.coe_list_dprod SetLike.coe_list_dProd
668-
667+
669668
/-- A version of `List.coe_dProd_set_like` with `Subtype.mk`. -/
670669
theorem SetLike.list_dProd_eq (A : ι → S) [SetLike.GradedMonoid A] (fι : α → ι) (fA : ∀ a, A (fι a))
671670
(l : List α) :
@@ -709,8 +708,7 @@ def SetLike.homogeneousSubmonoid [AddMonoid ι] [Monoid R] (A : ι → S) [SetLi
709708
Submonoid R where
710709
carrier := { a | SetLike.Homogeneous A a }
711710
one_mem' := SetLike.homogeneous_one A
712-
mul_mem' a b := SetLike.homogeneous_mul a b
711+
mul_mem' a b := SetLike.homogeneous_mul a b
713712
#align set_like.homogeneous_submonoid SetLike.homogeneousSubmonoid
714713

715714
end HomogeneousElements
716-

Mathlib/CategoryTheory/Category/Preorder.lean

Lines changed: 17 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -71,15 +71,15 @@ alias homOfLE ← _root_.LE.le.hom
7171
#align has_le.le.hom LE.le.hom
7272

7373
@[simp]
74-
theorem hom_of_le_refl {x : X} : (le_refl x).hom = 𝟙 x :=
74+
theorem homOfLE_refl {x : X} : (le_refl x).hom = 𝟙 x :=
7575
rfl
76-
#align category_theory.hom_of_le_refl CategoryTheory.hom_of_le_refl
76+
#align category_theory.hom_of_le_refl CategoryTheory.homOfLE_refl
7777

7878
@[simp]
79-
theorem hom_of_le_comp {x y z : X} (h : x ≤ y) (k : y ≤ z) :
79+
theorem homOfLE_comp {x y z : X} (h : x ≤ y) (k : y ≤ z) :
8080
homOfLE h ≫ homOfLE k = homOfLE (h.trans k) :=
8181
rfl
82-
#align category_theory.hom_of_le_comp CategoryTheory.hom_of_le_comp
82+
#align category_theory.hom_of_le_comp CategoryTheory.homOfLE_comp
8383

8484
/-- Extract the underlying inequality from a morphism in a preorder category.
8585
-/
@@ -90,29 +90,28 @@ theorem leOfHom {x y : X} (h : x ⟶ y) : x ≤ y :=
9090
alias leOfHom ← _root_.Quiver.Hom.le
9191
#align quiver.hom.le Quiver.Hom.le
9292

93-
-- porting note: linter seems to be wrong here
94-
@[simp, nolint simpNF]
95-
theorem le_of_hom_hom_of_le {x y : X} (h : x ≤ y) : h.hom.le = h :=
93+
-- porting note: why does this lemma exist? With proof irrelevance, we don't need to simplify proofs
94+
-- @[simp]
95+
theorem leOfHom_homOfLE {x y : X} (h : x ≤ y) : h.hom.le = h :=
9696
rfl
97-
#align category_theory.le_of_hom_hom_of_le CategoryTheory.le_of_hom_hom_of_le
97+
#align category_theory.le_of_hom_hom_of_le CategoryTheory.leOfHom_homOfLE
9898

99-
-- porting note: linter gives: "Left-hand side does not simplify, when using the simp lemma on
100-
-- itself. This usually means that it will never apply." removing simp? It doesn't fire
99+
-- porting note: why does this lemma exist? With proof irrelevance, we don't need to simplify proofs
101100
-- @[simp]
102-
theorem hom_of_le_le_of_hom {x y : X} (h : x ⟶ y) : h.le.hom = h := by
101+
theorem homOfLE_leOfHom {x y : X} (h : x ⟶ y) : h.le.hom = h := by
103102
cases' h with h
104103
cases h
105104
rfl
106-
#align category_theory.hom_of_le_le_of_hom CategoryTheory.hom_of_le_le_of_hom
105+
#align category_theory.hom_of_le_le_of_hom CategoryTheory.homOfLE_leOfHom
107106

108107
/-- Construct a morphism in the opposite of a preorder category from an inequality. -/
109108
def opHomOfLe {x y : Xᵒᵖ} (h : unop x ≤ unop y) : y ⟶ x :=
110109
(homOfLE h).op
111110
#align category_theory.op_hom_of_le CategoryTheory.opHomOfLe
112111

113-
theorem le_ofOp_hom {x y : Xᵒᵖ} (h : x ⟶ y) : unop y ≤ unop x :=
112+
theorem le_of_op_hom {x y : Xᵒᵖ} (h : x ⟶ y) : unop y ≤ unop x :=
114113
h.unop.le
115-
#align category_theory.le_of_op_hom CategoryTheory.le_ofOp_hom
114+
#align category_theory.le_of_op_hom CategoryTheory.le_of_op_hom
116115

117116
instance uniqueToTop [OrderTop X] {x : X} : Unique (x ⟶ ⊤) where
118117
default := homOfLE le_top
@@ -132,10 +131,9 @@ variable {X : Type u} {Y : Type v} [Preorder X] [Preorder Y]
132131

133132
/-- A monotone function between preorders induces a functor between the associated categories.
134133
-/
135-
def Monotone.functor {f : X → Y} (h : Monotone f) : X ⥤ Y
136-
where
134+
def Monotone.functor {f : X → Y} (h : Monotone f) : X ⥤ Y where
137135
obj := f
138-
map := @fun x₁ x₂ g => CategoryTheory.homOfLE (h g.le)
136+
map g := CategoryTheory.homOfLE (h g.le)
139137
#align monotone.functor Monotone.functor
140138

141139
@[simp]
@@ -169,13 +167,12 @@ theorem Iso.to_eq {x y : X} (f : x ≅ y) : x = y :=
169167

170168
/-- A categorical equivalence between partial orders is just an order isomorphism.
171169
-/
172-
def Equivalence.toOrderIso (e : X ≌ Y) : X ≃o Y
173-
where
170+
def Equivalence.toOrderIso (e : X ≌ Y) : X ≃o Y where
174171
toFun := e.functor.obj
175172
invFun := e.inverse.obj
176173
left_inv a := (e.unitIso.app a).to_eq.symm
177174
right_inv b := (e.counitIso.app b).to_eq
178-
map_rel_iff' := @fun a a' =>
175+
map_rel_iff' {a a'} :=
179176
fun h =>
180177
((Equivalence.unit e).app a ≫ e.inverse.map h.hom ≫ (Equivalence.unitInv e).app a').le,
181178
fun h : a ≤ a' => (e.functor.map h.hom).le⟩

Mathlib/CategoryTheory/Equivalence.lean

Lines changed: 9 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -697,9 +697,9 @@ namespace Equivalence
697697
698698
See <https://stacks.math.columbia.edu/tag/02C3>.
699699
-/
700-
theorem ess_surj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F :=
700+
theorem essSurj_of_equivalence (F : C ⥤ D) [IsEquivalence F] : EssSurj F :=
701701
fun Y => ⟨F.inv.obj Y, ⟨F.asEquivalence.counitIso.app Y⟩⟩⟩
702-
#align category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.ess_surj_of_equivalence
702+
#align category_theory.equivalence.ess_surj_of_equivalence CategoryTheory.Equivalence.essSurj_of_equivalence
703703

704704
-- see Note [lower instance priority]
705705
/-- An equivalence is faithful.
@@ -719,8 +719,8 @@ See <https://stacks.math.columbia.edu/tag/02C3>.
719719
-/
720720
instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : Full F
721721
where
722-
preimage := @fun X Y f => F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y
723-
witness := @fun X Y f =>
722+
preimage {X Y} f := F.asEquivalence.unit.app X ≫ F.inv.map f ≫ F.asEquivalence.unitInv.app Y
723+
witness {X Y} f :=
724724
F.inv.map_injective <| by
725725
simpa only [IsEquivalence.inv_fun_map, assoc, Iso.inv_hom_id_app_assoc,
726726
Iso.inv_hom_id_app] using comp_id _
@@ -730,9 +730,9 @@ instance (priority := 100) fullOfEquivalence (F : C ⥤ D) [IsEquivalence F] : F
730730
private noncomputable def equivalenceInverse (F : C ⥤ D) [Full F] [Faithful F] [EssSurj F] : D ⥤ C
731731
where
732732
obj X := F.objPreimage X
733-
map := @fun X Y f => F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv)
733+
map {X Y} f := F.preimage ((F.objObjPreimageIso X).hom ≫ f ≫ (F.objObjPreimageIso Y).inv)
734734
map_id X := by apply F.map_injective; aesop_cat
735-
map_comp := @fun X Y Z f g => by apply F.map_injective; simp
735+
map_comp {X Y Z} f g := by apply F.map_injective; simp
736736
-- #align category_theory.equivalence.equivalence_inverse CategoryTheory.Equivalence.equivalenceInverse
737737
/- Porting note: this is a private def in mathlib -/
738738

@@ -745,7 +745,7 @@ noncomputable def ofFullyFaithfullyEssSurj (F : C ⥤ D) [Full F] [Faithful F] [
745745
IsEquivalence F :=
746746
IsEquivalence.mk (equivalenceInverse F)
747747
(NatIso.ofComponents (fun X => (F.preimageIso <| F.objObjPreimageIso <| F.obj X).symm)
748-
@fun X Y f => by
748+
fun f => by
749749
apply F.map_injective
750750
aesop_cat)
751751
(NatIso.ofComponents F.objObjPreimageIso (by aesop_cat))
@@ -763,9 +763,8 @@ theorem inverse_map_inj_iff (e : C ≌ D) {X Y : D} (f g : X ⟶ Y) :
763763
functor_map_inj_iff e.symm f g
764764
#align category_theory.equivalence.inverse_map_inj_iff CategoryTheory.Equivalence.inverse_map_inj_iff
765765

766-
instance essSurjInducedFunctor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e)
767-
where mem_essImage Y :=
768-
⟨e.symm Y, by simpa using ⟨default⟩⟩
766+
instance essSurjInducedFunctor {C' : Type _} (e : C' ≃ D) : EssSurj (inducedFunctor e) where
767+
mem_essImage Y := ⟨e.symm Y, by simpa using ⟨default⟩⟩
769768
#align category_theory.equivalence.ess_surj_induced_functor CategoryTheory.Equivalence.essSurjInducedFunctor
770769

771770
noncomputable instance inducedFunctorOfEquiv {C' : Type _} (e : C' ≃ D) :

0 commit comments

Comments
 (0)