aboutsummaryrefslogtreecommitdiffhomepage
path: root/doc/refman/Universes.tex
blob: cd8222269d9cefbfb9f219d1fc1a82a6f8093482 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
\achapter{Polymorphic Universes}
\aauthor{Matthieu Sozeau}

\label{Universes-full}
\index{Universes!presentation}

\asection{General Presentation}

\begin{flushleft}
  \em The status of Universe Polymorphism is experimental. Some features
  are not compatible with it (yet): bytecode compilation does not handle
  polymorphic definitions, it treats them as opaque constants.
\end{flushleft}

This section describes the universe polymorphic extension of Coq.
Universe polymorphism allows writing generic definitions making use of
universes and reuse them at different and sometimes incompatible levels.

A standard example of the difference between universe \emph{polymorphic} and
\emph{monomorphic} definitions is given by the identity function:

\begin{coq_example*}
Definition identity {A : Type} (a : A) := a.
\end{coq_example*}

By default, constant declarations are monomorphic, hence the identity
function declares a global universe (say \texttt{Top.1}) for its
domain. Subsequently, if we try to self-apply the identity, we will get
an error:

\begin{coq_eval}
Set Printing Universes.
\end{coq_eval}
\begin{coq_example}
Fail Definition selfid := identity (@identity).
\end{coq_example}

Indeed, the global level \texttt{Top.1} would have to be strictly smaller than itself
for this self-application to typecheck, as the type of \texttt{(@identity)} is
\texttt{forall (A : Type@{Top.1}), A -> A} whose type is itself \texttt{Type@{Top.1+1}}.

A universe polymorphic identity function binds its domain universe level
at the definition level instead of making it global.

\begin{coq_example}
Polymorphic Definition pidentity {A : Type} (a : A) := a.
About pidentity.
\end{coq_example}

It is then possible to reuse the constant at different levels, like so:

\begin{coq_example}
Definition selfpid := pidentity (@pidentity).
\end{coq_example}

Of course, the two instances of \texttt{pidentity} in this definition
are different. This can be seen when \texttt{Set Printing Universes} is
on:

\begin{coq_example}
Print selfpid.
\end{coq_example}

Now \texttt{pidentity} is used at two different levels: at the head of
the application it is instantiated at \texttt{Top.3} while in the
argument position it is instantiated at \texttt{Top.4}. This definition
is only valid as long as \texttt{Top.4} is strictly smaller than
\texttt{Top.3}, as show by the constraints. Not that this definition is
monomorphic (not universe polymorphic), so in turn the two universes are
actually global levels.

Inductive types can also be declared universes polymorphic, on universes
appearing in their parameters or fields. A typical example is given by
monoids:

\begin{coq_example}
Polymorphic Record Monoid := { mon_car :> Type; mon_unit : mon_car; 
  mon_op : mon_car -> mon_car -> mon_car }.
Print Monoid.
\end{coq_example}

The \texttt{Monoid}'s carrier universe is polymorphic, hence it is
possible to instantiate it for example with \texttt{Monoid} itself.
First we build the trivial unit monoid, in \texttt{Set}:
\begin{coq_example}
Definition unit_monoid : Monoid := 
  {| mon_car := unit; mon_unit := tt; mon_op x y := tt |}.
\end{coq_example} 

From this we can build a definition for the monoid of
\texttt{Set}-monoids (where multiplication would be given by the product
of monoids).

\begin{coq_example*}
Polymorphic Definition monoid_monoid : Monoid.
  refine (@Build_Monoid Monoid unit_monoid (fun x y => x)).
Defined.
\end{coq_example*}
\begin{coq_example}
Print monoid_monoid.
\end{coq_example} 

As one can see from the constraints, this monoid is ``large'', it lives
in a universe strictly higher than \texttt{Set}.

\asection{\tt Polymorphic, Monomorphic}
\comindex{Polymorphic}
\comindex{Monomorphic}
\optindex{Universe Polymorphism}

As shown in the examples, polymorphic definitions and inductives can be
declared using the \texttt{Polymorphic} prefix. There also exists an
option \texttt{Set Universe Polymorphism} which will implicitly prepend
it to any definition of the user. In that case, to make a definition
producing global universe constraints, one can use the
\texttt{Monomorphic} prefix. Many other commands support the
\texttt{Polymorphic} flag, including:

