Skip to content

Commit fe3b2b2

Browse files
committed
feat(Data/Sigma): add Sigma.fst_surjective etc (#9914)
- Add `Sigma.fst_surjective`, `Sigma.fst_surjective_iff`, `Sigma.fst_injective`, and `Sigma.fst_injective_iff`. - Move `sigma_mk_injective` up. - Open `Function` namespace, drop `Function.`. - Fix indentation.
1 parent 8853442 commit fe3b2b2

File tree

1 file changed

+27
-16
lines changed

1 file changed

+27
-16
lines changed

Mathlib/Data/Sigma/Basic.lean

Lines changed: 27 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,8 @@ types. To that effect, we have `PSigma`, which takes value in `Sort*` and carrie
3232
complicated universe signature as a consequence.
3333
-/
3434

35+
open Function
36+
3537
section Sigma
3638

3739
variable {α α₁ α₂ : Type*} {β : α → Type*} {β₁ : α₁ → Type*} {β₂ : α₂ → Type*}
@@ -76,7 +78,7 @@ theorem _root_.Function.eq_of_sigmaMk_comp {γ : Type*} [Nonempty γ]
7678
a = b ∧ HEq f g := by
7779
rcases ‹Nonempty γ› with ⟨i⟩
7880
obtain rfl : a = b := congr_arg Sigma.fst (congr_fun h i)
79-
simpa [Function.funext_iff] using h
81+
simpa [funext_iff] using h
8082

8183
/-- A specialized ext lemma for equality of sigma types over an indexed subtype. -/
8284
@[ext]
@@ -101,10 +103,27 @@ theorem «exists» {p : (Σa, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a
101103
#align sigma.exists Sigma.exists
102104

103105
lemma exists' {p : ∀ a, β a → Prop} : (∃ a b, p a b) ↔ ∃ x : Σ a, β a, p x.1 x.2 :=
104-
(Sigma.exists (p := λ x ↦ p x.1 x.2)).symm
106+
(Sigma.exists (p := λ x ↦ p x.1 x.2)).symm
105107

106108
lemma forall' {p : ∀ a, β a → Prop} : (∀ a b, p a b) ↔ ∀ x : Σ a, β a, p x.1 x.2 :=
107-
(Sigma.forall (p := λ x ↦ p x.1 x.2)).symm
109+
(Sigma.forall (p := λ x ↦ p x.1 x.2)).symm
110+
111+
theorem _root_.sigma_mk_injective {i : α} : Injective (@Sigma.mk α β i)
112+
| _, _, rfl => rfl
113+
#align sigma_mk_injective sigma_mk_injective
114+
115+
theorem fst_surjective [h : ∀ a, Nonempty (β a)] : Surjective (fst : (Σ a, β a) → α) := fun a ↦
116+
let ⟨b⟩ := h a; ⟨⟨a, b⟩, rfl⟩
117+
118+
theorem fst_surjective_iff : Surjective (fst : (Σ a, β a) → α) ↔ ∀ a, Nonempty (β a) :=
119+
fun h a ↦ let ⟨x, hx⟩ := h a; hx ▸ ⟨x.2⟩, @fst_surjective _ _⟩
120+
121+
theorem fst_injective [h : ∀ a, Subsingleton (β a)] : Injective (fst : (Σ a, β a) → α) := by
122+
rintro ⟨a₁, b₁⟩ ⟨a₂, b₂⟩ (rfl : a₁ = a₂)
123+
exact congr_arg (mk a₁) <| Subsingleton.elim _ _
124+
125+
theorem fst_injective_iff : Injective (fst : (Σ a, β a) → α) ↔ ∀ a, Subsingleton (β a) :=
126+
fun h _ ↦ ⟨fun _ _ ↦ sigma_mk_injective <| h rfl⟩, @fst_injective _ _⟩
108127

109128
/-- Map the left and right components of a sigma -/
110129
def map (f₁ : α₁ → α₂) (f₂ : ∀ a, β₁ a → β₂ (f₁ a)) (x : Sigma β₁) : Sigma β₂ :=
@@ -115,35 +134,27 @@ lemma map_mk (f₁ : α₁ → α₂) (f₂ : ∀ a, β₁ a → β₂ (f₁ a))
115134
map f₁ f₂ ⟨x, y⟩ = ⟨f₁ x, f₂ x y⟩ := rfl
116135
end Sigma
117136

118-
theorem sigma_mk_injective {i : α} : Function.Injective (@Sigma.mk α β i)
119-
| _, _, rfl => rfl
120-
#align sigma_mk_injective sigma_mk_injective
121-
122137
theorem Function.Injective.sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)}
123-
(h₁ : Function.Injective f₁) (h₂ : ∀ a, Function.Injective (f₂ a)) :
124-
Function.Injective (Sigma.map f₁ f₂)
138+
(h₁ : Injective f₁) (h₂ : ∀ a, Injective (f₂ a)) : Injective (Sigma.map f₁ f₂)
125139
| ⟨i, x⟩, ⟨j, y⟩, h => by
126140
obtain rfl : i = j := h₁ (Sigma.mk.inj_iff.mp h).1
127141
obtain rfl : x = y := h₂ i (sigma_mk_injective h)
128142
rfl
129143
#align function.injective.sigma_map Function.Injective.sigma_map
130144

131145
theorem Function.Injective.of_sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)}
132-
(h : Function.Injective (Sigma.map f₁ f₂)) (a : α₁) : Function.Injective (f₂ a) :=
133-
fun x y hxy ↦
146+
(h : Injective (Sigma.map f₁ f₂)) (a : α₁) : Injective (f₂ a) := fun x y hxy ↦
134147
sigma_mk_injective <| @h ⟨a, x⟩ ⟨a, y⟩ (Sigma.ext rfl (heq_of_eq hxy))
135148
#align function.injective.of_sigma_map Function.Injective.of_sigma_map
136149

137150
theorem Function.Injective.sigma_map_iff {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)}
138-
(h₁ : Function.Injective f₁) :
139-
Function.Injective (Sigma.map f₁ f₂) ↔ ∀ a, Function.Injective (f₂ a) :=
151+
(h₁ : Injective f₁) : Injective (Sigma.map f₁ f₂) ↔ ∀ a, Injective (f₂ a) :=
140152
fun h ↦ h.of_sigma_map, h₁.sigma_map⟩
141153
#align function.injective.sigma_map_iff Function.Injective.sigma_map_iff
142154

143155
theorem Function.Surjective.sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)}
144-
(h₁ : Function.Surjective f₁) (h₂ : ∀ a, Function.Surjective (f₂ a)) :
145-
Function.Surjective (Sigma.map f₁ f₂) := by
146-
simp only [Function.Surjective, Sigma.forall, h₁.forall]
156+
(h₁ : Surjective f₁) (h₂ : ∀ a, Surjective (f₂ a)) : Surjective (Sigma.map f₁ f₂) := by
157+
simp only [Surjective, Sigma.forall, h₁.forall]
147158
exact fun i ↦ (h₂ _).forall.2 fun x ↦ ⟨⟨i, x⟩, rfl⟩
148159
#align function.surjective.sigma_map Function.Surjective.sigma_map
149160

0 commit comments

Comments
 (0)