Equivalence between product types #
In this file we continue the work on equivalences begun in Mathlib/Logic/Equiv/Defs.lean
,
focussing on product types.
Main definitions #
Equiv.prodCongr ea eb : α₁ × β₁ ≃ α₂ × β₂
: combine two equivalencesea : α₁ ≃ α₂
andeb : β₁ ≃ β₂
usingProd.map
.
Tags #
equivalence, congruence, bijective map
Combine two equivalences using PProd
in the domain and Prod
in the codomain.
Equations
- ea.pprodProd eb = (ea.pprodCongr eb).trans Equiv.pprodEquivProd
Instances For
Type product is associative up to an equivalence.
Equations
Instances For
γ
-valued functions on α × β
are equivalent to functions α → β → γ
.
Equations
- Equiv.curry α β γ = { toFun := Function.curry, invFun := Function.uncurry, left_inv := ⋯, right_inv := ⋯ }
Instances For
PUnit
is a right identity for type product up to an equivalence.
Equations
- Equiv.prodPUnit α = { toFun := fun (p : α × PUnit.{?u.14 + 1}) => p.fst, invFun := fun (a : α) => (a, PUnit.unit), left_inv := ⋯, right_inv := ⋯ }
Instances For
PUnit
is a right identity for dependent type product up to an equivalence.
Equations
- Equiv.sigmaPUnit α = { toFun := fun (p : (_ : α) × PUnit.{?u.16 + 1}) => p.fst, invFun := fun (a : α) => ⟨a, PUnit.unit⟩, left_inv := ⋯, right_inv := ⋯ }
Instances For
Any Unique
type is a right identity for type product up to equivalence.
Equations
- Equiv.prodUnique α β = ((Equiv.refl α).prodCongr (Equiv.equivPUnit β)).trans (Equiv.prodPUnit α)
Instances For
Any Unique
type is a left identity for type product up to equivalence.
Equations
- Equiv.uniqueProd α β = ((Equiv.equivPUnit β).prodCongr (Equiv.refl α)).trans (Equiv.punitProd α)
Instances For
Any family of Unique
types is a right identity for dependent type product up to
equivalence.
Equations
- Equiv.sigmaUnique α β = (Equiv.sigmaCongrRight fun (a : α) => Equiv.equivPUnit (β a)).trans (Equiv.sigmaPUnit α)
Instances For
Any Unique
type is a left identity for type sigma up to equivalence. Compare with uniqueProd
which is non-dependent.
Equations
Instances For
Empty
type is a right absorbing element for type product up to an equivalence.
Equations
- Equiv.prodEmpty α = Equiv.equivEmpty (α × Empty)
Instances For
Empty
type is a left absorbing element for type product up to an equivalence.
Equations
- Equiv.emptyProd α = Equiv.equivEmpty (Empty × α)
Instances For
PEmpty
type is a right absorbing element for type product up to an equivalence.
Equations
Instances For
PEmpty
type is a left absorbing element for type product up to an equivalence.
Equations
Instances For
A family of equivalences ∀ (a : α₁), β₁ ≃ β₂
generates an equivalence
between β₁ × α₁
and β₂ × α₁
.
Equations
Instances For
A family of equivalences ∀ (a : α₁), β₁ ≃ β₂
generates an equivalence
between α₁ × β₁
and α₁ × β₂
.
Equations
Instances For
A variation on Equiv.prodCongr
where the equivalence in the second component can depend
on the first component. A typical example is a shear mapping, explaining the name of this
declaration.
Equations
Instances For
prodExtendRight a e
extends e : Perm β
to Perm (α × β)
by sending (a, b)
to
(a, e b)
and keeping the other (a', b)
fixed.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The type of functions to a product β × γ
is equivalent to the type of pairs of functions
α → β
and β → γ
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The type of dependent functions on a sum type ι ⊕ ι'
is equivalent to the type of pairs of
functions on ι
and on ι'
. This is a dependent version of Equiv.sumArrowEquivProdArrow
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The equivalence between a product of two dependent functions types and a single dependent
function type. Basically a symmetric version of Equiv.sumPiEquivProdPi
.
Equations
- Equiv.prodPiEquivSumPi π π' = (Equiv.sumPiEquivProdPi (Sum.elim π π')).symm
Instances For
The type of functions on a sum type α ⊕ β
is equivalent to the type of pairs of functions
on α
and on β
.
Equations
Instances For
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 × β)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A subtype of a product defined by componentwise conditions is equivalent to a product of subtypes.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A subtype of a Prod
that depends only on the first component is equivalent to the
corresponding subtype of the first type times the second type.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A subtype of a Prod
is equivalent to a sigma type whose fibers are subtypes.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The type ∀ (i : α), β i
can be split as a product by separating the indices in α
depending on whether they satisfy a predicate p
or not.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A product of types can be split as the binary product of one of the types and the product of all the remaining types.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A product of copies of a type can be split as the binary product of one copy and the product of all the remaining copies.
Equations
- Equiv.funSplitAt i β = Equiv.piSplitAt i fun (a : α) => β
Instances For
If α
is a subsingleton, then it is equivalent to α × α
.