Commit 2022-04-05 12:54 da132ecf
View on Github →feat(*): define subobject classes from submonoid up to subfield (#11750)
The next part of my big refactoring plans: subobject classes in the same style as morphism classes.
This PR introduces the following subclasses of set_like
:
one_mem_class
,zero_mem_class
,mul_mem_class
,add_mem_class
,inv_mem_class
,neg_mem_class
submonoid_class
,add_submonoid_class
subgroup_class
,add_subgroup_class
subsemiring_class
,subring_class
,subfield_class
The main purpose of this refactor is that we can replace the wide variety of lemmas like{add_submonoid,add_subgroup,subring,subfield,submodule,subwhatever}.{prod,sum}_mem
with a singleprod_mem
lemma that is generic over all typesB
that extendsubmonoid
:
@[to_additive]
lemma prod_mem {M : Type*} [comm_monoid M] [set_like B M] [submonoid_class B M]
{ι : Type*} {t : finset ι} {f : ι → M} (h : ∀c ∈ t, f c ∈ S) : ∏ c in t, f c ∈ S
API changes
- When you extend a
struct subobject
, make sure to create a correspondingsubobject_class
instance.
Upcoming PRs
This PR splits out the first part of #11545, namely defining the subobject classes. I am planning these follow-up PRs for further parts of #11545:
- make the subobject consistently implicit in
{add,mul}_mem
#11758 - remove duplicate instances like
subgroup.to_group
(replaced by thesubgroup_class.to_subgroup
instances that are added by this PR) #11759 - further deduplication such as
finsupp_sum_mem
Subclassing set_like
Contrary to mathlib's typical subclass pattern, we don't extend set_like
, but take a set_like
instance parameter:
class one_mem_class (S : Type*) (M : out_param $ Type*) [has_one M] [set_like S M] :=
(one_mem : ∀ (s : S), (1 : M) ∈ s)
instead of:
class one_mem_class (S : Type*) (M : out_param $ Type*) [has_one M]
extends set_like S M :=
(one_mem : ∀ (s : S), (1 : M) ∈ s)
The main reason is that this avoids some big defeq checks when typechecking e.g. x * y : s
, where s : S
and [comm_group G] [subgroup_class S G]
. Namely, the type coe_sort s
could be given by subgroup_class → @@submonoid_class _ _ (comm_group.to_group.to_monoid) → set_like → has_coe_to_sort
or by subgroup_class → @@submonoid_class _ _ (comm_group.to_comm_monoid.to_monoid) → set_like → has_coe_to_sort
. When checking that has_mul
on the first type is the same as has_mul
on the second type, those two inheritance paths are unified many times over (sometimes exponentially many). So it's important to keep the size of types small, and therefore we avoid extends
-based inheritance.
Defeq fixes
Adding instances like subgroup_class.to_group
means that there are now two (defeq) group instances for subgroup
. This makes some code more fragile, until we can replace subgroup.to_group
with its more generic form in a follow-up PR. Especially when taking subgroups of subgroups I needed to help the elaborator in a few places. These should be minimally invasive for other uses of the code.
Timeout fixes
Some of the leaf files started timing out, so I made a couple of fixes. Generally these can be classed as:
squeeze_simps
- Give inheritance
subX_class S M
→X s
(wheres : S
) a lower prority thanY s
→X s
so thatsubY_class S M
→Y s
→X s
is preferred oversubY_class S M
→subX_class S M
→X s
. This addresses slow unifications whenx : s
,s
is a submonoid oft
, which is itself a subgroup ofG
: existing code expects to gosubgroup → group → monoid
, which got changed tosubgroup_class → submonoid_class → monoid
; when this kind of unification issue appears in your type this results in slow unification. By tweaking the priorities, we help the elaborator find our preferred instance, avoiding the big defeq checks. (The real fix should of course be to fix the unifier so it doesn't become exponential in these kinds of cases.) - Split a long proof with duplication into smaller parts. This was basically my last resort.
I decided to bump the limit for the
fails_quickly
linter formeasure_theory.Lp_meas.complete_space
, which apparently just barely goes over this limit now. The time difference was about 10%-20% for that specific instance.