aboutsummaryrefslogtreecommitdiffhomepage
path: root/doc/faq.tex
blob: c69c587b7924b277f90410e720279b54f39c9ef4 (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
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
697
\documentclass{article}
\usepackage{fullpage}
\usepackage{hevea}
\usepackage{url}

\input{./macros.tex}

\newcommand{\coqtt}[1]{{\tt #1}}
\newcommand{\coqimp}{{$->$}}
\newcommand{\coqequiv}{{$<->$}}

\newcommand{\question}[1]{
  \subsection*{\aname{no:number}{Question:~\sf{#1}}}
  \addcontentsline{toc}{subsection}{\ahrefloc{no:number}{#1}}
  \addcontentsline{htoc}{subsection}{\ahrefloc{no:number}{#1}}}
%{\subsubsection{Question:~\sf{#1}}}}
\newcommand{\answer}{{\bf Answer:~}}


\newcommand\vfile[2]{\ahref{#1}{\tt {#2}.v}}

\begin{document}
\urldef{\refman}{\url}{http://coq.inria.fr/doc/main.html}
\urldef{\InitWf}{\url}
  {http://coq.inria.fr/library/Coq.Init.Wf.html}
\urldef{\LogicBerardi}{\url}
  {http://coq.inria.fr/library/Coq.Logic.Berardi.html}
\urldef{\LogicClassical}{\url}
  {http://coq.inria.fr/library/Coq.Logic.Classical.html}
\urldef{\LogicClassicalFacts}{\url}
  {http://coq.inria.fr/library/Coq.Logic.ClassicalFacts.html}
\urldef{\LogicProofIrrelevance}{\url}
  {http://coq.inria.fr/library/Coq.Logic.ProofIrrelevance.html}
\urldef{\LogicEqdep}{\url}
  {http://coq.inria.fr/library/Coq.Logic.Eqdep.html}
\urldef{\LogicEqdepDec}{\url}
  {http://coq.inria.fr/library/Coq.Logic.Eqdep_dec.html}

\begin{center}
\begin{Large}
Frequently Asked Questions
\end{Large}

\medskip
{\Coq} version 7
\end{center}

\tableofcontents

\section{General}

\question{What is {\Coq}?}

\answer {\Coq} is a proof assistant with two main
application fields: proving program correctness and formalizing mathematical
theories. {\Coq} offers also automatic extraction of zero-defect programs
from a proof of their specification.

\question{What are the main applications done in {\Coq}?}

\answer Consequent developments include Java certification, safety of
cryptographic protocols, constructive and classical
analysis, algebraic topology, ...

\question{What is the logic of {\Coq}?}

\answer {\Coq} is based on an axiom-free type theory called
the Calculus of Inductive Constructions (see Coquand \cite{CoHu86}
and Coquand--Paulin-Mohring \cite{CoPa89}). It includes higher-order
functions and predicates, inductive and co-inductive datatypes and
predicates and a stratified hierarchy of sets.

\question{Is {\Coq}'s logic constructive?}

\answer {\Coq} theory is constructive, but classical
reasoning on proposition can be loaded as an external module.

\question{Is {\Coq} about proofs or provability?}

\answer {\Coq} is definitely about proofs. A statement
cannot be accepted without a proof expressed in the Calculus of
Inductive Constructions. The correctness of a proof relies on a
relatively small proof-checker at the kernel of {\Coq} (8000 lines of ML
code). Even {\Coq} decision procedures are producing proofs which are
double-checked by the kernel.

\question{Can I define non-terminating programs in {\Coq}?}

\answer No, all programs in {\Coq} are terminating. Especially, loops
must come with an evidence of their termination.

\question{Is {\Coq} a theorem prover?}

\answer {\Coq} comes with a few decision procedures (on propositional
calculus, Presburger arithmetic, ring and field simplification,
resolution, ...) but the main style for proving theorems is by using
LCF-style tactics.

\question{How is equational reasoning working in {\Coq}?}

\answer {\Coq} is reasoning modulo sequential evaluation of
(not necessarily closed) programs (it is decidable because {\Coq}
programs are terminating). This equality is called {\em conversion}.
Besides conversion, equations are to be treated by hand or using
specialized tactics (e.g. simplifications on fields).

\section{Equality}

%\subsection{Equality on Terms}

\question{How can I prove that 2 terms in an inductive sets are equal? Or different?}

\answer Cf "Decide Equality" and "Discriminate" in the \ahref{\refman}{Reference Manual}.

\question{Why is the proof of \coqtt{0+n=n} on natural numbers
trivial but the proof of \coqtt{n+0=n} is not?}

\answer Since \coqtt{+} (\coqtt{plus}) on natural numbers is defined by analysis on its first argument

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

the expression \coqtt{0+n} evaluates to \coqtt{n}. As {\Coq} reasons
modulo evaluation of expressions, \coqtt{0+n} and \coqtt{n} are
considered equal and the theorem \coqtt{0+n=n} is an instance of the
reflexivity of equality. On the other side, \coqtt{n+0} does not
evaluate to \coqtt{n} and a proof by induction on \coqtt{n} is
necessary to trigger the evaluation of \coqtt{+}.

\question{I have two proofs of the same proposition. Can I prove they are equal?}
\label{proof-irrelevance}

\answer In the base {\Coq} system, the answer is no. However, if
classical logic is set, the answer is yes for propositions in {\Prop}.

More precisely, the equality of all proofs of a given proposition is called
{\em proof-irrelevance}. Proof-irrelevance (in {\Prop}) can be assumed
  without leading to contradiction in {\Coq}.

  That classical logic (what can be assumed in {\Coq} by requiring the file
  \vfile{\LogicClassical}{Classical}) implies proof-irrelevance is shown in files
  \vfile{\LogicProofIrrelevance}{ProofIrrelevance} and \vfile{\LogicBerardi}{Berardi}.

  Alternatively, propositional extensionality (i.e.  \coqtt{(A
  \coqequiv B) \coqimp A=B}, defined in \vfile{\LogicClassicalFacts}{ClassicalFacts}),
  also implies proof-irrelevance.

  It is an ongoing work of research to natively include proof
  irrelevance in {\Coq}.

\question{I have two proofs of an equality statement. Can I prove they are 
equal?}

\answer Yes, if equality is decidable on the domain considered (which
is the case for {\tt nat}, {\tt bool}, etc): see {\Coq} file
\verb=Eqdep_dec.v=). No otherwise, unless
assuming Streicher's axiom $K$ (see \cite{HofStr98}) or a more general
assumption such as proof-irrelevance (see \ref{proof-irrelevance}) or
classical logic.

\question{Can I prove that the second components of equal dependent
pairs are equal?}

\answer The answer is the same as for proofs of equality
statements. It is provable if equality on the domain of the first
component is decidable (look at \verb=inj_right_pair= from file
\vfile{\LogicEqdepDec}{Eqdep\_dec}), but not provable in the general
case. However, it is consistent (with the Calculus of Constructions)
to assume it is true. The file \vfile{\LogicEqdep}{Eqdep} actually
provides an axiom (equivalent to Streicher's axiom $K$) which entails
the result (look at \verb=inj_pair2= in \vfile{\LogicEqdep}{Eqdep}).

\question{What is Streicher's axiom $K$}
\label{Streicher}

\answer Streicher's axiom $K$ \cite{HofStr98} asserts dependent
elimination of reflexive equality proofs.

\begin{coq_example*}
Axiom Streicher_K: (U:Type)(x:U)(P:x=x->Prop)
  (P (refl_equal ? x))->(p:x=x)(P p).
\end{coq_example*}

In the general case, axiom $K$ is an independent statement of the
Calculus of Inductive Constructions.  However, it is true on decidable
domains (see file \vfile{\LogicEqdepDec}{Eqdep\_dec}). It is also
trivially a consequence of proof-irrelevance (see
\ref{proof-irrelevance}) hence of classical logic.

Axiom $K$ is equivalent to {\em Uniqueness of Identity Proofs} \cite{HofStr98}
which states that

\begin{coq_example*}
Axiom UIP: (A:Set)(x,y:A)(p1,p2:x=y)p1==p2.
\end{coq_example*}

Axiom $K$ is also equivalent to {\em Uniqueness of Reflexive Identity Proofs} \cite{HofStr98}

\begin{coq_example*}
Axiom UIP_refl: (A:Set)(x:A)(p:x=x)p==(refl_equal A x).
\end{coq_example*}

Axiom $K$ is also equivalent to 

\begin{coq_example*}
Axiom eq_rec_eq : (A:Set)(a:A)(P:A->Set)(p:(P p))(h:a==a) p=(eq_rect A a P p a h).
\end{coq_example*}

It is also equivalent to the injectivity of dependent equality (dependent equality is itself equivalent to equality of dependent pairs).

\begin{coq_example*}
Inductive eq_dep [U:Set;P:U->Set;p:U;x:(P p)] : (q:U)(P q)->Prop :=
   eq_dep_intro : (eq_dep U P p x p x).
Axiom eq_dep_eq : (U:Set)(u:U)(P:U->Set)(p1,p2:(P u))(eq_dep U P u p1 u p2)->p1=p2.
\end{coq_example*}

All of these statements can be found in file \vfile{\LogicEqdep}{Eqdep}.

% \subsection{Equality on types}

\question{How to prove that 2 sets are differents?}

\answer You need to find a property true on one set and false on the
other one. As an example we show how to prove that {\tt bool} and {\tt
nat} are discriminable. As discrimination property we take the
property to have no more than 2 elements.

\begin{coq_example*}
Theorem nat_bool_discr : ~bool==nat.
Proof.
  Pose discr := [X:Set]~(a,b:X)~(x:X)~x=a->~x=b->False.
  Intro Heq; Assert H : (discr bool).
    Intro H; Apply (H true false); NewDestruct x; Auto.
  Rewrite Heq in H; Apply H; Clear H.
  NewDestruct a; NewDestruct b; Intro H0; EAuto.
  NewDestruct n; [Apply (H0 (S (S O))); Discriminate | EAuto].
Qed.
\end{coq_example*}

\question{How to exploit equalities on sets}

\answer To extract information from an equality on sets, you need to find a
predicate satisfied by the elements of the sets. As an example we show
the following theorem.

\begin{coq_example*}
Theorem le_le: (m,n:nat){x:nat | (le x m)} = {x:nat | (le x n)} -> m = n.

Inductive discr_int_n [n:nat;X:Set] : Prop :=
  intro : (fst:X->nat)(snd:(x:X)(le (fst x) n))
            ((x,y:X)(exist ? [z](le z n) (fst x) (snd x))=
                (exist ? [z](le z n) (fst y) (snd y))->x=y)
          -> (discr_int_n n X).

Proof.
Intros.
Assert Hdiscr : (discr_int_n m {x:nat | (le x m)}).
Split with (proj1_sig nat [x:?](le x m)) (proj2_sig nat [x:?](le x m)).
NewDestruct x; NewDestruct y.
Exact [x]x.
Rewrite H in Hdiscr.
Elim Hdiscr.
Assert H1:=(snd (exist nat [z](le z (S n)) (S n) (le_n (S n)))).
(* ... *)
\end{coq_example*}

\section{Axioms}

\question{Can I identify two functions that have the same behaviour (extensionality)?}

\answer Extensionality of functions is an axiom in, say set theory,
but from a programming point of view, extensionality cannot be {\em a
priori} accepted since it would identify, say programs as mergesort
and quicksort.

%\begin{coq_example*}
%  Axiom extensionality : (A,B:Set)(f,g:(A->B))(x:A)(f x)=(g x)->f=g.
%\end{coq_example*}

Let {\tt A}, {\tt B} be types. To deal with extensionality on 
\verb=A->B=, the recommended approach is to define his
own extensional equality on \verb=A->B=.

\begin{coq_example*}
  Definition ext_eq [f,g:A->B] := (x:A)(f x)=(g x).
\end{coq_example*}

and to reason on \verb=A->B= as a setoid (see the Chapter on
Setoids in the Reference Manual).

\question{Is {\Type} impredicative (that is, letting
\coqtt{T:=(X:Type)X->X}, can I apply an expression \coqtt{t} of type \coqtt{T} to \coqtt{T} itself)?}

\answer No, {\Type} is stratified. This is hidden for the
user, but {\Coq} internally maintains a set of constraints ensuring
stratification.

  If {\Type} were impredicative then Girard's systems U- and U could be
encoded in {\Coq} and it is known from Girard, Coquand, Hurkens and
Miquel that systems U- and U are inconsistent [cf Girard 1972, Coquand
1991, Hurkens 1993, Miquel 2001].

  For instance, when the user see {\tt (X:Type)X->X : Type}, each
occurrence of {\Type} is implicitly bound to a different level, say
$\alpha$ and $\beta$ and the actual statement is {\tt
(X:Type($\alpha$))X->X : Type($\beta$)} with the constraint
$\alpha<\beta$.

  When a statement violates a constraint, the message {\tt Universe
inconsistency} appears. Example: {\tt [x:=Type][y:(X:Type)X \coqimp X](z x x)}.

\section{Inductive types}

\question{What is a "large inductive definition"?}

\answer An inductive definition in {\Prop} pr {\Set} is called large
if its constructors embed sets or propositions. As an example, here is
a large inductive type:

\begin{coq_example*}
  Inductive sigST [P:Set->Set] : Set
      := existST : (X:Set)(P X) -> (sigST P).
\end{coq_example*}

   Large inductive definition have restricted elimination scheme to
prevent inconsistencies. Especially, projecting the set or the
proposition content of a large inductive definition is forbidden. If
it were allowed, we could be able to encode
e.g. Burali-Forti paradox \cite{Gir70,Coq85}.

\question{Why Injection does not work on impredicative types?}
 E.g. in this case:

\begin{coq_eval}
Reset Initial.
\end{coq_eval}

\begin{coq_example*}
  Inductive I : Set := intro : (k:Set)k -> I.
  Lemma eq_jdef : (x,y:nat) (intro ? x)=(intro ? y) -> x=y.
  Intros x y H; Injection H.
\end{coq_example*}

\answer Injectivity of constructors is restricted to predicative types. If
injectivity on large inductive types were not restricted, we would be
allowed to derive an inconsistency (e.g. following the lines of
Burali-Forti paradox). The question remains open whether injectivity
is consistent on some large inductive types not expressive enough to
encode known paradoxes (such as type I above).

\section{Recursion}

\question{Why can't I define a non terminating program?}

\answer Because otherwise the decidability of the type-checking
algorithm (which involves evaluation of programs) is not ensured.  On
another side, if non terminating proofs were allowed, we could get a
proof of {\tt False}:

\begin{coq_example*}
(* This is fortunately not allowed! *)
Fixpoint InfiniteProof [n:nat] : False := (InfiniteProof n).
Theorem Paradox : False.
Proof (InfiniteProof O).
\end{coq_example*}

\question{Why only structurally well-founded loops are allowed?}

\answer The structural order on inductive types is a simple and
powerful notion of termination. The consistency of the Calculus of
Inductive Constructions relies on it and another consistency proof
would have to be made for stronger termination arguments (such
as the termination of the evaluation of CIC programs themselves!).

In spite of this, all non-pathological termination orders can be mapped
to a structural order. Tools to do this are provided in the file 
\vfile{\InitWf}{Wf} of the standard library of {\Coq}.

\question{How to define loops based on non structurally smaller
recursive calls?}

\answer The procedure is as follows (we consider the definition of {\tt
mergesort} as an example).

\begin{itemize}

\item Define the termination order, say {\tt R} on the type {\tt A} of
the arguments of the loop.

\begin{coq_example*}
Definition R [a,b:(list nat)] := (length a) < (length b).
\end{coq_example*}

\item Prove that this order is well-founded (in fact that all elements in {\tt A} are accessible along {\tt R}).

\begin{coq_example*}
Lemma Rwf : (well_founded R).
\end{coq_example*}

\item Define the step function (which needs proofs that recursive
calls are on smaller arguments).

\begin{coq_example*}
Definition split :
   (l:(list nat)){l1:(list nat) | (R l1 l)} * {l2:(list nat) | (R l2 l)}
   := (* ... *) .
Definition concat : (l1,l2:(list nat))(list nat) := (* ... *) .
Definition merge_step [l:(list nat)][f:(l':(list nat))(R l' l)->(list nat)] :=
  let (lH1,lH2) = (split l) in
  let (l1,H1) = lH1 in
  let (l2,H2) = lH2 in
  (concat (f l1 H1) (f l2 H2)).
\end{coq_example*}

\item Define the recursive function by fixpoint on the step function.

\begin{coq_example*}
Definition merge := (fix (list nat) R Rwf [_](list nat) merge_step).
\end{coq_example*}

\end{itemize}

\question{What is behind these accessibility and well-foundedness proofs?}

\answer Well-foundedness of some relation {\tt R} on some type {\tt A}
is defined as the accessibility of all elements of {\tt A} along {\tt R}.

\begin{coq_example}
Print well_founded.
Print Acc.
\end{coq_example}

The structure of the accessibility predicate is a well-founded tree
branching at each node {\tt x} in {\tt A} along all the nodes {\tt x'}
less than {\tt x} along {\tt R}. Any sequence of elements of {\tt A}
decreasing along the order {\tt R} are branches in the accessibility
tree. Hence any decreasing along {\tt R} is mapped into a structural
decreasing in the accessibility tree of {\tt R}. This is emphasized in the definition of {\tt fix} which recurs not on its argument {\tt x:A} but on the accessibility of this argument along {\tt R}.

See file \vfile{\InitWf}{Wf}.

\section{Elimination, Case Analysis and Induction}

%\subsection{Parametric elimination}

\question{What is parametric elimination?}

\answer Parametric elimination is a 
generalisation of the substitutivity of equality. Only inductive types
with non-constant parameters are concerned with this kind of dependent
elimination. Typical examples of inductive types enjoying parametric
elimination in {\Coq} standard library are the equality predicate
\verb=eq= and the less than predicate on natural numbers \verb=le=.

\begin{coq_example}
Print eq.
Print le.
\end{coq_example}

%\subsection{Dependent elimination}

\question{What means dependent elimination?}

\answer Dependent elimination is a generalisation of the
standard Peano's induction principle on natural numbers.  This
elimination asserts that to prove a property P on say the set {\tt nat} of
natural numbers, it is enough to prove the property on O and to prove
it on successor numbers. Combined with fixpoints, this directly
provides with usual Peano's induction scheme.

Non dependent elimination is when elimination is used to build an
object whose type does not depend on the eliminated object. An example
of non dependent elimination is recursion, or the simpler form of
{\coqtt if $\ldots$ then $\ldots$ else}.

The relevant tactics to deal with this kind of elimination are {\tt
Elim}, {\tt Induction}, {\tt Destruct} and {\tt Case}, {\tt
NewInduction} and {\tt NewDestruct}.


% \question{Does the elimination in Prop follows a
%   different rule than the elimination in Set?}

% \answer There are two kinds of dependent case analysis. The result type of a
% case analysis may depend on the matched object (this is what is
% usually intended by dependent) but may also depend on the non-constant
% arguments of the type of the matched object without being dependent of
% the object itself.

%   This is typically what happens in a proof using equality. The result
% does not depend on the proof of the equality itself but crucially
% depends on the second compared argument of the equality.

%   It happens that proofs using elimination (i.e. case analysis or
% induction) on objects at the Set level are usually dependent on the
% matched term. This is because propositions precisely depend on the
% object on which an induction or case analysis is applied. Conversely,
% it happens that proofs using elimination on objects at the Prop level
% (typically elimination on equality proofs) are usually not dependent
% on the matched object. This is because proofs usually stay at the
% proof level and a few developments are stating things about proofs.
% But, again, to be not dependent on the matched objet does not mean not
% dependent at all. For equality, dependency on the second compared
% argument is precisely what provides with the substitutivity schema of
% equality.

\question{Why is dependent elimination in Prop not
available by default?}

\answer 
This is just because most of the time it is not needed. To derive a
dependent elimination principle in Prop, use the command {\tt Scheme} and
apply the elimination scheme using the \verb=using= option of
\verb=Elim=, \verb=NewDestruct= or \verb=NewInduction=. 


%\subsection{Induction}

\question{How to perform double induction?}

\answer To reason by induction on pairs, just reason by induction on the first component then by case analysis on the second component. Here is an example:

\begin{coq_eval}
Require Arith.
\end{coq_eval}

\begin{coq_example}
Infix "<" lt (at level 5, no associativity).
Infix "<=" le (at level 5, no associativity).
Lemma le_or_lt : (n,n0:nat) n0<n \/ n<=n0.
NewInduction n; NewDestruct n0; Auto with arith.
NewDestruct (IHn n0); Auto with arith.
\end{coq_example}

\question{How to define a function by double recursion?}

\answer The same trick applies, you can even use the patterm-matching
compilation algorithm to do the work for you. Here is an example:

\begin{coq_example}
Fixpoint minus [n:nat] : nat -> nat := 
  [m:nat]Cases n m of
            O    _     =>  O
         | (S k) O     => (S k)
         | (S k) (S l) => (minus k l)
        end. 
Print minus.
\end{coq_example}


\section{Coinductive types}

\question{I have a cofixpoint t:=F(t) and I want to prove t=F(t). What is
the simplest way?}

\answer Just case-expand F(t) then complete by a trivial case analysis.
Here is what it gives on e.g. the type of streams on naturals

\begin{coq_example}
  Set Implicit Arguments.
  CoInductive Stream [A : Set]  : Set := Cons : A->(Stream A)->(Stream A).
  CoFixpoint nats : nat -> (Stream nat) := [n](Cons n (nats (S n))).

  Lemma Stream_unfold : (n:nat)(nats n)= (Cons n (nats (S n))).
  Proof.
    Intro; Change (nats n) = (Cases (nats n) of (Cons x s) => (Cons x s) end).
    Case (nats n); Reflexivity.
  Qed.
\end{coq_example}

\section{Miscellaneous}

\question{I copy-paste a term and {\Coq} says it is not convertible
  to the original term. Sometimes it even says the copied term is not
well-typed.}

\answer This is probably due to invisible implicit information (implicit
arguments, coercions and Cases annotations) in the printed term, which
is not re-synthetized from the copied-pasted term in the same way as
it is in the original term.

  Consider for instance {\tt (eqT Type True True)}. This term is
printed as {\tt True==True} and re-parsed as {\tt (eqT Prop True
True)}. The two terms are not convertible (hence they fool tactics
like {\tt Pattern}.

  There is currently no satisfactory answer to the problem.
 
  Due to coercions, one may even face type-checking errors. In some
rare cases, the criterion to hide coercions is a bit too loosy, which
may result in a typing error message if the parser is not able to find
again the missing coercion.

\question{I have a problem of dependent elimination on
proofs, how to solve it?}

\begin{coq_eval}
Reset Initial.
\end{coq_eval}

\begin{coq_example*}
Inductive Def1 : Set
   := c1 : Def1.

Inductive DefProp : Def1 -> Set
   := c2 : (d:Def1)(DefProp d).

Inductive Comb : Set
   := c3 : (d:Def1)(DefProp d) -> Comb.

Lemma eq_comb
   : (d1,d1':Def1; d2:(DefProp d1); d2':(DefProp d1')) d1=d1' ->
     (c3 d1 d2)=(c3 d1' d2').
\end{coq_example*}

\answer You need to derive the dependent elimination
scheme for DefProp by hand using {\coqtt Scheme}.

\begin{coq_eval}
Abort.
\end{coq_eval}

\begin{coq_example*}
Scheme DefProp_elim := Induction for DefProp Sort Prop.

Lemma eq_comb
   : (d1,d1':Def1)d1=d1'->(d2:(DefProp d1); d2':(DefProp d1'))
     (c3 d1 d2)=(c3 d1' d2').
Intros.
NewDestruct H.
NewDestruct d2 using DefProp_elim.
NewDestruct d2' using DefProp_elim.
Reflexivity.
Qed.
\end{coq_example*}

\question{And what if I want to prove the following?}

\begin{coq_example*}
Inductive natProp : nat -> Prop
   := p0 : (natProp O)
   |  pS : (n:nat)(natProp n) -> (natProp (S n)).

Inductive package : Set
   := pack : (n:nat)(natProp n) -> package.

Lemma eq_pack
   : (n,n':nat)
     n=n' ->
     (np:(natProp n); np':(natProp n'))
     (pack n np)=(pack n' np').
\end{coq_example*}

\answer

\begin{coq_eval}
Abort.
\end{coq_eval}
\begin{coq_example*}
Scheme natProp_elim := Induction for natProp Sort Prop.

Definition pack_S : package -> package. 
NewDestruct 1.
Apply (pack (S n)).
Apply pS; Assumption.
Defined.

Lemma eq_pack
   : (n,n':nat)
     n=n' ->
     (np:(natProp n); np':(natProp n'))
     (pack n np)=(pack n' np').
Intros n n' Heq np np'.
Generalize Dependent n'.
NewInduction np using natProp_elim.
NewInduction np' using natProp_elim; Intros; Auto. 
Discriminate Heq.
NewInduction np' using natProp_elim; Intros; Auto. 
Discriminate Heq.
Change (pack_S (pack n  np))=(pack_S (pack n0 np')).
Apply (f_equal package).
Apply IHnp.
Auto.
Qed.
\end{coq_example*}

\question{Why does {\Coq} tell me that \texttt{\{x:A|(P x)\}} is not convertible with \texttt{(sig A P)}?}

\answer
This is because \texttt{\{x:A|(P x)\}} is a notation for \texttt{(sig A [x:A](P x))}. Since {\Coq} does not reason up to $\eta$-conversion, this is different from
\texttt{(sig A P)}.

\bibliographystyle{plain}
\bibliography{biblio}

\end{document}