\begin{itemize}
\item \texttt{Lemma}, \texttt{Axiom}, and all the other ``definition''
  keywords support polymorphism.
\item \texttt{Variables}, \texttt{Context} in a section support polymorphism.
  This means that the
  variables are discharged polymorphically over definitions that use
  them. In other words, two definitions in the section sharing a common
  variable will both get parameterized by the universes produced by the 
  variable declaration. This is in contrast to a ``mononorphic'' variable
  which introduces global universes, making the two definitions depend on
  the \emph{same} global universes associated to the variable.
\item \texttt{Hint \{Resolve, Rewrite\}} will use the auto/rewrite hint
  polymorphically, not at a single instance.
\end{itemize}

\asection{Conversion and unification}

The semantics of conversion and unification have to be modified a little
to account for the new universe instance arguments to polymorphic
references. The semantics respect the fact that definitions are
transparent, so indistinguishable from their bodies during conversion.

This is accomplished by changing one rule of unification, the
first-order approximation rule, which applies when two applicative terms
with the same head are compared. It tries to short-cut unfolding by
comparing the arguments directly. In case the constant is universe
polymorphic, we allow this rule to fire only when unifying the universes
results in instantiating a so-called flexible universe variables (not
given by the user). Similarly for conversion, if such an equation of
applicative terms fail due to a universe comparison not being satisfied,
the terms are unfolded. This change implies that conversion and
unification can have different unfolding behaviors on the same
development with universe polymorphism switched on or off.

\asection{Minimization}

Universe polymorphism with cumulativity tends to generate many useless
inclusion constraints in general. Typically at each application of a
polymorphic constant $f$, if an argument has expected type
\verb|Type@{i}| and is given a term of type \verb|Type@{j}|, a $j \le i$
constraint will be generated. It is however often the case that an
equation $j = i$ would be more appropriate, when $f$'s
universes are fresh for example. Consider the following example:

\begin{coq_eval}
Set Printing Universes.
\end{coq_eval}
\begin{coq_example}
Definition id0 := @pidentity nat 0.
Print id0.
\end{coq_example}

This definition is elaborated by minimizing the universe of id to level
\Set~while the more general definition would keep the fresh level i
generated at the application of id and a constraint that $\Set \le i$.
This minimization process is applied only to fresh universe
variables. It simply adds an equation between the variable and its lower
bound if it is an atomic universe (i.e. not an algebraic max()).

\asection{Explicit Universes}

\begin{flushleft}
  \em The design and implementation of explicit universes is very
  experimental and is likely to change in future versions.
\end{flushleft}

The syntax has been extended to allow users to explicitly bind names to
universes and explicitly instantiate polymorphic
definitions. Currently, binding is implicit at the first occurrence of a 
universe name. For example, i and j below are introduced by the
annotations attached to Types. 

\begin{coq_example*}
Polymorphic Definition le (A : Type@{i}) : Type@{j} := A.
\end{coq_example*}
\begin{coq_example}
Print le.
\end{coq_example}

During refinement we find that $j$ must be larger or equal than $i$, as
we are using $A : Type@{i} <= Type@{j}$, hence the generated
constraint. Note that the names here are not bound in the final
definition, they just allow to specify locally what relations should
hold. In the term and in general in proof mode, universe names
introduced in the types can be referred to in terms.

Definitions can also be instantiated explicitly, giving their full instance:
\begin{coq_example}
Check (pidentity@{Set}).
Check (le@{i j}).
\end{coq_example}

User-named universes are considered rigid for unification and are never
minimized.

Finally, two commands allow to name \emph{global} universes and constraints.

\subsection{\tt Universe {\ident}.
  \comindex{Universe}
  \label{UniverseCmd}}

This command declare a new global universe named {\ident}.

\subsection{\tt Constraint {\ident} {\textit{ord}} {\ident}.
  \comindex{Constraint}
  \label{ConstraintCmd}}

This command declare a new constraint between named universes. 
The order relation can be one of $<$, $<=$ or $=$. If consistent, 
the constraint is then enforced in the global environment.

\begin{ErrMsgs}
\item \errindex{Undeclared universe {\ident}}.
\item \errindex{Universe inconsistency}
\end{ErrMsgs}

%%% Local Variables: 
%%% mode: latex
%%% TeX-master: "Reference-Manual"
%%% End: