Skip to content

Commit

Permalink
Doc for sort poly inductives
Browse files Browse the repository at this point in the history
  • Loading branch information
SkySkimmer committed Nov 21, 2023
1 parent ac1662b commit be1dff5
Show file tree
Hide file tree
Showing 2 changed files with 54 additions and 0 deletions.
1 change: 1 addition & 0 deletions doc/changelog/01-kernel/17836-sort-poly.rst
Expand Up @@ -2,4 +2,5 @@
:ref:`sort-polymorphism` makes it possible to share common constructs
over `Type` `Prop` and `SProp`
(`#17836 <https://github.com/coq/coq/pull/17836>`_,
`#18331 <https://github.com/coq/coq/pull/18331>`_,
by Gaëtan Gilbert).
53 changes: 53 additions & 0 deletions doc/sphinx/addendum/universe-polymorphism.rst
Expand Up @@ -803,6 +803,59 @@ witness these temporary variables.
`α` followed by a number as printing will not distinguish between
your bound variables and temporary variables.

Sort polymorphic inductives may be declared when every instantiation
is valid. Elimination at a given universe instance requires that
elimination is allowed at every ground instantiation of the sort
variables in the instance. These restrictions ignore
:flag:`Definitional UIP`.

For instance

.. coqtop:: all reset

Set Universe Polymorphism.

Inductive Squash@{s|u|} (A:Type@{s|u}) : Prop := squash (_:A).

Elimination to `Prop` and `SProp` is always allowed, so `Squash_ind`
and `Squash_sind` are automatically defined.

Elimination to `Type` is not allowed with variable `s`, because the
instantiation `s := Type` does not allow elimination to `Type`.

However elimination to `Type` or to a polymorphic sort with `s := Prop` is allowed:

.. coqtop:: all

Definition Squash_Prop_rect A (P:Squash@{Prop|_} A -> Type)
(H:forall x, P (squash _ x))
: forall s, P s
:= fun s => match s with squash _ x => H x end.

Definition Squash_Prop_srect@{s|u +|} A (P:Squash@{Prop|_} A -> Type@{s|u})
(H:forall x, P (squash _ x))
: forall s, P s
:= fun s => match s with squash _ x => H x end.

.. note::

Since inductive types in `SProp` with one or more constructors have
elimination restricted to `SProp`, inductive types in a polymorphic
sort have the same restriction. Additionally for uniformity of the
implementation this restriction is not removed by setting
:flag:`Definitional UIP`.

Consequently sort polymorphic containers such as sigma types are
better defined as primitive records (which do not have this
restriction) when possible.

.. coqtop:: all

Set Primitive Projections.
Record sigma@{s|u v|} (A:Type@{s|u}) (B:A -> Type@{s|v})
: Type@{s|max(u,v)}
:= pair { pr1 : A; pr2 : B Rpr1 }.

.. _universe-polymorphism-in-sections:

Universe polymorphism and sections
Expand Down

0 comments on commit be1dff5

Please sign in to comment.