annotate src/Universes.v @ 378:6413675f8e04

Recursing under binders in reification
author Adam Chlipala <adam@chlipala.net>
date Thu, 29 Mar 2012 17:13:23 -0400
parents 252ba054a1cd
children 057a29f9c773
rev   line source
adam@377 1 (* Copyright (c) 2009-2012, Adam Chlipala
adamc@227 2 *
adamc@227 3 * This work is licensed under a
adamc@227 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@227 5 * Unported License.
adamc@227 6 * The license text is available at:
adamc@227 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@227 8 *)
adamc@227 9
adamc@227 10 (* begin hide *)
adam@377 11 Require Import List.
adam@377 12
adam@314 13 Require Import DepList CpdtTactics.
adamc@227 14
adamc@227 15 Set Implicit Arguments.
adamc@227 16 (* end hide *)
adamc@227 17
adamc@227 18
adamc@227 19 (** %\chapter{Universes and Axioms}% *)
adamc@227 20
adam@343 21 (** 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.
adamc@227 22
adam@343 23 %\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. *)
adamc@227 24
adamc@227 25
adamc@227 26 (** * The [Type] Hierarchy *)
adamc@227 27
adam@343 28 (** %\index{type hierarchy}%Every object in Gallina has a type. *)
adamc@227 29
adamc@227 30 Check 0.
adamc@227 31 (** %\vspace{-.15in}% [[
adamc@227 32 0
adamc@227 33 : nat
adamc@227 34
adamc@227 35 ]]
adamc@227 36
adamc@227 37 It is natural enough that zero be considered as a natural number. *)
adamc@227 38
adamc@227 39 Check nat.
adamc@227 40 (** %\vspace{-.15in}% [[
adamc@227 41 nat
adamc@227 42 : Set
adamc@227 43
adamc@227 44 ]]
adamc@227 45
adam@287 46 From a set theory perspective, it is unsurprising to consider the natural numbers as a %``%#"#set.#"#%''% *)
adamc@227 47
adamc@227 48 Check Set.
adamc@227 49 (** %\vspace{-.15in}% [[
adamc@227 50 Set
adamc@227 51 : Type
adamc@227 52
adamc@227 53 ]]
adamc@227 54
adam@343 55 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)}\textit{%#<i>#classes#</i>#%}%. In Coq, this more general notion is [Type]. *)
adamc@227 56
adamc@227 57 Check Type.
adamc@227 58 (** %\vspace{-.15in}% [[
adamc@227 59 Type
adamc@227 60 : Type
adamc@227 61
adamc@227 62 ]]
adamc@227 63
adam@343 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?
adamc@227 65
adam@343 66 Let us repeat some of our queries after toggling a flag related to Coq's printing behavior.%\index{Vernacular commands!Set Printing Universes}% *)
adamc@227 67
adamc@227 68 Set Printing Universes.
adamc@227 69
adamc@227 70 Check nat.
adamc@227 71 (** %\vspace{-.15in}% [[
adamc@227 72 nat
adamc@227 73 : Set
adam@302 74 ]]
adam@302 75 *)
adamc@227 76
adamc@227 77 (** printing $ %({}*% #(<a/>*# *)
adamc@227 78 (** printing ^ %*{})% #*<a/>)# *)
adamc@227 79
adamc@227 80 Check Set.
adamc@227 81 (** %\vspace{-.15in}% [[
adamc@227 82 Set
adamc@227 83 : Type $ (0)+1 ^
adamc@227 84
adam@302 85 ]]
adam@302 86 *)
adamc@227 87
adamc@227 88 Check Type.
adamc@227 89 (** %\vspace{-.15in}% [[
adamc@227 90 Type $ Top.3 ^
adamc@227 91 : Type $ (Top.3)+1 ^
adamc@227 92
adamc@227 93 ]]
adamc@227 94
adam@287 95 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].
adamc@227 96
adamc@227 97 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 %\textit{%#<i>#classifies#</i>#%}% [Set].
adamc@227 98
adam@343 99 In the second query's output, we see that the occurrence of [Type] that we check is assigned a fresh %\index{universe variable}\textit{%#<i>#universe variable#</i>#%}% [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.
adamc@227 100
adam@343 101 Another crucial concept in CIC is %\index{predicativity}\textit{%#<i>#predicativity#</i>#%}%. Consider these queries. *)
adamc@227 102
adamc@227 103 Check forall T : nat, fin T.
adamc@227 104 (** %\vspace{-.15in}% [[
adamc@227 105 forall T : nat, fin T
adamc@227 106 : Set
adam@302 107 ]]
adam@302 108 *)
adamc@227 109
adamc@227 110 Check forall T : Set, T.
adamc@227 111 (** %\vspace{-.15in}% [[
adamc@227 112 forall T : Set, T
adamc@227 113 : Type $ max(0, (0)+1) ^
adam@302 114 ]]
adam@302 115 *)
adamc@227 116
adamc@227 117 Check forall T : Type, T.
adamc@227 118 (** %\vspace{-.15in}% [[
adamc@227 119 forall T : Type $ Top.9 ^ , T
adamc@227 120 : Type $ max(Top.9, (Top.9)+1) ^
adamc@227 121
adamc@227 122 ]]
adamc@227 123
adamc@227 124 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.
adamc@227 125
adam@287 126 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]. *)
adamc@227 127
adamc@227 128 Definition id (T : Set) (x : T) : T := x.
adamc@227 129
adamc@227 130 Check id 0.
adamc@227 131 (** %\vspace{-.15in}% [[
adamc@227 132 id 0
adamc@227 133 : nat
adamc@227 134
adamc@227 135 Check id Set.
adam@343 136 ]]
adamc@227 137
adam@343 138 <<
adamc@227 139 Error: Illegal application (Type Error):
adamc@227 140 ...
adam@343 141 The 1st term has type "Type (* (Top.15)+1 *)" which should be coercible to "Set".
adam@343 142 >>
adamc@227 143
adam@343 144 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]. *)
adamc@227 145
adamc@227 146 Reset id.
adamc@227 147 Definition id (T : Type) (x : T) : T := x.
adamc@227 148 Check id 0.
adamc@227 149 (** %\vspace{-.15in}% [[
adamc@227 150 id 0
adamc@227 151 : nat
adam@302 152 ]]
adam@302 153 *)
adamc@227 154
adamc@227 155 Check id Set.
adamc@227 156 (** %\vspace{-.15in}% [[
adamc@227 157 id Set
adamc@227 158 : Type $ Top.17 ^
adam@302 159 ]]
adam@302 160 *)
adamc@227 161
adamc@227 162 Check id Type.
adamc@227 163 (** %\vspace{-.15in}% [[
adamc@227 164 id Type $ Top.18 ^
adamc@227 165 : Type $ Top.19 ^
adam@302 166 ]]
adam@302 167 *)
adamc@227 168
adamc@227 169 (** 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.
adamc@227 170 [[
adamc@227 171 Check id id.
adam@343 172 ]]
adamc@227 173
adam@343 174 <<
adamc@227 175 Error: Universe inconsistency (cannot enforce Top.16 < Top.16).
adam@343 176 >>
adamc@227 177
adam@343 178 %\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 %\textit{%#<i>#predicative#</i>#%}%, 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. *)
adamc@227 179
adamc@227 180
adamc@227 181 (** ** Inductive Definitions *)
adamc@227 182
adamc@227 183 (** 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].
adamc@227 184
adamc@227 185 [[
adamc@227 186 Inductive exp : Set -> Set :=
adamc@227 187 | Const : forall T : Set, T -> exp T
adamc@227 188 | Pair : forall T1 T2, exp T1 -> exp T2 -> exp (T1 * T2)
adamc@227 189 | Eq : forall T, exp T -> exp T -> exp bool.
adam@343 190 ]]
adamc@227 191
adam@343 192 <<
adamc@227 193 Error: Large non-propositional inductive types must be in Type.
adam@343 194 >>
adamc@227 195
adam@343 196 This definition is %\index{large inductive types}\textit{%#<i>#large#</i>#%}% 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. *)
adamc@227 197
adamc@227 198 Inductive exp : Type -> Type :=
adamc@227 199 | Const : forall T, T -> exp T
adamc@227 200 | Pair : forall T1 T2, exp T1 -> exp T2 -> exp (T1 * T2)
adamc@227 201 | Eq : forall T, exp T -> exp T -> exp bool.
adamc@227 202
adamc@228 203 (** 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].
adamc@228 204
adamc@228 205 Our new definition is accepted. We can build some sample expressions. *)
adamc@227 206
adamc@227 207 Check Const 0.
adamc@227 208 (** %\vspace{-.15in}% [[
adamc@227 209 Const 0
adamc@227 210 : exp nat
adam@302 211 ]]
adam@302 212 *)
adamc@227 213
adamc@227 214 Check Pair (Const 0) (Const tt).
adamc@227 215 (** %\vspace{-.15in}% [[
adamc@227 216 Pair (Const 0) (Const tt)
adamc@227 217 : exp (nat * unit)
adam@302 218 ]]
adam@302 219 *)
adamc@227 220
adamc@227 221 Check Eq (Const Set) (Const Type).
adamc@227 222 (** %\vspace{-.15in}% [[
adamc@228 223 Eq (Const Set) (Const Type $ Top.59 ^ )
adamc@227 224 : exp bool
adamc@227 225
adamc@227 226 ]]
adamc@227 227
adamc@227 228 We can check many expressions, including fancy expressions that include types. However, it is not hard to hit a type-checking wall.
adamc@227 229
adamc@227 230 [[
adamc@227 231 Check Const (Const O).
adam@343 232 ]]
adamc@227 233
adam@343 234 <<
adamc@227 235 Error: Universe inconsistency (cannot enforce Top.42 < Top.42).
adam@343 236 >>
adamc@227 237
adamc@227 238 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. *)
adamc@227 239
adamc@227 240 Print exp.
adamc@227 241 (** %\vspace{-.15in}% [[
adamc@227 242 Inductive exp
adamc@227 243 : Type $ Top.8 ^ ->
adamc@227 244 Type
adamc@227 245 $ max(0, (Top.11)+1, (Top.14)+1, (Top.15)+1, (Top.19)+1) ^ :=
adamc@227 246 Const : forall T : Type $ Top.11 ^ , T -> exp T
adamc@227 247 | Pair : forall (T1 : Type $ Top.14 ^ ) (T2 : Type $ Top.15 ^ ),
adamc@227 248 exp T1 -> exp T2 -> exp (T1 * T2)
adamc@227 249 | Eq : forall T : Type $ Top.19 ^ , exp T -> exp T -> exp bool
adamc@227 250
adamc@227 251 ]]
adamc@227 252
adamc@227 253 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] %\textit{%#<i>#must#</i>#%}% 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.
adamc@227 254
adam@343 255 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}% *)
adamc@227 256
adamc@227 257 Print Universes.
adamc@227 258 (** %\vspace{-.15in}% [[
adamc@227 259 Top.19 < Top.9 <= Top.8
adamc@227 260 Top.15 < Top.9 <= Top.8 <= Coq.Init.Datatypes.38
adamc@227 261 Top.14 < Top.9 <= Top.8 <= Coq.Init.Datatypes.37
adamc@227 262 Top.11 < Top.9 <= Top.8
adamc@227 263
adamc@227 264 ]]
adamc@227 265
adam@343 266 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.
adamc@227 267
adamc@227 268 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. *)
adamc@227 269
adamc@227 270 Print prod.
adamc@227 271 (** %\vspace{-.15in}% [[
adamc@227 272 Inductive prod (A : Type $ Coq.Init.Datatypes.37 ^ )
adamc@227 273 (B : Type $ Coq.Init.Datatypes.38 ^ )
adamc@227 274 : Type $ max(Coq.Init.Datatypes.37, Coq.Init.Datatypes.38) ^ :=
adamc@227 275 pair : A -> B -> A * B
adamc@227 276
adamc@227 277 ]]
adamc@227 278
adamc@227 279 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].
adamc@227 280
adamc@227 281 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.
adamc@227 282
adamc@227 283 %\medskip%
adamc@227 284
adamc@227 285 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 %\textit{%#<i>#one higher#</i>#%}% than the maximum. The critical difference is that, in the definition of [prod], [A] and [B] are defined as %\textit{%#<i>#parameters#</i>#%}%; that is, they appear named to the left of the main colon, rather than appearing (possibly unnamed) to the right.
adamc@227 286
adamc@231 287 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. *)
adamc@227 288
adamc@227 289 Check (nat, (Type, Set)).
adamc@227 290 (** %\vspace{-.15in}% [[
adamc@227 291 (nat, (Type $ Top.44 ^ , Set))
adamc@227 292 : Set * (Type $ Top.45 ^ * Type $ Top.46 ^ )
adamc@227 293 ]]
adamc@227 294
adamc@227 295 The same cannot be done with a counterpart to [prod] that does not use parameters. *)
adamc@227 296
adamc@227 297 Inductive prod' : Type -> Type -> Type :=
adamc@227 298 | pair' : forall A B : Type, A -> B -> prod' A B.
adamc@227 299 (** [[
adamc@227 300 Check (pair' nat (pair' Type Set)).
adam@343 301 ]]
adamc@227 302
adam@343 303 <<
adamc@227 304 Error: Universe inconsistency (cannot enforce Top.51 < Top.51).
adam@343 305 >>
adamc@227 306
adamc@233 307 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.
adamc@233 308
adam@343 309 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. *)
adamc@233 310
adamc@233 311 Inductive foo (A : Type) : Type :=
adamc@233 312 | Foo : A -> foo A.
adamc@229 313
adamc@229 314 (* begin hide *)
adamc@229 315 Unset Printing Universes.
adamc@229 316 (* end hide *)
adamc@229 317
adamc@233 318 Check foo nat.
adamc@233 319 (** %\vspace{-.15in}% [[
adamc@233 320 foo nat
adamc@233 321 : Set
adam@302 322 ]]
adam@302 323 *)
adamc@233 324
adamc@233 325 Check foo Set.
adamc@233 326 (** %\vspace{-.15in}% [[
adamc@233 327 foo Set
adamc@233 328 : Type
adam@302 329 ]]
adam@302 330 *)
adamc@233 331
adamc@233 332 Check foo True.
adamc@233 333 (** %\vspace{-.15in}% [[
adamc@233 334 foo True
adamc@233 335 : Prop
adamc@233 336
adamc@233 337 ]]
adamc@233 338
adam@287 339 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].
adamc@233 340
adamc@233 341 Imitation polymorphism can be confusing in some contexts. For instance, it is what is responsible for this weird behavior. *)
adamc@233 342
adamc@233 343 Inductive bar : Type := Bar : bar.
adamc@233 344
adamc@233 345 Check bar.
adamc@233 346 (** %\vspace{-.15in}% [[
adamc@233 347 bar
adamc@233 348 : Prop
adamc@233 349 ]]
adamc@233 350
adamc@233 351 The type that Coq comes up with may be used in strictly more contexts than the type one might have expected. *)
adamc@233 352
adamc@229 353
adamc@229 354 (** * The [Prop] Universe *)
adamc@229 355
adam@287 356 (** 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.
adamc@229 357
adamc@229 358 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]. *)
adamc@229 359
adamc@229 360 Print sig.
adamc@229 361 (** %\vspace{-.15in}% [[
adamc@229 362 Inductive sig (A : Type) (P : A -> Prop) : Type :=
adamc@229 363 exist : forall x : A, P x -> sig P
adam@302 364 ]]
adam@302 365 *)
adamc@229 366
adamc@229 367 Print ex.
adamc@229 368 (** %\vspace{-.15in}% [[
adamc@229 369 Inductive ex (A : Type) (P : A -> Prop) : Prop :=
adamc@229 370 ex_intro : forall x : A, P x -> ex P
adamc@229 371 ]]
adamc@229 372
adamc@229 373 It is natural to want a function to extract the first components of data structures like these. Doing so is easy enough for [sig]. *)
adamc@229 374
adamc@229 375 Definition projS A (P : A -> Prop) (x : sig P) : A :=
adamc@229 376 match x with
adamc@229 377 | exist v _ => v
adamc@229 378 end.
adamc@229 379
adamc@229 380 (** We run into trouble with a version that has been changed to work with [ex].
adamc@229 381 [[
adamc@229 382 Definition projE A (P : A -> Prop) (x : ex P) : A :=
adamc@229 383 match x with
adamc@229 384 | ex_intro v _ => v
adamc@229 385 end.
adam@343 386 ]]
adamc@229 387
adam@343 388 <<
adamc@229 389 Error:
adamc@229 390 Incorrect elimination of "x" in the inductive type "ex":
adamc@229 391 the return type has sort "Type" while it should be "Prop".
adamc@229 392 Elimination of an inductive object of sort Prop
adamc@229 393 is not allowed on a predicate in sort Type
adamc@229 394 because proofs can be eliminated only to build proofs.
adam@343 395 >>
adamc@229 396
adam@343 397 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.
adamc@229 398
adamc@229 399 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.
adamc@229 400
adam@343 401 Recall that %\index{program extraction}%extraction is Coq's facility for translating Coq developments into programs in general-purpose programming languages like OCaml. Extraction %\textit{%#<i>#erases#</i>#%}% proofs and leaves programs intact. A simple example with [sig] and [ex] demonstrates the distinction. *)
adamc@229 402
adamc@229 403 Definition sym_sig (x : sig (fun n => n = 0)) : sig (fun n => 0 = n) :=
adamc@229 404 match x with
adamc@229 405 | exist n pf => exist _ n (sym_eq pf)
adamc@229 406 end.
adamc@229 407
adamc@229 408 Extraction sym_sig.
adamc@229 409 (** <<
adamc@229 410 (** val sym_sig : nat -> nat **)
adamc@229 411
adamc@229 412 let sym_sig x = x
adamc@229 413 >>
adamc@229 414
adamc@229 415 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. *)
adamc@229 416
adamc@229 417 Definition sym_ex (x : ex (fun n => n = 0)) : ex (fun n => 0 = n) :=
adamc@229 418 match x with
adamc@229 419 | ex_intro n pf => ex_intro _ n (sym_eq pf)
adamc@229 420 end.
adamc@229 421
adamc@229 422 Extraction sym_ex.
adamc@229 423 (** <<
adamc@229 424 (** val sym_ex : __ **)
adamc@229 425
adamc@229 426 let sym_ex = __
adamc@229 427 >>
adamc@229 428
adam@302 429 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 %\texttt{\_\_}%#<tt>__</tt>#, whose single constructor is %\texttt{\_\_}%#<tt>__</tt>#. Not only are proofs replaced by [__], but proof arguments to functions are also removed completely, as we see here.
adamc@229 430
adam@343 431 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.
adamc@229 432
adam@343 433 Many fans of the %\index{Curry-Howard correspondence}%Curry-Howard correspondence support the idea of %\textit{%#<i>#extracting programs from proofs#</i>#%}%. 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.
adamc@229 434
adamc@229 435 %\medskip%
adamc@229 436
adam@343 437 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}\textit{%#<i>#impredicative#</i>#%}%, as this example shows. *)
adamc@229 438
adamc@229 439 Check forall P Q : Prop, P \/ Q -> Q \/ P.
adamc@229 440 (** %\vspace{-.15in}% [[
adamc@229 441 forall P Q : Prop, P \/ Q -> Q \/ P
adamc@229 442 : Prop
adamc@229 443
adamc@229 444 ]]
adamc@229 445
adamc@230 446 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.
adamc@230 447
adamc@230 448 Impredicativity also allows us to implement a version of our earlier [exp] type that does not suffer from the weakness that we found. *)
adamc@230 449
adamc@230 450 Inductive expP : Type -> Prop :=
adamc@230 451 | ConstP : forall T, T -> expP T
adamc@230 452 | PairP : forall T1 T2, expP T1 -> expP T2 -> expP (T1 * T2)
adamc@230 453 | EqP : forall T, expP T -> expP T -> expP bool.
adamc@230 454
adamc@230 455 Check ConstP 0.
adamc@230 456 (** %\vspace{-.15in}% [[
adamc@230 457 ConstP 0
adamc@230 458 : expP nat
adam@302 459 ]]
adam@302 460 *)
adamc@230 461
adamc@230 462 Check PairP (ConstP 0) (ConstP tt).
adamc@230 463 (** %\vspace{-.15in}% [[
adamc@230 464 PairP (ConstP 0) (ConstP tt)
adamc@230 465 : expP (nat * unit)
adam@302 466 ]]
adam@302 467 *)
adamc@230 468
adamc@230 469 Check EqP (ConstP Set) (ConstP Type).
adamc@230 470 (** %\vspace{-.15in}% [[
adamc@230 471 EqP (ConstP Set) (ConstP Type)
adamc@230 472 : expP bool
adam@302 473 ]]
adam@302 474 *)
adamc@230 475
adamc@230 476 Check ConstP (ConstP O).
adamc@230 477 (** %\vspace{-.15in}% [[
adamc@230 478 ConstP (ConstP 0)
adamc@230 479 : expP (expP nat)
adamc@230 480
adamc@230 481 ]]
adamc@230 482
adam@287 483 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 [=]. *)
adamc@230 484
adamc@230 485 Inductive eqPlus : forall T, T -> T -> Prop :=
adamc@230 486 | Base : forall T (x : T), eqPlus x x
adamc@230 487 | Func : forall dom ran (f1 f2 : dom -> ran),
adamc@230 488 (forall x : dom, eqPlus (f1 x) (f2 x))
adamc@230 489 -> eqPlus f1 f2.
adamc@230 490
adamc@230 491 Check (Base 0).
adamc@230 492 (** %\vspace{-.15in}% [[
adamc@230 493 Base 0
adamc@230 494 : eqPlus 0 0
adam@302 495 ]]
adam@302 496 *)
adamc@230 497
adamc@230 498 Check (Func (fun n => n) (fun n => 0 + n) (fun n => Base n)).
adamc@230 499 (** %\vspace{-.15in}% [[
adamc@230 500 Func (fun n : nat => n) (fun n : nat => 0 + n) (fun n : nat => Base n)
adamc@230 501 : eqPlus (fun n : nat => n) (fun n : nat => 0 + n)
adam@302 502 ]]
adam@302 503 *)
adamc@230 504
adamc@230 505 Check (Base (Base 1)).
adamc@230 506 (** %\vspace{-.15in}% [[
adamc@230 507 Base (Base 1)
adamc@230 508 : eqPlus (Base 1) (Base 1)
adam@302 509 ]]
adam@302 510 *)
adamc@230 511
adam@343 512 (** Stating equality facts about proofs may seem baroque, but we have already seen its utility in the chapter on reasoning about equality proofs. *)
adam@343 513
adamc@230 514
adamc@230 515 (** * Axioms *)
adamc@230 516
adam@343 517 (** 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}\textit{%#<i>#axioms#</i>#%}% without proof.
adamc@230 518
adamc@230 519 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. *)
adamc@230 520
adamc@230 521 (** ** The Basics *)
adamc@230 522
adam@343 523 (** One simple example of a useful axiom is the %\index{law of the excluded middle}%law of the excluded middle. *)
adamc@230 524
adamc@230 525 Require Import Classical_Prop.
adamc@230 526 Print classic.
adamc@230 527 (** %\vspace{-.15in}% [[
adamc@230 528 *** [ classic : forall P : Prop, P \/ ~ P ]
adamc@230 529 ]]
adamc@230 530
adam@343 531 In the implementation of module [Classical_Prop], this axiom was defined with the command%\index{Vernacular commands!Axiom}% *)
adamc@230 532
adamc@230 533 Axiom classic : forall P : Prop, P \/ ~ P.
adamc@230 534
adam@343 535 (** 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. *)
adamc@230 536
adamc@230 537 Parameter n : nat.
adamc@230 538 Axiom positive : n > 0.
adamc@230 539 Reset n.
adamc@230 540
adam@287 541 (** 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.
adamc@230 542
adam@343 543 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}\textit{%#<i>#inconsistent#</i>#%}%. That is, a set of axioms may imply [False], which allows any theorem to 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]. *)
adamc@230 544
adam@287 545 Axiom not_classic : ~ forall P : Prop, P \/ ~ P.
adamc@230 546
adamc@230 547 Theorem uhoh : False.
adam@287 548 generalize classic not_classic; tauto.
adamc@230 549 Qed.
adamc@230 550
adamc@230 551 Theorem uhoh_again : 1 + 1 = 3.
adamc@230 552 destruct uhoh.
adamc@230 553 Qed.
adamc@230 554
adamc@230 555 Reset not_classic.
adamc@230 556
adam@343 557 (** 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 %\textit{%#<i>#model#</i>#%}% 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.
adamc@230 558
adam@343 559 Recall that Coq implements %\index{constructive logic}\textit{%#<i>#constructive#</i>#%}% logic by default, where 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.
adamc@230 560
adamc@231 561 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] %\textit{%#<i>#would#</i>#%}% 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.
adamc@230 562
adam@343 563 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}% *)
adamc@230 564
adamc@230 565 Theorem t1 : forall P : Prop, P -> ~ ~ P.
adamc@230 566 tauto.
adamc@230 567 Qed.
adamc@230 568
adamc@230 569 Print Assumptions t1.
adam@343 570 (** <<
adamc@230 571 Closed under the global context
adam@343 572 >>
adam@302 573 *)
adamc@230 574
adamc@230 575 Theorem t2 : forall P : Prop, ~ ~ P -> P.
adamc@230 576 (** [[
adamc@230 577 tauto.
adam@343 578 ]]
adam@343 579 <<
adamc@230 580 Error: tauto failed.
adam@343 581 >>
adam@302 582 *)
adamc@230 583 intro P; destruct (classic P); tauto.
adamc@230 584 Qed.
adamc@230 585
adamc@230 586 Print Assumptions t2.
adamc@230 587 (** %\vspace{-.15in}% [[
adamc@230 588 Axioms:
adamc@230 589 classic : forall P : Prop, P \/ ~ P
adamc@230 590 ]]
adamc@230 591
adamc@231 592 It is possible to avoid this dependence in some specific cases, where excluded middle %\textit{%#<i>#is#</i>#%}% provable, for decidable families of propositions. *)
adamc@230 593
adam@287 594 Theorem nat_eq_dec : forall n m : nat, n = m \/ n <> m.
adamc@230 595 induction n; destruct m; intuition; generalize (IHn m); intuition.
adamc@230 596 Qed.
adamc@230 597
adamc@230 598 Theorem t2' : forall n m : nat, ~ ~ (n = m) -> n = m.
adam@287 599 intros n m; destruct (nat_eq_dec n m); tauto.
adamc@230 600 Qed.
adamc@230 601
adamc@230 602 Print Assumptions t2'.
adam@343 603 (** <<
adamc@230 604 Closed under the global context
adam@343 605 >>
adamc@230 606
adamc@230 607 %\bigskip%
adamc@230 608
adam@343 609 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}\textit{%#<i>#proof irrelevance#</i>#%}%, which simplifies proof issues that would not even arise in mainstream math. *)
adamc@230 610
adamc@230 611 Require Import ProofIrrelevance.
adamc@230 612 Print proof_irrelevance.
adamc@230 613 (** %\vspace{-.15in}% [[
adamc@230 614 *** [ proof_irrelevance : forall (P : Prop) (p1 p2 : P), p1 = p2 ]
adamc@230 615 ]]
adamc@230 616
adam@353 617 This axiom asserts that any two proofs of the same proposition are equal. If we replaced [p1 = p2] by [p1 <-> p2], then the statement would be provable. However, equality is a stronger notion than logical equivalence. Recall this example function from Chapter 6. *)
adamc@230 618
adamc@230 619 (* begin hide *)
adamc@230 620 Lemma zgtz : 0 > 0 -> False.
adamc@230 621 crush.
adamc@230 622 Qed.
adamc@230 623 (* end hide *)
adamc@230 624
adamc@230 625 Definition pred_strong1 (n : nat) : n > 0 -> nat :=
adamc@230 626 match n with
adamc@230 627 | O => fun pf : 0 > 0 => match zgtz pf with end
adamc@230 628 | S n' => fun _ => n'
adamc@230 629 end.
adamc@230 630
adam@343 631 (** We might want to prove that different proofs of [n > 0] do not lead to different results from our richly typed predecessor function. *)
adamc@230 632
adamc@230 633 Theorem pred_strong1_irrel : forall n (pf1 pf2 : n > 0), pred_strong1 pf1 = pred_strong1 pf2.
adamc@230 634 destruct n; crush.
adamc@230 635 Qed.
adamc@230 636
adamc@230 637 (** 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]. *)
adamc@230 638
adamc@230 639 Theorem pred_strong1_irrel' : forall n (pf1 pf2 : n > 0), pred_strong1 pf1 = pred_strong1 pf2.
adamc@230 640 intros; f_equal; apply proof_irrelevance.
adamc@230 641 Qed.
adamc@230 642
adamc@230 643
adamc@230 644 (** %\bigskip%
adamc@230 645
adamc@230 646 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: *)
adamc@230 647
adamc@230 648 Require Import Eqdep.
adamc@230 649 Import Eq_rect_eq.
adamc@230 650 Print eq_rect_eq.
adamc@230 651 (** %\vspace{-.15in}% [[
adamc@230 652 *** [ eq_rect_eq :
adamc@230 653 forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
adamc@230 654 x = eq_rect p Q x p h ]
adamc@230 655 ]]
adamc@230 656
adam@343 657 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.#"#%''% *)
adamc@230 658
adamc@230 659 Corollary UIP_refl : forall A (x : A) (pf : x = x), pf = refl_equal x.
adamc@230 660 intros; replace pf with (eq_rect x (eq x) (refl_equal x) x pf); [
adamc@230 661 symmetry; apply eq_rect_eq
adamc@230 662 | exact (match pf as pf' return match pf' in _ = y return x = y with
adamc@230 663 | refl_equal => refl_equal x
adamc@230 664 end = pf' with
adamc@230 665 | refl_equal => refl_equal _
adamc@230 666 end) ].
adamc@230 667 Qed.
adamc@230 668
adamc@230 669 Corollary UIP : forall A (x y : A) (pf1 pf2 : x = y), pf1 = pf2.
adamc@230 670 intros; generalize pf1 pf2; subst; intros;
adamc@230 671 match goal with
adamc@230 672 | [ |- ?pf1 = ?pf2 ] => rewrite (UIP_refl pf1); rewrite (UIP_refl pf2); reflexivity
adamc@230 673 end.
adamc@230 674 Qed.
adamc@230 675
adamc@231 676 (** 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.
adamc@230 677
adamc@230 678 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.
adamc@230 679
adamc@230 680 %\bigskip%
adamc@230 681
adamc@230 682 There are two more basic axioms that are often assumed, to avoid complications that do not arise in set theory. *)
adamc@230 683
adamc@230 684 Require Import FunctionalExtensionality.
adamc@230 685 Print functional_extensionality_dep.
adamc@230 686 (** %\vspace{-.15in}% [[
adamc@230 687 *** [ functional_extensionality_dep :
adamc@230 688 forall (A : Type) (B : A -> Type) (f g : forall x : A, B x),
adamc@230 689 (forall x : A, f x = g x) -> f = g ]
adamc@230 690
adamc@230 691 ]]
adamc@230 692
adamc@230 693 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.
adamc@230 694
adam@343 695 A simple corollary shows that the same property applies to predicates. *)
adamc@230 696
adamc@230 697 Corollary predicate_extensionality : forall (A : Type) (B : A -> Prop) (f g : forall x : A, B x),
adamc@230 698 (forall x : A, f x = g x) -> f = g.
adamc@230 699 intros; apply functional_extensionality_dep; assumption.
adamc@230 700 Qed.
adamc@230 701
adam@343 702 (** In some cases, one might prefer to assert this corollary as the axiom, to restrict the consequences to proofs and not programs. *)
adam@343 703
adamc@230 704
adamc@230 705 (** ** Axioms of Choice *)
adamc@230 706
adam@343 707 (** 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.
adamc@230 708
adamc@230 709 First, it is possible to implement a choice operator %\textit{%#<i>#without#</i>#%}% axioms in some potentially surprising cases. *)
adamc@230 710
adamc@230 711 Require Import ConstructiveEpsilon.
adamc@230 712 Check constructive_definite_description.
adamc@230 713 (** %\vspace{-.15in}% [[
adamc@230 714 constructive_definite_description
adamc@230 715 : forall (A : Set) (f : A -> nat) (g : nat -> A),
adamc@230 716 (forall x : A, g (f x) = x) ->
adamc@230 717 forall P : A -> Prop,
adamc@230 718 (forall x : A, {P x} + {~ P x}) ->
adamc@230 719 (exists! x : A, P x) -> {x : A | P x}
adam@302 720 ]]
adam@302 721 *)
adamc@230 722
adamc@230 723 Print Assumptions constructive_definite_description.
adam@343 724 (** <<
adamc@230 725 Closed under the global context
adam@343 726 >>
adamc@230 727
adamc@231 728 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 %\textit{%#<i>#unique#</i>#%}% existence [exists!], guarantees termination. The definition of this operator in Coq uses some interesting techniques, as seen in the implementation of the [ConstructiveEpsilon] module.
adamc@230 729
adamc@230 730 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. *)
adamc@230 731
adamc@230 732 Require Import ClassicalUniqueChoice.
adamc@230 733 Check dependent_unique_choice.
adamc@230 734 (** %\vspace{-.15in}% [[
adamc@230 735 dependent_unique_choice
adamc@230 736 : forall (A : Type) (B : A -> Type) (R : forall x : A, B x -> Prop),
adamc@230 737 (forall x : A, exists! y : B x, R x y) ->
adam@343 738 exists f : forall x : A, B x,
adam@343 739 forall x : A, R x (f x)
adamc@230 740 ]]
adamc@230 741
adamc@230 742 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. *)
adamc@230 743
adamc@230 744 Require Import ClassicalChoice.
adamc@230 745 Check choice.
adamc@230 746 (** %\vspace{-.15in}% [[
adamc@230 747 choice
adamc@230 748 : forall (A B : Type) (R : A -> B -> Prop),
adamc@230 749 (forall x : A, exists y : B, R x y) ->
adamc@230 750 exists f : A -> B, forall x : A, R x (f x)
adamc@230 751
adamc@230 752 ]]
adamc@230 753
adamc@230 754 This principle is proved as a theorem, based on the unique choice axiom and an additional axiom of relational choice from the [RelationalChoice] module.
adamc@230 755
adamc@230 756 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. *)
adamc@230 757
adamc@230 758 Definition choice_Set (A B : Type) (R : A -> B -> Prop) (H : forall x : A, {y : B | R x y})
adamc@230 759 : {f : A -> B | forall x : A, R x (f x)} :=
adamc@230 760 exist (fun f => forall x : A, R x (f x))
adamc@230 761 (fun x => proj1_sig (H x)) (fun x => proj2_sig (H x)).
adamc@230 762
adam@287 763 (** 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 subtlely different. In Gallina, the axioms implement a controlled relaxation of the restrictions on information flow from proofs to programs.
adamc@230 764
adam@287 765 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.
adamc@230 766
adamc@230 767 %\bigskip%
adamc@230 768
adam@343 769 The Coq tools support a command-line flag %\index{impredicative Set}\texttt{%#<tt>#-impredicative-set#</tt>#%}%, 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].
adamc@230 770
adamc@230 771 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. *)
adamc@230 772
adamc@230 773 (** ** Axioms and Computation *)
adamc@230 774
adamc@230 775 (** One additional axiom-related wrinkle arises from an aspect of Gallina that is very different from set theory: a notion of %\textit{%#<i>#computational equivalence#</i>#%}% 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. *)
adamc@230 776
adamc@230 777 Definition cast (x y : Set) (pf : x = y) (v : x) : y :=
adamc@230 778 match pf with
adamc@230 779 | refl_equal => v
adamc@230 780 end.
adamc@230 781
adamc@230 782 (** Computation over programs that use [cast] can proceed smoothly. *)
adamc@230 783
adamc@230 784 Eval compute in (cast (refl_equal (nat -> nat)) (fun n => S n)) 12.
adam@343 785 (** %\vspace{-.15in}%[[
adamc@230 786 = 13
adamc@230 787 : nat
adam@302 788 ]]
adam@302 789 *)
adamc@230 790
adamc@230 791 (** Things do not go as smoothly when we use [cast] with proofs that rely on axioms. *)
adamc@230 792
adamc@230 793 Theorem t3 : (forall n : nat, fin (S n)) = (forall n : nat, fin (n + 1)).
adamc@230 794 change ((forall n : nat, (fun n => fin (S n)) n) = (forall n : nat, (fun n => fin (n + 1)) n));
adamc@230 795 rewrite (functional_extensionality (fun n => fin (n + 1)) (fun n => fin (S n))); crush.
adamc@230 796 Qed.
adamc@230 797
adamc@230 798 Eval compute in (cast t3 (fun _ => First)) 12.
adamc@230 799 (** [[
adamc@230 800 = match t3 in (_ = P) return P with
adamc@230 801 | refl_equal => fun n : nat => First
adamc@230 802 end 12
adamc@230 803 : fin (12 + 1)
adamc@230 804 ]]
adamc@230 805
adamc@230 806 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 is easily fixed. *)
adamc@230 807
adamc@230 808 Reset t3.
adamc@230 809
adamc@230 810 Theorem t3 : (forall n : nat, fin (S n)) = (forall n : nat, fin (n + 1)).
adamc@230 811 change ((forall n : nat, (fun n => fin (S n)) n) = (forall n : nat, (fun n => fin (n + 1)) n));
adamc@230 812 rewrite (functional_extensionality (fun n => fin (n + 1)) (fun n => fin (S n))); crush.
adamc@230 813 Defined.
adamc@230 814
adamc@230 815 Eval compute in (cast t3 (fun _ => First)) 12.
adamc@230 816 (** [[
adamc@230 817 = match
adamc@230 818 match
adamc@230 819 match
adamc@230 820 functional_extensionality
adamc@230 821 ....
adamc@230 822 ]]
adamc@230 823
adamc@230 824 We elide most of the details. A very unwieldy tree of nested matches on equality proofs appears. This time evaluation really %\textit{%#<i>#is#</i>#%}% stuck on a use of an axiom.
adamc@230 825
adamc@230 826 If we are careful in using tactics to prove an equality, we can still compute with casts over the proof. *)
adamc@230 827
adamc@230 828 Lemma plus1 : forall n, S n = n + 1.
adamc@230 829 induction n; simpl; intuition.
adamc@230 830 Defined.
adamc@230 831
adamc@230 832 Theorem t4 : forall n, fin (S n) = fin (n + 1).
adamc@230 833 intro; f_equal; apply plus1.
adamc@230 834 Defined.
adamc@230 835
adamc@230 836 Eval compute in cast (t4 13) First.
adamc@230 837 (** %\vspace{-.15in}% [[
adamc@230 838 = First
adamc@230 839 : fin (13 + 1)
adam@302 840 ]]
adam@343 841
adam@343 842 This simple computational reduction hides the use of a recursive function to produce a suitable [refl_equal] proof term. The recursion originates in our use of [induction] in [t4]'s proof. *)
adam@343 843
adam@344 844
adam@344 845 (** ** Methods for Avoiding Axioms *)
adam@344 846
adam@344 847 (** 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}\emph{%#<i>#trusted code base#</i>#%}%. 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.
adam@344 848
adam@344 849 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!)
adam@344 850
adam@344 851 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: *)
adam@344 852
adam@344 853 Theorem fin_cases : forall n (f : fin (S n)), f = First \/ exists f', f = Next f'.
adam@344 854 intros; dep_destruct f; eauto.
adam@344 855 Qed.
adam@344 856
adam@344 857 Print Assumptions fin_cases.
adam@344 858 (** %\vspace{-.15in}%[[
adam@344 859 Axioms:
adam@344 860 JMeq.JMeq_eq : forall (A : Type) (x y : A), JMeq.JMeq x y -> x = y
adam@344 861 ]]
adam@344 862
adam@344 863 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. *)
adam@344 864
adam@344 865 (* begin thide *)
adam@344 866 Lemma fin_cases_again' : forall n (f : fin n),
adam@344 867 match n return fin n -> Prop with
adam@344 868 | O => fun _ => False
adam@344 869 | S n' => fun f => f = First \/ exists f', f = Next f'
adam@344 870 end f.
adam@344 871 destruct f; eauto.
adam@344 872 Qed.
adam@344 873
adam@344 874 (** 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. *)
adam@344 875
adam@344 876 Theorem fin_cases_again : forall n (f : fin (S n)), f = First \/ exists f', f = Next f'.
adam@344 877 intros; exact (fin_cases_again' f).
adam@344 878 Qed.
adam@344 879 (* end thide *)
adam@344 880
adam@344 881 Print Assumptions fin_cases_again.
adam@344 882 (** %\vspace{-.15in}%
adam@344 883 <<
adam@344 884 Closed under the global context
adam@344 885 >>
adam@344 886
adam@345 887 *)
adam@345 888
adam@345 889 (* begin thide *)
adam@345 890 (** 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.
adam@345 891
adam@345 892 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. *)
adam@345 893
adam@345 894 Definition finOut n (f : fin n) : match n return fin n -> Type with
adam@345 895 | O => fun _ => Empty_set
adam@345 896 | _ => fun f => {f' : _ | f = Next f'} + {f = First}
adam@345 897 end f :=
adam@345 898 match f with
adam@345 899 | First _ => inright _ (refl_equal _)
adam@345 900 | Next _ f' => inleft _ (exist _ f' (refl_equal _))
adam@345 901 end.
adam@345 902 (* end thide *)
adam@345 903
adam@345 904 (** 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!) *)
adam@344 905
adam@344 906 Inductive formula : list Type -> Type :=
adam@344 907 | Inject : forall Ts, Prop -> formula Ts
adam@344 908 | VarEq : forall T Ts, T -> formula (T :: Ts)
adam@344 909 | Lift : forall T Ts, formula Ts -> formula (T :: Ts)
adam@344 910 | Forall : forall T Ts, formula (T :: Ts) -> formula Ts
adam@344 911 | And : forall Ts, formula Ts -> formula Ts -> formula Ts.
adam@344 912
adam@344 913 (** 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. *)
adam@344 914
adam@344 915 Inductive proof : formula nil -> Prop :=
adam@344 916 | PInject : forall (P : Prop), P -> proof (Inject nil P)
adam@344 917 | PAnd : forall p q, proof p -> proof q -> proof (And p q).
adam@344 918
adam@344 919 (** Let us prove a lemma showing that a %``%#"#[P /\ Q -> P]#"#%''% rule is derivable within the rules of [proof]. *)
adam@344 920
adam@344 921 Theorem proj1 : forall p q, proof (And p q) -> proof p.
adam@344 922 destruct 1.
adam@344 923 (** %\vspace{-.15in}%[[
adam@344 924 p : formula nil
adam@344 925 q : formula nil
adam@344 926 P : Prop
adam@344 927 H : P
adam@344 928 ============================
adam@344 929 proof p
adam@344 930 ]]
adam@344 931 *)
adam@344 932
adam@344 933 (** 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].)
adam@344 934
adam@344 935 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. *)
adam@344 936
adam@344 937 Restart.
adam@344 938 Require Import Program.
adam@344 939 intros ? ? H; dependent destruction H; auto.
adam@344 940 Qed.
adam@344 941
adam@344 942 Print Assumptions proj1.
adam@344 943 (** %\vspace{-.15in}%[[
adam@344 944 Axioms:
adam@344 945 eq_rect_eq : forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
adam@344 946 x = eq_rect p Q x p h
adam@344 947 ]]
adam@344 948
adam@344 949 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. *)
adam@344 950
adam@344 951 Lemma proj1_again' : forall r, proof r
adam@344 952 -> forall p q, r = And p q -> proof p.
adam@344 953 destruct 1; crush.
adam@344 954 (** %\vspace{-.15in}%[[
adam@344 955 H0 : Inject [] P = And p q
adam@344 956 ============================
adam@344 957 proof p
adam@344 958 ]]
adam@344 959
adam@344 960 The first goal looks reasonable. Hypothesis [H0] is clearly contradictory, as [discriminate] can show. *)
adam@344 961
adam@344 962 discriminate.
adam@344 963 (** %\vspace{-.15in}%[[
adam@344 964 H : proof p
adam@344 965 H1 : And p q = And p0 q0
adam@344 966 ============================
adam@344 967 proof p0
adam@344 968 ]]
adam@344 969
adam@344 970 It looks like we are almost done. Hypothesis [H1] gives [p = p0] by injectivity of constructors, and then [H] finishes the case. *)
adam@344 971
adam@344 972 injection H1; intros.
adam@344 973
adam@344 974 (** Unfortunately, the %``%#"#equality#"#%''% that we expected between [p] and [p0] comes in a strange form:
adam@344 975
adam@344 976 [[
adam@344 977 H3 : existT (fun Ts : list Type => formula Ts) []%list p =
adam@344 978 existT (fun Ts : list Type => formula Ts) []%list p0
adam@344 979 ============================
adam@344 980 proof p0
adam@344 981 ]]
adam@344 982
adam@345 983 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.
adam@344 984
adam@344 985 How exactly does an axiom come into the picture here? Let us ask [crush] to finish the proof. *)
adam@344 986
adam@344 987 crush.
adam@344 988 Qed.
adam@344 989
adam@344 990 Print Assumptions proj1_again'.
adam@344 991 (** %\vspace{-.15in}%[[
adam@344 992 Axioms:
adam@344 993 eq_rect_eq : forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
adam@344 994 x = eq_rect p Q x p h
adam@344 995 ]]
adam@344 996
adam@344 997 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.
adam@344 998
adam@344 999 Hope is not lost, however. We can produce an even stranger looking lemma, which gives us the theorem without axioms. *)
adam@344 1000
adam@344 1001 Lemma proj1_again'' : forall r, proof r
adam@344 1002 -> match r with
adam@344 1003 | And Ps p _ => match Ps return formula Ps -> Prop with
adam@344 1004 | nil => fun p => proof p
adam@344 1005 | _ => fun _ => True
adam@344 1006 end p
adam@344 1007 | _ => True
adam@344 1008 end.
adam@344 1009 destruct 1; auto.
adam@344 1010 Qed.
adam@344 1011
adam@344 1012 Theorem proj1_again : forall p q, proof (And p q) -> proof p.
adam@344 1013 intros ? ? H; exact (proj1_again'' H).
adam@344 1014 Qed.
adam@344 1015
adam@344 1016 Print Assumptions proj1_again.
adam@344 1017 (** <<
adam@344 1018 Closed under the global context
adam@344 1019 >>
adam@344 1020
adam@377 1021 This example illustrates again how some of the same design patterns we learned for dependently typed programming can be used fruitfully in theorem statements.
adam@377 1022
adam@377 1023 %\medskip%
adam@377 1024
adam@377 1025 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 %\emph{%#<i>#compute#</i>#%}%. 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.
adam@377 1026
adam@377 1027 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. *)
adam@377 1028
adam@377 1029 Section withTypes.
adam@377 1030 Variable types : list Set.
adam@377 1031
adam@377 1032 (** To give the semantics of terms, we will need to represent value environments, which assign each variable a term of the proper type. *)
adam@377 1033
adam@377 1034 Variable values : hlist (fun x : Set => x) types.
adam@377 1035
adam@377 1036 (** 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. *)
adam@377 1037
adam@377 1038 Variable natIndex : nat.
adam@377 1039 Variable natIndex_ok : nth_error types natIndex = Some nat.
adam@377 1040
adam@377 1041 (** 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. *)
adam@377 1042
adam@377 1043 Lemma nth_error_nil : forall A n x,
adam@377 1044 nth_error (@nil A) n = Some x
adam@377 1045 -> False.
adam@377 1046 destruct n; simpl; unfold error; congruence.
adam@377 1047 Defined.
adam@377 1048
adam@377 1049 Implicit Arguments nth_error_nil [A n x].
adam@377 1050
adam@377 1051 Lemma Some_inj : forall A (x y : A),
adam@377 1052 Some x = Some y
adam@377 1053 -> x = y.
adam@377 1054 congruence.
adam@377 1055 Defined.
adam@377 1056
adam@377 1057 Fixpoint getNat (types' : list Set) (values' : hlist (fun x : Set => x) types')
adam@377 1058 (natIndex : nat) : (nth_error types' natIndex = Some nat) -> nat :=
adam@377 1059 match values' with
adam@377 1060 | HNil => fun pf => match nth_error_nil pf with end
adam@377 1061 | HCons t ts x values'' =>
adam@377 1062 match natIndex return nth_error (t :: ts) natIndex = Some nat -> nat with
adam@377 1063 | O => fun pf =>
adam@377 1064 match Some_inj pf in _ = T return T with
adam@377 1065 | refl_equal => x
adam@377 1066 end
adam@377 1067 | S natIndex' => getNat values'' natIndex'
adam@377 1068 end
adam@377 1069 end.
adam@377 1070 End withTypes.
adam@377 1071
adam@377 1072 (** The problem becomes apparent when we experiment with running [getNat] on a concrete [types] list. *)
adam@377 1073
adam@377 1074 Definition myTypes := unit :: nat :: bool :: nil.
adam@377 1075 Definition myValues : hlist (fun x : Set => x) myTypes :=
adam@377 1076 tt ::: 3 ::: false ::: HNil.
adam@377 1077
adam@377 1078 Definition myNatIndex := 1.
adam@377 1079
adam@377 1080 Theorem myNatIndex_ok : nth_error myTypes myNatIndex = Some nat.
adam@377 1081 reflexivity.
adam@377 1082 Defined.
adam@377 1083
adam@377 1084 Eval compute in getNat myValues myNatIndex myNatIndex_ok.
adam@377 1085 (** %\vspace{-.15in}%[[
adam@377 1086 = 3
adam@377 1087 ]]
adam@377 1088
adam@377 1089 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] %\emph{%#<i>#independently#</i>#%}% of a specific proof. *)
adam@377 1090
adam@377 1091 Theorem getNat_is_reasonable : forall pf, getNat myValues myNatIndex pf = 3.
adam@377 1092 intro; compute.
adam@377 1093 (**
adam@377 1094 <<
adam@377 1095 1 subgoal
adam@377 1096 >>
adam@377 1097 %\vspace{-.3in}%[[
adam@377 1098 pf : nth_error myTypes myNatIndex = Some nat
adam@377 1099 ============================
adam@377 1100 match
adam@377 1101 match
adam@377 1102 pf in (_ = y)
adam@377 1103 return (nat = match y with
adam@377 1104 | Some H => H
adam@377 1105 | None => nat
adam@377 1106 end)
adam@377 1107 with
adam@377 1108 | eq_refl => eq_refl
adam@377 1109 end in (_ = T) return T
adam@377 1110 with
adam@377 1111 | eq_refl => 3
adam@377 1112 end = 3
adam@377 1113 ]]
adam@377 1114
adam@377 1115 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. *)
adam@377 1116
adam@377 1117 Abort.
adam@377 1118
adam@377 1119 (** 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. *)
adam@377 1120
adam@377 1121 Fixpoint copies A (x : A) (n : nat) : list A :=
adam@377 1122 match n with
adam@377 1123 | O => nil
adam@377 1124 | S n' => x :: copies x n'
adam@377 1125 end.
adam@377 1126
adam@377 1127 Fixpoint update A (ls : list A) (n : nat) (x : A) : list A :=
adam@377 1128 match ls with
adam@377 1129 | nil => copies x n ++ x :: nil
adam@377 1130 | y :: ls' => match n with
adam@377 1131 | O => x :: ls'
adam@377 1132 | S n' => y :: update ls' n' x
adam@377 1133 end
adam@377 1134 end.
adam@377 1135
adam@377 1136 (** Now let us revisit the definition of [getNat]. *)
adam@377 1137
adam@377 1138 Section withTypes'.
adam@377 1139 Variable types : list Set.
adam@377 1140 Variable natIndex : nat.
adam@377 1141
adam@377 1142 (** Here is the trick: instead of asserting properties about the list [types], we build a %``%#"#new#"#%''% list that is %\emph{%#<i>#guaranteed by construction#</i>#%}% to have those properties. *)
adam@377 1143
adam@377 1144 Definition types' := update types natIndex nat.
adam@377 1145
adam@377 1146 Variable values : hlist (fun x : Set => x) types'.
adam@377 1147
adam@377 1148 (** Now a bit of dependent pattern matching helps us rewrite [getNat] in a way that avoids any use of equality proofs. *)
adam@377 1149
adam@378 1150 Fixpoint skipCopies (n : nat)
adam@378 1151 : hlist (fun x : Set => x) (copies nat n ++ nat :: nil) -> nat :=
adam@378 1152 match n with
adam@378 1153 | O => fun vs => hhd vs
adam@378 1154 | S n' => fun vs => skipCopies n' (htl vs)
adam@378 1155 end.
adam@378 1156
adam@377 1157 Fixpoint getNat' (types'' : list Set) (natIndex : nat)
adam@377 1158 : hlist (fun x : Set => x) (update types'' natIndex nat) -> nat :=
adam@377 1159 match types'' with
adam@378 1160 | nil => skipCopies natIndex
adam@377 1161 | t :: types0 =>
adam@377 1162 match natIndex return hlist (fun x : Set => x)
adam@377 1163 (update (t :: types0) natIndex nat) -> nat with
adam@377 1164 | O => fun vs => hhd vs
adam@377 1165 | S natIndex' => fun vs => getNat' types0 natIndex' (htl vs)
adam@377 1166 end
adam@377 1167 end.
adam@377 1168 End withTypes'.
adam@377 1169
adam@377 1170 (** Now the surprise comes in how easy it is to %\emph{%#<i>#use#</i>#%}% [getNat']. While typing works by modification of a types list, we can choose parameters so that the modification has no effect. *)
adam@377 1171
adam@377 1172 Theorem getNat_is_reasonable : getNat' myTypes myNatIndex myValues = 3.
adam@377 1173 reflexivity.
adam@377 1174 Qed.
adam@377 1175
adam@377 1176 (** The same parameters as before work without alteration, and we avoid use of axioms. *)