Mercurial > cpdt > repo
comparison src/Equality.v @ 119:8df59f0b3dc0
Proofs with equality
author | Adam Chlipala <adamc@hcoop.net> |
---|---|
date | Sat, 18 Oct 2008 13:52:09 -0400 |
parents | ee676bf3d681 |
children | 39c7894b3f7f |
comparison
equal
deleted
inserted
replaced
118:ee676bf3d681 | 119:8df59f0b3dc0 |
---|---|
262 | 262 |
263 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. *) | 263 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. *) |
264 | 264 |
265 End fhlist_map. | 265 End fhlist_map. |
266 | 266 |
267 | |
268 (** * Type-Casts in Theorem Statements *) | |
269 | |
270 (** 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. *) | |
271 | |
272 Section fhapp. | |
273 Variable A : Type. | |
274 Variable B : A -> Type. | |
275 | |
276 Fixpoint fhapp (ls1 ls2 : list A) {struct ls1} | |
277 : fhlist B ls1 -> fhlist B ls2 -> fhlist B (ls1 ++ ls2) := | |
278 match ls1 return fhlist _ ls1 -> _ -> fhlist _ (ls1 ++ ls2) with | |
279 | nil => fun _ hls2 => hls2 | |
280 | _ :: _ => fun hls1 hls2 => (fst hls1, fhapp _ _ (snd hls1) hls2) | |
281 end. | |
282 | |
283 Implicit Arguments fhapp [ls1 ls2]. | |
284 | |
285 (** We might like to prove that [fhapp] is associative. | |
286 | |
287 [[ | |
288 | |
289 Theorem fhapp_ass : forall ls1 ls2 ls3 | |
290 (hls1 : fhlist B ls1) (hls2 : fhlist B ls2) (hls3 : fhlist B ls3), | |
291 fhapp hls1 (fhapp hls2 hls3) = fhapp (fhapp hls1 hls2) hls3. | |
292 | |
293 [[ | |
294 | |
295 The term | |
296 "fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3) (fhapp (ls1:=ls1) (ls2:=ls2) hls1 hls2) | |
297 hls3" has type "fhlist B ((ls1 ++ ls2) ++ ls3)" | |
298 while it is expected to have type "fhlist B (ls1 ++ ls2 ++ ls3)" | |
299 ]] | |
300 | |
301 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. *) | |
302 | |
303 Theorem fhapp_ass : forall ls1 ls2 ls3 | |
304 (pf : (ls1 ++ ls2) ++ ls3 = ls1 ++ (ls2 ++ ls3)) | |
305 (hls1 : fhlist B ls1) (hls2 : fhlist B ls2) (hls3 : fhlist B ls3), | |
306 fhapp hls1 (fhapp hls2 hls3) | |
307 = match pf in (_ = ls) return fhlist _ ls with | |
308 | refl_equal => fhapp (fhapp hls1 hls2) hls3 | |
309 end. | |
310 induction ls1; crush. | |
311 | |
312 (** The first remaining subgoal looks trivial enough: | |
313 | |
314 [[ | |
315 | |
316 ============================ | |
317 fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3 = | |
318 match pf in (_ = ls) return (fhlist B ls) with | |
319 | refl_equal => fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3 | |
320 end | |
321 ]] | |
322 | |
323 We can try what worked in previous examples. | |
324 | |
325 [[ | |
326 case pf. | |
327 | |
328 [[ | |
329 | |
330 User error: Cannot solve a second-order unification problem | |
331 ]] | |
332 | |
333 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. *) | |
334 | |
335 rewrite (UIP_refl _ _ pf). | |
336 (** [[ | |
337 | |
338 ============================ | |
339 fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3 = | |
340 fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3 | |
341 ]] *) | |
342 | |
343 reflexivity. | |
344 | |
345 (** Our second subgoal is trickier. | |
346 | |
347 [[ | |
348 | |
349 pf : a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3 | |
350 ============================ | |
351 (a0, | |
352 fhapp (ls1:=ls1) (ls2:=ls2 ++ ls3) b | |
353 (fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3)) = | |
354 match pf in (_ = ls) return (fhlist B ls) with | |
355 | refl_equal => | |
356 (a0, | |
357 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3) | |
358 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3) | |
359 end | |
360 ]] | |
361 | |
362 | |
363 [[ | |
364 | |
365 rewrite (UIP_refl _ _ pf). | |
366 | |
367 [[ | |
368 The term "pf" has type "a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3" | |
369 while it is expected to have type "?556 = ?556" | |
370 ]] | |
371 | |
372 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. *) | |
373 | |
374 injection pf; intro pf'. | |
375 (** [[ | |
376 | |
377 pf : a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3 | |
378 pf' : (ls1 ++ ls2) ++ ls3 = ls1 ++ ls2 ++ ls3 | |
379 ============================ | |
380 (a0, | |
381 fhapp (ls1:=ls1) (ls2:=ls2 ++ ls3) b | |
382 (fhapp (ls1:=ls2) (ls2:=ls3) hls2 hls3)) = | |
383 match pf in (_ = ls) return (fhlist B ls) with | |
384 | refl_equal => | |
385 (a0, | |
386 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3) | |
387 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3) | |
388 end | |
389 ]] | |
390 | |
391 Now we can rewrite using the inductive hypothesis. *) | |
392 | |
393 rewrite (IHls1 _ _ pf'). | |
394 (** [[ | |
395 | |
396 ============================ | |
397 (a0, | |
398 match pf' in (_ = ls) return (fhlist B ls) with | |
399 | refl_equal => | |
400 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3) | |
401 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3 | |
402 end) = | |
403 match pf in (_ = ls) return (fhlist B ls) with | |
404 | refl_equal => | |
405 (a0, | |
406 fhapp (ls1:=ls1 ++ ls2) (ls2:=ls3) | |
407 (fhapp (ls1:=ls1) (ls2:=ls2) b hls2) hls3) | |
408 end | |
409 ]] | |
410 | |
411 We have made an important bit of progress, as now only a single call to [fhapp] appears in the conclusion. 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. *) | |
412 | |
413 generalize (fhapp (fhapp b hls2) hls3). | |
414 (** [[ | |
415 | |
416 forall f : fhlist B ((ls1 ++ ls2) ++ ls3), | |
417 (a0, | |
418 match pf' in (_ = ls) return (fhlist B ls) with | |
419 | refl_equal => f | |
420 end) = | |
421 match pf in (_ = ls) return (fhlist B ls) with | |
422 | refl_equal => (a0, f) | |
423 end | |
424 ]] | |
425 | |
426 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. | |
427 | |
428 In this case, it is helpful to generalize over our two proofs as well. *) | |
429 | |
430 generalize pf pf'. | |
431 (** [[ | |
432 | |
433 forall (pf0 : a :: (ls1 ++ ls2) ++ ls3 = a :: ls1 ++ ls2 ++ ls3) | |
434 (pf'0 : (ls1 ++ ls2) ++ ls3 = ls1 ++ ls2 ++ ls3) | |
435 (f : fhlist B ((ls1 ++ ls2) ++ ls3)), | |
436 (a0, | |
437 match pf'0 in (_ = ls) return (fhlist B ls) with | |
438 | refl_equal => f | |
439 end) = | |
440 match pf0 in (_ = ls) return (fhlist B ls) with | |
441 | refl_equal => (a0, f) | |
442 end | |
443 ]] | |
444 | |
445 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. | |
446 | |
447 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. *) | |
448 | |
449 rewrite app_ass. | |
450 (** [[ | |
451 | |
452 ============================ | |
453 forall (pf0 : a :: ls1 ++ ls2 ++ ls3 = a :: ls1 ++ ls2 ++ ls3) | |
454 (pf'0 : ls1 ++ ls2 ++ ls3 = ls1 ++ ls2 ++ ls3) | |
455 (f : fhlist B (ls1 ++ ls2 ++ ls3)), | |
456 (a0, | |
457 match pf'0 in (_ = ls) return (fhlist B ls) with | |
458 | refl_equal => f | |
459 end) = | |
460 match pf0 in (_ = ls) return (fhlist B ls) with | |
461 | refl_equal => (a0, f) | |
462 end | |
463 ]] | |
464 | |
465 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]. *) | |
466 | |
467 intros. | |
468 rewrite (UIP_refl _ _ pf0). | |
469 rewrite (UIP_refl _ _ pf'0). | |
470 reflexivity. | |
471 Qed. | |
472 End fhapp. |