@@ -127,7 +127,7 @@ theorem Injective.of_comp {g : γ → α} (I : Injective (f ∘ g)) : Injective
127
127
#align function.injective.of_comp Function.Injective.of_comp
128
128
129
129
@[simp]
130
- theorem Injective.of_comp_iff {f : α → β} (hf : Injective f) (g : γ → α) :
130
+ theorem Injective.of_comp_iff (hf : Injective f) (g : γ → α) :
131
131
Injective (f ∘ g) ↔ Injective g :=
132
132
⟨Injective.of_comp, hf.comp⟩
133
133
#align function.injective.of_comp_iff Function.Injective.of_comp_iff
@@ -138,6 +138,10 @@ theorem Injective.of_comp_right {g : γ → α} (I : Injective (f ∘ g)) (hg :
138
138
obtain ⟨y, rfl⟩ := hg y
139
139
exact congr_arg g (I h)
140
140
141
+ theorem Surjective.bijective₂_of_injective {g : γ → α} (hf : Surjective f) (hg : Surjective g)
142
+ (I : Injective (f ∘ g)) : Bijective f ∧ Bijective g :=
143
+ ⟨⟨I.of_comp_right hg, hf⟩, I.of_comp, hg⟩
144
+
141
145
@[simp]
142
146
theorem Injective.of_comp_iff' (f : α → β) {g : γ → α} (hg : Bijective g) :
143
147
Injective (f ∘ g) ↔ Injective f :=
@@ -182,6 +186,10 @@ theorem Surjective.of_comp_iff (f : α → β) {g : γ → α} (hg : Surjective
182
186
theorem Surjective.of_comp_left {g : γ → α} (S : Surjective (f ∘ g)) (hf : Injective f) :
183
187
Surjective g := fun a ↦ let ⟨c, hc⟩ := S (f a); ⟨c, hf hc⟩
184
188
189
+ theorem Injective.bijective₂_of_surjective {g : γ → α} (hf : Injective f) (hg : Injective g)
190
+ (S : Surjective (f ∘ g)) : Bijective f ∧ Bijective g :=
191
+ ⟨⟨hf, S.of_comp⟩, hg, S.of_comp_left hf⟩
192
+
185
193
@[simp]
186
194
theorem Surjective.of_comp_iff' (hf : Bijective f) (g : γ → α) :
187
195
Surjective (f ∘ g) ↔ Surjective g :=
0 commit comments