aboutsummaryrefslogtreecommitdiff
path: root/vernac
diff options
context:
space:
mode:
authorcoqbot-app[bot]2020-10-02 11:04:53 +0000
committerGitHub2020-10-02 11:04:53 +0000
commitbb2d0d56df08ca54764be5a3eb5c09ce00009d6c (patch)
tree2f10c490ef22db87d8b8c8c8929c17466bec298f /vernac
parent42a5e337c7a33bf0ec9530b6ce161a3053362b3d (diff)
parentf16290030b48dedf3091334af4cd21a7df157381 (diff)
Merge PR #13106: Remove the forward class hint feature.
Reviewed-by: SkySkimmer
Diffstat (limited to 'vernac')
-rw-r--r--vernac/classes.ml35
-rw-r--r--vernac/g_vernac.mlg7
-rw-r--r--vernac/ppvernac.ml3
-rw-r--r--vernac/record.ml22
-rw-r--r--vernac/vernacentries.ml2
-rw-r--r--vernac/vernacexpr.ml4
6 files changed, 40 insertions, 33 deletions
diff --git a/vernac/classes.ml b/vernac/classes.ml
index b38a249b73..a464eab127 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -58,13 +58,7 @@ let is_local_for_hint i =
let add_instance_base inst =
let locality = if is_local_for_hint inst then Goptions.OptLocal else Goptions.OptGlobal in
add_instance_hint (Hints.IsGlobRef inst.is_impl) [inst.is_impl] ~locality
- inst.is_info;
- List.iter (fun (path, pri, c) ->
- let h = Hints.IsConstr (EConstr.of_constr c, None) [@ocaml.warning "-3"] in
- add_instance_hint h path
- ~locality pri)
- (build_subclasses ~check:(not (isVarRef inst.is_impl))
- (Global.env ()) (Evd.from_env (Global.env ())) inst.is_impl inst.is_info)
+ inst.is_info
let mk_instance cl info glob impl =
let global =
@@ -161,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;
@@ -247,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))
| _ -> ())
@@ -430,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/g_vernac.mlg b/vernac/g_vernac.mlg
index e0550fd744..5b039e76f3 100644
--- a/vernac/g_vernac.mlg
+++ b/vernac/g_vernac.mlg
@@ -469,11 +469,8 @@ GRAMMAR EXTEND Gram
[ [ id = identref; c=constructor_type -> { c id } ] ]
;
of_type_with_opt_coercion:
- [ [ ":>>" -> { Some false }
- | ":>"; ">" -> { Some false }
- | ":>" -> { Some true }
- | ":"; ">"; ">" -> { Some false }
- | ":"; ">" -> { Some true }
+ [ [ ":>" -> { Some () }
+ | ":"; ">" -> { Some () }
| ":" -> { None } ] ]
;
END
diff --git a/vernac/ppvernac.ml b/vernac/ppvernac.ml
index b73e7c7515..8a98a43ba0 100644
--- a/vernac/ppvernac.ml
+++ b/vernac/ppvernac.ml
@@ -504,8 +504,7 @@ let pr_intarg n = spc () ++ int n
let pr_oc = function
| None -> str" :"
- | Some true -> str" :>"
- | Some false -> str" :>>"
+ | Some () -> str" :>"
let pr_record_field (x, { rf_subclass = oc ; rf_priority = pri ; rf_notation = ntn }) =
let prx = match x with
diff --git a/vernac/record.ml b/vernac/record.ml
index bd5b71cd6b..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 b -> Some ((if b then Backward else Forward), 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,14 +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 b ->
- if b then Backward, pri else Forward, 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
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index fba6800729..60c6d2ec0b 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -776,7 +776,7 @@ let vernac_inductive ~atts kind indl =
| _ -> CErrors.user_err Pp.(str "Definitional classes do not support the \"|\" syntax.")
in
let (coe, (lid, ce)) = l in
- let coe' = if coe then Some true else None in
+ let coe' = if coe then Some () else None in
let f = AssumExpr ((make ?loc:lid.loc @@ Name lid.v), ce),
{ rf_subclass = coe' ; rf_priority = None ; rf_notation = [] ; rf_canonical = true } in
vernac_record ~template udecl ~cumulative (Class true) ~poly finite [id, bl, c, None, [f]]
diff --git a/vernac/vernacexpr.ml b/vernac/vernacexpr.ml
index d8e17d00e3..721e710e1d 100644
--- a/vernac/vernacexpr.ml
+++ b/vernac/vernacexpr.ml
@@ -106,8 +106,8 @@ type search_restriction =
type verbose_flag = bool (* true = Verbose; false = Silent *)
type coercion_flag = bool (* true = AddCoercion false = NoCoercion *)
-type instance_flag = bool option
- (* Some true = Backward instance; Some false = Forward instance, None = NoInstance *)
+type instance_flag = unit option
+ (* Some () = Backward instance, None = NoInstance *)
type export_flag = bool (* true = Export; false = Import *)