Injective functions #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
- to_fun : α → β
- inj' : function.injective self.to_fun
α ↪ β is a bundled injective function.
Instances for function.embedding
- function.embedding.has_sizeof_inst
- function.embedding.has_coe_to_fun
- function.embedding_like
- function.injective.can_lift
- equiv.coe_embedding
- equiv.perm.coe_embedding
- function.embedding.is_empty
- function.embedding.fintype
- function.embedding.finite
- function.embedding.has_smul
- function.embedding.has_vadd
- function.embedding.is_scalar_tower
- function.embedding.smul_comm_class
- function.embedding.vadd_comm_class
- function.embedding.is_central_scalar
- function.embedding.mul_action
- function.embedding.add_action
Equations
Equations
- function.embedding_like = {coe := function.embedding.to_fun β, coe_injective' := _, injective' := _}
Equations
- function.injective.can_lift = {prf := _}
Convert an α ≃ β to α ↪ β.
This is also available as a coercion equiv.coe_embedding.
The explicit equiv.to_embedding version is preferred though, since the coercion can have issues
inferring the type of the resulting embedding. For example:
-- Works:
example (s : finset (fin 3)) (f : equiv.perm (fin 3)) : s.map f.to_embedding = s.map f := by simp
-- Error, `f` has type `fin 3 ≃ fin 3` but is expected to have type `fin 3 ↪ ?m_1 : Type ?`
example (s : finset (fin 3)) (f : equiv.perm (fin 3)) : s.map f = s.map f.to_embedding := by simp
Equations
Equations
The identity map as a function.embedding.
Transfer an embedding along a pair of equivalences.
Equations
- function.embedding.congr e₁ e₂ f = e₁.symm.to_embedding.trans (f.trans e₂.to_embedding)
A right inverse surj_inv of a surjective function as an embedding.
Equations
- function.embedding.of_surjective f hf = {to_fun := function.surj_inv hf, inj' := _}
Convert a surjective embedding to an equiv
Equations
- f.equiv_of_surjective hf = equiv.of_bijective ⇑f _
There is always an embedding from an empty type.
Equations
- function.embedding.of_is_empty = {to_fun := is_empty_elim (λ (a : α), β), inj' := _}
Change the value of an embedding f at one point. If the prescribed image
is already occupied by some f a', then swap the values at these two points.
Embedding into option α using some.
Equations
- function.embedding.some = {to_fun := option.some α, inj' := _}
Embedding into option α using coe. Usually the correct synctatical form for simp.
Equations
- function.embedding.coe_option = {to_fun := coe coe_to_lift, inj' := _}
A version of option.map for function.embeddings.
Equations
- f.option_map = {to_fun := option.map ⇑f, inj' := _}
Embedding of a subtype.
Equations
- function.embedding.subtype p = {to_fun := coe coe_to_lift, inj' := _}
quotient.out as an embedding.
Equations
- function.embedding.quotient_out α = {to_fun := quotient.out s, inj' := _}
An embedding e : α ↪ β defines an embedding (γ → α) ↪ (γ → β) that sends each f
to e ∘ f.
Equations
- e.arrow_congr_right = function.embedding.Pi_congr_right (λ (_x : γ), e)
An embedding e : α ↪ β defines an embedding (α → γ) ↪ (β → γ) for any inhabited type γ.
This embedding sends each f : α → γ to a function g : β → γ such that g ∘ e = f and
g y = default whenever y ∉ range e.
Equations
- e.arrow_congr_left = {to_fun := λ (f : α → γ), function.extend ⇑e f inhabited.default, inj' := _}
The type of embeddings α ↪ β is equivalent to
the subtype of all injective functions α → β.
If α₁ ≃ α₂ and β₁ ≃ β₂, then the type of embeddings α₁ ↪ β₁
is equivalent to the type of embeddings α₂ ↪ β₂.
Equations
- h.embedding_congr h' = {to_fun := λ (f : α ↪ γ), function.embedding.congr h h' f, inv_fun := λ (f : β ↪ δ), function.embedding.congr h.symm h'.symm f, left_inv := _, right_inv := _}
A subtype {x // p x ∨ q x} over a disjunction of p q : α → Prop can be injectively split
into a sum of subtypes {x // p x} ⊕ {x // q x} such that ¬ p x is sent to the right.