Skip to content

Commit

Permalink
Uniformize semantics of :> in class after deprecation phase of #16230
Browse files Browse the repository at this point in the history
This ends the deprecation phases of
#16230 and
#15802
that introduced :: instead of :> for subclasses in Class declarations.
Now :> means coercion in all record declarations (including typeclasses).
  • Loading branch information
proux01 committed Jan 30, 2024
1 parent a6d226f commit 9b95a02
Show file tree
Hide file tree
Showing 8 changed files with 18 additions and 74 deletions.
18 changes: 4 additions & 14 deletions doc/sphinx/addendum/type-classes.rst
Original file line number Diff line number Diff line change
Expand Up @@ -248,7 +248,7 @@ the implicit arguments mechanism if available, as shown in the example.
Substructures
~~~~~~~~~~~~~

.. index:: :> (substructure)
.. index:: :: (substructure)

Substructures are components of a typeclass which are themselves instances of a
typeclass. They often arise when using typeclasses for logical properties, e.g.:
Expand Down Expand Up @@ -309,9 +309,8 @@ Command summary
declared rigid during resolution so that the typeclass abstraction is
maintained.

The `>` in
:token:`record_definition` currently does nothing. In a future version, it will
create coercions as it does when used in :cmd:`Record` commands.
The `>` in :token:`record_definition`
creates coercions as it does when used in :cmd:`Record` commands.

Like any command declaring a record, this command supports the
:attr:`universes(polymorphic)`, :attr:`universes(template)`,
Expand All @@ -333,18 +332,9 @@ Command summary

This command has no effect when used on a typeclass.

.. _warn-future-coercion-class-field:

.. warn:: A coercion will be introduced instead of an instance in future versions when using ':>' in 'Class' declarations. Replace ':>' with '::' (or use '#[global] Existing Instance field.' for compatibility with Coq < 8.17).

In future versions, :g:`:>` in the :n:`@record_definition` or
:n:`@constructor` will declare a :ref:`coercion<coercions>`, as
it does for other :cmd:`Record` commands. To eliminate the warning, use
:g:`::` instead.

.. warn:: Ignored instance declaration for “@ident”: “@term” is not a class

Using the ``::`` (or deprecated ``:>``) syntax in the :n:`@record_definition`
Using the ``::`` syntax in the :n:`@record_definition`
or :n:`@constructor` with a right-hand-side that
is not itself a Class has no effect (apart from emitting this warning).

Expand Down
1 change: 0 additions & 1 deletion doc/sphinx/language/core/records.rst
Original file line number Diff line number Diff line change
Expand Up @@ -92,7 +92,6 @@ Defining record types
:n:`:>`
If specified, the field is declared as a coercion from the record name
to the class of the field type. See :ref:`coercions`.
Note that this currently does something else in :cmd:`Class` commands.

:n:`::`
If specified, the field is declared a typeclass instance of the class
Expand Down
3 changes: 1 addition & 2 deletions vernac/comInductive.ml
Original file line number Diff line number Diff line change
Expand Up @@ -671,8 +671,7 @@ let eq_params (up1,p1) (up2,p2) =
let extract_coercions indl =
let mkqid (_,({CAst.v=id},_)) = qualid_of_ident id in
let iscoe (_, coe, inst) = match inst with
(* remove BackInstanceWarning after deprecation phase *)
| Vernacexpr.(NoInstance | BackInstanceWarning) -> coe = Vernacexpr.AddCoercion
| Vernacexpr.NoInstance -> coe = Vernacexpr.AddCoercion
| _ -> user_err (Pp.str "'::' not allowed in inductives.") in
let extract lc = List.filter (fun (coe,_) -> iscoe coe) lc in
List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl))
Expand Down
4 changes: 2 additions & 2 deletions vernac/g_vernac.mlg
Original file line number Diff line number Diff line change
Expand Up @@ -576,8 +576,8 @@ GRAMMAR EXTEND Gram
| ":" -> { NoCoercion } ] ]
;
of_type_inst:
[ [ ":>" -> { (AddCoercion, BackInstanceWarning) (* replace with NoInstance at end of deprecation phase *) }
| ":"; ">" -> { (AddCoercion, BackInstanceWarning) (* replace with NoInstance at end of deprecation phase *) }
[ [ ":>" -> { (AddCoercion, NoInstance) }
| ":"; ">" -> { (AddCoercion, NoInstance) }
| "::" -> { (NoCoercion, BackInstance) }
| "::>" -> { (AddCoercion, BackInstance) }
| ":" -> { (NoCoercion, NoInstance) } ] ]
Expand Down
1 change: 0 additions & 1 deletion vernac/ppvernac.ml
Original file line number Diff line number Diff line change
Expand Up @@ -544,7 +544,6 @@ let pr_oc coe ins = match coe, ins with
| AddCoercion, NoInstance -> str" :>"
| NoCoercion, BackInstance -> str" ::"
| AddCoercion, BackInstance -> str" ::>"
| _, BackInstanceWarning -> str" :>" (* remove this line at end of deprecation phase *)

let pr_record_field (x, { rfu_attrs = attr ; rfu_coercion = coe ; rfu_instance = ins ; rfu_priority = pri ; rfu_notation = ntn }) =
let prx = match x with
Expand Down
56 changes: 8 additions & 48 deletions vernac/record.ml
Original file line number Diff line number Diff line change
Expand Up @@ -363,10 +363,8 @@ let instantiate_possibly_recursive_type ind u ntypes paramdecls fields =

(** Declare projection [ref] over [from] a coercion
or a typeclass instance according to [flags]. *)
(* remove the last argument (it will become alway true) after deprecation phase
(started in 8.17, c.f. https://github.com/coq/coq/pull/16230) *)
let declare_proj_coercion_instance ~flags ref from ~with_coercion =
if with_coercion && flags.Data.pf_coercion then begin
let declare_proj_coercion_instance ~flags ref from =
if flags.Data.pf_coercion then begin
let cl = ComCoercion.class_of_global from in
let local = flags.Data.pf_locality = Goptions.OptLocal in
ComCoercion.try_add_new_coercion_with_source ref ~local ~reversible:flags.Data.pf_reversible ~source:cl
Expand Down Expand Up @@ -438,7 +436,7 @@ let build_named_proj ~primitive ~flags ~poly ~univs ~uinstance ~kind env paramde
in
let refi = GlobRef.ConstRef kn in
Impargs.maybe_declare_manual_implicits false refi impls;
declare_proj_coercion_instance ~flags refi (GlobRef.IndRef indsp) ~with_coercion:true;
declare_proj_coercion_instance ~flags refi (GlobRef.IndRef indsp);
let i = if is_local_assum decl then i+1 else i in
(Some kn, i, Projection term::subst)

Expand Down Expand Up @@ -655,37 +653,18 @@ let implicits_of_context ctx =
| LocalAssum _ as d -> Some (RelDecl.get_name d))
ctx)))

