aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2020-09-28 11:14:23 +0200
committerPierre-Marie Pédrot2020-09-30 13:20:27 +0200
commite3a1cf35313bbc4eaca2a43f5fc95ca306bc45fa (patch)
tree288992d95abc5eadcdaa22867ebf0fd944e07a72
parent2c802aaf74c83274ae922c59081c01bfc267d31b (diff)
Remove the forward class hint feature.
It was not documented, not properly tested and thus likely buggy. Judging from the code alone I spotted already one potential bug. Further more it was prominently making use of the infamous "arbitrary term as hint" feature. Since the only user in our CI seems to be a math-classes file that introduced the feature under a claim of "cleanup", I believe we can safely remove it without anyone noticing.
-rw-r--r--pretyping/typeclasses.ml63
-rw-r--r--pretyping/typeclasses.mli10
-rw-r--r--tactics/class_tactics.ml12
-rw-r--r--test-suite/bugs/closed/bug_2928.v11
-rw-r--r--test-suite/output/UnivBinders.out6
-rw-r--r--vernac/classes.ml8
-rw-r--r--vernac/g_vernac.mlg7
-rw-r--r--vernac/ppvernac.ml3
-rw-r--r--vernac/record.ml5
-rw-r--r--vernac/vernacentries.ml2
-rw-r--r--vernac/vernacexpr.ml4
11 files changed, 15 insertions, 116 deletions
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index adb9c5299f..78f04c99e5 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -11,7 +11,6 @@
(*i*)
open Names
open Globnames
-open Term
open Constr
open Vars
open Evd
@@ -42,7 +41,7 @@ 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 = Forward | Backward
+type direction = Backward
(* This module defines type-classes *)
type typeclass = {
@@ -156,66 +155,6 @@ let load_class cl =
(** Build the subinstances hints. *)
-let check_instance env sigma c =
- try
- let (evd, c) = resolve_one_typeclass env sigma
- (Retyping.get_type_of env sigma c) in
- not (Evd.has_undefined evd)
- with e when CErrors.noncritical e -> false
-
-let build_subclasses ~check env sigma glob { hint_priority = pri } =
- let _id = Nametab.basename_of_global glob in
- let _next_id =
- let i = ref (-1) in
- (fun () -> incr i;
- Nameops.add_suffix _id ("_subinstance_" ^ string_of_int !i))
- in
- let ty, ctx = Typeops.type_of_global_in_context env glob in
- let inst, ctx = UnivGen.fresh_instance_from ctx None in
- let ty = Vars.subst_instance_constr inst ty in
- let ty = EConstr.of_constr ty in
- let sigma = Evd.merge_context_set Evd.univ_rigid sigma ctx in
- let rec aux pri c ty path =
- match class_of_constr env sigma ty with
- | None -> []
- | Some (rels, ((tc,u), args)) ->
- let instapp =
- Reductionops.whd_beta env sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect mkRel 0 rels)))
- in
- let instapp = EConstr.Unsafe.to_constr instapp in
- let projargs = Array.of_list (args @ [instapp]) in
- let projs = List.map_filter
- (fun (n, b, proj) ->
- match b with
- | None -> None
- | Some (Backward, _) -> None
- | Some (Forward, info) ->
- let proj = Option.get proj in
- let rels = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) rels in
- let u = EConstr.EInstance.kind sigma u in
- let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in
- if check && check_instance env sigma (EConstr.of_constr body) then None
- else
- let newpri =
- match pri, info.hint_priority with
- | Some p, Some p' -> Some (p + p')
- | Some p, None -> Some (p + 1)
- | _, _ -> None
- in
- Some (GlobRef.ConstRef proj, { info with hint_priority = newpri }, body)) tc.cl_projs
- in
- let declare_proj hints (cref, info, body) =
- let path' = cref :: path in
- let ty = Retyping.get_type_of env sigma (EConstr.of_constr body) in
- let rest = aux pri body ty path' in
- hints @ (path', info, body) :: rest
- in List.fold_left declare_proj [] projs
- in
- let term = Constr.mkRef (glob, inst) in
- (*FIXME subclasses should now get substituted for each particular instance of
- the polymorphic superclass *)
- aux pri term ty [glob]
-
(*
* interface functions
*)
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 9de8083276..c0e71f26c4 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -13,7 +13,7 @@ open Constr
open Evd
open Environ
-type direction = Forward | Backward
+type direction = Backward
(* Core typeclasses hints *)
type 'a hint_info_gen =
@@ -127,11 +127,3 @@ val classes_transparent_state : unit -> TransparentState.t
val solve_all_instances_hook : (env -> evar_map -> evar_filter -> bool -> bool -> bool -> evar_map) Hook.t
val solve_one_instance_hook : (env -> evar_map -> EConstr.types -> bool -> evar_map * EConstr.constr) Hook.t
-
-(** Build the subinstances hints for a given typeclass object.
- check tells if we should check for existence of the
- subinstances and add only the missing ones. *)
-
-val build_subclasses : check:bool -> env -> evar_map -> GlobRef.t ->
- hint_info ->
- (GlobRef.t list * hint_info * constr) list
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index b4d7e7d7f0..ed92a85a12 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -483,17 +483,7 @@ let make_resolve_hyp env sigma st only_classes pri decl =
if keep then
let id = GlobRef.VarRef id in
let name = PathHints [id] in
- let hints =
- if is_class then
- let hints = build_subclasses ~check:false env sigma id empty_hint_info in
- (List.map_append
- (fun (path,info,c) ->
- let h = IsConstr (EConstr.of_constr c, None) [@ocaml.warning "-3"] in
- make_resolves env sigma ~name:(PathHints path) info ~check:true h)
- hints)
- else []
- in
- (hints @ make_resolves env sigma pri ~name ~check:false (IsGlobRef id))
+ (make_resolves env sigma pri ~name ~check:false (IsGlobRef id))
else []
let make_hints g (modes,st) only_classes sign =
diff --git a/test-suite/bugs/closed/bug_2928.v b/test-suite/bugs/closed/bug_2928.v
deleted file mode 100644
index 21e92ae20c..0000000000
--- a/test-suite/bugs/closed/bug_2928.v
+++ /dev/null
@@ -1,11 +0,0 @@
-Class Equiv A := equiv: A -> A -> Prop.
-Infix "=" := equiv : type_scope.
-
-Class Associative {A} f `{Equiv A} := associativity x y z : f x (f y z) = f (f x y) z.
-
-Class SemiGroup A op `{Equiv A} := { sg_ass :>> Associative op }.
-
-Class SemiLattice A op `{Equiv A} :=
- { semilattice_sg :>> SemiGroup A op
- ; redundant : Associative op
- }.
diff --git a/test-suite/output/UnivBinders.out b/test-suite/output/UnivBinders.out
index 163ed15606..d8d3f696b7 100644
--- a/test-suite/output/UnivBinders.out
+++ b/test-suite/output/UnivBinders.out
@@ -67,9 +67,9 @@ mono
The command has indeed failed with message:
Universe u already exists.
bobmorane =
-let tt := Type@{UnivBinders.33} in
-let ff := Type@{UnivBinders.35} in tt -> ff
- : Type@{max(UnivBinders.32,UnivBinders.34)}
+let tt := Type@{UnivBinders.32} in
+let ff := Type@{UnivBinders.34} in tt -> ff
+ : Type@{max(UnivBinders.31,UnivBinders.33)}
The command has indeed failed with message:
Universe u already bound.
foo@{E M N} =
diff --git a/vernac/classes.ml b/vernac/classes.ml
index b38a249b73..3485d17951 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 =
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..79abf94726 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -557,7 +557,7 @@ 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 (Backward, List.hd priorities)
| None -> None
in
[cref, [Name proj_name, sub, Some proj_cst]]
@@ -568,8 +568,7 @@ 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 () -> Backward, pri) coe)
coers priorities
in
let map ind =
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 *)