Set of factors #
Main definitions #
Associates.FactorSet
: multiset of factors of an element, unique up to propositional equality.Associates.factors
: determine theFactorSet
for a given element.
TODO #
- set up the complete lattice structure on
FactorSet
.
FactorSet α
representation elements of unique factorization domain as multisets.
Multiset α
produced by normalizedFactors
are only unique up to associated elements, while the
multisets in FactorSet α
are unique by equality and restricted to irreducible elements. This
gives us a representation of each element as a unique multisets (or the added ⊤ for 0), which has a
complete lattice structure. Infimum is the greatest common divisor and supremum is the least common
multiple.
Equations
- Associates.FactorSet α = WithTop (Multiset { a : Associates α // Irreducible a })
Instances For
Evaluates the product of a FactorSet
to be the product of the corresponding multiset,
or 0
if there is none.
Equations
- Associates.FactorSet.prod none = 0
- Associates.FactorSet.prod (some s) = (Multiset.map Subtype.val s).prod
Instances For
bcount p s
is the multiplicity of p
in the FactorSet s
(with bundled p
)
Equations
- Associates.bcount p none = 0
- Associates.bcount p (some s) = Multiset.count p s
Instances For
count p s
is the multiplicity of the irreducible p
in the FactorSet s
.
If p
is not irreducible, count p s
is defined to be 0
.
Equations
- p.count = if hp : Irreducible p then Associates.bcount ⟨p, hp⟩ else 0
Instances For
membership in a FactorSet (bundled version)
Equations
- Associates.BfactorSetMem x none = True
- Associates.BfactorSetMem x (some l) = (x ∈ l)
Instances For
FactorSetMem p s
is the predicate that the irreducible p
is a member of
s : FactorSet α
.
If p
is not irreducible, p
is not a member of any FactorSet
.
Equations
- Associates.FactorSetMem s p = if hp : Irreducible p then Associates.BfactorSetMem ⟨p, hp⟩ s else False
Instances For
Equations
- Associates.instMembershipFactorSet = { mem := Associates.FactorSetMem }
This returns the multiset of irreducible factors as a FactorSet
,
a multiset of irreducible associates WithTop
.
Equations
- Associates.factors' a = Multiset.pmap (fun (a : α) (ha : Irreducible a) => ⟨Associates.mk a, ⋯⟩) (UniqueFactorizationMonoid.factors a) ⋯
Instances For
This returns the multiset of irreducible factors of an associate as a FactorSet
,
a multiset of irreducible associates WithTop
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Alias of Associates.factors_zero
.
Alias of Associates.factors_eq_top_iff_zero
.
Equations
- Associates.instMax = { max := fun (a b : Associates α) => (a.factors ⊔ b.factors).prod }
Equations
- Associates.instMin = { min := fun (a b : Associates α) => (a.factors ⊓ b.factors).prod }
Equations
- Associates.instLattice = Lattice.mk (fun (x1 x2 : Associates α) => x1 ⊓ x2) ⋯ ⋯ ⋯
The only divisors of prime powers are prime powers. See eq_pow_find_of_dvd_irreducible_pow
for an explicit expression as a p-power (without using count
).
The only divisors of prime powers are prime powers.