FinEnum instance for Option #
Provides a recursor for FinEnum types like Fintype.truncRecEmptyOption
, but capable of producing
non-truncated data.
TODO #
- recreate rest of
Mathlib.Data.Fintype.Option
Inserting an Option.none
anywhere in an enumeration yields another enumeration.
Equations
- FinEnum.insertNone α i = FinEnum.mk (FinEnum.card α + 1) (FinEnum.equiv.optionCongr.trans (finSuccEquiv' i).symm)
Instances For
This is an arbitrary choice of insertion rank for a default instance.
It keeps the mapping of the existing α
-inhabitants intact, modulo Fin.castSucc
.
Equations
A recursor principle for finite-and-enumerable types, analogous to Nat.rec
.
It effectively says that every FinEnum
is either Empty
or Option α
, up to an Equiv
mediated
by Fin
s of equal cardinality.
In contrast to the Fintype
case, data can be transported along such an Equiv
.
Also, since order matters, the choice of element that gets replaced by Option.none
has
to be provided for every step.
Since every FinEnum
instance implies a Fintype
instance and Prop
is squashed already,
Fintype.induction_empty_option
can be used if a Prop
needs to be constructed.
Cf. Data.Fintype.Option
Equations
- One or more equations did not get rendered due to their size.
Instances For
For an empty type, the recursion principle evaluates to whatever congr
makes of the base case.
For a type with positive card
, the recursion principle evaluates to whatever
congr
makes of the step result, where Option.none
has been inserted into the
(finChoice (card α - 1))
th rank of the enumeration.
A recursor principle for finite-and-enumerable types, analogous to Nat.recOn
.
It effectively says that every FinEnum
is either Empty
or Option α
, up to an Equiv
mediated
by Fin
s of equal cardinality.
In contrast to the Fintype
case, data can be transported along such an Equiv
.
Also, since order matters, the choice of element that gets replaced by Option.none
has
to be provided for every step.
Equations
- aenum.recOnEmptyOption finChoice congr empty option = FinEnum.recEmptyOption finChoice (fun {α β : Type ?u.96} => congr) empty (fun {α : Type ?u.96} => option) α