@[inline]
Transforms a function from pairs into an equivalent two-parameter function.
Examples:
Function.curry (fun (x, y) => x + y) 3 5 = 8
Function.curry Prod.swap 3 "five" = ("five", 3)
Equations
- Function.curry f a b = f (a, b)
Instances For
@[inline]
Transforms a two-parameter function into an equivalent function from pairs.
Examples:
Function.uncurry List.drop (1, ["a", "b", "c"]) = ["b", "c"]
[("orange", 2), ("android", 3) ].map (Function.uncurry String.take) = ["or", "and"]
Equations
- Function.uncurry f a = f a.fst a.snd
Instances For
A function f : α → β
is called injective if f x = f y
implies x = y
.
Equations
- Function.Injective f = ∀ ⦃a₁ a₂ : α⦄, f a₁ = f a₂ → a₁ = a₂
Instances For
theorem
Function.Surjective.comp
{α : Sort u_1}
{β : Sort u_2}
{γ : Sort u_3}
{g : β → γ}
{f : α → β}
(hg : Surjective g)
(hf : Surjective f)
:
Surjective (g ∘ f)
LeftInverse g f
means that g
is a left inverse to f
. That is, g ∘ f = id
.
Equations
- Function.LeftInverse g f = ∀ (x : α), g (f x) = x
Instances For
HasLeftInverse f
means that f
has an unspecified left inverse.
Equations
- Function.HasLeftInverse f = ∃ (finv : β → α), Function.LeftInverse finv f
Instances For
HasRightInverse f
means that f
has an unspecified right inverse.
Equations
- Function.HasRightInverse f = ∃ (finv : β → α), Function.RightInverse finv f
Instances For
theorem
Function.LeftInverse.injective
{α : Sort u_1}
{β : Sort u_2}
{g : β → α}
{f : α → β}
:
LeftInverse g f → Injective f
theorem
Function.HasLeftInverse.injective
{α : Sort u_1}
{β : Sort u_2}
{f : α → β}
:
HasLeftInverse f → Injective f
theorem
Function.rightInverse_of_injective_of_leftInverse
{α : Sort u_1}
{β : Sort u_2}
{f : α → β}
{g : β → α}
(injf : Injective f)
(lfg : LeftInverse f g)
:
RightInverse f g
theorem
Function.RightInverse.surjective
{α : Sort u_1}
{β : Sort u_2}
{f : α → β}
{g : β → α}
(h : RightInverse g f)
:
theorem
Function.HasRightInverse.surjective
{α : Sort u_1}
{β : Sort u_2}
{f : α → β}
:
HasRightInverse f → Surjective f
theorem
Function.leftInverse_of_surjective_of_rightInverse
{α : Sort u_1}
{β : Sort u_2}
{f : α → β}
{g : β → α}
(surjf : Surjective f)
(rfg : RightInverse f g)
:
LeftInverse f g
theorem
Function.LeftInverse.id
{α : Sort u_1}
{β : Sort u_2}
{g : β → α}
{f : α → β}
(h : LeftInverse g f)
:
theorem
Function.RightInverse.id
{α : Sort u_1}
{β : Sort u_2}
{g : β → α}
{f : α → β}
(h : RightInverse g f)
: