Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 40fc58c

Browse files
committed
feat(data/quot): quotient.out is injective (#13584)
1 parent 821e7c8 commit 40fc58c

File tree

1 file changed

+4
-1
lines changed

1 file changed

+4
-1
lines changed

src/data/quot.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -255,9 +255,12 @@ end
255255
x.out ≈ y.out ↔ x = y :=
256256
by rw [← quotient.eq_mk_iff_out, quotient.out_eq]
257257

258+
lemma quotient.out_injective {s : setoid α} : function.injective (@quotient.out α s) :=
259+
λ a b h, quotient.out_equiv_out.1 $ h ▸ setoid.refl _
260+
258261
@[simp] lemma quotient.out_inj {s : setoid α} {x y : quotient s} :
259262
x.out = y.out ↔ x = y :=
260-
⟨λ h, quotient.out_equiv_out.1 $ h ▸ setoid.refl _, λ h, h ▸ rfl⟩
263+
⟨λ h, quotient.out_injective h, λ h, h ▸ rfl⟩
261264

262265
section pi
263266

0 commit comments

Comments
 (0)