annotate src/Equality.v @ 314:d5787b70cf48

Rename Tactics; change 'principal typing' to 'principal types'
author Adam Chlipala <adam@chlipala.net>
date Wed, 07 Sep 2011 13:47:24 -0400
parents 7b38729be069
children 2fbb47fb02bd
rev   line source
adam@297 1 (* Copyright (c) 2008-2011, Adam Chlipala
adamc@118 2 *
adamc@118 3 * This work is licensed under a
adamc@118 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@118 5 * Unported License.
adamc@118 6 * The license text is available at:
adamc@118 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@118 8 *)
adamc@118 9
adamc@118 10 (* begin hide *)
adamc@120 11 Require Import Eqdep JMeq List.
adamc@118 12
adam@314 13 Require Import CpdtTactics.
adamc@118 14
adamc@118 15 Set Implicit Arguments.
adamc@118 16 (* end hide *)
adamc@118 17
adamc@118 18
adamc@118 19 (** %\chapter{Reasoning About Equality Proofs}% *)
adamc@118 20
adam@294 21 (** In traditional mathematics, the concept of equality is usually taken as a given. On the other hand, in type theory, equality is a very contentious subject. There are at least three different notions of equality that are important, and researchers are actively investigating new definitions of what it means for two terms to be equal. Even once we fix a notion of equality, there are inevitably tricky issues that arise in proving properties of programs that manipulate equality proofs explicitly. In this chapter, I will focus on design patterns for circumventing these tricky issues, and I will introduce the different notions of equality as they are germane. *)
adamc@118 22
adamc@118 23
adamc@122 24 (** * The Definitional Equality *)
adamc@122 25
adam@292 26 (** We have seen many examples so far where proof goals follow %``%#"#by computation.#"#%''% That is, we apply computational reduction rules to reduce the goal to a normal form, at which point it follows trivially. Exactly when this works and when it does not depends on the details of Coq's %\textit{%#<i>#definitional equality#</i>#%}%. This is an untyped binary relation appearing in the formal metatheory of CIC. CIC contains a typing rule allowing the conclusion $E : T$ from the premise $E : T'$ and a proof that $T$ and $T'$ are definitionally equal.
adamc@122 27
adamc@199 28 The [cbv] tactic will help us illustrate the rules of Coq's definitional equality. We redefine the natural number predecessor function in a somewhat convoluted way and construct a manual proof that it returns [0] when applied to [1]. *)
adamc@122 29
adamc@122 30 Definition pred' (x : nat) :=
adamc@122 31 match x with
adamc@122 32 | O => O
adamc@122 33 | S n' => let y := n' in y
adamc@122 34 end.
adamc@122 35
adamc@122 36 Theorem reduce_me : pred' 1 = 0.
adamc@218 37
adamc@124 38 (* begin thide *)
adamc@122 39 (** CIC follows the traditions of lambda calculus in associating reduction rules with Greek letters. Coq can certainly be said to support the familiar alpha reduction rule, which allows capture-avoiding renaming of bound variables, but we never need to apply alpha explicitly, since Coq uses a de Bruijn representation that encodes terms canonically.
adamc@122 40
adamc@131 41 The delta rule is for unfolding global definitions. We can use it here to unfold the definition of [pred']. We do this with the [cbv] tactic, which takes a list of reduction rules and makes as many call-by-value reduction steps as possible, using only those rules. There is an analogous tactic [lazy] for call-by-need reduction. *)
adamc@122 42
adamc@122 43 cbv delta.
adamc@122 44 (** [[
adamc@122 45 ============================
adamc@122 46 (fun x : nat => match x with
adamc@122 47 | 0 => 0
adamc@122 48 | S n' => let y := n' in y
adamc@122 49 end) 1 = 0
adamc@218 50
adamc@122 51 ]]
adamc@122 52
adamc@122 53 At this point, we want to apply the famous beta reduction of lambda calculus, to simplify the application of a known function abstraction. *)
adamc@122 54
adamc@122 55 cbv beta.
adamc@122 56 (** [[
adamc@122 57 ============================
adamc@122 58 match 1 with
adamc@122 59 | 0 => 0
adamc@122 60 | S n' => let y := n' in y
adamc@122 61 end = 0
adamc@218 62
adamc@122 63 ]]
adamc@122 64
adamc@122 65 Next on the list is the iota reduction, which simplifies a single [match] term by determining which pattern matches. *)
adamc@122 66
adamc@122 67 cbv iota.
adamc@122 68 (** [[
adamc@122 69 ============================
adamc@122 70 (fun n' : nat => let y := n' in y) 0 = 0
adamc@218 71
adamc@122 72 ]]
adamc@122 73
adamc@122 74 Now we need another beta reduction. *)
adamc@122 75
adamc@122 76 cbv beta.
adamc@122 77 (** [[
adamc@122 78 ============================
adamc@122 79 (let y := 0 in y) = 0
adamc@218 80
adamc@122 81 ]]
adamc@122 82
adam@296 83 The final reduction rule is zeta, which replaces a [let] expression by its body with the appropriate term substituted. *)
adamc@122 84
adamc@122 85 cbv zeta.
adamc@122 86 (** [[
adamc@122 87 ============================
adamc@122 88 0 = 0
adamc@218 89
adam@302 90 ]]
adam@302 91 *)
adamc@122 92
adamc@122 93 reflexivity.
adamc@122 94 Qed.
adamc@124 95 (* end thide *)
adamc@122 96
adamc@122 97 (** The standard [eq] relation is critically dependent on the definitional equality. [eq] is often called a %\textit{%#<i>#propositional equality#</i>#%}%, because it reifies definitional equality as a proposition that may or may not hold. Standard axiomatizations of an equality predicate in first-order logic define equality in terms of properties it has, like reflexivity, symmetry, and transitivity. In contrast, for [eq] in Coq, those properties are implicit in the properties of the definitional equality, which are built into CIC's metatheory and the implementation of Gallina. We could add new rules to the definitional equality, and [eq] would keep its definition and methods of use.
adamc@122 98
adam@294 99 This all may make it sound like the choice of [eq]'s definition is unimportant. To the contrary, in this chapter, we will see examples where alternate definitions may simplify proofs. Before that point, I will introduce proof methods for goals that use proofs of the standard propositional equality %``%#"#as data.#"#%''% *)
adamc@122 100
adamc@122 101
adamc@118 102 (** * Heterogeneous Lists Revisited *)
adamc@118 103
adam@292 104 (** One of our example dependent data structures from the last chapter was heterogeneous lists and their associated %``%#"#cursor#"#%''% type. The recursive version poses some special challenges related to equality proofs, since it uses such proofs in its definition of [member] types. *)
adamc@118 105
adamc@118 106 Section fhlist.
adamc@118 107 Variable A : Type.
adamc@118 108 Variable B : A -> Type.
adamc@118 109
adamc@118 110 Fixpoint fhlist (ls : list A) : Type :=
adamc@118 111 match ls with
adamc@118 112 | nil => unit
adamc@118 113 | x :: ls' => B x * fhlist ls'
adamc@118 114 end%type.
adamc@118 115
adamc@118 116 Variable elm : A.
adamc@118 117
adamc@118 118 Fixpoint fmember (ls : list A) : Type :=
adamc@118 119 match ls with
adamc@118 120 | nil => Empty_set
adamc@118 121 | x :: ls' => (x = elm) + fmember ls'
adamc@118 122 end%type.
adamc@118 123
adamc@118 124 Fixpoint fhget (ls : list A) : fhlist ls -> fmember ls -> B elm :=
adamc@118 125 match ls return fhlist ls -> fmember ls -> B elm with
adamc@118 126 | nil => fun _ idx => match idx with end
adamc@118 127 | _ :: ls' => fun mls idx =>
adamc@118 128 match idx with
adamc@118 129 | inl pf => match pf with
adamc@118 130 | refl_equal => fst mls
adamc@118 131 end
adamc@118 132 | inr idx' => fhget ls' (snd mls) idx'
adamc@118 133 end
adamc@118 134 end.
adamc@118 135 End fhlist.
adamc@118 136
adamc@118 137 Implicit Arguments fhget [A B elm ls].
adamc@118 138
adamc@118 139 (** We can define a [map]-like function for [fhlist]s. *)
adamc@118 140
adamc@118 141 Section fhlist_map.
adamc@118 142 Variables A : Type.
adamc@118 143 Variables B C : A -> Type.
adamc@118 144 Variable f : forall x, B x -> C x.
adamc@118 145
adamc@118 146 Fixpoint fhmap (ls : list A) : fhlist B ls -> fhlist C ls :=
adamc@118 147 match ls return fhlist B ls -> fhlist C ls with
adamc@118 148 | nil => fun _ => tt
adamc@118 149 | _ :: _ => fun hls => (f (fst hls), fhmap _ (snd hls))
adamc@118 150 end.
adamc@118 151
adamc@118 152 Implicit Arguments fhmap [ls].
adamc@118 153
adamc@118 154 (** For the inductive versions of the [ilist] definitions, we proved a lemma about the interaction of [get] and [imap]. It was a strategic choice not to attempt such a proof for the definitions that we just gave, because that sets us on a collision course with the problems that are the subject of this chapter. *)
adamc@118 155
adamc@118 156 Variable elm : A.
adamc@118 157
adamc@118 158 Theorem get_imap : forall ls (mem : fmember elm ls) (hls : fhlist B ls),
adamc@118 159 fhget (fhmap hls) mem = f (fhget hls mem).
adam@298 160 (* begin hide *)
adam@298 161 induction ls; crush; case a0; reflexivity.
adam@298 162 (* end hide *)
adam@298 163 (** [[
adamc@118 164 induction ls; crush.
adam@298 165
adam@298 166 ]]
adamc@118 167
adam@298 168 In Coq 8.2, one subgoal remains at this point. Coq 8.3 has added some tactic improvements that enable [crush] to complete all of both inductive cases. To introduce the basics of reasoning about equality, it will be useful to review what was necessary in Coq 8.2.
adam@297 169
adam@297 170 Part of our single remaining subgoal is:
adamc@118 171
adamc@118 172 [[
adamc@118 173 a0 : a = elm
adamc@118 174 ============================
adamc@118 175 match a0 in (_ = a2) return (C a2) with
adamc@118 176 | refl_equal => f a1
adamc@118 177 end = f match a0 in (_ = a2) return (B a2) with
adamc@118 178 | refl_equal => a1
adamc@118 179 end
adamc@218 180
adamc@118 181 ]]
adamc@118 182
adamc@118 183 This seems like a trivial enough obligation. The equality proof [a0] must be [refl_equal], since that is the only constructor of [eq]. Therefore, both the [match]es reduce to the point where the conclusion follows by reflexivity.
adamc@118 184
adamc@118 185 [[
adamc@118 186 destruct a0.
adamc@118 187
adamc@118 188 User error: Cannot solve a second-order unification problem
adamc@218 189
adamc@118 190 ]]
adamc@118 191
adamc@118 192 This is one of Coq's standard error messages for informing us that its heuristics for attempting an instance of an undecidable problem about dependent typing have failed. We might try to nudge things in the right direction by stating the lemma that we believe makes the conclusion trivial.
adamc@118 193
adamc@118 194 [[
adamc@118 195 assert (a0 = refl_equal _).
adamc@118 196
adamc@118 197 The term "refl_equal ?98" has type "?98 = ?98"
adamc@118 198 while it is expected to have type "a = elm"
adamc@218 199
adamc@118 200 ]]
adamc@118 201
adamc@118 202 In retrospect, the problem is not so hard to see. Reflexivity proofs only show [x = x] for particular values of [x], whereas here we are thinking in terms of a proof of [a = elm], where the two sides of the equality are not equal syntactically. Thus, the essential lemma we need does not even type-check!
adamc@118 203
adam@292 204 Is it time to throw in the towel? Luckily, the answer is %``%#"#no.#"#%''% In this chapter, we will see several useful patterns for proving obligations like this.
adamc@118 205
adam@297 206 For this particular example, the solution is surprisingly straightforward. [destruct] has a simpler sibling [case] which should behave identically for any inductive type with one constructor of no arguments.
adamc@118 207
adam@297 208 [[
adamc@118 209 case a0.
adam@297 210
adamc@118 211 ============================
adamc@118 212 f a1 = f a1
adamc@218 213
adamc@118 214 ]]
adamc@118 215
adam@297 216 It seems that [destruct] was trying to be too smart for its own good.
adamc@118 217
adam@297 218 [[
adamc@118 219 reflexivity.
adam@297 220
adam@302 221 ]]
adam@302 222 *)
adam@298 223
adamc@118 224 Qed.
adamc@124 225 (* end thide *)
adamc@118 226
adamc@118 227 (** It will be helpful to examine the proof terms generated by this sort of strategy. A simpler example illustrates what is going on. *)
adamc@118 228
adamc@118 229 Lemma lemma1 : forall x (pf : x = elm), O = match pf with refl_equal => O end.
adamc@124 230 (* begin thide *)
adamc@118 231 simple destruct pf; reflexivity.
adamc@118 232 Qed.
adamc@124 233 (* end thide *)
adamc@118 234
adamc@118 235 (** [simple destruct pf] is a convenient form for applying [case]. It runs [intro] to bring into scope all quantified variables up to its argument. *)
adamc@118 236
adamc@118 237 Print lemma1.
adamc@218 238 (** %\vspace{-.15in}% [[
adamc@118 239 lemma1 =
adamc@118 240 fun (x : A) (pf : x = elm) =>
adamc@118 241 match pf as e in (_ = y) return (0 = match e with
adamc@118 242 | refl_equal => 0
adamc@118 243 end) with
adamc@118 244 | refl_equal => refl_equal 0
adamc@118 245 end
adamc@118 246 : forall (x : A) (pf : x = elm), 0 = match pf with
adamc@118 247 | refl_equal => 0
adamc@118 248 end
adamc@218 249
adamc@118 250 ]]
adamc@118 251
adamc@118 252 Using what we know about shorthands for [match] annotations, we can write this proof in shorter form manually. *)
adamc@118 253
adamc@124 254 (* begin thide *)
adamc@118 255 Definition lemma1' :=
adamc@118 256 fun (x : A) (pf : x = elm) =>
adamc@118 257 match pf return (0 = match pf with
adamc@118 258 | refl_equal => 0
adamc@118 259 end) with
adamc@118 260 | refl_equal => refl_equal 0
adamc@118 261 end.
adamc@124 262 (* end thide *)
adamc@118 263
adamc@118 264 (** Surprisingly, what seems at first like a %\textit{%#<i>#simpler#</i>#%}% lemma is harder to prove. *)
adamc@118 265
adamc@118 266 Lemma lemma2 : forall (x : A) (pf : x = x), O = match pf with refl_equal => O end.
adamc@124 267 (* begin thide *)
adamc@118 268 (** [[
adamc@118 269 simple destruct pf.
adamc@205 270
adamc@118 271 User error: Cannot solve a second-order unification problem
adamc@218 272
adam@302 273 ]]
adam@302 274 *)
adamc@118 275 Abort.
adamc@118 276
adamc@118 277 (** Nonetheless, we can adapt the last manual proof to handle this theorem. *)
adamc@118 278
adamc@124 279 (* begin thide *)
adamc@124 280 Definition lemma2 :=
adamc@118 281 fun (x : A) (pf : x = x) =>
adamc@118 282 match pf return (0 = match pf with
adamc@118 283 | refl_equal => 0
adamc@118 284 end) with
adamc@118 285 | refl_equal => refl_equal 0
adamc@118 286 end.
adamc@124 287 (* end thide *)
adamc@118 288
adamc@118 289 (** We can try to prove a lemma that would simplify proofs of many facts like [lemma2]: *)
adamc@118 290
adamc@118 291 Lemma lemma3 : forall (x : A) (pf : x = x), pf = refl_equal x.
adamc@124 292 (* begin thide *)
adamc@118 293 (** [[
adamc@118 294 simple destruct pf.
adamc@205 295
adamc@118 296 User error: Cannot solve a second-order unification problem
adam@302 297 ]]
adam@302 298 *)
adamc@218 299
adamc@118 300 Abort.
adamc@118 301
adamc@118 302 (** This time, even our manual attempt fails.
adamc@118 303
adamc@118 304 [[
adamc@118 305 Definition lemma3' :=
adamc@118 306 fun (x : A) (pf : x = x) =>
adamc@118 307 match pf as pf' in (_ = x') return (pf' = refl_equal x') with
adamc@118 308 | refl_equal => refl_equal _
adamc@118 309 end.
adamc@118 310
adamc@118 311 The term "refl_equal x'" has type "x' = x'" while it is expected to have type
adamc@118 312 "x = x'"
adamc@218 313
adamc@118 314 ]]
adamc@118 315
adam@296 316 The type error comes from our [return] annotation. In that annotation, the [as]-bound variable [pf'] has type [x = x'], referring to the [in]-bound variable [x']. To do a dependent [match], we %\textit{%#<i>#must#</i>#%}% choose a fresh name for the second argument of [eq]. We are just as constrained to use the %``%#"#real#"#%''% value [x] for the first argument. Thus, within the [return] clause, the proof we are matching on %\textit{%#<i>#must#</i>#%}% equate two non-matching terms, which makes it impossible to equate that proof with reflexivity.
adamc@118 317
adamc@118 318 Nonetheless, it turns out that, with one catch, we %\textit{%#<i>#can#</i>#%}% prove this lemma. *)
adamc@118 319
adamc@118 320 Lemma lemma3 : forall (x : A) (pf : x = x), pf = refl_equal x.
adamc@118 321 intros; apply UIP_refl.
adamc@118 322 Qed.
adamc@118 323
adamc@118 324 Check UIP_refl.
adamc@218 325 (** %\vspace{-.15in}% [[
adamc@118 326 UIP_refl
adamc@118 327 : forall (U : Type) (x : U) (p : x = x), p = refl_equal x
adamc@218 328
adamc@118 329 ]]
adamc@118 330
adamc@118 331 [UIP_refl] comes from the [Eqdep] module of the standard library. Do the Coq authors know of some clever trick for building such proofs that we have not seen yet? If they do, they did not use it for this proof. Rather, the proof is based on an %\textit{%#<i>#axiom#</i>#%}%. *)
adamc@118 332
adamc@118 333 Print eq_rect_eq.
adamc@218 334 (** %\vspace{-.15in}% [[
adamc@118 335 eq_rect_eq =
adamc@118 336 fun U : Type => Eq_rect_eq.eq_rect_eq U
adamc@118 337 : forall (U : Type) (p : U) (Q : U -> Type) (x : Q p) (h : p = p),
adamc@118 338 x = eq_rect p Q x p h
adamc@218 339
adamc@118 340 ]]
adamc@118 341
adam@292 342 [eq_rect_eq] states a %``%#"#fact#"#%''% that seems like common sense, once the notation is deciphered. [eq_rect] is the automatically-generated recursion principle for [eq]. Calling [eq_rect] is another way of [match]ing on an equality proof. The proof we match on is the argument [h], and [x] is the body of the [match]. [eq_rect_eq] just says that [match]es on proofs of [p = p], for any [p], are superfluous and may be removed.
adamc@118 343
adam@292 344 Perhaps surprisingly, we cannot prove [eq_rect_eq] from within Coq. This proposition is introduced as an axiom; that is, a proposition asserted as true without proof. We cannot assert just any statement without proof. Adding [False] as an axiom would allow us to prove any proposition, for instance, defeating the point of using a proof assistant. In general, we need to be sure that we never assert %\textit{%#<i>#inconsistent#</i>#%}% sets of axioms. A set of axioms is inconsistent if its conjunction implies [False]. For the case of [eq_rect_eq], consistency has been verified outside of Coq via %``%#"#informal#"#%''% metatheory.
adamc@118 345
adamc@118 346 This axiom is equivalent to another that is more commonly known and mentioned in type theory circles. *)
adamc@118 347
adamc@118 348 Print Streicher_K.
adamc@124 349 (* end thide *)
adamc@218 350 (** %\vspace{-.15in}% [[
adamc@118 351 Streicher_K =
adamc@118 352 fun U : Type => UIP_refl__Streicher_K U (UIP_refl U)
adamc@118 353 : forall (U : Type) (x : U) (P : x = x -> Prop),
adamc@118 354 P (refl_equal x) -> forall p : x = x, P p
adamc@218 355
adamc@118 356 ]]
adamc@118 357
adam@292 358 This is the unfortunately-named %``%#"#Streicher's axiom K,#"#%''% which says that a predicate on properly-typed equality proofs holds of all such proofs if it holds of reflexivity. *)
adamc@118 359
adamc@118 360 End fhlist_map.
adamc@118 361
adamc@119 362
adamc@119 363 (** * Type-Casts in Theorem Statements *)
adamc@119 364
adamc@119 365 (** Sometimes we need to use tricks with equality just to state the theorems that we care about. To illustrate, we start by defining a concatenation function for [fhlist]s. *)
adamc@119 366
adamc@119 367 Section fhapp.
adamc@119 368 Variable A : Type.
adamc@119 369 Variable B : A -> Type.
adamc@119 370
adamc@218 371 Fixpoint fhapp (ls1 ls2 : list A)
adamc@119 372 : fhlist B ls1 -> fhlist B ls2 -> fhlist B (ls1 ++ ls2) :=
adamc@218 373 match ls1 with
adamc@119 374 | nil => fun _ hls2 => hls2
adamc@119 375 | _ :: _ => fun hls1 hls2 => (fst hls1, fhapp _ _ (snd hls1) hls2)
adamc@119 376 end.
adamc@119 377
adamc@119 378 Implicit Arguments fhapp [ls1 ls2].
adamc@119 379
adamc@124 380 (* EX: Prove that fhapp is associative. *)
adamc@124 381 (* begin thide *)
adamc@124 382
adamc@119 383 (** We might like to prove that [fhapp] is associative.
adamc@119 384
adamc@119 385 [[
adamc@119 386 Theorem fhapp_ass : forall ls1 ls2 ls3
adamc@119 387 (hls1 : fhlist B ls1) (hls2 : fhlist B ls2) (hls3 : fhlist B ls3),
adamc@119 388 fhapp hls1 (fhapp hls2 hls3) = fhapp (fhapp hls1 hls2) hls3.
adamc@119 389
adamc@119 390 The term
adamc@119 391 "fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3) (fhapp (ls1:=ls1) (ls2:=ls2) hls1 hls2)
adamc@119 392 hls3" has type "fhlist B ((ls1 ++ ls2) ++ ls3)"
adamc@119 393 while it is expected to have type "fhlist B (ls1 ++ ls2 ++ ls3)"
adamc@218 394
adamc@119 395 ]]
adamc@119 396
adamc@119 397 This first cut at the theorem statement does not even type-check. We know that the two [fhlist] types appearing in the error message are always equal, by associativity of normal list append, but this fact is not apparent to the type checker. This stems from the fact that Coq's equality is %\textit{%#<i>#intensional#</i>#%}%, in the sense that type equality theorems can never be applied after the fact to get a term to type-check. Instead, we need to make use of equality explicitly in the theorem statement. *)
adamc@119 398
adamc@119 399 Theorem fhapp_ass : forall ls1 ls2 ls3
adamc@119 400 (pf : (ls1 ++ ls2) ++ ls3 = ls1 ++ (ls2 ++ ls3))
adamc@119 401 (hls1 : fhlist B ls1) (hls2 : fhlist B ls2) (hls3 : fhlist B ls3),
adamc@119 402 fhapp hls1 (fhapp hls2 hls3)
adamc@119 403 = match pf in (_ = ls) return fhlist _ ls with
adamc@119 404 | refl_equal => fhapp (fhapp hls1 hls2) hls3
adamc@119 405 end.
adamc@119 406 induction ls1; crush.
adamc@119 407
adamc@119 408 (** The first remaining subgoal looks trivial enough:
adamc@119 409
adamc@119 410 [[
adamc@119 411 ============================
adamc@119 412 fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3 =
adamc@119 413 match pf in (_ = ls) return (fhlist B ls) with
adamc@119 414 | refl_equal => fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3
adamc@119 415 end
adamc@218 416
adamc@119 417 ]]
adamc@119 418
adamc@119 419 We can try what worked in previous examples.
adamc@119 420
adamc@119 421 [[
adamc@119 422 case pf.
adamc@119 423
adamc@119 424 User error: Cannot solve a second-order unification problem
adamc@218 425
adamc@119 426 ]]
adamc@119 427
adamc@119 428 It seems we have reached another case where it is unclear how to use a dependent [match] to implement case analysis on our proof. The [UIP_refl] theorem can come to our rescue again. *)
adamc@119 429
adamc@119 430 rewrite (UIP_refl _ _ pf).
adamc@119 431 (** [[
adamc@119 432 ============================
adamc@119 433 fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3 =
adamc@119 434 fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3
adamc@218 435
adam@302 436 ]]
adam@302 437 *)
adamc@119 438
adamc@119 439 reflexivity.
adamc@119 440
adamc@119 441 (** Our second subgoal is trickier.
adamc@119 442
adamc@119 443 [[
adamc@119 444 pf : a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3
adamc@119 445 ============================
adamc@119 446 (a0,
adamc@119 447 fhapp (ls1:=ls1) (ls2:=ls2 ++ ls3) b
adamc@119 448 (fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3)) =
adamc@119 449 match pf in (_ = ls) return (fhlist B ls) with
adamc@119 450 | refl_equal =>
adamc@119 451 (a0,
adamc@119 452 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3)
adamc@119 453 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3)
adamc@119 454 end
adamc@119 455
adamc@119 456 rewrite (UIP_refl _ _ pf).
adamc@119 457
adamc@119 458 The term "pf" has type "a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3"
adamc@119 459 while it is expected to have type "?556 = ?556"
adamc@218 460
adamc@119 461 ]]
adamc@119 462
adamc@119 463 We can only apply [UIP_refl] on proofs of equality with syntactically equal operands, which is not the case of [pf] here. We will need to manipulate the form of this subgoal to get us to a point where we may use [UIP_refl]. A first step is obtaining a proof suitable to use in applying the induction hypothesis. Inversion on the structure of [pf] is sufficient for that. *)
adamc@119 464
adamc@119 465 injection pf; intro pf'.
adamc@119 466 (** [[
adamc@119 467 pf : a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3
adamc@119 468 pf' : (ls1 ++ ls2) ++ ls3 = ls1 ++ ls2 ++ ls3
adamc@119 469 ============================
adamc@119 470 (a0,
adamc@119 471 fhapp (ls1:=ls1) (ls2:=ls2 ++ ls3) b
adamc@119 472 (fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3)) =
adamc@119 473 match pf in (_ = ls) return (fhlist B ls) with
adamc@119 474 | refl_equal =>
adamc@119 475 (a0,
adamc@119 476 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3)
adamc@119 477 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3)
adamc@119 478 end
adamc@218 479
adamc@119 480 ]]
adamc@119 481
adamc@119 482 Now we can rewrite using the inductive hypothesis. *)
adamc@119 483
adamc@119 484 rewrite (IHls1 _ _ pf').
adamc@119 485 (** [[
adamc@119 486 ============================
adamc@119 487 (a0,
adamc@119 488 match pf' in (_ = ls) return (fhlist B ls) with
adamc@119 489 | refl_equal =>
adamc@119 490 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3)
adamc@119 491 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3
adamc@119 492 end) =
adamc@119 493 match pf in (_ = ls) return (fhlist B ls) with
adamc@119 494 | refl_equal =>
adamc@119 495 (a0,
adamc@119 496 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3)
adamc@119 497 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3)
adamc@119 498 end
adamc@218 499
adamc@119 500 ]]
adamc@119 501
adam@294 502 We have made an important bit of progress, as now only a single call to [fhapp] appears in the conclusion, repeated twice. Trying case analysis on our proofs still will not work, but there is a move we can make to enable it. Not only does just one call to [fhapp] matter to us now, but it also %\textit{%#<i>#does not matter what the result of the call is#</i>#%}%. In other words, the subgoal should remain true if we replace this [fhapp] call with a fresh variable. The [generalize] tactic helps us do exactly that. *)
adamc@119 503
adamc@119 504 generalize (fhapp (fhapp b hls2) hls3).
adamc@119 505 (** [[
adamc@119 506 forall f : fhlist B ((ls1 ++ ls2) ++ ls3),
adamc@119 507 (a0,
adamc@119 508 match pf' in (_ = ls) return (fhlist B ls) with
adamc@119 509 | refl_equal => f
adamc@119 510 end) =
adamc@119 511 match pf in (_ = ls) return (fhlist B ls) with
adamc@119 512 | refl_equal => (a0, f)
adamc@119 513 end
adamc@218 514
adamc@119 515 ]]
adamc@119 516
adamc@119 517 The conclusion has gotten markedly simpler. It seems counterintuitive that we can have an easier time of proving a more general theorem, but that is exactly the case here and for many other proofs that use dependent types heavily. Speaking informally, the reason why this kind of activity helps is that [match] annotations only support variables in certain positions. By reducing more elements of a goal to variables, built-in tactics can have more success building [match] terms under the hood.
adamc@119 518
adamc@119 519 In this case, it is helpful to generalize over our two proofs as well. *)
adamc@119 520
adamc@119 521 generalize pf pf'.
adamc@119 522 (** [[
adamc@119 523 forall (pf0 : a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3)
adamc@119 524 (pf'0 : (ls1 ++ ls2) ++ ls3 = ls1 ++ ls2 ++ ls3)
adamc@119 525 (f : fhlist B ((ls1 ++ ls2) ++ ls3)),
adamc@119 526 (a0,
adamc@119 527 match pf'0 in (_ = ls) return (fhlist B ls) with
adamc@119 528 | refl_equal => f
adamc@119 529 end) =
adamc@119 530 match pf0 in (_ = ls) return (fhlist B ls) with
adamc@119 531 | refl_equal => (a0, f)
adamc@119 532 end
adamc@218 533
adamc@119 534 ]]
adamc@119 535
adamc@119 536 To an experienced dependent types hacker, the appearance of this goal term calls for a celebration. The formula has a critical property that indicates that our problems are over. To get our proofs into the right form to apply [UIP_refl], we need to use associativity of list append to rewrite their types. We could not do that before because other parts of the goal require the proofs to retain their original types. In particular, the call to [fhapp] that we generalized must have type [(ls1 ++ ls2) ++ ls3], for some values of the list variables. If we rewrite the type of the proof used to type-cast this value to something like [ls1 ++ ls2 ++ ls3 = ls1 ++ ls2 ++ ls3], then the lefthand side of the equality would no longer match the type of the term we are trying to cast.
adamc@119 537
adamc@119 538 However, now that we have generalized over the [fhapp] call, the type of the term being type-cast appears explicitly in the goal and %\textit{%#<i>#may be rewritten as well#</i>#%}%. In particular, the final masterstroke is rewriting everywhere in our goal using associativity of list append. *)
adamc@119 539
adamc@119 540 rewrite app_ass.
adamc@119 541 (** [[
adamc@119 542 ============================
adamc@119 543 forall (pf0 : a :: ls1 ++ ls2 ++ ls3 = a :: ls1 ++ ls2 ++ ls3)
adamc@119 544 (pf'0 : ls1 ++ ls2 ++ ls3 = ls1 ++ ls2 ++ ls3)
adamc@119 545 (f : fhlist B (ls1 ++ ls2 ++ ls3)),
adamc@119 546 (a0,
adamc@119 547 match pf'0 in (_ = ls) return (fhlist B ls) with
adamc@119 548 | refl_equal => f
adamc@119 549 end) =
adamc@119 550 match pf0 in (_ = ls) return (fhlist B ls) with
adamc@119 551 | refl_equal => (a0, f)
adamc@119 552 end
adamc@218 553
adamc@119 554 ]]
adamc@119 555
adamc@119 556 We can see that we have achieved the crucial property: the type of each generalized equality proof has syntactically equal operands. This makes it easy to finish the proof with [UIP_refl]. *)
adamc@119 557
adamc@119 558 intros.
adamc@119 559 rewrite (UIP_refl _ _ pf0).
adamc@119 560 rewrite (UIP_refl _ _ pf'0).
adamc@119 561 reflexivity.
adamc@119 562 Qed.
adamc@124 563 (* end thide *)
adamc@119 564 End fhapp.
adamc@120 565
adamc@120 566 Implicit Arguments fhapp [A B ls1 ls2].
adamc@120 567
adamc@120 568
adamc@120 569 (** * Heterogeneous Equality *)
adamc@120 570
adamc@120 571 (** There is another equality predicate, defined in the [JMeq] module of the standard library, implementing %\textit{%#<i>#heterogeneous equality#</i>#%}%. *)
adamc@120 572
adamc@120 573 Print JMeq.
adamc@218 574 (** %\vspace{-.15in}% [[
adamc@120 575 Inductive JMeq (A : Type) (x : A) : forall B : Type, B -> Prop :=
adamc@120 576 JMeq_refl : JMeq x x
adamc@218 577
adamc@120 578 ]]
adamc@120 579
adam@292 580 [JMeq] stands for %``%#"#John Major equality,#"#%''% a name coined by Conor McBride as a sort of pun about British politics. [JMeq] starts out looking a lot like [eq]. The crucial difference is that we may use [JMeq] %\textit{%#<i>#on arguments of different types#</i>#%}%. For instance, a lemma that we failed to establish before is trivial with [JMeq]. It makes for prettier theorem statements to define some syntactic shorthand first. *)
adamc@120 581
adamc@120 582 Infix "==" := JMeq (at level 70, no associativity).
adamc@120 583
adamc@124 584 (* EX: Prove UIP_refl' : forall (A : Type) (x : A) (pf : x = x), pf == refl_equal x *)
adamc@124 585 (* begin thide *)
adamc@121 586 Definition UIP_refl' (A : Type) (x : A) (pf : x = x) : pf == refl_equal x :=
adamc@120 587 match pf return (pf == refl_equal _) with
adamc@120 588 | refl_equal => JMeq_refl _
adamc@120 589 end.
adamc@124 590 (* end thide *)
adamc@120 591
adamc@120 592 (** There is no quick way to write such a proof by tactics, but the underlying proof term that we want is trivial.
adamc@120 593
adamc@271 594 Suppose that we want to use [UIP_refl'] to establish another lemma of the kind we have run into several times so far. *)
adamc@120 595
adamc@120 596 Lemma lemma4 : forall (A : Type) (x : A) (pf : x = x),
adamc@120 597 O = match pf with refl_equal => O end.
adamc@124 598 (* begin thide *)
adamc@121 599 intros; rewrite (UIP_refl' pf); reflexivity.
adamc@120 600 Qed.
adamc@124 601 (* end thide *)
adamc@120 602
adamc@120 603 (** All in all, refreshingly straightforward, but there really is no such thing as a free lunch. The use of [rewrite] is implemented in terms of an axiom: *)
adamc@120 604
adamc@120 605 Check JMeq_eq.
adamc@218 606 (** %\vspace{-.15in}% [[
adamc@120 607 JMeq_eq
adamc@120 608 : forall (A : Type) (x y : A), x == y -> x = y
adamc@218 609
adamc@218 610 ]]
adamc@120 611
adamc@218 612 It may be surprising that we cannot prove that heterogeneous equality implies normal equality. The difficulties are the same kind we have seen so far, based on limitations of [match] annotations.
adamc@120 613
adamc@120 614 We can redo our [fhapp] associativity proof based around [JMeq]. *)
adamc@120 615
adamc@120 616 Section fhapp'.
adamc@120 617 Variable A : Type.
adamc@120 618 Variable B : A -> Type.
adamc@120 619
adamc@120 620 (** This time, the naive theorem statement type-checks. *)
adamc@120 621
adamc@124 622 (* EX: Prove [fhapp] associativity using [JMeq]. *)
adamc@124 623
adamc@124 624 (* begin thide *)
adam@297 625 Theorem fhapp_ass' : forall ls1 ls2 ls3 (hls1 : fhlist B ls1) (hls2 : fhlist B ls2) (hls3 : fhlist B ls3),
adamc@120 626 fhapp hls1 (fhapp hls2 hls3) == fhapp (fhapp hls1 hls2) hls3.
adamc@120 627 induction ls1; crush.
adamc@120 628
adamc@120 629 (** Even better, [crush] discharges the first subgoal automatically. The second subgoal is:
adamc@120 630
adamc@120 631 [[
adamc@120 632 ============================
adam@297 633 (a0, fhapp b (fhapp hls2 hls3)) == (a0, fhapp (fhapp b hls2) hls3)
adamc@218 634
adamc@120 635 ]]
adamc@120 636
adam@297 637 It looks like one rewrite with the inductive hypothesis should be enough to make the goal trivial. Here is what happens when we try that in Coq 8.2:
adamc@120 638
adamc@120 639 [[
adamc@120 640 rewrite IHls1.
adamc@120 641
adamc@120 642 Error: Impossible to unify "fhlist B ((ls1 ++ ?1572) ++ ?1573)" with
adamc@120 643 "fhlist B (ls1 ++ ?1572 ++ ?1573)"
adamc@218 644
adamc@120 645 ]]
adamc@120 646
adam@297 647 Coq 8.3 currently gives an error message about an uncaught exception. Perhaps that will be fixed soon. In any case, it is educational to consider a more explicit approach.
adam@297 648
adamc@120 649 We see that [JMeq] is not a silver bullet. We can use it to simplify the statements of equality facts, but the Coq type-checker uses non-trivial heterogeneous equality facts no more readily than it uses standard equality facts. Here, the problem is that the form [(e1, e2)] is syntactic sugar for an explicit application of a constructor of an inductive type. That application mentions the type of each tuple element explicitly, and our [rewrite] tries to change one of those elements without updating the corresponding type argument.
adamc@120 650
adamc@120 651 We can get around this problem by another multiple use of [generalize]. We want to bring into the goal the proper instance of the inductive hypothesis, and we also want to generalize the two relevant uses of [fhapp]. *)
adamc@120 652
adamc@120 653 generalize (fhapp b (fhapp hls2 hls3))
adamc@120 654 (fhapp (fhapp b hls2) hls3)
adamc@120 655 (IHls1 _ _ b hls2 hls3).
adamc@120 656 (** [[
adamc@120 657 ============================
adamc@120 658 forall (f : fhlist B (ls1 ++ ls2 ++ ls3))
adamc@120 659 (f0 : fhlist B ((ls1 ++ ls2) ++ ls3)), f == f0 -> (a0, f) == (a0, f0)
adamc@218 660
adamc@120 661 ]]
adamc@120 662
adamc@120 663 Now we can rewrite with append associativity, as before. *)
adamc@120 664
adamc@120 665 rewrite app_ass.
adamc@120 666 (** [[
adamc@120 667 ============================
adamc@120 668 forall f f0 : fhlist B (ls1 ++ ls2 ++ ls3), f == f0 -> (a0, f) == (a0, f0)
adamc@218 669
adamc@120 670 ]]
adamc@120 671
adamc@120 672 From this point, the goal is trivial. *)
adamc@120 673
adamc@120 674 intros f f0 H; rewrite H; reflexivity.
adamc@120 675 Qed.
adamc@124 676 (* end thide *)
adamc@120 677 End fhapp'.
adamc@121 678
adamc@121 679
adamc@121 680 (** * Equivalence of Equality Axioms *)
adamc@121 681
adamc@124 682 (* EX: Show that the approaches based on K and JMeq are equivalent logically. *)
adamc@124 683
adamc@124 684 (* begin thide *)
adamc@272 685 (** Assuming axioms (like axiom K and [JMeq_eq]) is a hazardous business. The due diligence associated with it is necessarily global in scope, since two axioms may be consistent alone but inconsistent together. It turns out that all of the major axioms proposed for reasoning about equality in Coq are logically equivalent, so that we only need to pick one to assert without proof. In this section, we demonstrate this by showing how each of the previous two sections' approaches reduces to the other logically.
adamc@121 686
adamc@121 687 To show that [JMeq] and its axiom let us prove [UIP_refl], we start from the lemma [UIP_refl'] from the previous section. The rest of the proof is trivial. *)
adamc@121 688
adamc@121 689 Lemma UIP_refl'' : forall (A : Type) (x : A) (pf : x = x), pf = refl_equal x.
adamc@121 690 intros; rewrite (UIP_refl' pf); reflexivity.
adamc@121 691 Qed.
adamc@121 692
adamc@121 693 (** The other direction is perhaps more interesting. Assume that we only have the axiom of the [Eqdep] module available. We can define [JMeq] in a way that satisfies the same interface as the combination of the [JMeq] module's inductive definition and axiom. *)
adamc@121 694
adamc@121 695 Definition JMeq' (A : Type) (x : A) (B : Type) (y : B) : Prop :=
adamc@121 696 exists pf : B = A, x = match pf with refl_equal => y end.
adamc@121 697
adamc@121 698 Infix "===" := JMeq' (at level 70, no associativity).
adamc@121 699
adamc@121 700 (** We say that, by definition, [x] and [y] are equal if and only if there exists a proof [pf] that their types are equal, such that [x] equals the result of casting [y] with [pf]. This statement can look strange from the standpoint of classical math, where we almost never mention proofs explicitly with quantifiers in formulas, but it is perfectly legal Coq code.
adamc@121 701
adamc@121 702 We can easily prove a theorem with the same type as that of the [JMeq_refl] constructor of [JMeq]. *)
adamc@121 703
adamc@121 704 (** remove printing exists *)
adamc@121 705 Theorem JMeq_refl' : forall (A : Type) (x : A), x === x.
adamc@121 706 intros; unfold JMeq'; exists (refl_equal A); reflexivity.
adamc@121 707 Qed.
adamc@121 708
adamc@121 709 (** printing exists $\exists$ *)
adamc@121 710
adamc@121 711 (** The proof of an analogue to [JMeq_eq] is a little more interesting, but most of the action is in appealing to [UIP_refl]. *)
adamc@121 712
adamc@121 713 Theorem JMeq_eq' : forall (A : Type) (x y : A),
adamc@121 714 x === y -> x = y.
adamc@121 715 unfold JMeq'; intros.
adamc@121 716 (** [[
adamc@121 717 H : exists pf : A = A,
adamc@121 718 x = match pf in (_ = T) return T with
adamc@121 719 | refl_equal => y
adamc@121 720 end
adamc@121 721 ============================
adamc@121 722 x = y
adamc@218 723
adam@302 724 ]]
adam@302 725 *)
adamc@121 726
adamc@121 727 destruct H.
adamc@121 728 (** [[
adamc@121 729 x0 : A = A
adamc@121 730 H : x = match x0 in (_ = T) return T with
adamc@121 731 | refl_equal => y
adamc@121 732 end
adamc@121 733 ============================
adamc@121 734 x = y
adamc@218 735
adam@302 736 ]]
adam@302 737 *)
adamc@121 738
adamc@121 739 rewrite H.
adamc@121 740 (** [[
adamc@121 741 x0 : A = A
adamc@121 742 ============================
adamc@121 743 match x0 in (_ = T) return T with
adamc@121 744 | refl_equal => y
adamc@121 745 end = y
adamc@218 746
adam@302 747 ]]
adam@302 748 *)
adamc@121 749
adamc@121 750 rewrite (UIP_refl _ _ x0); reflexivity.
adamc@121 751 Qed.
adamc@121 752
adam@292 753 (** We see that, in a very formal sense, we are free to switch back and forth between the two styles of proofs about equality proofs. One style may be more convenient than the other for some proofs, but we can always interconvert between our results. The style that does not use heterogeneous equality may be preferable in cases where many results do not require the tricks of this chapter, since then the use of axioms is avoided altogether for the simple cases, and a wider audience will be able to follow those %``%#"#simple#"#%''% proofs. On the other hand, heterogeneous equality often makes for shorter and more readable theorem statements.
adamc@123 754
adamc@123 755 It is worth remarking that it is possible to avoid axioms altogether for equalities on types with decidable equality. The [Eqdep_dec] module of the standard library contains a parametric proof of [UIP_refl] for such cases. *)
adamc@124 756 (* end thide *)
adamc@123 757
adamc@123 758
adamc@123 759 (** * Equality of Functions *)
adamc@123 760
adamc@123 761 (** The following seems like a reasonable theorem to want to hold, and it does hold in set theory. [[
adamc@123 762 Theorem S_eta : S = (fun n => S n).
adamc@218 763
adamc@205 764 ]]
adamc@205 765
adamc@123 766 Unfortunately, this theorem is not provable in CIC without additional axioms. None of the definitional equality rules force function equality to be %\textit{%#<i>#extensional#</i>#%}%. That is, the fact that two functions return equal results on equal inputs does not imply that the functions are equal. We %\textit{%#<i>#can#</i>#%}% assert function extensionality as an axiom. *)
adamc@123 767
adamc@124 768 (* begin thide *)
adamc@123 769 Axiom ext_eq : forall A B (f g : A -> B),
adamc@123 770 (forall x, f x = g x)
adamc@123 771 -> f = g.
adamc@124 772 (* end thide *)
adamc@123 773
adamc@123 774 (** This axiom has been verified metatheoretically to be consistent with CIC and the two equality axioms we considered previously. With it, the proof of [S_eta] is trivial. *)
adamc@123 775
adamc@123 776 Theorem S_eta : S = (fun n => S n).
adamc@124 777 (* begin thide *)
adamc@123 778 apply ext_eq; reflexivity.
adamc@123 779 Qed.
adamc@124 780 (* end thide *)
adamc@123 781
adam@292 782 (** The same axiom can help us prove equality of types, where we need to %``%#"#reason under quantifiers.#"#%''% *)
adamc@123 783
adamc@123 784 Theorem forall_eq : (forall x : nat, match x with
adamc@123 785 | O => True
adamc@123 786 | S _ => True
adamc@123 787 end)
adamc@123 788 = (forall _ : nat, True).
adamc@123 789
adamc@123 790 (** There are no immediate opportunities to apply [ext_eq], but we can use [change] to fix that. *)
adamc@123 791
adamc@124 792 (* begin thide *)
adamc@123 793 change ((forall x : nat, (fun x => match x with
adamc@123 794 | 0 => True
adamc@123 795 | S _ => True
adamc@123 796 end) x) = (nat -> True)).
adamc@123 797 rewrite (ext_eq (fun x => match x with
adamc@123 798 | 0 => True
adamc@123 799 | S _ => True
adamc@123 800 end) (fun _ => True)).
adamc@123 801 (** [[
adamc@123 802 2 subgoals
adamc@123 803
adamc@123 804 ============================
adamc@123 805 (nat -> True) = (nat -> True)
adamc@123 806
adamc@123 807 subgoal 2 is:
adamc@123 808 forall x : nat, match x with
adamc@123 809 | 0 => True
adamc@123 810 | S _ => True
adamc@123 811 end = True
adam@302 812 ]]
adam@302 813 *)
adamc@123 814
adamc@123 815 reflexivity.
adamc@123 816
adamc@123 817 destruct x; constructor.
adamc@123 818 Qed.
adamc@124 819 (* end thide *)
adamc@127 820
adamc@127 821
adamc@127 822 (** * Exercises *)
adamc@127 823
adamc@127 824 (** %\begin{enumerate}%#<ol>#
adamc@127 825
adamc@127 826 %\item%#<li># Implement and prove correct a substitution function for simply-typed lambda calculus. In particular:
adamc@127 827 %\begin{enumerate}%#<ol>#
adamc@127 828 %\item%#<li># Define a datatype [type] of lambda types, including just booleans and function types.#</li>#
adamc@127 829 %\item%#<li># Define a type family [exp : list type -> type -> Type] of lambda expressions, including boolean constants, variables, and function application and abstraction.#</li>#
adamc@127 830 %\item%#<li># Implement a definitional interpreter for [exp]s, by way of a recursive function over expressions and substitutions for free variables, like in the related example from the last chapter.#</li>#
adam@292 831 %\item%#<li># Implement a function [subst : forall t' ts t, exp (t' :: ts) t -> exp ts t' -> exp ts t]. The type of the first expression indicates that its most recently bound free variable has type [t']. The second expression also has type [t'], and the job of [subst] is to substitute the second expression for every occurrence of the %``%#"#first#"#%''% variable of the first expression.#</li>#
adamc@127 832 %\item%#<li># Prove that [subst] preserves program meanings. That is, prove
adamc@127 833 [[
adamc@127 834 forall t' ts t (e : exp (t' :: ts) t) (e' : exp ts t') (s : hlist typeDenote ts),
adamc@127 835 expDenote (subst e e') s = expDenote e (expDenote e' s ::: s)
adamc@218 836
adamc@127 837 ]]
adam@292 838 where [:::] is an infix operator for heterogeneous %``%#"#cons#"#%''% that is defined in the book's [DepList] module.#</li>#
adamc@127 839 #</ol>#%\end{enumerate}%
adamc@127 840 The material presented up to this point should be sufficient to enable a good solution of this exercise, with enough ingenuity. If you get stuck, it may be helpful to use the following structure. None of these elements need to appear in your solution, but we can at least guarantee that there is a reasonable solution based on them.
adamc@127 841 %\begin{enumerate}%#<ol>#
adamc@127 842 %\item%#<li># The [DepList] module will be useful. You can get the standard dependent list definitions there, instead of copying-and-pasting from the last chapter. It is worth reading the source for that module over, since it defines some new helpful functions and notations that we did not use last chapter.#</li>#
adam@292 843 %\item%#<li># Define a recursive function [liftVar : forall ts1 ts2 t t', member t (ts1 ++ ts2) -> member t (ts1 ++ t' :: ts2)]. This function should %``%#"#lift#"#%''% a de Bruijn variable so that its type refers to a new variable inserted somewhere in the index list.#</li>#
adam@292 844 %\item%#<li># Define a recursive function [lift' : forall ts t (e : exp ts t) ts1 ts2 t', ts = ts1 ++ ts2 -> exp (ts1 ++ t' :: ts2) t] which performs a similar lifting on an [exp]. The convoluted type is to get around restrictions on [match] annotations. We delay %``%#"#realizing#"#%''% that the first index of [e] is built with list concatenation until after a dependent [match], and the new explicit proof argument must be used to cast some terms that come up in the [match] body.#</li>#
adamc@127 845 %\item%#<li># Define a function [lift : forall ts t t', exp ts t -> exp (t' :: ts) t], which handles simpler top-level lifts. This should be an easy one-liner based on [lift'].#</li>#
adamc@127 846 %\item%#<li># Define a recursive function [substVar : forall ts1 ts2 t t', member t (ts1 ++ t' :: ts2) -> (t' = t) + member t (ts1 ++ ts2)]. This function is the workhorse behind substitution applied to a variable. It returns [inl] to indicate that the variable we pass to it is the variable that we are substituting for, and it returns [inr] to indicate that the variable we are examining is %\textit{%#<i>#not#</i>#%}% the one we are substituting for. In the first case, we get a proof that the necessary typing relationship holds, and, in the second case, we get the original variable modified to reflect the removal of the substitutee from the typing context.#</li>#
adamc@127 847 %\item%#<li># Define a recursive function [subst' : forall ts t (e : exp ts t) ts1 t' ts2, ts = ts1 ++ t' :: ts2 -> exp (ts1 ++ ts2) t' -> exp (ts1 ++ ts2) t]. This is the workhorse of substitution in expressions, employing the same proof-passing trick as for [lift']. You will probably want to use [lift] somewhere in the definition of [subst'].#</li>#
adamc@127 848 %\item%#<li># Now [subst] should be a one-liner, defined in terms of [subst'].#</li>#
adamc@127 849 %\item%#<li># Prove a correctness theorem for each auxiliary function, leading up to the proof of [subst] correctness.#</li>#
adam@292 850 %\item%#<li># All of the reasoning about equality proofs in these theorems follows a regular pattern. If you have an equality proof that you want to replace with [refl_equal] somehow, run [generalize] on that proof variable. Your goal is to get to the point where you can [rewrite] with the original proof to change the type of the generalized version. To avoid type errors (the infamous %``%#"#second-order unification#"#%''% failure messages), it will be helpful to run [generalize] on other pieces of the proof context that mention the equality's lefthand side. You might also want to use [generalize dependent], which generalizes not just one variable but also all variables whose types depend on it. [generalize dependent] has the sometimes-helpful property of removing from the context all variables that it generalizes. Once you do manage the mind-bending trick of using the equality proof to rewrite its own type, you will be able to rewrite with [UIP_refl].#</li>#
adamc@127 851 %\item%#<li># A variant of the [ext_eq] axiom from the end of this chapter is available in the book module [Axioms], and you will probably want to use it in the [lift'] and [subst'] correctness proofs.#</li>#
adam@292 852 %\item%#<li># The [change] tactic should come in handy in the proofs about [lift] and [subst], where you want to introduce %``%#"#extraneous#"#%''% list concatenations with [nil] to match the forms of earlier theorems.#</li>#
adam@292 853 %\item%#<li># Be careful about [destruct]ing a term %``%#"#too early.#"#%''% You can use [generalize] on proof terms to bring into the proof context any important propositions about the term. Then, when you [destruct] the term, it is updated in the extra propositions, too. The [case_eq] tactic is another alternative to this approach, based on saving an equality between the original term and its new form.#</li>#
adamc@127 854 #</ol>#%\end{enumerate}%
adamc@127 855 #</li>#
adamc@127 856
adamc@127 857 #</ol>#%\end{enumerate}% *)