aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2020-09-30 13:21:02 +0200
committerPierre-Marie Pédrot2020-09-30 13:41:38 +0200
commitf16290030b48dedf3091334af4cd21a7df157381 (patch)
treed8719b40c6ab8e009816f1a6f4f74aa67c72eec3
parente3a1cf35313bbc4eaca2a43f5fc95ca306bc45fa (diff)
Further simplification of the typeclass registration API.
-rw-r--r--plugins/ltac/rewrite.ml2
-rw-r--r--pretyping/typeclasses.ml9
-rw-r--r--pretyping/typeclasses.mli10
-rw-r--r--vernac/classes.ml27
-rw-r--r--vernac/record.ml21
5 files changed, 47 insertions, 22 deletions
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index a1dbf9a439..8de6cbc0a6 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -181,7 +181,7 @@ end) = struct
fun env sigma -> class_info env sigma (Lazy.force r)
let proper_proj env sigma =
- mkConst (Option.get (pi3 (List.hd (proper_class env sigma).cl_projs)))
+ mkConst (Option.get (List.hd (proper_class env sigma).cl_projs).meth_const)
let proper_type env (sigma,cstrs) =
let l = (proper_class env sigma).cl_impl in
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 78f04c99e5..fc71254a46 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -41,7 +41,11 @@ let get_solve_one_instance, solve_one_instance_hook = Hook.make ()
let resolve_one_typeclass ?(unique=get_typeclasses_unique_solutions ()) env evm t =
Hook.get get_solve_one_instance env evm t unique
-type direction = Backward
+type class_method = {
+ meth_name : Name.t;
+ meth_info : hint_info option;
+ meth_const : Constant.t option;
+}
(* This module defines type-classes *)
type typeclass = {
@@ -58,8 +62,7 @@ type typeclass = {
cl_props : Constr.rel_context;
(* The method implementations as projections. *)
- cl_projs : (Name.t * (direction * hint_info) option
- * Constant.t option) list;
+ cl_projs : class_method list;
cl_strict : bool;
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index c0e71f26c4..3f84d08a7e 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -13,8 +13,6 @@ open Constr
open Evd
open Environ
-type direction = Backward
-
(* Core typeclasses hints *)
type 'a hint_info_gen =
{ hint_priority : int option;
@@ -22,6 +20,12 @@ type 'a hint_info_gen =
type hint_info = (Pattern.patvar list * Pattern.constr_pattern) hint_info_gen
+type class_method = {
+ meth_name : Name.t;
+ meth_info : hint_info option;
+ meth_const : Constant.t option;
+}
+
(** This module defines type-classes *)
type typeclass = {
cl_univs : Univ.AUContext.t;
@@ -39,7 +43,7 @@ type typeclass = {
cl_props : Constr.rel_context;
(** Context of definitions and properties on defs, will not be shared *)
- cl_projs : (Name.t * (direction * hint_info) option * Constant.t option) list;
+ cl_projs : class_method list;
(** The methods implementations of the typeclass as projections.
Some may be undefinable due to sorting restrictions or simply undefined if
no name is provided. The [int option option] indicates subclasses whose hint has
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 3485d17951..a464eab127 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -155,8 +155,17 @@ let subst_class (subst,cl) =
let do_subst_context (grs,ctx) =
List.Smart.map (Option.Smart.map do_subst_gr) grs,
do_subst_ctx ctx in
- let do_subst_projs projs = List.Smart.map (fun (x, y, z) ->
- (x, y, Option.Smart.map do_subst_con z)) projs in
+ let do_subst_meth m =
+ let c = Option.Smart.map do_subst_con m.meth_const in
+ if c == m.meth_const then m
+ else
+ {
+ meth_name = m.meth_name;
+ meth_info = m.meth_info;
+ meth_const = c;
+ }
+ in
+ let do_subst_projs projs = List.Smart.map do_subst_meth projs in
{ cl_univs = cl.cl_univs;
cl_impl = do_subst_gr cl.cl_impl;
cl_context = do_subst_context cl.cl_context;
@@ -241,10 +250,10 @@ let add_class cl =
let add_class env sigma cl =
add_class cl;
- List.iter (fun (n, inst, body) ->
- match inst with
- | Some (Backward, info) ->
- (match body with
+ List.iter (fun m ->
+ match m.meth_info with
+ | Some info ->
+ (match m.meth_const with
| None -> CErrors.user_err Pp.(str "Non-definable projection can not be declared as a subinstance")
| Some b -> declare_instance ~warn:true env sigma (Some info) false (GlobRef.ConstRef b))
| _ -> ())
@@ -424,9 +433,9 @@ let do_instance_type_ctx_instance props k env' ctx' sigma ~program_mode subst =
let rest' = List.filter (fun v -> not (is_id v)) rest
in
let {CAst.loc;v=mid} = get_id loc_mid in
- List.iter (fun (n, _, x) ->
- if Name.equal n (Name mid) then
- Option.iter (fun x -> Dumpglob.add_glob ?loc (GlobRef.ConstRef x)) x) k.cl_projs;
+ List.iter (fun m ->
+ if Name.equal m.meth_name (Name mid) then
+ Option.iter (fun x -> Dumpglob.add_glob ?loc (GlobRef.ConstRef x)) m.meth_const) k.cl_projs;
c :: props, rest'
with Not_found ->
((CAst.make @@ CHole (None(* Some Evar_kinds.GoalEvar *), Namegen.IntroAnonymous, None)) :: props), rest
diff --git a/vernac/record.ml b/vernac/record.ml
index 79abf94726..89acd79dda 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -557,10 +557,15 @@ let declare_class def cumulative ubinders univs id idbuild paramimpls params uni
Impargs.declare_manual_implicits false (GlobRef.ConstRef proj_cst) (List.hd fieldimpls);
Classes.set_typeclass_transparency (EvalConstRef cst) false false;
let sub = match List.hd coers with
- | Some () -> Some (Backward, List.hd priorities)
+ | Some () -> Some (List.hd priorities)
| None -> None
in
- [cref, [Name proj_name, sub, Some proj_cst]]
+ let m = {
+ meth_name = Name proj_name;
+ meth_info = sub;
+ meth_const = Some proj_cst;
+ } in
+ [cref, [m]]
| _ ->
let record_data = [id, idbuild, univ, arity, fieldimpls, fields, false,
List.map (fun _ -> { pf_subclass = false ; pf_canonical = true }) fields] in
@@ -568,13 +573,17 @@ let declare_class def cumulative ubinders univs id idbuild paramimpls params uni
params template ~kind:Decls.Method ~name:[|binder_name|] record_data
in
let coers = List.map2 (fun coe pri ->
- Option.map (fun () -> Backward, pri) coe)
+ Option.map (fun () -> pri) coe)
coers priorities
in
let map ind =
- let l = List.map3 (fun decl b y -> RelDecl.get_name decl, b, y)
- (List.rev fields) coers (Recordops.lookup_projections ind)
- in GlobRef.IndRef ind, l
+ let map decl b y = {
+ meth_name = RelDecl.get_name decl;
+ meth_info = b;
+ meth_const = y;
+ } in
+ let l = List.map3 map (List.rev fields) coers (Recordops.lookup_projections ind) in
+ GlobRef.IndRef ind, l
in
List.map map inds
in