@@ -594,8 +594,10 @@ mem_sUnion.2 ⟨z, hz, hy⟩
594
594
@[simp] theorem sUnion_singleton {x : Set.{u}} : ⋃₀ ({x} : Set) = x :=
595
595
ext $ λ y, by simp_rw [mem_sUnion, exists_prop, mem_singleton, exists_eq_left]
596
596
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
599
601
600
602
/-- The binary union operation -/
601
603
protected def union (x y : Set.{u}) : Set.{u} := ⋃₀ {x, y}
@@ -684,13 +686,13 @@ begin
684
686
{ rintro (rfl|rfl); [left, right]; assumption }
685
687
end
686
688
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
689
691
have ae := ext_iff.2 H,
690
692
simp only [pair, mem_pair] at ae,
691
693
obtain rfl : x = x',
692
694
{ cases (ae {x}).1 (by simp) with h h,
693
- { exact singleton_inj h },
695
+ { exact singleton_injective h },
694
696
{ have m : x' ∈ ({x} : Set),
695
697
{ simp [h] },
696
698
rw mem_singleton.mp m } },
@@ -708,6 +710,9 @@ begin
708
710
{ simp [yy'] } }
709
711
end
710
712
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
+
711
716
/-- The cartesian product, `{(a, b) | a ∈ x, b ∈ y}` -/
712
717
def prod : Set.{u} → Set.{u} → Set.{u} := pair_sep (λ a b, true)
713
718
@@ -716,7 +721,7 @@ by simp [prod]
716
721
717
722
@[simp] theorem pair_mem_prod {x y a b : Set.{u}} : pair a b ∈ prod x y ↔ a ∈ x ∧ b ∈ y :=
718
723
⟨λ 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 ,
720
725
λ ⟨ax, bY⟩, mem_prod.2 ⟨a, ax, b, bY, rfl⟩⟩
721
726
722
727
/-- `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
746
751
747
752
theorem map_unique {f : Set.{u} → Set.{u}} [H : definable 1 f] {x z : Set.{u}} (zx : z ∈ x) :
748
753
∃! 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
750
755
by rw[←fy, wz]⟩
751
756
752
757
@[simp] theorem map_is_func {f : Set → Set} [H : definable 1 f] {x y : Set} :
@@ -888,7 +893,7 @@ namespace Set
888
893
{x y : Set.{u}} (h : y ∈ x) :
889
894
(Set.map f x ′ y : Class.{u}) = f y :=
890
895
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],
892
897
λ e, by { subst e, exact ⟨_, h, rfl⟩ }⟩ })
893
898
894
899
variables (x : Set.{u}) (h : ∅ ∉ x)
0 commit comments