Injective functions #
α ↪ β is a bundled injective function.
- toFun : α → βAn embedding as a function. Use coercion instead. 
- An embedding is an injective function. Use - Function.Embedding.injectiveinstead.
Instances For
An embedding, a.k.a. a bundled injective function.
Equations
- Function.«term_↪_» = Lean.ParserDescr.trailingNode `Function.«term_↪_» 25 26 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ↪ ") (Lean.ParserDescr.cat `term 25))
Instances For
Equations
- Function.instFunLikeEmbedding = { coe := Function.Embedding.toFun, coe_injective' := ⋯ }
Convert an α ≃ β to α ↪ β.
This is also available as a coercion Equiv.coeEmbedding.
The explicit Equiv.toEmbedding 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.toEmbedding = 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.toEmbedding := by simp
Equations
- f.toEmbedding = { toFun := ⇑f, inj' := ⋯ }
Instances For
Equations
- Equiv.coeEmbedding = { coe := Equiv.toEmbedding }
Equations
Equations
- Function.Embedding.instUniqueOfIsEmpty = { default := { toFun := fun (a : α) => isEmptyElim a, inj' := ⋯ }, uniq := ⋯ }
The identity map as a Function.Embedding.
Equations
- Function.Embedding.refl α = { toFun := id, inj' := ⋯ }
Instances For
Equations
- Function.Embedding.instTrans = { trans := fun {a : Sort ?u.12} {b : Sort ?u.13} {c : Sort ?u.11} => Function.Embedding.trans }
Transfer an embedding along a pair of equivalences.
Equations
- Function.Embedding.congr e₁ e₂ f = e₁.symm.toEmbedding.trans (f.trans e₂.toEmbedding)
Instances For
A right inverse surjInv of a surjective function as an Embedding.
Equations
- Function.Embedding.ofSurjective f hf = { toFun := Function.surjInv hf, inj' := ⋯ }
Instances For
Convert a surjective Embedding to an Equiv
Equations
- f.equivOfSurjective hf = Equiv.ofBijective ⇑f ⋯
Instances For
There is always an embedding from an empty type.
Equations
- Function.Embedding.ofIsEmpty = { toFun := fun (a : α) => isEmptyElim a, inj' := ⋯ }
Instances For
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.
Equations
Instances For
A version of Option.map for Function.Embeddings.
Equations
- f.optionMap = { toFun := Option.map ⇑f, inj' := ⋯ }
Instances For
Embedding of a Subtype.
Equations
- Function.Embedding.subtype p = { toFun := Subtype.val, inj' := ⋯ }
Instances For
Quotient.out as an embedding.
Equations
- Function.Embedding.quotientOut α = { toFun := Quotient.out, inj' := ⋯ }
Instances For
Choosing an element b : β gives an embedding of PUnit into β.
Equations
- Function.Embedding.punit b = { toFun := fun (x : PUnit.{?u.11}) => b, inj' := ⋯ }
Instances For
The equivalence one ↪ α with α, for Unique one.
Equations
Instances For
Alias of Function.Embedding.sectL.
Fixing an element b : β gives an embedding α ↪ α × β.
Instances For
Alias of Function.Embedding.sectR.
Fixing an element a : α gives an embedding β ↪ α × β.
Equations
Instances For
Alias of Function.Embedding.sectL_apply.
Alias of Function.Embedding.sectR_apply.
Sigma.mk as a Function.Embedding.
Equations
- Function.Embedding.sigmaMk a = { toFun := Sigma.mk a, inj' := ⋯ }
Instances For
Define an embedding (Π a : α, β a) ↪ (Π a : α, γ a) from a family of embeddings
e : Π a, (β a ↪ γ a). This embedding sends f to fun a ↦ e a (f a).
Equations
- Function.Embedding.piCongrRight e = { toFun := fun (f : (a : α) → β a) (a : α) => (e a) (f a), inj' := ⋯ }
Instances For
An embedding e : α ↪ β defines an embedding (γ → α) ↪ (γ → β) that sends each f
to e ∘ f.
Equations
- e.arrowCongrRight = Function.Embedding.piCongrRight fun (x : γ) => e
Instances For
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.arrowCongrLeft = { toFun := fun (f : α → γ) => Function.extend (⇑e) f default, inj' := ⋯ }
Instances For
Restrict both domain and codomain of an embedding.
Equations
- f.subtypeMap h = { toFun := Subtype.map (⇑f) h, inj' := ⋯ }
Instances For
Given an equivalence to a subtype, produce an embedding to the elements of the corresponding set.
Equations
Instances For
If α₁ ≃ α₂ and β₁ ≃ β₂, then the type of embeddings α₁ ↪ β₁
is equivalent to the type of embeddings α₂ ↪ β₂.
Equations
- h.embeddingCongr h' = { toFun := fun (f : α ↪ γ) => Function.Embedding.congr h h' f, invFun := fun (f : β ↪ δ) => Function.Embedding.congr h.symm h'.symm f, left_inv := ⋯, right_inv := ⋯ }
Instances For
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.
Equations
Instances For
A subtype {x // p x} can be injectively sent to into a subtype {x // q x},
if p x → q x for all x : α.