Mercurial > cpdt > repo
comparison src/DataStruct.v @ 111:702e5c35cc89
Code for cond-folding example
author | Adam Chlipala <adamc@hcoop.net> |
---|---|
date | Tue, 14 Oct 2008 10:52:38 -0400 |
parents | 4627b9caac8b |
children | 623478074c2f |
comparison
equal
deleted
inserted
replaced
110:4627b9caac8b | 111:702e5c35cc89 |
---|---|
6 * The license text is available at: | 6 * The license text is available at: |
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/ | 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/ |
8 *) | 8 *) |
9 | 9 |
10 (* begin hide *) | 10 (* begin hide *) |
11 Require Import List. | 11 Require Import Arith List. |
12 | 12 |
13 Require Import Tactics. | 13 Require Import Tactics. |
14 | 14 |
15 Set Implicit Arguments. | 15 Set Implicit Arguments. |
16 (* end hide *) | 16 (* end hide *) |
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 | 421 |
422 Implicit Arguments fhget [A B elm ls]. | |
423 | |
422 | 424 |
423 (** * Data Structures as Index Functions *) | 425 (** * Data Structures as Index Functions *) |
424 | 426 |
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. *) | 427 (** 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 | 428 |
568 | 570 |
569 induction t; crush. | 571 induction t; crush. |
570 Qed. | 572 Qed. |
571 | 573 |
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. *) | 574 (** 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. *) |
575 | |
576 (** ** Another Interpreter Example *) | |
577 | |
578 Inductive type' : Type := | |
579 | Nat' : type' | |
580 | Bool' : type'. | |
581 | |
582 Inductive exp' : type' -> Type := | |
583 | NConst : nat -> exp' Nat' | |
584 | Plus : exp' Nat' -> exp' Nat' -> exp' Nat' | |
585 | Eq : exp' Nat' -> exp' Nat' -> exp' Bool' | |
586 | |
587 | BConst : bool -> exp' Bool' | |
588 | Cond : forall n t, (findex n -> exp' Bool') | |
589 -> (findex n -> exp' t) -> exp' t -> exp' t. | |
590 | |
591 Definition type'Denote (t : type') : Set := | |
592 match t with | |
593 | Nat' => nat | |
594 | Bool' => bool | |
595 end. | |
596 | |
597 Section cond. | |
598 Variable A : Set. | |
599 Variable default : A. | |
600 | |
601 Fixpoint cond (n : nat) : (findex n -> bool) -> (findex n -> A) -> A := | |
602 match n return (findex n -> bool) -> (findex n -> A) -> A with | |
603 | O => fun _ _ => default | |
604 | S n' => fun tests bodies => | |
605 if tests None | |
606 then bodies None | |
607 else cond n' | |
608 (fun idx => tests (Some idx)) | |
609 (fun idx => bodies (Some idx)) | |
610 end. | |
611 End cond. | |
612 | |
613 Implicit Arguments cond [A n]. | |
614 | |
615 Fixpoint exp'Denote t (e : exp' t) {struct e} : type'Denote t := | |
616 match e in exp' t return type'Denote t with | |
617 | NConst n => | |
618 n | |
619 | Plus e1 e2 => | |
620 exp'Denote e1 + exp'Denote e2 | |
621 | Eq e1 e2 => | |
622 if eq_nat_dec (exp'Denote e1) (exp'Denote e2) then true else false | |
623 | |
624 | BConst b => | |
625 b | |
626 | Cond _ _ tests bodies default => | |
627 cond | |
628 (exp'Denote default) | |
629 (fun idx => exp'Denote (tests idx)) | |
630 (fun idx => exp'Denote (bodies idx)) | |
631 end. | |
632 | |
633 Section cfoldCond. | |
634 Variable t : type'. | |
635 Variable default : exp' t. | |
636 | |
637 Fixpoint cfoldCond (n : nat) : (findex n -> exp' Bool') -> (findex n -> exp' t) -> exp' t := | |
638 match n return (findex n -> exp' Bool') -> (findex n -> exp' t) -> exp' t with | |
639 | O => fun _ _ => default | |
640 | S n' => fun tests bodies => | |
641 match tests None with | |
642 | BConst true => bodies None | |
643 | BConst false => cfoldCond n' | |
644 (fun idx => tests (Some idx)) | |
645 (fun idx => bodies (Some idx)) | |
646 | _ => | |
647 let e := cfoldCond n' | |
648 (fun idx => tests (Some idx)) | |
649 (fun idx => bodies (Some idx)) in | |
650 match e in exp' t return exp' Bool' -> exp' t -> exp' t with | |
651 | Cond n _ tests' bodies' default' => fun test body => | |
652 Cond | |
653 (S n) | |
654 (fun idx => match idx with | |
655 | None => test | |
656 | Some idx => tests' idx | |
657 end) | |
658 (fun idx => match idx with | |
659 | None => body | |
660 | Some idx => bodies' idx | |
661 end) | |
662 default' | |
663 | e => fun test body => | |
664 Cond | |
665 1 | |
666 (fun _ => test) | |
667 (fun _ => body) | |
668 e | |
669 end (tests None) (bodies None) | |
670 end | |
671 end. | |
672 End cfoldCond. | |
673 | |
674 Implicit Arguments cfoldCond [t n]. | |
675 | |
676 Fixpoint cfold t (e : exp' t) {struct e} : exp' t := | |
677 match e in exp' t return exp' t with | |
678 | NConst n => NConst n | |
679 | Plus e1 e2 => | |
680 let e1' := cfold e1 in | |
681 let e2' := cfold e2 in | |
682 match e1', e2' with | |
683 | NConst n1, NConst n2 => NConst (n1 + n2) | |
684 | _, _ => Plus e1' e2' | |
685 end | |
686 | Eq e1 e2 => | |
687 let e1' := cfold e1 in | |
688 let e2' := cfold e2 in | |
689 match e1', e2' with | |
690 | NConst n1, NConst n2 => BConst (if eq_nat_dec n1 n2 then true else false) | |
691 | _, _ => Eq e1' e2' | |
692 end | |
693 | |
694 | BConst b => BConst b | |
695 | Cond _ _ tests bodies default => | |
696 cfoldCond | |
697 (cfold default) | |
698 (fun idx => cfold (tests idx)) | |
699 (fun idx => cfold (bodies idx)) | |
700 end. | |
701 | |
702 Lemma cfoldCond_correct : forall t (default : exp' t) | |
703 n (tests : findex n -> exp' Bool') (bodies : findex n -> exp' t), | |
704 exp'Denote (cfoldCond default tests bodies) | |
705 = exp'Denote (Cond n tests bodies default). | |
706 induction n; crush; | |
707 match goal with | |
708 | [ IHn : forall tests bodies, _, tests : _ -> _, bodies : _ -> _ |- _ ] => | |
709 generalize (IHn (fun idx => tests (Some idx)) (fun idx => bodies (Some idx))); | |
710 clear IHn; intro IHn | |
711 end; | |
712 repeat (match goal with | |
713 | [ |- context[match ?E with | |
714 | NConst _ => _ | |
715 | Plus _ _ => _ | |
716 | Eq _ _ => _ | |
717 | BConst _ => _ | |
718 | Cond _ _ _ _ _ => _ | |
719 end] ] => dep_destruct E | |
720 | [ |- context[if ?B then _ else _] ] => destruct B | |
721 end; crush). | |
722 Qed. | |
723 | |
724 Lemma cond_ext : forall (A : Set) (default : A) n (tests tests' : findex n -> bool) | |
725 (bodies bodies' : findex n -> A), | |
726 (forall idx, tests idx = tests' idx) | |
727 -> (forall idx, bodies idx = bodies' idx) | |
728 -> cond default tests bodies | |
729 = cond default tests' bodies'. | |
730 induction n; crush; | |
731 match goal with | |
732 | [ |- context[if ?E then _ else _] ] => destruct E | |
733 end; crush. | |
734 Qed. | |
735 | |
736 Theorem cfold_correct : forall t (e : exp' t), | |
737 exp'Denote (cfold e) = exp'Denote e. | |
738 Hint Rewrite cfoldCond_correct : cpdt. | |
739 Hint Resolve cond_ext. | |
740 | |
741 induction e; crush; | |
742 repeat (match goal with | |
743 | [ |- context[cfold ?E] ] => dep_destruct (cfold E) | |
744 end; crush). | |
745 Qed. |