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

Commit 8050ac7

Browse files
committed
feat(set_theory/zfc/basic): inj lemmas (#15545)
We rename two existing lemmas from `*_inj` to `*_injective` to match mathlib convention, and add the corresponding `inj` lemmas.
1 parent c6357ee commit 8050ac7

File tree

1 file changed

+13
-8
lines changed

1 file changed

+13
-8
lines changed

src/set_theory/zfc/basic.lean

Lines changed: 13 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -594,8 +594,10 @@ mem_sUnion.2 ⟨z, hz, hy⟩
594594
@[simp] theorem sUnion_singleton {x : Set.{u}} : ⋃₀ ({x} : Set) = x :=
595595
ext $ λ y, by simp_rw [mem_sUnion, exists_prop, mem_singleton, exists_eq_left]
596596

597-
theorem singleton_inj {x y : Set.{u}} (H : ({x} : Set) = {y}) : x = y :=
598-
let this := congr_arg sUnion H in by rwa [sUnion_singleton, sUnion_singleton] at this
597+
theorem singleton_injective : function.injective (@singleton Set Set _) :=
598+
λ x y H, let this := congr_arg sUnion H in by rwa [sUnion_singleton, sUnion_singleton] at this
599+
600+
@[simp] theorem singleton_inj {x y : Set} : ({x} : Set) = {y} ↔ x = y := singleton_injective.eq_iff
599601

600602
/-- The binary union operation -/
601603
protected def union (x y : Set.{u}) : Set.{u} := ⋃₀ {x, y}
@@ -684,13 +686,13 @@ begin
684686
{ rintro (rfl|rfl); [left, right]; assumption }
685687
end
686688

687-
theorem pair_inj {x y x' y' : Set.{u}} (H : pair x y = pair x' y') : x = x' ∧ y = y' :=
688-
begin
689+
theorem pair_injective : function.injective2 pair :=
690+
λ x x' y y' H, begin
689691
have ae := ext_iff.2 H,
690692
simp only [pair, mem_pair] at ae,
691693
obtain rfl : x = x',
692694
{ cases (ae {x}).1 (by simp) with h h,
693-
{ exact singleton_inj h },
695+
{ exact singleton_injective h },
694696
{ have m : x' ∈ ({x} : Set),
695697
{ simp [h] },
696698
rw mem_singleton.mp m } },
@@ -708,6 +710,9 @@ begin
708710
{ simp [yy'] } }
709711
end
710712

713+
@[simp] theorem pair_inj {x y x' y' : Set} : pair x y = pair x' y' ↔ x = x' ∧ y = y' :=
714+
pair_injective.eq_iff
715+
711716
/-- The cartesian product, `{(a, b) | a ∈ x, b ∈ y}` -/
712717
def prod : Set.{u} → Set.{u} → Set.{u} := pair_sep (λ a b, true)
713718

@@ -716,7 +721,7 @@ by simp [prod]
716721

717722
@[simp] theorem pair_mem_prod {x y a b : Set.{u}} : pair a b ∈ prod x y ↔ a ∈ x ∧ b ∈ y :=
718723
⟨λ h, let ⟨a', a'x, b', b'y, e⟩ := mem_prod.1 h in
719-
match a', b', pair_inj e, a'x, b'y with ._, ._, ⟨rfl, rfl⟩, ax, bY := ⟨ax, bY⟩ end,
724+
match a', b', pair_injective e, a'x, b'y with ._, ._, ⟨rfl, rfl⟩, ax, bY := ⟨ax, bY⟩ end,
720725
λ ⟨ax, bY⟩, mem_prod.2 ⟨a, ax, b, bY, rfl⟩⟩
721726

722727
/-- `is_func x y f` is the assertion that `f` is a subset of `x × y` which relates to each element
@@ -746,7 +751,7 @@ mem_image
746751

747752
theorem map_unique {f : Set.{u} → Set.{u}} [H : definable 1 f] {x z : Set.{u}} (zx : z ∈ x) :
748753
∃! w, pair z w ∈ map f x :=
749-
⟨f z, image.mk _ _ zx, λ y yx, let ⟨w, wx, we⟩ := mem_image.1 yx, ⟨wz, fy⟩ := pair_inj we in
754+
⟨f z, image.mk _ _ zx, λ y yx, let ⟨w, wx, we⟩ := mem_image.1 yx, ⟨wz, fy⟩ := pair_injective we in
750755
by rw[←fy, wz]⟩
751756

752757
@[simp] theorem map_is_func {f : Set → Set} [H : definable 1 f] {x y : Set} :
@@ -888,7 +893,7 @@ namespace Set
888893
{x y : Set.{u}} (h : y ∈ x) :
889894
(Set.map f x ′ y : Class.{u}) = f y :=
890895
Class.iota_val _ _ (λ z, by { rw [Class.to_Set_of_Set, Class.mem_hom_right, mem_map], exact
891-
⟨λ ⟨w, wz, pr⟩, let ⟨wy, fw⟩ := Set.pair_inj pr in by rw[←fw, wy],
896+
⟨λ ⟨w, wz, pr⟩, let ⟨wy, fw⟩ := Set.pair_injective pr in by rw[←fw, wy],
892897
λ e, by { subst e, exact ⟨_, h, rfl⟩ }⟩ })
893898

894899
variables (x : Set.{u}) (h : ∅ ∉ x)

0 commit comments

Comments
 (0)