Mercurial > cpdt > repo
comparison src/GeneralRec.v @ 356:50e1d338728c
First draft of full prose for GeneralRec
author | Adam Chlipala <adam@chlipala.net> |
---|---|
date | Fri, 28 Oct 2011 17:43:53 -0400 |
parents | 62fdf0993e05 |
children | b01c7b3122cc |
comparison
equal
deleted
inserted
replaced
355:62fdf0993e05 | 356:50e1d338728c |
---|---|
398 Theorem run_Return : run Return v. | 398 Theorem run_Return : run Return v. |
399 run. | 399 run. |
400 Qed. | 400 Qed. |
401 End Return. | 401 End Return. |
402 | 402 |
403 (** The name [Return] was meant to be suggestive of the standard operations of monads%~\cite{Monads}%. The other standard operation is [Bind], which lets us run one computation and, if it terminates, pass its result off to another computation. *) | 403 (** The name [Return] was meant to be suggestive of the standard operations of %\index{monad}%monads%~\cite{Monads}%. The other standard operation is [Bind], which lets us run one computation and, if it terminates, pass its result off to another computation. *) |
404 | 404 |
405 Section Bind. | 405 Section Bind. |
406 Variables A B : Type. | 406 Variables A B : Type. |
407 Variable m1 : computation A. | 407 Variable m1 : computation A. |
408 Variable m2 : A -> computation B. | 408 Variable m2 : A -> computation B. |
595 There are other theorems that are important to prove about combinators like [Return], [Bind], and [Fix]. In general, for a computation [c], we sometimes have a hypothesis proving [run c v] for some [v], and we want to perform inversion to deduce what [v] must be. Each combinator should ideally have a theorem of that kind, for [c] built directly from that combinator. We have omitted such theorems here, but they are not hard to prove. In general, the domain theory-inspired approach avoids the type-theoretic %``%#"#gotchas#"#%''% that tend to show up in approaches that try to mix normal Coq computation with explicit syntax types. The next section of this chapter demonstrates two alternate approaches of that sort. *) | 595 There are other theorems that are important to prove about combinators like [Return], [Bind], and [Fix]. In general, for a computation [c], we sometimes have a hypothesis proving [run c v] for some [v], and we want to perform inversion to deduce what [v] must be. Each combinator should ideally have a theorem of that kind, for [c] built directly from that combinator. We have omitted such theorems here, but they are not hard to prove. In general, the domain theory-inspired approach avoids the type-theoretic %``%#"#gotchas#"#%''% that tend to show up in approaches that try to mix normal Coq computation with explicit syntax types. The next section of this chapter demonstrates two alternate approaches of that sort. *) |
596 | 596 |
597 | 597 |
598 (** * Co-Inductive Non-Termination Monads *) | 598 (** * Co-Inductive Non-Termination Monads *) |
599 | 599 |
600 (** There are two key downsides to both of the previous approaches: both require unusual syntax based on explicit calls to fixpoint combinators, and both generate immediate proof obligations about the bodies of recursive definitions. In Chapter 5, we have already seen how co-inductive types support recursive definitions that exhibit certain well-behaved varieties of non-termination. It turns out that we can leverage that co-induction support for encoding of general recursive definitions, by adding layers of co-inductive syntax. In effect, we mix elements of shallow and deep embeddings. | |
601 | |
602 Our first example of this kind, proposed by Capretta%~\cite{Capretta}%, defines a silly-looking type of thunks; that is, computations that may be forced to yield results, if they terminate. *) | |
603 | |
600 CoInductive thunk (A : Type) : Type := | 604 CoInductive thunk (A : Type) : Type := |
601 | Answer : A -> thunk A | 605 | Answer : A -> thunk A |
602 | Think : thunk A -> thunk A. | 606 | Think : thunk A -> thunk A. |
603 | 607 |
604 CoFixpoint TBind (A B : Type) (m1 : thunk A) (m2 : A -> thunk B) : thunk B := | 608 (** A computation is either an immediate [Answer] or another computation wrapped inside [Think]. Since [thunk] is co-inductive, every [thunk] type is inhabited by an infinite nesting of [Think]s, standing for non-termination. Terminating results are [Answer] wrapped inside some finite number of [Think]s. |
609 | |
610 Why bother to write such a strange definition? The definition of [thunk] is motivated by the ability it gives us to define a %``%#"#bind#"#%''% operation, similar to the one we defined in the previous section. *) | |
611 | |
612 CoFixpoint TBind A B (m1 : thunk A) (m2 : A -> thunk B) : thunk B := | |
605 match m1 with | 613 match m1 with |
606 | Answer x => m2 x | 614 | Answer x => m2 x |
607 | Think m1' => Think (TBind m1' m2) | 615 | Think m1' => Think (TBind m1' m2) |
608 end. | 616 end. |
609 | 617 |
618 (** Note that the definition would violate the co-recursion guardedness restriction if we left out the seemingly superfluous [Think] on the righthand side of the second [match] branch. | |
619 | |
620 We can prove that [Answer] and [TBind] form a monad for [thunk]. The proof is omitted here but present in the book source. As usual for this sort of proof, a key element is choosing an appropriate notion of equality for [thunk]s. *) | |
621 | |
622 (* begin hide *) | |
610 CoInductive thunk_eq A : thunk A -> thunk A -> Prop := | 623 CoInductive thunk_eq A : thunk A -> thunk A -> Prop := |
611 | EqAnswer : forall x, thunk_eq (Answer x) (Answer x) | 624 | EqAnswer : forall x, thunk_eq (Answer x) (Answer x) |
612 | EqThinkL : forall m1 m2, thunk_eq m1 m2 -> thunk_eq (Think m1) m2 | 625 | EqThinkL : forall m1 m2, thunk_eq m1 m2 -> thunk_eq (Think m1) m2 |
613 | EqThinkR : forall m1 m2, thunk_eq m1 m2 -> thunk_eq m1 (Think m2). | 626 | EqThinkR : forall m1 m2, thunk_eq m1 m2 -> thunk_eq m1 (Think m2). |
614 | 627 |
629 match goal with | 642 match goal with |
630 | [ H' : P _ _ |- _ ] => specialize (H H'); clear H' | 643 | [ H' : P _ _ |- _ ] => specialize (H H'); clear H' |
631 end; destruct m1; destruct m2; subst; repeat constructor; auto. | 644 end; destruct m1; destruct m2; subst; repeat constructor; auto. |
632 Qed. | 645 Qed. |
633 End thunk_eq_coind. | 646 End thunk_eq_coind. |
647 (* end hide *) | |
648 | |
649 (** In the proofs to follow, we will need a function similar to one we saw in Chapter 5, to pull apart and reassemble a [thunk] in a way that provokes reduction of co-recursive calls. *) | |
634 | 650 |
635 Definition frob A (m : thunk A) : thunk A := | 651 Definition frob A (m : thunk A) : thunk A := |
636 match m with | 652 match m with |
637 | Answer x => Answer x | 653 | Answer x => Answer x |
638 | Think m' => Think m' | 654 | Think m' => Think m' |
640 | 656 |
641 Theorem frob_eq : forall A (m : thunk A), frob m = m. | 657 Theorem frob_eq : forall A (m : thunk A), frob m = m. |
642 destruct m; reflexivity. | 658 destruct m; reflexivity. |
643 Qed. | 659 Qed. |
644 | 660 |
661 (* begin hide *) | |
645 Theorem thunk_eq_frob : forall A (m1 m2 : thunk A), | 662 Theorem thunk_eq_frob : forall A (m1 m2 : thunk A), |
646 thunk_eq (frob m1) (frob m2) | 663 thunk_eq (frob m1) (frob m2) |
647 -> thunk_eq m1 m2. | 664 -> thunk_eq m1 m2. |
648 intros; repeat rewrite frob_eq in *; auto. | 665 intros; repeat rewrite frob_eq in *; auto. |
649 Qed. | 666 Qed. |
688 intros; apply (thunk_eq_coind (fun m1 m2 => (exists m, | 705 intros; apply (thunk_eq_coind (fun m1 m2 => (exists m, |
689 m1 = TBind (TBind m f) g | 706 m1 = TBind (TBind m f) g |
690 /\ m2 = TBind m (fun x => TBind (f x) g)) | 707 /\ m2 = TBind m (fun x => TBind (f x) g)) |
691 \/ m1 = m2)); crush; eauto; repeat (findDestr; crush; eauto). | 708 \/ m1 = m2)); crush; eauto; repeat (findDestr; crush; eauto). |
692 Qed. | 709 Qed. |
710 (* end hide *) | |
711 | |
712 (** As a simple example, here is how we might define a tail-recursive factorial function. *) | |
693 | 713 |
694 CoFixpoint fact (n acc : nat) : thunk nat := | 714 CoFixpoint fact (n acc : nat) : thunk nat := |
695 match n with | 715 match n with |
696 | O => Answer acc | 716 | O => Answer acc |
697 | S n' => Think (fact n' (S n' * acc)) | 717 | S n' => Think (fact n' (S n' * acc)) |
698 end. | 718 end. |
699 | 719 |
720 (** To test our definition, we need an evaluation relation that characterizes results of evaluating [thunk]s. *) | |
721 | |
700 Inductive eval A : thunk A -> A -> Prop := | 722 Inductive eval A : thunk A -> A -> Prop := |
701 | EvalAnswer : forall x, eval (Answer x) x | 723 | EvalAnswer : forall x, eval (Answer x) x |
702 | EvalThink : forall m x, eval m x -> eval (Think m) x. | 724 | EvalThink : forall m x, eval m x -> eval (Think m) x. |
703 | 725 |
704 Hint Rewrite frob_eq : cpdt. | 726 Hint Rewrite frob_eq : cpdt. |
711 | 733 |
712 Theorem eval_fact : eval (fact 5 1) 120. | 734 Theorem eval_fact : eval (fact 5 1) 120. |
713 repeat (apply eval_frob; simpl; constructor). | 735 repeat (apply eval_frob; simpl; constructor). |
714 Qed. | 736 Qed. |
715 | 737 |
716 (** [[ | 738 (** We need to apply constructors of [eval] explicitly, but the process is easy to automate completely for concrete input programs. |
739 | |
740 Now consider another very similar definition, this time of a Fibonacci number funtion. | |
741 [[ | |
717 CoFixpoint fib (n : nat) : thunk nat := | 742 CoFixpoint fib (n : nat) : thunk nat := |
718 match n with | 743 match n with |
719 | 0 => Answer 1 | 744 | 0 => Answer 1 |
720 | 1 => Answer 1 | 745 | 1 => Answer 1 |
721 | _ => TBind (fib (pred n)) (fun n1 => | 746 | _ => TBind (fib (pred n)) (fun n1 => |
722 TBind (fib (pred (pred n))) (fun n2 => | 747 TBind (fib (pred (pred n))) (fun n2 => |
723 Answer (n1 + n2))) | 748 Answer (n1 + n2))) |
724 end. | 749 end. |
725 ]] | 750 ]] |
726 *) | 751 |
727 | 752 Coq complains that the guardedness condition is violated. The two recursive calls are immediate arguments to [TBind], but [TBind] is not a constructor of [thunk]. Rather, it is a defined function. This example shows a very serious limitation of [thunk] for traditional functional programming: it is not, in general, possible to make recursive calls and then make further recursive calls, depending on the first call's result. The [fact] example succeeded because it was already tail recursive, meaning no further computation is needed after a recursive call. |
753 | |
754 %\medskip% | |
755 | |
756 I know no easy fix for this problem of [thunk], but we can define an alternate co-inductive monad that avoids the problem, based on a proposal by Megacz%~\cite{Megacz}%. We ran into trouble because [TBind] was not a constructor of [thunk], so let us define a new type family where %``%#"#bind#"#%''% is a constructor. *) | |
728 | 757 |
729 CoInductive comp (A : Type) : Type := | 758 CoInductive comp (A : Type) : Type := |
730 | Ret : A -> comp A | 759 | Ret : A -> comp A |
731 | Bnd : forall B, comp B -> (B -> comp A) -> comp A. | 760 | Bnd : forall B, comp B -> (B -> comp A) -> comp A. |
761 | |
762 (** This example shows off Coq's support for %\index{recursively non-uniform parameters}\emph{%#<i>#recursively non-uniform parameters#</i>#%}%, as in the case of the parameter [A] declared above, where each constructor's type ends in [comp A], but there is a recursive use of [comp] with a different parameter [B]. Beside that technical wrinkle, we see the simplest possible definition of a monad, via a type whose two constructors are precisely the monad operators. | |
763 | |
764 It is easy to define the semantics of terminating [comp] computations. *) | |
732 | 765 |
733 Inductive exec A : comp A -> A -> Prop := | 766 Inductive exec A : comp A -> A -> Prop := |
734 | ExecRet : forall x, exec (Ret x) x | 767 | ExecRet : forall x, exec (Ret x) x |
735 | ExecBnd : forall B (c : comp B) (f : B -> comp A) x1 x2, exec (A := B) c x1 | 768 | ExecBnd : forall B (c : comp B) (f : B -> comp A) x1 x2, exec (A := B) c x1 |
736 -> exec (f x1) x2 | 769 -> exec (f x1) x2 |
737 -> exec (Bnd c f) x2. | 770 -> exec (Bnd c f) x2. |
738 | 771 |
772 (** We can also prove that [Ret] and [Bnd] form a monad according to a notion of [comp] equality based on [exec], but we omit details here; they are in the book source at this point. *) | |
773 | |
774 (* begin hide *) | |
739 Hint Constructors exec. | 775 Hint Constructors exec. |
740 | 776 |
741 Definition comp_eq A (c1 c2 : comp A) := forall r, exec c1 r <-> exec c2 r. | 777 Definition comp_eq A (c1 c2 : comp A) := forall r, exec c1 r <-> exec c2 r. |
742 | 778 |
743 Ltac inverter := repeat match goal with | 779 Ltac inverter := repeat match goal with |
792 | 828 |
793 Theorem cassociativity : forall A B C (m : comp A) (f : A -> comp B) (g : B -> comp C), | 829 Theorem cassociativity : forall A B C (m : comp A) (f : A -> comp B) (g : B -> comp C), |
794 comp_eq (Bnd (Bnd m f) g) (Bnd m (fun x => Bnd (f x) g)). | 830 comp_eq (Bnd (Bnd m f) g) (Bnd m (fun x => Bnd (f x) g)). |
795 red; crush; eauto. | 831 red; crush; eauto. |
796 Qed. | 832 Qed. |
833 (* end hide *) | |
834 | |
835 (** Not only can we define the Fibonacci function with the new monad, but even our running example of merge sort becomes definable. By shadowing our previous notation for %``%#"#bind,#"#%''%, we can write almost exactly the same code as in our previous [mergeSort'] definition, but with less syntactic clutter. *) | |
836 | |
837 Notation "x <- m1 ; m2" := (Bnd m1 (fun x => m2)). | |
797 | 838 |
798 CoFixpoint mergeSort'' A (le : A -> A -> bool) (ls : list A) : comp (list A) := | 839 CoFixpoint mergeSort'' A (le : A -> A -> bool) (ls : list A) : comp (list A) := |
799 if le_lt_dec 2 (length ls) | 840 if le_lt_dec 2 (length ls) |
800 then let lss := partition ls in | 841 then let lss := partition ls in |
801 Bnd (mergeSort'' le (fst lss)) (fun ls1 => | 842 ls1 <- mergeSort'' le (fst lss); |
802 Bnd (mergeSort'' le (snd lss)) (fun ls2 => | 843 ls2 <- mergeSort'' le (snd lss); |
803 Ret (merge le ls1 ls2))) | 844 Ret (merge le ls1 ls2) |
804 else Ret ls. | 845 else Ret ls. |
846 | |
847 (** To execute this function, we go through the usual exercise of writing a function to catalyze evaluation of co-recursive calls. *) | |
805 | 848 |
806 Definition frob' A (c : comp A) := | 849 Definition frob' A (c : comp A) := |
807 match c with | 850 match c with |
808 | Ret x => Ret x | 851 | Ret x => Ret x |
809 | Bnd _ c' f => Bnd c' f | 852 | Bnd _ c' f => Bnd c' f |
810 end. | 853 end. |
811 | 854 |
812 Lemma frob'_eq : forall A (c : comp A), frob' c = c. | |
813 destruct c; reflexivity. | |
814 Qed. | |
815 | |
816 Hint Rewrite frob'_eq : cpdt. | |
817 | |
818 Lemma exec_frob : forall A (c : comp A) x, | 855 Lemma exec_frob : forall A (c : comp A) x, |
819 exec (frob' c) x | 856 exec (frob' c) x |
820 -> exec c x. | 857 -> exec c x. |
821 crush. | 858 destruct c; crush. |
822 Qed. | 859 Qed. |
860 | |
861 (** Now the same sort of proof script that we applied for testing [thunk]s will get the job done. *) | |
823 | 862 |
824 Lemma test_mergeSort'' : exec (mergeSort'' leb (1 :: 2 :: 36 :: 8 :: 19 :: nil)) | 863 Lemma test_mergeSort'' : exec (mergeSort'' leb (1 :: 2 :: 36 :: 8 :: 19 :: nil)) |
825 (1 :: 2 :: 8 :: 19 :: 36 :: nil). | 864 (1 :: 2 :: 8 :: 19 :: 36 :: nil). |
826 repeat (apply exec_frob; simpl; econstructor). | 865 repeat (apply exec_frob; simpl; econstructor). |
827 Qed. | 866 Qed. |
828 | 867 |
868 (** Have we finally reached the ideal solution for encoding general recursive definitions, with minimal hassle in syntax and proof obligations? Unfortunately, we have not, as [comp] has a serious expressivity weakness. Consider the following definition of a curried addition function: *) | |
869 | |
829 Definition curriedAdd (n : nat) := Ret (fun m : nat => Ret (n + m)). | 870 Definition curriedAdd (n : nat) := Ret (fun m : nat => Ret (n + m)). |
830 | 871 |
831 (** [[ | 872 (** This definition works fine, but we run into trouble when we try to apply it in a trivial way. |
832 Definition testCurriedAdd := | 873 [[ |
833 Bnd (curriedAdd 2) (fun f => f 3). | 874 Definition testCurriedAdd := Bnd (curriedAdd 2) (fun f => f 3). |
834 ]] | 875 ]] |
835 | 876 |
836 << | 877 << |
837 Error: Universe inconsistency. | 878 Error: Universe inconsistency. |
838 >> | 879 >> |
839 *) | 880 |
881 The problem has to do with rules for inductive definitions that we still study in more detail in Chapter 12. Briefly, recall that the type of the constructor [Bnd] quantifies over a type [B]. To make [testCurriedAdd] work, we would need to instantiate [B] as [nat -> comp nat]. However, Coq enforces a %\emph{predicativity restriction}% that (roughly) no quantifier in an inductive or co-inductive type's definition may ever be instantiated with a term that contains the type being defined. Chapter 12 presents the exact mechanism by which this restriction is enforced, but for now our conclusion is that [comp] is fatally flawed as a way of encoding interesting higher-order functional programs that use general recursion. *) | |
840 | 882 |
841 | 883 |
842 (** * Comparing the Options *) | 884 (** * Comparing the Options *) |
843 | 885 |
886 (** We have seen four different approaches to encoding general recursive definitions in Coq. Among them there is no clear champion that dominates the others in every important way. Instead, we close the chapter by comparing the techniques along a number of dimensions. Every technique allows recursive definitions with terminaton arguments that go beyond Coq's built-in termination checking, so we must turn to subtler points to highlight differences. | |
887 | |
888 One useful property is automatic integration with normal Coq programming. That is, we would like the type of a function to be the same, whether or not that function is defined using an interesting recursion pattern. Only the first of the four techniques, well-founded recursion, meets this criterion. It is also the only one of the four to meet the related criterion that evaluation of function calls can take place entirely inside Coq's built-in computation machinery. The monad inspired by domain theory occupies some middle ground in this dimension, since generally standard computation is enough to evaluate a term once a high enough approximation level is provided. | |
889 | |
890 Another useful property is that a function and its termination argument may be developed separately. We may even want to define functions that fail to terminate on some or all inputs. The well-founded recursion technique does not have this property, but the other three do. | |
891 | |
892 One minor plus is the ability to write recursive definitions in natural syntax, rather than with calls to higher-order combinators. This downside of the first two techniques is actually rather easy to get around using Coq's notation mechanism, though we leave the details as an exercise for the reader. | |
893 | |
894 The first two techniques impose proof obligations that are more basic than terminaton arguments, where well-founded recursion requires a proof of extensionality and domain-theoretic recursion requires a proof of continuity. A function may not be defined, and thus may not be computed with, until these obligations are proved. The co-inductive techniques avoid this problem, as recursive definitions may be made without any proof obligations. | |
895 | |
896 We can also consider support for common idioms in functional programming. For instance, the [thunk] monad effectively only supports recursion that is tail recursion, while the others allow arbitrary recursion schemes. | |
897 | |
898 On the other hand, the [comp] monad does not support the effective mixing of higher-order functions and general recursion, while all the other techniques do. For instance, we can finish the failed [curriedAdd] example in the domain-theoretic monad. *) | |
899 | |
844 Definition curriedAdd' (n : nat) := Return (fun m : nat => Return (n + m)). | 900 Definition curriedAdd' (n : nat) := Return (fun m : nat => Return (n + m)). |
845 | 901 |
846 Definition testCurriedAdd := | 902 Definition testCurriedAdd := Bind (curriedAdd' 2) (fun f => f 3). |
847 Bind (curriedAdd' 2) (fun f => f 3). | 903 |
904 (** The same techniques also apply to more interesting higher-order functions like list map, and, as in all four techniques, we can mix regular and general recursion, preferring the former when possible to avoid proof obligations. *) | |
848 | 905 |
849 Fixpoint map A B (f : A -> computation B) (ls : list A) : computation (list B) := | 906 Fixpoint map A B (f : A -> computation B) (ls : list A) : computation (list B) := |
850 match ls with | 907 match ls with |
851 | nil => Return nil | 908 | nil => Return nil |
852 | x :: ls' => Bind (f x) (fun x' => | 909 | x :: ls' => Bind (f x) (fun x' => |
853 Bind (map f ls') (fun ls'' => | 910 Bind (map f ls') (fun ls'' => |
854 Return (x' :: ls''))) | 911 Return (x' :: ls''))) |
855 end. | 912 end. |
856 | 913 |
857 (** remove printing exists *) | 914 (** remove printing exists *) |
858 Theorem test_map : run (map (fun x => Return (S x)) (1 :: 2 :: 3 :: nil)) (2 :: 3 :: 4 :: nil). | 915 Theorem test_map : run (map (fun x => Return (S x)) (1 :: 2 :: 3 :: nil)) |
916 (2 :: 3 :: 4 :: nil). | |
859 exists 1; reflexivity. | 917 exists 1; reflexivity. |
860 Qed. | 918 Qed. |
919 | |
920 (** One further disadvantage of [comp] is that we cannot prove an inversion lemma for executions of [Bind] without appealing to an %\emph{axiom}%, a logical complication that we discuss at more length in Chapter 12. The other three techniques allow proof of all the important theorems within the normal logic of Coq. | |
921 | |
922 Perhaps one theme of our comparison is that one must trade off between on one hand, functional programming expressiveness and compatibility with normal Coq types and computation; and, on the other hand, the level of proof obligations one is willing to handle at function definition time. *) |