comparison src/DataStruct.v @ 110:4627b9caac8b

Index functions
author Adam Chlipala <adamc@hcoop.net>
date Mon, 13 Oct 2008 15:09:58 -0400
parents fe6cfbae86b9
children 702e5c35cc89
comparison
equal deleted inserted replaced
109:fe6cfbae86b9 110:4627b9caac8b
416 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x 416 Inductive eq (A : Type) (x : A) : A -> Prop := refl_equal : x = x
417 ]] 417 ]]
418 418
419 In a proposition [x = y], we see that [x] is a parameter and [y] is a regular argument. The type of the constructor [refl_equal] shows that [y] can only ever be instantiated to [x]. Thus, within a pattern-match with [refl_equal], occurrences of [y] can be replaced with occurrences of [x] for typing purposes. All examples of similar dependent pattern matching that we have seen before require explicit annotations, but Coq implements a special case of annotation inference for matches on equality proofs. *) 419 In a proposition [x = y], we see that [x] is a parameter and [y] is a regular argument. The type of the constructor [refl_equal] shows that [y] can only ever be instantiated to [x]. Thus, within a pattern-match with [refl_equal], occurrences of [y] can be replaced with occurrences of [x] for typing purposes. All examples of similar dependent pattern matching that we have seen before require explicit annotations, but Coq implements a special case of annotation inference for matches on equality proofs. *)
420 End fhlist. 420 End fhlist.
421
422
423 (** * Data Structures as Index Functions *)
424
425 (** Indexed lists can be useful in defining other inductive types with constructors that take variable numbers of arguments. In this section, we consider parameterized trees with arbitrary branching factor. *)
426
427 Section tree.
428 Variable A : Set.
429
430 Inductive tree : Set :=
431 | Leaf : A -> tree
432 | Node : forall n, ilist tree n -> tree.
433 End tree.
434
435 (** Every [Node] of a [tree] has a natural number argument, which gives the number of child trees in the second argument, typed with [ilist]. We can define two operations on trees of naturals: summing their elements and incrementing their elements. It is useful to define a generic fold function on [ilist]s first. *)
436
437 Section ifoldr.
438 Variables A B : Set.
439 Variable f : A -> B -> B.
440 Variable i : B.
441
442 Fixpoint ifoldr n (ls : ilist A n) {struct ls} : B :=
443 match ls with
444 | Nil => i
445 | Cons _ x ls' => f x (ifoldr ls')
446 end.
447 End ifoldr.
448
449 Fixpoint sum (t : tree nat) : nat :=
450 match t with
451 | Leaf n => n
452 | Node _ ls => ifoldr (fun t' n => sum t' + n) O ls
453 end.
454
455 Fixpoint inc (t : tree nat) : tree nat :=
456 match t with
457 | Leaf n => Leaf (S n)
458 | Node _ ls => Node (imap inc ls)
459 end.
460
461 (** Now we might like to prove that [inc] does not decrease a tree's [sum]. *)
462
463 Theorem sum_inc : forall t, sum (inc t) >= sum t.
464 induction t; crush.
465 (** [[
466
467 n : nat
468 i : ilist (tree nat) n
469 ============================
470 ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 (imap inc i) >=
471 ifoldr (fun (t' : tree nat) (n0 : nat) => sum t' + n0) 0 i
472 ]]
473
474 We are left with a single subgoal which does not seem provable directly. This is the same problem that we met in Chapter 3 with other nested inductive types. *)
475
476 Check tree_ind.
477 (** [[
478
479 tree_ind
480 : forall (A : Set) (P : tree A -> Prop),
481 (forall a : A, P (Leaf a)) ->
482 (forall (n : nat) (i : ilist (tree A) n), P (Node i)) ->
483 forall t : tree A, P t
484 ]]
485
486 The automatically-generated induction principle is too weak. For the [Node] case, it gives us no inductive hypothesis. We could write our own induction principle, as we did in Chapter 3, but there is an easier way, if we are willing to alter the definition of [tree]. *)
487 Abort.
488
489 Reset tree.
490
491 (** First, let us try using our recursive definition of [ilist]s instead of the inductive version. *)
492
493 Section tree.
494 Variable A : Set.
495
496 (** [[
497
498 Inductive tree : Set :=
499 | Leaf : A -> tree
500 | Node : forall n, filist tree n -> tree.
501
502 [[
503
504 Error: Non strictly positive occurrence of "tree" in
505 "forall n : nat, filist tree n -> tree"
506 ]]
507
508 The special-case rule for nested datatypes only works with nested uses of other inductive types, which could be replaced with uses of new mutually-inductive types. We defined [filist] recursively, so it may not be used for nested recursion.
509
510 Our final solution uses yet another of the inductive definition techniques introduced in Chapter 3, reflexive types. Instead of merely using [index] to get elements out of [ilist], we can %\textit{%#<i>#define#</i>#%}% [ilist] in terms of [index]. For the reasons outlined above, it turns out to be easier to work with [findex] in place of [index]. *)
511
512 Inductive tree : Set :=
513 | Leaf : A -> tree
514 | Node : forall n, (findex n -> tree) -> tree.
515
516 (** A [Node] is indexed by a natural number [n], and the node's [n] children are represented as a function from [findex n] to trees, which is isomorphic to the [ilist]-based representation that we used above. *)
517 End tree.
518
519 Implicit Arguments Node [A n].
520
521 (** We can redefine [sum] and [inc] for our new [tree] type. Again, it is useful to define a generic fold function first. This time, it takes in a function whose range is some [findex] type, and it folds another function over the results of calling the first function at every possible [findex] value. *)
522
523 Section rifoldr.
524 Variables A B : Set.
525 Variable f : A -> B -> B.
526 Variable i : B.
527
528 Fixpoint rifoldr (n : nat) : (findex n -> A) -> B :=
529 match n return (findex n -> A) -> B with
530 | O => fun _ => i
531 | S n' => fun get => f (get None) (rifoldr n' (fun idx => get (Some idx)))
532 end.
533 End rifoldr.
534
535 Implicit Arguments rifoldr [A B n].
536
537 Fixpoint sum (t : tree nat) : nat :=
538 match t with
539 | Leaf n => n
540 | Node _ f => rifoldr plus O (fun idx => sum (f idx))
541 end.
542
543 Fixpoint inc (t : tree nat) : tree nat :=
544 match t with
545 | Leaf n => Leaf (S n)
546 | Node _ f => Node (fun idx => inc (f idx))
547 end.
548
549 (** Now we are ready to prove the theorem where we got stuck before. We will not need to define any new induction principle, but it %\textit{%#<i>#will#</i>#%}% be helpful to prove some lemmas. *)
550
551 Lemma plus_ge : forall x1 y1 x2 y2,
552 x1 >= x2
553 -> y1 >= y2
554 -> x1 + y1 >= x2 + y2.
555 crush.
556 Qed.
557
558 Lemma sum_inc' : forall n (f1 f2 : findex n -> nat),
559 (forall idx, f1 idx >= f2 idx)
560 -> rifoldr plus 0 f1 >= rifoldr plus 0 f2.
561 Hint Resolve plus_ge.
562
563 induction n; crush.
564 Qed.
565
566 Theorem sum_inc : forall t, sum (inc t) >= sum t.
567 Hint Resolve sum_inc'.
568
569 induction t; crush.
570 Qed.
571
572 (** Even if Coq would generate complete induction principles automatically for nested inductive definitions like the one we started with, there would still be advantages to using this style of reflexive encoding. We see one of those advantages in the definition of [inc], where we did not need to use any kind of auxiliary function. In general, reflexive encodings often admit direct implementations of operations that would require recursion if performed with more traditional inductive data structures. *)