Typeclass for a type F
with an injective map to A → B
#
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This typeclass is primarily for use by homomorphisms like monoid_hom
and linear_map
.
Basic usage of fun_like
#
A typical type of morphisms should be declared as:
structure my_hom (A B : Type*) [my_class A] [my_class B] :=
(to_fun : A → B)
(map_op' : ∀ {x y : A}, to_fun (my_class.op x y) = my_class.op (to_fun x) (to_fun y))
namespace my_hom
variables (A B : Type*) [my_class A] [my_class B]
-- This instance is optional if you follow the "morphism class" design below:
instance : fun_like (my_hom A B) A (λ _, B) :=
{ coe := my_hom.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr' }
/-- Helper instance for when there's too many metavariables to apply
`fun_like.has_coe_to_fun` directly. -/
instance : has_coe_to_fun (my_hom A B) (λ _, A → B) := fun_like.has_coe_to_fun
@[simp] lemma to_fun_eq_coe {f : my_hom A B} : f.to_fun = (f : A → B) := rfl
@[ext] theorem ext {f g : my_hom A B} (h : ∀ x, f x = g x) : f = g := fun_like.ext f g h
/-- Copy of a `my_hom` with a new `to_fun` equal to the old one. Useful to fix definitional
equalities. -/
protected def copy (f : my_hom A B) (f' : A → B) (h : f' = ⇑f) : my_hom A B :=
{ to_fun := f',
map_op' := h.symm ▸ f.map_op' }
end my_hom
This file will then provide a has_coe_to_fun
instance and various
extensionality and simp lemmas.
Morphism classes extending fun_like
#
The fun_like
design provides further benefits if you put in a bit more work.
The first step is to extend fun_like
to create a class of those types satisfying
the axioms of your new type of morphisms.
Continuing the example above:
section
set_option old_structure_cmd true
/-- `my_hom_class F A B` states that `F` is a type of `my_class.op`-preserving morphisms.
You should extend this class when you extend `my_hom`. -/
class my_hom_class (F : Type*) (A B : out_param $ Type*) [my_class A] [my_class B]
extends fun_like F A (λ _, B) :=
(map_op : ∀ (f : F) (x y : A), f (my_class.op x y) = my_class.op (f x) (f y))
end
@[simp] lemma map_op {F A B : Type*} [my_class A] [my_class B] [my_hom_class F A B]
(f : F) (x y : A) : f (my_class.op x y) = my_class.op (f x) (f y) :=
my_hom_class.map_op
-- You can replace `my_hom.fun_like` with the below instance:
instance : my_hom_class (my_hom A B) A B :=
{ coe := my_hom.to_fun,
coe_injective' := λ f g h, by cases f; cases g; congr',
map_op := my_hom.map_op' }
-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here]
The second step is to add instances of your new my_hom_class
for all types extending my_hom
.
Typically, you can just declare a new class analogous to my_hom_class
:
structure cooler_hom (A B : Type*) [cool_class A] [cool_class B]
extends my_hom A B :=
(map_cool' : to_fun cool_class.cool = cool_class.cool)
section
set_option old_structure_cmd true
class cooler_hom_class (F : Type*) (A B : out_param $ Type*) [cool_class A] [cool_class B]
extends my_hom_class F A B :=
(map_cool : ∀ (f : F), f cool_class.cool = cool_class.cool)
end
@[simp] lemma map_cool {F A B : Type*} [cool_class A] [cool_class B] [cooler_hom_class F A B]
(f : F) : f cool_class.cool = cool_class.cool :=
my_hom_class.map_op
-- You can also replace `my_hom.fun_like` with the below instance:
instance : cool_hom_class (cool_hom A B) A B :=
{ coe := cool_hom.to_fun,
coe_injective' := λ f g h, by cases f; cases g; congr',
map_op := cool_hom.map_op',
map_cool := cool_hom.map_cool' }
-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here]
Then any declaration taking a specific type of morphisms as parameter can instead take the class you just defined:
-- Compare with: lemma do_something (f : my_hom A B) : sorry := sorry
lemma do_something {F : Type*} [my_hom_class F A B] (f : F) : sorry := sorry
This means anything set up for my_hom
s will automatically work for cool_hom_class
es,
and defining cool_hom_class
only takes a constant amount of effort,
instead of linearly increasing the work per my_hom
-related declaration.
- coe : F → Π (a : α), β a
- coe_injective' : function.injective fun_like.coe
The class fun_like F α β
expresses that terms of type F
have an
injective coercion to functions from α
to β
.
This typeclass is used in the definition of the homomorphism typeclasses,
such as zero_hom_class
, mul_hom_class
, monoid_hom_class
, ....
Instances of this typeclass
- embedding_like.to_fun_like
- zero_hom_class.to_fun_like
- add_hom_class.to_fun_like
- one_hom_class.to_fun_like
- mul_hom_class.to_fun_like
- rel_hom_class.to_fun_like
- nonneg_hom_class.to_fun_like
- subadditive_hom_class.to_fun_like
- submultiplicative_hom_class.to_fun_like
- mul_le_add_hom_class.to_fun_like
- nonarchimedean_hom_class.to_fun_like
- top_hom_class.to_fun_like
- bot_hom_class.to_fun_like
- sup_hom_class.to_fun_like
- inf_hom_class.to_fun_like
- smul_hom_class.to_fun_like
- star_hom_class.to_fun_like
- Sup_hom_class.to_fun_like
- Inf_hom_class.to_fun_like
- continuous_map_class.to_fun_like
- locally_bounded_map_class.to_fun_like
- spectral_map_class.to_fun_like
- dilation_class.to_fun_like
- slash_invariant_form_class.to_fun_like
- dfinsupp.fun_like
- finsupp.fun_like
- basis.fun_like
- pequiv.fun_like
- alternating_map.fun_like
- affine_map.fun_like
- affine_basis.fun_like
- pmf.fun_like
- gen_loop.fun_like
- schwartz_map.fun_like
- cont_mdiff_map.fun_like
- poly.fun_like
- pointed_smooth_map.fun_like
- freiman_hom.fun_like
- add_freiman_hom.fun_like
- quadratic_form.fun_like
- ssyt.fun_like
- first_order.language.hom.fun_like
- first_order.language.elementary_embedding.fun_like
Instances of other typeclasses for fun_like
- fun_like.has_sizeof_inst
Equations
This is not an instance to avoid slowing down every single subsingleton
typeclass search.