Skip to content

Commit 9b95a02

Browse files
committed
Uniformize semantics of :> in class after deprecation phase of #16230
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).
1 parent a6d226f commit 9b95a02

File tree

8 files changed

+18
-74
lines changed

8 files changed

+18
-74
lines changed

doc/sphinx/addendum/type-classes.rst

+4-14
Original file line numberDiff line numberDiff line change
@@ -248,7 +248,7 @@ the implicit arguments mechanism if available, as shown in the example.
248248
Substructures
249249
~~~~~~~~~~~~~
250250

251-
.. index:: :> (substructure)
251+
.. index:: :: (substructure)
252252

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

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

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

334333
This command has no effect when used on a typeclass.
335334

336-
.. _warn-future-coercion-class-field:
337-
338-
.. 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).
339-
340-
In future versions, :g:`:>` in the :n:`@record_definition` or
341-
:n:`@constructor` will declare a :ref:`coercion<coercions>`, as
342-
it does for other :cmd:`Record` commands. To eliminate the warning, use
343-
:g:`::` instead.
344-
345335
.. warn:: Ignored instance declaration for “@ident”: “@term” is not a class
346336

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

doc/sphinx/language/core/records.rst

-1
Original file line numberDiff line numberDiff line change
@@ -92,7 +92,6 @@ Defining record types
9292
:n:`:>`
9393
If specified, the field is declared as a coercion from the record name
9494
to the class of the field type. See :ref:`coercions`.
95-
Note that this currently does something else in :cmd:`Class` commands.
9695

9796
:n:`::`
9897
If specified, the field is declared a typeclass instance of the class

vernac/comInductive.ml

+1-2
Original file line numberDiff line numberDiff line change
@@ -671,8 +671,7 @@ let eq_params (up1,p1) (up2,p2) =
671671
let extract_coercions indl =
672672
let mkqid (_,({CAst.v=id},_)) = qualid_of_ident id in
673673
let iscoe (_, coe, inst) = match inst with
674-
(* remove BackInstanceWarning after deprecation phase *)
675-
| Vernacexpr.(NoInstance | BackInstanceWarning) -> coe = Vernacexpr.AddCoercion
674+
| Vernacexpr.NoInstance -> coe = Vernacexpr.AddCoercion
676675
| _ -> user_err (Pp.str "'::' not allowed in inductives.") in
677676
let extract lc = List.filter (fun (coe,_) -> iscoe coe) lc in
678677
List.map mkqid (List.flatten(List.map (fun (_,_,_,lc) -> extract lc) indl))

vernac/g_vernac.mlg

