Typeclasses¶
This chapter presents a quick reference of the commands related to type classes. For an actual introduction to typeclasses, there is a description of the system [SO08] and the literature on type classes in Haskell which also applies.
Class and Instance declarations¶
The syntax for class and instance declarations is the same as the record syntax of Coq:
The pi : ti
variables are called the parameters of the class and
the fi : ti
are called the methods. Each class definition gives
rise to a corresponding record declaration and each instance is a
regular definition whose name is given by instancename
and type is an
instantiation of the record type.
We’ll use the following example class in the rest of the chapter:
- Class EqDec (A : Type) := { eqb : A -> A -> bool ; eqb_leibniz : forall x y, eqb x y = true -> x = y }.
- EqDec is defined eqb is defined eqb_leibniz is defined
This class implements a boolean equality test which is compatible with Leibniz equality on some type. An example implementation is:
- Instance unit_EqDec : EqDec unit := { eqb x y := true ; eqb_leibniz x y H := match x, y return x = y with | tt, tt => eq_refl tt end }.
- unit_EqDec is defined
Using the refine
attribute, if the term is not sufficient to
finish the definition (e.g. due to a missing field or non-inferable
hole) it must be finished in proof mode. If it is sufficient a trivial
proof mode with no open goals is started.
- #[refine] Instance unit_EqDec' : EqDec unit := { eqb x y := true }.
- 1 subgoal ============================ forall x y : unit, true = true -> x = y
- Proof. intros [] [];reflexivity. Defined.
- No more subgoals.
Note that if you finish the proof with Qed
the entire instance
will be opaque, including the fields given in the initial term.
Alternatively, in Program Mode
if one does not give all the
members in the Instance declaration, Coq generates obligations for the
remaining fields, e.g.:
- Require Import Program.Tactics.
- Program Instance eq_bool : EqDec bool := { eqb x y := if x then y else negb y }.
- eq_bool has type-checked, generating 1 obligation Solving obligations automatically... 1 obligation remaining Obligation 1 of eq_bool: (forall x y : bool, (fun x0 y0 : bool => if x0 then y0 else negb y0) x y = true -> x = y).
- Next Obligation.
- 1 subgoal x, y : bool H : (if x then y else negb y) = true ============================ x = y
- destruct x ; destruct y ; (discriminate || reflexivity).
- No more subgoals.
- Defined.
One has to take care that the transparency of every field is
determined by the transparency of the Instance
proof. One can
use alternatively the program
attribute to get richer
facilities for dealing with obligations.
Binding classes¶
Once a typeclass is declared, one can use it in class binders:
- Definition neqb {A} {eqa : EqDec A} (x y : A) := negb (eqb x y).
- neqb is defined
When one calls a class method, a constraint is generated that is
satisfied only in contexts where the appropriate instances can be
found. In the example above, a constraint EqDec A
is generated and
satisfied by eqa : EqDec A
. In case no satisfying constraint can be
found, an error is raised:
- Fail Definition neqb' (A : Type) (x y : A) := negb (eqb x y).
- The command has indeed failed with message: The following term contains unresolved implicit arguments: (fun (A : Type) (x y : A) => negb (eqb x y)) More precisely: - ?EqDec: Cannot infer the implicit parameter EqDec of eqb whose type is "EqDec A" (no type class instance found) in environment: A : Type x, y : A
The algorithm used to solve constraints is a variant of the eauto
tactic that does proof search with a set of lemmas (the instances). It
will use local hypotheses as well as declared lemmas in
the typeclass_instances
database. Hence the example can also be
written:
- Definition neqb' A (eqa : EqDec A) (x y : A) := negb (eqb x y).
- neqb' is defined
However, the generalizing binders should be used instead as they have particular support for typeclasses:
They automatically set the maximally implicit status for typeclass arguments, making derived functions as easy to use as class methods. In the example above,
A
andeqa
should be set maximally implicit.They support implicit quantification on partially applied type classes (Implicit generalization). Any argument not given as part of a typeclass binder will be automatically generalized.
They also support implicit quantification on Superclasses.
Following the previous example, one can write:
- Generalizable Variables A B C.
- Definition neqb_implicit `{eqa : EqDec A} (x y : A) := negb (eqb x y).
- neqb_implicit is defined
Here A
is implicitly generalized, and the resulting function is
equivalent to the one above.
Parameterized Instances¶
One can declare parameterized instances as in Haskell simply by giving the constraints as a binding context before the instance, e.g.:
- Program Instance prod_eqb `(EA : EqDec A, EB : EqDec B) : EqDec (A * B) := { eqb x y := match x, y with | (la, ra), (lb, rb) => andb (eqb la lb) (eqb ra rb) end }.
- prod_eqb has type-checked, generating 1 obligation Solving obligations automatically... 1 obligation remaining Obligation 1 of prod_eqb: (forall (A : Type) (EA : EqDec A) (B : Type) (EB : EqDec B) (x y : A * B), (fun x0 y0 : A * B => let program_branch_0 := fun (la : A) (ra : B) (lb : A) (rb : B) (_ : (la, ra) = x0) (_ : (lb, rb) = y0) => (eqb la lb && eqb ra rb)%bool in (let (la, ra) as y' return (y' = x0 -> y0 = y0 -> bool) := x0 in let (lb, rb) as x' return ((la, ra) = x0 -> x' = y0 -> bool) := y0 in program_branch_0 la ra lb rb) eq_refl eq_refl) x y = true -> x = y).
- Admit Obligations.
- prod_eqb_obligation_1 is declared No more obligations remaining prod_eqb is defined
These instances are used just as well as lemmas in the instance hint database.
Sections and contexts¶
To ease developments parameterized by many instances, one can use the
Context
command to introduce the parameters into the local context,
it works similarly to the command Variable
, except it accepts any
binding context as an argument, so variables can be implicit, and
Implicit generalization can be used.
For example:
- Section EqDec_defs.
- Context `{EA : EqDec A}.
- A is declared EA is declared
- #[ global, program ] Instance option_eqb : EqDec (option A) := { eqb x y := match x, y with | Some x, Some y => eqb x y | None, None => true | _, _ => false end }.
- option_eqb has type-checked, generating 3 obligations Solving obligations automatically... 3 obligations remaining Obligation 1 of option_eqb: (forall x y : option A, let program_branch_0 := fun (x0 y0 : A) (_ : Some x0 = x) (_ : Some y0 = y) => eqb x0 y0 in let program_branch_1 := fun (_ : None = x) (_ : None = y) => true in let program_branch_2 := fun (wildcard' wildcard'0 : option A) (_ : ~ (None = wildcard' /\ None = wildcard'0) /\ (forall x0 y0 : A, ~ (Some x0 = wildcard' /\ Some y0 = wildcard'0))) (_ : wildcard' = x) (_ : wildcard'0 = y) => false in forall H : A, let wildcard' := Some H in let wildcard'0 := None in ~ (None = wildcard' /\ None = wildcard'0) /\ (forall x0 y0 : A, ~ (Some x0 = wildcard' /\ Some y0 = wildcard'0))). Obligation 2 of option_eqb: (forall x y : option A, let program_branch_0 := fun (x0 y0 : A) (_ : Some x0 = x) (_ : Some y0 = y) => eqb x0 y0 in let program_branch_1 := fun (_ : None = x) (_ : None = y) => true in let program_branch_2 := fun (wildcard' wildcard'0 : option A) (_ : ~ (None = wildcard' /\ None = wildcard'0) /\ (forall x0 y0 : A, ~ (Some x0 = wildcard' /\ Some y0 = wildcard'0))) (_ : wildcard' = x) (_ : wildcard'0 = y) => false in let wildcard' := None in forall H : A, let wildcard'0 := Some H in ~ (None = wildcard' /\ None = wildcard'0) /\ (forall x0 y0 : A, ~ (Some x0 = wildcard' /\ Some y0 = wildcard'0))). Obligation 3 of option_eqb: (forall x y : option A, (fun x0 y0 : option A => let program_branch_0 := fun (x1 y1 : A) (_ : Some x1 = x0) (_ : Some y1 = y0) => eqb x1 y1 in let program_branch_1 := fun (_ : None = x0) (_ : None = y0) => true in let program_branch_2 := fun (wildcard' wildcard'0 : option A) (_ : ~ (None = wildcard' /\ None = wildcard'0) /\ (forall x1 y1 : A, ~ (Some x1 = wildcard' /\ Some y1 = wildcard'0))) (_ : wildcard' = x0) (_ : wildcard'0 = y0) => false in match x0 as y' return (y' = x0 -> y0 = y0 -> bool) with | Some x1 => match y0 as x' return (Some x1 = x0 -> x' = y0 -> bool) with | Some y1 => program_branch_0 x1 y1 | None => program_branch_2 (Some x1) None (option_eqb_obligation_1 x0 y0 x1) end | None => match y0 as x' return (None = x0 -> x' = y0 -> bool) with | Some a => program_branch_2 None (Some a) (option_eqb_obligation_2 x0 y0 a) | None => program_branch_1 end end eq_refl eq_refl) x y = true -> x = y).
- Admit Obligations.
- option_eqb_obligation_1 is declared 2 obligations remaining Solving obligations automatically... 2 obligations remaining option_eqb_obligation_2 is declared 1 obligation remaining Solving obligations automatically... 1 obligation remaining option_eqb_obligation_3 is declared No more obligations remaining option_eqb is defined
- End EqDec_defs.
- About option_eqb.
- option_eqb : forall {A : Type}, EqDec A -> EqDec (option A) option_eqb is not universe polymorphic Arguments option_eqb {A}%type_scope {EA} option_eqb is transparent Expands to: Constant Top.option_eqb
Here the global
attribute redeclares the instance at the end of the
section, once it has been generalized by the context variables it
uses.
See also
Section Section mechanism
Building hierarchies¶
Superclasses¶
One can also parameterize classes by other classes, generating a hierarchy of classes and superclasses. In the same way, we give the superclasses as a binding context:
- Class Ord `(E : EqDec A) := { le : A -> A -> bool }.
- Ord is defined le is defined
Contrary to Haskell, we have no special syntax for superclasses, but this declaration is equivalent to:
This declaration means that any instance of the Ord
class must have
an instance of EqDec
. The parameters of the subclass contain at
least all the parameters of its superclasses in their order of
appearance (here A is the only one). As we have seen, Ord
is encoded
as a record type with two parameters: a type A
and an E
of type
EqDec A
. However, one can still use it as if it had a single
parameter inside generalizing binders: the generalization of
superclasses will be done automatically.
- Definition le_eqb `{Ord A} (x y : A) := andb (le x y) (le y x).
- le_eqb is defined
In some cases, to be able to specify sharing of structures, one may
want to give explicitly the superclasses. It is is possible to do it
directly in regular binders, and using the !
modifier in class
binders. For example:
- Definition lt `{eqa : EqDec A, ! Ord eqa} (x y : A) := andb (le x y) (neqb x y).
- lt is defined
The !
modifier switches the way a binder is parsed back to the usual
interpretation of Coq. In particular, it uses the implicit arguments
mechanism if available, as shown in the example.
Substructures¶
Substructures are components of a class which are instances of a class themselves. They often arise when using classes for logical properties, e.g.:
- Require Import Relation_Definitions.
- Class Reflexive (A : Type) (R : relation A) := reflexivity : forall x, R x x.
- Class Transitive (A : Type) (R : relation A) := transitivity : forall x y z, R x y -> R y z -> R x z.
This declares singleton classes for reflexive and transitive relations, (see the singleton class variant for an explanation). These may be used as parts of other classes:
- Class PreOrder (A : Type) (R : relation A) := { PreOrder_Reflexive :> Reflexive A R ; PreOrder_Transitive :> Transitive A R }.
- PreOrder is defined PreOrder_Reflexive is defined PreOrder_Transitive is defined
The syntax :>
indicates that each PreOrder
can be seen as a
Reflexive
relation. So each time a reflexive relation is needed, a
preorder can be used instead. This is very similar to the coercion
mechanism of Structure
declarations. The implementation simply
declares each projection as an instance.
-
Warning
Ignored instance declaration for “ident”: “term” is not a class
¶ Using this
:>
syntax with a right-hand-side that is not itself a Class has no effect (apart from emitting this warning). In particular, is does not declare a coercion.
One can also declare existing objects or structure projections using the Existing Instance command to achieve the same effect.
Summary of the commands¶
-
Command
Class record_definition
¶ -
Command
Class singleton_class_definition
- singleton_class_definition
::=
>? ident_decl binder* : sort? := constructorThe first form declares a record and makes the record a typeclass with parameters
binder*
and the listed record fields.The second form declares a singleton class with a single method. This singleton class is a so-called definitional class, represented simply as a definition
ident binders := term
and whose instances are themselves objects of this type. Definitional classes are not wrapped inside records, and the trivial projection of an instance of such a class is convertible to the instance itself. This can be useful to make instances of existing objects easily and to reduce proof size by not inserting useless projections. The class constant itself is declared rigid during resolution so that the class abstraction is maintained.Like any command declaring a record, this command supports the
universes(polymorphic)
,universes(template)
,universes(cumulative)
, andprivate(matching)
attributes.
-
Command
Instance ident_decl binder*? : type hint_info? := { field_def* }:= term?
¶ Declares a typeclass instance named
ident_decl
of the classtype
with the specified parameters and with fields defined byfield_def
, where each field must be a declared field of the class.Adds one or more
binder
s to declare a parameterized instance.hint_info
may be used to specify the hint priority, where 0 is the highest priority as forauto
hints. If the priority is not specified, the default is the number of non-dependent binders of the instance. Ifone_pattern
is given, terms matching that pattern will trigger use of the instance. Otherwise, use is triggered based on the conclusion of the type.This command supports the
global
attribute that can be used on instances declared in a section so that their generalization is automatically redeclared when the section is closed.Like
Definition
, it also supports theprogram
attribute to switch the type checking toProgram
(chapter Program) and to use the obligation mechanism to manage missing fields.Finally, it supports the lighter
refine
attribute:-
Attribute
refine
¶ This attribute can be used to leave holes or not provide all fields in the definition of an instance and open the tactic mode to fill them. It works exactly as if no body had been given and the
refine
tactic has been used first.
-
Command
Declare Instance ident_decl binder* : term hint_info?
¶ In a
Module Type
, declares that a corresponding concrete instance should exist in any implementation of thisModule Type
. This is similar to the distinction betweenParameter
vs.Definition
, or betweenDeclare Module
andModule
.
-
Command
Existing Instance qualid hint_info?
¶ -
Command
Existing Instances qualid+ | natural?
¶ Adds a constant whose type ends with an applied typeclass to the instance database with an optional priority
natural
. It can be used for redeclaring instances at the end of sections, or declaring structure projections as instances. This is equivalent toHint Resolve ident : typeclass_instances
, except it registers instances forPrint Instances
.
-
Flag
Instance Generalized Output
¶ Deprecated since version 8.13.
Disabled by default, this provides compatibility with Coq version 8.12 and earlier.
When enabled, the type of the instance is implicitly generalized over unbound and generalizable variables as though surrounded by
\`{}
.
-
Attribute
-
Command
Print Instances reference
¶ Shows the list of instances associated with the typeclass
reference
.
-
Tactic
typeclasses eauto bfs? nat_or_var? with ident+?
¶ This proof search tactic uses the resolution engine that is run implicitly during type checking. This tactic uses a different resolution engine than
eauto
andauto
. The main differences are the following:Unlike
eauto
andauto
, the resolution is done entirely in the proof engine, meaning that backtracking is available among dependent subgoals, and shelving goals is supported.typeclasses eauto
is a multi-goal tactic. It analyses the dependencies between subgoals to avoid backtracking on subgoals that are entirely independent.The transparency information of databases is used consistently for all hints declared in them. It is always used when calling the unifier. When considering local hypotheses, we use the transparent state of the first hint database given. Using an empty database (created with
Create HintDb
for example) with unfoldable variables and constants as the first argument oftypeclasses eauto
hence makes resolution with the local hypotheses use full conversion during unification.The mode hints (see
Hint Mode
) associated with a class are taken into account bytypeclasses eauto
. When a goal does not match any of the declared modes for its head (if any), instead of failing likeeauto
, the goal is suspended and resolution proceeds on the remaining goals. If after one run of resolution, there remains suspended goals, resolution is launched against on them, until it reaches a fixed point when the set of remaining suspended goals does not change. Usingsolve [typeclasses eauto]
can be used to ensure that no suspended goals remain.When considering local hypotheses, we use the union of all the modes declared in the given databases.
Use the
Typeclasses eauto
command to customize the behavior of this tactic.
nat_or_var
Specifies the maximum depth of the search.
Warning
The semantics for the limit
nat_or_var
are different than forauto
. By default, if no limit is given, the search is unbounded. Unlikeauto
, introduction steps count against the limit, which might result in larger limits being necessary when searching withtypeclasses eauto
than withauto
.with ident+
Runs resolution with the specified hint databases. It treats typeclass subgoals the same as other subgoals (no shelving of non-typeclass goals in particular), while allowing shelved goals to remain at any point during search.
When
with
is not specified,typeclasses eauto
uses thetypeclass_instances
database by default (instead ofcore
). Dependent subgoals are automatically shelved, and shelved goals can remain after resolution ends (following the behavior of Coq 8.5).Note
all:once (typeclasses eauto)
faithfully mimics what happens during typeclass resolution when it is called during refinement/type inference, except that only declared class subgoals are considered at the start of resolution during type inference, whileall
can select non-class subgoals as well. It might move toall:typeclasses eauto
in future versions when the refinement engine will be able to backtrack.
-
Tactic
autoapply one_term with ident
¶ The tactic
autoapply
appliesone_term
using the transparency information of the hint databaseident
, and does no typeclass resolution. This can be used inHint Extern
’s for typeclass instances (in the hint databasetypeclass_instances
) to allow backtracking on the typeclass subgoals created by the lemma application, rather than doing typeclass resolution locally at the hint application time.
Typeclasses Transparent, Typeclasses Opaque¶
-
Command
Typeclasses Transparent qualid+
¶ Makes
qualid
transparent during typeclass resolution. A shortcut forHint Transparent
qualid+ : typeclass_instances
-
Command
Typeclasses Opaque qualid+
¶ Make
qualid
opaque for typeclass search. A shortcut forHint Opaque
qualid+ : typeclass_instances
.It is useful when some constants prevent some unifications and make resolution fail. It is also useful to declare constants which should never be unfolded during proof search, like fixpoints or anything which does not look like an abbreviation. This can additionally speed up proof search as the typeclass map can be indexed by such rigid constants (see The hints databases for auto and eauto).
By default, all constants and local variables are considered transparent. One should take care not to make opaque any constant that is used to abbreviate a type, like:
Settings¶
-
Flag
Typeclasses Dependency Order
¶ This flag (off by default) respects the dependency order between subgoals, meaning that subgoals on which other subgoals depend come first, while the non-dependent subgoals were put before the dependent ones previously (Coq 8.5 and below). This can result in quite different performance behaviors of proof search.
-
Flag
Typeclasses Filtered Unification
¶ This flag, which is off by default, switches the hint application procedure to a filter-then-unify strategy. To apply a hint, we first check that the goal matches syntactically the inferred or specified pattern of the hint, and only then try to unify the goal with the conclusion of the hint. This can drastically improve performance by calling unification less often, matching syntactic patterns being very quick. This also provides more control on the triggering of instances. For example, forcing a constant to explicitly appear in the pattern will make it never apply on a goal where there is a hole in that place.
-
Flag
Typeclasses Limit Intros
¶ This flag (on by default) controls the ability to apply hints while avoiding (functional) eta-expansions in the generated proof term. It does so by allowing hints that conclude in a product to apply to a goal with a matching product directly, avoiding an introduction.
Warning
This can be expensive as it requires rebuilding hint clauses dynamically, and does not benefit from the invertibility status of the product introduction rule, resulting in potentially more expensive proof search (i.e. more useless backtracking).
-
Flag
Typeclass Resolution For Conversion
¶ This flag (on by default) controls the use of typeclass resolution when a unification problem cannot be solved during elaboration/type inference. With this flag on, when a unification fails, typeclass resolution is tried before launching unification once again.
-
Flag
Typeclasses Strict Resolution
¶ Typeclass declarations introduced when this flag is set have a stricter resolution behavior (the flag is off by default). When looking for unifications of a goal with an instance of this class, we “freeze” all the existentials appearing in the goals, meaning that they are considered rigid during unification and cannot be instantiated.
-
Flag
Typeclasses Unique Solutions
¶ When a typeclass resolution is launched we ensure that it has a single solution or fail. This ensures that the resolution is canonical, but can make proof search much more expensive.
-
Flag
Typeclasses Unique Instances
¶ Typeclass declarations introduced when this flag is set have a more efficient resolution behavior (the flag is off by default). When a solution to the typeclass goal of this class is found, we never backtrack on it, assuming that it is canonical.
-
Flag
Typeclasses Iterative Deepening
¶ When this flag is set, the proof search strategy is breadth-first search. Otherwise, the search strategy is depth-first search. The default is off.
Typeclasses eauto
is another way to set this flag.
-
Option
Typeclasses Depth natural
¶ Sets the maximum proof search depth. The default is unbounded.
Typeclasses eauto
is another way to set this option.
-
Flag
Typeclasses Debug
¶ Controls whether typeclass resolution steps are shown during search. Setting this flag also sets
Typeclasses Debug Verbosity
to 1.Typeclasses eauto
is another way to set this flag.
-
Option
Typeclasses Debug Verbosity natural
¶ Determines how much information is shown for typeclass resolution steps during search. 1 is the default level. 2 shows additional information such as tried tactics and shelving of goals. Setting this option to 1 or 2 turns on the
Typeclasses Debug
flag; setting this option to 0 turns that flag off.
Typeclasses eauto¶
-
Command
Typeclasses eauto := debug? ( bfsdfs )? natural?
¶ Allows more global customization of the
typeclasses eauto
tactic. The options are:debug
Sets debug mode. In debug mode, a trace of successfully applied tactics is printed. Debug mode can also be set with
Typeclasses Debug
.dfs
,bfs
Sets the search strategy to depth-first search (the default) or breadth-first search. The search strategy can also be set with
Typeclasses Iterative Deepening
.natural
Sets the depth limit for the search. The limit can also be set with
Typeclasses Depth
.