Mercurial > cpdt > repo
comparison src/Reflection.v @ 378:6413675f8e04
Recursing under binders in reification
author | Adam Chlipala <adam@chlipala.net> |
---|---|
date | Thu, 29 Mar 2012 17:13:23 -0400 |
parents | 549d604c3d16 |
children | e23d41ae63be |
comparison
equal
deleted
inserted
replaced
377:252ba054a1cd | 378:6413675f8e04 |
---|---|
739 *) | 739 *) |
740 Abort. | 740 Abort. |
741 | 741 |
742 (** More work would be needed to complete the reflective tactic, as we must connect our new syntax type with the real meanings of formulas, but the details are the same as in our prior implementation with [quote]. *) | 742 (** More work would be needed to complete the reflective tactic, as we must connect our new syntax type with the real meanings of formulas, but the details are the same as in our prior implementation with [quote]. *) |
743 (* end thide *) | 743 (* end thide *) |
744 | |
745 | |
746 (** * Building a Reification Tactic that Recurses Under Binders *) | |
747 | |
748 (** All of our examples so far have stayed away from reifying the syntax of terms that use such features as quantifiers and [fun] function abstractions. Such cases are complicated by the fact that different subterms may be allowed to reference different sets of free variables. Some cleverness is needed to clear this hurdle, but a few simple patterns will suffice. Consider this example of a simple dependently typed term language, where a function abstraction body is represented conveniently with a Coq function. *) | |
749 | |
750 Inductive type : Type := | |
751 | Nat : type | |
752 | NatFunc : type -> type. | |
753 | |
754 Inductive term : type -> Type := | |
755 | Const : nat -> term Nat | |
756 | Plus : term Nat -> term Nat -> term Nat | |
757 | Abs : forall t, (nat -> term t) -> term (NatFunc t). | |
758 | |
759 Fixpoint typeDenote (t : type) : Type := | |
760 match t with | |
761 | Nat => nat | |
762 | NatFunc t => nat -> typeDenote t | |
763 end. | |
764 | |
765 Fixpoint termDenote t (e : term t) : typeDenote t := | |
766 match e with | |
767 | Const n => n | |
768 | Plus e1 e2 => termDenote e1 + termDenote e2 | |
769 | Abs _ e1 => fun x => termDenote (e1 x) | |
770 end. | |
771 | |
772 (** Here is a naive first attempt at a reification tactic. *) | |
773 | |
774 Ltac refl' e := | |
775 match e with | |
776 | ?E1 + ?E2 => | |
777 let r1 := refl' E1 in | |
778 let r2 := refl' E2 in | |
779 constr:(Plus r1 r2) | |
780 | |
781 | fun x : nat => ?E1 => | |
782 let r1 := refl' E1 in | |
783 constr:(Abs (fun x => r1 x)) | |
784 | |
785 | _ => constr:(Const e) | |
786 end. | |
787 | |
788 (** Recall that a regular Ltac pattern variable [?X] only matches terms that %\emph{%#<i>#do not mention new variables introduced within the pattern#</i>#%}%. In our naive implementation, the case for matching function abstractions matches the function body in a way that prevents it from mentioning the function argument! Our code above plays fast and loose with the function body in a way that leads to independent problems, but we could change the code so that it indeed handles function abstractions that ignore their arguments. | |
789 | |
790 To handle functions in general, we will use the pattern variable form [@?X], which allows [X] to mention newly introduced variables that are declared explicitly. For instance: *) | |
791 | |
792 Reset refl'. | |
793 Ltac refl' e := | |
794 match e with | |
795 | ?E1 + ?E2 => | |
796 let r1 := refl' E1 in | |
797 let r2 := refl' E2 in | |
798 constr:(Plus r1 r2) | |
799 | |
800 | fun x : nat => @?E1 x => | |
801 let r1 := refl' E1 in | |
802 constr:(Abs r1) | |
803 | |
804 | _ => constr:(Const e) | |
805 end. | |
806 | |
807 (** Now, in the abstraction case, we bind [E1] as a function from an [x] value to the value of the abstraction body. Unfortunately, our recursive call there is not destined for success. It will match the same abstraction pattern and trigger another recursive call, and so on through infinite recursion. One last refactoring yields a working procedure. The key idea is to consider every input to [refl'] as %\emph{%#<i>#a function over the values of variables introduced during recursion#</i>#%}%. *) | |
808 | |
809 Reset refl'. | |
810 Ltac refl' e := | |
811 match eval simpl in e with | |
812 | fun x : ?T => @?E1 x + @?E2 x => | |
813 let r1 := refl' E1 in | |
814 let r2 := refl' E2 in | |
815 constr:(fun x => Plus (r1 x) (r2 x)) | |
816 | |
817 | fun (x : ?T) (y : nat) => @?E1 x y => | |
818 let r1 := refl' (fun p : T * nat => E1 (fst p) (snd p)) in | |
819 constr:(fun x => Abs (fun y => r1 (x, y))) | |
820 | |
821 | _ => constr:(fun x => Const (e x)) | |
822 end. | |
823 | |
824 (** Note how now even the addition case works in terms of functions, with [@?X] patterns. The abstraction case introduces a new variable by extending the type used to represent the free variables. In particular, the argument to [refl'] used type [T] to represent all free variables. We extend the type to [T * nat] for the type representing free variable values within the abstraction body. A bit of bookkeeping with pairs and their projections produces an appropriate version of the abstraction body to pass in a recursive call. To ensure that all this repackaging of terms does not interfere with pattern matching, we add an extra [simpl] reduction on the function argument, in the first line of the body of [refl']. | |
825 | |
826 Now one more tactic provides an example of how to apply reification. Let us consider goals that are equalities between terms that can be reified. We want to change such goals into equalities between appropriate calls to [termDenote]. *) | |
827 | |
828 Ltac refl := | |
829 match goal with | |
830 | [ |- ?E1 = ?E2 ] => | |
831 let E1' := refl' (fun _ : unit => E1) in | |
832 let E2' := refl' (fun _ : unit => E2) in | |
833 change (termDenote (E1' tt) = termDenote (E2' tt)); | |
834 cbv beta iota delta [fst snd] | |
835 end. | |
836 | |
837 Goal (fun (x y : nat) => x + y + 13) = (fun (_ z : nat) => z). | |
838 refl. | |
839 (** %\vspace{-.15in}%[[ | |
840 ============================ | |
841 termDenote | |
842 (Abs | |
843 (fun y : nat => | |
844 Abs (fun y0 : nat => Plus (Plus (Const y) (Const y0)) (Const 13)))) = | |
845 termDenote (Abs (fun _ : nat => Abs (fun y0 : nat => Const y0))) | |
846 ]] | |
847 *) | |
848 | |
849 Abort. | |
850 | |
851 (** Our encoding here uses Coq functions to represent binding within the terms we reify, which makes it difficult to implement certain functions over reified terms. An alternative would be to represent variables with numbers. This can be done by writing a slightly smarter reification function that detects variable references by detecting when term arguments are just compositions of [fst] and [snd]; from the order of the compositions we may read off the variable number. We leave the details as an exercise for the reader. *) |