+2-2
Original file line numberDiff line numberDiff line change
@@ -576,8 +576,8 @@ GRAMMAR EXTEND Gram
576576
| ":" -> { NoCoercion } ] ]
577577
;
578578
of_type_inst:
579-
[ [ ":>" -> { (AddCoercion, BackInstanceWarning) (* replace with NoInstance at end of deprecation phase *) }
580-
| ":"; ">" -> { (AddCoercion, BackInstanceWarning) (* replace with NoInstance at end of deprecation phase *) }
579+
[ [ ":>" -> { (AddCoercion, NoInstance) }
580+
| ":"; ">" -> { (AddCoercion, NoInstance) }
581581
| "::" -> { (NoCoercion, BackInstance) }
582582
| "::>" -> { (AddCoercion, BackInstance) }
583583
| ":" -> { (NoCoercion, NoInstance) } ] ]

vernac/ppvernac.ml

-1
Original file line numberDiff line numberDiff line change
@@ -544,7 +544,6 @@ let pr_oc coe ins = match coe, ins with
544544
| AddCoercion, NoInstance -> str" :>"
545545
| NoCoercion, BackInstance -> str" ::"
546546
| AddCoercion, BackInstance -> str" ::>"
547-
| _, BackInstanceWarning -> str" :>" (* remove this line at end of deprecation phase *)
548547

549548
let pr_record_field (x, { rfu_attrs = attr ; rfu_coercion = coe ; rfu_instance = ins ; rfu_priority = pri ; rfu_notation = ntn }) =
550549
let prx = match x with

vernac/record.ml

+8-48
Original file line numberDiff line numberDiff line change
@@ -363,10 +363,8 @@ let instantiate_possibly_recursive_type ind u ntypes paramdecls fields =
363363

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

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

658-
(* deprecated in 8.16, to be removed at the end of the deprecation phase
659-
(c.f., https://github.com/coq/coq/pull/15802 ) *)
660-
let warn_future_coercion_class_constructor =
661-
CWarnings.create ~name:"future-coercion-class-constructor" ~category:Deprecation.Version.v8_16
662-
~default:CWarnings.AsError
663-
Pp.(fun () -> str "'Class >' currently does nothing. Use 'Class' instead.")
664-
665-
(* deprecated in 8.17, to be removed at the end of the deprecation phase
666-
(c.f., https://github.com/coq/coq/pull/16230 ) *)
667-
let warn_future_coercion_class_field =
668-
CWarnings.create ~name:"future-coercion-class-field" ~category:Deprecation.Version.v8_17
669-
Pp.(fun definitional ->
670-
strbrk "A coercion will be introduced instead of an instance in future versions when using ':>' in 'Class' declarations. "
671-
++ 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."
672-
++ strbrk (if definitional then " Add an explicit #[global] attribute if you need to keep the current behavior. For example: \"Class foo := #[global] baz :: bar.\""
673-
else " Add an explicit #[global] attribute to the field if you need to keep the current behavior. For example: \"Class foo := { #[global] field :: bar }.\""))
674-
675656
let check_proj_flags kind rf =
676657
let open Vernacexpr in
677658
let pf_coercion, pf_reversible =
678659
match rf.rf_coercion with
679-
(* replace "kind_class kind = NotClass" with true after deprecation phase *)
680-
| AddCoercion -> kind_class kind = NotClass, Option.default true rf.rf_reversible
660+
| AddCoercion -> true, Option.default true rf.rf_reversible
681661
| NoCoercion ->
682662
if rf.rf_reversible <> None then
683663
Attributes.(unsupported_attributes
684664
[CAst.make ("reversible (without :>)",VernacFlagEmpty)]);
685665
false, false in
686666
let pf_instance =
687-
match rf.rf_instance with NoInstance -> false | BackInstance -> true
688-
| BackInstanceWarning -> kind_class kind <> NotClass in
667+
match rf.rf_instance with NoInstance -> false | BackInstance -> true in
689668
let pf_priority = rf.rf_priority in
690669
let pf_locality =
691670
begin match rf.rf_coercion, rf.rf_instance with
@@ -699,19 +678,10 @@ let check_proj_flags kind rf =
699678
[CAst.make ("export (without ::)",VernacFlagEmpty)])
700679
| _ -> ()
701680
end; rf.rf_locality in
702-
(* remove following let after deprecation phase (started in 8.17,
703-
c.f., https://github.com/coq/coq/pull/16230 ) *)
704-
let pf_locality =
705-
match rf.rf_instance, rf.rf_locality with
706-
| BackInstanceWarning, Goptions.OptDefault -> Goptions.OptGlobal
707-
| _ -> pf_locality in
708681
let pf_canonical = rf.rf_canonical in
709682
Data.{ pf_coercion; pf_reversible; pf_instance; pf_priority; pf_locality; pf_canonical }
710683

711-
(* remove the definitional argument at the end of the deprecation phase
712-
(started in 8.17)
713-
(c.f., https://github.com/coq/coq/pull/16230 ) *)
714-
let pre_process_structure ?(definitional=false) udecl kind ~poly (records : Ast.t list) =
684+
let pre_process_structure udecl kind ~poly (records : Ast.t list) =
715685
let indlocs = check_unique_names records in
716686
let () = check_priorities kind records in
717687
let ps, data = extract_record_data records in
@@ -737,11 +707,6 @@ let pre_process_structure ?(definitional=false) udecl kind ~poly (records : Ast.
737707
| Some n, _ -> n
738708
in
739709
let is_coercion = match is_coercion with AddCoercion -> true | NoCoercion -> false in
740-
if kind_class kind <> NotClass then begin
741-
if is_coercion then warn_future_coercion_class_constructor ();
742-
if List.exists (function (_, Vernacexpr.{ rf_instance = BackInstanceWarning; _ }) -> true | _ -> false) cfs then
743-
warn_future_coercion_class_field definitional
744-
end;
745710
{ Data.id = name.CAst.v; idbuild; rdata; is_coercion; proj_flags; inhabitant_id }
746711
in
747712
let data = List.map2 map data records in
@@ -883,7 +848,7 @@ let declare_class_constant ~univs paramimpls params data =
883848
Classes.set_typeclass_transparency ~locality:Hints.SuperGlobal
884849
[Evaluable.EvalConstRef cst] false;
885850
let () =
886-
declare_proj_coercion_instance ~flags:proj_flags (GlobRef.ConstRef proj_cst) cref ~with_coercion:false in
851+
declare_proj_coercion_instance ~flags:proj_flags (GlobRef.ConstRef proj_cst) cref in
887852
let m = {
888853
meth_name = Name proj_name;
889854
meth_info = None;
@@ -1017,16 +982,11 @@ let declare_existing_class g =
1017982
let definition_structure udecl kind ~template ~cumulative ~poly ~primitive_proj
1018983
finite (records : Ast.t list) : GlobRef.t list =
1019984
let impargs, params, univs, variances, projections_kind, data, indlocs =
1020-
let definitional = kind_class kind = DefClass in
1021-
pre_process_structure ~definitional udecl kind ~poly records
985+
pre_process_structure udecl kind ~poly records
1022986
in
1023987
let inds, def = match kind_class kind with
1024988
| DefClass -> declare_class_constant ~univs impargs params data
1025989
| RecordClass | NotClass ->
1026-
(* remove the following block after deprecation phase
1027-
(started in 8.16, c.f., https://github.com/coq/coq/pull/15802 ) *)
1028-
let data = if kind_class kind = NotClass then data else
1029-
List.map (fun d -> { d with Data.is_coercion = false }) data in
1030990
let structure =
1031991
interp_structure_core
1032992
~cumulative finite ~univs ~variances ~primitive_proj

vernac/vernacentries.ml

+2-2
Original file line numberDiff line numberDiff line change
@@ -937,7 +937,7 @@ let preprocess_inductive_decl ~atts kind indl =
937937
user_err Pp.(str "Definitional classes do not support the \">\" syntax.");
938938
let ((attr, rf_coercion, rf_instance), (lid, ce)) = l in
939939
let rf_locality = match rf_coercion, rf_instance with
940-
| AddCoercion, _ | _, (BackInstance | BackInstanceWarning) -> parse option_locality attr
940+
| AddCoercion, _ | _, BackInstance -> parse option_locality attr
941941
| _ -> let () = unsupported_attributes attr in Goptions.OptDefault in
942942
let f = AssumExpr ((make ?loc:lid.loc @@ Name lid.v), [], ce),
943943
{ rf_coercion ; rf_reversible = None ; rf_instance ; rf_priority = None ;
@@ -965,7 +965,7 @@ let preprocess_inductive_decl ~atts kind indl =
965965
| AddCoercion -> reversible
966966
| NoCoercion -> Notations.return None in
967967
let loc = match f.rfu_coercion, f.rfu_instance with
968-
| AddCoercion, _ | _, (BackInstance | BackInstanceWarning) -> option_locality
968+
| AddCoercion, _ | _, BackInstance -> option_locality
969969
| _ -> Notations.return Goptions.OptDefault in
970970
Notations.(rev ++ loc ++ canonical_field) in
971971
let (rf_reversible, rf_locality), rf_canonical = parse attr f.rfu_attrs in

vernac/vernacexpr.mli

+1-4
Original file line numberDiff line numberDiff line change
@@ -109,10 +109,7 @@ type 'a search_restriction =
109109

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

117114
type export_flag = Lib.export_flag = Export | Import
118115

0 commit comments

Comments
 (0)