(* deprecated in 8.16, to be removed at the end of the deprecation phase
(c.f., https://github.com/coq/coq/pull/15802 ) *)
let warn_future_coercion_class_constructor =
CWarnings.create ~name:"future-coercion-class-constructor" ~category:Deprecation.Version.v8_16
~default:CWarnings.AsError
Pp.(fun () -> str "'Class >' currently does nothing. Use 'Class' instead.")

(* deprecated in 8.17, to be removed at the end of the deprecation phase
(c.f., https://github.com/coq/coq/pull/16230 ) *)
let warn_future_coercion_class_field =
CWarnings.create ~name:"future-coercion-class-field" ~category:Deprecation.Version.v8_17
Pp.(fun definitional ->
strbrk "A coercion will be introduced instead of an instance in future versions when using ':>' in 'Class' declarations. "
++ strbrk "Replace ':>' with '::' (or use '#[global] Existing Instance field.' for compatibility with Coq < 8.18). Beware that the default locality for '::' is #[export], as opposed to #[global] for ':>' currently."
++ strbrk (if definitional then " Add an explicit #[global] attribute if you need to keep the current behavior. For example: \"Class foo := #[global] baz :: bar.\""
else " Add an explicit #[global] attribute to the field if you need to keep the current behavior. For example: \"Class foo := { #[global] field :: bar }.\""))

let check_proj_flags kind rf =
let open Vernacexpr in
let pf_coercion, pf_reversible =
match rf.rf_coercion with
(* replace "kind_class kind = NotClass" with true after deprecation phase *)
| AddCoercion -> kind_class kind = NotClass, Option.default true rf.rf_reversible
| AddCoercion -> true, Option.default true rf.rf_reversible
| NoCoercion ->
if rf.rf_reversible <> None then
Attributes.(unsupported_attributes
[CAst.make ("reversible (without :>)",VernacFlagEmpty)]);
false, false in
let pf_instance =
match rf.rf_instance with NoInstance -> false | BackInstance -> true
| BackInstanceWarning -> kind_class kind <> NotClass in
match rf.rf_instance with NoInstance -> false | BackInstance -> true in
let pf_priority = rf.rf_priority in
let pf_locality =
begin match rf.rf_coercion, rf.rf_instance with
Expand All @@ -699,19 +678,10 @@ let check_proj_flags kind rf =
[CAst.make ("export (without ::)",VernacFlagEmpty)])
| _ -> ()
end; rf.rf_locality in
(* remove following let after deprecation phase (started in 8.17,
c.f., https://github.com/coq/coq/pull/16230 ) *)
let pf_locality =
match rf.rf_instance, rf.rf_locality with
| BackInstanceWarning, Goptions.OptDefault -> Goptions.OptGlobal
| _ -> pf_locality in
let pf_canonical = rf.rf_canonical in
Data.{ pf_coercion; pf_reversible; pf_instance; pf_priority; pf_locality; pf_canonical }

(* remove the definitional argument at the end of the deprecation phase
(started in 8.17)
(c.f., https://github.com/coq/coq/pull/16230 ) *)
let pre_process_structure ?(definitional=false) udecl kind ~poly (records : Ast.t list) =
let pre_process_structure udecl kind ~poly (records : Ast.t list) =
let indlocs = check_unique_names records in
let () = check_priorities kind records in
let ps, data = extract_record_data records in
Expand All @@ -737,11 +707,6 @@ let pre_process_structure ?(definitional=false) udecl kind ~poly (records : Ast.
| Some n, _ -> n
in
let is_coercion = match is_coercion with AddCoercion -> true | NoCoercion -> false in
if kind_class kind <> NotClass then begin
if is_coercion then warn_future_coercion_class_constructor ();
if List.exists (function (_, Vernacexpr.{ rf_instance = BackInstanceWarning; _ }) -> true | _ -> false) cfs then
warn_future_coercion_class_field definitional
end;
{ Data.id = name.CAst.v; idbuild; rdata; is_coercion; proj_flags; inhabitant_id }
in
let data = List.map2 map data records in
Expand Down Expand Up @@ -883,7 +848,7 @@ let declare_class_constant ~univs paramimpls params data =
Classes.set_typeclass_transparency ~locality:Hints.SuperGlobal
[Evaluable.EvalConstRef cst] false;
let () =
declare_proj_coercion_instance ~flags:proj_flags (GlobRef.ConstRef proj_cst) cref ~with_coercion:false in
declare_proj_coercion_instance ~flags:proj_flags (GlobRef.ConstRef proj_cst) cref in
let m = {
meth_name = Name proj_name;
meth_info = None;
Expand Down Expand Up @@ -1017,16 +982,11 @@ let declare_existing_class g =
let definition_structure udecl kind ~template ~cumulative ~poly ~primitive_proj
finite (records : Ast.t list) : GlobRef.t list =
let impargs, params, univs, variances, projections_kind, data, indlocs =
let definitional = kind_class kind = DefClass in
pre_process_structure ~definitional udecl kind ~poly records
pre_process_structure udecl kind ~poly records
in
let inds, def = match kind_class kind with
| DefClass -> declare_class_constant ~univs impargs params data
| RecordClass | NotClass ->
(* remove the following block after deprecation phase
(started in 8.16, c.f., https://github.com/coq/coq/pull/15802 ) *)
let data = if kind_class kind = NotClass then data else
List.map (fun d -> { d with Data.is_coercion = false }) data in
let structure =
interp_structure_core
~cumulative finite ~univs ~variances ~primitive_proj
Expand Down
4 changes: 2 additions & 2 deletions vernac/vernacentries.ml
Original file line number Diff line number Diff line change
Expand Up @@ -937,7 +937,7 @@ let preprocess_inductive_decl ~atts kind indl =
user_err Pp.(str "Definitional classes do not support the \">\" syntax.");
let ((attr, rf_coercion, rf_instance), (lid, ce)) = l in
let rf_locality = match rf_coercion, rf_instance with
| AddCoercion, _ | _, (BackInstance | BackInstanceWarning) -> parse option_locality attr
| AddCoercion, _ | _, BackInstance -> parse option_locality attr
| _ -> let () = unsupported_attributes attr in Goptions.OptDefault in
let f = AssumExpr ((make ?loc:lid.loc @@ Name lid.v), [], ce),
{ rf_coercion ; rf_reversible = None ; rf_instance ; rf_priority = None ;
Expand Down Expand Up @@ -965,7 +965,7 @@ let preprocess_inductive_decl ~atts kind indl =
| AddCoercion -> reversible
| NoCoercion -> Notations.return None in
let loc = match f.rfu_coercion, f.rfu_instance with
| AddCoercion, _ | _, (BackInstance | BackInstanceWarning) -> option_locality
| AddCoercion, _ | _, BackInstance -> option_locality
| _ -> Notations.return Goptions.OptDefault in
Notations.(rev ++ loc ++ canonical_field) in
let (rf_reversible, rf_locality), rf_canonical = parse attr f.rfu_attrs in
Expand Down
5 changes: 1 addition & 4 deletions vernac/vernacexpr.mli
Original file line number Diff line number Diff line change
Expand Up @@ -109,10 +109,7 @@ type 'a search_restriction =

type verbose_flag = bool (* true = Verbose; false = Silent *)
type coercion_flag = AddCoercion | NoCoercion
(* Remove BackInstanceWarning at end of deprecation phase
(this is just to print a warning when :> is used instead of ::
to declare instances in classes) *)
type instance_flag = BackInstance | BackInstanceWarning | NoInstance
type instance_flag = BackInstance | NoInstance

type export_flag = Lib.export_flag = Export | Import

Expand Down

0 comments on commit 9b95a02

Please sign in to comment.