Universes.v 61.7 KB
Newer Older
1
(* Copyright (c) 2009-2012, Adam Chlipala
2 3 4 5 6 7 8 9 10
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)

(* begin hide *)
11 12
Require Import List.

13
Require Import DepList CpdtTactics.
14 15 16 17

Set Implicit Arguments.
(* end hide *)

18 19 20 21
(** printing $ %({}*% #(<a/>*# *)
(** printing ^ %*{})% #*<a/>)# *)


22 23 24

(** %\chapter{Universes and Axioms}% *)

Adam Chlipala's avatar
Adam Chlipala committed
25
(** Many traditional theorems can be proved in Coq without special knowledge of CIC, the logic behind the prover.  A development just seems to be using a particular ASCII notation for standard formulas based on %\index{set theory}%set theory.  Nonetheless, as we saw in Chapter 4, CIC differs from set theory in starting from fewer orthogonal primitives.  It is possible to define the usual logical connectives as derived notions.  The foundation of it all is a dependently typed functional programming language, based on dependent function types and inductive type families.  By using the facilities of this language directly, we can accomplish some things much more easily than in mainstream math.
26

Adam Chlipala's avatar
Adam Chlipala committed
27
   %\index{Gallina}%Gallina, which adds features to the more theoretical CIC%~\cite{CIC}%, is the logic implemented in Coq.  It has a relatively simple foundation that can be defined rigorously in a page or two of formal proof rules.  Still, there are some important subtleties that have practical ramifications.  This chapter focuses on those subtleties, avoiding formal metatheory in favor of example code. *)
28 29 30 31


(** * The [Type] Hierarchy *)

Adam Chlipala's avatar
Adam Chlipala committed
32
(** %\index{type hierarchy}%Every object in Gallina has a type. *)
33 34 35 36 37 38 39 40 41 42 43 44 45 46 47

Check 0.
(** %\vspace{-.15in}% [[
  0
     : nat
     ]]

  It is natural enough that zero be considered as a natural number. *)

Check nat.
(** %\vspace{-.15in}% [[
  nat
     : Set
     ]]

48
  From a set theory perspective, it is unsurprising to consider the natural numbers as a "set." *)
49 50 51 52 53 54 55

Check Set.
(** %\vspace{-.15in}% [[
  Set
     : Type
     ]]

56
  The type [Set] may be considered as the set of all sets, a concept that set theory handles in terms of%\index{class (in set theory)}% _classes_.  In Coq, this more general notion is [Type]. *)
57 58 59 60 61 62 63

Check Type.
(** %\vspace{-.15in}% [[
  Type
     : Type
     ]]

64
  Strangely enough, [Type] appears to be its own type.  It is known that polymorphic languages with this property are inconsistent, via %\index{Girard's paradox}%Girard's paradox%~\cite{GirardsParadox}%.  That is, using such a language to encode proofs is unwise, because it is possible to "prove" any proposition.  What is really going on here?
65

Adam Chlipala's avatar
Adam Chlipala committed
66
  Let us repeat some of our queries after toggling a flag related to Coq's printing behavior.%\index{Vernacular commands!Set Printing Universes}% *)
67 68 69 70 71 72 73

Set Printing Universes.

Check nat.
(** %\vspace{-.15in}% [[
  nat
     : Set
74
     ]]
75
*)
76 77 78 79 80

Check Set.
(** %\vspace{-.15in}% [[
  Set
     : Type $ (0)+1 ^
81 82
     ]]
     *)
83 84 85 86 87 88 89

Check Type.
(** %\vspace{-.15in}% [[
  Type $ Top.3 ^
     : Type $ (Top.3)+1 ^
     ]]

90
  Occurrences of [Type] are annotated with some additional information, inside comments.  These annotations have to do with the secret behind [Type]: it really stands for an infinite hierarchy of types.  The type of [Set] is [Type(0)], the type of [Type(0)] is [Type(1)], the type of [Type(1)] is [Type(2)], and so on.  This is how we avoid the "[Type : Type]" paradox.  As a convenience, the universe hierarchy drives Coq's one variety of subtyping.  Any term whose type is [Type] at level [i] is automatically also described by [Type] at level [j] when [j > i].
91

92
  In the outputs of our first [Check] query, we see that the type level of [Set]'s type is [(0)+1]. Here [0] stands for the level of [Set], and we increment it to arrive at the level that _classifies_ [Set].
93

94
  In the second query's output, we see that the occurrence of [Type] that we check is assigned a fresh%\index{universe variable}% _universe variable_ [Top.3].  The output type increments [Top.3] to move up a level in the universe hierarchy.  As we write code that uses definitions whose types mention universe variables, unification may refine the values of those variables.  Luckily, the user rarely has to worry about the details.
95

96
  Another crucial concept in CIC is%\index{predicativity}% _predicativity_.  Consider these queries. *)
97 98 99 100 101

Check forall T : nat, fin T.
(** %\vspace{-.15in}% [[
  forall T : nat, fin T
     : Set
102 103
     ]]
     *)
104 105 106 107 108

Check forall T : Set, T.
(** %\vspace{-.15in}% [[
  forall T : Set, T
     : Type $ max(0, (0)+1) ^
109 110
     ]]
     *)
111 112 113 114 115 116 117 118 119

Check forall T : Type, T.
(** %\vspace{-.15in}% [[
  forall T : Type $ Top.9 ^ , T
     : Type $ max(Top.9, (Top.9)+1) ^
     ]]

  These outputs demonstrate the rule for determining which universe a [forall] type lives in.  In particular, for a type [forall x : T1, T2], we take the maximum of the universes of [T1] and [T2].  In the first example query, both [T1] ([nat]) and [T2] ([fin T]) are in [Set], so the [forall] type is in [Set], too.  In the second query, [T1] is [Set], which is at level [(0)+1]; and [T2] is [T], which is at level [0].  Thus, the [forall] exists at the maximum of these two levels.  The third example illustrates the same outcome, where we replace [Set] with an occurrence of [Type] that is assigned universe variable [Top.9].  This universe variable appears in the places where [0] appeared in the previous query.

120
  The behind-the-scenes manipulation of universe variables gives us predicativity.  Consider this simple definition of a polymorphic identity function, where the first argument [T] will automatically be marked as implicit, since it can be inferred from the type of the second argument [x]. *)
121 122 123 124 125 126 127 128 129

Definition id (T : Set) (x : T) : T := x.

Check id 0.
(** %\vspace{-.15in}% [[
  id 0
     : nat
 
Check id Set.
Adam Chlipala's avatar
Adam Chlipala committed
130
]]
131

Adam Chlipala's avatar
Adam Chlipala committed
132
<<
133 134
Error: Illegal application (Type Error):
...
Adam Chlipala's avatar
Adam Chlipala committed
135 136
The 1st term has type "Type (* (Top.15)+1 *)" which should be coercible to "Set".
>>
137

Adam Chlipala's avatar
Adam Chlipala committed
138
  The parameter [T] of [id] must be instantiated with a [Set].  The type [nat] is a [Set], but [Set] is not.  We can try fixing the problem by generalizing our definition of [id]. *)
139 140 141 142 143 144 145

Reset id.
Definition id (T : Type) (x : T) : T := x.
Check id 0.
(** %\vspace{-.15in}% [[
  id 0
     : nat
146 147
     ]]
     *)
148 149 150 151 152

Check id Set.
(** %\vspace{-.15in}% [[
  id Set
     : Type $ Top.17 ^
153 154
  ]]
  *)
155 156 157 158 159

Check id Type.
(** %\vspace{-.15in}% [[
  id Type $ Top.18 ^
     : Type $ Top.19 ^
160 161
  ]]
  *)
162 163 164 165

(** So far so good.  As we apply [id] to different [T] values, the inferred index for [T]'s [Type] occurrence automatically moves higher up the type hierarchy.
   [[
Check id id.
Adam Chlipala's avatar
Adam Chlipala committed
166
]]
167

Adam Chlipala's avatar
Adam Chlipala committed
168
<<
169
Error: Universe inconsistency (cannot enforce Top.16 < Top.16).
Adam Chlipala's avatar
Adam Chlipala committed
170
>>
171

172
  %\index{universe inconsistency}%This error message reminds us that the universe variable for [T] still exists, even though it is usually hidden.  To apply [id] to itself, that variable would need to be less than itself in the type hierarchy.  Universe inconsistency error messages announce cases like this one where a term could only type-check by violating an implied constraint over universe variables.  Such errors demonstrate that [Type] is _predicative_, where this word has a CIC meaning closely related to its usual mathematical meaning.  A predicative system enforces the constraint that, for any object of quantified type, none of those quantifiers may ever be instantiated with the object itself.  %\index{impredicativity}%Impredicativity is associated with popular paradoxes in set theory, involving inconsistent constructions like "the set of all sets that do not contain themselves" (%\index{Russell's paradox}%Russell's paradox).  Similar paradoxes would result from uncontrolled impredicativity in Coq. *)
173 174 175 176 177 178 179 180 181 182


(** ** Inductive Definitions *)

(** Predicativity restrictions also apply to inductive definitions.  As an example, let us consider a type of expression trees that allows injection of any native Coq value.  The idea is that an [exp T] stands for a reflected expression of type [T].
   [[
Inductive exp : Set -> Set :=
| Const : forall T : Set, T -> exp T
| Pair : forall T1 T2, exp T1 -> exp T2 -> exp (T1 * T2)
| Eq : forall T, exp T -> exp T -> exp bool.
Adam Chlipala's avatar
Adam Chlipala committed
183
]]
184

Adam Chlipala's avatar
Adam Chlipala committed
185
<<
186
Error: Large non-propositional inductive types must be in Type.
Adam Chlipala's avatar
Adam Chlipala committed
187
>>
188

189
   This definition is%\index{large inductive types}% _large_ in the sense that at least one of its constructors takes an argument whose type has type [Type].  Coq would be inconsistent if we allowed definitions like this one in their full generality.  Instead, we must change [exp] to live in [Type].  We will go even further and move [exp]'s index to [Type] as well. *)
190 191 192 193 194 195

Inductive exp : Type -> Type :=
| Const : forall T, T -> exp T
| Pair : forall T1 T2, exp T1 -> exp T2 -> exp (T1 * T2)
| Eq : forall T, exp T -> exp T -> exp bool.

196 197 198
(** Note that before we had to include an annotation [: Set] for the variable [T] in [Const]'s type, but we need no annotation now.  When the type of a variable is not known, and when that variable is used in a context where only types are allowed, Coq infers that the variable is of type [Type].  That is the right behavior here, but it was wrong for the [Set] version of [exp].

   Our new definition is accepted.  We can build some sample expressions. *)
199 200 201 202 203

Check Const 0.
(** %\vspace{-.15in}% [[
  Const 0
     : exp nat
204 205
     ]]
     *)
206 207 208 209 210

Check Pair (Const 0) (Const tt).
(** %\vspace{-.15in}% [[
  Pair (Const 0) (Const tt)
     : exp (nat * unit)
211 212
     ]]
     *)
213 214 215

Check Eq (Const Set) (Const Type).
(** %\vspace{-.15in}% [[
216
  Eq (Const Set) (Const Type $ Top.59 ^ )
217 218 219 220 221 222
     : exp bool
     ]]

  We can check many expressions, including fancy expressions that include types.  However, it is not hard to hit a type-checking wall.
  [[
Check Const (Const O).
Adam Chlipala's avatar
Adam Chlipala committed
223
]]
224

Adam Chlipala's avatar
Adam Chlipala committed
225
<<
226
Error: Universe inconsistency (cannot enforce Top.42 < Top.42).
Adam Chlipala's avatar
Adam Chlipala committed
227
>>
228 229

  We are unable to instantiate the parameter [T] of [Const] with an [exp] type.  To see why, it is helpful to print the annotated version of [exp]'s inductive definition. *)
230
(** [[
231
Print exp.
232
]]
233
%\vspace{-.15in}%[[
234 235 236 237 238 239 240 241 242 243
Inductive exp
              : Type $ Top.8 ^ ->
                Type
                $ max(0, (Top.11)+1, (Top.14)+1, (Top.15)+1, (Top.19)+1) ^ :=
    Const : forall T : Type $ Top.11 ^ , T -> exp T
  | Pair : forall (T1 : Type $ Top.14 ^ ) (T2 : Type $ Top.15 ^ ),
           exp T1 -> exp T2 -> exp (T1 * T2)
  | Eq : forall T : Type $ Top.19 ^ , exp T -> exp T -> exp bool
  ]]

244
  We see that the index type of [exp] has been assigned to universe level [Top.8].  In addition, each of the four occurrences of [Type] in the types of the constructors gets its own universe variable.  Each of these variables appears explicitly in the type of [exp].  In particular, any type [exp T] lives at a universe level found by incrementing by one the maximum of the four argument variables.  A consequence of this is that [exp] _must_ live at a higher universe level than any type which may be passed to one of its constructors.  This consequence led to the universe inconsistency.
245

246
  Strangely, the universe variable [Top.8] only appears in one place.  Is there no restriction imposed on which types are valid arguments to [exp]?  In fact, there is a restriction, but it only appears in a global set of universe constraints that are maintained "off to the side," not appearing explicitly in types.  We can print the current database.%\index{Vernacular commands!Print Universes}% *)
247 248 249 250 251 252 253 254 255

Print Universes.
(** %\vspace{-.15in}% [[
Top.19 < Top.9 <= Top.8
Top.15 < Top.9 <= Top.8 <= Coq.Init.Datatypes.38
Top.14 < Top.9 <= Top.8 <= Coq.Init.Datatypes.37
Top.11 < Top.9 <= Top.8
]]

Adam Chlipala's avatar
Adam Chlipala committed
256
The command outputs many more constraints, but we have collected only those that mention [Top] variables.  We see one constraint for each universe variable associated with a constructor argument from [exp]'s definition.  Universe variable [Top.19] is the type argument to [Eq].  The constraint for [Top.19] effectively says that [Top.19] must be less than [Top.8], the universe of [exp]'s indices; an intermediate variable [Top.9] appears as an artifact of the way the constraint was generated.
257 258 259

The next constraint, for [Top.15], is more complicated.  This is the universe of the second argument to the [Pair] constructor.  Not only must [Top.15] be less than [Top.8], but it also comes out that [Top.8] must be less than [Coq.Init.Datatypes.38].  What is this new universe variable?  It is from the definition of the [prod] inductive family, to which types of the form [A * B] are desugared. *)

260
(* begin hide *)
261
(* begin thide *)
262 263
Inductive prod := pair.
Reset prod.
264
(* end thide *)
265 266
(* end hide *)

267
(** %\vspace{-.3in}%[[
268
Print prod.
269
]]
270
%\vspace{-.15in}%[[
271 272 273 274 275 276 277 278 279 280 281 282
Inductive prod (A : Type $ Coq.Init.Datatypes.37 ^ )
          (B : Type $ Coq.Init.Datatypes.38 ^ )
            : Type $ max(Coq.Init.Datatypes.37, Coq.Init.Datatypes.38) ^ :=
    pair : A -> B -> A * B
    ]]

  We see that the constraint is enforcing that indices to [exp] must not live in a higher universe level than [B]-indices to [prod].  The next constraint above establishes a symmetric condition for [A].

  Thus it is apparent that Coq maintains a tortuous set of universe variable inequalities behind the scenes.  It may look like some functions are polymorphic in the universe levels of their arguments, but what is really happening is imperative updating of a system of constraints, such that all uses of a function are consistent with a global set of universe levels.  When the constraint system may not be evolved soundly, we get a universe inconsistency error.

  %\medskip%

283
  Something interesting is revealed in the annotated definition of [prod].  A type [prod A B] lives at a universe that is the maximum of the universes of [A] and [B].  From our earlier experiments, we might expect that [prod]'s universe would in fact need to be _one higher_ than the maximum.  The critical difference is that, in the definition of [prod], [A] and [B] are defined as _parameters_; that is, they appear named to the left of the main colon, rather than appearing (possibly unnamed) to the right.
284

Adam Chlipala's avatar
Adam Chlipala committed
285
  Parameters are not as flexible as normal inductive type arguments.  The range types of all of the constructors of a parameterized type must share the same parameters.  Nonetheless, when it is possible to define a polymorphic type in this way, we gain the ability to use the new type family in more ways, without triggering universe inconsistencies.  For instance, nested pairs of types are perfectly legal. *)
286 287 288 289 290 291 292 293 294 295 296

Check (nat, (Type, Set)).
(** %\vspace{-.15in}% [[
  (nat, (Type $ Top.44 ^ , Set))
     : Set * (Type $ Top.45 ^ * Type $ Top.46 ^ )
  ]]

  The same cannot be done with a counterpart to [prod] that does not use parameters. *)

Inductive prod' : Type -> Type -> Type :=
| pair' : forall A B : Type, A -> B -> prod' A B.
297
(** %\vspace{-.15in}%[[
298
Check (pair' nat (pair' Type Set)).
Adam Chlipala's avatar
Adam Chlipala committed
299
]]
300

Adam Chlipala's avatar
Adam Chlipala committed
301
<<
302
Error: Universe inconsistency (cannot enforce Top.51 < Top.51).
Adam Chlipala's avatar
Adam Chlipala committed
303
>>
304

305 306
The key benefit parameters bring us is the ability to avoid quantifying over types in the types of constructors.  Such quantification induces less-than constraints, while parameters only introduce less-than-or-equal-to constraints.

Adam Chlipala's avatar
Adam Chlipala committed
307
Coq includes one more (potentially confusing) feature related to parameters.  While Gallina does not support real %\index{universe polymorphism}%universe polymorphism, there is a convenience facility that mimics universe polymorphism in some cases.  We can illustrate what this means with a simple example. *)
308 309 310

Inductive foo (A : Type) : Type :=
| Foo : A -> foo A.
Adam Chlipala's avatar
Adam Chlipala committed
311 312 313 314 315

(* begin hide *)
Unset Printing Universes.
(* end hide *)

316 317 318 319
Check foo nat.
(** %\vspace{-.15in}% [[
  foo nat
     : Set
320 321
     ]]
     *)
322 323 324 325 326

Check foo Set.
(** %\vspace{-.15in}% [[
  foo Set
     : Type
327 328
     ]]
     *)
329 330 331 332 333 334 335

Check foo True.
(** %\vspace{-.15in}% [[
  foo True
     : Prop
     ]]

336
  The basic pattern here is that Coq is willing to automatically build a "copied-and-pasted" version of an inductive definition, where some occurrences of [Type] have been replaced by [Set] or [Prop].  In each context, the type-checker tries to find the valid replacements that are lowest in the type hierarchy.  Automatic cloning of definitions can be much more convenient than manual cloning.  We have already taken advantage of the fact that we may re-use the same families of tuple and list types to form values in [Set] and [Type].
337 338 339 340 341 342 343 344 345 346 347 348 349

  Imitation polymorphism can be confusing in some contexts.  For instance, it is what is responsible for this weird behavior. *)

Inductive bar : Type := Bar : bar.

Check bar.
(** %\vspace{-.15in}% [[
  bar
     : Prop
     ]]

  The type that Coq comes up with may be used in strictly more contexts than the type one might have expected. *)

Adam Chlipala's avatar
Adam Chlipala committed
350

351 352 353 354 355 356 357 358 359 360 361 362 363
(** ** Deciphering Baffling Messages About Inability to Unify *)

(** One of the most confusing sorts of Coq error messages arises from an interplay between universes, syntax notations, and %\index{implicit arguments}%implicit arguments.  Consider the following innocuous lemma, which is symmetry of equality for the special case of types. *)

Theorem symmetry : forall A B : Type,
  A = B
  -> B = A.
  intros ? ? H; rewrite H; reflexivity.
Qed.

(** Let us attempt an admittedly silly proof of the following theorem. *)

Theorem illustrative_but_silly_detour : unit = unit.
364 365
  (** %\vspace{-.25in}%[[
  apply symmetry.
366 367 368 369 370
]]
<<
Error: Impossible to unify "?35 = ?34" with "unit = unit".
>>

371
Coq tells us that we cannot, in fact, apply our lemma [symmetry] here, but the error message seems defective.  In particular, one might think that [apply] should unify [?35] and [?34] with [unit] to ensure that the unification goes through.  In fact, the issue is in a part of the unification problem that is _not_ shown to us in this error message!
372 373 374 375

The following command is the secret to getting better error messages in such cases: *)

  Set Printing All.
376 377
  (** %\vspace{-.15in}%[[
   apply symmetry.
378 379 380 381 382
]]
<<
Error: Impossible to unify "@eq Type ?46 ?45" with "@eq Set unit unit".
>>

383
Now we can see the problem: it is the first, _implicit_ argument to the underlying equality function [eq] that disagrees across the two terms.  The universe [Set] may be both an element and a subtype of [Type], but the two are not definitionally equal. *)
384 385 386 387 388 389 390 391 392 393 394 395 396 397 398 399

Abort.

(** A variety of changes to the theorem statement would lead to use of [Type] as the implicit argument of [eq].  Here is one such change. *)

Theorem illustrative_but_silly_detour : (unit : Type) = unit.
  apply symmetry; reflexivity.
Qed.

(** There are many related issues that can come up with error messages, where one or both of notations and implicit arguments hide important details.  The [Set Printing All] command turns off all such features and exposes underlying CIC terms.

   For completeness, we mention one other class of confusing error message about inability to unify two terms that look obviously unifiable.  Each unification variable has a scope; a unification variable instantiation may not mention variables that were not already defined within that scope, at the point in proof search where the unification variable was introduced.  Consider this illustrative example: *)

Unset Printing All.

Theorem ex_symmetry : (exists x, x = 0) -> (exists x, 0 = x).
400
  eexists.
401 402 403 404 405 406 407 408 409 410 411 412 413 414 415 416
  (** %\vspace{-.15in}%[[
  H : exists x : nat, x = 0
  ============================
   0 = ?98
   ]]
   *)

  destruct H.
  (** %\vspace{-.15in}%[[
  x : nat
  H : x = 0
  ============================
   0 = ?99
   ]]
   *)

417 418
  (** %\vspace{-.2in}%[[
  symmetry; exact H.
419 420 421 422 423 424 425 426 427 428
]]

<<
Error: In environment
x : nat
H : x = 0
The term "H" has type "x = 0" while it is expected to have type 
"?99 = 0".
>>

429
  The problem here is that variable [x] was introduced by [destruct] _after_ we introduced [?99] with [eexists], so the instantiation of [?99] may not mention [x].  A simple reordering of the proof solves the problem. *)
430 431 432 433 434

  Restart.
  destruct 1 as [x]; apply ex_intro with x; symmetry; assumption.
Qed.

435
(** This restriction for unification variables may seem counterintuitive, but it follows from the fact that CIC contains no concept of unification variable.  Rather, to construct the final proof term, at the point in a proof where the unification variable is introduced, we replace it with the instantiation we eventually find for it.  It is simply syntactically illegal to refer there to variables that are not in scope.  Without such a restriction, we could trivially "prove" such non-theorems as [exists n : nat, forall m : nat, n = m] by [econstructor; intro; reflexivity]. *)
436 437


Adam Chlipala's avatar
Adam Chlipala committed
438 439
(** * The [Prop] Universe *)

440
(** In Chapter 4, we saw parallel versions of useful datatypes for "programs" and "proofs."  The convention was that programs live in [Set], and proofs live in [Prop].  We gave little explanation for why it is useful to maintain this distinction.  There is certainly documentation value from separating programs from proofs; in practice, different concerns apply to building the two types of objects.  It turns out, however, that these concerns motivate formal differences between the two universes in Coq.
Adam Chlipala's avatar
Adam Chlipala committed
441 442 443 444 445 446 447

   Recall the types [sig] and [ex], which are the program and proof versions of existential quantification.  Their definitions differ only in one place, where [sig] uses [Type] and [ex] uses [Prop]. *)

Print sig.
(** %\vspace{-.15in}% [[
  Inductive sig (A : Type) (P : A -> Prop) : Type :=
    exist : forall x : A, P x -> sig P
448 449
    ]]
    *)
Adam Chlipala's avatar
Adam Chlipala committed
450 451 452 453 454 455 456 457 458 459 460 461 462 463

Print ex.
(** %\vspace{-.15in}% [[
  Inductive ex (A : Type) (P : A -> Prop) : Prop :=
    ex_intro : forall x : A, P x -> ex P
    ]]

  It is natural to want a function to extract the first components of data structures like these.  Doing so is easy enough for [sig]. *)

Definition projS A (P : A -> Prop) (x : sig P) : A :=
  match x with
    | exist v _ => v
  end.

464
(* begin hide *)
465
(* begin thide *)
466
Definition projE := O.
467
(* end thide *)
468 469
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
470 471 472 473 474 475
(** We run into trouble with a version that has been changed to work with [ex].
   [[
Definition projE A (P : A -> Prop) (x : ex P) : A :=
  match x with
    | ex_intro v _ => v
  end.
Adam Chlipala's avatar
Adam Chlipala committed
476
]]
Adam Chlipala's avatar
Adam Chlipala committed
477

Adam Chlipala's avatar
Adam Chlipala committed
478
<<
Adam Chlipala's avatar
Adam Chlipala committed
479 480 481 482 483 484
Error:
Incorrect elimination of "x" in the inductive type "ex":
the return type has sort "Type" while it should be "Prop".
Elimination of an inductive object of sort Prop
is not allowed on a predicate in sort Type
because proofs can be eliminated only to build proofs.
Adam Chlipala's avatar
Adam Chlipala committed
485
>>
Adam Chlipala's avatar
Adam Chlipala committed
486

487
  In formal Coq parlance, %\index{elimination}%"elimination" means "pattern-matching."  The typing rules of Gallina forbid us from pattern-matching on a discriminee whose type belongs to [Prop], whenever the result type of the [match] has a type besides [Prop].  This is a sort of "information flow" policy, where the type system ensures that the details of proofs can never have any effect on parts of a development that are not also marked as proofs.
Adam Chlipala's avatar
Adam Chlipala committed
488 489 490

  This restriction matches informal practice.  We think of programs and proofs as clearly separated, and, outside of constructive logic, the idea of computing with proofs is ill-formed.  The distinction also has practical importance in Coq, where it affects the behavior of extraction.

491
  Recall that %\index{program extraction}%extraction is Coq's facility for translating Coq developments into programs in general-purpose programming languages like OCaml.  Extraction _erases_ proofs and leaves programs intact.  A simple example with [sig] and [ex] demonstrates the distinction. *)
Adam Chlipala's avatar
Adam Chlipala committed
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

Definition sym_sig (x : sig (fun n => n = 0)) : sig (fun n => 0 = n) :=
  match x with
    | exist n pf => exist _ n (sym_eq pf)
  end.

Extraction sym_sig.
(** <<
(** val sym_sig : nat -> nat **)

let sym_sig x = x
>>

Since extraction erases proofs, the second components of [sig] values are elided, making [sig] a simple identity type family.  The [sym_sig] operation is thus an identity function. *)

Definition sym_ex (x : ex (fun n => n = 0)) : ex (fun n => 0 = n) :=
  match x with
    | ex_intro n pf => ex_intro _ n (sym_eq pf)
  end.

Extraction sym_ex.
(** <<
(** val sym_ex : __ **)

let sym_ex = __
>>

519
In this example, the [ex] type itself is in [Prop], so whole [ex] packages are erased.  Coq extracts every proposition as the (Coq-specific) type <<__>>, whose single constructor is <<__>>.  Not only are proofs replaced by [__], but proof arguments to functions are also removed completely, as we see here.
Adam Chlipala's avatar
Adam Chlipala committed
520

521
Extraction is very helpful as an optimization over programs that contain proofs.  In languages like Haskell, advanced features make it possible to program with proofs, as a way of convincing the type checker to accept particular definitions.  Unfortunately, when proofs are encoded as values in GADTs%~\cite{GADT}%, these proofs exist at runtime and consume resources.  In contrast, with Coq, as long as all proofs are kept within [Prop], extraction is guaranteed to erase them.
Adam Chlipala's avatar
Adam Chlipala committed
522

523
Many fans of the %\index{Curry-Howard correspondence}%Curry-Howard correspondence support the idea of _extracting programs from proofs_.  In reality, few users of Coq and related tools do any such thing.  Instead, extraction is better thought of as an optimization that reduces the runtime costs of expressive typing.
Adam Chlipala's avatar
Adam Chlipala committed
524 525 526

%\medskip%

527
We have seen two of the differences between proofs and programs: proofs are subject to an elimination restriction and are elided by extraction.  The remaining difference is that [Prop] is%\index{impredicativity}% _impredicative_, as this example shows. *)
Adam Chlipala's avatar
Adam Chlipala committed
528 529 530 531 532 533 534

Check forall P Q : Prop, P \/ Q -> Q \/ P.
(** %\vspace{-.15in}% [[
  forall P Q : Prop, P \/ Q -> Q \/ P
     : Prop
     ]]

Adam Chlipala's avatar
Adam Chlipala committed
535 536 537 538 539 540 541 542 543 544 545 546 547
  We see that it is possible to define a [Prop] that quantifies over other [Prop]s.  This is fortunate, as we start wanting that ability even for such basic purposes as stating propositional tautologies.  In the next section of this chapter, we will see some reasons why unrestricted impredicativity is undesirable.  The impredicativity of [Prop] interacts crucially with the elimination restriction to avoid those pitfalls.

  Impredicativity also allows us to implement a version of our earlier [exp] type that does not suffer from the weakness that we found. *)

Inductive expP : Type -> Prop :=
| ConstP : forall T, T -> expP T
| PairP : forall T1 T2, expP T1 -> expP T2 -> expP (T1 * T2)
| EqP : forall T, expP T -> expP T -> expP bool.

Check ConstP 0.
(** %\vspace{-.15in}% [[
  ConstP 0
     : expP nat
548 549
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
550 551 552 553 554

Check PairP (ConstP 0) (ConstP tt).
(** %\vspace{-.15in}% [[
  PairP (ConstP 0) (ConstP tt)
     : expP (nat * unit)
555 556
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
557 558 559 560 561

Check EqP (ConstP Set) (ConstP Type).
(** %\vspace{-.15in}% [[
  EqP (ConstP Set) (ConstP Type)
     : expP bool
562 563
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
564 565 566 567 568 569 570

Check ConstP (ConstP O).
(** %\vspace{-.15in}% [[
  ConstP (ConstP 0)
     : expP (expP nat)
     ]]

571
  In this case, our victory is really a shallow one.  As we have marked [expP] as a family of proofs, we cannot deconstruct our expressions in the usual programmatic ways, which makes them almost useless for the usual purposes.  Impredicative quantification is much more useful in defining inductive families that we really think of as judgments.  For instance, this code defines a notion of equality that is strictly more permissive than the base equality [=]. *)
Adam Chlipala's avatar
Adam Chlipala committed
572 573 574 575 576 577 578 579 580 581 582

Inductive eqPlus : forall T, T -> T -> Prop :=
| Base : forall T (x : T), eqPlus x x
| Func : forall dom ran (f1 f2 : dom -> ran),
  (forall x : dom, eqPlus (f1 x) (f2 x))
  -> eqPlus f1 f2.

Check (Base 0).
(** %\vspace{-.15in}% [[
  Base 0
     : eqPlus 0 0
583 584
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
585 586 587 588 589

Check (Func (fun n => n) (fun n => 0 + n) (fun n => Base n)).
(** %\vspace{-.15in}% [[
  Func (fun n : nat => n) (fun n : nat => 0 + n) (fun n : nat => Base n)
     : eqPlus (fun n : nat => n) (fun n : nat => 0 + n)
590 591
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
592 593 594 595 596

Check (Base (Base 1)).
(** %\vspace{-.15in}% [[
  Base (Base 1)
     : eqPlus (Base 1) (Base 1)
597 598
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
599

Adam Chlipala's avatar
Adam Chlipala committed
600 601
(** Stating equality facts about proofs may seem baroque, but we have already seen its utility in the chapter on reasoning about equality proofs. *)

Adam Chlipala's avatar
Adam Chlipala committed
602 603 604

(** * Axioms *)

605
(** While the specific logic Gallina is hardcoded into Coq's implementation, it is possible to add certain logical rules in a controlled way.  In other words, Coq may be used to reason about many different refinements of Gallina where strictly more theorems are provable.  We achieve this by asserting%\index{axioms}% _axioms_ without proof.
Adam Chlipala's avatar
Adam Chlipala committed
606 607 608 609 610

   We will motivate the idea by touring through some standard axioms, as enumerated in Coq's online FAQ.  I will add additional commentary as appropriate. *)

(** ** The Basics *)

Adam Chlipala's avatar
Adam Chlipala committed
611
(** One simple example of a useful axiom is the %\index{law of the excluded middle}%law of the excluded middle. *)
Adam Chlipala's avatar
Adam Chlipala committed
612 613 614 615 616 617 618

Require Import Classical_Prop.
Print classic.
(** %\vspace{-.15in}% [[
  *** [ classic : forall P : Prop, P \/ ~ P ]
  ]]

Adam Chlipala's avatar
Adam Chlipala committed
619
  In the implementation of module [Classical_Prop], this axiom was defined with the command%\index{Vernacular commands!Axiom}% *)
Adam Chlipala's avatar
Adam Chlipala committed
620 621 622

Axiom classic : forall P : Prop, P \/ ~ P.

Adam Chlipala's avatar
Adam Chlipala committed
623
(** An [Axiom] may be declared with any type, in any of the universes.  There is a synonym %\index{Vernacular commands!Parameter}%[Parameter] for [Axiom], and that synonym is often clearer for assertions not of type [Prop].  For instance, we can assert the existence of objects with certain properties. *)
Adam Chlipala's avatar
Adam Chlipala committed
624

625 626 627
Parameter num : nat.
Axiom positive : num > 0.
Reset num.
Adam Chlipala's avatar
Adam Chlipala committed
628

629
(** This kind of "axiomatic presentation" of a theory is very common outside of higher-order logic.  However, in Coq, it is almost always preferable to stick to defining your objects, functions, and predicates via inductive definitions and functional programming.
Adam Chlipala's avatar
Adam Chlipala committed
630

631
   In general, there is a significant burden associated with any use of axioms.  It is easy to assert a set of axioms that together is%\index{inconsistent axioms}% _inconsistent_.   That is, a set of axioms may imply [False], which allows any theorem to be proved, which defeats the purpose of a proof assistant.  For example, we could assert the following axiom, which is consistent by itself but inconsistent when combined with [classic]. *)
Adam Chlipala's avatar
Adam Chlipala committed
632

633
Axiom not_classic : ~ forall P : Prop, P \/ ~ P.
Adam Chlipala's avatar
Adam Chlipala committed
634 635

Theorem uhoh : False.
636
  generalize classic not_classic; tauto.
Adam Chlipala's avatar
Adam Chlipala committed
637 638 639 640 641 642 643 644
Qed.

Theorem uhoh_again : 1 + 1 = 3.
  destruct uhoh.
Qed.

Reset not_classic.

645
(** On the subject of the law of the excluded middle itself, this axiom is usually quite harmless, and many practical Coq developments assume it.  It has been proved metatheoretically to be consistent with CIC.  Here, "proved metatheoretically" means that someone proved on paper that excluded middle holds in a _model_ of CIC in set theory%~\cite{SetsInTypes}%.  All of the other axioms that we will survey in this section hold in the same model, so they are all consistent together.
Adam Chlipala's avatar
Adam Chlipala committed
646

647
   Recall that Coq implements%\index{constructive logic}% _constructive_ logic by default, where the law of the excluded middle is not provable.  Proofs in constructive logic can be thought of as programs.  A [forall] quantifier denotes a dependent function type, and a disjunction denotes a variant type.  In such a setting, excluded middle could be interpreted as a decision procedure for arbitrary propositions, which computability theory tells us cannot exist.  Thus, constructive logic with excluded middle can no longer be associated with our usual notion of programming.
Adam Chlipala's avatar
Adam Chlipala committed
648

649
   Given all this, why is it all right to assert excluded middle as an axiom?  The intuitive justification is that the elimination restriction for [Prop] prevents us from treating proofs as programs.  An excluded middle axiom that quantified over [Set] instead of [Prop] _would_ be problematic.  If a development used that axiom, we would not be able to extract the code to OCaml (soundly) without implementing a genuine universal decision procedure.  In contrast, values whose types belong to [Prop] are always erased by extraction, so we sidestep the axiom's algorithmic consequences.
Adam Chlipala's avatar
Adam Chlipala committed
650

Adam Chlipala's avatar
Adam Chlipala committed
651
   Because the proper use of axioms is so precarious, there are helpful commands for determining which axioms a theorem relies on.%\index{Vernacular commands!Print Assumptions}% *)
Adam Chlipala's avatar
Adam Chlipala committed
652 653 654 655 656 657

Theorem t1 : forall P : Prop, P -> ~ ~ P.
  tauto.
Qed.

Print Assumptions t1.
Adam Chlipala's avatar
Adam Chlipala committed
658
(** <<
Adam Chlipala's avatar
Adam Chlipala committed
659
  Closed under the global context
Adam Chlipala's avatar
Adam Chlipala committed
660
>>
661
*)
Adam Chlipala's avatar
Adam Chlipala committed
662 663

Theorem t2 : forall P : Prop, ~ ~ P -> P.
664
  (** %\vspace{-.25in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
665
  tauto.
666
]]
Adam Chlipala's avatar
Adam Chlipala committed
667 668 669
<<
Error: tauto failed.
>>
670
*)
Adam Chlipala's avatar
Adam Chlipala committed
671 672 673 674 675 676 677 678 679
  intro P; destruct (classic P); tauto.
Qed.

Print Assumptions t2.
(** %\vspace{-.15in}% [[
  Axioms:
  classic : forall P : Prop, P \/ ~ P
  ]]

680
  It is possible to avoid this dependence in some specific cases, where excluded middle _is_ provable, for decidable families of propositions. *)
Adam Chlipala's avatar
Adam Chlipala committed
681

682
Theorem nat_eq_dec : forall n m : nat, n = m \/ n <> m.
Adam Chlipala's avatar
Adam Chlipala committed
683 684 685 686
  induction n; destruct m; intuition; generalize (IHn m); intuition.
Qed.

Theorem t2' : forall n m : nat, ~ ~ (n = m) -> n = m.
687
  intros n m; destruct (nat_eq_dec n m); tauto.
Adam Chlipala's avatar
Adam Chlipala committed
688 689 690
Qed.

Print Assumptions t2'.
Adam Chlipala's avatar
Adam Chlipala committed
691
(** <<
Adam Chlipala's avatar
Adam Chlipala committed
692
Closed under the global context
Adam Chlipala's avatar
Adam Chlipala committed
693
>>
Adam Chlipala's avatar
Adam Chlipala committed
694 695 696

  %\bigskip%

697
  Mainstream mathematical practice assumes excluded middle, so it can be useful to have it available in Coq developments, though it is also nice to know that a theorem is proved in a simpler formal system than classical logic.  There is a similar story for%\index{proof irrelevance}% _proof irrelevance_, which simplifies proof issues that would not even arise in mainstream math. *)
Adam Chlipala's avatar
Adam Chlipala committed
698 699 700

Require Import ProofIrrelevance.
Print proof_irrelevance.
701

Adam Chlipala's avatar
Adam Chlipala committed
702 703 704 705
(** %\vspace{-.15in}% [[
  *** [ proof_irrelevance : forall (P : Prop) (p1 p2 : P), p1 = p2 ]
  ]]

706
  This axiom asserts that any two proofs of the same proposition are equal.  Recall this example function from Chapter 6. *)
Adam Chlipala's avatar
Adam Chlipala committed
707 708 709 710 711 712 713 714 715 716 717 718 719

(* begin hide *)
Lemma zgtz : 0 > 0 -> False.
  crush.
Qed.
(* end hide *)

Definition pred_strong1 (n : nat) : n > 0 -> nat :=
  match n with
    | O => fun pf : 0 > 0 => match zgtz pf with end
    | S n' => fun _ => n'
  end.

Adam Chlipala's avatar
Adam Chlipala committed
720
(** We might want to prove that different proofs of [n > 0] do not lead to different results from our richly typed predecessor function. *)
Adam Chlipala's avatar
Adam Chlipala committed
721 722 723 724 725 726 727 728 729 730 731 732 733 734 735 736 737 738 739 740 741 742 743 744 745

Theorem pred_strong1_irrel : forall n (pf1 pf2 : n > 0), pred_strong1 pf1 = pred_strong1 pf2.
  destruct n; crush.
Qed.

(** The proof script is simple, but it involved peeking into the definition of [pred_strong1].  For more complicated function definitions, it can be considerably more work to prove that they do not discriminate on details of proof arguments.  This can seem like a shame, since the [Prop] elimination restriction makes it impossible to write any function that does otherwise.  Unfortunately, this fact is only true metatheoretically, unless we assert an axiom like [proof_irrelevance].  With that axiom, we can prove our theorem without consulting the definition of [pred_strong1]. *)

Theorem pred_strong1_irrel' : forall n (pf1 pf2 : n > 0), pred_strong1 pf1 = pred_strong1 pf2.
  intros; f_equal; apply proof_irrelevance.
Qed.


(** %\bigskip%

   In the chapter on equality, we already discussed some axioms that are related to proof irrelevance.  In particular, Coq's standard library includes this axiom: *)

Require Import Eqdep.
Import Eq_rect_eq.
Print eq_rect_eq.
(** %\vspace{-.15in}% [[
  *** [ eq_rect_eq : 
  forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
  x = eq_rect p Q x p h ]
  ]]

746
  This axiom says that it is permissible to simplify pattern matches over proofs of equalities like [e = e].  The axiom is logically equivalent to some simpler corollaries.  In the theorem names, "UIP" stands for %\index{unicity of identity proofs}%"unicity of identity proofs", where "identity" is a synonym for "equality." *)
Adam Chlipala's avatar
Adam Chlipala committed
747

748 749
Corollary UIP_refl : forall A (x : A) (pf : x = x), pf = eq_refl x.
  intros; replace pf with (eq_rect x (eq x) (eq_refl x) x pf); [
Adam Chlipala's avatar
Adam Chlipala committed
750 751
    symmetry; apply eq_rect_eq
    | exact (match pf as pf' return match pf' in _ = y return x = y with
752
                                      | eq_refl => eq_refl x
Adam Chlipala's avatar
Adam Chlipala committed
753
                                    end = pf' with
754
               | eq_refl => eq_refl _
Adam Chlipala's avatar
Adam Chlipala committed
755 756 757 758 759 760 761 762 763 764
             end) ].
Qed.

Corollary UIP : forall A (x y : A) (pf1 pf2 : x = y), pf1 = pf2.
  intros; generalize pf1 pf2; subst; intros;
    match goal with
      | [ |- ?pf1 = ?pf2 ] => rewrite (UIP_refl pf1); rewrite (UIP_refl pf2); reflexivity
    end.
Qed.

765
(* begin hide *)
766
(* begin thide *)
767
Require Eqdep_dec.
768
(* end thide *)
769 770
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
771
(** These corollaries are special cases of proof irrelevance.  In developments that only need proof irrelevance for equality, there is no need to assert full irrelevance.
Adam Chlipala's avatar
Adam Chlipala committed
772 773 774 775 776 777 778 779 780 781 782 783 784 785 786 787 788 789

   Another facet of proof irrelevance is that, like excluded middle, it is often provable for specific propositions.  For instance, [UIP] is provable whenever the type [A] has a decidable equality operation.  The module [Eqdep_dec] of the standard library contains a proof.  A similar phenomenon applies to other notable cases, including less-than proofs.  Thus, it is often possible to use proof irrelevance without asserting axioms.

   %\bigskip%

   There are two more basic axioms that are often assumed, to avoid complications that do not arise in set theory. *)

Require Import FunctionalExtensionality.
Print functional_extensionality_dep.
(** %\vspace{-.15in}% [[
  *** [ functional_extensionality_dep : 
  forall (A : Type) (B : A -> Type) (f g : forall x : A, B x),
  (forall x : A, f x = g x) -> f = g ]
 
  ]]

  This axiom says that two functions are equal if they map equal inputs to equal outputs.  Such facts are not provable in general in CIC, but it is consistent to assume that they are.

Adam Chlipala's avatar
Adam Chlipala committed
790
  A simple corollary shows that the same property applies to predicates. *)
Adam Chlipala's avatar
Adam Chlipala committed
791 792 793 794 795 796

Corollary predicate_extensionality : forall (A : Type) (B : A -> Prop) (f g : forall x : A, B x),
  (forall x : A, f x = g x) -> f = g.
  intros; apply functional_extensionality_dep; assumption.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
797 798
(** In some cases, one might prefer to assert this corollary as the axiom, to restrict the consequences to proofs and not programs. *)

Adam Chlipala's avatar
Adam Chlipala committed
799 800 801

(** ** Axioms of Choice *)

Adam Chlipala's avatar
Adam Chlipala committed
802
(** Some Coq axioms are also points of contention in mainstream math.  The most prominent example is the %\index{axiom of choice}%axiom of choice.  In fact, there are multiple versions that we might consider, and, considered in isolation, none of these versions means quite what it means in classical set theory.
Adam Chlipala's avatar
Adam Chlipala committed
803

804
   First, it is possible to implement a choice operator _without_ axioms in some potentially surprising cases. *)
Adam Chlipala's avatar
Adam Chlipala committed
805 806 807 808 809 810 811 812 813 814

Require Import ConstructiveEpsilon.
Check constructive_definite_description.
(** %\vspace{-.15in}% [[
  constructive_definite_description
     : forall (A : Set) (f : A -> nat) (g : nat -> A),
       (forall x : A, g (f x) = x) ->
       forall P : A -> Prop,
       (forall x : A, {P x} + {~ P x}) ->
       (exists! x : A, P x) -> {x : A | P x}
815 816
       ]]
       *)
Adam Chlipala's avatar
Adam Chlipala committed
817 818

Print Assumptions constructive_definite_description.
Adam Chlipala's avatar
Adam Chlipala committed
819
(** <<
Adam Chlipala's avatar
Adam Chlipala committed
820
Closed under the global context
Adam Chlipala's avatar
Adam Chlipala committed
821
>>
Adam Chlipala's avatar
Adam Chlipala committed
822

823
  This function transforms a decidable predicate [P] into a function that produces an element satisfying [P] from a proof that such an element exists.  The functions [f] and [g], in conjunction with an associated injectivity property, are used to express the idea that the set [A] is countable.  Under these conditions, a simple brute force algorithm gets the job done: we just enumerate all elements of [A], stopping when we find one satisfying [P].  The existence proof, specified in terms of _unique_ existence [exists!], guarantees termination.  The definition of this operator in Coq uses some interesting techniques, as seen in the implementation of the [ConstructiveEpsilon] module.
Adam Chlipala's avatar
Adam Chlipala committed
824 825 826 827 828 829 830 831 832

  Countable choice is provable in set theory without appealing to the general axiom of choice.  To support the more general principle in Coq, we must also add an axiom.  Here is a functional version of the axiom of unique choice. *)

Require Import ClassicalUniqueChoice.
Check dependent_unique_choice.
(** %\vspace{-.15in}% [[
  dependent_unique_choice
     : forall (A : Type) (B : A -> Type) (R : forall x : A, B x -> Prop),
       (forall x : A, exists! y : B x, R x y) ->
Adam Chlipala's avatar
Adam Chlipala committed
833 834
       exists f : forall x : A, B x,
         forall x : A, R x (f x)
Adam Chlipala's avatar
Adam Chlipala committed
835 836 837 838
  ]]

  This axiom lets us convert a relational specification [R] into a function implementing that specification.  We need only prove that [R] is truly a function.  An alternate, stronger formulation applies to cases where [R] maps each input to one or more outputs.  We also simplify the statement of the theorem by considering only non-dependent function types. *)

839
(* begin hide *)
840
(* begin thide *)
841
Require RelationalChoice.
842
(* end thide *)
843 844
(* end hide *)

Adam Chlipala's avatar
Adam Chlipala committed
845 846 847 848 849 850 851
Require Import ClassicalChoice.
Check choice.
(** %\vspace{-.15in}% [[
choice
     : forall (A B : Type) (R : A -> B -> Prop),
       (forall x : A, exists y : B, R x y) ->
       exists f : A -> B, forall x : A, R x (f x)
852
   ]]
Adam Chlipala's avatar
Adam Chlipala committed
853 854 855 856 857 858 859 860 861 862

  This principle is proved as a theorem, based on the unique choice axiom and an additional axiom of relational choice from the [RelationalChoice] module.

  In set theory, the axiom of choice is a fundamental philosophical commitment one makes about the universe of sets.  In Coq, the choice axioms say something weaker.  For instance, consider the simple restatement of the [choice] axiom where we replace existential quantification by its Curry-Howard analogue, subset types. *)

Definition choice_Set (A B : Type) (R : A -> B -> Prop) (H : forall x : A, {y : B | R x y})
  : {f : A -> B | forall x : A, R x (f x)} :=
  exist (fun f => forall x : A, R x (f x))
  (fun x => proj1_sig (H x)) (fun x => proj2_sig (H x)).

863
(** %\smallskip{}%Via the Curry-Howard correspondence, this "axiom" can be taken to have the same meaning as the original.  It is implemented trivially as a transformation not much deeper than uncurrying.  Thus, we see that the utility of the axioms that we mentioned earlier comes in their usage to build programs from proofs.  Normal set theory has no explicit proofs, so the meaning of the usual axiom of choice is subtly different.  In Gallina, the axioms implement a controlled relaxation of the restrictions on information flow from proofs to programs.
Adam Chlipala's avatar
Adam Chlipala committed
864

865
   However, when we combine an axiom of choice with the law of the excluded middle, the idea of "choice" becomes more interesting.  Excluded middle gives us a highly non-computational way of constructing proofs, but it does not change the computational nature of programs.  Thus, the axiom of choice is still giving us a way of translating between two different sorts of "programs," but the input programs (which are proofs) may be written in a rich language that goes beyond normal computability.  This truly is more than repackaging a function with a different type.
Adam Chlipala's avatar
Adam Chlipala committed
866 867 868

   %\bigskip%

869
   The Coq tools support a command-line flag %\index{impredicative Set}%<<-impredicative-set>>, which modifies Gallina in a more fundamental way by making [Set] impredicative.  A term like [forall T : Set, T] has type [Set], and inductive definitions in [Set] may have constructors that quantify over arguments of any types.  To maintain consistency, an elimination restriction must be imposed, similarly to the restriction for [Prop].  The restriction only applies to large inductive types, where some constructor quantifies over a type of type [Type].  In such cases, a value in this inductive type may only be pattern-matched over to yield a result type whose type is [Set] or [Prop].  This contrasts with [Prop], where the restriction applies even to non-large inductive types, and where the result type may only have type [Prop].
Adam Chlipala's avatar
Adam Chlipala committed
870 871 872 873 874

   In old versions of Coq, [Set] was impredicative by default.  Later versions make [Set] predicative to avoid inconsistency with some classical axioms.  In particular, one should watch out when using impredicative [Set] with axioms of choice.  In combination with excluded middle or predicate extensionality, this can lead to inconsistency.  Impredicative [Set] can be useful for modeling inherently impredicative mathematical concepts, but almost all Coq developments get by fine without it. *)

(** ** Axioms and Computation *)

875
(** One additional axiom-related wrinkle arises from an aspect of Gallina that is very different from set theory: a notion of _computational equivalence_ is central to the definition of the formal system.  Axioms tend not to play well with computation.  Consider this example.  We start by implementing a function that uses a type equality proof to perform a safe type-cast. *)
Adam Chlipala's avatar
Adam Chlipala committed
876 877 878

Definition cast (x y : Set) (pf : x = y) (v : x) : y :=
  match pf with
879
    | eq_refl => v
Adam Chlipala's avatar
Adam Chlipala committed
880 881 882 883
  end.

(** Computation over programs that use [cast] can proceed smoothly. *)

884
Eval compute in (cast (eq_refl (nat -> nat)) (fun n => S n)) 12.
Adam Chlipala's avatar
Adam Chlipala committed
885
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
886 887
     = 13
     : nat
888 889
     ]]
     *)
Adam Chlipala's avatar
Adam Chlipala committed
890 891 892 893 894 895 896 897 898

(** Things do not go as smoothly when we use [cast] with proofs that rely on axioms. *)

Theorem t3 : (forall n : nat, fin (S n)) = (forall n : nat, fin (n + 1)). 
  change ((forall n : nat, (fun n => fin (S n)) n) = (forall n : nat, (fun n => fin (n + 1)) n));
    rewrite (functional_extensionality (fun n => fin (n + 1)) (fun n => fin (S n))); crush.
Qed.

Eval compute in (cast t3 (fun _ => First)) 12.
899
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
900
     = match t3 in (_ = P) return P with
901
       | eq_refl => fun n : nat => First
Adam Chlipala's avatar
Adam Chlipala committed
902 903 904 905
       end 12
     : fin (12 + 1)
     ]]

906
  Computation gets stuck in a pattern-match on the proof [t3].  The structure of [t3] is not known, so the match cannot proceed.  It turns out a more basic problem leads to this particular situation.  We ended the proof of [t3] with [Qed], so the definition of [t3] is not available to computation.  That mistake is easily fixed. *)
Adam Chlipala's avatar
Adam Chlipala committed
907 908 909 910 911 912 913 914 915

Reset t3.

Theorem t3 : (forall n : nat, fin (S n)) = (forall n : nat, fin (n + 1)). 
  change ((forall n : nat, (fun n => fin (S n)) n) = (forall n : nat, (fun n => fin (n + 1)) n));
    rewrite (functional_extensionality (fun n => fin (n + 1)) (fun n => fin (S n))); crush.
Defined.

Eval compute in (cast t3 (fun _ => First)) 12.
916
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
917 918 919 920 921 922 923
     = match
         match
           match
             functional_extensionality
     ....
     ]]

924
  We elide most of the details.  A very unwieldy tree of nested matches on equality proofs appears.  This time evaluation really _is_ stuck on a use of an axiom.
Adam Chlipala's avatar
Adam Chlipala committed
925 926 927 928 929 930 931 932 933 934 935 936 937 938 939

  If we are careful in using tactics to prove an equality, we can still compute with casts over the proof. *)

Lemma plus1 : forall n, S n = n + 1.
  induction n; simpl; intuition.
Defined.

Theorem t4 : forall n, fin (S n) = fin (n + 1).
  intro; f_equal; apply plus1.
Defined.

Eval compute in cast (t4 13) First.
(** %\vspace{-.15in}% [[
     = First
     : fin (13 + 1)
940
     ]]
Adam Chlipala's avatar
Adam Chlipala committed
941

942
   This simple computational reduction hides the use of a recursive function to produce a suitable [eq_refl] proof term.  The recursion originates in our use of [induction] in [t4]'s proof. *)
Adam Chlipala's avatar
Adam Chlipala committed
943

944 945 946

(** ** Methods for Avoiding Axioms *)

947
(** The last section demonstrated one reason to avoid axioms: they interfere with computational behavior of terms.  A further reason is to reduce the philosophical commitment of a theorem.  The more axioms one assumes, the harder it becomes to convince oneself that the formal system corresponds appropriately to one's intuitions.  A refinement of this last point, in applications like %\index{proof-carrying code}%proof-carrying code%~\cite{PCC}% in computer security, has to do with minimizing the size of a%\index{trusted code base}% _trusted code base_.  To convince ourselves that a theorem is true, we must convince ourselves of the correctness of the program that checks the theorem.  Axioms effectively become new source code for the checking program, increasing the effort required to perform a correctness audit.
948

949
   An earlier section gave one example of avoiding an axiom.  We proved that [pred_strong1] is agnostic to details of the proofs passed to it as arguments, by unfolding the definition of the function.  A "simpler" proof keeps the function definition opaque and instead applies a proof irrelevance axiom.  By accepting a more complex proof, we reduce our philosophical commitment and trusted base.  (By the way, the less-than relation that the proofs in question here prove turns out to admit proof irrelevance as a theorem provable within normal Gallina!)
950 951 952 953 954 955 956

   One dark secret of the [dep_destruct] tactic that we have used several times is reliance on an axiom.  Consider this simple case analysis principle for [fin] values: *)

Theorem fin_cases : forall n (f : fin (S n)), f = First \/ exists f', f = Next f'.
  intros; dep_destruct f; eauto.
Qed.

957 958
(* begin hide *)
Require Import JMeq.
959
(* begin thide *)
960
Definition jme := (JMeq, JMeq_eq).
961
(* end thide *)
962 963
(* end hide *)

964 965 966
Print Assumptions fin_cases.
(** %\vspace{-.15in}%[[
Axioms:
967
JMeq_eq : forall (A : Type) (x y : A), JMeq x y -> x = y
968 969 970 971 972 973 974 975 976 977 978 979 980 981 982 983 984 985 986 987 988 989 990 991 992 993
]]

  The proof depends on the [JMeq_eq] axiom that we met in the chapter on equality proofs.  However, a smarter tactic could have avoided an axiom dependence.  Here is an alternate proof via a slightly strange looking lemma. *)

(* begin thide *)
Lemma fin_cases_again' : forall n (f : fin n),
  match n return fin n -> Prop with
    | O => fun _ => False
    | S n' => fun f => f = First \/ exists f', f = Next f'
  end f.
  destruct f; eauto.
Qed.

(** We apply a variant of the %\index{convoy pattern}%convoy pattern, which we are used to seeing in function implementations.  Here, the pattern helps us state a lemma in a form where the argument to [fin] is a variable.  Recall that, thanks to basic typing rules for pattern-matching, [destruct] will only work effectively on types whose non-parameter arguments are variables.  The %\index{tactics!exact}%[exact] tactic, which takes as argument a literal proof term, now gives us an easy way of proving the original theorem. *)

Theorem fin_cases_again : forall n (f : fin (S n)), f = First \/ exists f', f = Next f'.
  intros; exact (fin_cases_again' f).
Qed.
(* end thide *)

Print Assumptions fin_cases_again.
(** %\vspace{-.15in}%
<<
Closed under the global context
>>

994 995 996 997 998
*)

(* begin thide *)
(** As the Curry-Howard correspondence might lead us to expect, the same pattern may be applied in programming as in proving.  Axioms are relevant in programming, too, because, while Coq includes useful extensions like [Program] that make dependently typed programming more straightforward, in general these extensions generate code that relies on axioms about equality.  We can use clever pattern matching to write our code axiom-free.

999
As an example, consider a [Set] version of [fin_cases].  We use [Set] types instead of [Prop] types, so that return values have computational content and may be used to guide the behavior of algorithms.  Beside that, we are essentially writing the same "proof" in a more explicit way. *)
1000 1001 1002 1003 1004 1005

Definition finOut n (f : fin n) : match n return fin n -> Type with
                                    | O => fun _ => Empty_set
                                    | _ => fun f => {f' : _ | f = Next f'} + {f = First}
                                  end f :=
  match f with
1006 1007
    | First _ => inright _ (eq_refl _)
    | Next _ f' => inleft _ (exist _ f' (eq_refl _))
1008 1009 1010 1011
  end.
(* end thide *)

(** As another example, consider the following type of formulas in first-order logic.  The intent of the type definition will not be important in what follows, but we give a quick intuition for the curious reader.  Our formulas may include [forall] quantification over arbitrary [Type]s, and we index formulas by environments telling which variables are in scope and what their types are; such an environment is a [list Type].  A constructor [Inject] lets us include any Coq [Prop] as a formula, and [VarEq] and [Lift] can be used for variable references, in what is essentially the de Bruijn index convention.  (Again, the detail in this paragraph is not important to understand the discussion that follows!) *)
1012 1013 1014 1015 1016 1017 1018 1019 1020 1021 1022 1023 1024 1025

Inductive formula : list Type -> Type :=
| Inject : forall Ts, Prop -> formula Ts
| VarEq : forall T Ts, T -> formula (T :: Ts)
| Lift : forall T Ts, formula Ts -> formula (T :: Ts)
| Forall : forall T Ts, formula (T :: Ts) -> formula Ts
| And : forall Ts, formula Ts -> formula Ts -> formula Ts.

(** This example is based on my own experiences implementing variants of a program logic called XCAP%~\cite{XCAP}%, which also includes an inductive predicate for characterizing which formulas are provable.  Here I include a pared-down version of such a predicate, with only two constructors, which is sufficient to illustrate certain tricky issues. *)

Inductive proof : formula nil -> Prop :=
| PInject : forall (P : Prop), P -> proof (Inject nil P)
| PAnd : forall p q, proof p -> proof q -> proof (And p q).

1026
(** Let us prove a lemma showing that a "[P /\ Q -> P]" rule is derivable within the rules of [proof]. *)
1027 1028 1029 1030 1031 1032 1033 1034 1035 1036 1037 1038 1039 1040 1041 1042 1043 1044 1045 1046 1047 1048 1049 1050 1051 1052 1053 1054 1055 1056 1057 1058 1059 1060 1061 1062 1063 1064 1065 1066 1067 1068 1069 1070 1071 1072 1073 1074 1075 1076 1077 1078 1079 1080

Theorem proj1 : forall p q, proof (And p q) -> proof p.
  destruct 1.
(** %\vspace{-.15in}%[[
  p : formula nil
  q : formula nil
  P : Prop
  H : P
  ============================
   proof p
]]
*)

(** We are reminded that [induction] and [destruct] do not work effectively on types with non-variable arguments.  The first subgoal, shown above, is clearly unprovable.  (Consider the case where [p = Inject nil False].)

   An application of the %\index{tactics!dependent destruction}%[dependent destruction] tactic (the basis for [dep_destruct]) solves the problem handily.  We use a shorthand with the %\index{tactics!intros}%[intros] tactic that lets us use question marks for variable names that do not matter. *)

  Restart.
  Require Import Program.
  intros ? ? H; dependent destruction H; auto.
Qed.

Print Assumptions proj1.
(** %\vspace{-.15in}%[[
Axioms:
eq_rect_eq : forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
             x = eq_rect p Q x p h
]]

Unfortunately, that built-in tactic appeals to an axiom.  It is still possible to avoid axioms by giving the proof via another odd-looking lemma.  Here is a first attempt that fails at remaining axiom-free, using a common equality-based trick for supporting induction on non-variable arguments to type families.  The trick works fine without axioms for datatypes more traditional than [formula], but we run into trouble with our current type. *)

Lemma proj1_again' : forall r, proof r
  -> forall p q, r = And p q -> proof p.
  destruct 1; crush.
(** %\vspace{-.15in}%[[
  H0 : Inject [] P = And p q
  ============================
   proof p
]]

  The first goal looks reasonable.  Hypothesis [H0] is clearly contradictory, as [discriminate] can show. *)

  discriminate.
(** %\vspace{-.15in}%[[
  H : proof p
  H1 : And p q = And p0 q0
  ============================
   proof p0
]]

  It looks like we are almost done.  Hypothesis [H1] gives [p = p0] by injectivity of constructors, and then [H] finishes the case. *)

  injection H1; intros.

1081
(* begin hide *)
1082
(* begin thide *)
1083
  Definition existT' := existT.
1084
(* end thide *)
1085 1086 1087
(* end hide *)

(** Unfortunately, the "equality" that we expected between [p] and [p0] comes in a strange form:
1088 1089 1090 1091 1092 1093 1094 1095

[[
  H3 : existT (fun Ts : list Type => formula Ts) []%list p =
       existT (fun Ts : list Type => formula Ts) []%list p0
  ============================
   proof p0
]]

1096
It may take a bit of tinkering, but, reviewing Chapter 3's discussion of writing injection principles manually, it makes sense that an [existT] type is the most direct way to express the output of [injection] on a dependently typed constructor.  The constructor [And] is dependently typed, since it takes a parameter [Ts] upon which the types of [p] and [q] depend.  Let us not dwell further here on why this goal appears; the reader may like to attempt the (impossible) exercise of building a better injection lemma for [And], without using axioms.
1097 1098 1099 1100 1101 1102 1103 1104 1105 1106 1107 1108 1109 1110 1111 1112 1113 1114 1115 1116 1117 1118 1119 1120 1121 1122 1123 1124 1125 1126 1127 1128 1129 1130 1131 1132 1133

How exactly does an axiom come into the picture here?  Let us ask [crush] to finish the proof. *)

  crush.
Qed.

Print Assumptions proj1_again'.
(** %\vspace{-.15in}%[[
Axioms:
eq_rect_eq : forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
             x = eq_rect p Q x p h
]]

It turns out that this familiar axiom about equality (or some other axiom) is required to deduce [p = p0] from the hypothesis [H3] above.  The soundness of that proof step is neither provable nor disprovable in Gallina.

Hope is not lost, however.  We can produce an even stranger looking lemma, which gives us the theorem without axioms. *)

Lemma proj1_again'' : forall r, proof r
  -> match r with
       | And Ps p _ => match Ps return formula Ps -> Prop with
                         | nil => fun p => proof p
                         | _ => fun _ => True
                       end p
       | _ => True
     end.
  destruct 1; auto.
Qed.

Theorem proj1_again : forall p q, proof (And p q) -> proof p.
  intros ? ? H; exact (proj1_again'' H).
Qed.

Print Assumptions proj1_again.
(** <<
Closed under the global context
>>

1134 1135 1136 1137
This example illustrates again how some of the same design patterns we learned for dependently typed programming can be used fruitfully in theorem statements.

%\medskip%

1138
To close the chapter, we consider one final way to avoid dependence on axioms.  Often this task is equivalent to writing definitions such that they _compute_.  That is, we want Coq's normal reduction to be able to run certain programs to completion.  Here is a simple example where such computation can get stuck.  In proving properties of such functions, we would need to apply axioms like %\index{axiom K}%K manually to make progress.
1139 1140 1141 1142 1143 1144 1145 1146 1147 1148 1149 1150 1151 1152 1153 1154 1155 1156 1157 1158 1159 1160 1161 1162 1163 1164 1165 1166 1167 1168 1169 1170 1171 1172 1173 1174 1175 1176 1177

Imagine we are working with %\index{deep embedding}%deeply embedded syntax of some programming language, where each term is considered to be in the scope of a number of free variables that hold normal Coq values.  To enforce proper typing, we will need to model a Coq typing environment somehow.  One natural choice is as a list of types, where variable number [i] will be treated as a reference to the [i]th element of the list. *)

Section withTypes.
  Variable types : list Set.

  (** To give the semantics of terms, we will need to represent value environments, which assign each variable a term of the proper type. *)

  Variable values : hlist (fun x : Set => x) types.

  (** Now imagine that we are writing some procedure that operates on a distinguished variable of type [nat].  A hypothesis formalizes this assumption, using the standard library function [nth_error] for looking up list elements by position. *)

  Variable natIndex : nat.
  Variable natIndex_ok : nth_error types natIndex = Some nat.

  (** It is not hard to use this hypothesis to write a function for extracting the [nat] value in position [natIndex] of [values], starting with two helpful lemmas, each of which we finish with [Defined] to mark the lemma as transparent, so that its definition may be expanded during evaluation. *)

  Lemma nth_error_nil : forall A n x,
    nth_error (@nil A) n = Some x
    -> False.
    destruct n; simpl; unfold error; congruence.
  Defined.

  Implicit Arguments nth_error_nil [A n x].

  Lemma Some_inj : forall A (x y : A),
    Some x = Some y
    -> x = y.
    congruence.
  Defined.

  Fixpoint getNat (types' : list Set) (values' : hlist (fun x : Set => x) types')
    (natIndex : nat) : (nth_error types' natIndex = Some nat) -> nat :=
    match values' with
      | HNil => fun pf => match nth_error_nil pf with end
      | HCons t ts x values'' =>
        match natIndex return nth_error (t :: ts) natIndex = Some nat -> nat with
          | O => fun pf =>
            match Some_inj pf in _ = T return T with
1178
              | eq_refl => x
1179 1180 1181 1182 1183 1184 1185 1186 1187 1188 1189 1190 1191 1192 1193 1194 1195 1196 1197 1198 1199 1200 1201
            end
          | S natIndex' => getNat values'' natIndex'
        end
    end.
End withTypes.

(** The problem becomes apparent when we experiment with running [getNat] on a concrete [types] list. *)

Definition myTypes := unit :: nat :: bool :: nil.
Definition myValues : hlist (fun x : Set => x) myTypes :=
  tt ::: 3 ::: false ::: HNil.

Definition myNatIndex := 1.

Theorem myNatIndex_ok : nth_error myTypes myNatIndex = Some nat.
  reflexivity.
Defined.

Eval compute in getNat myValues myNatIndex myNatIndex_ok.
(** %\vspace{-.15in}%[[
     = 3
]]

1202
We have not hit the problem yet, since we proceeded with a concrete equality proof for [myNatIndex_ok].  However, consider a case where we want to reason about the behavior of [getNat] _independently_ of a specific proof. *)
1203 1204 1205 1206 1207 1208 1209 1210 1211 1212 1213 1214 1215 1216 1217 1218 1219 1220 1221 1222 1223 1224 1225 1226 1227 1228 1229 1230 1231 1232 1233 1234 1235 1236 1237 1238 1239 1240 1241 1242 1243 1244 1245 1246 1247 1248 1249 1250 1251 1252 1253 1254

Theorem getNat_is_reasonable : forall pf, getNat myValues myNatIndex pf = 3.
  intro; compute.
(**
<<
1 subgoal
>>
%\vspace{-.3in}%[[
  pf : nth_error myTypes myNatIndex = Some nat
  ============================
   match
     match
       pf in (_ = y)
       return (nat = match y with
                     | Some H => H
                     | None => nat
                     end)
     with
     | eq_refl => eq_refl
     end in (_ = T) return T
   with
   | eq_refl => 3
   end = 3
]]

Since the details of the equality proof [pf] are not known, computation can proceed no further.  A rewrite with axiom K would allow us to make progress, but we can rethink the definitions a bit to avoid depending on axioms. *)

Abort.

(** Here is a definition of a function that turns out to be useful, though no doubt its purpose will be mysterious for now.  A call [update ls n x] overwrites the [n]th position of the list [ls] with the value [x], padding the end of the list with extra [x] values as needed to ensure sufficient length. *)

Fixpoint copies A (x : A) (n : nat) : list A :=
  match n with
    | O => nil
    | S n' => x :: copies x n'
  end.

Fixpoint update A (ls : list A) (n : nat) (x : A) : list A :=
  match ls with
    | nil => copies x n ++ x :: nil
    | y :: ls' => match n with
                    | O => x :: ls'
                    | S n' => y :: update ls' n' x
                  end
  end.

(** Now let us revisit the definition of [getNat]. *)

Section withTypes'.
  Variable types : list Set.
  Variable natIndex : nat.

1255
  (** Here is the trick: instead of asserting properties about the list [types], we build a "new" list that is _guaranteed by construction_ to have those properties. *)
1256 1257 1258 1259 1260 1261 1262

  Definition types' := update types natIndex nat.

  Variable values : hlist (fun x : Set => x) types'.

  (** Now a bit of dependent pattern matching helps us rewrite [getNat] in a way that avoids any use of equality proofs. *)

1263 1264 1265 1266 1267 1268 1269
  Fixpoint skipCopies (n : nat)
    : hlist (fun x : Set => x) (copies nat n ++ nat :: nil) -> nat :=
    match n with
      | O => fun vs => hhd vs
      | S n' => fun vs => skipCopies n' (htl vs)
    end.

1270 1271 1272
  Fixpoint getNat' (types'' : list Set) (natIndex : nat)
    : hlist (fun x : Set => x) (update types'' natIndex nat) -> nat :=
    match types'' with
1273
      | nil => skipCopies natIndex
1274 1275 1276 1277 1278 1279 1280 1281 1282
      | t :: types0 =>
        match natIndex return hlist (fun x : Set => x)
          (update (t :: types0) natIndex nat) -> nat with
          | O => fun vs => hhd vs
          | S natIndex' => fun vs => getNat' types0 natIndex' (htl vs)
        end
    end.
End withTypes'.

1283
(** Now the surprise comes in how easy it is to _use_ [getNat'].  While typing works by modification of a types list, we can choose parameters so that the modification has no effect. *)
1284 1285 1286 1287 1288 1289

Theorem getNat_is_reasonable : getNat' myTypes myNatIndex myValues = 3.
  reflexivity.
Qed.

(** The same parameters as before work without alteration, and we avoid use of axioms. *)