aboutsummaryrefslogtreecommitdiffhomepage
path: root/doc/RefMan-cic.tex
diff options
context:
space:
mode:
Diffstat (limited to 'doc/RefMan-cic.tex')
-rwxr-xr-xdoc/RefMan-cic.tex1233
1 files changed, 1233 insertions, 0 deletions
diff --git a/doc/RefMan-cic.tex b/doc/RefMan-cic.tex
new file mode 100755
index 000000000..51bbc506a
--- /dev/null
+++ b/doc/RefMan-cic.tex
@@ -0,0 +1,1233 @@
+\chapter{The Calculus of Inductive Constructions}
+\label{Cic}
+\index{Cic@\textsc{CIC}}
+\index{Calculus of (Co)Inductive Constructions}
+
+The underlying formal language of {\Coq} is the {\em Calculus of
+(Co)Inductive Constructions}\index{Calculus of (Co)Inductive Constructions}
+(\CIC\ in short). It is presented in this chapter.
+
+In \CIC\, all objects have a {\em type}. There are types for functions (or
+programs), there are atomic types (especially datatypes)... but also
+types for proofs and types for the types themselves.
+Especially, any object handled in the formalism must belong to a
+type. For instance, the statement {\it ``for all x, P''} is not
+allowed in type theory; you must say instead: {\it ``for all x
+belonging to T, P''}. The expression {\it ``x belonging to T''} is
+written {\it ``x:T''}. One also says: {\it ``x has type T''}.
+The terms of {\CIC} are detailed in section \ref{Terms}.
+
+In \CIC\, there is an internal reduction mechanism. In particular, it
+allows to decide if two programs are {\em intentionally} equal (one
+says {\em convertible}). Convertibility is presented in section
+\ref{convertibility}.
+
+The remaining sections are concerned with the type-checking of terms.
+The beginner can skip them.
+
+The reader seeking a background on the {\CIC} may read several
+papers. Giménez~\cite{Gim98}
+provides an introduction to inductive and coinductive
+definitions in Coq, Werner~\cite{Wer94} and Paulin-Mohring~\cite{Moh97} are the
+most recent theses on the
+{\CIC}. Coquand-Huet~\cite{CoHu85a,CoHu85b,CoHu86} introduces the
+Calculus of Constructions. Coquand-Paulin~\cite{CoPa89} introduces
+inductive definitions. The {\CIC} is a formulation of type theory
+including the possibility of inductive
+constructions. Barendregt~\cite{Bar91} studies the modern form of type
+theory.
+
+\section{The terms}\label{Terms}
+
+In most type theories, one usually makes a syntactic distinction
+between types and terms. This is not the case for \CIC\ which defines
+both types and terms in the same syntactical structure. This is
+because the type-theory itself forces terms and types to be defined in
+a mutual recursive way and also because similar constructions can be
+applied to both terms and types and consequently can share the same
+syntactic structure.
+
+For instance the type of functions will have several meanings. Assume
+\nat\ is the type of natural numbers then $\nat\ra\nat$ is the type of
+functions from \nat\ to \nat, $\nat \ra \Prop$ is the type of unary
+predicates over the natural numbers. For instance $[x:\nat](x=x)$ will
+represent a predicate $P$, informally written in mathematics
+$P(x)\equiv x=x$. If $P$ has type $\nat \ra \Prop$, $(P~x)$ is a
+proposition, furthermore $(x:\nat)(P~x)$ will represent the type of
+functions which associate to each natural number $n$ an object of
+type $(P~n)$ and consequently represent proofs of the formula
+``$\forall x.P(x)$''.
+
+\subsection{Sorts}\label{Sorts}
+\index{Sorts}
+Types are seen as terms of the language and then should belong to
+another type. The type of a type is always a constant of the language
+called a sort.
+
+The two basic sorts in the language of \CIC\ are \Set\ and \Prop.
+
+The sort \Prop\ intends to be the type of logical propositions. If
+$M$ is a logical proposition then it denotes a class, namely the class
+of terms representing proofs of $M$. An object $m$ belonging to $M$
+witnesses the fact that $M$ is true. An object of type \Prop\ is
+called a {\em proposition}.
+
+The sort \Set\ intends to be the type of specifications. This includes
+programs and the usual sets such as booleans, naturals, lists
+etc.
+
+These sorts themselves can be manipulated as ordinary terms.
+Consequently sorts also should be given a type. Because assuming
+simply that \Set\ has type \Set\ leads to an inconsistent theory, we
+have infinitely many sorts in the language of \CIC\ . These are, in
+addition to \Set\ and \Prop\, a hierarchy of universes \Type$(i)$
+for any integer $i$. We call \Sort\ the set of sorts
+which is defined by:
+\[\Sort \equiv \{\Prop,\Set,\Type(i)| i \in \NN\} \]
+\index{Type@{\Type}}
+\index{Prop@{\Prop}}
+\index{Set@{\Set}}
+The sorts enjoy the following properties: {\Prop:\Type(0)} and
+ {\Type$(i)$:\Type$(i+1)$}.
+
+The user will never mention explicitly the index $i$ when referring to
+the universe \Type$(i)$. One only writes \Type. The
+system itself generates for each instance of \Type\ a new
+index for the universe and checks that the constraints between these
+indexes can be solved. From the user point of view we consequently
+have {\sf Type :Type}.
+
+We shall make precise in the typing rules the constraints between the
+indexes.
+
+\medskip
+\Rem
+The extraction mechanism is not compatible with this universe
+hierarchy. It is supposed to work only on terms which are explicitly
+typed in the Calculus of Constructions without universes and with
+Inductive Definitions at the Set level and only a small elimination.
+In other cases, extraction may generate a dummy answer and sometimes
+failed. To avoid failure when developing proofs, an error while
+extracting the computational contents of a proof will not stop the
+proof but only give a warning.
+
+
+\subsection{Constants}
+Besides the sorts, the language also contains constants denoting
+objects in the environment. These constants may denote previously
+defined objects but also objects related to inductive definitions
+(either the type itself or one of its constructors or destructors).
+
+\medskip\noindent {\bf Remark. } In other presentations of \CIC,
+the inductive objects are not seen as
+external declarations but as first-class terms. Usually the
+definitions are also completely ignored. This is a nice theoretical
+point of view but not so practical. An inductive definition is
+specified by a possibly huge set of declarations, clearly we want to
+share this specification among the various inductive objects and not
+to duplicate it. So the specification should exist somewhere and the
+various objects should refer to it. We choose one more level of
+indirection where the objects are just represented as constants and
+the environment gives the information on the kind of object the
+constant refers to.
+
+\medskip
+Our inductive objects will be manipulated as constants declared in the
+environment. This roughly corresponds to the way they are actually
+implemented in the \Coq\ system. It is simple to map this presentation
+in a theory where inductive objects are represented by terms.
+
+\subsection{Language}
+
+\paragraph{Types.}
+
+Roughly speaking types can be separated into atomic and composed
+types.
+
+An atomic type of the {\em Calculus of Inductive Constructions} is
+either a sort or is built from a type variable or an inductive
+definition applied to some terms.
+
+A composed type will be a product $(x:T)U$ with $T$ and $U$ two types.
+
+\paragraph{Terms.}
+A term is either a type or a term variable or a term constant of the
+environment.
+
+As usual in $\lambda$-calculus, we combine objects using abstraction
+and application.
+
+More precisely the language of the {\em Calculus of Inductive
+ Constructions} is built with the following rules:
+
+\begin{enumerate}
+\item the sorts {\sf Set, Prop, Type} are terms.
+\item constants of the environment are terms.
+\item variables are terms.
+\item if $x$ is a variable and $T$, $U$ are terms then $(x:T)U$ is a
+ term. If $x$ occurs in $U$, $(x:T)U$ reads as {\it ``for all x of
+ type T, U''}. As $U$ depends on $x$, one says that $(x:T)U$ is a
+ {\em dependent product}. If $x$ doesn't occurs in $U$ then $(x:T)U$
+ reads as {\it ``if T then U''}. A non dependent product can be
+ written: $T \rightarrow U$.
+\item if $x$ is a variable and $T$, $U$ are terms then $[x:T]U$ is a
+ term. This is a notation for the $\lambda$-abstraction of
+ $\lambda$-calculus\index{lambda-calculus@$\lambda$-calculus}
+ \cite{Bar81}. The term $[x:T]U$ is a function which maps elements of
+ $T$ to $U$.
+\item if $T$ and $U$ are terms then $(T\ U)$ is a term. The term $(T\
+ U)$ reads as {\it ``T applied to U''}.
+\end{enumerate}
+
+\paragraph{Notations.} Application associates to the left such that
+$(t~t_1\ldots t_n)$ represents $(\ldots (t~t_1)\ldots t_n)$. The
+products and arrows associate to the right such that $(x:A)B\ra C\ra
+D$ represents $(x:A)(B\ra (C\ra D))$. One uses sometimes $(x,y:A)B$ or
+$[x,y:A]B$ to denote the abstraction or product of several variables
+of the same type. The equivalent formulation is $(x:A)(y:A)B$ or
+$[x:A][y:A]B$
+\paragraph{Free variables.}
+The notion of free variables is defined as usual. In the expressions
+$[x:T]U$ and $(x:T)U$ the occurrences of $x$ in $U$ are bound. They
+are represented by de Bruijn indexes in the internal structure of
+terms.
+
+\paragraph{Substitution.} \index{Substitution}
+The notion of substituting a term $T$ to free occurrences of a
+variable $x$ in a term $U$ is defined as usual. The resulting term
+will be written $\subst{U}{x}{T}$.
+
+
+\section{Typed terms}\label{Typed-terms}
+As objects of type theory, terms are subjected to {\em type
+ discipline}. The well typing of a term depends on a set of
+declarations of variables we call a {\em context}. A context $\Gamma$
+is written $[x_1:T_1;..; x_n:T_n]$ where the $x_i$'s are distinct
+variables and the $T_i$'s are terms. If $\Gamma$ contains some $x:T$,
+we write $(x:T)\in\Gamma$ and also $x \in\Gamma$. Contexts must be
+themselves {\em well formed}. The notation $\Gamma::(y:T)$ denotes
+the context $[x_1:T_1;..;x_n:T_n;y:T]$. The notation $[]$ denotes the
+empty context.
+\index{Context}
+
+We define the inclusion of two contexts $\Gamma$ and $\Delta$ (written
+as $\Gamma \subset \Delta$) as the property, for all variable $x$ and
+type $T$, if $(x:T) \in \Gamma$ then $(x:T)\in \Delta$. We write
+$|\Delta|$ for the length of the context $\Delta$ which is $n$ if
+$\Delta$ is $[x_1:T_1;..; x_n:T_n]$.
+
+A variable $x$ is said to be free in $\Gamma$ if $\Gamma$ contains a
+declaration $y:T$ such that $x$ is free in $T$.
+
+\paragraph{Environment.}\index{Environment}
+Because we are manipulating constants, we also need to consider an
+environment $E$. We shall give afterwards the rules for introducing
+new objects in the environment. For the typing relation of terms, it
+is enough to introduce two notions. One which says if a name is
+defined in the environment we shall write $c \in E$ and the other one
+which gives the type of this constant in $E$. We shall write $(c : T)
+\in E$.
+
+In the following, we assume $E$ is a valid environment. We define
+simultaneously two judgments. The first one \WTEG{t}{T} means the
+term $t$ is well-typed and has type $T$ in the environment $E$ and
+context $\Gamma$. The second judgment \WFE{\Gamma} means that the
+environment $E$ is well-formed and the context $\Gamma$ is a valid
+context in this environment. It also means a third property which
+makes sure that any constant in $E$ was defined in an environment
+which is included in $\Gamma$
+\footnote{This requirement could be relaxed if we instead introduced
+ an explicit mechanism for instantiating constants. At the external
+ level, the Coq engine works accordingly to this view that all the
+ definitions in the environment were built in a sub-context of the
+ current context.}.
+
+A term $t$ is well typed in an environment $E$ iff there exists a
+context $\Gamma$ and a term $T$ such that the judgment \WTEG{t}{T} can
+be derived from the following rules.
+\begin{itemize}\label{Typing-rules}\index{Typing rules}
+\item [W-E] \inference{\WF{[]}{[]}}
+\item [W-$s$]
+\inference{\frac{\WTEG{T}{s}~~~~s\in \Sort~~~~x \not\in
+ \Gamma \cup E}{\WFE{\Gamma::(x:T)}}}
+\item [Ax] \index{Typing rules!Ax}
+\inference{\frac{\WFE{\Gamma}}{\WTEG{\Prop}{\Type(p)}}~~~~~
+\frac{\WFE{\Gamma}}{\WTEG{\Set}{\Type(q)}}}\\[3mm]
+\inference{\frac{\WFE{\Gamma}~~~~i<j}{\WTEG{\Type(i)}{\Type(j)}}}
+\item[Var]\index{Typing rules!Var}
+ \inference{\frac{ \WFE{\Gamma}~~~~~(x:T)\in\Gamma}{\WTEG{x}{T}}}
+\item[Const] \index{Typing rules!Const}
+\inference{\frac{\WFE{\Gamma}~~~~(c:T) \in E}{\WTEG{c}{T}}}
+\item [Prod] \index{Typing rules!Prod}
+\inference{\frac{\WTEG{T}{s_1}~~~~
+ \WTE{\Gamma::(x:T)}{U}{s_2}~~~s_1\in\{\Prop, \Set\}~\mbox{or}~
+ s_2\in\{\Prop, \Set\}}
+ { \WTEG{(x:T)U}{s_2}}} \\[3mm]
+\inference{\frac{\WTEG{T}{\Type(i)}~~~~
+ \WTE{\Gamma::(x:T)}{U}{\Type(j)}~~~i\leq
+ k~~~j \leq k}{ \WTEG{(x:T)U}{\Type(k)}}}
+\item [Lam]\index{Typing rules!Lam}
+\inference{\frac{\WTEG{(x:T)U}{s}~~~~ \WTE{\Gamma::(x:T)}{t}{U}}
+ {\WTEG{[x:T]t}{(x:T)U}}}
+\item [App]\index{Typing rules!App}
+ \inference{\frac{\WTEG{t}{(x:U)T}~~~~\WTEG{u}{U}}
+ {\WTEG{(t\ u)}{\subst{T}{x}{u}}}}
+\end{itemize}
+
+\section{Conversion rules}
+\index{Conversion rules}
+\label{conv-rules}
+\paragraph{$\beta$-reduction.}
+\label{beta}\index{beta-reduction@$\beta$-reduction}
+
+We want to be able to identify some terms as we can identify the
+application of a function to a given argument with its result. For
+instance the identity function over a given type $T$ can be written
+$[x:T]x$. We want to identify any object $a$ (of type $T$) with the
+application $([x:T]x~a)$. We define for this a {\em reduction} (or a
+{\em conversion}) rule we call $\beta$:
+\[ ([x:T]t~u) \triangleright_{\beta} \subst{t}{x}{u} \]
+
+We say that $\subst{t}{x}{u}$ is the {\em $\beta$-contraction} of
+$([x:T]t~u)$ and, conversely, that $([x:T]t~u)$ is the {\em
+ $\beta$-expansion} of $\subst{t}{x}{u}$.
+
+According to $\beta$-reduction, terms of the {\em Calculus of
+ Inductive Constructions} enjoy some fundamental properties such as
+confluence, strong normalization, subject reduction. These results are
+theoretically of great importance but we will not detail them here and
+refer the interested reader to \cite{Coq85}.
+
+\paragraph{$\iota$-reduction.}
+\label{iota}\index{iota-reduction@$\iota$-reduction}
+A specific conversion rule is associated to the inductive objects in
+the environment. We shall give later on (section \ref{iotared}) the
+precise rules but it just says that a destructor applied to an object
+built from a constructor behaves as expected. This reduction is
+called $\iota$-reduction and is more precisely studied in
+\cite{Moh93,Wer94}.
+
+
+\paragraph{$\delta$-reduction.}
+\label{convertibility}
+\label{delta}\index{delta-reduction@$\delta$-reduction}
+In the environment we also have constants representing abbreviations
+for terms. It is legal to identify a constant with its value. This
+reduction will be precised in section \ref{deltared} where we define
+well-formed environments. This reduction will be called
+$\delta$-reduction.
+
+\paragraph{Convertibility.}
+\index{beta-conversion@$\beta$-conversion}\index{iota-conversion@$\iota$-conversion}\index{delta-conversion@$\delta$-conversion}
+
+
+Let us write $t \triangleright u$ for the relation $t$ reduces to $u$
+with one of the previous reduction $\beta,\iota$ or $\delta$.
+
+We say that two terms $t_1$ and $t_2$ are {\em convertible} (or {\em
+ equivalent)} iff there exists a term $u$ such that $t_1
+\triangleright \ldots \triangleright u$ and $t_2 \triangleright \ldots
+\triangleright u$. We note $t_1 \convert t_2$.
+
+The convertibility relation allows to introduce a new typing rule
+which says that two convertible well-formed types have the same
+inhabitants.
+
+At the moment, we did not take into account one rule between universes
+which says that any term in a universe of index $i$ is also a term in
+the universe of index $i+1$. This property is included into the
+conversion rule by extending the equivalence relation of
+convertibility into an order inductively defined by:
+\begin{enumerate}
+\item if $M\convert N$ then $M\leconvert N$,
+\item if $i \leq j$ then $\Type(i)\leconvert\Type(j)$,
+\item if $T \convert U$ and $M\leconvert N$ then $(x:T)M\leconvert
+ (x:U)N$.
+\end{enumerate}
+
+The conversion rule is now exactly:
+
+\begin{itemize}\label{Conv}
+\item [Conv]\index{Typing rules!Conv}
+ \inference{
+ \frac{\WTEG{U}{S}~~~~\WTEG{t}{T}~~~~T \leconvert U}{\WTEG{t}{U}}}
+\end{itemize}
+
+
+\paragraph{$\eta$-conversion.}\label{eta}\index{eta-conversion@$\eta$-conversion}\index{eta-reduction@$\eta$-reduction}
+An other important rule is the $\eta$-conversion. It is to identify
+terms over a dummy abstraction of a variable followed by an
+application of this variable. Let $T$ be a type, $t$ be a term in
+which the variable $x$ doesn't occurs free. We have
+\[ [x:T](t\ x) \triangleright t \]
+Indeed, as $x$ doesn't occurs free in $t$, for any $u$ one
+applies to $[x:T](t\ x)$, it $\beta$-reduces to $(t\ u)$. So
+$[x:T](t\ x)$ and $t$ can be identified.
+
+\Rem The $\eta$-reduction is not taken into account in the
+convertibility rule of \Coq.
+
+\paragraph{Normal form.}\index{Normal form}\label{Normal-form}\label{Head-normal-form}\index{Head normal form}
+A term which cannot be any more reduced is said to be in {\em normal
+ form}. There are several ways (or strategies) to apply the reduction
+rule. Among them, we have to mention the {\em head reduction} which
+will play an important role (see chapter \ref{Tactics}). Any term can
+be written as $[x_1:T_1]\ldots[x_k:T_k](t_0\ t_1\ldots t_n)$ where
+$t_0$ is not an application. We say then that $t_0$ is the {\em head
+ of $t$}. If we assume that $t_0$ is $[x:T]u_0$ then one step of
+$\beta$-head reduction of $t$ is:
+\[[x_1:T_1]\ldots[x_k:T_k]([x:T]u_0\ t_1\ldots t_n)
+~\triangleright ~ [x_1:T_1]\ldots[x_k:T_k](\subst{u_0}{x}{t_1}\ t_2 \ldots t_n)\]
+Iterating the process of head reduction until the head of the reduced
+term is no more an abstraction leads to the {\em $\beta$-head normal
+ form} of $t$:
+\[ t \triangleright \ldots \triangleright [x_1:T_1]\ldots[x_k:T_k](v\ u_1
+\ldots u_m)\]
+where $v$ is not an abstraction (nor an application). Note that the
+head normal form must not be confused with the normal form since some
+$u_i$ can be reducible.
+
+Similar notions of head-normal forms involving $\delta$ and $\iota$
+reductions or any combination of those can also be defined.
+
+\section{Definitions in environments}\label{Cic-definitions}
+We now give the rules for manipulating objects in the environment.
+Because a constant can depend on previously introduced constants, the
+environment will be an ordered list of declarations. When specifying
+an inductive definition, several objects will be introduced at the
+same time. So any object in the environment will define one or more
+constants.
+
+In this presentation we introduce two different sorts of objects in
+the environment. The first one is ordinary definitions which give a
+name to a particular well-formed term, the second one is inductive
+definitions which introduce new inductive objects.
+
+\subsection{Rules for definitions}
+
+\paragraph{Adding a new definition.}
+The simplest objects in the environment are definitions which can be
+seen as one possible mechanism for abbreviation.
+
+A definition will be represented in the environment as
+\Def{\Gamma}{c}{t}{T} which means that $c$ is a constant which is
+valid in the context $\Gamma$ whose value is $t$ and type is $T$.
+
+\paragraph{$\delta$-reduction.} \label{deltared}
+If \Def{\Gamma}{c}{t}{T} is in the environment $E$ then in this
+environment the $\delta$-reduc\-tion $c \triangleright_{\delta} t$ is
+introduced.
+
+The rule for adding a new definition is simple:
+
+\begin{description}
+\item[Def] \inference{\frac{\WTEG{t}{T}~~~c \notin E\cup \Gamma}
+ {\WF{E;\Def{\Gamma}{c}{t}{T}}{\Gamma}}}
+\end{description}
+
+\subsection{Derived rules}
+
+From the original rules of the type system, one can derive new rules
+which change the context of definition of objects in the environment.
+Because these rules correspond to elementary operations in the \Coq\
+engine used in the discharge mechanism at the end of a section, we
+state them explicitly.
+
+\paragraph{Mechanism of substitution.}
+
+One rule which can be proved valid, is to replace a term $c$ by its
+value in the environment. As we defined the substitution of a term for
+a variable in a term, one can define the substitution of a term for a
+constant. One easily extends this substitution to contexts and
+environments.
+
+\paragraph{Substitution Property:}
+\inference{\frac{\WF{E;\Def{\Gamma}{c}{t}{T}; F}{\Delta}}
+ {\WF{E; \subst{F}{c}{t}}{\subst{\Delta}{c}{t}}}}
+
+
+\paragraph{Abstraction.}
+
+One can modify the context of definition of a constant $c$ by
+abstracting a constant with respect to the last variable $x$ of its
+defining context. For doing that, we need to check that the constants
+appearing in the body of the declaration do not depend on $x$, we need
+also to modify the reference to the constant $c$ in the environment
+and context by explicitly applying this constant to the variable $x$.
+Because of the rules for building environments and terms we know the
+variable $x$ is available at each stage where $c$ is mentioned.
+
+\paragraph{Abstracting property:}
+ \inference{\frac{\WF{E; \Def{\Gamma::(x:U)}{c}{t}{T};
+ F}{\Delta}~~~~\WFE{\Gamma}}
+ {\WF{E;\Def{\Gamma}{c}{[x:U]t}{(x:U)T};
+ \subst{F}{c}{(c~x)}}{\subst{\Delta}{c}{(c~x)}}}}
+
+\paragraph{Pruning the context.}
+We said the judgment \WFE{\Gamma} means that the defining contexts of
+constants in $E$ are included in $\Gamma$. If one abstracts or
+substitutes the constants with the above rules then it may happen
+that the context $\Gamma$ is now bigger than the one needed for
+defining the constants in $E$. Because defining contexts are growing
+in $E$, the minimum context needed for defining the constants in $E$
+is the same as the one for the last constant. One can consequently
+derive the following property.
+
+\paragraph{Pruning property:}
+\inference{\frac{\WF{E; \Def{\Delta}{c}{t}{T}}{\Gamma}}
+ {\WF{E;\Def{\Delta}{c}{t}{T}}{\Delta}}}
+
+
+\section{Inductive Definitions}\label{Cic-inductive-definitions}
+
+A (possibly mutual) inductive definition is specified by giving the
+names and the type of the inductive sets or families to be
+defined and the names and types of the constructors of the inductive
+predicates. An inductive declaration in the environment can
+consequently be represented with two contexts (one for inductive
+definitions, one for constructors).
+
+Stating the rules for inductive definitions in their general form
+needs quite tedious definitions. We shall try to give a concrete
+understanding of the rules by precising them on running examples. We
+take as examples the type of natural numbers, the type of
+parameterized lists over a type $A$, the relation which state that
+a list has some given length and the mutual inductive definition of trees and
+forests.
+
+\subsection{Representing an inductive definition}
+\subsubsection{Inductive definitions without parameters}
+As for constants, inductive definitions can be defined in a non-empty
+context. \\
+We write \NInd{\Gamma}{\Gamma_I}{\Gamma_C} an inductive
+definition valid in a context $\Gamma$, a
+context of definitions $\Gamma_I$ and a context of constructors
+$\Gamma_C$.
+\paragraph{Examples.}
+The inductive declaration for the type of natural numbers will be:
+\[\NInd{}{\nat:\Set}{\nO:\nat,\nS:\nat\ra\nat}\]
+In a context with a variable $A:\Set$, the lists of elements in $A$ is
+represented by:
+\[\NInd{A:\Set}{\List:\Set}{\Nil:\List,\cons : A \ra \List \ra
+ \List}\]
+ Assuming
+ $\Gamma_I$ is $[I_1:A_1;\ldots;I_k:A_k]$, and $\Gamma_C$ is
+ $[c_1:C_1;\ldots;c_n:C_n]$, the general typing rules are:
+
+\bigskip
+\inference{\frac{\NInd{\Gamma}{\Gamma_I}{\Gamma_C} \in E
+ ~~j=1\ldots k}{(I_j:A_j) \in E}}
+
+\inference{\frac{\NInd{\Gamma}{\Gamma_I}{\Gamma_C} \in E
+ ~~~~i=1.. n}
+ {(c_i:\subst{C_i}{I_j}{I_j}_{j=1\ldots k})\in E}}
+
+\subsubsection{Inductive definitions with parameters}
+
+We have to slightly complicate the representation above in order to handle
+the delicate problem of parameters.
+Let us explain that on the example of \List. As they were defined
+above, the type \List\ can only be used in an environment where we
+have a variable $A:\Set$. Generally one want to consider lists of
+elements in different types. For constants this is easily done by abstracting
+the value over the parameter. In the case of inductive definitions we
+have to handle the abstraction over several objects.
+
+One possible way to do that would be to define the type \List\
+inductively as being an inductive family of type $\Set\ra\Set$:
+\[\NInd{}{\List:\Set\ra\Set}{\Nil:(A:\Set)(\List~A),\cons : (A:\Set)A
+ \ra (\List~A) \ra (\List~A)}\]
+There are drawbacks to this point of view. The
+information which says that $(\List~\nat)$ is an inductively defined
+\Set\ has been lost.
+%\footnote{
+%The interested reader may look at the compare the above definition with the two
+%following ones which have very different logical meaning:\\
+%$\NInd{}{\List:\Set}{\Nil:\List,\cons : (A:\Set)A
+% \ra \List \ra \List}$ \\
+%$\NInd{}{\List:\Set\ra\Set}{\Nil:(A:\Set)(\List~A),\cons : (A:\Set)A
+% \ra (\List~A\ra A) \ra (\List~A)}$.}
+
+In the system, we keep track in the syntax of the context of
+parameters. The idea of these parameters is that they can be
+instantiated and still we have an inductive definition for which we
+know the specification.
+
+Formally the representation of an inductive declaration
+will be
+\Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C} for an inductive
+definition valid in a context $\Gamma$ with parameters $\Gamma_P$, a
+context of definitions $\Gamma_I$ and a context of constructors
+$\Gamma_C$.
+The occurrences of the variables of $\Gamma_P$ in the contexts
+$\Gamma_I$ and $\Gamma_C$ are bound.
+
+The definition \Ind{\Gamma}{\Gamma_P}{\;\Gamma_I}{\Gamma_C\;} will be
+well-formed exactly when \NInd{\Gamma,\Gamma_P}{\;\Gamma_I}{\Gamma_C\;} is.
+If $\Gamma_P$ is $[p_1:P_1;\ldots;p_r:P_r]$, an object in
+\Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C} applied to $q_1,\ldots,q_r$
+will behave as the corresponding object of
+\NInd{\Gamma}{\substs{\Gamma_I}{p_i}{q_i}{i=1..r}}{\substs{\Gamma_C}{p_i}{q_i}{i=1..r}}.
+
+\paragraph{Examples}
+The declaration for parameterized lists is:
+\[\Ind{}{A:\Set}{\List:\Set}{\Nil:\List,\cons : A \ra \List \ra
+ \List}\]
+
+The declaration for the length of lists is:
+\[\Ind{}{A:\Set}{\Length:(\List~A)\ra \nat\ra\Prop}
+ {\LNil:(\Length~(\Nil~A)~\nO),\\
+ \LCons :(a:A)(l:(\List~A))(n:\nat)(\Length~l~n)\ra (\Length~(\cons~A~a~l)~(\nS~n))}\]
+
+The declaration for a mutual inductive definition of forests and trees is:
+\[\NInd{[]}{\tree:\Set,\forest:\Set}
+ {\node:\forest \ra \tree,
+ \emptyf:\forest,\consf:\tree \ra \forest \ra \forest\-}\]
+
+These representations are the ones obtained as the result of the \Coq\
+declaration:
+\begin{coq_example*}
+Inductive Set nat := O : nat | S : nat -> nat.
+Inductive list [A : Set] : Set :=
+ nil : (list A) | cons : A -> (list A) -> (list A).
+\end{coq_example*}
+\begin{coq_example*}
+Inductive Length [A:Set] : (list A) -> nat -> Prop :=
+ Lnil : (Length A (nil A) O)
+ | Lcons : (a:A)(l:(list A))(n:nat)
+ (Length A l n)->(Length A (cons A a l) (S n)).
+Mutual Inductive tree : Set := node : forest -> tree
+with forest : Set := emptyf : forest | consf : tree -> forest -> forest.
+\end{coq_example*}
+The inductive declaration in \Coq\ is slightly different from the one
+we described theoretically. The difference is that in the type of
+constructors the inductive definition is explicitly applied to the
+parameters variables. The \Coq\ type-checker verifies that all
+parameters are applied in the correct manner in each recursive call.
+In particular, the following definition will not be accepted because
+there is an occurrence of \List\ which is not applied to the parameter
+variable:
+\begin{coq_example}
+Inductive list [A : Set] : Set :=
+ nil : (list A) | cons : A -> (list A->A) -> (list A).
+\end{coq_example}
+
+\subsection{Types of inductive objects}
+We have to give the type of constants in an environment $E$ which
+contains an inductive declaration.
+
+\begin{description}
+\item[Ind-Const] Assuming $\Gamma_P$ is $[p_1:P_1;\ldots;p_r:P_r]$,
+ $\Gamma_I$ is $[I_1:A_1;\ldots;I_k:A_k]$, and $\Gamma_C$ is
+ $[c_1:C_1;\ldots;c_n:C_n]$,
+
+\inference{\frac{\Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C} \in E
+ ~~j=1\ldots k}{(I_j:(p_1:P_1)\ldots(p_r:P_r)A_j) \in E}}
+
+\inference{\frac{\Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C} \in E
+ ~~~~i=1.. n}
+ {(c_i:(p_1:P_1)\ldots(p_r:P_r)\subst{C_i}{I_j}{(I_j~p_1\ldots
+ p_r)}_{j=1\ldots k})\in E}}
+\end{description}
+
+\paragraph{Example.}
+We have $(\List:\Set \ra \Set), (\cons:(A:\Set)A\ra(\List~A)\ra
+(\List~A))$, \\
+$(\Length:(A:\Set)(\List~A)\ra\nat\ra\Prop)$, $\tree:\Set$ and $\forest:\Set$.
+
+From now on, we write $\ListA$ instead of $(\List~A)$ and $\LengthA$
+for $(\Length~A)$.
+
+%\paragraph{Parameters.}
+%%The parameters introduce a distortion between the inside specification
+%%of the inductive declaration where parameters are supposed to be
+%%instantiated (this representation is appropriate for checking the
+%%correctness or deriving the destructor principle) and the outside
+%%typing rules where the inductive objects are seen as objects
+%%abstracted with respect to the parameters.
+
+%In the definition of \List\ or \Length\, $A$ is a parameter because
+%what is effectively inductively defined is $\ListA$ or $\LengthA$ for
+%a given $A$ which is constant in the type of constructors. But when
+%we define $(\LengthA~l~n)$, $l$ and $n$ are not parameters because the
+%constructors manipulate different instances of this family.
+
+\subsection{Well-formed inductive definitions}
+We cannot accept any inductive declaration because some of them lead
+to inconsistent systems. We restrict ourselves to definitions which
+satisfy a syntactic criterion of positivity. Before giving the formal
+rules, we need a few definitions:
+
+\paragraph{Definitions}\index{Positivity}\label{Positivity}
+
+A type $T$ is an {\em arity of sort $s$}\index{Arity}
+if it converts to the sort $s$ or to a
+product $(x:T)U$ with $U$ an arity of sort $s$. (For instance $A\ra
+\Set$ or $(A:\Prop)A\ra \Prop$ are arities of sort respectively \Set\
+and \Prop).
+A {\em type of constructor of $I$}\index{Type of constructor}
+ is either a term $(I~t_1\ldots ~t_n)$ or
+$(x:T)C$ with $C$ a {\em type of constructor of $I$}.
+
+\smallskip
+
+The type of constructor $T$ will be said to {\em satisfy the positivity
+condition} for a constant $X$ in the following cases:
+
+\begin{itemize}
+\item $T=(X~t_1\ldots ~t_n)$ and $X$ does not occur free in
+any $t_i$
+\item $T=(x:T)U$ and $X$ occurs only strictly positively in $T$ and
+the type $U$ satisfies the positivity condition for $X$
+\end{itemize}
+
+The constant $X$ {\em occurs strictly positively} in $T$ in the
+following cases:
+
+\begin{itemize}
+\item $X$ does not occur in $T$
+\item $T$ converts to $(X~t_1 \ldots ~t_n)$ and $X$ does not occur in any of $t_i$
+\item $T$ converts to $(x:U)V$ and $X$ does not occur in type $U$ but occurs strictly
+positively in type $V$
+\item $T$ converts to $(I~a_1 \ldots ~a_m ~ t_1 \ldots ~t_p)$ where $I$ is the name of an
+inductive declaration of the form $\Ind{\Gamma}{[p_1:P_1;\ldots;p_m:P_m]}{[I:A]}{[c_1:C_1;\ldots;c_n:C_n]}$
+(in particular, it is not mutually defined and it has $m$ parameters)
+and $X$ does not occur in any of the $t_i$, and
+the types of constructor $C_i\{p_j/a_j\}_{j=1\ldots m}$ of $I$
+satisfy the imbricated positivity condition for $X$
+%\item more generally, when $T$ is not a type, $X$ occurs strictly
+%positively in $T[x:U]u$ if $X$ does not occur in $U$ but occurs
+%strictly positively in $u$
+\end{itemize}
+
+The type constructor $T$ of $I$ {\em satisfies the imbricated
+positivity condition} for a constant $X$ in the following
+cases:
+
+\begin{itemize}
+\item $T=(I~t_1\ldots ~t_n)$ and $X$ does not occur in
+any $t_i$
+\item $T=(x:T)U$ and $X$ occurs only strictly positively in $T$ and
+the type $U$ satisfies the imbricated positivity condition for $X$
+\end{itemize}
+
+\paragraph{Example}
+$X$ occurs strictly positively in $A\ra X$ or $X*A$ or $({\tt list}
+X)$ but not in $X \ra A$ or $(X \ra A)\ra A$ assuming the notion of
+product and lists were already defined. Assuming $X$ has arity ${\tt
+nat \ra Prop}$ and {\tt ex} is inductively defined existential
+quantifier, the occurrence of $X$ in ${\tt (ex~ nat~ [n:nat](X~ n))}$ is
+also strictly positive.\\
+
+\paragraph{Correctness rules.}
+We shall now describe the rules allowing the introduction of a new
+inductive definition.
+
+\begin{description}
+\item[W-Ind] Let $E$ be an environment and
+ $\Gamma,\Gamma_P,\Gamma_I,\Gamma_C$ are contexts such that
+ $\Gamma_I$ is $[I_1:A_1;\ldots;I_k:A_k]$ and $\Gamma_C$ is
+ $[c_1:C_1;\ldots;c_n:C_n]$. \\[3mm]
+\inference{
+ \frac{
+ (\WTE{\Gamma;\Gamma_P}{A_j}{s'_j})_{j=1\ldots k}
+ ~~ (\WTE{\Gamma;\Gamma_P;\Gamma_I}{C_i}{s_{p_i}})_{i=1\ldots n}
+}
+ {\WF{E;\Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C}}{\Gamma}}}\\
+providing the following side conditions hold:
+\begin{itemize}
+\item $k>0$, $I_j$, $c_i$ are different names for $j=1\ldots k$ and $i=1\ldots n$,
+\item for $j=1\ldots k$ we have $A_j$ is an arity of sort $s_j$ and $I_j
+ \notin \Gamma \cup E$,
+\item for $i=1\ldots n$ we have $C_i$ is a type of constructor of
+ $I_{p_i}$ which satisfies the positivity condition for $I_1 \ldots I_k$
+ and $c_i \notin \Gamma \cup E$.
+\end{itemize}
+\end{description}
+One can remark that there is a constraint between the sort of the
+arity of the inductive type and the sort of the type of its
+constructors which will always be satisfied for impredicative sorts
+(\Prop\ or \Set) but may generate constraints between universes.
+
+
+%We shall assume for the following definitions that, if necessary, we
+%annotated the type of constructors such that we know if the argument
+%is recursive or not. We shall write the type $(x:_R T)C$ if it is
+%a recursive argument and $(x:_P T)C$ if the argument is not recursive.
+
+\subsection{Destructors}
+The specification of inductive definitions with arities and
+constructors is quite natural. But we still have to say how to use an
+object in an inductive type.
+
+This problem is rather delicate. There are actually several different
+ways to do that. Some of them are logically equivalent but not always
+equivalent from the computational point of view or from the user point
+of view.
+
+From the computational point of view, we want to be able to define a
+function whose domain is an inductively defined type by using a
+combination of case analysis over the possible constructors of the
+object and recursion.
+
+Because we need to keep a consistent theory and also we prefer to keep
+a strongly normalising reduction, we cannot accept any sort of
+recursion (even terminating). So the basic idea is to restrict
+ourselves to primitive recursive functions and functionals.
+
+For instance, assuming a parameter $A:\Set$ exists in the context, we
+want to build a function \length\ of type $\ListA\ra \nat$ which
+computes the length of the list, so such that $(\length~\Nil) = \nO$
+and $(\length~(\cons~A~a~l)) = (\nS~(\length~l))$. We want these
+equalities to be recognized implicitly and taken into account in the
+conversion rule.
+
+From the logical point of view, we have built a type family by giving
+a set of constructors. We want to capture the fact that we do not
+have any other way to build an object in this type. So when trying to
+prove a property $(P~m)$ for $m$ in an inductive definition it is
+enough to enumerate all the cases where $m$ starts with a different
+constructor.
+
+In case the inductive definition is effectively a recursive one, we
+want to capture the extra property that we have built the smallest
+fixed point of this recursive equation. This says that we are only
+manipulating finite objects. This analysis provides induction
+principles.
+
+For instance, in order to prove $(l:\ListA)(\LengthA~l~(\length~l))$
+it is enough to prove:
+
+\noindent $(\LengthA~\Nil~(\length~\Nil))$ and
+
+\smallskip
+$(a:A)(l:\ListA)(\LengthA~l~(\length~l)) \ra
+(\LengthA~(\cons~A~a~l)~(\length~(\cons~A~a~l)))$.
+\smallskip
+
+\noindent which given the conversion equalities satisfied by \length\ is the
+same as proving:
+$(\LengthA~\Nil~\nO)$ and $(a:A)(l:\ListA)(\LengthA~l~(\length~l)) \ra
+(\LengthA~(\cons~A~a~l)~(\nS~(\length~l)))$.
+
+One conceptually simple way to do that, following the basic scheme
+proposed by Martin-L\"of in his Intuitionistic Type Theory, is to
+introduce for each inductive definition an elimination operator. At
+the logical level it is a proof of the usual induction principle and
+at the computational level it implements a generic operator for doing
+primitive recursion over the structure.
+
+But this operator is rather tedious to implement and use. We choose in
+this version of Coq to factorize the operator for primitive recursion
+into two more primitive operations as was first suggested by Th. Coquand
+in~\cite{Coq92}. One is the definition by case
+analysis. The second one is a definition by guarded fixpoints.
+
+\subsubsection{The {\tt Cases\ldots of \ldots end} construction.}
+\label{Caseexpr}
+\index{Cases@{\tt Cases\ldots of\ldots end}}
+
+The basic idea of this destructor operation is that we have an object
+$m$ in an inductive type $I$ and we want to prove a property $(P~m)$
+which in general depends on $m$. For this, it is enough to prove the
+property for $m = (c_i~u_1\ldots u_{p_i})$ for each constructor of $I$.
+
+This proof will be denoted by a generic term:
+\[\Case{P}{m}{(c_1~x_{11}~...~x_{1p_1}) \mbox{\tt =>} f_1 \ldots
+(c_n~x_{n1}~|~...~|~x_{np_n}) \mbox{\tt =>} f_n }\]
+In this expression, if $m$ is a term built from a constructor
+$(c_i~u_1\ldots u_{p_i})$ then the expression will behave as it is specified
+with $i$-th branch and will reduce to $f_i$ where the $x_{i1}$\ldots $x_{ip_i}$ are
+replaced by the $u_1\ldots u_p$ according to
+the $\iota$-reduction.\\
+
+This is the basic idea which is generalized to the case where $I$ is
+an inductively defined $n$-ary relation (in which case the property
+$P$ to be proved will be a $n+1$-ary relation).
+
+\paragraph{Non-dependent elimination.}
+When defining a function by case analysis, we build an object of type $I
+\ra C$ and the minimality principle on an inductively defined logical
+predicate of type $A \ra \Prop$ is often used to prove a property
+$(x:A)(I~x)\ra (C~x)$. This is a particular case of the dependent
+principle that we stated before with a predicate which does not depend
+explicitly on the object in the inductive definition.
+
+For instance, a function testing whether a list is empty
+can be
+defined as:
+
+\[[l:\ListA]\Case{[H:\ListA]\bool}{l}{\Nil~ \mbox{\tt =>}~\true~ |~ (\cons~a~m)~ \mbox{\tt =>}~\false}\]
+\noindent {\bf Remark. } In the system \Coq\ the expression above, can be
+written without mentioning
+the dummy abstraction:
+\Case{\bool}{l}{\Nil~ \mbox{\tt =>}~\true~ |~ (\cons~a~m)~
+ \mbox{\tt =>}~ \false}
+
+\paragraph{Allowed elimination sorts.}
+\index{Elimination sorts}
+
+An important question for building the typing rule for {\tt Case} is
+what can be the type of $P$ with respect to the type of the inductive
+definitions.
+
+Remembering that the elimination builds an object in $(P~m)$ from an
+object in $m$ in type $I$ it is clear that we cannot allow any combination.
+
+For instance we cannot in general have $I$ has type \Prop\
+and $P$ has type $I\ra \Set$, because it will mean to build an informative
+proof of type $(P~m)$ doing a case analysis over a non-computational
+object that will disappear in the extracted program.
+But the other way is safe
+with respect to our interpretation we can have $I$ a computational
+object and $P$ a non-computational one, it just corresponds to proving a
+logical property of a computational object.
+
+Also if $I$ is in one of the sorts \{\Prop, \Set\}, one cannot in
+general allow an elimination over a bigger sort such as \Type.
+But this operation is safe whenever $I$ is a {\em small
+ inductive} type, which means that all the types of constructors of
+$I$ are small with the following definition:\\
+$(I~t_1\ldots t_s)$ is a {\em small type of constructor} and $(x:T)C$ is a
+small type of constructor if $C$ is and if $T$ has type \Prop\ or
+\Set.
+\index{Small inductive type}
+
+We call this particular elimination which gives the possibility to
+compute a type by induction on the structure of a term, a {\em strong
+ elimination}\index{Strong elimination}.
+
+
+We define now a relation \compat{I:A}{B} between an inductive
+definition $I$ of type $A$, an arity $B$ which says that an object in
+the inductive definition $I$ can be eliminated for proving a property
+$P$ of type $B$.
+
+The \compat{I:A}{B} is defined as the smallest relation satisfying the
+following rules:
+
+\begin{description}
+\item[Prod] \inference{\frac{\compat{(I~x):A'}{B'}}
+ {\compat{I:(x:A)A'}{(x:A)B'}}}
+\item[\Prop] \inference{\compat{I:\Prop}{I\ra\Prop}~~~~~
+ \frac{I \mbox{~is a singleton definition}}{\compat{I:\Prop}{I\ra \Set}}}
+
+\item[\Set] \inference{\frac{s \in
+ \{\Prop, \Set\}}{\compat{I:\Set}{I\ra s}}
+~~~~\frac{I \mbox{~is a small inductive definition}~~~~s \in
+ \{\Type(i)\}}
+ {\compat{I:\Set}{I\ra s}}}
+\item[\Type] \inference{\frac{
+ s \in \{\Prop,\Set,\Type(j)\}}{\compat{I:\Type(i)}{I\ra s}}}
+\end{description}
+
+\paragraph{Notations.}
+We write \compat{I}{B} for \compat{I:A}{B} where $A$ is the type of
+$I$.
+
+%\paragraph{Warning: strong elimination}
+%\index{Elimination!Strong elimination}
+%In previous versions of Coq, for a small inductive definition, only the
+%non-informative strong elimination on \Type\ was allowed, because
+%strong elimination on \Typeset\ was not compatible with the current
+%extraction procedure. In this version, strong elimination on \Typeset\
+%is accepted but a dummy element is extracted from it and may generate
+%problems if extracted terms are explicitly used such as in the
+%{\tt Program} tactic or when extracting ML programs.
+
+\paragraph{Singleton elimination}
+\index{Elimination!Singleton elimination}
+
+A {\em singleton definition} has always an informative content,
+even if it is a proposition.
+
+A {\em singleton
+definition} has only one constructor and all the argument of this
+constructor are non informative. In that case, there is a canonical
+way to interpret the informative extraction on an object in that type,
+such that the elimination on sort $s$ is legal. Typical examples are
+the conjunction of non-informative propositions and the equality. In
+that case, the term \verb!eq_rec! which was defined as an axiom, is
+now a term of the calculus.
+\begin{coq_example}
+Print eq_rec.
+Extraction eq_rec.
+\end{coq_example}
+
+\paragraph{Type of branches.}
+Let $c$ be a term of type $C$, we assume $C$ is a type of constructor
+for an inductive definition $I$. Let $P$ be a term that represents the
+property to be proved.
+We assume $r$ is the number of parameters.
+
+We define a new type \CI{c:C}{P} which represents the type of the
+branch corresponding to the $c:C$ constructor.
+\[
+\begin{array}{ll}
+\CI{c:(I_i~p_1\ldots p_r\ t_1 \ldots t_p)}{P} &\equiv (P~t_1\ldots ~t_p~c) \\[2mm]
+\CI{c:(x: T)C}{P} &\equiv (x:T)\CI{(c~x):C}{P}
+\end{array}
+\]
+We write \CI{c}{P} for \CI{c:C}{P} with $C$ the type of $c$.
+
+\paragraph{Examples.}
+For $\ListA$ the type of $P$ will be $\ListA\ra s$ for $s \in
+\{\Prop,\Set,\Type(i)\}$. \\
+$ \CI{(\cons~A)}{P} \equiv
+(a:A)(l:\ListA)(P~(\cons~A~a~l))$.
+
+For $\LengthA$, the type of $P$ will be
+$(l:\ListA)(n:\nat)(\LengthA~l~n)\ra \Prop$ and the expression
+\CI{(\LCons~A)}{P} is defined as:\\
+$(a:A)(l:\ListA)(n:\nat)(h:(\LengthA~l~n))(P~(\cons~A~a~l)~(\nS~n)~(\LCons~A~a~l~n~l))$.\\
+If $P$ does not depend on its third argument, we find the more natural
+expression:\\
+$(a:A)(l:\ListA)(n:\nat)(\LengthA~l~n)\ra(P~(\cons~A~a~l)~(\nS~n))$.
+
+\paragraph{Typing rule.}
+
+Our very general destructor for inductive definition enjoys the
+following typing rule (we write {\Case{P}{c}{[x_{11}:T_{11}]\ldots[x_{1p_1}:T_{1p_1}]f_1\ldots [x_{n1}:T_{n1}]\ldots[x_{np_n}:T_{np_n}f_n} for \Case{P}{m}{(c_1~x_{11}~...~x_{1p_1}) \mbox{\tt =>} f_1 \ldots
+(c_n~x_{n1}~|~...~|~x_{np_n}) \mbox{\tt =>} f_n }}):
+
+\begin{description}
+\item[Case] \label{elimdep} \index{Typing rules!Case}
+\inference{
+\frac{\WTEG{c}{(I~q_1\ldots q_r~t_1\ldots t_s)}~~
+ \WTEG{P}{B}~~\compat{(I~q_1\ldots q_r)}{B}
+ ~~
+(\WTEG{f_i}{\CI{(c_{p_i}~q_1\ldots q_r)}{P}})_{i=1\ldots l}}
+{\WTEG{\Case{P}{c}{f_1\ldots f_l}}{(P\ t_1\ldots t_s\ c)}}}%\\[3mm]
+
+provided $I$ is an inductive type in a declaration
+\Ind{\Delta}{\Gamma_P}{\Gamma_I}{\Gamma_C} with $|\Gamma_P| = r$,
+$\Gamma_C = [c_1:C_1;\ldots;c_n:C_n]$ and $c_{p_1}\ldots c_{p_l}$ are the
+only constructors of $I$.
+\end{description}
+
+\paragraph{Example.}
+For \List\ and \Length\ the typing rules for the {\tt Case} expression
+are (writing just $t:M$ instead of \WTEG{t}{M}, the environment and
+context being the same in all the judgments).
+
+\[\frac{l:\ListA~~P:\ListA\ra s~~~f_1:(P~(\Nil~A))~~
+ f_2:(a:A)(l:\ListA)(P~(\cons~A~a~l))}
+ {\Case{P}{l}{f_1~f_2}:(P~l)}\]
+
+\[\frac{
+ \begin{array}[b]{c}
+H:(\LengthA~L~N) \\ P:(l:\ListA)(n:\nat)(\LengthA~l~n)\ra
+ \Prop\\
+ f_1:(P~(\Nil~A)~\nO~\LNil) \\
+ f_2:(a:A)(l:\ListA)(n:\nat)(h:(\LengthA~l~n))(P~(\cons~A~a~n)~(\nS~n)~(\LCons~A~a~l~n~h))
+ \end{array}}
+ {\Case{P}{H}{f_1~f_2}:(P~L~N~H)}\]
+
+\paragraph{Definition of $\iota$-reduction.}\label{iotared}
+\index{iota-reduction@$\iota$-reduction}
+We still have to define the $\iota$-reduction in the general case.
+
+A $\iota$-redex is a term of the following form:
+\[\Case{P}{(c_{p_i}~q_1\ldots q_r~a_1\ldots a_m)}{f_1\ldots
+ f_l}\]
+with $c_{p_i}$ the $i$-th constructor of the inductive type $I$ with $r$
+parameters.
+
+The $\iota$-contraction of this term is $(f_i~a_1\ldots a_m)$ leading
+to the general reduction rule:
+\[ \Case{P}{(c_{p_i}~q_1\ldots q_r~a_1\ldots a_m)}{f_1\ldots
+ f_n} \triangleright_{\iota} (f_i~a_1\ldots a_m) \]
+
+\subsection{Fixpoint definitions}
+\label{Fix-term} \index{Fix@{\tt Fix}}
+The second operator for elimination is fixpoint definition.
+This fixpoint may involve several mutually recursive definitions.
+The basic syntax for a recursive set of declarations is
+\[\Fix{}{f_1:A_1:=t_1 \ldots f_n:A_n:=t_n}\]
+The terms are obtained by projections from this set of declarations
+and are written $\Fix{f_i}{f_1:A_1:=t_1 \ldots f_n:A_n:=t_n}$
+
+\subsubsection{Typing rule}
+The typing rule is the expected one for a fixpoint.
+
+\begin{description}
+\item[Fix] \index{Typing rules!Fix}
+\inference{\frac{(\WTEG{A_i}{s_i})_{i=1\ldots n}~~~~
+ (\WTE{\Gamma,f_1:A_1,\ldots,f_n:A_n}{t_i}{A_i})_{i=1\ldots n}}
+ {\WTEG{\Fix{f_i}{f_1:A_1:=t_1 \ldots f_n:A_n:=t_n}}{A_i}}}
+\end{description}
+
+Any fixpoint definition cannot be accepted because non-normalizing terms
+will lead to proofs of absurdity.
+
+The basic scheme of recursion that should be allowed is the one needed for
+defining primitive
+recursive functionals. In that case the fixpoint enjoys special
+syntactic restriction, namely one of the arguments belongs to an
+inductive type, the function starts with a case analysis and recursive
+calls are done on variables coming from patterns and representing subterms.
+
+For instance in the case of natural numbers, a proof of the induction
+principle of type
+\[(P:\nat\ra\Prop)(P~\nO)\ra((n:\nat)(P~n)\ra(P~(\nS~n)))\ra(n:\nat)(P~n)\]
+can be represented by the term:
+\[\begin{array}{l}
+[P:\nat\ra\Prop][f:(P~\nO)][g:(n:\nat)(P~n)\ra(P~(\nS~n))]\\
+\Fix{h}{h:(n:\nat)(P~n):=[n:\nat]\Case{P}{n}{f~[p:\nat](g~p~(h~p))}}
+\end{array}
+\]
+
+Before accepting a fixpoint definition as being correctly typed, we
+check that the definition is ``guarded''. A precise analysis of this
+notion can be found in~\cite{Gim94}.
+
+The first stage is to precise on which argument the fixpoint will be
+decreasing. The type of this argument should be an inductive
+definition.
+
+For doing this the syntax of fixpoints is extended and becomes
+ \[\Fix{f_i}{f_1/k_1:A_1:=t_1 \ldots f_n/k_n:A_n:=t_n}\]
+where $k_i$ are positive integers.
+Each $A_i$ should be a type (reducible to a term) starting with at least
+$k_i$ products $(y_1:B_1)\ldots (y_{k_i}:B_{k_i}) A'_i$ and $B_{k_i}$
+being an instance of an inductive definition.
+
+Now in the definition $t_i$, if $f_j$ occurs then it should be applied
+to at least $k_j$ arguments and the $k_j$-th argument should be
+syntactically recognized as structurally smaller than $y_{k_i}$
+
+
+The definition of being structurally smaller is a bit technical.
+One needs first to define the notion of
+{\em recursive arguments of a constructor}\index{Recursive arguments}.
+For an inductive definition \Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C},
+the type of a constructor $c$ have the form
+$(p_1:P_1)\ldots(p_r:P_r)(x_1:T_1)\ldots (x_r:T_r)(I_j~p_1\ldots
+p_r~t_1\ldots t_s)$ the recursive arguments will correspond to $T_i$ in
+which one of the $I_l$ occurs.
+
+
+The main rules for being structurally smaller are the following:\\
+Given a variable $y$ of type an inductive
+definition in a declaration
+\Ind{\Gamma}{\Gamma_P}{\Gamma_I}{\Gamma_C}
+where $\Gamma_I$ is $[I_1:A_1;\ldots;I_k:A_k]$, and $\Gamma_C$ is
+ $[c_1:C_1;\ldots;c_n:C_n]$.
+The terms structurally smaller than $y$ are:
+\begin{itemize}
+\item $(t~u), [x:u]t$ when $t$ is structurally smaller than $y$ .
+\item \Case{P}{c}{f_1\ldots f_n} when each $f_i$ is structurally
+ smaller than $y$. \\
+ If $c$ is $y$ or is structurally smaller than $y$, its type is an inductive
+ definition $I_p$ part of the inductive
+ declaration corresponding to $y$.
+ Each $f_i$ corresponds to a type of constructor $C_q \equiv
+ (y_1:B_1)\ldots (y_k:B_k)(I~a_1\ldots a_k)$ and can consequently be
+ written $[y_1:B'_1]\ldots [y_k:B'_k]g_i$.
+ ($B'_i$ is obtained from $B_i$ by substituting parameters variables)
+ the variables $y_j$ occurring
+ in $g_i$ corresponding to recursive arguments $B_i$ (the ones in
+ which one of the $I_l$ occurs) are structurally smaller than $y$.
+\end{itemize}
+The following definitions are correct, we enter them using the
+{\tt Fixpoint} command as described in section~\ref{Fixpoint} and show
+the internal representation.
+\begin{coq_example}
+Fixpoint plus [n:nat] : nat -> nat :=
+ [m:nat]Case n of m [p:nat](S (plus p m)) end.
+Print plus.
+Fixpoint lgth [A:Set;l:(list A)] : nat :=
+ Case l of O [a:A][l':(list A)](S (lgth A l')) end.
+Print lgth.
+Fixpoint sizet [t:tree] : nat
+ := Case t of [f:forest](S (sizef f)) end
+with sizef [f:forest] : nat
+ := Case f of O [t:tree][f:forest](plus (sizet t) (sizef f)) end.
+Print sizet.
+\end{coq_example}
+
+
+\subsubsection{Reduction rule}
+\index{iota-reduction@$\iota$-reduction}
+Let $F$ be the set of declarations: $f_1/k_1:A_1:=t_1 \ldots
+f_n/k_n:A_n:=t_n$.
+The reduction for fixpoints is:
+\[ (\Fix{f_i}{F}~a_1\ldots
+a_{k_i}) \triangleright_{\iota} \substs{t_i}{f_k}{\Fix{f_k}{F}}{k=1\ldots n}\]
+when $a_{k_i}$ starts with a constructor.
+This last restriction is needed in order to keep strong normalization
+and corresponds to the reduction for primitive recursive operators.
+
+We can illustrate this behavior on examples.
+\begin{coq_example}
+Goal (n,m:nat)(plus (S n) m)=(S (plus n m)).
+Reflexivity.
+Abort.
+Goal (f:forest)(sizet (node f))=(S (sizef f)).
+Reflexivity.
+Abort.
+\end{coq_example}
+But assuming the definition of a son function from \tree\ to \forest:
+\begin{coq_example}
+ Definition sont : tree -> forest := [t]Case t of [f]f end.
+\end{coq_example}
+The following is not a conversion but can be proved after a case analysis.
+\begin{coq_example}
+Goal (t:tree)(sizet t)=(S (sizef (sont t))).
+(* this one fails *)
+Reflexivity.
+Destruct t.
+Reflexivity.
+\end{coq_example}
+\begin{coq_eval}
+Abort.
+\end{coq_eval}
+
+\subsubsection{The {\tt Match \ldots with \ldots end} expression}
+\label{Matchexpr}
+%\paragraph{A unary {\tt Match\ldots with \ldots end}.}
+\index{Match...with...end@{\tt Match \ldots with \ldots end}}
+The {\tt Match} operator which was a primitive notion in older
+presentations of the Calculus of Inductive Constructions is now just a
+macro definition which generates the good combination of {\tt Case}
+and {\tt Fix} operators in order to generate an operator for primitive
+recursive definitions. It always considers an inductive definition as
+a single inductive definition.
+
+The following examples illustrates this feature.
+\begin{coq_example}
+Definition nat_pr : (C:Set)C->(nat->C->C)->nat->C
+ :=[C,x,g,n]Match n with x g end.
+Print nat_pr.
+\end{coq_example}
+\begin{coq_example}
+Definition forest_pr
+ : (C:Set)C->(tree->forest->C->C)->forest->C
+ := [C,x,g,n]Match n with x g end.
+\end{coq_example}
+
+% Cet exemple faisait error (HH le 12/12/96), j'ai change pour une
+% version plus simple
+%\begin{coq_example}
+%Definition forest_pr
+% : (P:forest->Set)(P emptyf)->((t:tree)(f:forest)(P f)->(P (consf t f)))
+% ->(f:forest)(P f)
+% := [C,x,g,n]Match n with x g end.
+%\end{coq_example}
+
+The principles of mutual induction can be automatically generated
+using the {\tt Scheme} command described in section~\ref{Scheme}.
+
+\section{Coinductive types}
+The implementation contains also coinductive definitions, which are
+types inhabited by infinite objects.
+%They are described inchapter~\ref{Coinductives}.
+
+% $Id$
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: "Reference-Manual"
+%%% End:
+
+