- to_fun : α → β
- inv_fun : β → α
- left_inv : function.left_inverse c.inv_fun c.to_fun
- right_inv : function.right_inverse c.inv_fun c.to_fun
α ≃ β
is the type of functions from α → β
with a two-sided inverse.
Convert an involutive function f
to an equivalence with to_fun = inv_fun = f
.
perm α
is the type of bijections from α
to itself.
Equations
- equiv.perm α = (α ≃ α)
Equations
- equiv.has_coe_to_fun = {F := λ (x : α ≃ β), α → β, coe := equiv.to_fun β}
The map coe_fn : (r ≃ s) → (r → s)
is injective. We can't use function.injective
here but mimic its signature by using ⦃e₁ e₂⦄
.
Transfer decidable_eq
across an equivalence.
Equations
If α
is equivalent to β
and γ
is equivalent to δ
, then the type of equivalences α ≃ γ
is equivalent to the type of equivalences β ≃ δ
.
If α
is equivalent to β
, then perm α
is equivalent to perm β
.
Equations
- e.perm_congr = e.equiv_congr e
Equations
- equiv.perm.perm_group = {mul := λ (f g : equiv.perm α), equiv.trans g f, mul_assoc := _, one := equiv.refl α, one_mul := _, mul_one := _, inv := equiv.symm α, mul_left_inv := _}
Equations
If α
is an empty type, then it is equivalent to the pempty
type in any universe.
empty
is equivalent to pempty
.
Equations
- equiv.empty_equiv_pempty = equiv.equiv_pempty equiv.empty_equiv_pempty._proof_1
pempty
types from any two universes are equivalent.
Equations
ulift α
is equivalent to α
.
Equations
- equiv.ulift = {to_fun := ulift.down α, inv_fun := ulift.up α, left_inv := _, right_inv := _}
plift α
is equivalent to α
.
Equations
- equiv.plift = {to_fun := plift.down α, inv_fun := plift.up α, left_inv := _, right_inv := _}
If α₁
is equivalent to α₂
and β₁
is equivalent to β₂
, then the type of maps α₁ → β₁
is equivalent to the type of maps α₂ → β₂
.
A version of equiv.arrow_congr
in Type
, rather than Sort
.
The equiv_rw
tactic is not able to use the default Sort
level equiv.arrow_congr
,
because Lean's universe rules will not unify ?l_1
with imax (1 ?m_1)
.
Equations
- hα.arrow_congr' hβ = hα.arrow_congr hβ
Conjugate a map f : α → α
by an equivalence α ≃ β
.
Equations
- e.conj = e.arrow_congr e
punit
sorts in any two universes are equivalent.
Equations
- equiv.punit_equiv_punit = {to_fun := λ (_x : punit), punit.star, inv_fun := λ (_x : punit), punit.star, left_inv := equiv.punit_equiv_punit._proof_1, right_inv := equiv.punit_equiv_punit._proof_2}
The sort of maps to punit.{v}
is equivalent to punit.{w}
.
Equations
- equiv.arrow_punit_equiv_punit α = {to_fun := λ (f : α → punit), punit.star, inv_fun := λ (u : punit) (f : α), punit.star, left_inv := _, right_inv := _}
Product of two equivalences. If α₁ ≃ α₂
and β₁ ≃ β₂
, then α₁ × β₁ ≃ α₂ × β₂
.
Type product is associative up to an equivalence.
punit
is a left identity for type product up to an equivalence.
Equations
- equiv.punit_prod α = (equiv.prod_comm punit α).trans (equiv.prod_punit α)
empty
type is a right absorbing element for type product up to an equivalence.
Equations
empty
type is a left absorbing element for type product up to an equivalence.
Equations
pempty
type is a right absorbing element for type product up to an equivalence.
Equations
pempty
type is a left absorbing element for type product up to an equivalence.
Equations
If α ≃ α'
and β ≃ β'
, then α ⊕ β ≃ α' ⊕ β'
.
bool
is equivalent the sum of two punit
s.
Equations
- equiv.bool_equiv_punit_sum_punit = {to_fun := λ (b : bool), cond b (sum.inr punit.star) (sum.inl punit.star), inv_fun := λ (s : punit ⊕ punit), s.rec_on (λ (_x : punit), ff) (λ (_x : punit), tt), left_inv := equiv.bool_equiv_punit_sum_punit._proof_1, right_inv := equiv.bool_equiv_punit_sum_punit._proof_2}
Sum of types is associative up to an equivalence.
The sum of empty
with any Sort*
is equivalent to the right summand.
Equations
- equiv.empty_sum α = (equiv.sum_comm empty α).trans (equiv.sum_empty α)
The sum of pempty
with any Sort*
is equivalent to the right summand.
Equations
- equiv.pempty_sum α = (equiv.sum_comm pempty α).trans (equiv.sum_pempty α)
option α
is equivalent to α ⊕ punit
Equations
- equiv.option_equiv_sum_punit α = {to_fun := λ (o : option α), equiv.option_equiv_sum_punit._match_1 α o, inv_fun := λ (s : α ⊕ punit), equiv.option_equiv_sum_punit._match_2 α s, left_inv := _, right_inv := _}
- equiv.option_equiv_sum_punit._match_1 α (some a) = sum.inl a
- equiv.option_equiv_sum_punit._match_1 α none = sum.inr punit.star
- equiv.option_equiv_sum_punit._match_2 α (sum.inr _x) = none
- equiv.option_equiv_sum_punit._match_2 α (sum.inl a) = some a
The set of x : option α
such that is_some x
is equivalent to α
.
α ⊕ β
is equivalent to a sigma
-type over bool
.
Equations
- equiv.sum_equiv_sigma_bool α β = {to_fun := λ (s : α ⊕ β), equiv.sum_equiv_sigma_bool._match_1 α β s, inv_fun := λ (s : Σ (b : bool), cond b α β), equiv.sum_equiv_sigma_bool._match_2 α β s, left_inv := _, right_inv := _}
- equiv.sum_equiv_sigma_bool._match_1 α β (sum.inr b) = ⟨ff, b⟩
- equiv.sum_equiv_sigma_bool._match_1 α β (sum.inl a) = ⟨tt, a⟩
- equiv.sum_equiv_sigma_bool._match_2 α β ⟨tt, a⟩ = sum.inl a
- equiv.sum_equiv_sigma_bool._match_2 α β ⟨ff, b⟩ = sum.inr b
sigma_preimage_equiv f
for f : α → β
is the natural equivalence between
the type of all fibres of f
and the total space α
.
For any predicate p
on α
,
the sum of the two subtypes {a // p a}
and its complement {a // ¬ p a}
is naturally equivalent to α
.
For a fixed function x₀ : {a // p a} → β
defined on a subtype of α
,
the subtype of functions x : α → β
that agree with x₀
on the subtype {a // p a}
is naturally equivalent to the type of functions {a // ¬ p a} → β
.
A family of equivalences Π a, β₁ a ≃ β₂ a
generates an equivalence between Π a, β₁ a
and
Π a, β₂ a
.
Dependent curry
equivalence: the type of dependent functions on Σ i, β i
is equivalent
to the type of dependent functions of two arguments (i.e., functions to the space of functions).
A family of equivalences Π a, β₁ a ≃ β₂ a
generates an equivalence between Σ a, β₁ a
and
Σ a, β₂ a
.
An equivalence f : α₁ ≃ α₂
generates an equivalence between Σ a, β (f a)
and Σ a, β a
.
Transporting a sigma type through an equivalence of the base
Equations
Transporting a sigma type through an equivalence of the base and a family of equivalences of matching fibers
Equations
- f.sigma_congr F = (equiv.sigma_congr_right F).trans f.sigma_congr_left
sigma
type with a constant fiber is equivalent to the product.
If each fiber of a sigma
type is equivalent to a fixed type, then the sigma type
is equivalent to the product.
Equations
The type of functions to a product α × β
is equivalent to the type of pairs of functions
γ → α
and γ → β
.
Functions α → β → γ
are equivalent to functions on α × β
.
Equations
- equiv.arrow_arrow_equiv_prod_arrow α β γ = {to_fun := function.uncurry γ, inv_fun := function.curry γ, left_inv := _, right_inv := _}
The type of functions on a sum type α ⊕ β
is equivalent to the type of pairs of functions
on α
and on β
.
Type product is right distributive with respect to type sum up to an equivalence.
Equations
- equiv.sum_prod_distrib α β γ = {to_fun := λ (p : (α ⊕ β) × γ), equiv.sum_prod_distrib._match_1 α β γ p, inv_fun := λ (s : α × γ ⊕ β × γ), equiv.sum_prod_distrib._match_2 α β γ s, left_inv := _, right_inv := _}
- equiv.sum_prod_distrib._match_1 α β γ (sum.inr b, c) = sum.inr (b, c)
- equiv.sum_prod_distrib._match_1 α β γ (sum.inl a, c) = sum.inl (a, c)
- equiv.sum_prod_distrib._match_2 α β γ (sum.inr q) = (sum.inr q.fst, q.snd)
- equiv.sum_prod_distrib._match_2 α β γ (sum.inl q) = (sum.inl q.fst, q.snd)
Type product is left distributive with respect to type sum up to an equivalence.
Equations
- equiv.prod_sum_distrib α β γ = ((equiv.prod_comm α (β ⊕ γ)).trans (equiv.sum_prod_distrib β γ α)).trans ((equiv.prod_comm β α).sum_congr (equiv.prod_comm γ α))
The product of an indexed sum of types (formally, a sigma
-type Σ i, α i
) by a type β
is
equivalent to the sum of products Σ i, (α i × β)
.
The product bool × α
is equivalent to α ⊕ α
.
Equations
The set of natural numbers is equivalent to ℕ ⊕ punit
.
Equations
- equiv.nat_equiv_nat_sum_punit = {to_fun := λ (n : ℕ), equiv.nat_equiv_nat_sum_punit._match_1 n, inv_fun := λ (s : ℕ ⊕ punit), equiv.nat_equiv_nat_sum_punit._match_2 s, left_inv := equiv.nat_equiv_nat_sum_punit._proof_1, right_inv := equiv.nat_equiv_nat_sum_punit._proof_2}
- equiv.nat_equiv_nat_sum_punit._match_1 a.succ = sum.inl a
- equiv.nat_equiv_nat_sum_punit._match_1 0 = sum.inr punit.star
- equiv.nat_equiv_nat_sum_punit._match_2 (sum.inr punit.star) = 0
- equiv.nat_equiv_nat_sum_punit._match_2 (sum.inl n) = n.succ
ℕ ⊕ punit
is equivalent to ℕ
.
The type of integer numbers is equivalent to ℕ ⊕ ℕ
.
Equations
If α
is equivalent to β
and the predicates p : α → Prop
and q : β → Prop
are equivalent
at corresponding points, then {a // p a}
is equivalent to {b // q b}
.
If two predicates p
and q
are pointwise equivalent, then {x // p x}
is equivalent to
{x // q x}
.
Equations
If α ≃ β
, then for any predicate p : β → Prop
the subtype {a // p (e a)}
is equivalent
to the subtype {b // p b}
.
Equations
If α ≃ β
, then for any predicate p : α → Prop
the subtype {a // p a}
is equivalent
to the subtype {b // p (e.symm b)}
. This version is used by equiv_rw
.
Equations
If two predicates are equal, then the corresponding subtypes are equivalent.
Equations
The subtypes corresponding to equal sets are equivalent.
Equations
A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. This
version allows the “inner” predicate to depend on h : p a
.
Equations
- equiv.subtype_subtype_equiv_subtype_exists p q = {to_fun := λ (_x : subtype q), equiv.subtype_subtype_equiv_subtype_exists._match_1 p q _x, inv_fun := λ (_x : {a // ∃ (h : p a), q ⟨a, h⟩}), equiv.subtype_subtype_equiv_subtype_exists._match_2 p q _x, left_inv := _, right_inv := _}
- equiv.subtype_subtype_equiv_subtype_exists._match_1 p q ⟨⟨a, ha⟩, ha'⟩ = ⟨a, _⟩
- equiv.subtype_subtype_equiv_subtype_exists._match_2 p q ⟨a, ha⟩ = ⟨⟨a, _⟩, _⟩
A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates.
Equations
- equiv.subtype_subtype_equiv_subtype_inter p q = (equiv.subtype_subtype_equiv_subtype_exists p (λ (x : subtype p), q x.val)).trans (equiv.subtype_congr_right _)
If the outer subtype has more restrictive predicate than the inner one, then we can drop the latter.
If a proposition holds for all elements, then the subtype is equivalent to the original type.
A subtype of a sigma-type is a sigma-type over a subtype.
A sigma type over a subtype is equivalent to the sigma set over the original type, if the fiber is empty outside of the subset
Equations
If a predicate p : β → Prop
is true on the range of a map f : α → β
, then
Σ y : {y // p y}, {x // f x = y}
is equivalent to α
.
Equations
- equiv.sigma_subtype_preimage_equiv f p h = (equiv.sigma_subtype_equiv_of_subset (λ (y : β), {x // f x = y}) p _).trans (equiv.sigma_preimage_equiv f)
If for each x
we have p x ↔ q (f x)
, then Σ y : {y // q y}, f ⁻¹' {y}
is equivalent
to {x // p x}
.
Equations
- equiv.sigma_subtype_preimage_equiv_subtype f h = (equiv.sigma_congr_right (λ (y : subtype q), ((equiv.subtype_subtype_equiv_subtype_exists p (λ (x : subtype p), ⟨f ↑x, _⟩ = y)).trans (equiv.subtype_congr_right _)).symm)).trans (equiv.sigma_preimage_equiv (λ (x : subtype p), ⟨f ↑x, _⟩))
The pi
-type Π i, π i
is equivalent to the type of sections f : ι → Σ i, π i
of the
sigma
type such that for all i
we have (f i).fst = i
.
The set of functions f : Π a, β a
such that for all a
we have p a (f a)
is equivalent
to the set of functions Π a, {b : β a // p a b}
.
A subtype of a product defined by componentwise conditions is equivalent to a product of subtypes.
The type of all functions X → Y
with prescribed values for all x' ≠ x
is equivalent to the codomain Y
.
Equations
- equiv.subtype_equiv_codomain f = (equiv.subtype_preimage (λ (x' : X), x' ≠ x) f).trans (equiv.fun_unique {x' // ¬x' ≠ x} Y)
An empty set is equivalent to the empty
type.
Equations
An empty set is equivalent to a pempty
type.
Equations
If sets s
and t
are separated by a decidable predicate, then s ∪ t
is equivalent to
s ⊕ t
.
Equations
- equiv.set.union' p hs ht = {to_fun := λ (x : ↥(s ∪ t)), dite (p ↑x) (λ (hp : p ↑x), sum.inl ⟨x.val, _⟩) (λ (hp : ¬p ↑x), sum.inr ⟨x.val, _⟩), inv_fun := λ (o : ↥s ⊕ ↥t), equiv.set.union'._match_1 o, left_inv := _, right_inv := _}
- equiv.set.union'._match_1 (sum.inr x) = ⟨↑x, _⟩
- equiv.set.union'._match_1 (sum.inl x) = ⟨↑x, _⟩
If sets s
and t
are disjoint, then s ∪ t
is equivalent to s ⊕ t
.
Equations
- equiv.set.union H = equiv.set.union' (λ (x : α), x ∈ s) equiv.set.union._proof_1 _
If a ∉ s
, then insert a s
is equivalent to s ⊕ punit
.
Equations
- equiv.set.insert H = ((equiv.set.of_eq equiv.set.insert._proof_1).trans (equiv.set.union _)).trans ((equiv.refl ↥s).sum_congr (equiv.set.singleton a))
If s : set α
is a set with decidable membership, then s ⊕ sᶜ
is equivalent to α
.
Equations
- equiv.set.sum_compl s = ((equiv.set.union _).symm.trans (equiv.set.of_eq _)).trans (equiv.set.univ α)
sum_diff_subset s t
is the natural equivalence between
s ⊕ (t \ s)
and t
, where s
and t
are two sets.
Equations
- equiv.set.sum_diff_subset h = (equiv.set.union equiv.set.sum_diff_subset._proof_1).symm.trans (equiv.set.of_eq _)
If s
is a set with decidable membership, then the sum of s ∪ t
and s ∩ t
is equivalent
to s ⊕ t
.
Equations
- equiv.set.union_sum_inter s t = ((((_.mpr (equiv.refl (↥(s ∪ t) ⊕ ↥(s ∩ t)))).trans ((equiv.set.union _).sum_congr (equiv.refl ↥(s ∩ t)))).trans (equiv.sum_assoc ↥s ↥(t \ s) ↥(s ∩ t))).trans ((equiv.refl ↥s).sum_congr (equiv.set.union' (λ (_x : α), _x ∉ s) _ _).symm)).trans (_.mpr (equiv.refl (↥s ⊕ ↥t)))
The set product of two sets is equivalent to the type product of their coercions to types.
Equations
If a function f
is injective on a set s
, then s
is equivalent to f '' s
.
If f
is an injective function, then s
is equivalent to f '' s
.
Equations
- equiv.set.image f s H = equiv.set.image_of_inj_on f s _
If f : α → β
is an injective function, then α
is equivalent to the range of f
.
The set {x ∈ s | t x}
is equivalent to the set of x : s
such that t x
.
Equations
If f
is a bijective function, then its domain is equivalent to its codomain.
Equations
- equiv.of_bijective f hf = (equiv.set.range f _).trans ((equiv.set_congr _).trans (equiv.set.univ β))
If f
is an injective function, then its domain is equivalent to its range.
Equations
- equiv.of_injective f hf = equiv.of_bijective (λ (x : α), ⟨f x, _⟩) _
A helper function for equiv.swap
.
swap a b
is the permutation that swaps a
and b
and
leaves other values as is.
Equations
- equiv.swap a b = {to_fun := equiv.swap_core a b, inv_fun := equiv.swap_core a b, left_inv := _, right_inv := _}
Augment an equivalence with a prescribed mapping f a = b
Equations
- f.set_value a b = equiv.trans (equiv.swap a (⇑(f.symm) b)) f
Transport dependent functions through an equivalence of the base space.
Transporting dependent functions through an equivalence of the base, expressed as a "simplification".
Equations
- equiv.Pi_congr_left P e = (equiv.Pi_congr_left' P e.symm).symm
Transport dependent functions through an equivalence of the base spaces and a family of equivalences of the matching fibers.
Equations
- h₁.Pi_congr h₂ = (equiv.Pi_congr_right h₂).trans (equiv.Pi_congr_left Z h₁)
Equations
Equations
If α
is a singleton, then it is equivalent to any punit
.
Equations
If α
is a subsingleton, then it is equivalent to α × α
.
To give an equivalence between two subsingleton types, it is sufficient to give any two functions between them.
Quotients are congruent on equivalences under equality of their relation.
An alternative is just to use rewriting with eq
, but then computational proofs get stuck.
Equations
- quot.congr_right eq = quot.congr (equiv.refl α) eq
An equivalence e : α ≃ β
generates an equivalence between the quotient space of α
by a relation ra
and the quotient space of β
by the image of this relation under e
.
Equations
- quot.congr_left e = quot.congr e _
An equivalence e : α ≃ β
generates an equivalence between quotient spaces,
if `ra a₁ a₂ ↔ rb (e a₁) (e a₂).
Equations
- quotient.congr e eq = quot.congr e eq
If a function is a bijection between univ
and a set s
in the target type, it induces an
equivalence between the original type and the type ↑s
.
Equations
- set.bij_on.equiv f h = equiv.of_bijective (λ (x : α), ⟨f x, _⟩) _
The composition of an updated function with an equiv on a subset can be expressed as an updated function.