Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: subsingleton tactic #12525

Closed
wants to merge 30 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
30 commits
Select commit Hold shift + click to select a range
ec95b7f
feat: `subsingleton` tactic
kmill Apr 29, 2024
ca956b5
fix bug, add tests
kmill Apr 29, 2024
09bd5eb
factor out meta
kmill Apr 29, 2024
93ab853
arguments to Subsingleton
kmill Apr 30, 2024
2dd4958
uses of `subsingleton`
kmill Apr 30, 2024
215cb5a
use of subsingleton
kmill Apr 30, 2024
aa33827
doc
kmill Apr 30, 2024
7956d7f
more
kmill Apr 30, 2024
b5acff0
more
kmill Apr 30, 2024
744067a
revert
kmill Apr 30, 2024
5264706
handle universe level abstraction
kmill Apr 30, 2024
f96601d
more
kmill Apr 30, 2024
6d8450d
remove warning for unused local instances
kmill Apr 30, 2024
6ac7cfe
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Apr 30, 2024
181de1b
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill May 6, 2024
9e5cefc
remove `Subsingleton.helim` support
kmill May 22, 2024
fb98761
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill May 22, 2024
c534bdb
fixes
kmill May 22, 2024
1d8604e
noshake
kmill May 22, 2024
48ebbc8
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill May 23, 2024
ffad8a7
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill May 25, 2024
e64ed01
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill May 25, 2024
2552727
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill May 31, 2024
0b389b9
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Jun 3, 2024
b29c8b5
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Jun 7, 2024
40085f5
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Jun 9, 2024
8669357
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Jun 17, 2024
f9a92ed
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Jun 22, 2024
4e86a4a
Merge branch 'master' of https://github.com/leanprover-community/math…
kmill Jul 1, 2024
3861d3d
fix test
kmill Jul 1, 2024
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4023,6 +4023,7 @@ import Mathlib.Tactic.Simps.NotationClass
import Mathlib.Tactic.SlimCheck
import Mathlib.Tactic.SplitIfs
import Mathlib.Tactic.Spread
import Mathlib.Tactic.Subsingleton
import Mathlib.Tactic.Substs
import Mathlib.Tactic.SuccessIfFailWithMsg
import Mathlib.Tactic.SudoSetOption
Expand Down
7 changes: 3 additions & 4 deletions Mathlib/Algebra/BigOperators/Fin.lean
Original file line number Diff line number Diff line change
Expand Up @@ -342,8 +342,7 @@ def finFunctionFinEquiv {m n : ℕ} : (Fin n → Fin m) ≃ Fin (m ^ n) :=
fun a => by
dsimp
induction' n with n ih
· haveI : Subsingleton (Fin (m ^ 0)) := (finCongr <| pow_zero _).subsingleton
exact Subsingleton.elim _ _
· subsingleton [(finCongr <| pow_zero _).subsingleton]
simp_rw [Fin.forall_iff, Fin.ext_iff] at ih
ext
simp_rw [Fin.sum_univ_succ, Fin.val_zero, Fin.val_succ, pow_zero, Nat.div_one,
Expand Down Expand Up @@ -398,9 +397,9 @@ def finPiFinEquiv {m : ℕ} {n : Fin m → ℕ} : (∀ i : Fin m, Fin (n i)) ≃
intro a; revert a; dsimp only [Fin.val_mk]
refine Fin.consInduction ?_ ?_ n
· intro a
haveI : Subsingleton (Fin (∏ i : Fin 0, i.elim0)) :=
have : Subsingleton (Fin (∏ i : Fin 0, i.elim0)) :=
(finCongr <| prod_empty).subsingleton
exact Subsingleton.elim _ _
subsingleton
· intro n x xs ih a
simp_rw [Fin.forall_iff, Fin.ext_iff] at ih
ext
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/BigOperators/Group/Finset.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2324,7 +2324,7 @@ theorem prod_empty {α β : Type*} [CommMonoid β] [IsEmpty α] [Fintype α] (f
@[to_additive]
theorem prod_subsingleton {α β : Type*} [CommMonoid β] [Subsingleton α] [Fintype α] (f : α → β)
(a : α) : ∏ x : α, f x = f a := by
haveI : Unique α := uniqueOfSubsingleton a
have : Unique α := uniqueOfSubsingleton a
rw [prod_unique f, Subsingleton.elim default a]
#align fintype.prod_subsingleton Fintype.prod_subsingleton
#align fintype.sum_subsingleton Fintype.sum_subsingleton
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Algebra/Category/Grp/Zero.lean
Original file line number Diff line number Diff line change
Expand Up @@ -31,7 +31,7 @@ theorem isZero_of_subsingleton (G : Grp) [Subsingleton G] : IsZero G := by
have : x = 1 := Subsingleton.elim _ _
rw [this, map_one, map_one]
· ext
apply Subsingleton.elim
subsingleton
set_option linter.uppercaseLean3 false in
#align Group.is_zero_of_subsingleton Grp.isZero_of_subsingleton
set_option linter.uppercaseLean3 false in
Expand All @@ -52,7 +52,7 @@ theorem isZero_of_subsingleton (G : CommGrp) [Subsingleton G] : IsZero G := by
have : x = 1 := Subsingleton.elim _ _
rw [this, map_one, map_one]
· ext
apply Subsingleton.elim
subsingleton
set_option linter.uppercaseLean3 false in
#align CommGroup.is_zero_of_subsingleton CommGrp.isZero_of_subsingleton
set_option linter.uppercaseLean3 false in
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/ModuleCat/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -206,7 +206,7 @@ theorem isZero_of_subsingleton (M : ModuleCat R) [Subsingleton M] : IsZero M whe
simp⟩⟩
unique_from X := ⟨⟨⟨(0 : X →ₗ[R] M)⟩, fun f => by
ext x
apply Subsingleton.elim⟩⟩
subsingleton⟩⟩
#align Module.is_zero_of_subsingleton ModuleCat.isZero_of_subsingleton

instance : HasZeroObject (ModuleCat.{v} R) :=
Expand Down
12 changes: 6 additions & 6 deletions Mathlib/Algebra/DirectLimit.lean
Original file line number Diff line number Diff line change
Expand Up @@ -200,7 +200,7 @@ def map (g : (i : ι) → G i →ₗ[R] G' i) (hg : ∀ i j h, g j ∘ₗ f i j
DirectLimit G f →ₗ[R] DirectLimit G' f' :=
lift _ _ _ _ (fun i ↦ of _ _ _ _ _ ∘ₗ g i) fun i j h g ↦ by
cases isEmpty_or_nonempty ι
· exact Subsingleton.elim _ _
· subsingleton
· have eq1 := LinearMap.congr_fun (hg i j h) g
simp only [LinearMap.coe_comp, Function.comp_apply] at eq1 ⊢
rw [eq1, of_f]
Expand All @@ -213,7 +213,7 @@ def map (g : (i : ι) → G i →ₗ[R] G' i) (hg : ∀ i j h, g j ∘ₗ f i j

@[simp] lemma map_id [IsDirected ι (· ≤ ·)] :
map (fun i ↦ LinearMap.id) (fun _ _ _ ↦ rfl) = LinearMap.id (R := R) (M := DirectLimit G f) :=
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦
x.induction_on fun i g ↦ by simp

lemma map_comp [IsDirected ι (· ≤ ·)]
Expand All @@ -225,7 +225,7 @@ lemma map_comp [IsDirected ι (· ≤ ·)]
(map (fun i ↦ g₂ i ∘ₗ g₁ i) fun i j h ↦ by
rw [LinearMap.comp_assoc, hg₁ i, ← LinearMap.comp_assoc, hg₂ i, LinearMap.comp_assoc] :
DirectLimit G f →ₗ[R] DirectLimit G'' f'') :=
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦
x.induction_on fun i g ↦ by simp

open LinearEquiv LinearMap in
Expand Down Expand Up @@ -473,7 +473,7 @@ def map (g : (i : ι) → G i →+ G' i)
DirectLimit G f →+ DirectLimit G' f' :=
lift _ _ _ (fun i ↦ (of _ _ _).comp (g i)) fun i j h g ↦ by
cases isEmpty_or_nonempty ι
· exact Subsingleton.elim _ _
· subsingleton
· have eq1 := DFunLike.congr_fun (hg i j h) g
simp only [AddMonoidHom.coe_comp, Function.comp_apply] at eq1 ⊢
rw [eq1, of_f]
Expand All @@ -486,7 +486,7 @@ def map (g : (i : ι) → G i →+ G' i)

@[simp] lemma map_id [IsDirected ι (· ≤ ·)] :
map (fun i ↦ AddMonoidHom.id _) (fun _ _ _ ↦ rfl) = AddMonoidHom.id (DirectLimit G f) :=
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦
x.induction_on fun i g ↦ by simp

lemma map_comp [IsDirected ι (· ≤ ·)]
Expand All @@ -499,7 +499,7 @@ lemma map_comp [IsDirected ι (· ≤ ·)]
rw [AddMonoidHom.comp_assoc, hg₁ i, ← AddMonoidHom.comp_assoc, hg₂ i,
AddMonoidHom.comp_assoc] :
DirectLimit G f →+ DirectLimit G'' f'') :=
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (fun _ ↦ Subsingleton.elim _ _) fun _ ↦
DFunLike.ext _ _ fun x ↦ (isEmpty_or_nonempty ι).elim (by subsingleton) fun _ ↦
x.induction_on fun i g ↦ by simp

/--
Expand Down
8 changes: 4 additions & 4 deletions Mathlib/Algebra/Group/Units.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,7 @@ import Mathlib.Algebra.Group.Commute.Defs
import Mathlib.Logic.Unique
import Mathlib.Tactic.Nontriviality
import Mathlib.Tactic.Lift
import Mathlib.Tactic.Subsingleton

#align_import algebra.group.units from "leanprover-community/mathlib"@"e8638a0fcaf73e4500469f368ef9494e495099b3"

Expand Down Expand Up @@ -672,7 +673,7 @@ theorem isUnit_iff_exists_and_exists [Monoid M] {a : M} :

@[to_additive (attr := nontriviality)]
theorem isUnit_of_subsingleton [Monoid M] [Subsingleton M] (a : M) : IsUnit a :=
⟨⟨a, a, Subsingleton.elim _ _, Subsingleton.elim _ _⟩, rfl⟩
⟨⟨a, a, by subsingleton, by subsingleton⟩, rfl⟩
#align is_unit_of_subsingleton isUnit_of_subsingleton
#align is_add_unit_of_subsingleton isAddUnit_of_subsingleton

Expand All @@ -683,7 +684,7 @@ instance [Monoid M] : CanLift M Mˣ Units.val IsUnit :=
/-- A subsingleton `Monoid` has a unique unit. -/
@[to_additive "A subsingleton `AddMonoid` has a unique additive unit."]
instance [Monoid M] [Subsingleton M] : Unique Mˣ where
uniq a := Units.val_eq_one.mp <| Subsingleton.elim (a : M) 1
uniq _ := Units.val_eq_one.mp (by subsingleton)

@[to_additive (attr := simp)]
protected theorem Units.isUnit [Monoid M] (u : Mˣ) : IsUnit (u : M) :=
Expand Down Expand Up @@ -735,8 +736,7 @@ lemma IsUnit.exists_left_inv {a : M} (h : IsUnit a) : ∃ b, b * a = 1 := by
#align is_unit.pow IsUnit.pow
#align is_add_unit.nsmul IsAddUnit.nsmul

theorem units_eq_one [Unique Mˣ] (u : Mˣ) : u = 1 :=
Subsingleton.elim u 1
theorem units_eq_one [Unique Mˣ] (u : Mˣ) : u = 1 := by subsingleton
#align units_eq_one units_eq_one

@[to_additive] lemma isUnit_iff_eq_one [Unique Mˣ] {x : M} : IsUnit x ↔ x = 1 :=
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Algebra/Homology/HomologicalComplex.lean
Original file line number Diff line number Diff line change
Expand Up @@ -317,8 +317,8 @@ theorem isZero_zero [HasZeroObject V] : IsZero (zero : HomologicalComplex V c) :
refine ⟨fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩⟩
all_goals
ext
dsimp [zero]
apply Subsingleton.elim
dsimp only [zero]
subsingleton
#align homological_complex.is_zero_zero HomologicalComplex.isZero_zero

instance [HasZeroObject V] : HasZeroObject (HomologicalComplex V c) :=
Expand Down
3 changes: 1 addition & 2 deletions Mathlib/Algebra/Lie/Engel.lean
Original file line number Diff line number Diff line change
Expand Up @@ -166,8 +166,7 @@ theorem LieAlgebra.isEngelian_of_subsingleton [Subsingleton L] : LieAlgebra.IsEn
intro M _i1 _i2 _i3 _i4 _h
use 1
suffices (⊤ : LieIdeal R L) = ⊥ by simp [this]
haveI := (LieSubmodule.subsingleton_iff R L L).mpr inferInstance
apply Subsingleton.elim
subsingleton [(LieSubmodule.subsingleton_iff R L L).mpr inferInstance]
#align lie_algebra.is_engelian_of_subsingleton LieAlgebra.isEngelian_of_subsingleton

theorem Function.Surjective.isEngelian {f : L →ₗ⁅R⁆ L₂} (hf : Function.Surjective f)
Expand Down
8 changes: 4 additions & 4 deletions Mathlib/Algebra/Order/Antidiag/Prod.lean
Original file line number Diff line number Diff line change
Expand Up @@ -64,17 +64,17 @@ attribute [simp] mem_antidiagonal
variable {A : Type*}

/-- All `HasAntidiagonal` instances are equal -/
instance [AddMonoid A] : Subsingleton (HasAntidiagonal A) :=
by
instance [AddMonoid A] : Subsingleton (HasAntidiagonal A) where
allEq := by
rintro ⟨a, ha⟩ ⟨b, hb⟩
congr with n xy
rw [ha, hb]
rw [ha, hb]

-- The goal of this lemma is to allow to rewrite antidiagonal
-- when the decidability instances obsucate Lean
lemma hasAntidiagonal_congr (A : Type*) [AddMonoid A]
[H1 : HasAntidiagonal A] [H2 : HasAntidiagonal A] :
H1.antidiagonal = H2.antidiagonal := by congr!; apply Subsingleton.elim
H1.antidiagonal = H2.antidiagonal := by congr!; subsingleton

theorem swap_mem_antidiagonal [AddCommMonoid A] [HasAntidiagonal A] {n : A} {xy : A × A}:
xy.swap ∈ antidiagonal n ↔ xy ∈ antidiagonal n := by
Expand Down
30 changes: 15 additions & 15 deletions Mathlib/Algebra/PUnitInstances.lean
Original file line number Diff line number Diff line change
Expand Up @@ -87,10 +87,10 @@ instance normalizedGCDMonoid : NormalizedGCDMonoid PUnit where
normUnit_zero := rfl
normUnit_mul := by intros; rfl
normUnit_coe_units := by intros; rfl
gcd_dvd_left _ _ := ⟨unit, Subsingleton.elim _ _
gcd_dvd_right _ _ := ⟨unit, Subsingleton.elim _ _
dvd_gcd {_ _} _ _ _ := ⟨unit, Subsingleton.elim _ _
gcd_mul_lcm _ _ := ⟨1, Subsingleton.elim _ _
gcd_dvd_left _ _ := ⟨unit, by subsingleton
gcd_dvd_right _ _ := ⟨unit, by subsingleton
dvd_gcd {_ _} _ _ _ := ⟨unit, by subsingleton
gcd_mul_lcm _ _ := ⟨1, by subsingleton
lcm_zero_left := by intros; rfl
lcm_zero_right := by intros; rfl
normalize_gcd := by intros; rfl
Expand All @@ -112,7 +112,7 @@ theorem norm_unit_eq {x : PUnit} : normUnit x = 1 :=
#align punit.norm_unit_eq PUnit.norm_unit_eq

instance canonicallyOrderedAddCommMonoid : CanonicallyOrderedAddCommMonoid PUnit where
exists_add_of_le {_ _} _ := ⟨unit, Subsingleton.elim _ _
exists_add_of_le {_ _} _ := ⟨unit, by subsingleton
add_le_add_left _ _ _ _ := trivial
le_self_add _ _ := trivial

Expand Down Expand Up @@ -150,23 +150,23 @@ instance instIsScalarTowerOfSMul [SMul R S] : IsScalarTower R S PUnit :=

instance smulWithZero [Zero R] : SMulWithZero R PUnit where
__ := PUnit.smul
smul_zero _ := Subsingleton.elim _ _
zero_smul _ := Subsingleton.elim _ _
smul_zero := by subsingleton
zero_smul := by subsingleton

instance mulAction [Monoid R] : MulAction R PUnit where
__ := PUnit.smul
one_smul _ := Subsingleton.elim _ _
mul_smul _ _ _ := Subsingleton.elim _ _
one_smul := by subsingleton
mul_smul := by subsingleton

instance distribMulAction [Monoid R] : DistribMulAction R PUnit where
__ := PUnit.mulAction
smul_zero _ := Subsingleton.elim _ _
smul_add _ _ _ := Subsingleton.elim _ _
smul_zero := by subsingleton
smul_add := by subsingleton

instance mulDistribMulAction [Monoid R] : MulDistribMulAction R PUnit where
__ := PUnit.mulAction
smul_mul _ _ _ := Subsingleton.elim _ _
smul_one _ := Subsingleton.elim _ _
smul_mul := by subsingleton
smul_one := by subsingleton

instance mulSemiringAction [Semiring R] : MulSemiringAction R PUnit :=
{ PUnit.distribMulAction, PUnit.mulDistribMulAction with }
Expand All @@ -176,7 +176,7 @@ instance mulActionWithZero [MonoidWithZero R] : MulActionWithZero R PUnit :=

instance module [Semiring R] : Module R PUnit where
__ := PUnit.distribMulAction
add_smul _ _ _ := Subsingleton.elim _ _
zero_smul _ := Subsingleton.elim _ _
add_smul := by subsingleton
zero_smul := by subsingleton

end PUnit
7 changes: 3 additions & 4 deletions Mathlib/AlgebraicTopology/CechNerve.lean
Original file line number Diff line number Diff line change
Expand Up @@ -378,12 +378,11 @@ def wideCospan.limitCone [Finite ι] (X : C) : LimitCone (wideCospan ι X) where
cases f
· cases i
all_goals dsimp; simp
· dsimp
simp only [terminal.comp_from]
exact Subsingleton.elim _ _ } }
· simp only [Functor.const_obj_obj, Functor.const_obj_map, terminal.comp_from]
subsingleton } }
isLimit :=
{ lift := fun s => Limits.Pi.lift fun j => s.π.app (some j)
fac := fun s j => Option.casesOn j (Subsingleton.elim _ _) fun j => limit.lift_π _ _
fac := fun s j => Option.casesOn j (by subsingleton) fun j => limit.lift_π _ _
uniq := fun s f h => by
dsimp
ext j
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/AlgebraicTopology/ExtraDegeneracy.lean
Original file line number Diff line number Diff line change
Expand Up @@ -198,7 +198,7 @@ protected noncomputable def extraDegeneracy (Δ : SimplexCategory) :
(shift (objEquiv _ _ f))
s'_comp_ε := by
dsimp
apply Subsingleton.elim
subsingleton
s₀_comp_δ₁ := by
dsimp
ext1 x
Expand Down
22 changes: 8 additions & 14 deletions Mathlib/CategoryTheory/Filtered/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -97,9 +97,7 @@ class IsFiltered extends IsFilteredOrEmpty C : Prop where
instance (priority := 100) isFilteredOrEmpty_of_semilatticeSup (α : Type u) [SemilatticeSup α] :
IsFilteredOrEmpty α where
cocone_objs X Y := ⟨X ⊔ Y, homOfLE le_sup_left, homOfLE le_sup_right, trivial⟩
cocone_maps X Y f g := ⟨Y, 𝟙 _, by
apply ULift.ext
apply Subsingleton.elim⟩
cocone_maps X Y f g := ⟨Y, 𝟙 _, by subsingleton⟩
#align category_theory.is_filtered_or_empty_of_semilattice_sup CategoryTheory.isFilteredOrEmpty_of_semilatticeSup

instance (priority := 100) isFiltered_of_semilatticeSup_nonempty (α : Type u) [SemilatticeSup α]
Expand All @@ -111,9 +109,7 @@ instance (priority := 100) isFilteredOrEmpty_of_directed_le (α : Type u) [Preor
cocone_objs X Y :=
let ⟨Z, h1, h2⟩ := exists_ge_ge X Y
⟨Z, homOfLE h1, homOfLE h2, trivial⟩
cocone_maps X Y f g := ⟨Y, 𝟙 _, by
apply ULift.ext
apply Subsingleton.elim⟩
cocone_maps X Y f g := ⟨Y, 𝟙 _, by subsingleton⟩
#align category_theory.is_filtered_or_empty_of_directed_le CategoryTheory.isFilteredOrEmpty_of_directed_le

instance (priority := 100) isFiltered_of_directed_le_nonempty (α : Type u) [Preorder α]
Expand All @@ -126,10 +122,8 @@ example (α : Type u) [SemilatticeSup α] [OrderBot α] : IsFiltered α := by in
example (α : Type u) [SemilatticeSup α] [OrderTop α] : IsFiltered α := by infer_instance

instance : IsFiltered (Discrete PUnit) where
cocone_objs X Y := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, ⟨⟨Subsingleton.elim _ _⟩⟩, trivial⟩
cocone_maps X Y f g := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, by
apply ULift.ext
apply Subsingleton.elim⟩
cocone_objs X Y := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, ⟨⟨by subsingleton⟩⟩, trivial⟩
cocone_maps X Y f g := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, by subsingleton⟩

namespace IsFiltered

Expand Down Expand Up @@ -568,7 +562,7 @@ instance (priority := 100) isCofilteredOrEmpty_of_semilatticeInf (α : Type u) [
cone_objs X Y := ⟨X ⊓ Y, homOfLE inf_le_left, homOfLE inf_le_right, trivial⟩
cone_maps X Y f g := ⟨X, 𝟙 _, by
apply ULift.ext
apply Subsingleton.elim
subsingleton
#align category_theory.is_cofiltered_or_empty_of_semilattice_inf CategoryTheory.isCofilteredOrEmpty_of_semilatticeInf

instance (priority := 100) isCofiltered_of_semilatticeInf_nonempty (α : Type u) [SemilatticeInf α]
Expand All @@ -582,7 +576,7 @@ instance (priority := 100) isCofilteredOrEmpty_of_directed_ge (α : Type u) [Pre
⟨Z, homOfLE hX, homOfLE hY, trivial⟩
cone_maps X Y f g := ⟨X, 𝟙 _, by
apply ULift.ext
apply Subsingleton.elim
subsingleton
#align category_theory.is_cofiltered_or_empty_of_directed_ge CategoryTheory.isCofilteredOrEmpty_of_directed_ge

instance (priority := 100) isCofiltered_of_directed_ge_nonempty (α : Type u) [Preorder α]
Expand All @@ -595,10 +589,10 @@ example (α : Type u) [SemilatticeInf α] [OrderBot α] : IsCofiltered α := by
example (α : Type u) [SemilatticeInf α] [OrderTop α] : IsCofiltered α := by infer_instance

instance : IsCofiltered (Discrete PUnit) where
cone_objs X Y := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, ⟨⟨Subsingleton.elim _ _⟩⟩, trivial⟩
cone_objs X Y := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, ⟨⟨by subsingleton⟩⟩, trivial⟩
cone_maps X Y f g := ⟨⟨PUnit.unit⟩, ⟨⟨by trivial⟩⟩, by
apply ULift.ext
apply Subsingleton.elim
subsingleton

namespace IsCofiltered

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/CategoryTheory/Filtered/Final.lean
Original file line number Diff line number Diff line change
Expand Up @@ -182,7 +182,7 @@ theorem Functor.final_iff_of_isFiltered [IsFilteredOrEmpty C] :
· intro d c s s'
have : colimit.ι (F ⋙ coyoneda.obj (op d)) c s = colimit.ι (F ⋙ coyoneda.obj (op d)) c s' := by
apply (Final.colimitCompCoyonedaIso F d).toEquiv.injective
exact Subsingleton.elim _ _
subsingleton
obtain ⟨c', t₁, t₂, h⟩ := (Types.FilteredColimit.colimit_eq_iff.{v₁, v₁, v₁} _).mp this
refine ⟨IsFiltered.coeq t₁ t₂, t₁ ≫ IsFiltered.coeqHom t₁ t₂, ?_⟩
conv_rhs => rw [IsFiltered.coeq_condition t₁ t₂]
Expand Down
6 changes: 3 additions & 3 deletions Mathlib/CategoryTheory/Functor/Flat.lean
Original file line number Diff line number Diff line change
Expand Up @@ -218,7 +218,7 @@ noncomputable def preservesFiniteLimitsIffFlat [HasFiniteLimits C] (F : C ⥤ D)
dsimp only [preservesFiniteLimitsOfPreservesFiniteLimitsOfSize]
congr
-- Porting note: this next line wasn't needed in lean 3
apply Subsingleton.elim
subsingleton

#align category_theory.preserves_finite_limits_iff_flat CategoryTheory.preservesFiniteLimitsIffFlat

Expand Down Expand Up @@ -319,15 +319,15 @@ noncomputable def preservesFiniteLimitsIffLanPreservesFiniteLimits (F : C ⥤ D)
-- indicates that it was doing the same as `dsimp only`
dsimp only [preservesFiniteLimitsOfPreservesFiniteLimitsOfSize]; congr
-- Porting note: next line wasn't necessary in lean 3
apply Subsingleton.elim
subsingleton
right_inv x := by
-- cases x; -- Porting note: not necessary in lean 4
dsimp only [lanPreservesFiniteLimitsOfPreservesFiniteLimits,
lanPreservesFiniteLimitsOfFlat,
preservesFiniteLimitsOfPreservesFiniteLimitsOfSize]
congr
-- Porting note: next line wasn't necessary in lean 3
apply Subsingleton.elim
subsingleton
set_option linter.uppercaseLean3 false in
#align category_theory.preserves_finite_limits_iff_Lan_preserves_finite_limits CategoryTheory.preservesFiniteLimitsIffLanPreservesFiniteLimits

Expand Down
Loading
Loading